diff options
| author | Enrico Tassi | 2015-03-09 11:07:53 +0100 |
|---|---|---|
| committer | Enrico Tassi | 2015-03-09 11:24:38 +0100 |
| commit | fc84c27eac260dffd8f2fb1cb56d599f1e3486d9 (patch) | |
| tree | c16205f1637c80833a4c4598993c29fa0fd8c373 /mathcomp/odd_order | |
Initial commit
Diffstat (limited to 'mathcomp/odd_order')
34 files changed, 40156 insertions, 0 deletions
diff --git a/mathcomp/odd_order/BGappendixAB.v b/mathcomp/odd_order/BGappendixAB.v new file mode 100644 index 0000000..4cbafd0 --- /dev/null +++ b/mathcomp/odd_order/BGappendixAB.v @@ -0,0 +1,508 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div. +Require Import fintype bigop prime finset ssralg fingroup morphism. +Require Import automorphism quotient gfunctor commutator zmodp center pgroup. +Require Import sylow gseries nilpotent abelian maximal. +Require Import matrix mxalgebra mxrepresentation mxabelem. +Require Import BGsection1 BGsection2. + +(******************************************************************************) +(* This file contains the useful material in B & G, appendices A and B, i.e., *) +(* the proof of the p-stability properties and the ZL-Theorem (the Puig *) +(* replacement for the Glaubermann ZJ-theorem). The relevant definitions are *) +(* given in BGsection1. *) +(* Theorem A.4(a) has not been formalised: it is a result on external *) +(* p-stability, which concerns faithful representations of group with a *) +(* trivial p-core on a field of characteristic p. It's the historical concept *) +(* that was studied by Hall and Higman, but it's not used for FT. Note that *) +(* the finite field case can be recovered from A.4(c) with a semi-direct *) +(* product. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Open Local Scope ring_scope. +Import GroupScope GRing.Theory. + +Section AppendixA. + +Implicit Type gT : finGroupType. +Implicit Type p : nat. + +Import MatrixGenField. + +(* This is B & G, Theorem A.4(c) (in Appendix A, not section 16!). We follow *) +(* both B & G and Gorenstein in using the general form of the p-stable *) +(* property. We could simplify the property because the conditions under *) +(* which we prove p-stability are inherited by sections (morphic image in our *) +(* framework), and restrict to the case where P is normal in G. (Clearly the *) +(* 'O_p^'(G) * P <| G premise plays no part in the proof.) *) +(* Theorems A.1-A.3 are essentially inlined in this proof. *) + +Theorem odd_p_stable gT p (G : {group gT}) : odd #|G| -> p.-stable G. +Proof. +move: gT G. +pose p_xp gT (E : {group gT}) x := p.-elt x && (x \in 'C([~: E, [set x]])). +suffices IH gT (E : {group gT}) x y (G := <<[set x; y]>>) : + [&& odd #|G|, p.-group E & G \subset 'N(E)] -> p_xp gT E x && p_xp gT E y -> + p.-group (G / 'C(E)). +- move=> gT G oddG P A pP /andP[/mulGsubP[_ sPG] _] /andP[sANG pA] cRA. + apply/subsetP=> _ /morphimP[x Nx Ax ->]; have NGx := subsetP sANG x Ax. + apply: Baer_Suzuki => [|_ /morphimP[y Ny NGy ->]]; first exact: mem_quotient. + rewrite -morphJ // -!morphim_set1 -?[<<_>>]morphimY ?sub1set ?groupJ //. + set G1 := _ <*> _; rewrite /pgroup -(card_isog (second_isog _)); last first. + by rewrite join_subG !sub1set Nx groupJ. + have{Nx NGx Ny NGy} [[Gx Nx] [Gy Ny]] := (setIP NGx, setIP NGy). + have sG1G: G1 \subset G by rewrite join_subG !sub1set groupJ ?andbT. + have nPG1: G1 \subset 'N(P) by rewrite join_subG !sub1set groupJ ?andbT. + rewrite -setIA setICA (setIidPr sG1G). + rewrite (card_isog (second_isog _)) ?norms_cent //. + apply: IH => //; first by rewrite pP nPG1 (oddSg sG1G). + rewrite /p_xp -{2}(normP Ny) -conjg_set1 -conjsRg centJ memJ_conjg. + rewrite p_eltJ andbb (mem_p_elt pA) // -sub1set centsC (sameP commG1P trivgP). + by rewrite -cRA !commgSS ?sub1set. +move: {2}_.+1 (ltnSn #|E|) => n; elim: n => // n IHn in gT E x y G *. +rewrite ltnS => leEn /and3P[oddG pE nEG] /and3P[/andP[p_x cRx] p_y cRy]. +have [Gx Gy]: x \in G /\ y \in G by apply/andP; rewrite -!sub1set -join_subG. +apply: wlog_neg => p'Gc; apply/pgroupP=> q q_pr qGc; apply/idPn => p'q. +have [Q sylQ] := Sylow_exists q [group of G]. +have [sQG qQ]: Q \subset G /\ q.-group Q by case/and3P: sylQ. +have{qQ p'q} p'Q: p^'.-group Q by apply: sub_in_pnat qQ => q' _ /eqnP->. +have{q q_pr sylQ qGc} ncEQ: ~~ (Q \subset 'C(E)). + apply: contraL qGc => cEQ; rewrite -p'natE // -partn_eq1 //. + have nCQ: Q \subset 'N('C(E)) by exact: subset_trans (normG _). + have sylQc: q.-Sylow(G / 'C(E)) (Q / 'C(E)) by rewrite morphim_pSylow. + by rewrite -(card_Hall sylQc) -trivg_card1 (sameP eqP trivgP) quotient_sub1. +have solE: solvable E := pgroup_sol pE. +have ntE: E :!=: 1 by apply: contra ncEQ; move/eqP->; rewrite cents1. +have{Q ncEQ p'Q sQG} minE_EG: minnormal E (E <*> G). + apply/mingroupP; split=> [|D]; rewrite join_subG ?ntE ?normG //. + case/and3P=> ntD nDE nDG sDE; have nDGi := subsetP nDG. + apply/eqP; rewrite eqEcard sDE leqNgt; apply: contra ncEQ => ltDE. + have nDQ: Q \subset 'N(D) by rewrite (subset_trans sQG). + have cDQ: Q \subset 'C(D). + rewrite -quotient_sub1 ?norms_cent // ?[_ / _]card1_trivg //. + apply: pnat_1 (morphim_pgroup _ p'Q); apply: pgroupS (quotientS _ sQG) _. + apply: IHn (leq_trans ltDE leEn) _ _; first by rewrite oddG (pgroupS sDE). + rewrite /p_xp p_x p_y /=; apply/andP. + by split; [move: cRx | move: cRy]; apply: subsetP; rewrite centS ?commSg. + apply: (stable_factor_cent cDQ) solE; rewrite ?(pnat_coprime pE) //. + apply/and3P; split; rewrite // -quotient_cents2 // centsC. + rewrite -quotient_sub1 ?norms_cent ?quotient_norms ?(subset_trans sQG) //=. + rewrite [(_ / _) / _]card1_trivg //=. + apply: pnat_1 (morphim_pgroup _ (morphim_pgroup _ p'Q)). + apply: pgroupS (quotientS _ (quotientS _ sQG)) _. + have defGq: G / D = <<[set coset D x; coset D y]>>. + by rewrite quotient_gen -1?gen_subG ?quotientU ?quotient_set1 ?nDGi. + rewrite /= defGq IHn ?(leq_trans _ leEn) ?ltn_quotient // -?defGq. + by rewrite quotient_odd // quotient_pgroup // quotient_norms. + rewrite /p_xp -!sub1set !morph_p_elt -?quotient_set1 ?nDGi //=. + by rewrite -!quotientR ?quotient_cents ?sub1set ?nDGi. +have abelE: p.-abelem E. + by rewrite -is_abelem_pgroup //; case: (minnormal_solvable minE_EG _ solE). +have cEE: abelian E by case/and3P: abelE. +have{minE_EG} minE: minnormal E G. + case/mingroupP: minE_EG => _ minE; apply/mingroupP; rewrite ntE. + split=> // D ntD sDE; apply: minE => //; rewrite join_subG cents_norm //. + by rewrite centsC (subset_trans sDE). +have nCG: G \subset 'N('C_G(E)) by rewrite normsI ?normG ?norms_cent. +suffices{p'Gc} pG'c: p.-group (G / 'C_G(E))^`(1). + have [Pc sylPc sGc'Pc]:= Sylow_superset (der_subS _ _) pG'c. + have nsPc: Pc <| G / 'C_G(E) by rewrite sub_der1_normal ?(pHall_sub sylPc). + case/negP: p'Gc; rewrite /pgroup -(card_isog (second_isog _)) ?norms_cent //. + rewrite setIC; apply: pgroupS (pHall_pgroup sylPc) => /=. + rewrite sub_quotient_pre // join_subG !sub1set !(subsetP nCG, inE) //=. + by rewrite !(mem_normal_Hall sylPc) ?mem_quotient ?morph_p_elt ?(subsetP nCG). +have defC := rker_abelem abelE ntE nEG; rewrite /= -/G in defC. +set rG := abelem_repr _ _ _ in defC. +case ncxy: (rG x *m rG y == rG y *m rG x). + have Cxy: [~ x, y] \in 'C_G(E). + rewrite -defC inE groupR //= !repr_mxM ?groupM ?groupV // mul1mx -/rG. + by rewrite (eqP ncxy) -!repr_mxM ?groupM ?groupV // mulKg mulVg repr_mx1. + rewrite [_^`(1)](commG1P _) ?pgroup1 //= quotient_gen -gen_subG //= -/G. + rewrite !gen_subG centsC gen_subG quotient_cents2r ?gen_subG //= -/G. + rewrite /commg_set imset2Ul !imset2_set1l !imsetU !imset_set1. + by rewrite !subUset andbC !sub1set !commgg group1 /= -invg_comm groupV Cxy. +pose Ax : 'M(E) := rG x - 1; pose Ay : 'M(E) := rG y - 1. +have Ax2: Ax *m Ax = 0. + apply/row_matrixP=> i; apply/eqP; rewrite row_mul mulmxBr mulmx1. + rewrite row0 subr_eq0 -(inj_eq (@rVabelem_inj _ _ _ abelE ntE)). + rewrite rVabelemJ // conjgE -(centP cRx) ?mulKg //. + rewrite linearB /= addrC row1 rowE rVabelemD rVabelemN rVabelemJ //=. + by rewrite mem_commg ?set11 ?mem_rVabelem. +have Ay2: Ay *m Ay = 0. + apply/row_matrixP=> i; apply/eqP; rewrite row_mul mulmxBr mulmx1. + rewrite row0 subr_eq0 -(inj_eq (@rVabelem_inj _ _ _ abelE ntE)). + rewrite rVabelemJ // conjgE -(centP cRy) ?mulKg //. + rewrite linearB /= addrC row1 rowE rVabelemD rVabelemN rVabelemJ //=. + by rewrite mem_commg ?set11 ?mem_rVabelem. +pose A := Ax *m Ay + Ay *m Ax. +have cAG: centgmx rG A. + rewrite /centgmx gen_subG subUset !sub1set !inE Gx Gy /=; apply/andP. + rewrite -[rG x](subrK 1%R) -[rG y](subrK 1%R) -/Ax -/Ay. + rewrite 2!(mulmxDl _ 1 A) 2!(mulmxDr A _ 1) !mulmx1 !mul1mx. + rewrite !(inj_eq (addIr A)) ![_ *m A]mulmxDr ![A *m _]mulmxDl. + by rewrite -!mulmxA Ax2 Ay2 !mulmx0 !mulmxA Ax2 Ay2 !mul0mx !addr0 !add0r. +have irrG: mx_irreducible rG by exact/abelem_mx_irrP. +pose rAG := gen_repr irrG cAG; pose inFA := in_gen irrG cAG. +pose valFA := @val_gen _ _ _ _ _ _ irrG cAG. +set dA := gen_dim A in rAG inFA valFA. +rewrite -(rker_abelem abelE ntE nEG) -/rG -(rker_gen irrG cAG) -/rAG. +have dA_gt0: dA > 0 by rewrite (gen_dim_gt0 irrG cAG). +have irrAG: mx_irreducible rAG by exact: gen_mx_irr. +have: dA <= 2. + case Ax0: (Ax == 0). + by rewrite subr_eq0 in Ax0; case/eqP: ncxy; rewrite (eqP Ax0) mulmx1 mul1mx. + case/rowV0Pn: Ax0 => v; case/submxP => u def_v nzv. + pose U := col_mx v (v *m Ay); pose UA := <<inFA _ U>>%MS. + pose rvalFA := @rowval_gen _ _ _ _ _ _ irrG cAG. + have Umod: mxmodule rAG UA. + rewrite /mxmodule gen_subG subUset !sub1set !inE Gx Gy /= andbC. + apply/andP; split; rewrite (eqmxMr _ (genmxE _)) -in_genJ // genmxE. + rewrite submx_in_gen // -[rG y](subrK 1%R) -/Ay mulmxDr mulmx1. + rewrite addmx_sub // mul_col_mx -mulmxA Ay2 mulmx0. + by rewrite -!addsmxE addsmx0 addsmxSr. + rewrite -[rG x](subrK 1%R) -/Ax mulmxDr mulmx1 in_genD mul_col_mx. + rewrite -mulmxA -[Ay *m Ax](addKr (Ax *m Ay)) (mulmxDr v _ A) -mulmxN. + rewrite mulmxA {1 2}def_v -(mulmxA u) Ax2 mulmx0 mul0mx add0r. + pose B := A; rewrite -(mul0mx _ B) -mul_col_mx -[B](mxval_groot irrG cAG). + rewrite {B} -[_ 0 v](in_genK irrG cAG) -val_genZ val_genK. + rewrite addmx_sub ?scalemx_sub ?submx_in_gen //. + by rewrite -!addsmxE adds0mx addsmxSl. + have nzU: UA != 0. + rewrite -mxrank_eq0 genmxE mxrank_eq0; apply/eqP. + move/(canRL ((in_genK _ _) _)); rewrite val_gen0; apply/eqP. + by rewrite -submx0 -addsmxE addsmx_sub submx0 negb_and nzv. + case/mx_irrP: irrAG => _ /(_ UA Umod nzU)/eqnP <-. + by rewrite genmxE rank_leq_row. +rewrite leq_eqVlt ltnS leq_eqVlt ltnNge dA_gt0 orbF orbC; case/pred2P=> def_dA. + rewrite [_^`(1)](commG1P _) ?pgroup1 // quotient_cents2r // gen_subG. + apply/subsetP=> zt; case/imset2P=> z t Gz Gt ->{zt}. + rewrite !inE groupR //= mul1mx; have Gtz := groupM Gt Gz. + rewrite -(inj_eq (can_inj (mulKmx (repr_mx_unit rAG Gtz)))) mulmx1. + rewrite [eq_op]lock -repr_mxM ?groupR ?groupM // -commgC !repr_mxM // -lock. + apply/eqP; move: (rAG z) (rAG t); rewrite /= -/dA def_dA => Az At. + by rewrite [Az]mx11_scalar scalar_mxC. +move: (kquo_repr _) (kquo_mx_faithful rAG) => /=; set K := rker _. +rewrite def_dA => r2G; move/der1_odd_GL2_charf; move/implyP. +rewrite quotient_odd //= -/G; apply: etrans; apply: eq_pgroup => p'. +have [p_pr _ _] := pgroup_pdiv pE ntE. +by rewrite (fmorph_char (gen_rmorphism _ _)) (charf_eq (char_Fp _)). +Qed. + +Section A5. + +Variables (gT : finGroupType) (p : nat) (G P X : {group gT}). + +Hypotheses (oddG : odd #|G|) (solG : solvable G) (pP : p.-group P). +Hypotheses (nsPG : P <| G) (sXG : X \subset G). +Hypotheses (genX : generated_by (p_norm_abelian p P) X). + +Let C := 'C_G(P). +Let defN : 'N_G(P) = G. Proof. by rewrite (setIidPl _) ?normal_norm. Qed. +Let nsCG : C <| G. Proof. by rewrite -defN subcent_normal. Qed. +Let nCG := normal_norm nsCG. +Let nCX := subset_trans sXG nCG. + +(* This is B & G, Theorem A.5.1; it does not depend on the solG assumption. *) +Theorem odd_abelian_gen_stable : X / C \subset 'O_p(G / C). +Proof. +case/exists_eqP: genX => gX defX. +rewrite -defN sub_quotient_pre // -defX gen_subG. +apply/bigcupsP=> A gX_A; have [_ pA nAP cAA] := and4P gX_A. +have{gX_A} sAX: A \subset X by rewrite -defX sub_gen ?bigcup_sup. +rewrite -sub_quotient_pre ?(subset_trans sAX nCX) //=. +rewrite odd_p_stable ?normalM ?pcore_normal //. + by rewrite /psubgroup pA defN (subset_trans sAX sXG). +by apply/commG1P; rewrite (subset_trans _ cAA) // commg_subr. +Qed. + +(* This is B & G, Theorem A.5.2. *) +Theorem odd_abelian_gen_constrained : + 'O_p^'(G) = 1 -> 'C_('O_p(G))(P) \subset P -> X \subset 'O_p(G). +Proof. +set Q := 'O_p(G) => p'G1 sCQ_P. +have sPQ: P \subset Q by rewrite pcore_max. +have defQ: 'O_{p^', p}(G) = Q by rewrite pseries_pop2. +have pQ: p.-group Q by exact: pcore_pgroup. +have sCQ: 'C_G(Q) \subset Q. + by rewrite -{2}defQ solvable_p_constrained //= defQ /pHall pQ indexgg subxx. +have pC: p.-group C. + apply/pgroupP=> q q_pr; case/Cauchy=> // u Cu q_u; apply/idPn=> p'q. + suff cQu: u \in 'C_G(Q). + case/negP: p'q; have{q_u}: q %| #[u] by rewrite q_u. + by apply: pnatP q q_pr => //; apply: mem_p_elt pQ _; exact: (subsetP sCQ). + have [Gu cPu] := setIP Cu; rewrite inE Gu /= -cycle_subG. + rewrite coprime_nil_faithful_cent_stab ?(pgroup_nil pQ) //= -/C -/Q. + - by rewrite cycle_subG; apply: subsetP Gu; rewrite normal_norm ?pcore_normal. + - by rewrite (pnat_coprime pQ) // [#|_|]q_u pnatE. + have sPcQu: P \subset 'C_Q(<[u]>) by rewrite subsetI sPQ centsC cycle_subG. + by apply: subset_trans (subset_trans sCQ_P sPcQu); rewrite setIS // centS. +rewrite -(quotientSGK nCX) ?pcore_max // -pquotient_pcore //. +exact: odd_abelian_gen_stable. +Qed. + +End A5. + +End AppendixA. + +Section AppendixB. + +Local Notation "X --> Y" := (generated_by (norm_abelian X) Y) + (at level 70, no associativity) : group_scope. + +Variable gT : finGroupType. +Implicit Types G H A : {group gT}. +Implicit Types D E : {set gT}. +Implicit Type p : nat. + +Lemma Puig_char G : 'L(G) \char G. +Proof. exact: gFchar. Qed. + +Lemma center_Puig_char G : 'Z('L(G)) \char G. +Proof. exact: char_trans (center_char _) (Puig_char _). Qed. + +(* This is B & G, Lemma B.1(a). *) +Lemma Puig_succS G D E : D \subset E -> 'L_[G](E) \subset 'L_[G](D). +Proof. +move=> sDE; apply: Puig_max (Puig_succ_sub _ _). +exact: norm_abgenS sDE (Puig_gen _ _). +Qed. + +(* This is part of B & G, Lemma B.1(b) (see also BGsection1.Puig1). *) +Lemma Puig_sub_even m n G : m <= n -> 'L_{m.*2}(G) \subset 'L_{n.*2}(G). +Proof. +move/subnKC <-; move: {n}(n - m)%N => n. +by elim: m => [|m IHm] /=; rewrite ?sub1G ?Puig_succS. +Qed. + +(* This is part of B & G, Lemma B.1(b). *) +Lemma Puig_sub_odd m n G : m <= n -> 'L_{n.*2.+1}(G) \subset 'L_{m.*2.+1}(G). +Proof. by move=> le_mn; rewrite Puig_succS ?Puig_sub_even. Qed. + +(* This is part of B & G, Lemma B.1(b). *) +Lemma Puig_sub_even_odd m n G : 'L_{m.*2}(G) \subset 'L_{n.*2.+1}(G). +Proof. +elim: n m => [|n IHn] m; first by rewrite Puig1 Puig_at_sub. +by case: m => [|m]; rewrite ?sub1G ?Puig_succS ?IHn. +Qed. + +(* This is B & G, Lemma B.1(c). *) +Lemma Puig_limit G : + exists m, forall k, m <= k -> + 'L_{k.*2}(G) = 'L_*(G) /\ 'L_{k.*2.+1}(G) = 'L(G). +Proof. +pose L2G m := 'L_{m.*2}(G); pose n := #|G|. +have []: #|L2G n| <= n /\ n <= n by rewrite subset_leq_card ?Puig_at_sub. +elim: {1 2 3}n => [| m IHm leLm1 /ltnW]; first by rewrite leqNgt cardG_gt0. +have [eqLm le_mn|] := eqVneq (L2G m.+1) (L2G m); last first. + rewrite eq_sym eqEcard Puig_sub_even ?leqnSn // -ltnNge => lt_m1_m. + exact: IHm (leq_trans lt_m1_m leLm1). +have{eqLm} eqLm k: m <= k -> 'L_{k.*2}(G) = L2G m. + rewrite leq_eqVlt => /predU1P[-> // |]; elim: k => // k IHk. + by rewrite leq_eqVlt => /predU1P[<- //| ltmk]; rewrite -eqLm !PuigS IHk. +by exists m => k le_mk; rewrite Puig_def PuigS /Puig_inf /= !eqLm. +Qed. + +(* This is B & G, Lemma B.1(d), second part; the first part is covered by *) +(* BGsection1.Puig_inf_sub. *) +Lemma Puig_inf_sub_Puig G : 'L_*(G) \subset 'L(G). +Proof. exact: Puig_sub_even_odd. Qed. + +(* This is B & G, Lemma B.1(e). *) +Lemma abelian_norm_Puig n G A : + n > 0 -> abelian A -> A <| G -> A \subset 'L_{n}(G). +Proof. +case: n => // n _ cAA /andP[sAG nAG]. +rewrite PuigS sub_gen // bigcup_sup // inE sAG /norm_abelian cAA andbT. +exact: subset_trans (Puig_at_sub n G) nAG. +Qed. + +(* This is B & G, Lemma B.1(f), first inclusion. *) +Lemma sub_cent_Puig_at n p G : + n > 0 -> p.-group G -> 'C_G('L_{n}(G)) \subset 'L_{n}(G). +Proof. +move=> n_gt0 pG. +have /ex_maxgroup[M /(max_SCN pG)SCN_M]: exists M, (gval M <| G) && abelian M. + by exists 1%G; rewrite normal1 abelian1. +have{SCN_M} [cMM [nsMG defCM]] := (SCN_abelian SCN_M, SCN_P SCN_M). +have sML: M \subset 'L_{n}(G) by apply: abelian_norm_Puig. +by apply: subset_trans (sML); rewrite -defCM setIS // centS. +Qed. + +(* This is B & G, Lemma B.1(f), second inclusion. *) +Lemma sub_center_cent_Puig_at n G : 'Z(G) \subset 'C_G('L_{n}(G)). +Proof. by rewrite setIS ?centS ?Puig_at_sub. Qed. + +(* This is B & G, Lemma B.1(f), third inclusion (the fourth is trivial). *) +Lemma sub_cent_Puig_inf p G : p.-group G -> 'C_G('L_*(G)) \subset 'L_*(G). +Proof. by apply: sub_cent_Puig_at; rewrite double_gt0. Qed. + +(* This is B & G, Lemma B.1(f), fifth inclusion (the sixth is trivial). *) +Lemma sub_cent_Puig p G : p.-group G -> 'C_G('L(G)) \subset 'L(G). +Proof. exact: sub_cent_Puig_at. Qed. + +(* This is B & G, Lemma B.1(f), final remark (we prove the contrapositive). *) +Lemma trivg_center_Puig_pgroup p G : p.-group G -> 'Z('L(G)) = 1 -> G :=: 1. +Proof. +move=> pG LG1; apply/(trivg_center_pgroup pG)/trivgP. +rewrite -(trivg_center_pgroup (pgroupS (Puig_sub _) pG) LG1). +by apply: subset_trans (sub_cent_Puig pG); apply: sub_center_cent_Puig_at. +Qed. + +(* This is B & G, Lemma B.1(g), second part; the first part is simply the *) +(* definition of 'L(G) in terms of 'L_*(G). *) +Lemma Puig_inf_def G : 'L_*(G) = 'L_[G]('L(G)). +Proof. +have [k defL] := Puig_limit G. +by case: (defL k) => // _ <-; case: (defL k.+1) => [|<- //]; apply: leqnSn. +Qed. + +(* This is B & G, Lemma B.2. *) +Lemma sub_Puig_eq G H : H \subset G -> 'L(G) \subset H -> 'L(H) = 'L(G). +Proof. +move=> sHG sLG_H; apply/setP/subset_eqP/andP. +have sLH_G := subset_trans (Puig_succ_sub _ _) sHG. +have gPuig := norm_abgenS _ (Puig_gen _ _). +have [[kG defLG] [kH defLH]] := (Puig_limit G, Puig_limit H). +have [/defLG[_ {1}<-] /defLH[_ <-]] := (leq_maxl kG kH, leq_maxr kG kH). +split; do [elim: (maxn _ _) => [|k IHk /=]; first by rewrite !Puig1]. + rewrite doubleS !(PuigS _.+1) Puig_max ?gPuig // Puig_max ?gPuig //. + exact: subset_trans (Puig_sub_even_odd _.+1 _ _) sLG_H. +rewrite doubleS Puig_max // -!PuigS Puig_def gPuig //. +by rewrite Puig_inf_def Puig_max ?gPuig ?sLH_G. +Qed. + +Lemma norm_abgen_pgroup p X G : + p.-group G -> X --> G -> generated_by (p_norm_abelian p X) G. +Proof. +move=> pG /exists_eqP[gG defG]. +have:= subxx G; rewrite -{1 3}defG gen_subG /= => /bigcupsP-sGG. +apply/exists_eqP; exists gG; congr <<_>>; apply: eq_bigl => A. +by rewrite andbA andbAC andb_idr // => /sGG/pgroupS->. +Qed. + +Variables (p : nat) (G S : {group gT}). +Hypotheses (oddG : odd #|G|) (solG : solvable G) (sylS : p.-Sylow(G) S). +Hypothesis p'G1 : 'O_p^'(G) = 1. + +Let T := 'O_p(G). +Let nsTG : T <| G := pcore_normal _ _. +Let pT : p.-group T := pcore_pgroup _ _. +Let pS : p.-group S := pHall_pgroup sylS. +Let sSG := pHall_sub sylS. + +(* This is B & G, Lemma B.3. *) +Lemma pcore_Sylow_Puig_sub : 'L_*(S) \subset 'L_*(T) /\ 'L(T) \subset 'L(S). +Proof. +have [[kS defLS] [kT defLT]] := (Puig_limit S, Puig_limit [group of T]). +have [/defLS[<- <-] /defLT[<- <-]] := (leq_maxl kS kT, leq_maxr kS kT). +have sL_ := subset_trans (Puig_succ_sub _ _). +elim: (maxn kS kT) => [|k [_ sL1]]; first by rewrite !Puig1 pcore_sub_Hall. +have{sL1} gL: 'L_{k.*2.+1}(T) --> 'L_{k.*2.+2}(S). + exact: norm_abgenS sL1 (Puig_gen _ _). +have sCT_L: 'C_T('L_{k.*2.+1}(T)) \subset 'L_{k.*2.+1}(T). + exact: sub_cent_Puig_at pT. +have{sCT_L} sLT: 'L_{k.*2.+2}(S) \subset T. + apply: odd_abelian_gen_constrained sCT_L => //. + - exact: pgroupS (Puig_at_sub _ _) pT. + - by apply: char_normal_trans nsTG; apply: gFchar. + - exact: sL_ sSG. + by rewrite norm_abgen_pgroup // (pgroupS _ pS) ?Puig_at_sub. +have sL2: 'L_{k.*2.+2}(S) \subset 'L_{k.*2.+2}(T) by apply: Puig_max. +split; [exact: sL2 | rewrite doubleS; apply: subset_trans (Puig_succS _ sL2) _]. +by rewrite Puig_max -?PuigS ?Puig_gen // sL_ // pcore_sub_Hall. +Qed. + +Let Y := 'Z('L(T)). +Let L := 'L(S). + +(* This is B & G, Theorem B.4(b). *) +Theorem Puig_center_normal : 'Z(L) <| G. +Proof. +have [sLiST sLTS] := pcore_Sylow_Puig_sub. +have sLiLT: 'L_*(T) \subset 'L(T) by exact: Puig_sub_even_odd. +have sZY: 'Z(L) \subset Y. + rewrite subsetI andbC subIset ?centS ?orbT //=. + suffices: 'C_S('L_*(S)) \subset 'L(T). + by apply: subset_trans; rewrite setISS ?Puig_sub ?centS ?Puig_sub_even_odd. + apply: subset_trans (subset_trans sLiST sLiLT). + by apply: sub_cent_Puig_at pS; rewrite double_gt0. +have chY: Y \char G := char_trans (center_Puig_char _) (pcore_char _ _). +have nsCY_G: 'C_G(Y) <| G by rewrite char_normal 1?subcent_char ?char_refl. +have [C defC sCY_C nsCG] := inv_quotientN nsCY_G (pcore_normal p _). +have sLG: L \subset G by rewrite (subset_trans _ (pHall_sub sylS)) ?Puig_sub. +have nsL_nCS: L <| 'N_G(C :&: S). + have sYLiS: Y \subset 'L_*(S). + rewrite abelian_norm_Puig ?double_gt0 ?center_abelian //. + apply: normalS (pHall_sub sylS) (char_normal chY). + by rewrite subIset // (subset_trans sLTS) ?Puig_sub. + have gYL: Y --> L := norm_abgenS sYLiS (Puig_gen _ _). + have sLCS: L \subset C :&: S. + rewrite subsetI Puig_sub andbT. + rewrite -(quotientSGK _ sCY_C) ?(subset_trans sLG) ?normal_norm // -defC. + rewrite odd_abelian_gen_stable ?char_normal ?norm_abgen_pgroup //. + by rewrite (pgroupS _ pT) ?subIset // Puig_sub. + by rewrite (pgroupS _ pS) ?Puig_sub. + rewrite -[L](sub_Puig_eq _ sLCS) ?subsetIr //. + by rewrite (char_normal_trans (Puig_char _)) ?normalSG // subIset // sSG orbT. +have sylCS: p.-Sylow(C) (C :&: S) := Sylow_setI_normal nsCG sylS. +have{defC} defC: 'C_G(Y) * (C :&: S) = C. + apply/eqP; rewrite eqEsubset mulG_subG sCY_C subsetIl /=. + have nCY_C: C \subset 'N('C_G(Y)). + exact: subset_trans (normal_sub nsCG) (normal_norm nsCY_G). + rewrite -quotientSK // -defC /= -pseries1. + rewrite -(pseries_catr_id [:: p : nat_pred]) (pseries_rcons_id [::]) /=. + rewrite pseries1 /= pseries1 defC pcore_sub_Hall // morphim_pHall //. + by rewrite subIset ?nCY_C. +have defG: 'C_G(Y) * 'N_G(C :&: S) = G. + have sCS_N: C :&: S \subset 'N_G(C :&: S). + by rewrite subsetI normG subIset // sSG orbT. + by rewrite -(mulSGid sCS_N) mulgA defC (Frattini_arg _ sylCS). +have nsZ_N: 'Z(L) <| 'N_G(C :&: S) := char_normal_trans (center_char _) nsL_nCS. +rewrite /normal subIset ?sLG //= -{1}defG mulG_subG /=. +rewrite cents_norm ?normal_norm // centsC. +by rewrite (subset_trans sZY) // centsC subsetIr. +Qed. + +End AppendixB. + +Section Puig_factorization. + +Variables (gT : finGroupType) (p : nat) (G S : {group gT}). +Hypotheses (oddG : odd #|G|) (solG : solvable G) (sylS : p.-Sylow(G) S). + +(* This is B & G, Theorem B.4(a). *) +Theorem Puig_factorization : 'O_p^'(G) * 'N_G('Z('L(S))) = G. +Proof. +set D := 'O_p^'(G); set Z := 'Z(_); have [sSG pS _] := and3P sylS. +have sSN: S \subset 'N(D) by rewrite (subset_trans sSG) ?gFnorm. +have p'D: p^'.-group D := pcore_pgroup _ _. +have tiSD: S :&: D = 1 := coprime_TIg (pnat_coprime pS p'D). +have def_Zq: Z / D = 'Z('L(S / D)). + rewrite !quotientE -(setIid S) -(morphim_restrm sSN); set f := restrm _ _. + have injf: 'injm f by rewrite ker_restrm ker_coset tiSD. + rewrite -!(injmF _ injf) ?Puig_sub //= morphim_restrm. + by rewrite (setIidPr _) // subIset ?Puig_sub. +have{def_Zq} nZq: Z / D <| G / D. + have sylSq: p.-Sylow(G / D) (S / D) by exact: morphim_pHall. + rewrite def_Zq (Puig_center_normal _ _ sylSq) ?quotient_odd ?quotient_sol //. + exact: trivg_pcore_quotient. +have sZS: Z \subset S by rewrite subIset ?Puig_sub. +have sZN: Z \subset 'N_G(Z) by rewrite subsetI normG (subset_trans sZS). +have nDZ: Z \subset 'N(D) by rewrite (subset_trans sZS). +rewrite -(mulSGid sZN) mulgA -(norm_joinEr nDZ) (@Frattini_arg p) //= -/D -/Z. + rewrite -cosetpre_normal quotientK ?quotientGK ?pcore_normal // in nZq. + by rewrite norm_joinEr. +rewrite /pHall -divgS joing_subr ?(pgroupS sZS) /= ?norm_joinEr //= -/Z. +by rewrite TI_cardMg ?mulnK //; apply/trivgP; rewrite /= setIC -tiSD setSI. +Qed. + +End Puig_factorization. + + + + + diff --git a/mathcomp/odd_order/BGappendixC.v b/mathcomp/odd_order/BGappendixC.v new file mode 100644 index 0000000..31bccfb --- /dev/null +++ b/mathcomp/odd_order/BGappendixC.v @@ -0,0 +1,749 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype choice ssrnat seq div fintype. +Require Import tuple finfun bigop ssralg finset prime binomial poly polydiv. +Require Import fingroup morphism quotient automorphism action finalg zmodp. +Require Import gproduct cyclic commutator pgroup abelian frobenius BGsection1. +Require Import matrix mxalgebra mxabelem vector falgebra fieldext galois. +Require Import finfield ssrnum algC classfun character integral_char inertia. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory FinRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Section AppendixC. + +Variables (gT : finGroupType) (p q : nat) (H P P0 U Q : {group gT}). +Let nU := ((p ^ q).-1 %/ p.-1)%N. + +(* External statement of the finite field assumption. *) +CoInductive finFieldImage : Prop := + FinFieldImage (F : finFieldType) (sigma : {morphism P >-> F}) of + isom P [set: F] sigma & sigma @*^-1 <[1 : F]> = P0 + & exists2 sigmaU : {morphism U >-> {unit F}}, + 'injm sigmaU & {in P & U, morph_act 'J 'U sigma sigmaU}. + +(* These correspond to hypothesis (A) of B & G, Appendix C, Theorem C. *) +Hypotheses (pr_p : prime p) (pr_q : prime q) (coUp1 : coprime nU p.-1). +Hypotheses (defH : P ><| U = H) (fieldH : finFieldImage). +Hypotheses (oP : #|P| = (p ^ q)%N) (oU : #|U| = nU). + +(* These correspond to hypothesis (B) of B & G, Appendix C, Theorem C. *) +Hypotheses (abelQ : q.-abelem Q) (nQP0 : P0 \subset 'N(Q)). +Hypothesis nU_P0Q : exists2 y, y \in Q & P0 :^ y \subset 'N(U). + +Section ExpandHypotheses. + +(* Negation of the goal of B & G, Appendix C, Theorem C. *) +Hypothesis ltqp : (q < p)%N. + +(* From the fieldH assumption. *) +Variables (fT : finFieldType) (charFp : p \in [char fT]). +Local Notation F := (PrimeCharType charFp). +Local Notation galF := [splittingFieldType 'F_p of F]. +Let Fpq : {vspace F} := fullv. +Let Fp : {vspace F} := 1%VS. + +Hypothesis oF : #|F| = (p ^ q)%N. +Let oF_p : #|'F_p| = p. Proof. exact: card_Fp. Qed. +Let oFp : #|Fp| = p. Proof. by rewrite card_vspace1. Qed. +Let oFpq : #|Fpq| = (p ^ q)%N. Proof. by rewrite card_vspacef. Qed. +Let dimFpq : \dim Fpq = q. Proof. by rewrite primeChar_dimf oF pfactorK. Qed. + +Variables (sigma : {morphism P >-> F}) (sigmaU : {morphism U >-> {unit F}}). +Hypotheses (inj_sigma : 'injm sigma) (inj_sigmaU : 'injm sigmaU). +Hypothesis im_sigma : sigma @* P = [set: F]. +Variable s : gT. +Hypotheses (sP0P : P0 \subset P) (sigma_s : sigma s = 1) (defP0 : <[s]> = P0). + +Let psi u : F := val (sigmaU u). +Let inj_psi : {in U &, injective psi}. +Proof. by move=> u v Uu Uv /val_inj/(injmP inj_sigmaU)->. Qed. + +Hypothesis sigmaJ : {in P & U, forall x u, sigma (x ^ u) = sigma x * psi u}. + +Let Ps : s \in P. Proof. by rewrite -cycle_subG defP0. Qed. +Let P0s : s \in P0. Proof. by rewrite -defP0 cycle_id. Qed. + +Let nz_psi u : psi u != 0. Proof. by rewrite -unitfE (valP (sigmaU u)). Qed. + +Let sigma1 : sigma 1%g = 0. Proof. exact: morph1. Qed. +Let sigmaM : {in P &, {morph sigma : x1 x2 / (x1 * x2)%g >-> x1 + x2}}. +Proof. exact: morphM. Qed. +Let sigmaV : {in P, {morph sigma : x / x^-1%g >-> - x}}. +Proof. exact: morphV. Qed. +Let sigmaX n : {in P, {morph sigma : x / (x ^+ n)%g >-> x *+ n}}. +Proof. exact: morphX. Qed. + +Let psi1 : psi 1%g = 1. Proof. by rewrite /psi morph1. Qed. +Let psiM : {in U &, {morph psi : u1 u2 / (u1 * u2)%g >-> u1 * u2}}. +Proof. by move=> u1 u2 Uu1 Uu2; rewrite /psi morphM. Qed. +Let psiV : {in U, {morph psi : u / u^-1%g >-> u^-1}}. +Proof. by move=> u Uu; rewrite /psi morphV. Qed. +Let psiX n : {in U, {morph psi : u / (u ^+ n)%g >-> u ^+ n}}. +Proof. by move=> u Uu; rewrite /psi morphX // val_unitX. Qed. + +Let sigmaE := (sigma1, sigma_s, mulr1, mul1r, + (sigmaJ, sigmaX, sigmaM, sigmaV), (psi1, psiX, psiM, psiV)). + +Let psiE u : u \in U -> psi u = sigma (s ^ u). +Proof. by move=> Uu; rewrite !sigmaE. Qed. + +Let nPU : U \subset 'N(P). Proof. by have [] := sdprodP defH. Qed. +Let memJ_P : {in P & U, forall x u, x ^ u \in P}. +Proof. by move=> x u Px Uu; rewrite /= memJ_norm ?(subsetP nPU). Qed. +Let in_PU := (memJ_P, in_group). + +Let sigmaP0 : sigma @* P0 =i Fp. +Proof. +rewrite -defP0 morphim_cycle // sigma_s => x. +apply/cycleP/vlineP=> [] [n ->]; first by exists n%:R; rewrite scaler_nat. +by exists (val n); rewrite -{1}[n]natr_Zp -in_algE rmorph_nat zmodXgE. +Qed. + +Let nt_s : s != 1%g. +Proof. by rewrite -(morph_injm_eq1 inj_sigma) // sigmaE oner_eq0. Qed. + +Let p_gt0 : (0 < p)%N. Proof. exact: prime_gt0. Qed. +Let q_gt0 : (0 < q)%N. Proof. exact: prime_gt0. Qed. +Let p1_gt0 : (0 < p.-1)%N. Proof. by rewrite -subn1 subn_gt0 prime_gt1. Qed. + +(* This is B & G, Appendix C, Remark I. *) +Let not_dvd_q_p1 : ~~ (q %| p.-1)%N. +Proof. +rewrite -prime_coprime // -[q]card_ord -sum1_card -coprime_modl -modn_summ. +have:= coUp1; rewrite /nU predn_exp mulKn //= -coprime_modl -modn_summ. +congr (coprime (_ %% _) _); apply: eq_bigr => i _. +by rewrite -{1}[p](subnK p_gt0) subn1 -modnXm modnDl modnXm exp1n. +Qed. + +(* This is the first assertion of B & G, Appendix C, Remark V. *) +Let odd_p : odd p. +Proof. +by apply: contraLR ltqp => /prime_oddPn-> //; rewrite -leqNgt prime_gt1. +Qed. + +(* This is the second assertion of B & G, Appendix C, Remark V. *) +Let odd_q : odd q. +Proof. +apply: contraR not_dvd_q_p1 => /prime_oddPn-> //. +by rewrite -subn1 dvdn2 odd_sub ?odd_p. +Qed. + +Let qgt2 : (2 < q)%N. Proof. by rewrite odd_prime_gt2. Qed. +Let pgt4 : (4 < p)%N. Proof. by rewrite odd_geq ?(leq_ltn_trans qgt2). Qed. +Let qgt1 : (1 < q)%N. Proof. exact: ltnW. Qed. + +Local Notation Nm := (galNorm Fp Fpq). +Local Notation uval := (@FinRing.uval _). + +Let cycFU (FU : {group {unit F}}) : cyclic FU. +Proof. exact: field_unit_group_cyclic. Qed. +Let cUU : abelian U. +Proof. by rewrite cyclic_abelian // -(injm_cyclic inj_sigmaU) ?cycFU. Qed. + +(* This is B & G, Appendix C, Remark VII. *) +Let im_psi (x : F) : (x \in psi @: U) = (Nm x == 1). +Proof. +have /cyclicP[u0 defFU]: cyclic [set: {unit F}] by exact: cycFU. +have o_u0: #[u0] = (p ^ q).-1 by rewrite orderE -defFU card_finField_unit oF. +have ->: psi @: U = uval @: (sigmaU @* U) by rewrite morphimEdom -imset_comp. +have /set1P[->]: (sigmaU @* U)%G \in [set <[u0 ^+ (#[u0] %/ nU)]>%G]. + rewrite -cycle_sub_group ?inE; last first. + by rewrite o_u0 -(divnK (dvdn_pred_predX p q)) dvdn_mulr. + by rewrite -defFU subsetT card_injm //= oU. +rewrite divnA ?dvdn_pred_predX // -o_u0 mulKn //. +have [/= alpha alpha_gen Dalpha] := finField_galois_generator (subvf Fp). +have{Dalpha} Dalpha x1: x1 != 0 -> x1 / alpha x1 = x1^-1 ^+ p.-1. + move=> nz_x1; rewrite -[_ ^+ _](mulVKf nz_x1) -exprS Dalpha ?memvf // exprVn. + by rewrite dimv1 oF_p prednK ?prime_gt0. +apply/idP/(Hilbert's_theorem_90 alpha_gen (memvf _)) => [|[u [_ nz_u] ->]]. + case/imsetP=> /= _ /cycleP[n ->] ->; rewrite expgAC; set u := (u0 ^+ n)%g. + have nz_u: (val u)^-1 != 0 by rewrite -unitfE unitrV (valP u). + by exists (val u)^-1; rewrite ?memvf ?Dalpha //= invrK val_unitX. +have /cycleP[n Du]: (insubd u0 u)^-1%g \in <[u0]> by rewrite -defFU inE. +have{Du} Du: u^-1 = val (u0 ^+ n)%g by rewrite -Du /= insubdK ?unitfE. +by rewrite Dalpha // Du -val_unitX mem_imset // expgAC mem_cycle. +Qed. + +(* This is B & G, Appendix C, Remark VIII. *) +Let defFU : sigmaU @* U \x [set u | uval u \in Fp] = [set: {unit F}]. +Proof. +have fP v: in_alg F (uval v) \is a GRing.unit by rewrite rmorph_unit ?(valP v). +pose f (v : {unit 'F_p}) := FinRing.unit F (fP v). +have fM: {in setT &, {morph f: v1 v2 / (v1 * v2)%g}}. + by move=> v1 v2 _ _; apply: val_inj; rewrite /= -in_algE rmorphM. +pose galFpU := Morphism fM @* [set: {unit 'F_p}]. +have ->: [set u | uval u \in Fp] = galFpU. + apply/setP=> u; rewrite inE /galFpU morphimEdom. + apply/idP/imsetP=> [|[v _ ->]]; last by rewrite /= rpredZ // memv_line. + case/vlineP=> v Du; have nz_v: v != 0. + by apply: contraTneq (valP u) => v0; rewrite unitfE /= Du v0 scale0r eqxx. + exists (insubd (1%g : {unit 'F_p}) v); rewrite ?inE //. + by apply: val_inj; rewrite /= insubdK ?unitfE. +have oFpU: #|galFpU| = p.-1. + rewrite card_injm ?card_finField_unit ?oF_p //. + by apply/injmP=> v1 v2 _ _ []/(fmorph_inj [rmorphism of in_alg F])/val_inj. +have oUU: #|sigmaU @* U| = nU by rewrite card_injm. +rewrite dprodE ?coprime_TIg ?oUU ?oFpU //; last first. + by rewrite (sub_abelian_cent2 (cyclic_abelian (cycFU [set: _]))) ?subsetT. +apply/eqP; rewrite eqEcard subsetT coprime_cardMg oUU oFpU //=. +by rewrite card_finField_unit oF divnK ?dvdn_pred_predX. +Qed. + +(* This is B & G, Appendix C, Remark IX. *) +Let frobH : [Frobenius H = P ><| U]. +Proof. +apply/Frobenius_semiregularP=> // [||u /setD1P[ntu Uu]]. +- by rewrite -(morphim_injm_eq1 inj_sigma) // im_sigma finRing_nontrivial. +- rewrite -cardG_gt1 oU ltn_divRL ?dvdn_pred_predX // mul1n -!subn1. + by rewrite ltn_sub2r ?(ltn_exp2l 0) ?(ltn_exp2l 1) ?prime_gt1. +apply/trivgP/subsetP=> x /setIP[Px /cent1P/commgP]. +rewrite inE -!(morph_injm_eq1 inj_sigma) ?(sigmaE, in_PU) //. +rewrite -mulrN1 addrC -mulrDr mulf_eq0 subr_eq0 => /orP[] // /idPn[]. +by rewrite (inj_eq val_inj (sigmaU u) 1%g) morph_injm_eq1. +Qed. + +(* From the abelQ assumption of Peterfalvi, Theorem (14.2) to the assumptions *) +(* of part (B) of the assumptions of Theorem C. *) +Let p'q : q != p. Proof. by rewrite ltn_eqF. Qed. +Let cQQ : abelian Q. Proof. exact: abelem_abelian abelQ. Qed. +Let p'Q : p^'.-group Q. Proof. exact: pi_pgroup (abelem_pgroup abelQ) _. Qed. + +Let pP : p.-group P. Proof. by rewrite /pgroup oP pnat_exp ?pnat_id. Qed. +Let coQP : coprime #|Q| #|P|. Proof. exact: p'nat_coprime p'Q pP. Qed. +Let sQP0Q : [~: Q, P0] \subset Q. Proof. by rewrite commg_subl. Qed. + +(* This is B & G, Appendix C, Remark X. *) +Let defQ : 'C_Q(P0) \x [~: Q, P0] = Q. +Proof. by rewrite dprodC coprime_abelian_cent_dprod // (coprimegS sP0P). Qed. + +(* This is B & G, Appendix C, Remark XI. *) +Let nU_P0QP0 : exists2 y, y \in [~: Q, P0] & P0 :^ y \subset 'N(U). +Proof. +have [_ /(mem_dprod defQ)[z [y [/setIP[_ cP0z] QP0y -> _]]]] := nU_P0Q. +by rewrite conjsgM (normsP (cent_sub P0)) //; exists y. +Qed. + +Let E := [set x : galF | Nm x == 1 & Nm (2%:R - x) == 1]. + +Let E_1 : 1 \in E. +Proof. by rewrite !inE -addrA subrr addr0 galNorm1 eqxx. Qed. + +(* This is B & G, Appendix C, Lemma C.1. *) +Let Einv_gt1_le_pq : E = [set x^-1 | x in E] -> (1 < #|E|)%N -> (p <= q)%N. +Proof. +rewrite (cardsD1 1) E_1 ltnS card_gt0 => Einv /set0Pn[/= a /setD1P[not_a1 Ea]]. +pose tau (x : F) := (2%:R - x)^-1. +have Etau x: x \in E -> tau x \in E. + rewrite inE => Ex; rewrite Einv (mem_imset (fun y => y^-1)) //. + by rewrite inE andbC opprD addNKr opprK. +pose Pa := \prod_(beta in 'Gal(Fpq / Fp)) (beta (1 - a) *: 'X + 1). +have galPoly_roots: all (root (Pa - 1)) (enum Fp). + apply/allP=> x; rewrite mem_enum => /vlineP[b ->]. + rewrite rootE !hornerE horner_prod subr_eq0 /=; apply/eqP. + pose h k := (1 - a) *+ k + 1; transitivity (Nm (h b)). + apply: eq_bigr => beta _; rewrite !(rmorphB, rmorphD, rmorphMn) rmorph1 /=. + by rewrite -mulr_natr -scaler_nat natr_Zp hornerD hornerZ hornerX hornerC. + elim: (b : nat) => [|k IHk]; first by rewrite /h add0r galNorm1. + suffices{IHk}: h k / h k.+1 \in E. + rewrite inE -invr_eq1 => /andP[/eqP <- _]. + by rewrite galNormM galNormV /= [galNorm _ _ (h k)]IHk mul1r invrK. + elim: k => [|k IHk]; first by rewrite /h add0r mul1r addrAC Etau. + have nz_hk1: h k.+1 != 0. + apply: contraTneq IHk => ->; rewrite invr0 mulr0. + by rewrite inE galNorm0 eq_sym oner_eq0. + congr (_ \in E): (Etau _ IHk); apply: canLR (@invrK _) _; rewrite invfM invrK. + apply: canRL (mulKf nz_hk1) _; rewrite mulrC mulrBl divfK // mulrDl mul1r. + by rewrite {2}/h mulrS -2!addrA addrK addrAC -mulrSr. +have sizePa: size Pa = q.+1. + have sizePaX (beta : {rmorphism F -> F}) : size (beta (1 - a) *: 'X + 1) = 2. + rewrite -mul_polyC size_MXaddC oner_eq0 andbF size_polyC fmorph_eq0. + by rewrite subr_eq0 eq_sym (negbTE not_a1). + rewrite size_prod => [|i _]; last by rewrite -size_poly_eq0 sizePaX. + rewrite (eq_bigr (fun _ => 2)) => [|beta _]; last by rewrite sizePaX. + rewrite sum_nat_const muln2 -addnn -addSn addnK. + by rewrite -galois_dim ?finField_galois ?subvf // dimv1 divn1 dimFpq. +have sizePa1: size (Pa - 1) = q.+1. + by rewrite size_addl // size_opp size_poly1 sizePa. +have nz_Pa1 : Pa - 1 != 0 by rewrite -size_poly_eq0 sizePa1. +by rewrite -ltnS -oFp -sizePa1 cardE max_poly_roots ?enum_uniq. +Qed. + +(* This is B & G, Appendix C, Lemma C.2. *) +Let E_gt1 : (1 < #|E|)%N. +Proof. +have [q_gt4 | q_le4] := ltnP 4 q. + pose inK x := enum_rank_in (classes1 H) (x ^: H). + have inK_E x: x \in H -> enum_val (inK x) = x ^: H. + by move=> Hx; rewrite enum_rankK_in ?mem_classes. + pose j := inK s; pose k := inK (s ^+ 2)%g; pose e := gring_classM_coef j j k. + have cPP: abelian P by rewrite -(injm_abelian inj_sigma) ?zmod_abelian. + have Hs: s \in H by rewrite -(sdprodW defH) -[s]mulg1 mem_mulg. + have DsH n: (s ^+ n) ^: H = (s ^+ n) ^: U. + rewrite -(sdprodW defH) classM (abelian_classP _ cPP) ?groupX //. + by rewrite class_support_set1l. + have injJU: {in U &, injective (conjg s)}. + by move=> u v Uu Uv eq_s_uv; apply/inj_psi; rewrite ?psiE ?eq_s_uv. + have ->: #|E| = e. + rewrite /e /gring_classM_coef !inK_E ?groupX //. + transitivity #|[set u in U | s^-1 ^ u * s ^+ 2 \in s ^: U]%g|. + rewrite -(card_in_imset (sub_in2 _ inj_psi)) => [|u /setIdP[] //]. + apply: eq_card => x; rewrite inE -!im_psi. + apply/andP/imsetP=> [[/imsetP[u Uu ->] /imsetP[v Uv Dv]]{x} | ]. + exists u; rewrite // inE Uu /=; apply/imsetP; exists v => //. + by apply: (injmP inj_sigma); rewrite ?(sigmaE, in_PU) // mulN1r addrC. + case=> u /setIdP[Uu /imsetP[v Uv /(congr1 sigma)]]. + rewrite ?(sigmaE, in_PU) // mulN1r addrC => Dv ->. + by rewrite Dv !mem_imset. + rewrite DsH (DsH 1%N) expg1; have [w Uw ->] := repr_class U (s ^+ 2). + pose f u := (s ^ (u * w), (s^-1 ^ u * s ^+ 2) ^ w). + rewrite -(@card_in_imset _ _ f) => [|u v]; last first. + by move=> /setIdP[Uu _] /setIdP[Uv _] [/injJU/mulIg-> //]; apply: groupM. + apply: eq_card => [[x1 x2]]; rewrite inE -andbA. + apply/imsetP/and3P=> [[u /setIdP[Uu sUs2u'] [-> ->]{x1 x2}] | []]. + rewrite /= conjgM -(rcoset_id Uw) class_rcoset !memJ_conjg mem_orbit //. + by rewrite sUs2u' -conjMg conjVg mulKVg. + case/imsetP=> u Uu /= -> sUx2 /eqP/(canRL (mulKg _)) Dx2. + exists (u * w^-1)%g; last first. + by rewrite /f /= conjMg -conjgM mulgKV conjVg -Dx2. + rewrite inE !in_PU // Uw -(memJ_conjg _ w) -class_rcoset rcoset_id //. + by rewrite conjMg -conjgM mulgKV conjVg -Dx2. + pose chi_s2 i := 'chi[H]_i s ^+ 2 * ('chi_i (s ^+ 2)%g)^* / 'chi_i 1%g. + have De: e%:R = #|U|%:R / #|P|%:R * (\sum_i chi_s2 i). + have Ks: s \in enum_val j by rewrite inK_E ?class_refl. + have Ks2: (s ^+ 2)%g \in enum_val k by rewrite inK_E ?groupX ?class_refl. + rewrite (gring_classM_coef_sum_eq Ks Ks Ks2) inK_E //; congr (_ * _). + have ->: #|s ^: H| = #|U| by rewrite (DsH 1%N) (card_in_imset injJU). + by rewrite -(sdprod_card defH) mulnC !natrM invfM mulrA mulfK ?neq0CG. + pose linH := [pred i | P \subset cfker 'chi[H]_i]. + have nsPH: P <| H by have [] := sdprod_context defH. + have sum_linH: \sum_(i in linH) chi_s2 i = #|U|%:R. + have isoU: U \isog H / P := sdprod_isog defH. + have abHbar: abelian (H / P) by rewrite -(isog_abelian isoU). + rewrite (card_isog isoU) -(card_Iirr_abelian abHbar) -sumr_const. + rewrite (reindex _ (mod_Iirr_bij nsPH)) /chi_s2 /=. + apply: eq_big => [i | i _]; rewrite ?mod_IirrE ?cfker_mod //. + have lin_i: ('chi_i %% P)%CF \is a linear_char. + exact/cfMod_lin_char/char_abelianP. + rewrite lin_char1 // divr1 -lin_charX // -normCK. + by rewrite normC_lin_char ?groupX ?expr1n. + have degU i: i \notin linH -> 'chi_i 1%g = #|U|%:R. + case/(Frobenius_Ind_irrP (FrobeniusWker frobH)) => {i}i _ ->. + rewrite cfInd1 ?normal_sub // -(index_sdprod defH) lin_char1 ?mulr1 //. + exact/char_abelianP. + have ub_linH' m (s_m := (s ^+ m)%g): + (0 < m < 5)%N -> \sum_(i in predC linH) `|'chi_i s_m| ^+ 2 <= #|P|%:R. + - case/andP=> m_gt0 m_lt5; have{m_gt0 m_lt5} P1sm: s_m \in P^#. + rewrite !inE groupX // -order_dvdn -(order_injm inj_sigma) // sigmaE. + by rewrite andbT order_primeChar ?oner_neq0 ?gtnNdvd ?(leq_trans m_lt5). + have ->: #|P| = (#|P| * (s_m \in s_m ^: H))%N by rewrite class_refl ?muln1. + have{P1sm} /eqP <-: 'C_H[s ^+ m] == P. + rewrite eqEsubset (Frobenius_cent1_ker frobH) // subsetI normal_sub //=. + by rewrite sub_cent1 groupX // (subsetP cPP). + rewrite mulrnA -second_orthogonality_relation ?groupX // big_mkcond. + by apply: ler_sum => i _; rewrite normCK; case: ifP; rewrite ?mul_conjC_ge0. + have sqrtP_gt0: 0 < sqrtC #|P|%:R by rewrite sqrtC_gt0 ?gt0CG. + have{De ub_linH'}: `|(#|P| * e)%:R - #|U|%:R ^+ 2| <= #|P|%:R * sqrtC #|P|%:R. + rewrite natrM De mulrCA mulrA divfK ?neq0CG // (bigID linH) /= sum_linH. + rewrite mulrDr addrC addKr mulrC mulr_suml /chi_s2. + rewrite (ler_trans (ler_norm_sum _ _ _)) // -ler_pdivr_mulr // mulr_suml. + apply: ler_trans (ub_linH' 1%N isT); apply: ler_sum => i linH'i. + rewrite ler_pdivr_mulr // degU ?divfK ?neq0CG //. + rewrite normrM -normrX norm_conjC ler_wpmul2l ?normr_ge0 //. + rewrite -ler_sqr ?qualifE ?normr_ge0 ?(@ltrW _ 0) // sqrtCK. + apply: ler_trans (ub_linH' 2 isT); rewrite (bigD1 i) ?ler_paddr //=. + by apply: sumr_ge0 => i1 _; rewrite exprn_ge0 ?normr_ge0. + rewrite natrM real_ler_distl ?rpredB ?rpredM ?rpred_nat // => /andP[lb_Pe _]. + rewrite -ltC_nat -(ltr_pmul2l (gt0CG P)) {lb_Pe}(ltr_le_trans _ lb_Pe) //. + rewrite ltr_subr_addl (@ler_lt_trans _ ((p ^ q.-1)%:R ^+ 2)) //; last first. + rewrite -!natrX ltC_nat ltn_sqr oU ltn_divRL ?dvdn_pred_predX //. + rewrite -(subnKC qgt1) /= -!subn1 mulnBr muln1 -expnSr. + by rewrite ltn_sub2l ?(ltn_exp2l 0) // prime_gt1. + rewrite -mulrDr -natrX -expnM muln2 -subn1 doubleB -addnn -addnBA // subn2. + rewrite expnD natrM -oP ler_wpmul2l ?ler0n //. + apply: ler_trans (_ : 2%:R * sqrtC #|P|%:R <= _). + rewrite mulrDl mul1r ler_add2l -(@expr_ge1 _ 2) ?(ltrW sqrtP_gt0) // sqrtCK. + by rewrite oP natrX expr_ge1 ?ler0n ?ler1n. + rewrite -ler_sqr ?rpredM ?rpred_nat ?qualifE ?(ltrW sqrtP_gt0) //. + rewrite exprMn sqrtCK -!natrX -natrM leC_nat -expnM muln2 oP. + rewrite -(subnKC q_gt4) doubleS (expnS p _.*2.+1) -(subnKC pgt4) leq_mul //. + by rewrite ?leq_exp2l // !doubleS !ltnS -addnn leq_addl. +have q3: q = 3 by apply/eqP; rewrite eqn_leq qgt2 andbT -ltnS -(odd_ltn 5). +rewrite (cardsD1 1) E_1 ltnS card_gt0; apply/set0Pn => /=. +pose f (c : 'F_p) : {poly 'F_p} := 'X * ('X - 2%:R%:P) * ('X - c%:P) + ('X - 1). +have fc0 c: (f c).[0] = -1 by rewrite !hornerE. +have fc2 c: (f c).[2%:R] = 1 by rewrite !(subrr, hornerE) /= addrK. +have /existsP[c nz_fc]: [exists c, ~~ [exists d, root (f c) d]]. + have nz_f_0 c: ~~ root (f c) 0 by rewrite /root fc0 oppr_eq0. + rewrite -negb_forall; apply/negP=> /'forall_existsP/fin_all_exists[/= rf rfP]. + suffices inj_rf: injective rf. + by have /negP[] := nz_f_0 (invF inj_rf 0); rewrite -{2}[0](f_invF inj_rf). + move=> a b eq_rf_ab; apply/oppr_inj/(addrI (rf a)). + have: (f a).[rf a] = (f b).[rf a] by rewrite {2}eq_rf_ab !(rootP _). + rewrite !(hornerXsubC, hornerD, hornerM) hornerX => /addIr/mulfI-> //. + rewrite mulf_neq0 ?subr_eq0 1?(contraTneq _ (rfP a)) // => -> //. + by rewrite /root fc2. +have{nz_fc} /= nz_fc: ~~ root (f c) _ by apply/forallP; rewrite -negb_exists. +have sz_fc_lhs: size ('X * ('X - 2%:R%:P) * ('X - c%:P)) = 4. + by rewrite !(size_mul, =^~ size_poly_eq0) ?size_polyX ?size_XsubC. +have sz_fc: size (f c) = 4 by rewrite size_addl ?size_XsubC sz_fc_lhs. +have irr_fc: irreducible_poly (f c) by apply: cubic_irreducible; rewrite ?sz_fc. +have fc_monic : f c \is monic. + rewrite monicE lead_coefDl ?size_XsubC ?sz_fc_lhs // -monicE. + by rewrite !monicMl ?monicXsubC ?monicX. +pose inF := [rmorphism of in_alg F]; pose fcF := map_poly inF (f c). +have /existsP[a fcFa_0]: [exists a : F, root fcF a]. + suffices: ~~ coprimep (f c) ('X ^+ #|F| - 'X). + apply: contraR; rewrite -(coprimep_map inF) negb_exists => /forallP-nz_fcF. + rewrite -/fcF rmorphB rmorphX /= map_polyX finField_genPoly. + elim/big_rec: _ => [|x gF _ co_fcFg]; first exact: coprimep1. + by rewrite coprimep_mulr coprimep_XsubC nz_fcF. + have /irredp_FAdjoin[L dimL [z /coprimep_root fcz0 _]] := irr_fc. + pose finL := [vectType 'F_p of FinFieldExtType L]. + set fcL := map_poly _ _ in fcz0; pose inL := [rmorphism of in_alg L]. + rewrite -(coprimep_map inL) -/fcL rmorphB rmorphX /= map_polyX. + apply: contraL (fcz0 _) _; rewrite hornerD hornerN hornerXn hornerX subr_eq0. + have ->: #|F| = #|{: finL}%VS| by rewrite oF card_vspace dimL sz_fc oF_p q3. + by rewrite card_vspacef (expf_card (z : finL)). +have Fp_fcF: fcF \is a polyOver Fp. + by apply/polyOverP => i; rewrite coef_map /= memvZ ?memv_line. +pose G := 'Gal(Fpq / Fp). +have galG: galois Fp Fpq by rewrite finField_galois ?subvf. +have oG: #|G| = 3 by rewrite -galois_dim // dimv1 dimFpq q3. +have Fp'a: a \notin Fp. + by apply: contraL fcFa_0 => /vlineP[d ->]; rewrite fmorph_root. +have DfcF: fcF = \prod_(beta in G) ('X - (beta a)%:P). + pose Pa : {poly F} := minPoly Fp a. + have /eqP szPa: size Pa == 4. + rewrite size_minPoly eqSS. + rewrite (sameP eqP (prime_nt_dvdP _ _)) ?adjoin_deg_eq1 //. + by rewrite adjoin_degreeE dimv1 divn1 -q3 -dimFpq field_dimS ?subvf. + have dvd_Pa_fcF: Pa %| fcF by apply: minPoly_dvdp fcFa_0. + have{dvd_Pa_fcF} /eqP <-: Pa == fcF. + rewrite -eqp_monic ?monic_minPoly ?monic_map // -dvdp_size_eqp //. + by rewrite szPa size_map_poly sz_fc. + have [r [srG /map_uniq Ur defPa]]:= galois_factors (subvf _) galG a (memvf a). + rewrite -/Pa big_map in defPa; rewrite defPa big_uniq //=. + apply/eq_bigl/subset_cardP=> //; apply/eqP. + by rewrite -eqSS (card_uniqP Ur) oG -szPa defPa size_prod_XsubC. +exists a; rewrite !inE; apply/and3P; split. +- by apply: contraNneq Fp'a => ->; apply: mem1v. +- apply/eqP; transitivity ((- 1) ^+ #|G| * fcF.[inF 0]). + rewrite DfcF horner_prod -prodrN; apply: eq_bigr => beta _. + by rewrite rmorph0 hornerXsubC add0r opprK. + by rewrite -signr_odd mulr_sign oG horner_map fc0 rmorphN1 opprK. +apply/eqP; transitivity (fcF.[inF 2%:R]); last by rewrite horner_map fc2 rmorph1. +rewrite DfcF horner_prod; apply: eq_bigr => beta _. +by rewrite hornerXsubC rmorphB !rmorph_nat. +Qed. + +Section AppendixC3. + +Import GroupScope. + +Variables y : gT. +Hypotheses (QP0y : y \in [~: Q, P0]) (nUP0y : P0 :^ y \subset 'N(U)). +Let Qy : y \in Q. Proof. by rewrite (subsetP sQP0Q). Qed. + +Let t := s ^ y. +Let P1 := P0 :^ y. + +(* This is B & G, Appendix C, Lemma C.3, Step 1. *) +Let splitH x : + x \in H -> + exists2 u, u \in U & exists2 v, v \in U & exists2 s1, s1 \in P0 + & x = u * s1 * v. +Proof. +case/(mem_sdprod defH) => z [v [Pz Uv -> _]]. +have [-> | nt_z] := eqVneq z 1. + by exists 1 => //; exists v => //; exists 1; rewrite ?mulg1. +have nz_z: sigma z != 0 by rewrite (morph_injm_eq1 inj_sigma). +have /(mem_dprod defFU)[]: finField_unit nz_z \in setT := in_setT _. +move=> _ [w [/morphimP[u Uu _ ->] Fp_w /(congr1 val)/= Dz _]]. +have{Fp_w Dz} [n Dz]: exists n, sigma z = sigma ((s ^+ n) ^ u). + move: Fp_w; rewrite {}Dz inE => /vlineP[n ->]; exists n. + by rewrite -{1}(natr_Zp n) scaler_nat mulr_natr conjXg !sigmaE ?in_PU. +exists u^-1; last exists (u * v); rewrite ?groupV ?groupM //. +exists (s ^+ n); rewrite ?groupX // mulgA; congr (_ * _). +by apply: (injmP inj_sigma); rewrite -?mulgA ?in_PU. +Qed. + +(* This is B & G, Appendix C, Lemma C.3, Step 2. *) +Let not_splitU s1 s2 u : + s1 \in P0 -> s2 \in P0 -> u \in U -> s1 * u * s2 \in U -> + (s1 == 1) && (s2 == 1) || (u == 1) && (s1 * s2 == 1). +Proof. +move=> P0s1 P0s2 Uu; have [_ _ _ tiPU] := sdprodP defH. +have [Ps1 Ps2]: s1 \in P /\ s2 \in P by rewrite !(subsetP sP0P). +have [-> | nt_s1 /=] := altP (s1 =P 1). + by rewrite mul1g groupMl // -in_set1 -set1gE -tiPU inE Ps2 => ->. +have [-> | nt_u /=] := altP (u =P 1). + by rewrite mulg1 -in_set1 -set1gE -tiPU inE (groupM Ps1). +rewrite (conjgC _ u) -mulgA groupMl // => Us12; case/negP: nt_u. +rewrite -(morph_injm_eq1 inj_sigmaU) // -in_set1 -set1gE. +have [_ _ _ <-] := dprodP defFU; rewrite !inE mem_morphim //= -/(psi u). +have{Us12}: s1 ^ u * s2 == 1. + by rewrite -in_set1 -set1gE -tiPU inE Us12 andbT !in_PU. +rewrite -(morph_injm_eq1 inj_sigma) ?(in_PU, sigmaE) // addr_eq0. +move/eqP/(canRL (mulKf _))->; rewrite ?morph_injm_eq1 //. +by rewrite mulrC rpred_div ?rpredN //= -sigmaP0 mem_morphim. +Qed. + +(* This is B & G, Appendix C, Lemma C.3, Step 3. *) +Let tiH_P1 t1 : t1 \in P1^# -> H :&: H :^ t1 = U. +Proof. +case/setD1P=>[nt_t1 P1t1]; set X := H :&: _. +have [nsPH sUH _ _ tiPU] := sdprod_context defH. +have sUX: U \subset X. + by rewrite subsetI sUH -(normsP nUP0y t1 P1t1) conjSg. +have defX: (P :&: X) * U = X. + by rewrite setIC group_modr // (sdprodW defH) setIAC setIid. +have [tiPX | ntPX] := eqVneq (P :&: X) 1; first by rewrite -defX tiPX mul1g. +have irrPU: acts_irreducibly U P 'J. + apply/mingroupP; (split=> [|V /andP[ntV]]; rewrite astabsJ) => [|nVU sVP]. + by have [_ ->] := Frobenius_context frobH. + apply/eqP; rewrite eqEsubset sVP; apply/subsetP=> x Px. + have [-> // | ntx] := eqVneq x 1. + have [z Vz ntz] := trivgPn _ ntV; have Pz := subsetP sVP z Vz. + have nz_z: sigma z != 0%R by rewrite morph_injm_eq1. + have uP: (sigma x / sigma z)%R \is a GRing.unit. + by rewrite unitfE mulf_neq0 ?invr_eq0 ?morph_injm_eq1. + have: FinRing.unit F uP \in setT := in_setT _. + case/(mem_dprod defFU)=> _ [s1 [/morphimP[u Uu _ ->]]]. + rewrite inE => /vlineP[n Ds1] /(congr1 val)/= Dx _. + suffices ->: x = (z ^ u) ^+ n by rewrite groupX ?memJ_norm ?(subsetP nVU). + apply: (injmP inj_sigma); rewrite ?(in_PU, sigmaE) //. + by rewrite -mulr_natr -scaler_nat natr_Zp -Ds1 -mulrA -Dx mulrC divfK. +have{ntPX defX irrPU} defX: X :=: H. + rewrite -(sdprodW defH) -defX; congr (_ * _). + have [_ -> //] := mingroupP irrPU; rewrite ?subsetIl //= -/X astabsJ ntPX. + by rewrite normsI // normsG. +have nHt1: t1 \in 'N(H) by rewrite -groupV inE sub_conjgV; apply/setIidPl. +have oP0: #|P0| = p by rewrite -(card_injm inj_sigma) // (eq_card sigmaP0) oFp. +have{nHt1} nHP1: P1 \subset 'N(H). + apply: prime_meetG; first by rewrite cardJg oP0. + by apply/trivgPn; exists t1; rewrite // inE P1t1. +have{nHP1} nPP1: P1 \subset 'N(P). + have /Hall_pi hallP: Hall H P by apply: Frobenius_ker_Hall frobH. + by rewrite -(normal_Hall_pcore hallP nsPH) (char_norm_trans (pcore_char _ _)). +have sylP0: p.-Sylow(Q <*> P0) P0. + rewrite /pHall -divgS joing_subr ?(pgroupS sP0P) //=. + by rewrite norm_joinEr // coprime_cardMg ?(coprimegS sP0P) ?mulnK. +have sP1QP0: P1 \subset Q <*> P0. + by rewrite conj_subG ?joing_subr ?mem_gen // inE Qy. +have nP10: P1 \subset 'N(P0). + have: P1 \subset 'N(P :&: (Q <*> P0)) by rewrite normsI // normsG. + by rewrite norm_joinEr // -group_modr // setIC coprime_TIg // mul1g. +have eqP10: P1 :=: P0. + apply/eqP; rewrite eqEcard cardJg leqnn andbT. + rewrite (comm_sub_max_pgroup (Hall_max sylP0)) //; last exact: normC. + by rewrite pgroupJ (pHall_pgroup sylP0). +have /idPn[] := prime_gt1 pr_p. +rewrite -oP0 cardG_gt1 negbK -subG1 -(Frobenius_trivg_cent frobH) subsetI sP0P. +apply/commG1P/trivgP; rewrite -tiPU commg_subI // subsetI ?subxx //. +by rewrite sP0P -eqP10. +Qed. + +(* This is B & G, Appendix C, Lemma C.3, Step 4. *) +Fact BGappendixC3_Ediv : E = [set x^-1 | x in E]%R. +Proof. +suffices sEV_E: [set x^-1 | x in E]%R \subset E. + by apply/esym/eqP; rewrite eqEcard sEV_E card_imset //=; apply: invr_inj. +have /mulG_sub[/(subset_trans sP0P)/subsetP sP0H /subsetP sUH] := sdprodW defH. +have Hs := sP0H s P0s; have P1t: t \in P1 by rewrite memJ_conjg. +have nUP1 t1: t1 \in P1 -> U :^ t1 = U by move/(subsetP nUP0y)/normP. +have nUtn n u: u \in U -> u ^ (t ^+ n) \in U. + by rewrite -mem_conjgV nUP1 ?groupV ?groupX. +have nUtVn n u: u \in U -> u ^ (t ^- n) \in U. + by rewrite -mem_conjg nUP1 ?groupX. +have Qsti i: s ^- i * t ^+ i \in Q. + by rewrite -conjXg -commgEl (subsetP sQP0Q) // commGC mem_commg ?groupX. +pose is_sUs m a j n u s1 v := + [/\ a \in U, u \in U, v \in U, s1 \in P0 + & s ^+ m * a ^ t ^+ j * s ^- n = u * s1 * v]. +have split_sUs m a j n: + a \in U -> exists u, exists s1, exists v, is_sUs m a j n u s1 v. +- move=> Ua; suffices: s ^+ m * a ^ t ^+ j * s ^- n \in H. + by case/splitH=> u Uu [v Uv [s1 P0s1 Dusv1]]; exists u, s1, v. + by rewrite 2?groupM ?groupV ?groupX // sUH ?nUtn. +have nt_sUs m j n a u s1 v: + (m == n.+1) || (n == m.+1) -> is_sUs m a j n u s1 v -> s1 != 1. +- move/pred2P=> Dmn [Ua Uu Uv _ Dusv]; have{Dmn}: s ^+ m != s ^+ n. + by case: Dmn => ->; last rewrite eq_sym; rewrite expgS eq_mulgV1 ?mulgK. + apply: contraNneq => s1_1; rewrite {s1}s1_1 mulg1 in Dusv. + have:= groupM Uu Uv; rewrite -Dusv => /(not_splitU _ _ (nUtn j a Ua))/orP. + by rewrite !in_group // eq_invg1 -eq_mulgV1 => -[]// /andP[? /eqP->]. +have sUs_modP m a j n u s1 v: is_sUs m a j n u s1 v -> a ^ t ^+ j = u * v. + have [nUP /isom_inj/injmP/=quoUP_inj] := sdprod_isom defH. + case=> Ua Uu Uv P0s1 /(congr1 (coset P)); rewrite (conjgCV u) -(mulgA _ u). + rewrite coset_kerr ?groupV 2?coset_kerl ?groupX //; last first. + by rewrite -mem_conjg (normsP nUP) // (subsetP sP0P). + by move/quoUP_inj->; rewrite ?nUtn ?groupM. +have expUMp u v Uu Uv := expgMn p (centsP cUU u v Uu Uv). +have sUsXp m a j n u s1 v: + is_sUs m a j n u s1 v -> is_sUs m (a ^+ p) j n (u ^+ p) s1 (v ^+ p). +- move=> Dusv; have{Dusv} [/sUs_modP Duv [Ua Uu Vv P0s1 Dusv]] := (Dusv, Dusv). + split; rewrite ?groupX //; move: P0s1 Dusv; rewrite -defP0 => /cycleP[k ->]. + rewrite conjXg -!(mulgA _ (s ^+ k)) ![s ^+ k * _]conjgC 2!mulgA -expUMp //. + rewrite {}Duv ![s ^+ m * _]conjgC !conjXg -![_ * _ * s ^- n]mulgA. + move/mulgI/(congr1 (Frobenius_aut charFp \o sigma))=> /= Duv_p. + congr (_ * _); apply/(injmP inj_sigma); rewrite ?in_PU //. + by rewrite !{1}sigmaE ?in_PU // rmorphB !rmorphMn rmorph1 in Duv_p *. +have odd_P: odd #|P| by rewrite oP odd_exp odd_p orbT. +suffices EpsiV a: a \in U -> psi a \in E -> psi (a^-1 ^ t ^+ 3) \in E. + apply/subsetP => _ /imsetP[x Ex ->]. + have /imsetP[a Ua Dx]: x \in psi @: U by rewrite im_psi; case/setIdP: Ex. + suffices: psi (a^-1 ^ t ^+ (3 * #|P|)) \in E. + rewrite Dx -psiV // -{2}(conjg1 a^-1); congr (psi (_ ^ _) \in E). + by apply/eqP; rewrite -order_dvdn orderJ dvdn_mull ?order_dvdG. + rewrite -(odd_double_half #|P|) odd_P addnC. + elim: _./2 => [|n /EpsiV/EpsiV/=]; first by rewrite EpsiV -?Dx. + by rewrite conjVg invgK -!conjgM -!expgD -!mulnSr !(groupV, nUtn) //; apply. +move=> Ua Ea; have{Ea} [b Ub Dab]: exists2 b, b \in U & psi a + psi b = 2%:R. + case/setIdP: Ea => _; rewrite -im_psi => /imsetP[b Ub Db]; exists b => //. + by rewrite -Db addrC subrK. +(* In the book k is arbitrary in Fp; however only k := 3 is used. *) +have [u2 [s2 [v2 usv2P]]] := split_sUs 3 (a * _) 2 1%N (groupM Ua (groupVr Ub)). +have{Ua} [u1 [s1 [v1 usv1P]]] := split_sUs 1%N a^-1 3 2 (groupVr Ua). +have{Ub} [u3 [s3 [v3 usv3P]]] := split_sUs 2 b 1%N 3 Ub. +pose s2def w1 w2 w3 := t * s2^-1 * t = w1 * s3 * w2 * t ^+ 2 * s1 * w3. +pose w1 := v2 ^ t^-1 * u3; pose w2 := v3 * u1 ^ t ^- 2; pose w3 := v1 * u2 ^ t. +have stXC m n: (m <= n)%N -> s ^- n ^ t ^+ m = s ^- m ^ t ^+ n * s ^- (n - m). + move/subnK=> Dn; apply/(mulgI (s ^- (n - m) * t ^+ n))/(mulIg (t ^+ (n - m))). + rewrite -{1}[in t ^+ n]Dn expgD !mulgA !mulgK -invMg -2!mulgA -!expgD. + by rewrite addnC Dn (centsP (abelem_abelian abelQ)) ?mulgA. +wlog suffices Ds2: a b u1 v1 u2 v2 u3 v3 @w1 @w2 @w3 Dab usv1P usv2P usv3P / + s2def w1 w2 w3; last first. +- apply/esym; rewrite -[_ * t]mulgA [_ * t]conjgC mulgA -(expgS _ 1) conjVg. + rewrite /w2 mulgA; apply: (canRL (mulKVg _)); rewrite 2!mulgA -conjgE. + rewrite conjMg conjgKV /w3 mulgA; apply: (canLR (mulgKV _)). + rewrite /w1 -4!mulgA (mulgA u1) (mulgA u3) conjMg -conjgM mulKg -mulgA. + have [[[Ua _ _ _ <-] [_ _ _ _ Ds2]] [Ub _ _ _ <-]] := (usv1P, usv2P, usv3P). + apply: (canLR (mulKVg _)); rewrite -!invMg -!conjMg -{}Ds2 groupV in Ua *. + rewrite -[t]expg1 2!conjMg -conjgM -expgS 2!conjMg -conjgM -expgSr mulgA. + apply: (canLR (mulgK _)); rewrite 2!invMg -!conjVg invgK invMg invgK -4!mulgA. + rewrite (mulgA _ s) stXC // mulgKV -!conjMg stXC // mulgKV -conjMg conjgM. + apply: (canLR (mulKVg _)); rewrite -2!conjVg 2!mulgA -conjMg (stXC 1%N) //. + rewrite mulgKV -conjgM -expgSr -mulgA -!conjMg; congr (_ ^ t ^+ 3). + apply/(canLR (mulKVg _))/(canLR (mulgK _))/(canLR invgK). + rewrite -!mulgA (mulgA _ b) mulgA invMg -!conjVg !invgK. + by apply/(injmP inj_sigma); rewrite 1?groupM ?sigmaE ?memJ_P. +have [[Ua Uu1 Uv1 P0s1 Dusv1] /sUs_modP-Duv1] := (usv1P, usv1P). +have [[_ Uu2 Uv2 P0s2 _] [Ub Uu3 Uv3 P0s3 _]] := (usv2P, usv3P). +suffices /(congr1 sigma): s ^+ 2 = s ^ v1 * s ^ a^-1 ^ t ^+ 3. + rewrite inE sigmaX // sigma_s sigmaM ?memJ_P -?psiE ?nUtn // => ->. + by rewrite addrK -!im_psi !mem_imset ?nUtn. +rewrite groupV in Ua; have [Hs1 Hs3]: s1 \in H /\ s3 \in H by rewrite !sP0H. +have nt_s1: s1 != 1 by apply: nt_sUs usv1P. +have nt_s3: s3 != 1 by apply: nt_sUs usv3P. +have{sUsXp} Ds2p: s2def (w1 ^+ p) (w2 ^+ p) (w3 ^+ p). + have [/sUsXp-usv1pP /sUsXp-usv2pP /sUsXp-usv3pP] := And3 usv1P usv2P usv3P. + rewrite expUMp ?groupV // !expgVn in usv1pP usv2pP. + rewrite !(=^~ conjXg _ _ p, expUMp) ?groupV -1?[t]expg1 ?nUtn ?nUtVn //. + apply: Ds2 usv1pP usv2pP usv3pP => //. + by rewrite !psiX // -!Frobenius_autE -rmorphD Dab rmorph_nat. +have{Ds2} Ds2: s2def w1 w2 w3 by apply: Ds2 usv1P usv2P usv3P. +wlog [Uw1 Uw2 Uw3]: w1 w2 w3 Ds2p Ds2 / [/\ w1 \in U, w2 \in U & w3 \in U]. + by move/(_ w1 w2 w3)->; rewrite ?(nUtVn, nUtVn 1%N, nUtn 1%N, in_group). +have{Ds2p} Dw3p: (w2 ^- p * w1 ^- p.-1 ^ s3 * w2) ^ t ^+ 2 = w3 ^+ p.-1 ^ s1^-1. + rewrite -[w1 ^+ _](mulKg w1) -[w3 ^+ _](mulgK w3) -expgS -expgSr !prednK //. + rewrite -(canLR (mulKg _) Ds2p) -(canLR (mulKg _) Ds2) 6!invMg !invgK. + by rewrite mulgA mulgK [2]lock /conjg !mulgA mulVg mul1g mulgK. +have w_id w: w \in U -> w ^+ p.-1 == 1 -> w = 1. + by move=> Uw /eqP/(canRL_in (expgK _) Uw)->; rewrite ?expg1n ?oU. +have{Uw3} Dw3: w3 = 1. + apply: w_id => //; have:= @not_splitU s1^-1^-1 s1^-1 (w3 ^+ p.-1). + rewrite !groupV mulVg eqxx andbT {2}invgK (negPf nt_s1) groupX //= => -> //. + have /tiH_P1 <-: t ^+ 2 \in P1^#. + rewrite 2!inE groupX // andbT -order_dvdn gtnNdvd // orderJ. + by rewrite odd_gt2 ?order_gt1 // orderE defP0 (oddSg sP0P). + by rewrite -mulgA -conjgE inE -{2}Dw3p memJ_conjg !in_group ?Hs1 // sUH. +have{Dw3p} Dw2p: w2 ^+ p.-1 = w1 ^- p.-1 ^ s3. + apply/(mulIg w2)/eqP; rewrite -expgSr prednK // eq_mulVg1 mulgA. + by rewrite (canRL (conjgK _) Dw3p) Dw3 expg1n !conj1g. +have{Uw1} Dw1: w1 = 1. + apply: w_id => //; have:= @not_splitU s3^-1 s3 (w1 ^- p.-1). + rewrite mulVg (negPf nt_s3) andbF -mulgA -conjgE -Dw2p !in_group //=. + by rewrite eqxx andbT eq_invg1 /= => ->. +have{w1 w2 w3 Dw1 Dw3 w_id Uw2 Dw2p Ds2} Ds2: t * s2^-1 * t = s3 * t ^+ 2 * s1. + by rewrite Ds2 Dw3 [w2]w_id ?mulg1 ?Dw2p ?Dw1 ?mul1g // expg1n invg1 conj1g. +have /centsP abP0: abelian P0 by rewrite -defP0 cycle_abelian. +have QP0ys := memJ_norm y (subsetP (commg_normr P0 Q) _ _). +have{QP0ys} memQP0 := (QP0ys, groupV, groupM); have nQ_P0 := subsetP nQP0. +have sQP0_Q: [~: Q, P0] \subset Q by rewrite commg_subl. +have /centsP abQP0 := abelianS sQP0_Q (abelem_abelian abelQ). +have{s2def} Ds312: s3 * s1 * s2 = 1. + apply/set1P; rewrite -set1gE -(coprime_TIg coQP) inE. + rewrite coset_idr ?(subsetP sP0P) ?nQ_P0 ?groupM //. + rewrite -mulgA -[s2](mulgK s) [_ * s]abP0 // -[s2](mulKVg s). + rewrite -!mulgA [s * _]mulgA [s1 * _]mulgA [s1 * _]abP0 ?groupM //. + rewrite 2!(mulgA s3) [s^-1 * _]mulgA !(morphM, morphV) ?nQ_P0 ?in_group //=. + have ->: coset Q s = coset Q t by rewrite coset_kerl ?groupV ?coset_kerr. + have nQt: t \in 'N(Q) by rewrite -(conjGid Qy) normJ memJ_conjg nQ_P0. + rewrite -morphV // -!morphM ?(nQt, groupM) ?groupV // ?nQ_P0 //= -Ds2. + by rewrite 2!mulgA mulgK mulgKV mulgV morph1. +pose x := (y ^ s3)^-1 * y ^ s^-1 * (y ^ (s * s1)^-1)^-1 * y. +have{abP0} Dx: x ^ s^-1 = x. + rewrite 3!conjMg !conjVg -!conjgM -!invMg (mulgA s) -(expgS _ 1). + rewrite [x]abQP0 ?memQP0 // [rhs in y * rhs]abQP0 ?memQP0 //. + apply/(canRL (mulKVg _)); rewrite 4!mulgA; congr (_ * _). + rewrite [RHS]abQP0 ?memQP0 //; apply/(canRL (mulgK _))/eqP. + rewrite -3!mulgA [rhs in y^-1 * rhs]abQP0 ?memQP0 // -eq_invg_sym eq_invg_mul. + apply/eqP; transitivity (t ^+ 2 * s1 * (t^-1 * s2 * t^-1) * s3); last first. + by rewrite -[s2]invgK -!invMg mulgA Ds2 -(mulgA s3) invMg mulKVg mulVg. + rewrite (canRL (mulKg _) Ds312) -2![_ * t^-1]mulgA. + have Dt1 si: si \in P0 -> t^-1 = (s^-1 ^ si) ^ y. + by move=> P0si; rewrite {2}/conjg -conjVg -(abP0 si) ?groupV ?mulKg. + by rewrite {1}(Dt1 s1) // (Dt1 s3^-1) ?groupV // -conjXg /conjg !{1}gnorm. +have{Dx memQP0} Dx1: x = 1. + apply/set1P; rewrite -set1gE; have [_ _ _ <-] := dprodP defQ. + rewrite setIAC (setIidPr sQP0_Q) inE -{2}defP0 -cycleV cent_cycle. + by rewrite (sameP cent1P commgP) commgEl Dx mulVg eqxx !memQP0. +pose t1 := s1 ^ y; pose t3 := s3 ^ y. +have{x Dx1} Ds13: s1 * (t * t1)^-1 = (t3 * t)^-1 * s3. + by apply/eqP; rewrite eq_sym eq_mulVg1 invMg invgK -Dx1 /x /conjg !gnorm. +suffices Ds1: s1 = s^-1. + rewrite -(canLR (mulKg _) (canRL (mulgKV _) Dusv1)) Ds1 Duv1. + by rewrite !invMg invgK /conjg !gnorm. +have [_ _ /trivgPn[u Uu nt_u] _ _] := Frobenius_context frobH. +apply: (conjg_inj y); apply: contraNeq nt_u. +rewrite -/t1 conjVg -/t eq_mulVg1 -invMg => nt_tt1. +have Hu := sUH u Uu; have P1tt1: t * t1 \in P1 by rewrite groupM ?memJ_conjg. +have /tiH_P1 defU: (t * t1)^-1 \in P1^# by rewrite 2!inE nt_tt1 groupV. +suffices: (u ^ s1) ^ (t * t1)^-1 \in U. + rewrite -mem_conjg nUP1 // conjgE mulgA => /(not_splitU _ _ Uu). + by rewrite groupV (negPf nt_s1) andbF mulVg eqxx andbT /= => /(_ _ _)/eqP->. +rewrite -defU inE memJ_conjg -conjgM Ds13 conjgM groupJ ?(groupJ _ Hs1) //. +by rewrite sUH // -mem_conjg nUP1 // groupM ?memJ_conjg. +Qed. + +End AppendixC3. + +Fact BGappendixC_inner_subproof : (p <= q)%N. +Proof. +have [y QP0y nUP0y] := nU_P0QP0. +by apply: Einv_gt1_le_pq E_gt1; apply: BGappendixC3_Ediv nUP0y. +Qed. + +End ExpandHypotheses. + +(* This is B & G, Appendix C, Theorem C. *) +Theorem prime_dim_normed_finField : (p <= q)%N. +Proof. +apply: wlog_neg; rewrite -ltnNge => ltqp. +have [F sigma /isomP[inj_sigma im_sigma] defP0] := fieldH. +case=> sigmaU inj_sigmaU sigmaJ. +have oF: #|F| = (p ^ q)%N by rewrite -cardsT -im_sigma card_injm. +have charFp: p \in [char F] := card_finCharP oF pr_p. +have sP0P: P0 \subset P by rewrite -defP0 subsetIl. +pose s := invm inj_sigma 1%R. +have sigma_s: sigma s = 1%R by rewrite invmK ?im_sigma ?inE. +have{defP0} defP0: <[s]> = P0. + by rewrite -morphim_cycle /= ?im_sigma ?inE // morphim_invmE. +exact: BGappendixC_inner_subproof defP0 sigmaJ. +Qed. + +End AppendixC. diff --git a/mathcomp/odd_order/BGsection1.v b/mathcomp/odd_order/BGsection1.v new file mode 100644 index 0000000..98a2d08 --- /dev/null +++ b/mathcomp/odd_order/BGsection1.v @@ -0,0 +1,1340 @@ +Set Printing Width 50. +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div fintype. +Require Import bigop prime binomial finset fingroup morphism perm automorphism. +Require Import quotient action gproduct gfunctor commutator. +Require Import ssralg finalg zmodp cyclic center pgroup finmodule gseries. +Require Import nilpotent sylow abelian maximal hall extremal. +Require Import matrix mxalgebra mxrepresentation mxabelem. + +(******************************************************************************) +(* This file contains most of the material in B & G, section 1, including the *) +(* definitions: *) +(* p.-length_1 G == the upper p-series of G has length <= 1, i.e., *) +(* 'O_{p^',p,p^'}(G) = G *) +(* p_elt_gen p G == the subgroup of G generated by its p-elements. *) +(* This file currently covers B & G 1.3-4, 1.6, 1.8-1.21, and also *) +(* Gorenstein 8.1.3 and 2.8.1 (maximal order of a p-subgroup of GL(2,p)). *) +(* This file also provides, mostly for future reference, the following *) +(* definitions, drawn from Gorenstein, Chapter 8, and B & G, Appendix B: *) +(* p.-constrained G <-> the p',p core of G contains the centralisers of *) +(* all its Sylow p-subgroups. The Hall-Higman Lemma *) +(* 1.2.3 (B & G, 1.15a) asserts that this holds for *) +(* all solvable groups. *) +(* p.-stable G <-> a rather group theoretic generalization of the *) +(* Hall-Higman type condition that in a faithful *) +(* p-modular linear representation of G no p-element *) +(* has a quadratic minimal polynomial, to groups G *) +(* with a non-trivial p-core. *) +(* p.-abelian_constrained <-> the p',p core of G contains all the normal *) +(* abelian subgroups of the Sylow p-subgroups of G. *) +(* It is via this property and the ZL theorem (the *) +(* substitute for the ZJ theorem) that the *) +(* p-stability of groups of odd order is exploited *) +(* in the proof of the Odd Order Theorem. *) +(* generated_by p G == G is generated by a set of subgroups satisfying *) +(* p : pred {group gT} *) +(* norm_abelian X A == A is abelian and normalised by X. *) +(* p_norm_abelian p X A == A is an abelian p-group normalised by X. *) +(* 'L_[G](X) == the group generated by the abelian subgroups of G *) +(* normalized by X. *) +(* 'L_{n}(G) == the Puig group series, defined by the recurrence *) +(* 'L_{0}(G) = 1, 'L_{n.+1}(G) = 'L_[G]('L_{n}(G)). *) +(* 'L_*(G) == the lower limit of the Puig series. *) +(* 'L(G) == the Puig subgroup: the upper limit of the Puig *) +(* series: 'L(G) = 'L_[G]('L_*(G)) and conversely. *) +(* The following notation is used locally here and in AppendixB, but is NOT *) +(* exported: *) +(* D --> G == G is generated by abelian groups normalised by D *) +(* := generated_by (norm_abelian D) G *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Section Definitions. + +Variables (n : nat) (gT : finGroupType). +Implicit Type p : nat. + +Definition plength_1 p (G : {set gT}) := 'O_{p^', p, p^'}(G) == G. + +Definition p_elt_gen p (G : {set gT}) := <<[set x in G | p.-elt x]>>. + +Definition p_constrained p (G : {set gT}) := + forall P : {group gT}, + p.-Sylow('O_{p^',p}(G)) P -> + 'C_G(P) \subset 'O_{p^',p}(G). + +Definition p_abelian_constrained p (G : {set gT}) := + forall S A : {group gT}, + p.-Sylow(G) S -> abelian A -> A <| S -> + A \subset 'O_{p^',p}(G). + +Definition p_stable p (G : {set gT}) := + forall P A : {group gT}, + p.-group P -> 'O_p^'(G) * P <| G -> + p.-subgroup('N_G(P)) A -> [~: P, A, A] = 1 -> + A / 'C_G(P) \subset 'O_p('N_G(P) / 'C_G(P)). + +Definition generated_by (gp : pred {group gT}) (E : {set gT}) := + [exists gE : {set {group gT}}, <<\bigcup_(G in gE | gp G) G>> == E]. + +Definition norm_abelian (D : {set gT}) : pred {group gT} := + fun A => (D \subset 'N(A)) && abelian A. + +Definition p_norm_abelian p (D : {set gT}) : pred {group gT} := + fun A => p.-group A && norm_abelian D A. + +Definition Puig_succ (D E : {set gT}) := + <<\bigcup_(A in subgroups D | norm_abelian E A) A>>. + +Definition Puig_rec D := iter n (Puig_succ D) 1. + +End Definitions. + +(* This must be defined outside a Section to avoid spurrious delta-reduction *) +Definition Puig_at := nosimpl Puig_rec. + +Definition Puig_inf (gT : finGroupType) (G : {set gT}) := Puig_at #|G|.*2 G. + +Definition Puig (gT : finGroupType) (G : {set gT}) := Puig_at #|G|.*2.+1 G. + +Notation "p .-length_1" := (plength_1 p) + (at level 2, format "p .-length_1") : group_scope. + +Notation "p .-constrained" := (p_constrained p) + (at level 2, format "p .-constrained") : group_scope. +Notation "p .-abelian_constrained" := (p_abelian_constrained p) + (at level 2, format "p .-abelian_constrained") : group_scope. +Notation "p .-stable" := (p_stable p) + (at level 2, format "p .-stable") : group_scope. + +Notation "''L_[' G ] ( L )" := (Puig_succ G L) + (at level 8, format "''L_[' G ] ( L )") : group_scope. +Notation "''L_{' n } ( G )" := (Puig_at n G) + (at level 8, format "''L_{' n } ( G )") : group_scope. +Notation "''L_*' ( G )" := (Puig_inf G) + (at level 8, format "''L_*' ( G )") : group_scope. +Notation "''L' ( G )" := (Puig G) + (at level 8, format "''L' ( G )") : group_scope. + +Section BGsection1. + +Implicit Types (gT : finGroupType) (p : nat). + +(* This is B & G, Lemma 1.1, first part. *) +Lemma minnormal_solvable_abelem gT (M G : {group gT}) : + minnormal M G -> solvable M -> is_abelem M. +Proof. by move=> minM solM; case: (minnormal_solvable minM (subxx _) solM). Qed. + +(* This is B & G, Lemma 1.2, second part. *) +Lemma minnormal_solvable_Fitting_center gT (M G : {group gT}) : + minnormal M G -> M \subset G -> solvable M -> M \subset 'Z('F(G)). +Proof. +have nZG: 'Z('F(G)) <| G. + by apply: (char_normal_trans (center_char _)); exact: Fitting_normal. +move=> minM sMG solM; have[/andP[ntM nMG] minM'] := mingroupP minM. +apply/setIidPl/minM'; last exact: subsetIl. +apply/andP; split; last by rewrite normsI // normal_norm. +apply: meet_center_nil => //; first by apply: Fitting_nil. +apply/andP; split; last by apply: subset_trans nMG; exact: Fitting_sub. +apply: Fitting_max; rewrite // /normal ?sMG //; apply: abelian_nil. +by move: (minnormal_solvable_abelem minM solM) => /abelem_abelian. +Qed. + +Lemma sol_chief_abelem gT (G V U : {group gT}) : + solvable G -> chief_factor G V U -> is_abelem (U / V). +Proof. +move=> solG chiefUV; have minUV := chief_factor_minnormal chiefUV. +have [|//] := minnormal_solvable minUV (quotientS _ _) (quotient_sol _ solG). +by case/and3P: chiefUV. +Qed. + +Section HallLemma. + +Variables (gT : finGroupType) (G G' : {group gT}). + +Hypothesis solG : solvable G. +Hypothesis nsG'G : G' <| G. +Let sG'G : G' \subset G := normal_sub nsG'G. +Let nG'G : G \subset 'N(G') := normal_norm nsG'G. + +Let nsF'G : 'F(G') <| G := char_normal_trans (Fitting_char G') nsG'G. + +Let Gchief (UV : {group gT} * {group gT}) := chief_factor G UV.2 UV.1. +Let H := \bigcap_(UV | Gchief UV) 'C(UV.1 / UV.2 | 'Q). +Let H' := + G' :&: \bigcap_(UV | Gchief UV && (UV.1 \subset 'F(G'))) 'C(UV.1 / UV.2 | 'Q). + +(* This is B & G Proposition 1.2, non trivial inclusion of the first equality.*) +Proposition Fitting_stab_chief : 'F(G') \subset H. +Proof. +apply/bigcapsP=> [[U V] /= chiefUV]. +have minUV: minnormal (U / V) (G / V) := chief_factor_minnormal chiefUV. +have{chiefUV} [/=/maxgroupp/andP[_ nVG] sUG nUG] := and3P chiefUV. +have solUV: solvable (U / V) by rewrite quotient_sol // (solvableS sUG). +have{solUV minUV}: U / V \subset 'Z('F(G / V)). + exact: minnormal_solvable_Fitting_center minUV (quotientS V sUG) solUV. +rewrite sub_astabQ (subset_trans (normal_sub nsF'G) nVG) /=. +rewrite subsetI centsC => /andP[_]; apply: subset_trans. +by rewrite Fitting_max ?quotient_normal ?quotient_nil ?Fitting_nil. +Qed. + +(* This is B & G Proposition 1.2, non trivial inclusion of the second *) +(* equality. *) +Proposition chief_stab_sub_Fitting : H' \subset 'F(G'). +Proof. +without loss: / {K | [min K | K <| G & ~~ (K \subset 'F(G'))] & K \subset H'}. + move=> IH; apply: wlog_neg => s'H'F; apply/IH/mingroup_exists=> {IH}/=. + rewrite /normal subIset ?sG'G ?normsI ?norms_bigcap {s'H'F}//. + apply/bigcapsP=> /= U /andP[/and3P[/maxgroupp/andP/=[_ nU2G] _ nU1G] _]. + exact: subset_trans (actsQ nU2G nU1G) (astab_norm 'Q (U.1 / U.2)). +case=> K /mingroupP[/andP[nsKG s'KF] minK] /subsetIP[sKG' nFK]. +have [[Ks chiefKs defK] sKG]:= (chief_series_exists nsKG, normal_sub nsKG). +suffices{nsKG s'KF} cKsK: (K.-central).-series 1%G Ks. + by rewrite Fitting_max ?(normalS _ sG'G) ?(centrals_nil cKsK) in s'KF. +move: chiefKs; rewrite -!(rev_path _ _ Ks) {}defK. +case: {Ks}(rev _) => //= K1 Kr /andP[chiefK1 chiefKr]. +have [/maxgroupp/andP[/andP[sK1K ltK1K] nK1G] _] := andP chiefK1. +suffices{chiefK1} cKrK: [rel U V | central_factor K V U].-series K1 Kr. + have cKK1: abelian (K / K1) := abelem_abelian (sol_chief_abelem solG chiefK1). + by rewrite /central_factor subxx sK1K der1_min //= (subset_trans sKG). +have{minK ltK1K nK1G} sK1F: K1 \subset 'F(G'). + have nsK1G: K1 <| G by rewrite /normal (subset_trans sK1K). + by apply: contraR ltK1K => s'K1F; rewrite (minK K1) ?nsK1G. +elim: Kr K1 chiefKr => //= K2 Kr IHr K1 /andP[chiefK2 chiefKr] in sK1F sK1K *. +have [/maxgroupp/andP[/andP[sK21 _] /(subset_trans sKG)nK2K] _] := andP chiefK2. +rewrite /central_factor sK1K {}IHr ?(subset_trans sK21) {chiefKr}// !andbT. +rewrite commGC -sub_astabQR ?(subset_trans _ nK2K) //. +exact/(subset_trans nFK)/(bigcap_inf (K1, K2))/andP. +Qed. + +End HallLemma. + +(* This is B & G, Proposition 1.3 (due to P. Hall). *) +Proposition cent_sub_Fitting gT (G : {group gT}) : + solvable G -> 'C_G('F(G)) \subset 'F(G). +Proof. +move=> solG; apply: subset_trans (chief_stab_sub_Fitting solG _) => //. +rewrite subsetI subsetIl; apply/bigcapsP=> [[U V]] /=. +case/andP=> /andP[/maxgroupp/andP[_ nVG] _] sUF. +by rewrite astabQ (subset_trans _ (morphpre_cent _ _)) // setISS ?centS. +Qed. + +(* This is B & G, Proposition 1.4, for internal actions. *) +Proposition coprime_trivg_cent_Fitting gT (A G : {group gT}) : + A \subset 'N(G) -> coprime #|G| #|A| -> solvable G -> + 'C_A(G) = 1 -> 'C_A('F(G)) = 1. +Proof. +move=> nGA coGA solG regAG; without loss cycA: A nGA coGA regAG / cyclic A. + move=> IH; apply/trivgP/subsetP=> a; rewrite -!cycle_subG subsetI. + case/andP=> saA /setIidPl <-. + rewrite {}IH ?cycle_cyclic ?(coprimegS saA) ?(subset_trans saA) //. + by apply/trivgP; rewrite -regAG setSI. +pose X := G <*> A; pose F := 'F(X); pose pi := \pi(A); pose Q := 'O_pi(F). +have pi'G: pi^'.-group G by rewrite /pgroup -coprime_pi' //= coprime_sym. +have piA: pi.-group A by exact: pgroup_pi. +have oX: #|X| = (#|G| * #|A|)%N by rewrite [X]norm_joinEr ?coprime_cardMg. +have hallG: pi^'.-Hall(X) G. + by rewrite /pHall -divgS joing_subl //= pi'G pnatNK oX mulKn. +have nsGX: G <| X by rewrite /normal joing_subl join_subG normG. +have{oX pi'G piA} hallA: pi.-Hall(X) A. + by rewrite /pHall -divgS joing_subr //= piA oX mulnK. +have nsQX: Q <| X := char_normal_trans (pcore_char _ _) (Fitting_normal _). +have{solG cycA} solX: solvable X. + rewrite (series_sol nsGX) {}solG /= norm_joinEr // quotientMidl //. + by rewrite morphim_sol // abelian_sol // cyclic_abelian. +have sQA: Q \subset A. + by apply: normal_sub_max_pgroup (Hall_max hallA) (pcore_pgroup _ _) nsQX. +have pi'F: 'O_pi(F) = 1. + suff cQG: G \subset 'C(Q) by apply/trivgP; rewrite -regAG subsetI sQA centsC. + apply/commG1P/trivgP; rewrite -(coprime_TIg coGA) subsetI commg_subl. + rewrite (subset_trans sQA) // (subset_trans _ sQA) // commg_subr. + by rewrite (subset_trans _ (normal_norm nsQX)) ?joing_subl. +have sFG: F \subset G. + have /dprodP[_ defF _ _]: _ = F := nilpotent_pcoreC pi (Fitting_nil _). + by rewrite (sub_normal_Hall hallG) ?gFsub //= -defF pi'F mul1g pcore_pgroup. +have <-: F = 'F(G). + apply/eqP; rewrite eqEsubset -{1}(setIidPr sFG) FittingS ?joing_subl //=. + by rewrite Fitting_max ?Fitting_nil // (char_normal_trans (Fitting_char _)). +apply/trivgP; rewrite /= -(coprime_TIg coGA) subsetI subsetIl andbT. +apply: subset_trans (subset_trans (cent_sub_Fitting solX) sFG). +by rewrite setSI ?joing_subr. +Qed. + +(* A "contrapositive" of Proposition 1.4 above. *) +Proposition coprime_cent_Fitting gT (A G : {group gT}) : + A \subset 'N(G) -> coprime #|G| #|A| -> solvable G -> + 'C_A('F(G)) \subset 'C(G). +Proof. +move=> nGA coGA solG; apply: subset_trans (subsetIr A _); set C := 'C_A(G). +rewrite -quotient_sub1 /= -/C; last first. + by rewrite subIset // normsI ?normG // norms_cent. +apply: subset_trans (quotient_subcent _ _ _) _; rewrite /= -/C. +have nCG: G \subset 'N(C) by rewrite cents_norm // centsC subsetIr. +rewrite /= -(setIidPr (Fitting_sub _)) -[(G :&: _) / _](morphim_restrm nCG). +rewrite injmF //=; last first. + by rewrite ker_restrm ker_coset setIA (coprime_TIg coGA) subIset ?subxx. +rewrite morphim_restrm -quotientE setIid. +rewrite coprime_trivg_cent_Fitting ?quotient_norms ?coprime_morph //=. + exact: morphim_sol. +rewrite -strongest_coprime_quotient_cent ?trivg_quotient ?solG ?orbT //. + by rewrite -setIA subsetIl. +by rewrite coprime_sym -setIA (coprimegS (subsetIl _ _)). +Qed. + +(* B & G Proposition 1.5 is covered by several lemmas in hall.v : *) +(* 1.5a -> coprime_Hall_exists (internal action) *) +(* ext_coprime_Hall_exists (general group action) *) +(* 1.5b -> coprime_Hall_subset (internal action) *) +(* ext_coprime_Hall_subset (general group action) *) +(* 1.5c -> coprime_Hall_trans (internal action) *) +(* ext_coprime_Hall_trans (general group action) *) +(* 1.5d -> coprime_quotient_cent (internal action) *) +(* ext_coprime_quotient_cent (general group action) *) +(* several stronger variants are proved for internal action *) +(* 1.5e -> coprime_comm_pcore (internal action only) *) + +(* A stronger variant of B & G, Proposition 1.6(a). *) +Proposition coprimeR_cent_prod gT (A G : {group gT}) : + A \subset 'N(G) -> coprime #|[~: G, A]| #|A| -> solvable [~: G, A] -> + [~: G, A] * 'C_G(A) = G. +Proof. +move=> nGA coRA solR; apply/eqP; rewrite eqEsubset mulG_subG commg_subl nGA. +rewrite subsetIl -quotientSK ?commg_norml //=. +rewrite coprime_norm_quotient_cent ?commg_normr //=. +by rewrite subsetI subxx quotient_cents2r. +Qed. + +(* This is B & G, Proposition 1.6(a). *) +Proposition coprime_cent_prod gT (A G : {group gT}) : + A \subset 'N(G) -> coprime #|G| #|A| -> solvable G -> + [~: G, A] * 'C_G(A) = G. +Proof. +move=> nGA; have sRG: [~: G, A] \subset G by rewrite commg_subl. +rewrite -(Lagrange sRG) coprime_mull => /andP[coRA _] /(solvableS sRG). +exact: coprimeR_cent_prod. +Qed. + +(* This is B & G, Proposition 1.6(b). *) +Proposition coprime_commGid gT (A G : {group gT}) : + A \subset 'N(G) -> coprime #|G| #|A| -> solvable G -> + [~: G, A, A] = [~: G, A]. +Proof. +move=> nGA coGA solG; apply/eqP; rewrite eqEsubset commSg ?commg_subl //. +have nAC: 'C_G(A) \subset 'N(A) by rewrite subIset ?cent_sub ?orbT. +rewrite -{1}(coprime_cent_prod nGA) // commMG //=; first 1 last. + by rewrite !normsR // subIset ?normG. +by rewrite (commG1P (subsetIr _ _)) mulg1. +Qed. + +(* This is B & G, Proposition 1.6(c). *) +Proposition coprime_commGG1P gT (A G : {group gT}) : + A \subset 'N(G) -> coprime #|G| #|A| -> solvable G -> + [~: G, A, A] = 1 -> A \subset 'C(G). +Proof. +by move=> nGA coGA solG; rewrite centsC coprime_commGid // => /commG1P. +Qed. + +(* This is B & G, Proposition 1.6(d), TI-part, from finmod.v *) +Definition coprime_abel_cent_TI := coprime_abel_cent_TI. + +(* This is B & G, Proposition 1.6(d) (direct product) *) +Proposition coprime_abelian_cent_dprod gT (A G : {group gT}) : + A \subset 'N(G) -> coprime #|G| #|A| -> abelian G -> + [~: G, A] \x 'C_G(A) = G. +Proof. +move=> nGA coGA abelG; rewrite dprodE ?coprime_cent_prod ?abelian_sol //. + by rewrite subIset 1?(subset_trans abelG) // centS // commg_subl. +by apply/trivgP; rewrite /= setICA coprime_abel_cent_TI ?subsetIr. +Qed. + +(* This is B & G, Proposition 1.6(e), which generalises Aschbacher (24.3). *) +Proposition coprime_abelian_faithful_Ohm1 gT (A G : {group gT}) : + A \subset 'N(G) -> coprime #|G| #|A| -> abelian G -> + A \subset 'C('Ohm_1(G)) -> A \subset 'C(G). +Proof. +move=> nGA coGA cGG; rewrite !(centsC A) => cAG1. +have /dprodP[_ defG _ tiRC] := coprime_abelian_cent_dprod nGA coGA cGG. +have sRG: [~: G, A] \subset G by rewrite commg_subl. +rewrite -{}defG -(setIidPl sRG) TI_Ohm1 ?mul1g ?subsetIr //. +by apply/trivgP; rewrite -{}tiRC setIS // subsetI Ohm_sub. +Qed. + +(* B & G, Lemma 1.7 is covered by several lemmas in maximal.v : *) +(* 1.7a -> Phi_nongen *) +(* 1.7b -> Phi_quotient_abelem *) +(* 1.7c -> trivg_Phi *) +(* 1.7d -> Phi_joing *) + +(* This is B & G, Proposition 1.8, or Aschbacher 24.1. Note that the coprime *) +(* assumption is slightly weaker than requiring that A be a p'-group. *) +Proposition coprime_cent_Phi gT p (A G : {group gT}) : + p.-group G -> coprime #|G| #|A| -> [~: G, A] \subset 'Phi(G) -> + A \subset 'C(G). +Proof. +move=> pG coGA sRphi; rewrite centsC; apply/setIidPl. +rewrite -['C_G(A)]genGid; apply/Phi_nongen/eqP. +rewrite eqEsubset join_subG Phi_sub subsetIl -genM_join sub_gen //=. +rewrite -{1}(coprime_cent_prod _ coGA) ?(pgroup_sol pG) ?mulSg //. +by rewrite -commg_subl (subset_trans sRphi) ?Phi_sub. +Qed. + +(* This is B & G, Proposition 1.9, base (and most common) case, for internal *) +(* coprime action. *) +Proposition stable_factor_cent gT (A G H : {group gT}) : + A \subset 'C(H) -> stable_factor A H G -> + coprime #|G| #|A| -> solvable G -> + A \subset 'C(G). +Proof. +move=> cHA /and3P[sRH sHG nHG] coGA solG. +suffices: G \subset 'C_G(A) by rewrite subsetI subxx centsC. +rewrite -(quotientSGK nHG) ?subsetI ?sHG 1?centsC //. +by rewrite coprime_quotient_cent ?cents_norm ?subsetI ?subxx ?quotient_cents2r. +Qed. + +(* This is B & G, Proposition 1.9 (for internal coprime action) *) +Proposition stable_series_cent gT (A G : {group gT}) s : + last 1%G s :=: G -> (A.-stable).-series 1%G s -> + coprime #|G| #|A| -> solvable G -> + A \subset 'C(G). +Proof. +move=> <-{G}; elim/last_ind: s => /= [|s G IHs]; first by rewrite cents1. +rewrite last_rcons rcons_path /= => /andP[/IHs{IHs}]. +move: {s}(last _ _) => H IH_H nHGA coGA solG; have [_ sHG _] := and3P nHGA. +by rewrite (stable_factor_cent _ nHGA) ?IH_H ?(solvableS sHG) ?(coprimeSg sHG). +Qed. + +(* This is B & G, Proposition 1.10. *) +Proposition coprime_nil_faithful_cent_stab gT (A G : {group gT}) : + A \subset 'N(G) -> coprime #|G| #|A| -> nilpotent G -> + let C := 'C_G(A) in 'C_G(C) \subset C -> A \subset 'C(G). +Proof. +move=> nGA coGA nilG C; rewrite subsetI subsetIl centsC /= -/C => cCA. +pose N := 'N_G(C); have sNG: N \subset G by rewrite subsetIl. +have sCG: C \subset G by rewrite subsetIl. +suffices cNA : A \subset 'C(N). + rewrite centsC (sameP setIidPl eqP) -(nilpotent_sub_norm nilG sCG) //= -/C. + by rewrite subsetI subsetIl centsC. +have{nilG} solN: solvable N by rewrite(solvableS sNG) ?nilpotent_sol. +rewrite (stable_factor_cent cCA) ?(coprimeSg sNG) /stable_factor //= -/N -/C. +rewrite subcent_normal subsetI (subset_trans (commSg A sNG)) ?commg_subl //=. +rewrite comm_norm_cent_cent 1?centsC ?subsetIr // normsI // !norms_norm //. +by rewrite cents_norm 1?centsC ?subsetIr. +Qed. + +(* B & G, Theorem 1.11, via Aschbacher 24.7 rather than Gorenstein 5.3.10. *) +Theorem coprime_odd_faithful_Ohm1 gT p (A G : {group gT}) : + p.-group G -> A \subset 'N(G) -> coprime #|G| #|A| -> odd #|G| -> + A \subset 'C('Ohm_1(G)) -> A \subset 'C(G). +Proof. +move=> pG nGA coGA oddG; rewrite !(centsC A) => cAG1. +have [-> | ntG] := eqsVneq G 1; first exact: sub1G. +have{oddG ntG} [p_pr oddp]: prime p /\ odd p. + have [p_pr p_dv_G _] := pgroup_pdiv pG ntG. + by rewrite !odd_2'nat in oddG *; rewrite pnatE ?(pgroupP oddG). +without loss defR: G pG nGA coGA cAG1 / [~: G, A] = G. + move=> IH; have solG := pgroup_sol pG. + rewrite -(coprime_cent_prod nGA) ?mul_subG ?subsetIr //=. + have sRG: [~: G, A] \subset G by rewrite commg_subl. + rewrite IH ?coprime_commGid ?(pgroupS sRG) ?commg_normr ?(coprimeSg sRG) //. + by rewrite (subset_trans (OhmS 1 sRG)). +have [|[defPhi defG'] defC] := abelian_charsimple_special pG coGA defR. + apply/bigcupsP=> H /andP[chH abH]; have sHG := char_sub chH. + have nHA := char_norm_trans chH nGA. + rewrite centsC coprime_abelian_faithful_Ohm1 ?(coprimeSg sHG) //. + by rewrite centsC (subset_trans (OhmS 1 sHG)). +have abelZ: p.-abelem 'Z(G) by exact: center_special_abelem. +have cAZ: {in 'Z(G), centralised A} by apply/centsP; rewrite -defC subsetIr. +have cGZ: {in 'Z(G), centralised G} by apply/centsP; rewrite subsetIr. +have defG1: 'Ohm_1(G) = 'Z(G). + apply/eqP; rewrite eqEsubset -{1}defC subsetI Ohm_sub cAG1 /=. + by rewrite -(Ohm1_id abelZ) OhmS ?center_sub. +rewrite (subset_trans _ (subsetIr G _)) // defC -defG1 -{1}defR gen_subG /=. +apply/subsetP=> xa; case/imset2P=> x a Gx Aa ->{xa}; rewrite commgEl. +set u := x^-1; set v := x ^ a; pose w := [~ v, u]. +have [Gu Gv]: u \in G /\ v \in G by rewrite groupV memJ_norm ?(subsetP nGA). +have Zw: w \in 'Z(G) by rewrite -defG' mem_commg. +rewrite (OhmE 1 pG) mem_gen // !inE expn1 groupM //=. +rewrite expMg_Rmul /commute ?(cGZ w) // bin2odd // expgM. +case/(abelemP p_pr): abelZ => _ /(_ w)-> //. +rewrite expg1n mulg1 expgVn -conjXg (sameP commgP eqP) cAZ // -defPhi. +by rewrite (Phi_joing pG) joingC mem_gen // inE (Mho_p_elt 1) ?(mem_p_elt pG). +Qed. + +(* This is B & G, Corollary 1.12. *) +Corollary coprime_odd_faithful_cent_abelem gT p (A G E : {group gT}) : + E \in 'E_p(G) -> p.-group G -> + A \subset 'N(G) -> coprime #|G| #|A| -> odd #|G| -> + A \subset 'C('Ldiv_p('C_G(E))) -> A \subset 'C(G). +Proof. +case/pElemP=> sEG abelE pG nGA coGA oddG cCEA. +have [-> | ntG] := eqsVneq G 1; first by rewrite cents1. +have [p_pr _ _] := pgroup_pdiv pG ntG. +have{cCEA} cCEA: A \subset 'C('Ohm_1('C_G(E))). + by rewrite (OhmE 1 (pgroupS _ pG)) ?subsetIl ?cent_gen. +apply: coprime_nil_faithful_cent_stab (pgroup_nil pG) _ => //. +rewrite subsetI subsetIl centsC /=; set CC := 'C_G(_). +have sCCG: CC \subset G := subsetIl _ _; have pCC := pgroupS sCCG pG. +rewrite (coprime_odd_faithful_Ohm1 pCC) ?(coprimeSg sCCG) ?(oddSg sCCG) //. + by rewrite !(normsI, norms_cent, normG). +rewrite (subset_trans cCEA) // centS // OhmS // setIS // centS //. +rewrite subsetI sEG /= centsC (subset_trans cCEA) // centS //. +have cEE: abelian E := abelem_abelian abelE. +by rewrite -{1}(Ohm1_id abelE) OhmS // subsetI sEG. +Qed. + +(* This is B & G, Theorem 1.13. *) +Theorem critical_odd gT p (G : {group gT}) : + p.-group G -> odd #|G| -> G :!=: 1 -> + {H : {group gT} | + [/\ H \char G, [~: H, G] \subset 'Z(H), nil_class H <= 2, exponent H = p + & p.-group 'C(H | [Aut G])]}. +Proof. +move=> pG oddG ntG; have [H krH]:= Thompson_critical pG. +have [chH sPhiZ sGH_Z scH] := krH; have clH := critical_class2 krH. +have sHG := char_sub chH; set D := 'Ohm_1(H)%G; exists D. +have chD: D \char G := char_trans (Ohm_char 1 H) chH. +have sDH: D \subset H := Ohm_sub 1 H. +have sDG_Z: [~: D, G] \subset 'Z(D). + rewrite subsetI commg_subl char_norm // commGC. + apply: subset_trans (subset_trans sGH_Z _); first by rewrite commgS. + by rewrite subIset // orbC centS. +rewrite nil_class2 !(subset_trans (commgS D _) sDG_Z) ?(char_sub chD) {sDH}//. +have [p_pr p_dv_G _] := pgroup_pdiv pG ntG; have odd_p := dvdn_odd p_dv_G oddG. +split=> {chD sDG_Z}//. + apply/prime_nt_dvdP=> //; last by rewrite exponent_Ohm1_class2 ?(pgroupS sHG). + rewrite -dvdn1 -trivg_exponent /= Ohm1_eq1; apply: contraNneq ntG => H1. + by rewrite -(setIidPl (cents1 G)) -{1}H1 scH H1 center1. +apply/pgroupP=> q q_pr /Cauchy[] //= f. +rewrite astab_ract => /setIdP[Af cDf] ofq; apply: wlog_neg => p'q. +suffices: f \in 'C(H | [Aut G]). + move/(mem_p_elt (critical_p_stab_Aut krH pG))/pnatP=> -> //. + by rewrite ofq. +rewrite astab_ract inE Af; apply/astabP=> x Hx; rewrite /= /aperm /=. +rewrite nil_class2 in clH; have pH := pgroupS sHG pG. +have /p_natP[i ox]: p.-elt x by apply: mem_p_elt Hx. +have{ox}: x ^+ (p ^ i) = 1 by rewrite -ox expg_order. +elim: i x Hx => [|[|i] IHi] x Hx xp1. +- by rewrite [x]xp1 -(autmE Af) morph1. +- by apply: (astabP cDf); rewrite (OhmE 1 pH) mem_gen // !inE Hx xp1 eqxx. +have expH': {in H &, forall y z, [~ y, z] ^+ p = 1}. + move=> y z Hy Hz; apply/eqP. + have /setIP[_ cHyz]: [~ y, z] \in 'Z(H) by rewrite (subsetP clH) // mem_commg. + rewrite -commXg; last exact/commute_sym/(centP cHyz). + suffices /setIP[_ cHyp]: y ^+ p \in 'Z(H) by exact/commgP/(centP cHyp). + rewrite (subsetP sPhiZ) // (Phi_joing pH) mem_gen // inE orbC. + by rewrite (Mho_p_elt 1) ?(mem_p_elt pH). +have Hfx: f x \in H. + case/charP: chH => _ /(_ _ (injm_autm Af) (im_autm Af)) <-. + by rewrite -{1}(autmE Af) mem_morphim // (subsetP sHG). +set y := x^-1 * f x; set z := [~ f x, x^-1]. +have Hy: y \in H by rewrite groupM ?groupV. +have /centerP[_ Zz]: z \in 'Z(H) by rewrite (subsetP clH) // mem_commg ?groupV. +have fy: f y = y. + apply: (IHi); first by rewrite groupM ?groupV. + rewrite expMg_Rmul; try by apply: commute_sym; apply: Zz; rewrite ?groupV. + rewrite -/z bin2odd ?odd_exp // {3}expnS -mulnA expgM expH' ?groupV //. + rewrite expg1n mulg1 expgVn -(autmE Af) -morphX ?(subsetP sHG) //= autmE. + rewrite IHi ?mulVg ?groupX // {2}expnS expgM -(expgM x _ p) -expnSr. + by rewrite xp1 expg1n. +have /eqP: (f ^+ q) x = x * y ^+ q. + elim: (q) => [|j IHj]; first by rewrite perm1 mulg1. + rewrite expgSr permM {}IHj -(autmE Af). + rewrite morphM ?morphX ?groupX ?(subsetP sHG) //= autmE. + by rewrite fy expgS mulgA mulKVg. +rewrite -{1}ofq expg_order perm1 eq_mulVg1 mulKg -order_dvdn. +case: (primeP q_pr) => _ dv_q /dv_q; rewrite order_eq1 -eq_mulVg1. +case/pred2P=> // oyq; case/negP: p'q. +by apply: (pgroupP pH); rewrite // -oyq order_dvdG. +Qed. + +Section CoprimeQuotientPgroup. + +(* This is B & G, Lemma 1.14, which we divide in four lemmas, each one giving *) +(* the (sub)centraliser or (sub)normaliser of a quotient by a coprime p-group *) +(* acting on it. Note that we weaken the assumptions of B & G -- M does not *) +(* need to be normal in G, T need not be a subgroup of G, p need not be a *) +(* prime, and M only needs to be coprime with T. Note also that the subcenter *) +(* quotient lemma is special case of a lemma in coprime_act. *) + +Variables (gT : finGroupType) (p : nat) (T M G : {group gT}). +Hypothesis pT : p.-group T. +Hypotheses (nMT : T \subset 'N(M)) (coMT : coprime #|M| #|T|). + +(* This is B & G, Lemma 1.14, for a global normaliser. *) +Lemma coprime_norm_quotient_pgroup : 'N(T / M) = 'N(T) / M. +Proof. +have [-> | ntT] := eqsVneq T 1; first by rewrite quotient1 !norm1 quotientT. +have [p_pr _ [m oMpm]] := pgroup_pdiv pT ntT. +apply/eqP; rewrite eqEsubset morphim_norms // andbT; apply/subsetP=> Mx. +case: (cosetP Mx) => x Nx ->{Mx} nTqMx. +have sylT: p.-Sylow(M <*> T) T. + rewrite /pHall pT -divgS joing_subr //= norm_joinEr ?coprime_cardMg //. + rewrite mulnK // ?p'natE -?prime_coprime // coprime_sym. + by rewrite -(@coprime_pexpr m.+1) -?oMpm. +have sylTx: p.-Sylow(M <*> T) (T :^ x). + have nMTx: x \in 'N(M <*> T). + rewrite norm_joinEr // inE -quotientSK ?conj_subG ?mul_subG ?normG //. + by rewrite quotientJ // quotientMidl (normP nTqMx). + by rewrite pHallE /= -{1}(normP nMTx) conjSg cardJg -pHallE. +have{sylT sylTx} [ay] := Sylow_trans sylT sylTx. +rewrite /= joingC norm_joinEl //; case/imset2P=> a y Ta. +rewrite -groupV => My ->{ay} defTx; rewrite -(coset_kerr x My). +rewrite mem_morphim //; first by rewrite groupM // (subsetP (normG M)). +by rewrite inE !(conjsgM, defTx) conjsgK conjGid. +Qed. + +(* This is B & G, Lemma 1.14, for a global centraliser. *) +Lemma coprime_cent_quotient_pgroup : 'C(T / M) = 'C(T) / M. +Proof. +symmetry; rewrite -quotientInorm -quotientMidl -['C(T / M)]cosetpreK. +congr (_ / M); set Cq := _ @*^-1 _; set C := 'N_('C(T))(M). +suffices <-: 'N_Cq(T) = C. + rewrite setIC group_modl ?sub_cosetpre //= -/Cq; apply/setIidPr. + rewrite -quotientSK ?subsetIl // cosetpreK. + by rewrite -coprime_norm_quotient_pgroup cent_sub. +apply/eqP; rewrite eqEsubset subsetI -sub_quotient_pre ?subsetIr //. +rewrite quotientInorm quotient_cents //= andbC subIset ?cent_sub //=. +have nMC': 'N_Cq(T) \subset 'N(M) by rewrite subIset ?subsetIl. +rewrite subsetI nMC' andbT (sameP commG1P trivgP) /=. +rewrite -(coprime_TIg coMT) subsetI commg_subr subsetIr andbT. +by rewrite -quotient_cents2 ?sub_quotient_pre ?subsetIl. +Qed. + +Hypothesis sMG : M \subset G. + +(* This is B & G, Lemma 1.14, for a local normaliser. *) +Lemma coprime_subnorm_quotient_pgroup : 'N_(G / M)(T / M) = 'N_G(T) / M. +Proof. by rewrite quotientGI -?coprime_norm_quotient_pgroup. Qed. + +(* This is B & G, Lemma 1.14, for a local centraliser. *) +Lemma coprime_subcent_quotient_pgroup : 'C_(G / M)(T / M) = 'C_G(T) / M. +Proof. by rewrite quotientGI -?coprime_cent_quotient_pgroup. Qed. + +End CoprimeQuotientPgroup. + +Section Constrained. + +Variables (gT : finGroupType) (p : nat) (G : {group gT}). + +(* This is B & G, Proposition 1.15a (Lemma 1.2.3 of P. Hall & G. Higman). *) +Proposition solvable_p_constrained : solvable G -> p.-constrained G. +Proof. +move=> solG P sylP; have [sPO pP _] := and3P sylP; pose K := 'O_p^'(G). +have nKG: G \subset 'N(K) by rewrite normal_norm ?pcore_normal. +have nKC: 'C_G(P) \subset 'N(K) by rewrite subIset ?nKG. +rewrite -(quotientSGK nKC) //; last first. + by rewrite /= -pseries1 (pseries_sub_catl [::_]). +apply: subset_trans (quotient_subcent _ _ _) _ ;rewrite /= -/K. +suffices ->: P / K = 'O_p(G / K). + rewrite quotient_pseries2 -Fitting_eq_pcore ?trivg_pcore_quotient // -/K. + by rewrite cent_sub_Fitting ?morphim_sol. +apply/eqP; rewrite eqEcard -(part_pnat_id (pcore_pgroup _ _)). +have sylPK: p.-Sylow('O_p(G / K)) (P / K). + rewrite -quotient_pseries2 morphim_pHall //. + exact: subset_trans (subset_trans sPO (pseries_sub _ _)) nKG. +by rewrite -(card_Hall sylPK) leqnn -quotient_pseries2 quotientS. +Qed. + +(* This is Gorenstein, Proposition 8.1.3. *) +Proposition p_stable_abelian_constrained : + p.-constrained G -> p.-stable G -> p.-abelian_constrained G. +Proof. +move=> constrG stabG P A sylP cAA /andP[sAP nAP]. +have [sPG pP _] := and3P sylP; have sAG := subset_trans sAP sPG. +set K2 := 'O_{p^', p}(G); pose K1 := 'O_p^'(G); pose Q := P :&: K2. +have sQG: Q \subset G by rewrite subIset ?sPG. +have nK1G: G \subset 'N(K1) by rewrite normal_norm ?pcore_normal. +have nsK2G: K2 <| G := pseries_normal _ _; have [sK2G nK2G] := andP nsK2G. +have sylQ: p.-Sylow(K2) Q by rewrite /Q setIC (Sylow_setI_normal nsK2G). +have defK2: K1 * Q = K2. + have sK12: K1 \subset K2 by rewrite /K1 -pseries1 (pseries_sub_catl [::_]). + apply/eqP; rewrite eqEsubset mulG_subG /= sK12 subsetIr /=. + rewrite -quotientSK ?(subset_trans sK2G) //= quotientIG //= -/K1 -/K2. + rewrite subsetI subxx andbT quotient_pseries2. + by rewrite pcore_sub_Hall // morphim_pHall // ?(subset_trans sPG). +have{cAA} rQAA_1: [~: Q, A, A] = 1. + by apply/commG1P; apply: subset_trans cAA; rewrite commg_subr subIset // nAP. +have nK2A := subset_trans sAG nK2G. +have sAN: A \subset 'N_G(Q) by rewrite subsetI sAG normsI // normsG. +have{stabG rQAA_1 defK2 sQG} stabA: A / 'C_G(Q) \subset 'O_p('N_G(Q) / 'C_G(Q)). + apply: stabG; rewrite //= /psubgroup -/Q ?sAN ?(pgroupS _ pP) ?subsetIl //. + by rewrite defK2 pseries_normal. +rewrite -quotient_sub1 //= -/K2 -(setIidPr sAN). +have nK2N: 'N_G(Q) \subset 'N(K2) by rewrite subIset ?nK2G. +rewrite -[_ / _](morphim_restrm nK2N); set qK2 := restrm _ _. +have{constrG} fqKp: 'ker (coset 'C_G(Q)) \subset 'ker qK2. + by rewrite ker_restrm !ker_coset subsetI subcent_sub constrG. +rewrite -(morphim_factm fqKp (subcent_norm _ _)) -(quotientE A _). +apply: subset_trans {stabA}(morphimS _ stabA) _. +apply: subset_trans (morphim_pcore _ _ _) _. +rewrite morphim_factm morphim_restrm setIid -quotientE. +rewrite /= -quotientMidl /= -/K2 (Frattini_arg _ sylQ) ?pseries_normal //. +by rewrite -quotient_pseries //= (pseries_rcons_id [::_]) trivg_quotient. +Qed. + +End Constrained. + +(* This is B & G, Proposition 1.15b (due to D. Goldschmith). *) +Proposition p'core_cent_pgroup gT p (G R : {group gT}) : + p.-subgroup(G) R -> solvable G -> 'O_p^'('C_G(R)) \subset 'O_p^'(G). +Proof. +case/andP=> sRG pR solG. +without loss p'G1: gT G R sRG pR solG / 'O_p^'(G) = 1. + have nOG_CR: 'C_G(R) \subset 'N('O_p^'(G)) by rewrite subIset ?gFnorm. + move=> IH; rewrite -quotient_sub1 ?(subset_trans (pcore_sub _ _)) //. + apply: subset_trans (morphimF _ _ nOG_CR) _; rewrite /= -quotientE. + rewrite -(coprime_subcent_quotient_pgroup pR) ?pcore_sub //; first 1 last. + - by rewrite (subset_trans sRG) ?gFnorm. + - by rewrite coprime_sym (pnat_coprime _ (pcore_pgroup _ _)). + have p'Gq1 : 'O_p^'(G / 'O_p^'(G)) = 1 := trivg_pcore_quotient p^' G. + by rewrite -p'Gq1 IH ?morphimS ?morphim_pgroup ?morphim_sol. +set M := 'O_p^'('C_G(R)); pose T := 'O_p(G). +have /subsetIP[sMG cMR]: M \subset 'C_G(R) by exact: pcore_sub. +have [p'M pT]: p^'.-group M /\ p.-group T by rewrite !pcore_pgroup. +have nRT: R \subset 'N(T) by rewrite (subset_trans sRG) ?gFnorm. +have pRT: p.-group (R <*> T). + rewrite -(pquotient_pgroup pT) ?join_subG ?nRT ?normG //=. + by rewrite norm_joinEl // quotientMidr morphim_pgroup. +have nRT_M: M \subset 'N(R <*> T). + by rewrite normsY ?(cents_norm cMR) // (subset_trans sMG) ?gFnorm. +have coRT_M: coprime #|R <*> T| #|M| := pnat_coprime pRT p'M. +have cMcR: 'C_(R <*> T)(R) \subset 'C(M). + apply/commG1P; apply/trivgP; rewrite -(coprime_TIg coRT_M) subsetI commg_subr. + rewrite (subset_trans (commSg _ (subsetIl _ _))) ?commg_subl //= -/M. + by apply: subset_trans (gFnorm _ _); rewrite setSI // join_subG sRG pcore_sub. +have cRT_M: M \subset 'C(R <*> T). + rewrite coprime_nil_faithful_cent_stab ?(pgroup_nil pRT) //= -/M. + rewrite subsetI subsetIl (subset_trans _ cMcR) // ?setIS ?centS //. + by rewrite subsetI joing_subl centsC. +have sMT: M \subset T. + have defT: 'F(G) = T := Fitting_eq_pcore p'G1. + rewrite -defT (subset_trans _ (cent_sub_Fitting solG)) // defT subsetI sMG. + by rewrite (subset_trans cRT_M) // centY subsetIr. +by rewrite -(setIidPr sMT) p'G1 coprime_TIg // (pnat_coprime pT). +Qed. + +(* This is B & G, Proposition 1.16, second assertion. Contrary to the text, *) +(* we derive this directly, rather than by induction on the first, because *) +(* this is actually how the proof is done in Gorenstein. Note that the non *) +(* cyclic assumption for A is not needed here. *) +Proposition coprime_abelian_gen_cent gT (A G : {group gT}) : + abelian A -> A \subset 'N(G) -> coprime #|G| #|A| -> + <<\bigcup_(B : {group gT} | cyclic (A / B) && (B <| A)) 'C_G(B)>> = G. +Proof. +move=> abelA nGA coGA; symmetry; move: {2}_.+1 (ltnSn #|G|) => n. +elim: n gT => // n IHn gT in A G abelA nGA coGA *; rewrite ltnS => leGn. +without loss nilG: G nGA coGA leGn / nilpotent G. + move=> {IHn} IHn; apply/eqP; rewrite eqEsubset gen_subG. + apply/andP; split; last by apply/bigcupsP=> B _; exact: subsetIl. + pose T := [set P : {group gT} | Sylow G P & A \subset 'N(P)]. + rewrite -{1}(@Sylow_transversal_gen _ T G) => [|P | p _]; first 1 last. + - by rewrite inE -!andbA; case/and4P. + - have [//|P sylP nPA] := sol_coprime_Sylow_exists p (abelian_sol abelA) nGA. + by exists P; rewrite ?inE ?(p_Sylow sylP). + rewrite gen_subG; apply/bigcupsP=> P {T}/setIdP[/SylowP[p _ sylP] nPA]. + have [sPG pP _] := and3P sylP. + rewrite (IHn P) ?(pgroup_nil pP) ?(coprimeSg sPG) ?genS //. + by apply/bigcupsP=> B cycBq; rewrite (bigcup_max B) ?setSI. + by rewrite (leq_trans (subset_leq_card sPG)). +apply/eqP; rewrite eqEsubset gen_subG. +apply/andP; split; last by apply/bigcupsP=> B _; exact: subsetIl. +have [Z1 | ntZ] := eqsVneq 'Z(G) 1. + by rewrite (TI_center_nil _ (normal_refl G)) ?Z1 ?(setIidPr _) ?sub1G. +have nZA: A \subset 'N('Z(G)) := char_norm_trans (center_char G) nGA. +have{ntZ nZA} [M /= minM] := minnormal_exists ntZ nZA. +rewrite subsetI centsC => /andP[sMG /cents_norm nMG]. +have coMA := coprimeSg sMG coGA; have{nilG} solG := nilpotent_sol nilG. +have [nMA ntM abelM] := minnormal_solvable minM sMG solG. +set GC := <<_>>; have sMGC: M \subset GC. + rewrite sub_gen ?(bigcup_max 'C_A(M)%G) //=; last first. + by rewrite subsetI sMG centsC subsetIr. + case/is_abelemP: abelM => p _ abelM; rewrite -(rker_abelem abelM ntM nMA). + rewrite rker_normal -(setIidPl (quotient_abelian _ _)) ?center_kquo_cyclic //. + exact/abelem_mx_irrP. +rewrite -(quotientSGK nMG sMGC). +have: A / M \subset 'N(G / M) by rewrite morphim_norms. +move/IHn->; rewrite ?morphim_abelian ?coprime_morph {IHn}//; first 1 last. + by rewrite (leq_trans _ leGn) ?ltn_quotient. +rewrite gen_subG; apply/bigcupsP=> Bq; rewrite andbC => /andP[]. +have: M :&: A = 1 by rewrite coprime_TIg. +move/(quotient_isom nMA); case/isomP=> /=; set qM := restrm _ _ => injqM <-. +move=> nsBqA; have sBqA := normal_sub nsBqA. +rewrite -(morphpreK sBqA) /= -/qM; set B := qM @*^-1 Bq. +move: nsBqA; rewrite -(morphpre_normal sBqA) ?injmK //= -/B => nsBA. +rewrite -(morphim_quotm _ nsBA) /= -/B injm_cyclic ?injm_quotm //= => cycBA. +rewrite morphim_restrm -quotientE morphpreIdom -/B; have sBA := normal_sub nsBA. +rewrite -coprime_quotient_cent ?(coprimegS sBA, subset_trans sBA) //= -/B. +by rewrite quotientS ?sub_gen // (bigcup_max [group of B]) ?cycBA. +Qed. + +(* B & G, Proposition 1.16, first assertion. *) +Proposition coprime_abelian_gen_cent1 gT (A G : {group gT}) : + abelian A -> ~~ cyclic A -> A \subset 'N(G) -> coprime #|G| #|A| -> + <<\bigcup_(a in A^#) 'C_G[a]>> = G. +Proof. +move=> abelA ncycA nGA coGA. +apply/eqP; rewrite eq_sym eqEsubset /= gen_subG. +apply/andP; split; last by apply/bigcupsP=> B _; exact: subsetIl. +rewrite -{1}(coprime_abelian_gen_cent abelA nGA) ?genS //. +apply/bigcupsP=> B; have [-> | /trivgPn[a Ba n1a]] := eqsVneq B 1. + by rewrite injm_cyclic ?coset1_injm ?norms1 ?(negbTE ncycA). +case/and3P=> _ sBA _; rewrite (bigcup_max a) ?inE ?n1a ?(subsetP sBA) //. +by rewrite setIS // -cent_set1 centS // sub1set. +Qed. + +Section Focal_Subgroup. + +Variables (gT : finGroupType) (G S : {group gT}) (p : nat). +Hypothesis sylS : p.-Sylow(G) S. + +Import finalg FiniteModule GRing.Theory. + +(* This is B & G, Theorem 1.17 ("Focal Subgroup Theorem", D. G. Higman), also *) +(* Gorenstein Theorem 7.3.4 and Aschbacher (37.4). *) +Theorem focal_subgroup_gen : + S :&: G^`(1) = <<[set [~ x, u] | x in S, u in G & x ^ u \in S]>>. +Proof. +set K := <<_>>; set G' := G^`(1); have [sSG coSiSG] := andP (pHall_Hall sylS). +apply/eqP; rewrite eqEsubset gen_subG andbC; apply/andP; split. + apply/subsetP=> _ /imset2P[x u Sx /setIdP[Gu Sxu] ->]. + by rewrite inE groupM ?groupV // mem_commg // (subsetP sSG). +apply/subsetP=> g /setIP[Sg G'g]; have Gg := subsetP sSG g Sg. +have nKS: S \subset 'N(K). + rewrite norms_gen //; apply/subsetP=> y Sy; rewrite inE. + apply/subsetP=> _ /imsetP[_ /imset2P[x u Sx /setIdP[Gu Sxu] ->] ->]. + have Gy: y \in G := subsetP sSG y Sy. + by rewrite conjRg mem_imset2 ?groupJ // inE -conjJg /= 2?groupJ. +set alpha := restrm_morphism nKS (coset_morphism K). +have alphim: (alpha @* S) = (S / K) by rewrite morphim_restrm setIid. +have abelSK : abelian (alpha @* S). + rewrite alphim sub_der1_abelian // genS //. + apply/subsetP=> _ /imset2P[x y Sx Sy ->]. + by rewrite mem_imset2 // inE (subsetP sSG) ?groupJ. +set ker_trans := 'ker (transfer G abelSK). +have G'ker : G' \subset ker_trans. + rewrite gen_subG; apply/subsetP=> h; case/imset2P=> h1 h2 Gh1 Gh2 ->{h}. + by rewrite !inE groupR // morphR //; apply/commgP; exact: addrC. +have transg0: transfer G abelSK g = 0%R. + by move/kerP: (subsetP G'ker g G'g); exact. +have partX := rcosets_cycle_partition sSG Gg. +have trX := transversalP partX; set X := transversal _ _ in trX. +have /and3P[_ sXG _] := trX. +have gGSeq0: (fmod abelSK (alpha g) *+ #|G : S| = 0)%R. + rewrite -transg0 (transfer_cycle_expansion sSG abelSK Gg trX). + rewrite -(sum_index_rcosets_cycle sSG Gg trX) -sumrMnr /restrm. + apply: eq_bigr=> x Xx; rewrite -[(_ *+ _)%R]morphX ?mem_morphim //=. + rewrite -morphX //= /restrm; congr fmod. + apply/rcoset_kercosetP; rewrite /= -/K. + - by rewrite (subsetP nKS) ?groupX. + - rewrite (subsetP nKS) // conjgE invgK mulgA -mem_rcoset. + exact: mulg_exp_card_rcosets. + rewrite mem_rcoset -{1}[g ^+ _]invgK -conjVg -commgEl mem_gen ?mem_imset2 //. + by rewrite groupV groupX. + rewrite inE conjVg !groupV (subsetP sXG) //= conjgE invgK mulgA -mem_rcoset. + exact: mulg_exp_card_rcosets. +move: (congr_fmod gGSeq0). +rewrite fmval0 morphX ?inE //= fmodK ?mem_morphim // /restrm /=. +move/((congr1 (expgn^~ (expg_invn (S / K) #|G : S|))) _). +rewrite expg1n expgK ?mem_quotient ?coprime_morphl // => Kg1. +by rewrite coset_idr ?(subsetP nKS). +Qed. + +(* This is B & G, Theorem 1.18 (due to Burnside). *) +Theorem Burnside_normal_complement : + 'N_G(S) \subset 'C(S) -> 'O_p^'(G) ><| S = G. +Proof. +move=> cSN; set K := 'O_p^'(G); have [sSG pS _] := and3P sylS. +have /andP[sKG nKG]: K <| G by exact: pcore_normal. +have{nKG} nKS := subset_trans sSG nKG. +have p'K: p^'.-group K by exact: pcore_pgroup. +have{pS p'K} tiKS: K :&: S = 1 by rewrite setIC coprime_TIg ?(pnat_coprime pS). +suffices{tiKS nKS} hallK: p^'.-Hall(G) K. + rewrite sdprodE //= -/K; apply/eqP; rewrite eqEcard ?mul_subG //=. + by rewrite TI_cardMg //= (card_Hall sylS) (card_Hall hallK) mulnC partnC. +pose G' := G^`(1); have nsG'G : G' <| G by rewrite der_normalS. +suffices{K sKG} p'G': p^'.-group G'. + have nsG'K: G' <| K by rewrite (normalS _ sKG) ?pcore_max. + rewrite -(pquotient_pHall p'G') -?pquotient_pcore //= -/G'. + by rewrite nilpotent_pcore_Hall ?abelian_nil ?der_abelian. +suffices{nsG'G} tiSG': S :&: G' = 1. + have sylG'S : p.-Sylow(G') (G' :&: S) by rewrite (Sylow_setI_normal _ sylS). + rewrite /pgroup -[#|_|](partnC p) ?cardG_gt0 // -{sylG'S}(card_Hall sylG'S). + by rewrite /= setIC tiSG' cards1 mul1n part_pnat. +apply/trivgP; rewrite /= focal_subgroup_gen ?(p_Sylow sylS) // gen_subG. +apply/subsetP=> _ /imset2P[x u Sx /setIdP[Gu Sxu] ->]. +have cSS y: y \in S -> S \subset 'C_G[y]. + rewrite subsetI sSG -cent_set1 centsC sub1set; apply: subsetP. + by apply: subset_trans cSN; rewrite subsetI sSG normG. +have{cSS} [v]: exists2 v, v \in 'C_G[x ^ u | 'J] & S :=: (S :^ u) :^ v. + have sylSu : p.-Sylow(G) (S :^ u) by rewrite pHallJ. + have [sSC sCG] := (cSS _ Sxu, subsetIl G 'C[x ^ u]). + rewrite astab1J; apply: (@Sylow_trans p); apply: pHall_subl sCG _ => //=. + by rewrite -conjg_set1 normJ -(conjGid Gu) -conjIg conjSg cSS. +rewrite in_set1 -conjsgM => /setIP[Gv /astab1P cx_uv] nSuv. +apply/conjg_fixP; rewrite -cx_uv /= -conjgM; apply: astabP Sx. +by rewrite astabJ (subsetP cSN) // !inE -nSuv groupM /=. +Qed. + +(* This is B & G, Corollary 1.19(a). *) +Corollary cyclic_Sylow_tiVsub_der1 : + cyclic S -> S :&: G^`(1) = 1 \/ S \subset G^`(1). +Proof. +move=> cycS; have [sSG pS _] := and3P sylS. +have nsSN: S <| 'N_G(S) by rewrite normalSG. +have hallSN: Hall 'N_G(S) S. + by apply: pHall_Hall (pHall_subl _ _ sylS); rewrite ?subsetIl ?normal_sub. +have /splitsP[K /complP[tiSK /= defN]] := SchurZassenhaus_split hallSN nsSN. +have sKN: K \subset 'N_G(S) by rewrite -defN mulG_subr. +have [sKG nSK] := subsetIP sKN. +have coSK: coprime #|S| #|K|. + by case/andP: hallSN => sSN; rewrite -divgS //= -defN TI_cardMg ?mulKn. +have:= coprime_abelian_cent_dprod nSK coSK (cyclic_abelian cycS). +case/(cyclic_pgroup_dprod_trivg pS cycS) => [[_ cSK]|[_ <-]]; last first. + by right; rewrite commgSS. +have cSN: 'N_G(S) \subset 'C(S). + by rewrite -defN mulG_subG -abelianE cyclic_abelian // centsC -cSK subsetIr. +have /sdprodP[_ /= defG _ _] := Burnside_normal_complement cSN. +set Q := 'O_p^'(G) in defG; have nQG: G \subset 'N(Q) := gFnorm _ _. +left; rewrite coprime_TIg ?(pnat_coprime pS) //. +apply: pgroupS (pcore_pgroup _ G); rewrite /= -/Q. +rewrite -quotient_sub1 ?(subset_trans (der_sub _ _)) ?quotientR //= -/Q. +rewrite -defG quotientMidl (sameP trivgP commG1P) -abelianE. +by rewrite morphim_abelian ?cyclic_abelian. +Qed. + +End Focal_Subgroup. + +(* This is B & G, Corollary 1.19(b). *) +Corollary Zgroup_der1_Hall gT (G : {group gT}) : + Zgroup G -> Hall G G^`(1). +Proof. +move=> ZgG; set G' := G^`(1). +rewrite /Hall der_sub coprime_sym coprime_pi' ?cardG_gt0 //=. +apply/pgroupP=> p p_pr pG'; have [P sylP] := Sylow_exists p G. +have cycP: cyclic P by have:= forallP ZgG P; rewrite (p_Sylow sylP). +case: (cyclic_Sylow_tiVsub_der1 sylP cycP) => [tiPG' | sPG']. + have: p.-Sylow(G') (P :&: G'). + by rewrite setIC (Sylow_setI_normal _ sylP) ?gFnormal. + move/card_Hall/eqP; rewrite /= tiPG' cards1 eq_sym. + by rewrite partn_eq1 ?cardG_gt0 // p'natE ?pG'. +rewrite inE /= mem_primes p_pr indexg_gt0 -?p'natE // -partn_eq1 //. +have sylPq: p.-Sylow(G / G') (P / G') by rewrite morphim_pHall ?normsG. +rewrite -card_quotient ?gFnorm // -(card_Hall sylPq) -trivg_card1. +by rewrite /= -quotientMidr mulSGid ?trivg_quotient. +Qed. + +(* This is Aschbacher (39.2). *) +Lemma cyclic_pdiv_normal_complement gT (S G : {group gT}) : + (pdiv #|G|).-Sylow(G) S -> cyclic S -> exists H : {group gT}, H ><| S = G. +Proof. +set p := pdiv _ => sylS cycS; have cSS := cyclic_abelian cycS. +exists 'O_p^'(G)%G; apply: Burnside_normal_complement => //. +have [-> | ntS] := eqsVneq S 1; first exact: cents1. +have [sSG pS p'iSG] := and3P sylS; have [pr_p _ _] := pgroup_pdiv pS ntS. +rewrite -['C(S)]mulg1 -ker_conj_aut -morphimSK ?subsetIr // setIC morphimIdom. +set A_G := _ @* _; pose A := Aut S. +have [_ [_ [cAA _ oAp' _]] _] := cyclic_pgroup_Aut_structure pS cycS ntS. +have{cAA cSS p'iSG} /setIidPl <-: A_G \subset 'O_p^'(A). + rewrite pcore_max -?sub_abelian_normal ?Aut_conj_aut //=. + apply: pnat_dvd p'iSG; rewrite card_morphim ker_conj_aut /= setIC. + have sSN: S \subset 'N_G(S) by rewrite subsetI sSG normG. + by apply: dvdn_trans (indexSg sSN (subsetIl G 'N(S))); apply: indexgS. +rewrite coprime_TIg ?sub1G // coprime_morphl // coprime_sym coprime_pi' //. +apply/pgroupP=> q pr_q q_dv_G; rewrite !inE mem_primes gtnNdvd ?andbF // oAp'. +by rewrite prednK ?prime_gt0 ?pdiv_min_dvd ?prime_gt1. +Qed. + +(* This is Aschbacher (39.3). *) +Lemma Zgroup_metacyclic gT (G : {group gT}) : Zgroup G -> metacyclic G. +Proof. +elim: {G}_.+1 {-2}G (ltnSn #|G|) => // n IHn G; rewrite ltnS => leGn ZgG. +have{n IHn leGn} solG: solvable G. + have [-> | ntG] := eqsVneq G 1; first exact: solvable1. + have [S sylS] := Sylow_exists (pdiv #|G|) G. + have cycS: cyclic S := forall_inP ZgG S (p_Sylow sylS). + have [H defG] := cyclic_pdiv_normal_complement sylS cycS. + have [nsHG _ _ _ _] := sdprod_context defG; rewrite (series_sol nsHG) andbC. + rewrite -(isog_sol (sdprod_isog defG)) (abelian_sol (cyclic_abelian cycS)). + rewrite metacyclic_sol ?IHn ?(ZgroupS _ ZgG) ?normal_sub //. + rewrite (leq_trans _ leGn) // -(sdprod_card defG) ltn_Pmulr // cardG_gt1. + by rewrite -rank_gt0 (rank_Sylow sylS) p_rank_gt0 pi_pdiv cardG_gt1. +pose K := 'F(G)%G; apply/metacyclicP; exists K. +have nsKG: K <| G := Fitting_normal G; have [sKG nKG] := andP nsKG. +have cycK: cyclic K by rewrite nil_Zgroup_cyclic ?Fitting_nil ?(ZgroupS sKG). +have cKK: abelian K := cyclic_abelian cycK. +have{solG cKK} defK: 'C_G(K) = K. + by apply/setP/subset_eqP; rewrite cent_sub_Fitting // subsetI sKG. +rewrite cycK nil_Zgroup_cyclic ?morphim_Zgroup ?abelian_nil //. +rewrite -defK -ker_conj_aut (isog_abelian (first_isog_loc _ _)) //. +exact: abelianS (Aut_conj_aut K G) (Aut_cyclic_abelian cycK). +Qed. + +(* This is B & G, Theorem 1.20 (Maschke's Theorem) for internal action on *) +(* elementary abelian subgroups; a more general case, for linear *) +(* represenations on matrices, can be found in mxrepresentation.v. *) +Theorem Maschke_abelem gT p (G V U : {group gT}) : + p.-abelem V -> p^'.-group G -> U \subset V -> + G \subset 'N(V) -> G \subset 'N(U) -> + exists2 W : {group gT}, U \x W = V & G \subset 'N(W). +Proof. +move=> pV p'G sUV nVG nUG. +have splitU: [splits V, over U] := abelem_splits pV sUV. +case/and3P: pV => pV abV; have cUV := subset_trans sUV abV. +have sVVG := joing_subl V G. +have{nUG} nUVG: U <| V <*> G. + by rewrite /(U <| _) join_subG (subset_trans sUV) // cents_norm // centsC. +rewrite -{nUVG}(Gaschutz_split nUVG) ?(abelianS sUV) // in splitU; last first. + rewrite -divgS ?joing_subl //= norm_joinEr //. + have coVG: coprime #|V| #|G| := pnat_coprime pV p'G. + by rewrite coprime_cardMg // mulnC mulnK // (coprimeSg sUV). +case/splitsP: splitU => WG /complP[tiUWG /= defVG]. +exists (WG :&: V)%G. + rewrite dprodE; last by rewrite setIA tiUWG (setIidPl _) ?sub1G. + by rewrite group_modl // defVG (setIidPr _). + by rewrite subIset // orbC centsC cUV. +rewrite (subset_trans (joing_subr V _)) // -defVG mul_subG //. + by rewrite cents_norm ?(subset_trans cUV) ?centS ?subsetIr. +rewrite normsI ?normG // (subset_trans (mulG_subr U _)) //. +by rewrite defVG join_subG normG. +Qed. + +Section Plength1. + +Variables (gT : finGroupType) (p : nat). +Implicit Types G H : {group gT}. + +(* Some basic properties of p.-length_1 that are direct consequences of their *) +(* definition using p-series. *) + +Lemma plength1_1 : p.-length_1 (1 : {set gT}). +Proof. by rewrite -[_ 1]subG1 pseries_sub. Qed. + +Lemma plength1_p'group G : p^'.-group G -> p.-length_1 G. +Proof. +move=> p'G; rewrite [p.-length_1 G]eqEsubset pseries_sub /=. +by rewrite -{1}(pcore_pgroup_id p'G) -pseries1 pseries_sub_catl. +Qed. + +Lemma plength1_nonprime G : ~~ prime p -> p.-length_1 G. +Proof. +move=> not_p_pr; rewrite plength1_p'group // p'groupEpi mem_primes. +by rewrite (negPf not_p_pr). +Qed. + +Lemma plength1_pcore_quo_Sylow G (Gb := G / 'O_p^'(G)) : + p.-length_1 G = p.-Sylow(Gb) 'O_p(Gb). +Proof. +rewrite /plength_1 eqEsubset pseries_sub /=. +rewrite (pseries_rcons _ [:: _; _]) -sub_quotient_pre ?gFnorm //=. +rewrite /pHall pcore_sub pcore_pgroup /= -card_quotient ?gFnorm //=. +rewrite -quotient_pseries2 /= {}/Gb -(pseries1 _ G). +rewrite (card_isog (third_isog _ _ _)) ?pseries_normal ?pseries_sub_catl //. +apply/idP/idP=> p'Gbb; last by rewrite (pcore_pgroup_id p'Gbb). +exact: pgroupS p'Gbb (pcore_pgroup _ _). +Qed. + +Lemma plength1_pcore_Sylow G : + 'O_p^'(G) = 1 -> p.-length_1 G = p.-Sylow(G) 'O_p(G). +Proof. +move=> p'G1; rewrite plength1_pcore_quo_Sylow -quotient_pseries2. +by rewrite p'G1 pseries_pop2 // pquotient_pHall ?normal1 ?pgroup1. +Qed. + +(* This is the characterization given in Section 10 of B & G, p. 75, just *) +(* before Theorem 10.6. *) +Lemma plength1_pseries2_quo G : p.-length_1 G = p^'.-group (G / 'O_{p^', p}(G)). +Proof. +rewrite /plength_1 eqEsubset pseries_sub lastI pseries_rcons /=. +rewrite -sub_quotient_pre ?gFnorm //. +by apply/idP/idP=> pl1G; rewrite ?pcore_pgroup_id ?(pgroupS pl1G) ?pcore_pgroup. +Qed. + +(* This is B & G, Lemma 1.21(a). *) +Lemma plength1S G H : H \subset G -> p.-length_1 G -> p.-length_1 H. +Proof. +rewrite /plength_1 => sHG pG1; rewrite eqEsubset pseries_sub. +by apply: subset_trans (pseriesS _ sHG); rewrite (eqP pG1) (setIidPr _). +Qed. + +Lemma plength1_quo G H : p.-length_1 G -> p.-length_1 (G / H). +Proof. +rewrite /plength_1 => pG1; rewrite eqEsubset pseries_sub. +by rewrite -{1}(eqP pG1) morphim_pseries. +Qed. + +(* This is B & G, Lemma 1.21(b). *) +Lemma p'quo_plength1 G H : + H <| G -> p^'.-group H -> p.-length_1 (G / H) = p.-length_1 G. +Proof. +rewrite /plength_1 => nHG p'H; apply/idP/idP; last exact: plength1_quo. +move=> pGH1; rewrite eqEsubset pseries_sub. +have nOG: 'O_{p^'}(G) <| G by exact: pseries_normal. +rewrite -(quotientSGK (normal_norm nOG)) ?(pseries_sub_catl [:: _]) //. +have [|f f_inj im_f] := third_isom _ nHG nOG. + by rewrite /= pseries1 pcore_max. +rewrite (quotient_pseries_cat [:: _]) -{}im_f //= -injmF //. +rewrite {f f_inj}morphimS // pseries1 -pquotient_pcore // -pseries1 /=. +by rewrite -quotient_pseries_cat /= (eqP pGH1). +Qed. + +(* This is B & G, Lemma 1.21(c). *) +Lemma pquo_plength1 G H : + H <| G -> p.-group H -> 'O_p^'(G / H) = 1-> + p.-length_1 (G / H) = p.-length_1 G. +Proof. +rewrite /plength_1 => nHG pH trO; apply/idP/idP; last exact: plength1_quo. +rewrite (pseries_pop _ trO) => pGH1; rewrite eqEsubset pseries_sub /=. +rewrite pseries_pop //; last first. + apply/eqP; rewrite -subG1; have <-: H :&: 'O_p^'(G) = 1. + by apply: coprime_TIg; exact: pnat_coprime (pcore_pgroup _ _). + rewrite setIC subsetI subxx -quotient_sub1. + by rewrite -trO morphim_pcore. + apply: subset_trans (normal_norm nHG); exact: pcore_sub. +have nOG: 'O_{p}(G) <| G by exact: pseries_normal. +rewrite -(quotientSGK (normal_norm nOG)) ?(pseries_sub_catl [:: _]) //. +have [|f f_inj im_f] := third_isom _ nHG nOG. + by rewrite /= pseries1 pcore_max. +rewrite (quotient_pseries [::_]) -{}im_f //= -injmF //. +rewrite {f f_inj}morphimS // pseries1 -pquotient_pcore // -(pseries1 p) /=. +by rewrite -quotient_pseries /= (eqP pGH1). +Qed. + +Canonical p_elt_gen_group A : {group gT} := + Eval hnf in [group of p_elt_gen p A]. + +(* Note that p_elt_gen could be a functor. *) +Lemma p_elt_gen_normal G : p_elt_gen p G <| G. +Proof. +apply/normalP; split=> [|x Gx]. + by rewrite gen_subG; apply/subsetP=> x; rewrite inE; case/andP. +rewrite -genJ; congr <<_>>; apply/setP=> y; rewrite mem_conjg !inE. +by rewrite p_eltJ -mem_conjg conjGid. +Qed. + +(* This is B & G, Lemma 1.21(d). *) +Lemma p_elt_gen_length1 G : + p.-length_1 G = p^'.-Hall(p_elt_gen p G) 'O_p^'(p_elt_gen p G). +Proof. +rewrite /pHall pcore_sub pcore_pgroup pnatNK /= /plength_1. +have nUG := p_elt_gen_normal G; have [sUG nnUG]:= andP nUG. +apply/idP/idP=> [p1G | pU]. + apply: (@pnat_dvd _ #|p_elt_gen p G : 'O_p^'(G)|). + by rewrite -[#|_ : 'O_p^'(G)|]indexgI indexgS ?pcoreS. + apply: (@pnat_dvd _ #|'O_p(G / 'O_{p^'}(G))|); last exact: pcore_pgroup. + rewrite -card_quotient; last first. + by rewrite (subset_trans sUG) // normal_norm ?pcore_normal. + rewrite -quotient_pseries pseries1 cardSg ?morphimS //=. + rewrite gen_subG; apply/subsetP=> x; rewrite inE; case/andP=> Gx p_x. + have nOx: x \in 'N('O_{p^',p}(G)). + by apply: subsetP Gx; rewrite normal_norm ?pseries_normal. + rewrite coset_idr //; apply/eqP; rewrite -[coset _ x]expg1 -order_dvdn. + rewrite [#[_]](@pnat_1 p) //; first exact: morph_p_elt. + apply: mem_p_elt (pcore_pgroup _ (G / _)) _. + by rewrite /= -quotient_pseries /= (eqP p1G); apply/morphimP; exists x. +have nOG: 'O_{p^', p}(G) <| G by exact: pseries_normal. +rewrite eqEsubset pseries_sub. +rewrite -(quotientSGK (normal_norm nOG)) ?(pseries_sub_catl [:: _; _]) //=. +rewrite (quotient_pseries [::_;_]) pcore_max //. +rewrite /pgroup card_quotient ?normal_norm //. +apply: (@pnat_dvd _ #|G : p_elt_gen p G|); last first. + case p_pr: (prime p); last by rewrite p'natEpi // mem_primes p_pr. + rewrite -card_quotient // p'natE //; apply/negP=> /Cauchy[] // Ux. + case/morphimP=> x Nx Gx -> /= oUx_p; have:= prime_gt1 p_pr. + rewrite -(part_pnat_id (pnat_id p_pr)) -{1}oUx_p {oUx_p} -order_constt. + rewrite -morph_constt //= coset_id ?order1 //. + by rewrite mem_gen // inE groupX // p_elt_constt. +apply: indexgS. +have nOU: p_elt_gen p G \subset 'N('O_{p^'}(G)). + by rewrite (subset_trans sUG) // normal_norm ?pseries_normal. +rewrite -(quotientSGK nOU) ?(pseries_sub_catl [:: _]) //=. +rewrite (quotient_pseries [::_]) pcore_max ?morphim_normal //. +rewrite /pgroup card_quotient //= pseries1; apply: pnat_dvd pU. +apply: indexgS; rewrite pcore_max ?pcore_pgroup //. +apply: char_normal_trans nUG; exact: pcore_char. +Qed. + +End Plength1. + +(* This is B & G, Lemma 1.21(e). *) +Lemma quo2_plength1 gT p (G H K : {group gT}) : + H <| G -> K <| G -> H :&: K = 1 -> + p.-length_1 (G / H) && p.-length_1 (G / K) = p.-length_1 G. +Proof. +move=> nHG nKG trHK. +have [p_pr | p_nonpr] := boolP (prime p); last by rewrite !plength1_nonprime. +apply/andP/idP=> [[pH1 pK1] | pG1]; last by rewrite !plength1_quo. +pose U := p_elt_gen p G; have nU : U <| G by exact: p_elt_gen_normal. +have exB (N : {group gT}) : + N <| G -> p.-length_1 (G / N) -> + exists B : {group gT}, + [/\ U \subset 'N(B), + forall x, x \in B -> #[x] = p -> x \in N + & forall Q : {group gT}, p^'.-subgroup(U) Q -> Q \subset B]. +- move=> nsNG; have [sNG nNG] := andP nsNG. + rewrite p_elt_gen_length1 // (_ : p_elt_gen _ _ = U / N); last first. + rewrite /quotient morphim_gen -?quotientE //; last first. + by rewrite setIdE subIset ?nNG. + congr <<_>>; apply/setP=> Nx; rewrite inE setIdE quotientGI // inE. + apply: andb_id2l => /morphimP[x NNx Gx ->{Nx}] /=. + apply/idP/idP=> [pNx | /morphimP[y NNy]]; last first. + by rewrite inE => p_y ->; exact: morph_p_elt. + rewrite -(constt_p_elt pNx) -morph_constt // mem_morphim ?groupX //. + by rewrite inE p_elt_constt. + have nNU: U \subset 'N(N) := subset_trans (normal_sub nU) nNG. + have nN_UN: U <*> N \subset 'N(N) by rewrite gen_subG subUset normG nNU. + case/(inv_quotientN _): (pcore_normal p^' [group of U <*> N / N]) => /= [|B]. + by rewrite /normal sub_gen ?subsetUr. + rewrite /= quotientYidr //= /U => defB sNB; case/andP=> sB nB hallB. + exists B; split=> [| x Ux p_x | Q /andP[sQU p'Q]]. + - by rewrite (subset_trans (sub_gen _) nB) ?subsetUl. + - have nNx: x \in 'N(N) by rewrite (subsetP nN_UN) ?(subsetP sB). + apply: coset_idr => //; rewrite -[coset N x](consttC p). + rewrite !(constt1P _) ?mulg1 // ?p_eltNK. + by rewrite morph_p_elt // /p_elt p_x pnat_id. + have: coset N x \in B / N by apply/morphimP; exists x. + by apply: mem_p_elt; rewrite /= -defB pcore_pgroup. + rewrite -(quotientSGK (subset_trans sQU nNU) sNB). + by rewrite -defB (sub_Hall_pcore hallB) ?quotientS ?quotient_pgroup. +have{pH1} [A [nAU pA p'A]] := exB H nHG pH1. +have{pK1 exB} [B [nBU pB p'B]] := exB K nKG pK1. +rewrite p_elt_gen_length1 //; apply: normal_max_pgroup_Hall (pcore_normal _ _). +apply/maxgroupP; split; first by rewrite /psubgroup pcore_sub pcore_pgroup. +move=> Q p'Q sOQ; apply/eqP; rewrite eqEsubset sOQ andbT. +apply: subset_trans (_ : U :&: (A :&: B) \subset _); last rewrite /U. + by rewrite !subsetI p'A ?p'B //; case/andP: p'Q => ->. +apply: pcore_max; last by rewrite /normal subsetIl !normsI ?normG. +rewrite /pgroup p'natE //. +apply/negP=> /Cauchy[] // x /setIP[_ /setIP[Ax Bx]] oxp. +suff: x \in 1%G by move/set1P=> x1; rewrite -oxp x1 order1 in p_pr. +by rewrite /= -trHK inE pA ?pB. +Qed. + +(* B & G Lemma 1.22 is covered by sylow.normal_pgroup. *) + +(* Encapsulation of the use of the order of GL_2(p), via abelem groups. *) +Lemma logn_quotient_cent_abelem gT p (A E : {group gT}) : + A \subset 'N(E) -> p.-abelem E -> logn p #|E| <= 2 -> + logn p #|A : 'C_A(E)| <= 1. +Proof. +move=> nEA abelE maxdimE; have [-> | ntE] := eqsVneq E 1. + by rewrite (setIidPl (cents1 _)) indexgg logn1. +pose rP := abelem_repr abelE ntE nEA. +have [p_pr _ _] := pgroup_pdiv (abelem_pgroup abelE) ntE. +have ->: 'C_A(E) = 'ker (reprGLm rP) by rewrite ker_reprGLm rker_abelem. +rewrite -card_quotient ?ker_norm // (card_isog (first_isog _)). +apply: leq_trans (dvdn_leq_log _ _ (cardSg (subsetT _))) _ => //. +rewrite logn_card_GL_p ?(dim_abelemE abelE) //. +by case: logn maxdimE; do 2?case. +Qed. + +End BGsection1. + +Section PuigSeriesGroups. + +Implicit Type gT : finGroupType. + +Canonical Puig_succ_group gT (D E : {set gT}) := [group of 'L_[D](E)]. + +Fact Puig_at_group_set n gT D : @group_set gT 'L_{n}(D). +Proof. case: n => [|n]; exact: groupP. Qed. + +Canonical Puig_at_group n gT D := Group (@Puig_at_group_set n gT D). +Canonical Puig_inf_group gT (D : {set gT}) := [group of 'L_*(D)]. +Canonical Puig_group gT (D : {set gT}) := [group of 'L(D)]. + +End PuigSeriesGroups. + +Notation "''L_[' G ] ( L )" := (Puig_succ_group G L) : Group_scope. +Notation "''L_{' n } ( G )" := (Puig_at_group n G) + (at level 8, format "''L_{' n } ( G )") : Group_scope. +Notation "''L_*' ( G )" := (Puig_inf_group G) : Group_scope. +Notation "''L' ( G )" := (Puig_group G) : Group_scope. + +(* Elementary properties of the Puig series. *) +Section PuigBasics. + +Variable gT : finGroupType. +Implicit Types (D E : {set gT}) (G H : {group gT}). + +Lemma Puig0 D : 'L_{0}(D) = 1. Proof. by []. Qed. +Lemma PuigS n D : 'L_{n.+1}(D) = 'L_[D]('L_{n}(D)). Proof. by []. Qed. +Lemma Puig_recE n D : Puig_rec n D = 'L_{n}(D). Proof. by []. Qed. +Lemma Puig_def D : 'L(D) = 'L_[D]('L_*(D)). Proof. by []. Qed. + +Local Notation "D --> E" := (generated_by (norm_abelian D) E) + (at level 70, no associativity) : group_scope. + +Lemma Puig_gen D E : E --> 'L_[D](E). +Proof. by apply/existsP; exists (subgroups D). Qed. + +Lemma Puig_max G D E : D --> E -> E \subset G -> E \subset 'L_[G](D). +Proof. +case/existsP=> gE /eqP <-{E}; rewrite !gen_subG. +move/bigcupsP=> sEG; apply/bigcupsP=> A gEA; have [_ abnA]:= andP gEA. +by rewrite sub_gen // bigcup_sup // inE sEG. +Qed. + +Lemma norm_abgenS D1 D2 E : D1 \subset D2 -> D2 --> E -> D1 --> E. +Proof. +move=> sD12 /exists_eqP[gE <-{E}]. +apply/exists_eqP; exists [set A in gE | norm_abelian D2 A]. +congr <<_>>; apply: eq_bigl => A; rewrite !inE. +apply: andb_idr => /and3P[_ nAD cAA]. +by apply/andP; rewrite (subset_trans sD12). +Qed. + +Lemma Puig_succ_sub G D : 'L_[G](D) \subset G. +Proof. by rewrite gen_subG; apply/bigcupsP=> A /andP[]; rewrite inE. Qed. + +Lemma Puig_at_sub n G : 'L_{n}(G) \subset G. +Proof. by case: n => [|n]; rewrite ?sub1G ?Puig_succ_sub. Qed. + +(* This is B & G, Lemma B.1(d), first part. *) +Lemma Puig_inf_sub G : 'L_*(G) \subset G. +Proof. exact: Puig_at_sub. Qed. + +Lemma Puig_sub G : 'L(G) \subset G. +Proof. exact: Puig_at_sub. Qed. + +(* This is part of B & G, Lemma B.1(b). *) +Lemma Puig1 G : 'L_{1}(G) = G. +Proof. +apply/eqP; rewrite eqEsubset Puig_at_sub; apply/subsetP=> x Gx. +rewrite -cycle_subG sub_gen // -[<[x]>]/(gval _) bigcup_sup //=. +by rewrite inE cycle_subG Gx /= /norm_abelian cycle_abelian sub1G. +Qed. + +End PuigBasics. + +(* Functoriality of the Puig series. *) + +Fact Puig_at_cont n : GFunctor.iso_continuous (Puig_at n). +Proof. +elim: n => [|n IHn] aT rT G f injf; first by rewrite morphim1. +have IHnS := Puig_at_sub n; pose func_n := [igFun by IHnS & !IHn]. +rewrite !PuigS sub_morphim_pre ?Puig_succ_sub // gen_subG; apply/bigcupsP=> A. +rewrite inE => /and3P[sAG nAL cAA]; rewrite -sub_morphim_pre ?sub_gen //. +rewrite -[f @* A]/(gval _) bigcup_sup // inE morphimS // /norm_abelian. +rewrite morphim_abelian // -['L_{n}(_)](injmF func_n injf) //=. +by rewrite morphim_norms. +Qed. + +Canonical Puig_at_igFun n := [igFun by Puig_at_sub^~ n & !Puig_at_cont n]. + +Fact Puig_inf_cont : GFunctor.iso_continuous Puig_inf. +Proof. +by move=> aT rT G f injf; rewrite /Puig_inf card_injm // Puig_at_cont. +Qed. + +Canonical Puig_inf_igFun := [igFun by Puig_inf_sub & !Puig_inf_cont]. + +Fact Puig_cont : GFunctor.iso_continuous Puig. +Proof. by move=> aT rT G f injf; rewrite /Puig card_injm // Puig_at_cont. Qed. + +Canonical Puig_igFun := [igFun by Puig_sub & !Puig_cont]. diff --git a/mathcomp/odd_order/BGsection10.v b/mathcomp/odd_order/BGsection10.v new file mode 100644 index 0000000..88b4c39 --- /dev/null +++ b/mathcomp/odd_order/BGsection10.v @@ -0,0 +1,1497 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import bigop finset prime fingroup morphism perm automorphism quotient. +Require Import action gproduct gfunctor pgroup cyclic center commutator. +Require Import gseries nilpotent sylow abelian maximal hall. +Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +Require Import BGsection7 BGsection9. + +(******************************************************************************) +(* This file covers B & G, section 10, including with the definitions: *) +(* \alpha(M) == the primes p such that M has p-rank at least 3. *) +(* \beta(M) == the primes p in \alpha(M) such that Sylow p-subgroups of M *) +(* are not narrow (see BGsection5), i.e., such that M contains *) +(* no maximal elementary abelian subgroups of rank 2. In a *) +(* minimal counter-example G, \beta(M) is the intersection of *) +(* \alpha(M) and \beta(G). Note that B & G refers to primes in *) +(* \beta(G) as ``ideal'' primes, somewhat inconsistently. *) +(* \sigma(M) == the primes p such that there exists a p-Sylow subgroup P *) +(* of M whose normaliser (in the minimal counter-example) is *) +(* contained in M. *) +(* M`_\alpha == the \alpha(M)-core of M. *) +(* M`_\beta == the \beta(M)-core of M. *) +(* M`_\sigma == the \sigma(M)-core of M. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Reserved Notation "\alpha ( M )" (at level 2, format "\alpha ( M )"). +Reserved Notation "\beta ( M )" (at level 2, format "\beta ( M )"). +Reserved Notation "\sigma ( M )" (at level 2, format "\sigma ( M )"). + +Reserved Notation "M `_ \alpha" (at level 3, format "M `_ \alpha"). +Reserved Notation "M `_ \beta" (at level 3, format "M `_ \beta"). +Reserved Notation "M `_ \sigma" (at level 3, format "M `_ \sigma"). + +Section Def. + +Variable gT : finGroupType. +Implicit Type p : nat. + +Variable M : {set gT}. + +Definition alpha := [pred p | 2 < 'r_p(M)]. +Definition alpha_core := 'O_alpha(M). +Canonical Structure alpha_core_group := Eval hnf in [group of alpha_core]. + +Definition beta := + [pred p | [forall (P : {group gT} | p.-Sylow(M) P), ~~ p.-narrow P]]. +Definition beta_core := 'O_beta(M). +Canonical Structure beta_core_group := Eval hnf in [group of beta_core]. + +Definition sigma := + [pred p | [exists (P : {group gT} | p.-Sylow(M) P), 'N(P) \subset M]]. +Definition sigma_core := 'O_sigma(M). +Canonical Structure sigma_core_group := Eval hnf in [group of sigma_core]. + +End Def. + +Notation "\alpha ( M )" := (alpha M) : group_scope. +Notation "M `_ \alpha" := (alpha_core M) : group_scope. +Notation "M `_ \alpha" := (alpha_core_group M) : Group_scope. + +Notation "\beta ( M )" := (beta M) : group_scope. +Notation "M `_ \beta" := (beta_core M) : group_scope. +Notation "M `_ \beta" := (beta_core_group M) : Group_scope. + +Notation "\sigma ( M )" := (sigma M) : group_scope. +Notation "M `_ \sigma" := (sigma_core M) : group_scope. +Notation "M `_ \sigma" := (sigma_core_group M) : Group_scope. + +Section CoreTheory. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Type x : gT. +Implicit Type P : {group gT}. + +Section GenericCores. + +Variables H K : {group gT}. + +Lemma sigmaJ x : \sigma(H :^ x) =i \sigma(H). +Proof. +move=> p; apply/exists_inP/exists_inP=> [] [P sylP sNH]; last first. + by exists (P :^ x)%G; rewrite ?pHallJ2 // normJ conjSg. +by exists (P :^ x^-1)%G; rewrite ?normJ ?sub_conjgV // -(pHallJ2 _ _ _ x) actKV. +Qed. + +Lemma MsigmaJ x : (H :^ x)`_\sigma = H`_\sigma :^ x. +Proof. by rewrite /sigma_core -(eq_pcore H (sigmaJ x)) pcoreJ. Qed. + +Lemma alphaJ x : \alpha(H :^ x) =i \alpha(H). +Proof. by move=> p; rewrite !inE /= p_rankJ. Qed. + +Lemma MalphaJ x : (H :^ x)`_\alpha = H`_\alpha :^ x. +Proof. by rewrite /alpha_core -(eq_pcore H (alphaJ x)) pcoreJ. Qed. + +Lemma betaJ x : \beta(H :^ x) =i \beta(H). +Proof. +move=> p; apply/forall_inP/forall_inP=> nnSylH P sylP. + by rewrite -(@narrowJ _ _ _ x) nnSylH ?pHallJ2. +by rewrite -(@narrowJ _ _ _ x^-1) nnSylH // -(pHallJ2 _ _ _ x) actKV. +Qed. + +Lemma MbetaJ x : (H :^ x)`_\beta = H`_\beta :^ x. +Proof. by rewrite /beta_core -(eq_pcore H (betaJ x)) pcoreJ. Qed. + +End GenericCores. + +(* This remark appears at the start (p. 70) of B & G, Section 10, just after *) +(* the definition of ideal, which we do not include, since it is redundant *) +(* with the notation p \in \beta(G) that is used later. *) +Remark not_narrow_ideal p P : p.-Sylow(G) P -> ~~ p.-narrow P -> p \in \beta(G). +Proof. +move=> sylP nnP; apply/forall_inP=> Q sylQ. +by have [x _ ->] := Sylow_trans sylP sylQ; rewrite narrowJ. +Qed. + +Section MaxCores. + +Variables M : {group gT}. +Hypothesis maxM : M \in 'M. + +(* This is the first inclusion in the remark following the preliminary *) +(* definitions in B & G, p. 70. *) +Remark beta_sub_alpha : {subset \beta(M) <= \alpha(M)}. +Proof. +move=> p; rewrite !inE /= => /forall_inP nnSylM. +have [P sylP] := Sylow_exists p M; have:= nnSylM P sylP. +by rewrite negb_imply (p_rank_Sylow sylP) => /andP[]. +Qed. + +Remark alpha_sub_sigma : {subset \alpha(M) <= \sigma(M)}. +Proof. +move=> p a_p; have [P sylP] := Sylow_exists p M; have [sPM pP _ ] := and3P sylP. +have{a_p} rP: 2 < 'r(P) by rewrite (rank_Sylow sylP). +apply/exists_inP; exists P; rewrite ?uniq_mmax_norm_sub //. +exact: def_uniq_mmax (rank3_Uniqueness (mFT_pgroup_proper pP) rP) maxM sPM. +Qed. + +Remark beta_sub_sigma : {subset \beta(M) <= \sigma(M)}. +Proof. by move=> p; move/beta_sub_alpha; exact: alpha_sub_sigma. Qed. + +Remark Mbeta_sub_Malpha : M`_\beta \subset M`_\alpha. +Proof. exact: sub_pcore beta_sub_alpha. Qed. + +Remark Malpha_sub_Msigma : M`_\alpha \subset M`_\sigma. +Proof. exact: sub_pcore alpha_sub_sigma. Qed. + +Remark Mbeta_sub_Msigma : M`_\beta \subset M`_\sigma. +Proof. exact: sub_pcore beta_sub_sigma. Qed. + +(* This is the first part of the remark just above B & G, Theorem 10.1. *) +Remark norm_sigma_Sylow p P : + p \in \sigma(M) -> p.-Sylow(M) P -> 'N(P) \subset M. +Proof. +case/exists_inP=> Q sylQ sNPM sylP. +by case: (Sylow_trans sylQ sylP) => m mM ->; rewrite normJ conj_subG. +Qed. + +(* This is the second part of the remark just above B & G, Theorem 10.1. *) +Remark sigma_Sylow_G p P : p \in \sigma(M) -> p.-Sylow(M) P -> p.-Sylow(G) P. +Proof. +move=> sMp sylP; apply: (mmax_sigma_Sylow maxM) => //. +exact: norm_sigma_Sylow sMp sylP. +Qed. + +Lemma sigma_Sylow_neq1 p P : p \in \sigma(M) -> p.-Sylow(M) P -> P :!=: 1. +Proof. +move=> sMp /(norm_sigma_Sylow sMp); apply: contraTneq => ->. +by rewrite norm1 subTset -properT mmax_proper. +Qed. + +Lemma sigma_sub_pi : {subset \sigma(M) <= \pi(M)}. +Proof. +move=> p sMp; have [P sylP]:= Sylow_exists p M. +by rewrite -p_rank_gt0 -(rank_Sylow sylP) rank_gt0 (sigma_Sylow_neq1 sMp sylP). +Qed. + +Lemma predI_sigma_alpha : [predI \sigma(M) & \alpha(G)] =i \alpha(M). +Proof. +move=> p; rewrite inE /= -(andb_idl (@alpha_sub_sigma p)). +apply: andb_id2l => sMp; have [P sylP] := Sylow_exists p M. +by rewrite !inE -(rank_Sylow sylP) -(rank_Sylow (sigma_Sylow_G sMp sylP)). +Qed. + +Lemma predI_sigma_beta : [predI \sigma(M) & \beta(G)] =i \beta(M). +Proof. +move=> p; rewrite inE /= -(andb_idl (@beta_sub_sigma p)). +apply: andb_id2l => sMp; apply/idP/forall_inP=> [bGp P sylP | nnSylM]. + exact: forall_inP bGp P (sigma_Sylow_G sMp sylP). +have [P sylP] := Sylow_exists p M. +exact: not_narrow_ideal (sigma_Sylow_G sMp sylP) (nnSylM P sylP). +Qed. + +End MaxCores. + +End CoreTheory. + +Section Ten. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). + +Implicit Type p : nat. +Implicit Type A E H K M N P Q R S V W X Y : {group gT}. + +(* This is B & G, Theorem 10.1(d); note that we do not assume M is maximal. *) +Theorem sigma_Sylow_trans M p X g : + p \in \sigma(M) -> p.-Sylow(M) X -> X :^ g \subset M -> g \in M. +Proof. +move=> sMp sylX sXgM; have pX := pHall_pgroup sylX. +have [|h hM /= sXghX] := Sylow_Jsub sylX sXgM; first by rewrite pgroupJ. +by rewrite -(groupMr _ hM) (subsetP (norm_sigma_Sylow _ sylX)) ?inE ?conjsgM. +Qed. + +(* This is B & G, Theorem 10.1 (a, b, c). *) +(* Part (e) of Theorem 10.1 is obviously stated incorrectly, and this is *) +(* difficult to correct because it is not used in the rest of the proof. *) +Theorem sigma_group_trans M p X : + M \in 'M -> p \in \sigma(M) -> p.-group X -> + [/\ (*a*) forall g, X \subset M -> X :^ g \subset M -> + exists2 c, c \in 'C(X) & exists2 m, m \in M & g = c * m, + (*b*) [transitive 'C(X), on [set Mg in M :^: G | X \subset Mg] | 'Js ] + & (*c*) X \subset M -> 'C(X) * 'N_M(X) = 'N(X)]. +Proof. +move=> maxM sMp pX; have defNM := norm_mmax maxM. +pose OM (Y : {set gT}) : {set {set gT}} := [set Mg in M :^: G | Y \subset Mg]. +pose trM (Y : {set gT}) := [transitive 'C(Y), on OM Y | 'Js]. +have actsOM Y: [acts 'N(Y), on OM Y | 'Js]. + apply/actsP=> z nYz Q; rewrite !inE -{1}(normP nYz) conjSg. + by rewrite (acts_act (acts_orbit _ _ _)) ?inE. +have OMid Y: (gval M \in OM Y) = (Y \subset M) by rewrite inE orbit_refl. +have ntOM Y: p.-group Y -> exists B, gval B \in OM Y. + have [S sylS] := Sylow_exists p M; have sSM := pHall_sub sylS. + have sylS_G := sigma_Sylow_G maxM sMp sylS. + move=> pY; have [g Gg sXSg] := Sylow_subJ sylS_G (subsetT Y) pY. + by exists (M :^ g)%G; rewrite inE mem_orbit // (subset_trans sXSg) ?conjSg. +have maxOM Y H: gval H \in OM Y -> H \in 'M. + by case/setIdP=> /imsetP[g _ /val_inj->]; rewrite mmaxJ. +have part_c Y H: trM Y -> gval H \in OM Y -> 'C(Y) * 'N_H(Y) = 'N(Y). + move=> trMY O_H; rewrite -(norm_mmax (maxOM Y H O_H)) -(astab1Js H) setIC. + have [sCN nCN] := andP (cent_normal Y); rewrite -normC 1?subIset ?nCN //. + by apply/(subgroup_transitiveP O_H); rewrite ?(atrans_supgroup sCN) ?actsOM. +suffices trMX: trM X. + do [split; rewrite // -OMid] => [g O_M sXgM |]; last exact: part_c. + have O_Mg': M :^ g^-1 \in OM X by rewrite inE mem_orbit -?sub_conjg ?inE. + have [c Cc /= Mc] := atransP2 trMX O_M O_Mg'; exists c^-1; rewrite ?groupV //. + by exists (c * g); rewrite ?mulKg // -defNM inE conjsgM -Mc conjsgKV. +elim: {X}_.+1 {-2}X (ltnSn (#|G| - #|X|)) => // n IHn X geXn in pX *. +have{n IHn geXn} IHX Y: X \proper Y -> p.-group Y -> trM Y. + move=> ltXY; apply: IHn; rewrite -ltnS (leq_trans _ geXn) // ltnS. + by rewrite ltn_sub2l ?(leq_trans (proper_card ltXY)) // cardsT max_card. +have [-> | ntX] := eqsVneq X 1. + rewrite /trM cent1T /OM setIdE (setIidPl _) ?atrans_orbit //. + by apply/subsetP=> Mg; case/imsetP=> g _ ->; rewrite inE sub1G. +pose L := 'N(X)%G; have ltLG := mFT_norm_proper ntX (mFT_pgroup_proper pX). +have IH_L: {in OM X &, forall B B', + B != B' -> exists2 X1, X \proper gval X1 & p.-Sylow(B :&: L) X1}. +- move=> _ _ /setIdP[/imsetP[a Ga ->] sXMa] /setIdP[/imsetP[b Gb ->] sXMb]. + move=> neqMab. + have [S sylS sXS] := Sylow_superset sXMa pX; have [sSMa pS _] := and3P sylS. + have [defS | ltXS] := eqVproper sXS. + case/eqP: neqMab; apply: (canRL (actKV _ _)); apply: (act_inj 'Js a). + rewrite /= -conjsgM [_ :^ _]conjGid ?(sigma_Sylow_trans _ sylS) ?sigmaJ //. + by rewrite -defS conjsgM conjSg sub_conjgV. + have pSL: p.-group (S :&: L) := pgroupS (subsetIl _ _) pS. + have [X1 sylX1 sNX1] := Sylow_superset (setSI L sSMa) pSL; exists X1 => //. + by rewrite (proper_sub_trans (nilpotent_proper_norm (pgroup_nil pS) _)). +have [M1 O_M1] := ntOM X pX; apply/imsetP; exists (gval M1) => //; apply/eqP. +rewrite eqEsubset andbC acts_sub_orbit ?(subset_trans (cent_sub X)) // O_M1 /=. +apply/subsetP=> M2 O_M2. +have [-> | neqM12] := eqsVneq M1 M2; first exact: orbit_refl. +have [|X2 ltXX2 sylX2] := IH_L _ _ O_M2 O_M1; first by rewrite eq_sym. +have{IH_L neqM12} [X1 ltXX1 sylX1] := IH_L _ _ O_M1 O_M2 neqM12. +have [[sX1L1 pX1 _] [sX2L2 pX2 _]] := (and3P sylX1, and3P sylX2). +have [[sX1M1 sX1L] [sX2M2 sX2L]] := (subsetIP sX1L1, subsetIP sX2L2). +have [P sylP sX1P] := Sylow_superset sX1L pX1; have [sPL pP _] := and3P sylP. +have [M0 O_M0] := ntOM P pP; have [MG_M0 sPM0] := setIdP O_M0. +have [t Lt sX2Pt] := Sylow_subJ sylP sX2L pX2. +have [sX1M0 ltXP] := (subset_trans sX1P sPM0, proper_sub_trans ltXX1 sX1P). +have M0C_M1: gval M1 \in orbit 'Js 'C(X) M0. + rewrite (subsetP (imsetS _ (centS (proper_sub ltXX1)))) // -orbitE. + by rewrite (atransP (IHX _ ltXX1 pX1)) inE ?MG_M0 //; case/setIdP: O_M1 => ->. +have M0tC_M2: M2 \in orbit 'Js 'C(X) (M0 :^ t). + rewrite (subsetP (imsetS _ (centS (proper_sub ltXX2)))) // -orbitE. + rewrite (atransP (IHX _ ltXX2 pX2)) inE; first by case/setIdP: O_M2 => ->. + rewrite (acts_act (acts_orbit _ _ _)) ?inE ?MG_M0 //. + by rewrite (subset_trans sX2Pt) ?conjSg. +rewrite (orbit_transl M0C_M1) (orbit_transr _ M0tC_M2). +have maxM0 := maxOM _ _ O_M0; have ltMG := mmax_proper maxM0. +have [rPgt2 | rPle2] := ltnP 2 'r(P). + have uP: P \in 'U by rewrite rank3_Uniqueness ?(mFT_pgroup_proper pP). + have uP_M0: 'M(P) = [set M0] := def_uniq_mmax uP maxM0 sPM0. + by rewrite conjGid ?orbit_refl ?(subsetP (sub_uniq_mmax uP_M0 sPL ltLG)). +have pl1L: p.-length_1 L. + have [oddL]: odd #|L| /\ 'r_p(L) <= 2 by rewrite mFT_odd -(rank_Sylow sylP). + by case/rank2_der1_complement; rewrite ?mFT_sol ?plength1_pseries2_quo. +have [|u v nLPu Lp'_v ->] := imset2P (_ : t \in 'N_L(P) * 'O_p^'(L)). + by rewrite normC ?plength1_Frattini // subIset ?gFnorm. +rewrite actM (orbit_transr _ (mem_orbit _ _ _)); last first. + have coLp'X: coprime #|'O_p^'(L)| #|X| := p'nat_coprime (pcore_pgroup _ _) pX. + apply: subsetP Lp'_v; have [sLp'L nLp'L] := andP (pcore_normal p^' L). + rewrite -subsetIidl -coprime_norm_cent ?subsetIidl //. + exact: subset_trans (normG X) nLp'L. +have [|w x nM0Pw cPx ->] := imset2P (_ : u \in 'N_M0(P) * 'C(P)). + rewrite normC ?part_c ?IHX //; first by case/setIP: nLPu. + by rewrite setIC subIset ?cent_norm. +rewrite actM /= conjGid ?mem_orbit //; last by case/setIP: nM0Pw. +by rewrite (subsetP (centS (subset_trans (proper_sub ltXX1) sX1P))). +Qed. + +Section OneMaximal. + +Variable M : {group gT}. +Hypothesis maxM : M \in 'M. + +Let ltMG := mmax_proper maxM. +Let solM := mmax_sol maxM. + +Let aMa : \alpha(M).-group (M`_\alpha). Proof. exact: pcore_pgroup. Qed. +Let nsMaM : M`_\alpha <| M. Proof. exact: pcore_normal. Qed. +Let sMaMs : M`_\alpha \subset M`_\sigma. Proof. exact: Malpha_sub_Msigma. Qed. + +Let F := 'F(M / M`_\alpha). +Let nsFMa : F <| M / M`_\alpha. Proof. exact: Fitting_normal. Qed. + +Let alpha'F : \alpha(M)^'.-group F. +Proof. +rewrite -[F](nilpotent_pcoreC \alpha(M) (Fitting_nil _)) -Fitting_pcore /=. +by rewrite trivg_pcore_quotient (trivgP (Fitting_sub 1)) dprod1g pcore_pgroup. +Qed. + +Let Malpha_quo_sub_Fitting : M^`(1) / M`_\alpha \subset F. +Proof. +have [/= K defF sMaK nsKM] := inv_quotientN nsMaM nsFMa; rewrite -/F in defF. +have [sKM _] := andP nsKM; have nsMaK: M`_\alpha <| K := normalS sMaK sKM nsMaM. +have [[_ nMaK] [_ nMaM]] := (andP nsMaK, andP nsMaM). +have hallMa: \alpha(M).-Hall(K) M`_\alpha. + by rewrite /pHall sMaK pcore_pgroup -card_quotient -?defF. +have [H hallH] := Hall_exists \alpha(M)^' (solvableS sKM solM). +have{hallH} defK := sdprod_normal_p'HallP nsMaK hallH hallMa. +have{defK} [_ sHK defK nMaH tiMaH] := sdprod_context defK. +have{defK} isoHF: H \isog F by rewrite [F]defF -defK quotientMidl quotient_isog. +have{sHK nMaH} sHM := subset_trans sHK sKM. +have{tiMaH isoHF sHM H} rF: 'r(F) <= 2. + rewrite -(isog_rank isoHF); have [p p_pr -> /=] := rank_witness H. + have [|a_p] := leqP 'r_p(M) 2; first exact: leq_trans (p_rankS p sHM). + rewrite 2?leqW // leqNgt p_rank_gt0 /= (card_isog isoHF) /= -/F. + exact: contraL (pnatPpi alpha'F) a_p. +by rewrite quotient_der // rank2_der1_sub_Fitting ?mFT_quo_odd ?quotient_sol. +Qed. + +Let sigma_Hall_sub_der1 H : \sigma(M).-Hall(M) H -> H \subset M^`(1). +Proof. +move=> hallH; have [sHM sH _] := and3P hallH. +rewrite -(Sylow_gen H) gen_subG; apply/bigcupsP=> P /SylowP[p p_pr sylP]. +have [-> | ntP] := eqsVneq P 1; first by rewrite sub1G. +have [sPH pP _] := and3P sylP; have{ntP} [_ p_dv_P _] := pgroup_pdiv pP ntP. +have{p_dv_P} s_p: p \in \sigma(M) := pgroupP (pgroupS sPH sH) p p_pr p_dv_P. +have{sylP} sylP: p.-Sylow(M) P := subHall_Sylow hallH s_p sylP. +have [sPM nMP] := (pHall_sub sylP, norm_sigma_Sylow s_p sylP). +have sylP_G := sigma_Sylow_G maxM s_p sylP. +have defG': G^`(1) = G. + have [_ simpG] := simpleP _ (mFT_simple gT). + by have [?|//] := simpG _ (der_normal 1 _); case/derG1P: (mFT_nonAbelian gT). +rewrite -subsetIidl -{1}(setIT P) -defG'. +rewrite (focal_subgroup_gen sylP_G) (focal_subgroup_gen sylP) genS //. +apply/subsetP=> _ /imset2P[x g Px /setIdP[Gg Pxg] ->]. +pose X := <[x]>; have sXM : X \subset M by rewrite cycle_subG (subsetP sPM). +have sXgM: X :^ g \subset M by rewrite -cycleJ cycle_subG (subsetP sPM). +have [trMX _ _] := sigma_group_trans maxM s_p (mem_p_elt pP Px). +have [c cXc [m Mm def_g]] := trMX _ sXM sXgM; rewrite cent_cycle in cXc. +have def_xg: x ^ g = x ^ m by rewrite def_g conjgM /conjg -(cent1P cXc) mulKg. +by rewrite commgEl def_xg -commgEl mem_imset2 // inE Mm -def_xg. +Qed. + +(* This is B & G, Theorem 10.2(a1). *) +Theorem Malpha_Hall : \alpha(M).-Hall(M) M`_\alpha. +Proof. +have [H hallH] := Hall_exists \sigma(M) solM; have [sHM sH _] := and3P hallH. +rewrite (subHall_Hall hallH (alpha_sub_sigma maxM)) // /pHall pcore_pgroup /=. +rewrite -(card_quotient (subset_trans sHM (normal_norm nsMaM))) -pgroupE. +rewrite (subset_trans sMaMs) ?pcore_sub_Hall ?(pgroupS _ alpha'F) //=. +exact: subset_trans (quotientS _ (sigma_Hall_sub_der1 hallH)) _. +Qed. + +(* This is B & G, Theorem 10.2(b1). *) +Theorem Msigma_Hall : \sigma(M).-Hall(M) M`_\sigma. +Proof. +have [H hallH] := Hall_exists \sigma(M) solM; have [sHM sH _] := and3P hallH. +rewrite /M`_\sigma (normal_Hall_pcore hallH) // -(quotientGK nsMaM). +rewrite -(quotientGK (normalS _ sHM nsMaM)) ?cosetpre_normal //; last first. + by rewrite (subset_trans sMaMs) ?pcore_sub_Hall. +have hallHa: \sigma(M).-Hall(F) (H / M`_\alpha). + apply: pHall_subl (subset_trans _ Malpha_quo_sub_Fitting) (Fitting_sub _) _. + by rewrite quotientS ?sigma_Hall_sub_der1. + exact: quotient_pHall (subset_trans sHM (normal_norm nsMaM)) hallH. +rewrite (nilpotent_Hall_pcore (Fitting_nil _) hallHa) /=. +exact: char_normal_trans (pcore_char _ _) nsFMa. +Qed. + +Lemma pi_Msigma : \pi(M`_\sigma) =i \sigma(M). +Proof. +move=> p; apply/idP/idP=> [|s_p /=]; first exact: pnatPpi (pcore_pgroup _ _). +by rewrite (card_Hall Msigma_Hall) pi_of_part // inE /= sigma_sub_pi. +Qed. + +(* This is B & G, Theorem 10.2(b2). *) +Theorem Msigma_Hall_G : \sigma(M).-Hall(G) M`_\sigma. +Proof. +rewrite pHallE subsetT /= eqn_dvd {1}(card_Hall Msigma_Hall). +rewrite partn_dvd ?cardG_gt0 ?cardSg ?subsetT //=. +apply/dvdn_partP; rewrite ?part_gt0 // => p. +rewrite pi_of_part ?cardG_gt0 // => /andP[_ s_p]. +rewrite partn_part => [|q /eqnP-> //]. +have [P sylP] := Sylow_exists p M; have [sPM pP _] := and3P sylP. +rewrite -(card_Hall (sigma_Sylow_G _ _ sylP)) ?cardSg //. +by rewrite (sub_Hall_pcore Msigma_Hall) ?(pi_pgroup pP). +Qed. + +(* This is B & G, Theorem 10.2(a2). *) +Theorem Malpha_Hall_G : \alpha(M).-Hall(G) M`_\alpha. +Proof. +apply: subHall_Hall Msigma_Hall_G (alpha_sub_sigma maxM) _. +exact: pHall_subl sMaMs (pcore_sub _ _) Malpha_Hall. +Qed. + +(* This is B & G, Theorem 10.2(c). *) +Theorem Msigma_der1 : M`_\sigma \subset M^`(1). +Proof. exact: sigma_Hall_sub_der1 Msigma_Hall. Qed. + +(* This is B & G, Theorem 10.2(d1). *) +Theorem Malpha_quo_rank2 : 'r(M / M`_\alpha) <= 2. +Proof. +have [p p_pr ->] := rank_witness (M / M`_\alpha). +have [P sylP] := Sylow_exists p M; have [sPM pP _] := and3P sylP. +have nMaP := subset_trans sPM (normal_norm nsMaM). +rewrite -(rank_Sylow (quotient_pHall nMaP sylP)) /= leqNgt. +have [a_p | a'p] := boolP (p \in \alpha(M)). + by rewrite quotientS1 ?rank1 ?(sub_Hall_pcore Malpha_Hall) ?(pi_pgroup pP). +rewrite -(isog_rank (quotient_isog _ _)) ?coprime_TIg ?(rank_Sylow sylP) //. +exact: pnat_coprime aMa (pi_pnat pP _). +Qed. + +(* This is B & G, Theorem 10.2(d2). *) +Theorem Malpha_quo_nil : nilpotent (M^`(1) / M`_\alpha). +Proof. exact: nilpotentS Malpha_quo_sub_Fitting (Fitting_nil _). Qed. + +(* This is B & G, Theorem 10.2(e). *) +Theorem Msigma_neq1 : M`_\sigma :!=: 1. +Proof. +without loss Ma1: / M`_\alpha = 1. + by case: eqP => // Ms1 -> //; apply/trivgP; rewrite -Ms1 Malpha_sub_Msigma. +have{Ma1} rFM: 'r('F(M)) <= 2. + rewrite (leq_trans _ Malpha_quo_rank2) // Ma1. + by rewrite -(isog_rank (quotient1_isog _)) rankS ?Fitting_sub. +pose q := max_pdiv #|M|; pose Q := 'O_q(M)%G. +have sylQ: q.-Sylow(M) Q := rank2_max_pcore_Sylow (mFT_odd M) solM rFM. +have piMq: q \in \pi(M) by rewrite pi_max_pdiv cardG_gt1 mmax_neq1. +have{piMq} ntQ: Q :!=: 1 by rewrite -rank_gt0 (rank_Sylow sylQ) p_rank_gt0. +rewrite (subG1_contra _ ntQ) ?(sub_Hall_pcore Msigma_Hall) ?pcore_sub //. +rewrite (pi_pgroup (pcore_pgroup _ _)) //; apply/exists_inP; exists Q => //. +by rewrite (mmax_normal maxM) ?pcore_normal. +Qed. + +(* This is B & G, Lemma 10.3. *) +Theorem cent_alpha'_uniq X : + X \subset M -> \alpha(M)^'.-group X -> 'r('C_(M`_\alpha)(X)) >= 2 -> + 'C_M(X)%G \in 'U. +Proof. +have ltM_G := sub_proper_trans (subsetIl M _) ltMG. +move=> sXM a'X; have [p p_pr -> rCX] := rank_witness 'C_(M`_\alpha)(X). +have{rCX} [B EpB] := p_rank_geP rCX; have{EpB} [sBCX abelB dimB] := pnElemP EpB. +have [[sBMa cXB] [pB cBB _]] := (subsetIP sBCX, and3P abelB). +have rMa: 1 < 'r_p(M`_\alpha) by rewrite -dimB -p_rank_abelem ?p_rankS. +have{rMa} a_p: p \in \alpha(M) by rewrite (pnatPpi aMa) // -p_rank_gt0 ltnW. +have nBX: X \subset 'N(B) by rewrite cents_norm // centsC. +have coMaX: coprime #|M`_\alpha| #|X| := pnat_coprime aMa a'X. +have [sMaM nMaM] := andP nsMaM; have solMa := solvableS sMaM solM. +have nMaX := subset_trans sXM nMaM. +have [P [sylP nPX sBP]] := coprime_Hall_subset nMaX coMaX solMa sBMa pB nBX. +have [sPMa pP _] := and3P sylP; have sPM := subset_trans sPMa sMaM. +have EpCB: B \in 'E_p^2('C_P(B)) by rewrite !inE subsetI sBP abelB dimB !andbT. +have: 1 < 'r_p('C_P(B)) by apply/p_rank_geP; exists B. +rewrite leq_eqVlt; case: ltngtP => // rCPB _. + apply: (uniq_mmaxS (subset_trans sBCX (setSI _ sMaM))) => //=. + have pCPB := pgroupS (subsetIl P 'C(B)) pP; rewrite -rank_pgroup // in rCPB. + have: 2 < 'r('C(B)) by rewrite (leq_trans rCPB) ?rankS ?subsetIr. + by apply: cent_rank3_Uniqueness; rewrite -dimB -rank_abelem. +have cPX: P \subset 'C(X). + have EpPB: B \in 'E_p(P) by exact/pElemP. + have coPX: coprime #|P| #|X| := coprimeSg sPMa coMaX. + rewrite centsC (coprime_odd_faithful_cent_abelem EpPB) ?mFT_odd //. + rewrite -(setIid 'C(B)) setIA (pmaxElem_LdivP p_pr _) 1?centsC //. + by rewrite (subsetP (p_rankElem_max _ _)) -?rCPB. +have sylP_M := subHall_Sylow Malpha_Hall a_p sylP. +have{sylP_M} rP: 2 < 'r(P) by rewrite (rank_Sylow sylP_M). +by rewrite rank3_Uniqueness ?(leq_trans rP (rankS _)) //= subsetI sPM. +Qed. + +Variable p : nat. + +(* This is B & G, Lemma 10.4(a). *) +(* We omit the redundant assumption p \in \pi(M). *) +Lemma der1_quo_sigma' : p %| #|M / M^`(1)| -> p \in \sigma(M)^'. +Proof. +apply: contraL => /= s_p; have piMp := sigma_sub_pi maxM s_p. +have p_pr: prime p by move: piMp; rewrite mem_primes; case/andP. +rewrite -p'natE ?(pi'_p'nat _ s_p) // -pgroupE -partG_eq1. +rewrite -(card_Hall (quotient_pHall _ Msigma_Hall)) /=; last first. + exact: subset_trans (pcore_sub _ _) (der_norm _ _). +by rewrite quotientS1 ?cards1 // Msigma_der1. +Qed. + +Hypothesis s'p : p \in \sigma(M)^'. + +(* This is B & G, Lemma 10.4(b). *) +(* We do not need the assumption M`_\alpha != 1; the assumption p \in \pi(M) *) +(* is restated as P != 1. *) +Lemma cent1_sigma'_Zgroup P : + p.-Sylow(M) P -> P :!=: 1 -> + exists x, + [/\ x \in 'Ohm_1('Z(P))^#, 'M('C[x]) != [set M] & Zgroup 'C_(M`_\alpha)[x]]. +Proof. +move=> sylP ntP; set T := 'Ohm_1(_); have [sPM pP _] := and3P sylP. +have [charT nilP] := (char_trans (Ohm_char 1 _) (center_char P), pgroup_nil pP). +suffices [x Tx not_uCx]: exists2 x, x \in T^# & 'M('C[x]) != [set M]. + exists x; split=> //; rewrite odd_rank1_Zgroup ?mFT_odd //= leqNgt. + apply: contra not_uCx; rewrite -cent_cycle; set X := <[x]> => rCMaX. + have{Tx} [ntX Tx] := setD1P Tx; rewrite -cycle_eq1 -/X in ntX. + have sXP: X \subset P by rewrite cycle_subG (subsetP (char_sub charT)). + rewrite (@def_uniq_mmaxS _ M 'C_M(X)) ?subsetIr ?mFT_cent_proper //. + apply: def_uniq_mmax; rewrite ?subsetIl //. + rewrite cent_alpha'_uniq ?(subset_trans sXP) ?(pi_pgroup (pgroupS sXP pP)) //. + by apply: contra s'p; apply: alpha_sub_sigma. +apply/exists_inP; rewrite -negb_forall_in; apply: contra s'p. +move/forall_inP => uCT; apply/exists_inP; exists P => //. +apply/subsetP=> u nPu; have [y Ty]: exists y, y \in T^#. + by apply/set0Pn; rewrite setD_eq0 subG1 Ohm1_eq1 center_nil_eq1. +rewrite -(norm_mmax maxM) (sameP normP eqP) (inj_eq (@group_inj gT)) -in_set1. +have Tyu: y ^ u \in T^#. + by rewrite memJ_norm // normD1 (subsetP (char_norms charT)). +by rewrite -(eqP (uCT _ Tyu)) -conjg_set1 normJ mmax_ofJ (eqP (uCT _ Ty)) set11. +Qed. + +(* This is B & G, Lemma 10.4(c), part 1. *) +(* The redundant assumption p \in \pi(M) is omitted. *) +Lemma sigma'_rank2_max : 'r_p(M) = 2 -> 'E_p^2(M) \subset 'E*_p(G). +Proof. +move=> rpM; apply: contraR s'p => /subsetPn[A Ep2A not_maxA]. +have{Ep2A} [sAM abelA dimA] := pnElemP Ep2A; have [pA _ _] := and3P abelA. +have [P sylP sAP] := Sylow_superset sAM pA; have [_ pP _] := and3P sylP. +apply/exists_inP; exists P; rewrite ?uniq_mmax_norm_sub //. +apply: def_uniq_mmaxS (mFT_pgroup_proper pP) (def_uniq_mmax _ _ sAM) => //. +by rewrite (@nonmaxElem2_Uniqueness _ p) // !(not_maxA, inE) abelA dimA subsetT. +Qed. + +(* This is B & G, Lemma 10.4(c), part 2 *) +(* The redundant assumption p \in \pi(M) is omitted. *) +Lemma sigma'_rank2_beta' : 'r_p(M) = 2 -> p \notin \beta(G). +Proof. +move=> rpM; rewrite -[p \in _]negb_exists_in negbK; apply/exists_inP. +have [A Ep2A]: exists A, A \in 'E_p^2(M) by apply/p_rank_geP; rewrite rpM. +have [_ abelA dimA] := pnElemP Ep2A; have [pA _] := andP abelA. +have [P sylP sAP] := Sylow_superset (subsetT _) pA. +exists P; rewrite ?inE //; apply/implyP=> _; apply/set0Pn. +exists A; rewrite 3!inE abelA dimA sAP (subsetP (pmaxElemS _ (subsetT P))) //. +by rewrite inE (subsetP (sigma'_rank2_max rpM)) // inE. +Qed. + +(* This is B & G, Lemma 10.5, part 1; the condition on X has been weakened, *) +(* because the proof of Lemma 12.2(a) requires the stronger result. *) +Lemma sigma'_norm_mmax_rank2 X : p.-group X -> 'N(X) \subset M -> 'r_p(M) = 2. +Proof. +move=> pX sNX_M; have sXM: X \subset M := subset_trans (normG X) sNX_M. +have [P sylP sXP] := Sylow_superset sXM pX; have [sPM pP _] := and3P sylP. +apply: contraNeq s'p; case: ltngtP => // rM _; last exact: alpha_sub_sigma. +apply/exists_inP; exists P; rewrite ?(subset_trans _ sNX_M) ?char_norms //. +rewrite sub_cyclic_char // (odd_pgroup_rank1_cyclic pP) ?mFT_odd //. +by rewrite (p_rank_Sylow sylP). +Qed. + +(* This is B & G, Lemma 10.5, part 2. We omit the second claim of B & G 10.5 *) +(* as it is an immediate consequence of sigma'_rank2_beta' (i.e., 10.4(c)). *) +Lemma sigma'1Elem_sub_p2Elem X : + X \in 'E_p^1(G) -> 'N(X) \subset M -> + exists2 A, A \in 'E_p^2(G) & X \subset A. +Proof. +move=> EpX sNXM; have sXM := subset_trans (normG X) sNXM. +have [[_ abelX dimX] p_pr] := (pnElemP EpX, pnElem_prime EpX). +have pX := abelem_pgroup abelX; have rpM2 := sigma'_norm_mmax_rank2 pX sNXM. +have [P sylP sXP] := Sylow_superset sXM pX; have [sPM pP _] := and3P sylP. +pose T := 'Ohm_1('Z(P)); pose A := X <*> T; have nilP := pgroup_nil pP. +have charT: T \char P := char_trans (Ohm_char 1 _) (center_char P). +have neqTX: T != X. + apply: contraNneq s'p => defX; apply/exists_inP; exists P => //. + by rewrite (subset_trans _ sNXM) // -defX char_norms. +have rP: 'r(P) = 2 by rewrite (rank_Sylow sylP) rpM2. +have ntT: T != 1 by rewrite Ohm1_eq1 center_nil_eq1 // -rank_gt0 rP. +have sAP: A \subset P by rewrite join_subG sXP char_sub. +have cTX: T \subset 'C(X) := centSS (Ohm_sub 1 _) sXP (subsetIr P _). +have{cTX} defA: X \* T = A by rewrite cprodEY. +have{defA} abelA : p.-abelem A. + have pZ: p.-group 'Z(P) := pgroupS (center_sub P) pP. + by rewrite (cprod_abelem _ defA) abelX Ohm1_abelem ?center_abelian. +exists [group of A]; last exact: joing_subl. +rewrite !inE subsetT abelA eqn_leq -{1}rP -{1}(rank_abelem abelA) rankS //=. +rewrite -dimX (properG_ltn_log (pgroupS sAP pP)) // /proper join_subG subxx. +rewrite joing_subl /=; apply: contra ntT => sTX; rewrite eqEsubset sTX in neqTX. +by rewrite -(setIidPr sTX) prime_TIg ?(card_pnElem EpX). +Qed. + +End OneMaximal. + +(* This is B & G, Theorem 10.6. *) +Theorem mFT_proper_plength1 p H : H \proper G -> p.-length_1 H. +Proof. +case/mmax_exists=> M /setIdP[maxM sHM]. +suffices{H sHM}: p.-length_1 M by apply: plength1S. +have [solM oddM] := (mmax_sol maxM, mFT_odd M). +have [rpMle2 | a_p] := leqP 'r_p(M) 2. + by rewrite plength1_pseries2_quo; case/rank2_der1_complement: rpMle2. +pose Ma := M`_\alpha; have hallMa: \alpha(M).-Hall(M) Ma := Malpha_Hall maxM. +have [[K hallK] [sMaM aMa _]] := (Hall_exists \alpha(M)^' solM, and3P hallMa). +have defM: Ma ><| K = M by apply/sdprod_Hall_pcoreP. +have{aMa} coMaK: coprime #|Ma| #|K| := pnat_coprime aMa (pHall_pgroup hallK). +suffices{a_p hallMa}: p.-length_1 Ma. + rewrite !p_elt_gen_length1 /p_elt_gen setIdE /= -/Ma -(setIidPl sMaM) -setIA. + rewrite -(setIdE M) (setIidPr _) //; apply/subsetP=> x; case/setIdP=> Mx p_x. + by rewrite (mem_Hall_pcore hallMa) /p_elt ?(pi_pnat p_x). +have{sMaM} <-: [~: Ma, K] = Ma. + have sMaMs: Ma \subset M`_\sigma := Malpha_sub_Msigma maxM. + have sMaM': Ma \subset M^`(1) := subset_trans sMaMs (Msigma_der1 maxM). + by have [] := coprime_der1_sdprod defM coMaK (solvableS sMaM solM) sMaM'. +have [q q_pr q_dv_Mq]: {q | prime q & q %| #|M / M^`(1)| }. + apply: pdivP; rewrite card_quotient ?der_norm // indexg_gt1 proper_subn //. + by rewrite (sol_der1_proper solM) ?mmax_neq1. +have s'q: q \in \sigma(M)^' by apply: der1_quo_sigma' q_dv_Mq. +have [Q sylQ] := Sylow_exists q K; have [sQK qQ _] := and3P sylQ. +have a'q: q \in \alpha(M)^' by apply: contra s'q; apply: alpha_sub_sigma. +have{a'q sylQ hallK} sylQ: q.-Sylow(M) Q := subHall_Sylow hallK a'q sylQ. +have{q_dv_Mq} ntQ: Q :!=: 1. + rewrite -rank_gt0 (rank_Sylow sylQ) p_rank_gt0 mem_primes q_pr cardG_gt0. + exact: dvdn_trans q_dv_Mq (dvdn_quotient _ _). +have{s'q sylQ ntQ} [x [Q1x _ ZgCx]] := cent1_sigma'_Zgroup maxM s'q sylQ ntQ. +have{Q1x} [ntx Q1x] := setD1P Q1x. +have sZQ := center_sub Q; have{sQK} sZK := subset_trans sZQ sQK. +have{sZK} Kx: x \in K by rewrite (subsetP sZK) // (subsetP (Ohm_sub 1 _)). +have{sZQ qQ} abelQ1 := Ohm1_abelem (pgroupS sZQ qQ) (center_abelian Q). +have{q q_pr Q abelQ1 Q1x} ox: prime #[x] by rewrite (abelem_order_p abelQ1). +move: Kx ox ZgCx; rewrite -cycle_subG -cent_cycle. +exact: odd_sdprod_Zgroup_cent_prime_plength1 solM oddM defM coMaK. +Qed. + +Section OneSylow. + +Variables (p : nat) (P : {group gT}). +Hypothesis sylP_G: p.-Sylow(G) P. +Let pP : p.-group P := pHall_pgroup sylP_G. + +(* This is an B & G, Corollary 10.7(a), second part (which does not depend on *) +(* a particular complement). *) +Corollary mFT_Sylow_der1 : P \subset 'N(P)^`(1). +Proof. +have [-> | ntP] := eqsVneq P 1; first exact: sub1G. +have ltNG: 'N(P) \proper G := mFT_norm_proper ntP (mFT_pgroup_proper pP). +have [M] := mmax_exists ltNG; case/setIdP=> /= maxM sNM. +have [ltMG solM] := (mmax_proper maxM, mmax_sol maxM). +have [pl1M sPM] := (mFT_proper_plength1 p ltMG, subset_trans (normG P) sNM). +have sylP := pHall_subl sPM (subsetT M) sylP_G. +have sMp: p \in \sigma(M) by apply/exists_inP; exists P. +apply: subset_trans (dergS 1 (subsetIr M 'N(P))) => /=. +apply: plength1_Sylow_sub_der1 sylP pl1M (subset_trans _ (Msigma_der1 maxM)). +by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) ?(pi_pgroup pP). +Qed. + +(* This is B & G, Corollary 10.7(a), first part. *) +Corollary mFT_Sylow_sdprod_commg V : P ><| V = 'N(P) -> [~: P, V] = P. +Proof. +move=> defV; have sPN' := mFT_Sylow_der1. +have sylP := pHall_subl (normG P) (subsetT 'N(P)) sylP_G. +have [|//] := coprime_der1_sdprod defV _ (pgroup_sol pP) sPN'. +by rewrite (coprime_sdprod_Hall_l defV) // (pHall_Hall sylP). +Qed. + +(* This is B & G, Corollary 10.7(b). *) +Corollary mFT_rank2_Sylow_cprod : + 'r(P) < 3 -> ~~ abelian P -> + exists2 S, [/\ ~~ abelian (gval S), logn p #|S| = 3 & exponent S %| p] + & exists2 C, cyclic (gval C) & S \* C = P /\ 'Ohm_1(C) = 'Z(S). +Proof. +move=> rP not_cPP; have sylP := pHall_subl (normG P) (subsetT 'N(P)) sylP_G. +have ntP: P :!=: 1 by apply: contraNneq not_cPP => ->; apply: abelian1. +have ltNG: 'N(P) \proper G := mFT_norm_proper ntP (mFT_pgroup_proper pP). +have [V hallV] := Hall_exists p^' (mFT_sol ltNG); have [_ p'V _] := and3P hallV. +have defNP: P ><| V = 'N(P) := sdprod_normal_p'HallP (normalG P) hallV sylP. +have defP: [~: P, V] = P := mFT_Sylow_sdprod_commg defNP. +have [_] := rank2_coprime_comm_cprod pP (mFT_odd _) ntP rP defP p'V (mFT_odd _). +case=> [/idPn// | [S esS [C [mulSC cycC defC1]]]]. +exists S => //; exists C => //; split=> //; rewrite defC1. +have sSP: S \subset P by case/cprodP: mulSC => _ /mulG_sub[]. +have [[not_cSS dimS _] pS] := (esS, pgroupS sSP pP). +by have [||[]] := p3group_extraspecial pS; rewrite ?dimS. +Qed. + +(* This is B & G, Corollary 10.7(c). *) +Corollary mFT_sub_Sylow_trans : forall Q x, + Q \subset P -> Q :^ x \subset P -> exists2 y, y \in 'N(P) & Q :^ x = Q :^ y. +Proof. +move=> Q x; have [-> /trivgP-> /trivgP-> | ntP sQP sQxP] := eqsVneq P 1. + by exists 1; rewrite ?group1 ?conjs1g. +have ltNG: 'N(P) \proper G := mFT_norm_proper ntP (mFT_pgroup_proper pP). +have [M /=] := mmax_exists ltNG; case/setIdP=> maxM sNM. +have [ltMG solM] := (mmax_proper maxM, mmax_sol maxM). +have [pl1M sPM] := (mFT_proper_plength1 p ltMG, subset_trans (normG P) sNM). +have sylP := pHall_subl sPM (subsetT M) sylP_G. +have sMp: p \in \sigma(M) by apply/exists_inP; exists P. +have [transCQ _ _] := sigma_group_trans maxM sMp (pgroupS sQP pP). +have [||q cQq [u Mu defx]] := transCQ x; try exact: subset_trans _ sPM. +have nQC := normP (subsetP (cent_sub Q) _ _). +have [|q' cMQq' [y nMPy defu]] := plength1_Sylow_trans sylP pl1M solM sQP Mu. + by rewrite defx conjsgM nQC in sQxP. +have [[_ nPy] [_ cQq']] := (setIP nMPy, setIP cMQq'). +by exists y; rewrite // defx defu !conjsgM 2?nQC. +Qed. + +(* This is B & G, Corollary 10.7(d). *) +Corollary mFT_subnorm_Sylow Q : Q \subset P -> p.-Sylow('N(Q)) 'N_P(Q). +Proof. +move=> sQP; have pQ := pgroupS sQP pP. +have [S /= sylS] := Sylow_exists p 'N(Q); have [sNS pS _] := and3P sylS. +have sQS: Q \subset S := normal_sub_max_pgroup (Hall_max sylS) pQ (normalG Q). +have [x _ sSxP] := Sylow_Jsub sylP_G (subsetT S) pS. +have sQxP: Q :^ x \subset P by rewrite (subset_trans _ sSxP) ?conjSg. +have [y nPy defQy] := mFT_sub_Sylow_trans sQP sQxP. +have nQxy: x * y^-1 \in 'N(Q) by rewrite inE conjsgM defQy actK. +have sSxyP: S :^ (x * y^-1) \subset P by rewrite conjsgM sub_conjgV (normP nPy). +have sylSxy: p.-Sylow('N(Q)) (S :^ (x * y^-1)) by rewrite pHallJ. +have pNPQ: p.-group 'N_P(Q) := pgroupS (subsetIl P 'N(Q)) pP. +by rewrite (sub_pHall sylSxy pNPQ) ?subsetIr // subsetI sSxyP (@pHall_sub _ p). +Qed. + +(* This is B & G, Corollary 10.7(e). *) +Corollary mFT_Sylow_normalS Q R : + p.-group R -> Q \subset P :&: R -> Q <| 'N(P) -> Q <| 'N(R). +Proof. +move=> pR /subsetIP[sQP sQR] /andP[nQP nQ_NP]. +have [x _ sRxP] := Sylow_Jsub sylP_G (subsetT R) pR. +rewrite /normal normsG //; apply/subsetP=> y nRy. +have sQxP: Q :^ x \subset P by rewrite (subset_trans _ sRxP) ?conjSg. +have sQyxP: Q :^ (y * x) \subset P. + by rewrite actM (subset_trans _ sRxP) // -(normP nRy) !conjSg. +have [t tNP defQx] := mFT_sub_Sylow_trans sQP sQxP. +have [z zNP defQxy] := mFT_sub_Sylow_trans sQP sQyxP. +by rewrite inE -(conjSg _ _ x) -actM /= defQx defQxy !(normsP nQ_NP). +Qed. + +End OneSylow. + +Section AnotherMaximal. + +Variable M : {group gT}. +Hypothesis maxM : M \in 'M. + +Let solM : solvable M := mmax_sol maxM. +Let ltMG : M \proper G := mmax_proper maxM. + +Let sMbMs : M`_\beta \subset M`_\sigma := Mbeta_sub_Msigma maxM. +Let nsMbM : M`_\beta <| M := pcore_normal _ _. + +Let hallMs : \sigma(M).-Hall(M) M`_\sigma := Msigma_Hall maxM. +Let nsMsM : M`_\sigma <| M := pcore_normal _ M. +Let sMsM' : M`_\sigma \subset M^`(1) := Msigma_der1 maxM. + +Lemma Mbeta_der1 : M`_\beta \subset M^`(1). +Proof. exact: subset_trans sMbMs sMsM'. Qed. + +Let sM'M : M^`(1) \subset M := der_sub 1 M. +Let nsMsM' : M`_\sigma <| M^`(1) := normalS sMsM' sM'M nsMsM. +Let nsMbM' : M`_\beta <| M^`(1) := normalS Mbeta_der1 sM'M nsMbM. +Let nMbM' := normal_norm nsMbM'. + +(* This is B & G, Lemma 10.8(c). *) +Lemma beta_max_pdiv p : + p \notin \beta(M) -> + [/\ p^'.-Hall(M^`(1)) 'O_p^'(M^`(1)), + p^'.-Hall(M`_\sigma) 'O_p^'(M`_\sigma) + & forall q, q \in \pi(M / 'O_p^'(M)) -> q <= p]. +Proof. +rewrite !inE -negb_exists_in negbK => /exists_inP[P sylP nnP]. +have [|ncM' p_max] := narrow_der1_complement_max_pdiv (mFT_odd M) solM sylP nnP. + by rewrite mFT_proper_plength1 ?implybT. +by rewrite -(pcore_setI_normal _ nsMsM') (Hall_setI_normal nsMsM'). +Qed. + +(* This is B & G, Lemma 10.8(a), first part. *) +Lemma Mbeta_Hall : \beta(M).-Hall(M) M`_\beta. +Proof. +have [H hallH] := Hall_exists \beta(M) solM; have [sHM bH _]:= and3P hallH. +rewrite [M`_\beta](sub_pHall hallH) ?pcore_pgroup ?pcore_sub //=. +rewrite -(setIidPl sMbMs) pcore_setI_normal ?pcore_normal //. +have sH: \sigma(M).-group H := sub_pgroup (beta_sub_sigma maxM) bH. +have sHMs: H \subset M`_\sigma by rewrite (sub_Hall_pcore hallMs). +rewrite -pcoreNK -bigcap_p'core subsetI sHMs. +apply/bigcapsP=> p b'p; have [_ hallKp' _] := beta_max_pdiv b'p. +by rewrite (sub_Hall_pcore hallKp') ?(pi_p'group bH). +Qed. + +(* This is B & G, Lemma 10.8(a), second part. *) +Lemma Mbeta_Hall_G : \beta(M).-Hall(G) M`_\beta. +Proof. +apply: (subHall_Hall (Msigma_Hall_G maxM) (beta_sub_sigma maxM)). +exact: pHall_subl sMbMs (pcore_sub _ _) Mbeta_Hall. +Qed. + +(* This is an equivalent form of B & G, Lemma 10.8(b), which is used directly *) +(* later in the proof (e.g., Corollary 10.9a below, and Lemma 12.11), and is *) +(* proved as an intermediate step of the proof of of 12.8(b). *) +Lemma Mbeta_quo_nil : nilpotent (M^`(1) / M`_\beta). +Proof. +have /and3P[_ bMb b'M'Mb] := pHall_subl Mbeta_der1 sM'M Mbeta_Hall. +apply: nilpotentS (Fitting_nil (M^`(1) / M`_\beta)) => /=. +rewrite -{1}[_ / _]Sylow_gen gen_subG. +apply/bigcupsP=> Q /SylowP[q _ /and3P[sQM' qQ _]]. +apply: subset_trans (pcore_sub q _). +rewrite p_core_Fitting -pcoreNK -bigcap_p'core subsetI sQM' /=. +apply/bigcapsP=> [[p /= _] q'p]; have [b_p | b'p] := boolP (p \in \beta(M)). + by rewrite pcore_pgroup_id ?(pi'_p'group _ b_p) // /pgroup card_quotient. +have p'Mb: p^'.-group M`_\beta := pi_p'group bMb b'p. +rewrite sub_Hall_pcore ?(pi_p'group qQ) {Q qQ sQM'}//. +rewrite pquotient_pcore ?quotient_pHall ?(subset_trans (pcore_sub _ _)) //. +by have [-> _ _] := beta_max_pdiv b'p. +Qed. + +(* This is B & G, Lemma 10.8(b), generalized to arbitrary beta'-subgroups of *) +(* M^`(1) (which includes Hall beta'-subgroups of M^`(1) and M`_\beta). *) +Lemma beta'_der1_nil H : \beta(M)^'.-group H -> H \subset M^`(1) -> nilpotent H. +Proof. +move=> b'H sHM'; have [_ bMb _] := and3P Mbeta_Hall. +have{b'H} tiMbH: M`_\beta :&: H = 1 := coprime_TIg (pnat_coprime bMb b'H). +rewrite {tiMbH}(isog_nil (quotient_isog (subset_trans sHM' nMbM') tiMbH)). +exact: nilpotentS (quotientS _ sHM') Mbeta_quo_nil. +Qed. + +(* This is B & G, Corollary 10.9(a). *) +Corollary beta'_cent_Sylow p q X : + p \notin \beta(M) -> q \notin \beta(M) -> q.-group X -> + (p != q) && (X \subset M^`(1)) || (p < q) && (X \subset M) -> + [/\ (*a1*) exists2 P, p.-Sylow(M`_\sigma) (gval P) & X \subset 'C(P), + (*a2*) p \in \alpha(M) -> 'C_M(X)%G \in 'U + & (*a3*) q.-Sylow(M^`(1)) X -> + exists2 P, p.-Sylow(M^`(1)) (gval P) & P \subset 'N_M(X)^`(1)]. +Proof. +move=> b'p b'q qX q'p_sXM'; pose pq : nat_pred := pred2 p q. +have [q'p sXM]: p \in q^' /\ X \subset M. + case/orP: q'p_sXM' => /andP[q'p /subset_trans-> //]. + by rewrite !inE neq_ltn q'p. +have sXM'M: X <*> M^`(1) \subset M by rewrite join_subG sXM. +have solXM': solvable (X <*> M^`(1)) := solvableS sXM'M solM. +have pqX: pq.-group X by rewrite (pi_pgroup qX) ?inE ?eqxx ?orbT. +have{solXM' pqX} [W /= hallW sXW] := Hall_superset solXM' (joing_subl _ _) pqX. +have [sWXM' pqW _] := and3P hallW; have sWM := subset_trans sWXM' sXM'M. +have{b'q} b'W: \beta(M)^'.-group W. (* GG -- Coq diverges on b'p <> b'q *) + by apply: sub_pgroup pqW => r /pred2P[]->; [exact: b'p | exact: b'q]. +have nilM'W: nilpotent (M^`(1) :&: W). + by rewrite beta'_der1_nil ?subsetIl ?(pgroupS (subsetIr _ _)). +have{nilM'W} nilW: nilpotent W. + do [case/orP: q'p_sXM'=> /andP[]] => [_ sXM' | lt_pq _]. + by rewrite -(setIidPr sWXM') (joing_idPr sXM'). + pose Wq := 'O_p^'(M) :&: W; pose Wp := 'O_p(M^`(1) :&: W). + have nMp'M := char_norm (pcore_char p^' M). + have nMp'W := subset_trans sWM nMp'M. + have sylWq: q.-Sylow(W) Wq. + have [sWqMp' sWp'W] := subsetIP (subxx Wq). + have [Q sylQ] := Sylow_exists q W; have [sQW qQ _] := and3P sylQ. + rewrite [Wq](sub_pHall sylQ _ _ (subsetIr _ W)) //= -/Wq. + apply/pgroupP=> r r_pr r_dv_Wp'. + have:= pgroupP (pgroupS sWqMp' (pcore_pgroup _ _)) r r_pr r_dv_Wp'. + by apply/implyP; rewrite implyNb; exact: (pgroupP (pgroupS sWp'W pqW)). + have [[_ _ max_p] sQM] := (beta_max_pdiv b'p, subset_trans sQW sWM). + rewrite subsetI sQW -quotient_sub1 ?(subset_trans sQM nMp'M) //. + apply: contraLR lt_pq; rewrite -leqNgt andbT subG1 -rank_gt0. + rewrite (rank_pgroup (quotient_pgroup _ qQ)) p_rank_gt0 => piQb_q. + exact: max_p (piSg (quotientS _ sQM) piQb_q). + have nM'W: W \subset 'N(M^`(1)) by rewrite (subset_trans sWM) ?der_norm. + have qWWM': q.-group (W / (M^`(1) :&: W)). + rewrite (isog_pgroup _ (second_isog _)) ?(pgroupS (quotientS _ sWXM')) //=. + by rewrite (quotientYidr (subset_trans sXW nM'W)) quotient_pgroup. + have{qWWM'} sylWp: p.-Sylow(W) Wp. + rewrite /pHall pcore_pgroup (subset_trans (pcore_sub _ _)) ?subsetIr //=. + rewrite -(Lagrange_index (subsetIr _ _) (pcore_sub _ _)) pnat_mul //. + rewrite -(divgS (pcore_sub _ _)) -card_quotient ?normsI ?normG //= -pgroupE. + rewrite (pi_p'group qWWM') //= -(dprod_card (nilpotent_pcoreC p nilM'W)). + by rewrite mulKn ?cardG_gt0 // -pgroupE pcore_pgroup. + have [[sWqW qWq _] [sWpW pWp _]] := (and3P sylWq, and3P sylWp). + have <-: Wp * Wq = W. + apply/eqP; rewrite eqEcard mul_subG //= -(partnC q (cardG_gt0 W)). + rewrite (coprime_cardMg (p'nat_coprime (pi_pnat pWp _) qWq)) //. + rewrite (card_Hall sylWp) (card_Hall sylWq) -{2}(part_pnat_id pqW) -partnI. + rewrite mulnC (@eq_partn _ p) // => r. + by rewrite !inE andb_orl andbN orbF; apply: andb_idr; move/eqP->. + apply: nilpotentS (mul_subG _ _) (Fitting_nil W). + rewrite Fitting_max ?(pgroup_nil pWp) //. + by rewrite (char_normal_trans (pcore_char _ _)) //= setIC norm_normalI. + by rewrite Fitting_max ?(pgroup_nil qWq) //= setIC norm_normalI. +have part1: exists2 P : {group gT}, p.-Sylow(M`_\sigma) P & X \subset 'C(P). + have sMsXM' := subset_trans sMsM' (joing_subr X _). + have nsMsXM': M`_\sigma <| X <*> M^`(1) := normalS sMsXM' sXM'M nsMsM. + have sylWp: p.-Hall(M`_\sigma) ('O_p(W) :&: M`_\sigma). + rewrite setIC (Sylow_setI_normal nsMsXM') //. + exact: subHall_Sylow hallW (predU1l _ _) (nilpotent_pcore_Hall p nilW). + have [_ _ cWpWp' _] := dprodP (nilpotent_pcoreC p nilW). + exists ('O_p(W) :&: M`_\sigma)%G; rewrite ?(centSS _ _ cWpWp') ?subsetIl //. + by rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ _)) ?(pi_p'group qX). +split=> // [a_p | {part1}sylX]. + have ltCMX_G := sub_proper_trans (subsetIl M 'C(X)) ltMG. + have [P sylP cPX] := part1; have s_p := alpha_sub_sigma maxM a_p. + have{sylP} sylP := subHall_Sylow hallMs s_p sylP. + apply: rank3_Uniqueness ltCMX_G (leq_trans a_p _). + by rewrite -(rank_Sylow sylP) rankS //= subsetI (pHall_sub sylP) // centsC. +do [move: sWXM'; rewrite (joing_idPr (pHall_sub sylX)) => sWM'] in hallW. +have nMbX: X \subset 'N(M`_\beta) := subset_trans sXM (normal_norm nsMbM). +have nsMbXM : M`_\beta <*> X <| M. + rewrite -{2}(quotientGK nsMbM) -quotientYK ?cosetpre_normal //=. + rewrite (eq_Hall_pcore _ (quotient_pHall nMbX sylX)); last first. + exact: nilpotent_pcore_Hall Mbeta_quo_nil. + by rewrite (char_normal_trans (pcore_char _ _)) ?quotient_normal ?der_normal. +pose U := 'N_M(X); have defM: M`_\beta * U = M. + have sXU : X \subset U by rewrite subsetI sXM normG. + rewrite -[U](mulSGid sXU) /= -/U mulgA -norm_joinEr //. + apply: Frattini_arg nsMbXM (pHall_subl (joing_subr _ X) _ sylX). + by rewrite join_subG Mbeta_der1 (pHall_sub sylX). +have sWpU: 'O_p(W) \subset U. + rewrite (subset_trans (pcore_sub _ _)) // subsetI sWM normal_norm //=. + have sylX_W: q.-Sylow(W) X := pHall_subl sXW sWM' sylX. + by rewrite (eq_Hall_pcore (nilpotent_pcore_Hall q nilW) sylX_W) pcore_normal. +have sylWp: p.-Sylow(M^`(1)) 'O_p(W). + exact: subHall_Sylow hallW (predU1l _ _) (nilpotent_pcore_Hall p nilW). +exists 'O_p(W)%G; rewrite //= -(setIidPl (pHall_sub sylWp)). +rewrite (pprod_focal_coprime defM) ?pcore_normal ?subsetIr //. +exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat (pcore_pgroup _ _) _). +Qed. + +(* This is B & G, Corollary 10.9(b). *) +Corollary nonuniq_norm_Sylow_pprod p H S : + H \in 'M -> H :!=: M -> p.-Sylow(G) S -> 'N(S) \subset H :&: M -> + M`_\beta * (H :&: M) = M /\ \alpha(M) =i \beta(M). +Proof. +move=> maxH neqHM sylS_G sN_HM; have [sNH sNM] := subsetIP sN_HM. +have [sSM sSH] := (subset_trans (normG S) sNM, subset_trans (normG S) sNH). +have [sylS pS] := (pHall_subl sSM (subsetT M) sylS_G, pHall_pgroup sylS_G). +have sMp: p \in \sigma(M) by apply/exists_inP; exists S. +have aM'p: p \in \alpha(M)^'. + apply: contra neqHM; rewrite !inE -(rank_Sylow sylS) => rS. + have uniqS: S \in 'U := rank3_Uniqueness (mFT_pgroup_proper pS) rS. + by rewrite (eq_uniq_mmax (def_uniq_mmax uniqS maxM sSM) maxH sSH). +have sSM': S \subset M^`(1). + by rewrite (subset_trans _ sMsM') ?(sub_Hall_pcore hallMs) ?(pi_pgroup pS). +have nMbS := subset_trans sSM (normal_norm nsMbM). +have nMbSM: M`_\beta <*> S <| M. + rewrite -{2}(quotientGK nsMbM) -quotientYK ?cosetpre_normal //=. + have sylS_M' := pHall_subl sSM' sM'M sylS. + rewrite (eq_Hall_pcore _ (quotient_pHall nMbS sylS_M')); last first. + exact: nilpotent_pcore_Hall Mbeta_quo_nil. + by rewrite (char_normal_trans (pcore_char _ _)) ?quotient_normal ?der_normal. +have defM: M`_\beta * 'N_M(S) = M. + have sSNM: S \subset 'N_M(S) by rewrite subsetI sSM normG. + rewrite -(mulSGid sSNM) /= mulgA -norm_joinEr //. + by rewrite (Frattini_arg _ (pHall_subl _ _ sylS_G)) ?joing_subr ?subsetT. +split=> [|q]. + apply/eqP; rewrite setIC eqEsubset mulG_subG subsetIl pcore_sub /=. + by rewrite -{1}defM mulgS ?setIS. +apply/idP/idP=> [aMq|]; last exact: beta_sub_alpha. +apply: contraR neqHM => bM'q; have bM'p := contra (@beta_sub_alpha _ M p) aM'p. +have [|_ uniqNM _] := beta'_cent_Sylow bM'q bM'p pS. + by apply: contraR aM'p; rewrite sSM'; case: eqP => //= <- _. +rewrite (eq_uniq_mmax (def_uniq_mmax (uniqNM aMq) maxM (subsetIl _ _)) maxH) //. +by rewrite subIset ?(subset_trans (cent_sub _)) ?orbT. +Qed. + +(* This is B & G, Proposition 10.10. *) +Proposition max_normed_2Elem_signaliser p q (A Q : {group gT}) : + p != q -> A \in 'E_p^2(G) :&: 'E*_p(G) -> Q \in |/|*(A; q) -> + q %| #|'C(A)| -> + exists2 P : {group gT}, p.-Sylow(G) P /\ A \subset P + & [/\ (*a*) 'O_p^'('C(P)) * ('N(P) :&: 'N(Q)) = 'N(P), + (*b*) P \subset 'N(Q)^`(1) + & (*c*) q.-narrow Q -> P \subset 'C(Q)]. +Proof. +move=> neq_pq /setIP[Ep2A EpmA] maxQ piCAq. +have [_ abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have [p_pr oA] := (pnElem_prime Ep2A, card_pnElem Ep2A). +have{dimA} rA2: 'r(A) = 2 by rewrite (rank_abelem abelA). +have{EpmA} ncA: normed_constrained A. + have ntA: A :!=: 1 by rewrite -rank_gt0 rA2. + exact: plength_1_normed_constrained ntA EpmA (mFT_proper_plength1 _). +pose pi := \pi(A); pose K := 'O_pi^'('C(A)). +have def_pi : pi =i (p : nat_pred). + by move=> r; rewrite !inE /= oA primes_exp ?primes_prime ?inE. +have pi'q : q \notin pi by rewrite def_pi !inE eq_sym. +have transKA: [transitive K, on |/|*(A; q) | 'JG]. + by rewrite normed_constrained_rank2_trans // (center_idP cAA) rA2. +have [P0 sylP0 sAP0] := Sylow_superset (subsetT _) pA. +have pP0: p.-group P0 := pHall_pgroup sylP0. +have piP0: pi.-group P0 by rewrite (eq_pgroup _ def_pi). +have{pP0} snAP0: A <|<| P0 := nilpotent_subnormal (pgroup_nil pP0) sAP0. +have{pi'q snAP0 ncA piP0} [//|] := normed_trans_superset ncA pi'q snAP0 piP0. +rewrite /= -/pi -/K => -> transKP submaxPA maxPfactoring. +have{transKP} [Q0 maxQ0 _] := imsetP transKP. +have{transKA} [k Kk defQ] := atransP2 transKA (subsetP submaxPA _ maxQ0) maxQ. +set P := P0 :^ k; have{sylP0} sylP: p.-Sylow(G) P by rewrite pHallJ ?in_setT. +have nAK: K \subset 'N(A) by rewrite cents_norm ?pcore_sub. +have{sAP0 nAK K Kk} sAP: A \subset P by rewrite -(normsP nAK k Kk) conjSg. +exists [group of P] => //. +have{maxPfactoring} [sPNQ' defNP] := maxPfactoring _ maxQ0. +move/(congr1 ('Js%act^~ k)): defNP sPNQ'; rewrite -(conjSg _ _ k) /=. +rewrite conjsMg !conjIg !conjsRg -!derg1 -!normJ -pcoreJ -centJ -/P. +rewrite -(congr_group defQ) (eq_pcore _ (eq_negn def_pi)) => defNP sPNQ'. +have{sPNQ'} sPNQ': P \subset 'N(Q)^`(1). + by rewrite (setIidPl (mFT_Sylow_der1 sylP)) in sPNQ'. +split=> // narrowQ; have [-> | ntQ] := eqsVneq Q 1; first exact: cents1. +pose AutQ := conj_aut Q @* 'N(Q). +have qQ: q.-group Q by case/mem_max_normed: maxQ. +have ltNG: 'N(Q) \proper G by rewrite mFT_norm_proper // (mFT_pgroup_proper qQ). +have{ltNG} qAutQ': q.-group AutQ^`(1). + have qAutQq: q.-group 'O_q(AutQ) := pcore_pgroup _ _. + rewrite (pgroupS _ qAutQq) // der1_min ?gFnorm //. + have solAutQ: solvable AutQ by rewrite morphim_sol -?mFT_sol_proper. + have [oddQ oddAutQ]: odd #|Q| /\ odd #|AutQ| by rewrite morphim_odd mFT_odd. + by have /(Aut_narrow qQ)[] := Aut_conj_aut Q 'N(Q). +have nQP: P \subset 'N(Q) := subset_trans sPNQ' (der_sub 1 _). +rewrite (sameP setIidPl eqP) eqEsubset subsetIl /=. +rewrite -quotient_sub1 ?normsI ?normG ?norms_cent //= -ker_conj_aut subG1. +rewrite trivg_card1 (card_isog (first_isog_loc _ _)) //= -trivg_card1 -subG1. +have q'AutP: q^'.-group (conj_aut Q @* P). + by rewrite morphim_pgroup //; apply: pi_pnat (pHall_pgroup sylP) _. +rewrite -(coprime_TIg (pnat_coprime qAutQ' q'AutP)) subsetI subxx. +by rewrite /= -morphim_der // morphimS. +Qed. + +(* Notation for Proposition 11, which is the last to appear in this segment. *) +Local Notation sigma' := \sigma(gval M)^'. + +(* This is B & G, Proposition 10.11(a). *) +Proposition sigma'_not_uniq K : K \subset M -> sigma'.-group K -> K \notin 'U. +Proof. +move=> sKM sg'K; have [E hallE sKE] := Hall_superset solM sKM sg'K. +have [sEM sg'E _] := and3P hallE. +have rEle2: 'r(E) <= 2. + have [q _ ->] := rank_witness E; rewrite leqNgt; apply/negP=> rEgt2. + have: q \in sigma' by rewrite (pnatPpi sg'E) // -p_rank_gt0 -(subnKC rEgt2). + by rewrite inE /= alpha_sub_sigma //; apply: leq_trans (p_rankS q sEM). +have [E1 | ntE]:= eqsVneq E 1. + by apply: contraL (@uniq_mmax_neq1 _ K) _; rewrite -subG1 -E1. +pose p := max_pdiv #|E|; pose P := 'O_p(E). +have piEp: p \in \pi(E) by rewrite pi_max_pdiv cardG_gt1. +have sg'p: p \in sigma' by rewrite (pnatPpi sg'E). +have sylP: p.-Sylow(E) P. + rewrite rank2_max_pcore_Sylow ?mFT_odd ?(solvableS sEM solM) //. + exact: leq_trans (rankS (Fitting_sub E)) rEle2. +apply: contra (sg'p) => uniqK; apply/existsP; exists [group of P]. +have defMK := def_uniq_mmax uniqK maxM (subset_trans sKE sEM). +rewrite (subHall_Sylow hallE) // (sub_uniq_mmax defMK) //; last first. + rewrite mFT_norm_proper ?(mFT_pgroup_proper (pcore_pgroup _ _)) //. + by rewrite -cardG_gt1 (card_Hall sylP) p_part_gt1. +by rewrite (subset_trans sKE) // gFnorm. +Qed. + +(* This is B & G, Proposition 10.11(b). *) +Proposition sub'cent_sigma_rank1 K : + K \subset M -> sigma'.-group K -> 'r('C_K(M`_\sigma)) <= 1. +Proof. +move=> sKM sg'K; rewrite leqNgt; apply/rank_geP=> [[A /nElemP[p Ep2A]]]. +have p_pr := pnElem_prime Ep2A. +have [sACKMs abelA dimA] := pnElemP Ep2A; rewrite subsetI centsC in sACKMs. +have{sACKMs} [sAK cAMs]: A \subset K /\ M`_\sigma \subset 'C(A) := andP sACKMs. +have sg'p: p \in sigma'. + by rewrite (pgroupP (pgroupS sAK sg'K)) // (card_pnElem Ep2A) dvdn_mull. +have [Ms1 | [q q_pr q_dvd_Ms]] := trivgVpdiv M`_\sigma. + by case/eqP: (Msigma_neq1 maxM). +have sg_q: q \in \sigma(M) := pgroupP (pcore_pgroup _ _) _ q_pr q_dvd_Ms. +have neq_pq: p != q by apply: contraNneq sg'p => ->. +have [Q sylQ] := Sylow_exists q M`_\sigma; have [sQMs qQ _] := and3P sylQ. +have cAQ: Q \subset 'C(A) := subset_trans sQMs cAMs. +have{q_dvd_Ms} q_dv_CA: q %| #|'C(A)|. + rewrite (dvdn_trans _ (cardSg cAQ)) // -(part_pnat_id (pnat_id q_pr)). + by rewrite (card_Hall sylQ) partn_dvd. +have{cAQ} maxQ: Q \in |/|*(A; q). + rewrite inE; apply/maxgroupP; rewrite qQ cents_norm 1?centsC //. + split=> // Y /andP[qY _] sQY; apply: sub_pHall qY sQY (subsetT Y). + by rewrite (subHall_Sylow (Msigma_Hall_G maxM)). +have sNQM: 'N(Q) \subset M. + by rewrite (norm_sigma_Sylow sg_q) // (subHall_Sylow hallMs). +have rCAle2: 'r('C(A)) <= 2. + apply: contraR (sigma'_not_uniq sKM sg'K); rewrite -ltnNge => rCAgt2. + apply: uniq_mmaxS sAK (sub_mmax_proper maxM sKM) _. + by apply: cent_rank3_Uniqueness rCAgt2; rewrite (rank_abelem abelA) dimA. +have max2A: A \in 'E_p^2(G) :&: 'E*_p(G). + rewrite 3!inE subsetT abelA dimA; apply/pmaxElemP; rewrite inE subsetT. + split=> // Y /pElemP[_ abelY /eqVproper[]//ltAY]. + have [pY cYY _] := and3P abelY. + suffices: 'r_p('C(A)) > 2 by rewrite ltnNge (leq_trans (p_rank_le_rank p _)). + rewrite -dimA (leq_trans (properG_ltn_log pY ltAY)) //. + by rewrite logn_le_p_rank // inE centsC (subset_trans (proper_sub ltAY)). +have{rCAle2 cAMs} Ma1: M`_\alpha = 1. + apply: contraTeq rCAle2; rewrite -rank_gt0 -ltnNge. + have [r _ ->] := rank_witness M`_\alpha; rewrite p_rank_gt0. + move/(pnatPpi (pcore_pgroup _ _))=> a_r; apply: (leq_trans a_r). + have [R sylR] := Sylow_exists r M`_\sigma. + have sylR_M: r.-Sylow(M) R. + by rewrite (subHall_Sylow (Msigma_Hall maxM)) ?alpha_sub_sigma. + rewrite -(p_rank_Sylow sylR_M) (p_rank_Sylow sylR). + by rewrite (leq_trans (p_rank_le_rank r _)) // rankS // centsC. +have{Ma1} nilM': nilpotent M^`(1). + by rewrite (isog_nil (quotient1_isog _)) -Ma1 Malpha_quo_nil. +have{max2A maxQ neq_pq q_dv_CA} [P [sylP sAP] sPNQ']: + exists2 P : {group gT}, p.-Sylow(G) P /\ A \subset P & P \subset 'N(Q)^`(1). +- by case/(max_normed_2Elem_signaliser neq_pq): maxQ => // P ? []; exists P. +have{sNQM} defP: 'O_p(M^`(1)) = P. + rewrite (nilpotent_Hall_pcore nilM' (pHall_subl _ _ sylP)) ?subsetT //. + by rewrite (subset_trans sPNQ') ?dergS. +have charP: P \char M by rewrite -defP (char_trans (pcore_char p _)) ?der_char. +have [sPM nsPM] := (char_sub charP, char_normal charP). +case/exists_inP: sg'p; exists P; first exact: pHall_subl (subsetT M) sylP. +by rewrite (mmax_normal maxM) // -rank_gt0 ltnW // -dimA -rank_abelem ?rankS. +Qed. + +(* This is B & G, Proposition 10.11(c). *) +Proposition sub'cent_sigma_cyclic K (Y := 'C_K(M`_\sigma) :&: M^`(1)) : + K \subset M -> sigma'.-group K -> cyclic Y /\ Y <| M. +Proof. +move=> sKM sg'K; pose Z := 'O_sigma'('F(M)). +have nsZM: Z <| M := char_normal_trans (pcore_char _ _) (Fitting_normal M). +have [sZM nZM] := andP nsZM; have Fnil := Fitting_nil M. +have rZle1: 'r(Z) <= 1. + apply: leq_trans (rankS _) (sub'cent_sigma_rank1 sZM (pcore_pgroup _ _)). + rewrite subsetI subxx (sameP commG1P trivgP) /=. + rewrite -(TI_pcoreC \sigma(M) M 'F(M)) subsetI commg_subl commg_subr. + by rewrite (subset_trans sZM) ?gFnorm ?(subset_trans (pcore_sub _ _)). +have{rZle1} cycZ: cyclic Z. + have nilZ: nilpotent Z := nilpotentS (pcore_sub _ _) Fnil. + by rewrite nil_Zgroup_cyclic // odd_rank1_Zgroup // mFT_odd. +have cZM': M^`(1) \subset 'C_M(Z). + rewrite der1_min ?normsI ?normG ?norms_cent //= -ker_conj_aut. + rewrite (isog_abelian (first_isog_loc _ _)) //. + by rewrite (abelianS (Aut_conj_aut _ _)) // Aut_cyclic_abelian. +have sYF: Y \subset 'F(M). + apply: subset_trans (cent_sub_Fitting (mmax_sol maxM)). + have [_ /= <- _ _] := dprodP (nilpotent_pcoreC \sigma(M) Fnil). + by rewrite centM setICA setISS // setIC subIset ?centS // pcore_Fitting. +have{sYF} sYZ: Y \subset Z. + rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ Fnil)) //=. + by rewrite -setIA (pgroupS (subsetIl K _)). +by rewrite (cyclicS sYZ cycZ) (char_normal_trans _ nsZM) // sub_cyclic_char. +Qed. + +(* This is B & G, Proposition 10.11(d). *) +Proposition commG_sigma'_1Elem_cyclic p K P (K0 := [~: K, P]) : + K \subset M -> sigma'.-group K -> p \in sigma' -> P \in 'E_p^1('N_M(K)) -> + 'C_(M`_\sigma)(P) = 1 -> p^'.-group K -> abelian K -> + [/\ K0 \subset 'C(M`_\sigma), cyclic K0 & K0 <| M]. +Proof. +move=> sKM sg'K sg'p EpP regP p'K cKK. +have nK0P: P \subset 'N(K0) := commg_normr P K. +have p_pr := pnElem_prime EpP; have [sPMN _ oP] := pnElemPcard EpP. +have [sPM nKP]: P \subset M /\ P \subset 'N(K) by apply/subsetIP. +have /andP[sMsM nMsM]: M`_\sigma <| M := pcore_normal _ _. +have sK0K: K0 \subset K by rewrite commg_subl. +have [sK0M sg'K0]:= (subset_trans sK0K sKM, pgroupS sK0K sg'K). +have [nMsK0 nMsP] := (subset_trans sK0M nMsM, subset_trans sPM nMsM). +have coKP: coprime #|K| #|P| by rewrite oP coprime_sym prime_coprime -?p'natE. +have coK0Ms: coprime #|K0| #|M`_\sigma|. + by rewrite coprime_sym (pnat_coprime (pcore_pgroup _ _)). +have nilK0Ms: nilpotent (K0 <*> M`_\sigma). + have mulK0MsP: K0 <*> M`_\sigma ><| P = K0 <*> M`_\sigma <*> P. + rewrite sdprodEY ?normsY // coprime_TIg //= norm_joinEl //. + rewrite coprime_cardMg // coprime_mull (coprimeSg sK0K) //. + by rewrite oP (pnat_coprime (pcore_pgroup _ _)) ?pnatE. + apply: (prime_Frobenius_sol_kernel_nil mulK0MsP); rewrite ?oP //=. + by rewrite (solvableS _ solM) // !join_subG sK0M pcore_sub. + rewrite norm_joinEl // -subcent_TImulg ?subsetI ?nK0P //. + by rewrite coprime_abel_cent_TI ?mul1g. + exact: coprime_TIg. +have cMsK0: K0 \subset 'C(M`_\sigma). + rewrite (sub_nilpotent_cent2 nilK0Ms) ?joing_subl ?joing_subr //. + exact: pnat_coprime (pcore_pgroup _ _) sg'K0. +have [cycY nsYM] := sub'cent_sigma_cyclic sK0M sg'K0. +set Y := _ :&: _ in cycY nsYM. +have sK0Y: K0 \subset Y by rewrite !subsetI subxx cMsK0 commgSS. +split=> //; first exact: cyclicS sK0Y cycY. +by apply: char_normal_trans nsYM; rewrite sub_cyclic_char. +Qed. + +End AnotherMaximal. + +(* This is B & G, Lemma 10.12. *) +Lemma sigma_disjoint M H : + M \in 'M -> H \in 'M -> gval H \notin M :^: G -> + [/\ (*a*) M`_\alpha :&: H`_\sigma = 1, + [predI \alpha(M) & \sigma(H)] =i pred0 + & (*b*) nilpotent M`_\sigma -> + M`_\sigma :&: H`_\sigma = 1 + /\ [predI \sigma(M) & \sigma(H)] =i pred0]. +Proof. +move=> maxM maxH notjMH. +suffices sigmaMHnil p: p \in [predI \sigma(M) & \sigma(H)] -> + p \notin \alpha(M) /\ ~~ nilpotent M`_\sigma. +- have a2: [predI \alpha(M) & \sigma(H)] =i pred0. + move=> p; apply/andP=> [[/= aMp sHp]]. + by case: (sigmaMHnil p); rewrite /= ?aMp // inE /= alpha_sub_sigma. + split=> // [|nilMs]. + rewrite coprime_TIg // (pnat_coprime (pcore_pgroup _ _)) //. + apply: sub_in_pnat (pcore_pgroup _ _) => p _ sHp. + by apply: contraFN (a2 p) => aMp; rewrite inE /= sHp andbT. + have b2: [predI \sigma(M) & \sigma(H)] =i pred0. + by move=> p; apply/negP; case/sigmaMHnil => _; rewrite nilMs. + rewrite coprime_TIg // (pnat_coprime (pcore_pgroup _ _)) //. + apply: sub_in_pnat (pcore_pgroup _ _) => p _ sHp. + by apply: contraFN (b2 p) => bMp; rewrite inE /= sHp andbT. +case/andP=> sMp sHp; have [S sylS]:= Sylow_exists p M. +have [sSM pS _] := and3P sylS. +have sylS_G: p.-Sylow(G) S := sigma_Sylow_G maxM sMp sylS. +have [g sSHg]: exists g, S \subset H :^ g. + have [Sg' sylSg']:= Sylow_exists p H. + have [g _ ->] := Sylow_trans (sigma_Sylow_G maxH sHp sylSg') sylS_G. + by exists g; rewrite conjSg (pHall_sub sylSg'). +have{notjMH} neqHgM: H :^ g != M. + by apply: contraNneq notjMH => <-; rewrite orbit_sym mem_orbit ?in_setT. +do [split; apply: contra neqHgM] => [|nilMs]. + rewrite !inE -(p_rank_Sylow sylS) -rank_pgroup //= => rS_gt3. + have uniqS: S \in 'U := rank3_Uniqueness (mFT_pgroup_proper pS) rS_gt3. + have defUS: 'M(S) = [set M] := def_uniq_mmax uniqS maxM sSM. + by rewrite (eq_uniq_mmax defUS _ sSHg) ?mmaxJ. +have nsSM: S <| M. + have nsMsM: M`_\sigma <| M by exact: pcore_normal. + have{sylS} sylS: p.-Sylow(M`_\sigma) S. + apply: pHall_subl (pcore_sub _ _) sylS => //. + by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) ?(pi_pgroup pS). + rewrite (nilpotent_Hall_pcore nilMs sylS). + by rewrite (char_normal_trans (pcore_char _ _)). +have sNS_Hg: 'N(S) \subset H :^ g. + rewrite -sub_conjgV -normJ (norm_sigma_Sylow sHp) //. + by rewrite (pHall_subl _ (subsetT _)) ?sub_conjgV // pHallJ ?in_setT. +have ltHg: H :^ g \proper G by rewrite mmax_proper ?mmaxJ //. +rewrite (mmax_max maxM ltHg) // -(mmax_normal maxM nsSM) //. +by apply: contraTneq sNS_Hg => ->; rewrite norm1 proper_subn. +Qed. + +(* This is B & G, Lemma 10.13. *) +Lemma basic_p2maxElem_structure p A P : + A \in 'E_p^2(G) :&: 'E*_p(G) -> p.-group P -> A \subset P -> ~~ abelian P -> + let Z0 := ('Ohm_1('Z(P)))%G in + [/\ (*a*) Z0 \in 'E_p^1(A), + (*b*) exists Y : {group gT}, + [/\ cyclic Y, Z0 \subset Y + & forall A0, A0 \in 'E_p^1(A) :\ Z0 -> A0 \x Y = 'C_P(A)] + & (*c*) [transitive 'N_P(A), on 'E_p^1(A) :\ Z0| 'JG]]. +Proof. +case/setIP=> Ep2A maxA pP sAP not_cPP Z0; set E1A := 'E_p^1(A). +have p_pr: prime p := pnElem_prime Ep2A; have [_ abelA dimA] := pnElemP Ep2A. +have [oA [pA cAA _]] := (card_pnElem Ep2A, and3P abelA). +have [p_gt0 p_gt1] := (prime_gt0 p_pr, prime_gt1 p_pr). +have{maxA} maxA S: + p.-group S -> A \subset S -> A \in 'E*_p(S) /\ 'Ohm_1('C_S(A)) = A. +- move=> pS sAS; suff maxAS: A \in 'E*_p(S) by rewrite (Ohm1_cent_max maxAS). + by rewrite (subsetP (pmaxElemS p (subsetT S))) // inE maxA inE. +have [S sylS sPS] := Sylow_superset (subsetT P) pP. +pose Z1 := 'Ohm_1('Z(S))%G; have sZ1Z: Z1 \subset 'Z(S) := Ohm_sub 1 _. +have [pS sAS] := (pHall_pgroup sylS, subset_trans sAP sPS). +have [maxAS defC1] := maxA S pS sAS; set C := 'C_S(A) in defC1. +have sZ0A: Z0 \subset A by rewrite -defC1 OhmS // setISS // centS. +have sZ1A: Z1 \subset A by rewrite -defC1 OhmS // setIS // centS. +have [pZ0 pZ1]: p.-group Z0 /\ p.-group Z1 by split; exact: pgroupS pA. +have sZ10: Z1 \subset Z0. + rewrite -[gval Z1]Ohm_id OhmS // subsetI (subset_trans sZ1A) //=. + by rewrite (subset_trans sZ1Z) // subIset // centS ?orbT. +have ntZ1: Z1 :!=: 1. + have: A :!=: 1 by rewrite -cardG_gt1 oA (ltn_exp2l 0). + apply: contraNneq; rewrite -subG1 -(setIidPr sZ1Z) => /TI_Ohm1. + by rewrite setIid => /(trivg_center_pgroup pS) <-. +have EpZ01: abelian C -> Z1 = Z0 /\ Z0 \in E1A. + move=> cCC; have [eqZ0A | ltZ0A] := eqVproper sZ0A. + rewrite (abelianS _ cCC) // in not_cPP. + by rewrite subsetI sPS centsC -eqZ0A (subset_trans (Ohm_sub _ _)) ?subsetIr. + have leZ0p: #|Z0| <= p ^ 1. + by rewrite (card_pgroup pZ0) leq_exp2l // -ltnS -dimA properG_ltn_log. + have [_ _ [e oZ1]] := pgroup_pdiv pZ1 ntZ1. + have{e oZ1}: #|Z1| >= p by rewrite oZ1 (leq_exp2l 1). + rewrite (geq_leqif (leqif_trans (subset_leqif_card sZ10) (leqif_eq leZ0p))). + rewrite [E1A]p1ElemE // !inE sZ0A; case/andP=> sZ01 ->. + by split=> //; apply/eqP; rewrite -val_eqE eqEsubset sZ10. +have [A1 neqA1Z EpA1]: exists2 A1, A1 != Z1 & #|Z1| = p -> A1 \in E1A. + have [oZ1 |] := #|Z1| =P p; last by exists 1%G; rewrite // eq_sym. + have [A1 defA]:= abelem_split_dprod abelA sZ1A. + have{defA} [_ defA _ tiA1Z1] := dprodP defA. + have EpZ1: Z1 \in E1A by rewrite [E1A]p1ElemE // !inE sZ1A /= oZ1. + suffices: A1 \in E1A by exists A1; rewrite // eq_sym; exact/(TIp1ElemP EpZ1). + rewrite [E1A]p1ElemE // !inE -defA mulG_subr /=. + by rewrite -(mulKn #|A1| p_gt0) -{1}oZ1 -TI_cardMg // defA oA mulKn. +pose cplA1C Y := [/\ cyclic Y, Z0 \subset Y, A1 \x Y = C & abelian C]. +have [Y [{cplA1C} cycY sZ0Y defC cCC]]: exists Y, cplA1C Y. + have [rSgt2 | rSle2] := ltnP 2 'r(S). + rewrite (rank_pgroup pS) in rSgt2; have oddS := mFT_odd S. + have max2AS: A \in 'E_p^2(S) :&: 'E*_p(S) by rewrite 3!inE sAS abelA dimA. + have oZ1: #|Z1| = p by case/Ohm1_ucn_p2maxElem: max2AS => // _ []. + have{EpA1} EpA1 := EpA1 oZ1; have [sA1A abelA1 oA1] := pnElemPcard EpA1. + have EpZ1: Z1 \in E1A by rewrite [E1A]p1ElemE // !inE sZ1A /= oZ1. + have [_ defA cA1Z tiA1Z] := dprodP (p2Elem_dprodP Ep2A EpA1 EpZ1 neqA1Z). + have defC: 'C_S(A1) = C. + rewrite /C -defA centM setICA setIC ['C_S(Z1)](setIidPl _) // centsC. + by rewrite (subset_trans sZ1Z) ?subsetIr. + have rCSA1: 'r_p('C_S(A1)) <= 2. + by rewrite defC -p_rank_Ohm1 defC1 (p_rank_abelem abelA) dimA. + have sA1S := subset_trans sA1A sAS. + have nnS: p.-narrow S by apply/implyP=> _; apply/set0Pn; exists A. + have [] := narrow_cent_dprod pS oddS rSgt2 nnS oA1 sA1S rCSA1. + set Y := _ :&: _; rewrite {}defC => cycY _ _ defC; exists [group of Y]. + have cCC: abelian C; last split=> //. + apply/center_idP; rewrite -(center_dprod defC). + rewrite (center_idP (abelem_abelian abelA1)). + by rewrite (center_idP (cyclic_abelian cycY)). + have{EpZ01} [<- _] := EpZ01 cCC; rewrite subsetI (subset_trans sZ1Z) //. + by rewrite setIS ?centS // (subset_trans (Ohm_sub 1 _)) ?ucn_sub. + have not_cSS := contra (abelianS sPS) not_cPP. + have:= mFT_rank2_Sylow_cprod sylS rSle2 not_cSS. + case=> E [_ dimE3 eE] [Y cycY [defS defY1]]. + have [[_ mulEY cEY] cYY] := (cprodP defS, cyclic_abelian cycY). + have defY: 'Z(S) = Y. + case/cprodP: (center_cprod defS) => _ <- _. + by rewrite (center_idP cYY) -defY1 mulSGid ?Ohm_sub. + have pY: p.-group Y by rewrite -defY (pgroupS (center_sub S)). + have sES: E \subset S by rewrite -mulEY mulG_subl. + have pE := pgroupS sES pS. + have defS1: 'Ohm_1(S) = E. + apply/eqP; rewrite (OhmE 1 pS) eqEsubset gen_subG andbC. + rewrite sub_gen; last by rewrite subsetI sES sub_LdivT. + apply/subsetP=> ey /LdivP[]; rewrite -mulEY. + case/imset2P=> e y Ee Yy -> eyp; rewrite groupM //. + rewrite (subsetP (center_sub E)) // -defY1 (OhmE 1 pY) mem_gen //. + rewrite expgMn in eyp; last by red; rewrite -(centsP cEY). + by rewrite (exponentP eE) // mul1g in eyp; rewrite !inE Yy eyp eqxx. + have sAE: A \subset E by rewrite -defS1 -(Ohm1_id abelA) OhmS. + have defC: A * Y = C. + rewrite /C -mulEY setIC -group_modr; last first. + by rewrite -defY subIset // orbC centS. + congr (_ * _); apply/eqP; rewrite /= setIC eqEcard subsetI sAE. + have pCEA: p.-group 'C_E(A) := pgroupS (subsetIl E _) pE. + rewrite -abelianE cAA (card_pgroup pCEA) oA leq_exp2l //= leqNgt. + apply: contraL cycY => dimCEA3. + have sAZE: A \subset 'Z(E). + rewrite subsetI sAE // centsC (sameP setIidPl eqP) eqEcard subsetIl /=. + by rewrite (card_pgroup pE) (card_pgroup pCEA) dimE3 leq_exp2l. + rewrite abelian_rank1_cyclic // -ltnNge (rank_pgroup pY). + by rewrite (p_rank_abelian p cYY) defY1 -dimA lognSg. + have cAY: Y \subset 'C(A) by apply: centSS cEY. + have cCC: abelian C by rewrite -defC abelianM cAA cYY. + have{EpZ01} [eqZ10 EpZ1] := EpZ01 cCC; rewrite -eqZ10 in EpZ1. + have sZ0Y: Z0 \subset Y by rewrite -eqZ10 -defY Ohm_sub. + have{EpA1} EpA1 := EpA1 (card_pnElem EpZ1). + have [sA1A _ oA1] := pnElemPcard EpA1. + have [_ defA _ tiA1Z] := dprodP (p2Elem_dprodP Ep2A EpA1 EpZ1 neqA1Z). + exists Y; split; rewrite // dprodE ?(centSS _ sA1A cAY) ?prime_TIg ?oA1 //. + by rewrite -(mulSGid sZ0Y) -eqZ10 mulgA defA. + apply: contraL cycY => sA1Y; rewrite abelian_rank1_cyclic // -ltnNge. + by rewrite -dimA -rank_abelem ?rankS // -defA eqZ10 mul_subG. +have{EpZ01} [eqZ10 EpZ0] := EpZ01 cCC; have oZ0 := card_pnElem EpZ0. +have{EpA1} EpA1: A1 \in E1A by rewrite EpA1 ?eqZ10. +have [sA1A _ oA1] := pnElemPcard EpA1; rewrite {}eqZ10 in neqA1Z. +have [_ defA _ tiA1Z] := dprodP (p2Elem_dprodP Ep2A EpA1 EpZ0 neqA1Z). +split=> //; first exists (P :&: Y)%G. + have sPY_Y := subsetIr P Y; rewrite (cyclicS sPY_Y) //. + rewrite subsetI (subset_trans sZ0A) //= sZ0Y. + split=> // A0 /setD1P[neqA0Z EpA0]; have [sA0A _ _] := pnElemP EpA0. + have [_ mulA0Z _ tiA0Z] := dprodP (p2Elem_dprodP Ep2A EpA0 EpZ0 neqA0Z). + have{defC} [_ defC cA1Y tiA1Y] := dprodP defC. + rewrite setIC -{2}(setIidPr sPS) setIAC. + apply: dprod_modl (subset_trans sA0A sAP); rewrite -defC dprodE /=. + - by rewrite -(mulSGid sZ0Y) !mulgA mulA0Z defA. + - rewrite (centSS (subxx Y) sA0A) // -defA centM subsetI cA1Y /=. + by rewrite sub_abelian_cent ?cyclic_abelian. + rewrite setIC -(setIidPr sA0A) setIA -defA -group_modr //. + by rewrite (setIC Y) tiA1Y mul1g setIC. +apply/imsetP; exists A1; first by rewrite 2!inE neqA1Z. +apply/eqP; rewrite eq_sym eqEcard; apply/andP; split. + apply/subsetP=> _ /imsetP[x /setIP[Px nAx] ->]. + rewrite 2!inE /E1A -(normP nAx) pnElemJ EpA1 andbT -val_eqE /=. + have nZ0P: P \subset 'N(Z0). + by rewrite (char_norm_trans (Ohm_char 1 _)) // gFnorm. + by rewrite -(normsP nZ0P x Px) (inj_eq (@conjsg_inj _ x)). +have pN: p.-group 'N_P(_) := pgroupS (subsetIl P _) pP. +have defCPA: 'N_('N_P(A))(A1) = 'C_P(A). + apply/eqP; rewrite eqEsubset andbC subsetI setIS ?cent_sub //. + rewrite subIset /=; last by rewrite orbC cents_norm ?centS. + rewrite setIAC (subset_trans (subsetIl _ _)) //= subsetI subsetIl /=. + rewrite -defA centM subsetI andbC subIset /=; last first. + by rewrite centsC (subset_trans (Ohm_sub 1 _)) ?subsetIr. + have nC_NP: 'N_P(A1) \subset 'N('C(A1)) by rewrite norms_cent ?subsetIr. + rewrite -quotient_sub1 // subG1 trivg_card1. + rewrite (pnat_1 (quotient_pgroup _ (pN _))) //. + rewrite -(card_isog (second_isog nC_NP)) /= (setIC 'C(A1)). + by apply: p'group_quotient_cent_prime; rewrite ?subsetIr ?oA1. +have sCN: 'C_P(A) \subset 'N_P(A) by rewrite setIS ?cent_sub. +have nA_NCPA: 'N_P('C_P(A)) \subset 'N_P(A). + have [_ defCPA1] := maxA P pP sAP. + by rewrite -{2}defCPA1 setIS // (char_norm_trans (Ohm_char 1 _)). +rewrite card_orbit astab1JG /= {}defCPA. +rewrite -(leq_add2l (Z0 \in E1A)) -cardsD1 EpZ0 (card_p1Elem_p2Elem Ep2A) ltnS. +rewrite dvdn_leq ?(pfactor_dvdn 1) ?indexg_gt0 // -divgS // logn_div ?cardSg //. +rewrite subn_gt0 properG_ltn_log ?pN //= (proper_sub_trans _ nA_NCPA) //. +rewrite (nilpotent_proper_norm (pgroup_nil pP)) // properEneq subsetIl andbT. +by apply: contraNneq not_cPP => <-; rewrite (abelianS (setSI _ sPS)). +Qed. + +(* This is B & G, Proposition 10.14(a). *) +Proposition beta_not_narrow p : p \in \beta(G) -> + [disjoint 'E_p^2(G) & 'E*_p(G)] + /\ (forall P, p.-Sylow(G) P -> [disjoint 'E_p^2(P) & 'E*_p(P)]). +Proof. +move/forall_inP=> nnG. +have nnSyl P: p.-Sylow(G) P -> [disjoint 'E_p^2(P) & 'E*_p(P)]. + by move/nnG; rewrite negb_imply negbK setI_eq0 => /andP[]. +split=> //; apply/pred0Pn=> [[E /andP[/= Ep2E EpmE]]]. +have [_ abelE dimE]:= pnElemP Ep2E; have pE := abelem_pgroup abelE. +have [P sylP sEP] := Sylow_superset (subsetT E) pE. +case/pred0Pn: (nnSyl P sylP); exists E; rewrite /= 2!inE sEP abelE dimE /=. +by rewrite (subsetP (pmaxElemS p (subsetT P))) // inE EpmE inE. +Qed. + +(* This is B & G, Proposition 10.14(b). *) +Proposition beta_noncyclic_uniq p R : + p \in \beta(G) -> p.-group R -> 'r(R) > 1 -> R \in 'U. +Proof. +move=> b_p pR rRgt1; have [P sylP sRP] := Sylow_superset (subsetT R) pR. +rewrite (rank_pgroup pR) in rRgt1; have [A Ep2A] := p_rank_geP rRgt1. +have [sAR abelA dimA] := pnElemP Ep2A; have p_pr := pnElem_prime Ep2A. +case: (pickP [pred F in 'E_p(P) | A \proper F]) => [F | maxA]; last first. + have [_ nnSyl] := beta_not_narrow b_p; case/pred0Pn: (nnSyl P sylP). + exists A; rewrite /= (subsetP (pnElemS p 2 sRP)) //. + apply/pmaxElemP; split=> [|F EpF]; first by rewrite inE (subset_trans sAR). + by case/eqVproper=> [// | ltAF]; case/andP: (maxA F). +case/andP=> /pElemP[_ abelF] ltAF; have [pF cFF _] := and3P abelF. +apply: uniq_mmaxS sAR (mFT_pgroup_proper pR) _. +have rCAgt2: 'r('C(A)) > 2. + rewrite -dimA (leq_trans (properG_ltn_log pF ltAF)) // -(rank_abelem abelF). + by rewrite rankS // centsC (subset_trans (proper_sub ltAF)). +by apply: cent_rank3_Uniqueness rCAgt2; rewrite (rank_abelem abelA) dimA. +Qed. + +(* This is B & G, Proposition 10.14(c). *) +Proposition beta_subnorm_uniq p P X : + p \in \beta(G) -> p.-Sylow(G) P -> X \subset P -> 'N_P(X)%G \in 'U. +Proof. +move=> b_p sylP sXP; set Q := 'N_P(X)%G. +have pP := pHall_pgroup sylP; have pQ: p.-group Q := pgroupS (subsetIl _ _) pP. +have [| rQle1] := ltnP 1 'r(Q); first exact: beta_noncyclic_uniq pQ. +have cycQ: cyclic Q. + by rewrite (odd_pgroup_rank1_cyclic pQ) ?mFT_odd -?rank_pgroup. +have defQ: P :=: Q. + apply: (nilpotent_sub_norm (pgroup_nil pP) (subsetIl _ _)). + by rewrite setIS // char_norms // sub_cyclic_char // subsetI sXP normG. +have:= forall_inP b_p P; rewrite inE negb_imply ltnNge; move/(_ sylP). +by rewrite defQ -(rank_pgroup pQ) (leq_trans rQle1). +Qed. + +(* This is B & G, Proposition 10.14(d). *) +Proposition beta_norm_sub_mmax M Y : + M \in 'M -> \beta(M).-subgroup(M) Y -> Y :!=: 1 -> 'N(Y) \subset M. +Proof. +move=> maxM /andP[sYM bY] ntY. +have [F1 | [q q_pr q_dv_FY]] := trivgVpdiv 'F(Y). + by rewrite -(trivg_Fitting (solvableS sYM (mmax_sol maxM))) F1 eqxx in ntY. +pose X := 'O_q(Y); have qX: q.-group X := pcore_pgroup q _. +have ntX: X != 1. + apply: contraTneq q_dv_FY => X1; rewrite -p'natE // -partn_eq1 //. + rewrite -(card_Hall (nilpotent_pcore_Hall q (Fitting_nil Y))). + by rewrite /= p_core_Fitting -/X X1 cards1. +have bMq: q \in \beta(M) by apply: (pgroupP (pgroupS (Fitting_sub Y) bY)). +have b_q: q \in \beta(G) by move: bMq; rewrite -predI_sigma_beta //; case/andP. +have sXM: X \subset M := subset_trans (pcore_sub q Y) sYM. +have [P sylP sXP] := Sylow_superset sXM qX; have [sPM qP _] := and3P sylP. +have sylPG: q.-Sylow(G) P by rewrite (sigma_Sylow_G maxM) ?beta_sub_sigma. +have uniqNX: 'M('N_P(X)) = [set M]. + apply: def_uniq_mmax => //; last by rewrite subIset ?sPM. + exact: (beta_subnorm_uniq b_q). +rewrite (subset_trans (char_norms (pcore_char q Y))) //. +rewrite (sub_uniq_mmax uniqNX) ?subsetIr // mFT_norm_proper //. +by rewrite (sub_mmax_proper maxM). +Qed. + +End Ten. + + diff --git a/mathcomp/odd_order/BGsection11.v b/mathcomp/odd_order/BGsection11.v new file mode 100644 index 0000000..6fccf96 --- /dev/null +++ b/mathcomp/odd_order/BGsection11.v @@ -0,0 +1,438 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import bigop finset prime fingroup morphism perm automorphism quotient. +Require Import action gproduct gfunctor pgroup cyclic center commutator. +Require Import gseries nilpotent sylow abelian maximal hall. +Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +Require Import BGsection7 BGsection10. + +(******************************************************************************) +(* This file covers B & G, section 11; it has only one definition: *) +(* exceptional_FTmaximal p M A0 A <=> *) +(* p, M and A0 satisfy the conditions of Hypothesis 11.1 in B & G, i.e., *) +(* M is an "exceptional" maximal subgroup in the terminology of B & G. *) +(* In addition, A is elementary abelian p-subgroup of M of rank 2, that *) +(* contains A0. The existence of A is guaranteed by Lemma 10.5, but as *) +(* in the only two lemmas that make use of the results in this section *) +(* (Lemma 12.3 and Theorem 12.5) A is known, we elected to make the *) +(* dependency on A explicit. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Section Section11. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). + +Implicit Types p q r : nat. +Implicit Types A E H K M N P Q R S T U V W X Y : {group gT}. + +Variables (p : nat) (M A0 A P : {group gT}). + +(* This definition corresponsd to Hypothesis 11.1, where the condition on A *) +(* has been made explicit. *) +Definition exceptional_FTmaximal := + [/\ p \in \sigma(M)^', A \in 'E_p^2(M), A0 \in 'E_p^1(A) & 'N(A0) \subset M]. + +Hypotheses (maxM : M \in 'M) (excM : exceptional_FTmaximal). +Hypotheses (sylP : p.-Sylow(M) P) (sAP : A \subset P). + +(* Splitting the excM hypothesis. *) +Let sM'p : p \in \sigma(M)^'. Proof. by case: excM. Qed. +Let Ep2A : A \in 'E_p^2(M). Proof. by case: excM. Qed. +Let Ep1A0 : A0 \in 'E_p^1(A). Proof. by case: excM. Qed. +Let sNA0_M : 'N(A0) \subset M. Proof. by case: excM. Qed. + +(* Arithmetics of p. *) +Let p_pr : prime p := pnElem_prime Ep2A. +Let p_gt1 : p > 1 := prime_gt1 p_pr. +Let p_gt0 : p > 0 := prime_gt0 p_pr. + +(* Group orders. *) +Let oA : #|A| = (p ^ 2)%N := card_pnElem Ep2A. +Let oA0 : #|A0| = p := card_pnElem Ep1A0. + +(* Structure of A. *) +Let abelA : p.-abelem A. Proof. by case/pnElemP: Ep2A. Qed. +Let pA : p.-group A := abelem_pgroup abelA. +Let cAA : abelian A := abelem_abelian abelA. + +(* Various set inclusions. *) +Let sA0A : A0 \subset A. Proof. by case/pnElemP: Ep1A0. Qed. +Let sPM : P \subset M := pHall_sub sylP. +Let sAM : A \subset M := subset_trans sAP sPM. +Let sCA0_M : 'C(A0) \subset M := subset_trans (cent_sub A0) sNA0_M. +Let sCA_M : 'C(A) \subset M := subset_trans (centS sA0A) sCA0_M. + +(* Alternative E_p^1 memberships for A0; the first is the one used to state *) +(* Hypothesis 11.1 in B & G, the second is the form expected by Lemma 10.5. *) +(* Note that #|A0| = p (oA0 above) would wokr just as well. *) +Let Ep1A0_M : A0 \in 'E_p^1(M) := subsetP (pnElemS p 1 sAM) A0 Ep1A0. +Let Ep1A0_G : A0 \in 'E_p^1(G) := subsetP (pnElemS p 1 (subsetT M)) A0 Ep1A0_M. + +(* This does not depend on exceptionalM, and could move to Section 10. *) +Lemma sigma'_Sylow_contra : p \in \sigma(M)^' -> ~~ ('N(P) \subset M). +Proof. by apply: contra => sNM; apply/exists_inP; exists P. Qed. + +(* First preliminary remark of Section 11; only depends on sM'p and sylP. *) +Let not_sNP_M: ~~ ('N(P) \subset M) := sigma'_Sylow_contra sM'p. + +(* Second preliminary remark of Section 11; only depends on sM'p, Ep1A0_M, *) +(* and sNA0_M. *) +Lemma p_rank_exceptional : 'r_p(M) = 2. +Proof. exact: sigma'_norm_mmax_rank2 (pgroupS sA0A pA) _. Qed. +Let rM := p_rank_exceptional. + +(* Third preliminary remark of Section 11. *) +Lemma exceptional_pmaxElem : A \in 'E*_p(G). +Proof. +have [_ _ dimA]:= pnElemP Ep2A. +apply/pmaxElemP; split=> [|E EpE sAE]; first by rewrite !inE subsetT. +have [//|ltAE]: A :=: E \/ A \proper E := eqVproper sAE. +have [_ abelE] := pElemP EpE; have [pE cEE _] := and3P abelE. +suffices: logn p #|E| <= 'r_p(M) by rewrite leqNgt rM -dimA properG_ltn_log. +by rewrite logn_le_p_rank // inE (subset_trans cEE) ?(subset_trans (centS sAE)). +Qed. +Let EpmA := exceptional_pmaxElem. + +(* This is B & G, Lemma 11.1. *) +Lemma exceptional_TIsigmaJ g q Q1 Q2 : + g \notin M -> A \subset M :^ g -> + q.-Sylow(M`_\sigma) Q1 -> A \subset 'N(Q1) -> + q.-Sylow(M`_\sigma :^ g) Q2 -> A \subset 'N(Q2) -> + (*a*) Q1 :&: Q2 = 1 + /\ (*b*) (forall X, X \in 'E_p^1(A) -> 'C_Q1(X) = 1 \/ 'C_Q2(X) = 1). +Proof. +move=> notMg sAMg sylQ1 nQ1A sylQ2 nQ2A. +have [-> | ntQ1] := eqsVneq Q1 1. + by split=> [|? _]; last left; apply: (setIidPl (sub1G _)). +have [sQ1Ms qQ1 _] := and3P sylQ1. +have{qQ1} [q_pr q_dv_Q1 _] := pgroup_pdiv qQ1 ntQ1. +have{sQ1Ms q_dv_Q1} sMq: q \in \sigma(M). + exact: pgroupP (pgroupS sQ1Ms (pcore_pgroup _ _)) q q_pr q_dv_Q1. +have{sylQ1} sylQ1: q.-Sylow(M) Q1. + by rewrite (subHall_Sylow (Msigma_Hall maxM)). +have sQ1M := pHall_sub sylQ1. +have{sylQ2} sylQ2g': q.-Sylow(M) (Q2 :^ g^-1). + by rewrite (subHall_Sylow (Msigma_Hall _)) // -(pHallJ2 _ _ _ g) actKV. +have sylQ2: q.-Sylow(G) Q2. + by rewrite -(pHallJ _ _ (in_setT g^-1)) (sigma_Sylow_G maxM). +suffices not_Q1_CA_Q2: gval Q2 \notin Q1 :^: 'O_\pi(A)^'('C(A)). + have ncA: normed_constrained A. + have ntA: A :!=: 1 by rewrite -cardG_gt1 oA (ltn_exp2l 0). + exact: plength_1_normed_constrained ntA EpmA (mFT_proper_plength1 _). + have q'A: q \notin \pi(A). + by apply: contraL sMq; move/(pnatPpi pA); move/eqnP->. + have maxnAq Q: q.-Sylow(G) Q -> A \subset 'N(Q) -> Q \in |/|*(A; q). + move=> sylQ; case/(max_normed_exists (pHall_pgroup sylQ)) => R maxR sQR. + have [qR _] := mem_max_normed maxR. + by rewrite -(group_inj (sub_pHall sylQ qR sQR (subsetT R))). + have maxQ1 := maxnAq Q1 (sigma_Sylow_G maxM sMq sylQ1) nQ1A. + have maxQ2 := maxnAq Q2 sylQ2 nQ2A. + have transCAQ := normed_constrained_meet_trans ncA q'A _ _ maxQ1 maxQ2. + split=> [|X EpX]. + apply: contraNeq not_Q1_CA_Q2 => ntQ12; apply/imsetP. + apply: transCAQ (sAM) (mmax_proper maxM) _ _. + by rewrite (setIidPl sQ1M). + by apply: contraNneq ntQ12 => tiQ2M; rewrite setIC -subG1 -tiQ2M setIS. + apply/pred2P; apply: contraR not_Q1_CA_Q2; case/norP=> ntCQ1 ntCQ2. + have [sXA _ oX] := pnElemPcard EpX. + apply/imsetP; apply: transCAQ (centSS _ sXA cAA) _ ntCQ1 ntCQ2 => //. + by rewrite mFT_cent_proper // -cardG_gt1 oX prime_gt1. +apply: contra notMg; case/imsetP=> k cAk defQ2. +have{cAk} Mk := subsetP sCA_M k (subsetP (pcore_sub _ _) k cAk). +have{k Mk defQ2} sQ2M: Q2 \subset M by rewrite defQ2 conj_subG. +have [sQ2g'M qQ2g' _] := and3P sylQ2g'. +by rewrite (sigma_Sylow_trans _ sylQ2g') // actKV. +Qed. + +(* This is B & G, Corollary 11.2. *) +Corollary exceptional_TI_MsigmaJ g : + g \notin M -> A \subset M :^ g -> + (*a*) M`_\sigma :&: M :^ g = 1 + /\ (*b*) M`_\sigma :&: 'C(A0 :^ g) = 1. +Proof. +move=> notMg sAMg; set Ms := M`_\sigma; set H := [group of Ms :&: M :^ g]. +have [H1 | ntH] := eqsVneq H 1. + by split=> //; apply/trivgP; rewrite -H1 setIS //= centJ conjSg. +pose q := pdiv #|H|. +suffices: #|H|`_q == 1%N by rewrite p_part_eq1 pi_pdiv cardG_gt1 ntH. +have nsMsM: Ms <| M := pcore_normal _ _; have [_ nMsM] := andP nsMsM. +have sHMs: H \subset Ms := subsetIl _ _. +have sHMsg: H \subset Ms :^ g. + rewrite -sub_conjgV (sub_Hall_pcore (Msigma_Hall _)) //. + by rewrite pgroupJ (pgroupS sHMs) ?pcore_pgroup. + by rewrite sub_conjgV subsetIr. +have nMsA := subset_trans sAM nMsM. +have nHA: A \subset 'N(H) by rewrite normsI // normsG. +have nMsgA: A \subset 'N(Ms :^ g) by rewrite normJ (subset_trans sAMg) ?conjSg. +have coMsA: coprime #|Ms| #|A|. + by rewrite oA coprime_expr ?(pnat_coprime (pcore_pgroup _ _)) ?pnatE. +have coHA: coprime #|H| #|A| := coprimeSg sHMs coMsA. +have coMsgA: coprime #|Ms :^ g| #|A| by rewrite cardJg. +have solA: solvable A := abelian_sol cAA. +have [Q0 sylQ0 nQ0A] := sol_coprime_Sylow_exists q solA nHA coHA. +have [sQ0H qQ0 _] := and3P sylQ0. +have supQ0 := sol_coprime_Sylow_subset _ _ solA (subset_trans sQ0H _) qQ0 nQ0A. +have [Q1 [sylQ1 nQ1A sQ01]] := supQ0 _ nMsA coMsA sHMs. +have [Q2 [sylQ2 nQ2A sQ02]] := supQ0 _ nMsgA coMsgA sHMsg. +have tiQ12: Q1 :&: Q2 = 1. + by have [-> _] := exceptional_TIsigmaJ notMg sAMg sylQ1 nQ1A sylQ2 nQ2A. +by rewrite -(card_Hall sylQ0) -trivg_card1 -subG1 -tiQ12 subsetI sQ01. +Qed. + +(* This is B & G, Theorem 11.3. *) +Theorem exceptional_sigma_nil : nilpotent M`_\sigma. +Proof. +have [g nPg notMg] := subsetPn not_sNP_M. +set Ms := M`_\sigma; set F := Ms <*> A0 :^ g. +have sA0gM: A0 :^ g \subset M. + by rewrite (subset_trans _ sPM) // -(normP nPg) conjSg (subset_trans sA0A). +have defF: Ms ><| A0 :^ g = F. + rewrite sdprodEY ?coprime_TIg //. + by rewrite (subset_trans sA0gM) ?gFnorm. + by rewrite cardJg oA0 (pnat_coprime (pcore_pgroup _ _)) ?pnatE. +have regA0g: 'C_Ms(A0 :^ g) = 1. + case/exceptional_TI_MsigmaJ: notMg => //. + by rewrite -sub_conjgV (subset_trans _ sPM) // sub_conjgV (normP _). +rewrite (prime_Frobenius_sol_kernel_nil defF) ?cardJg ?oA0 //. +by rewrite (solvableS _ (mmax_sol maxM)) // join_subG pcore_sub. +Qed. + +(* This is B & G, Corollary 11.4. *) +Corollary exceptional_sigma_uniq H : + H \in 'M(A) -> H`_\sigma :&: M `_\sigma != 1 -> H :=: M. +Proof. +rewrite setIC => /setIdP[maxH sAH] ntMsHs. +have [g _ defH]: exists2 g, g \in G & H :=: M :^ g. + apply/imsetP; apply: contraR ntMsHs => /sigma_disjoint[] // _ _. + by case/(_ exceptional_sigma_nil)=> ->. +rewrite defH conjGid //; apply: contraR ntMsHs => notMg. +have [|tiMsMg _] := exceptional_TI_MsigmaJ notMg; first by rewrite -defH. +by rewrite -subG1 -tiMsMg -defH setIS ?pcore_sub. +Qed. + +(* This is B & G, Theorem 11.5. *) +Theorem exceptional_Sylow_abelian P1 : p.-Sylow(M) P1 -> abelian P1. +Proof. +have nregA Q: gval Q != 1 -> A \subset 'N(Q) -> coprime #|Q| #|A| -> + exists2 X, X \in 'E_p^1(A) & 'C_Q(X) != 1. +- move=> ntQ nQA coQA; apply/exists_inP; apply: contraR ntQ. + rewrite negb_exists_in -subG1; move/forall_inP=> regA. + have ncycA: ~~ cyclic A by rewrite (abelem_cyclic abelA) oA pfactorK. + rewrite -(coprime_abelian_gen_cent1 _ _ nQA) // gen_subG. + apply/bigcupsP=> x /setD1P[ntx Ax]. + apply/negPn; rewrite /= -cent_cycle subG1 regA // p1ElemE // !inE. + by rewrite cycle_subG Ax /= -orderE (abelem_order_p abelA). +suffices{P1} cPP: abelian P. + by move=> sylP1; have [m _ ->] := Sylow_trans sylP sylP1; rewrite abelianJ. +have [g nPg notMg] := subsetPn not_sNP_M. +pose Ms := M`_\sigma; pose q := pdiv #|Ms|; have pP := pHall_pgroup sylP. +have nMsP: P \subset 'N(Ms) by rewrite (subset_trans sPM) ?gFnorm. +have coMsP: coprime #|Ms| #|P|. + exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat pP sM'p). +have [Q1 sylQ1 nQ1P]:= sol_coprime_Sylow_exists q (pgroup_sol pP) nMsP coMsP. +have ntQ1: Q1 :!=: 1. + rewrite -cardG_gt1 (card_Hall sylQ1) p_part_gt1 pi_pdiv cardG_gt1. + by rewrite Msigma_neq1. +have nQ1A: A \subset 'N(Q1) := subset_trans sAP nQ1P. +have coQ1A: coprime #|Q1| #|A|. + by rewrite (coprimeSg (pHall_sub sylQ1)) // (coprimegS sAP). +have [X1 EpX1 nregX11] := nregA _ ntQ1 nQ1A coQ1A. +pose Q2 := Q1 :^ g; have sylQ2: q.-Sylow(Ms :^ g) Q2 by rewrite pHallJ2. +have{ntQ1} ntQ2: Q2 != 1 by rewrite -!cardG_gt1 cardJg in ntQ1 *. +have nQ2A: A \subset 'N(Q2) by rewrite (subset_trans sAP) ?norm_conj_norm. +have{coQ1A} coQ2A: coprime #|Q2| #|A| by rewrite cardJg. +have{nregA ntQ2 coQ2A} [X2 EpX2 nregX22] := nregA _ ntQ2 nQ2A coQ2A. +have [|_ regA]:= exceptional_TIsigmaJ notMg _ sylQ1 nQ1A sylQ2 nQ2A. + by rewrite (subset_trans sAP) // -(normP nPg) conjSg. +have regX21: 'C_Q1(X2) = 1 by case: (regA X2) nregX22 => // ->; rewrite eqxx. +have regX12: 'C_Q2(X1) = 1 by case: (regA X1) nregX11 => // ->; rewrite eqxx. +pose X := 'Ohm_1('Z(P))%G. +have eqCQ12_X: ('C_Q1(X) == 1) = ('C_Q2(X) == 1). + rewrite -(inj_eq (@conjsg_inj _ g)) conjs1g conjIg -/Q2 -centJ (normP _) //. + rewrite (subsetP (char_norm_trans (Ohm_char 1 _) _) g nPg) //. + by rewrite char_norms ?center_char. +have{EpX1} EpX1: X1 \in 'E_p^1(A) :\ X. + rewrite 2!inE EpX1 andbT; apply: contraNneq nregX11 => defX1. + by rewrite defX1 eqCQ12_X -defX1 regX12. +have{EpX2 eqCQ12_X} EpX2: X2 \in 'E_p^1(A) :\ X. + rewrite 2!inE EpX2 andbT; apply: contraNneq nregX22 => defX2. + by rewrite defX2 -eqCQ12_X -defX2 regX21. +apply: contraR nregX11 => not_cPP. +have{not_cPP} transNPA: [transitive 'N_P(A), on 'E_p^1(A) :\ X | 'JG]. + have [|_ _]:= basic_p2maxElem_structure _ pP sAP not_cPP; last by []. + by rewrite inE (subsetP (pnElemS p 2 (subsetT M))). +have [y PnAy ->] := atransP2 transNPA EpX2 EpX1; have [Py _] := setIP PnAy. +by rewrite centJ -(normsP nQ1P y Py) -conjIg regX21 conjs1g. +Qed. + +(* This is B & G, Corollary 11.6. *) +Corollary exceptional_structure (Ms := M`_\sigma) : + [/\ (*a*) A :=: 'Ohm_1(P), + (*b*) 'C_Ms(A) = 1 + & (*c*) exists2 A1, A1 \in 'E_p^1(A) & exists2 A2, A2 \in 'E_p^1(A) & + [/\ A1 :!=: A2, 'C_Ms(A1) = 1 & 'C_Ms(A2) = 1]]. +Proof. +pose iMNA := #|'N(A) : M|. +have defA: A :=: 'Ohm_1(P). + apply/eqP; rewrite eqEcard -{1}(Ohm1_id abelA) OhmS //= oA -rM. + rewrite -(p_rank_Sylow sylP) p_rank_abelian ?exceptional_Sylow_abelian //. + by rewrite -card_pgroup // (pgroupS _ (pHall_pgroup sylP)) ?Ohm_sub. +have iMNAgt1: iMNA > 1. + rewrite indexg_gt1 defA; apply: contra (subset_trans _) not_sNP_M. + by rewrite char_norms ?Ohm_char. +have iMNAgt2: iMNA > 2. + pose q := pdiv iMNA; have q_iMNA: q %| iMNA := pdiv_dvd iMNA. + rewrite (leq_trans _ (dvdn_leq (ltnW _) q_iMNA)) // ltn_neqAle eq_sym. + rewrite (sameP eqP (prime_oddPn _)) ?prime_gt1 ?pdiv_prime //. + by rewrite (dvdn_odd q_iMNA) // (dvdn_odd (dvdn_indexg _ _)) ?mFT_odd. +rewrite [iMNA](cardD1 (gval M)) orbit_refl !ltnS lt0n in iMNAgt1 iMNAgt2. +have{iMNAgt1} [Mg1 /= NM_Mg1] := pred0Pn iMNAgt1. +rewrite (cardD1 Mg1) inE /= NM_Mg1 ltnS lt0n in iMNAgt2. +have{iMNAgt2} [Mg2 /= NM_Mg2] := pred0Pn iMNAgt2. +case/andP: NM_Mg1 => neM_Mg1 /rcosetsP[g1 nAg1 defMg1]. +have{neM_Mg1} notMg1: g1 \notin M. + by apply: contra neM_Mg1 => M_g1; rewrite defMg1 rcoset_id. +case/and3P: NM_Mg2 => neMg12 neM_Mg2 /rcosetsP[g2 nAg2 defMg2]. +have{neM_Mg2} notMg2: g2 \notin M. + by apply: contra neM_Mg2 => M_g2; rewrite defMg2 rcoset_id. +pose A1 := (A0 :^ g1)%G; pose A2 := (A0 :^ g2)%G. +have EpA1: A1 \in 'E_p^1(A) by rewrite -(normP nAg1) pnElemJ. +have EpA2: A2 \in 'E_p^1(A) by rewrite -(normP nAg2) pnElemJ. +have{neMg12} neqA12: A1 :!=: A2. + rewrite -(canF_eq (conjsgKV g2)) -conjsgM (sameP eqP normP). + rewrite (contra (subsetP sNA0_M _)) // -mem_rcoset. + by apply: contra neMg12 => g1Mg2; rewrite defMg1 defMg2 (rcoset_transl g1Mg2). +have{notMg1 nAg1} regA1: 'C_Ms(A1) = 1. + by case/exceptional_TI_MsigmaJ: notMg1; rewrite // -(normP nAg1) conjSg. +have{notMg2 nAg2} regA2: 'C_Ms(A2) = 1. + by case/exceptional_TI_MsigmaJ: notMg2; rewrite // -(normP nAg2) conjSg. +split=> //; last by exists A1 => //; exists A2 => //. +by apply/trivgP; rewrite -regA1 setIS ?centS //; case/pnElemP: EpA1. +Qed. + +(* This is B & G, Theorem 11.7 (the main result on exceptional subgroups). *) +Theorem exceptional_mul_sigma_normal : M`_\sigma <*> A <| M. +Proof. +set Ms := M`_\sigma; have pP := pHall_pgroup sylP; have solM := mmax_sol maxM. +have [E hallE sPE] := Hall_superset solM sPM (pi_pnat pP sM'p). +have sAE := subset_trans sAP sPE; have [sEM s'E _] := and3P hallE. +have [_ _ dimA] := pnElemP Ep2A. +have rE: 'r(E) = 2. + apply/eqP; rewrite eqn_leq -{2}dimA -rank_abelem ?rankS // andbT leqNgt. + have [q q_pr ->]:= rank_witness E; apply/negP=> rqEgt2. + have piEq: q \in \pi(E) by rewrite -p_rank_gt0 -(subnKC rqEgt2). + case/negP: (pnatPpi s'E piEq); rewrite /= alpha_sub_sigma // !inE. + by rewrite (leq_trans rqEgt2) ?p_rankS. +have rFEle2: 'r('F(E)) <= 2 by rewrite -rE rankS ?Fitting_sub. +have solE := solvableS sEM solM; have oddE := mFT_odd E. +pose tau : nat_pred := [pred q | q > p]; pose K := 'O_tau(E). +have hallK: tau.-Hall(E) K by rewrite rank2_ge_pcore_Hall. +pose ptau : nat_pred := [pred q | q >= p]; pose KP := K <*> P. +have nKP: P \subset 'N(K) by rewrite (subset_trans sPE) ?gFnorm. +have coKP: coprime #|K| #|P|. + by rewrite (pnat_coprime (pcore_pgroup _ _)) ?(pi_pnat pP) //= !inE ltnn. +have hallKP: ptau.-Hall(E) KP. + rewrite pHallE join_subG pcore_sub sPE /= norm_joinEr ?coprime_cardMg //. + apply/eqP; rewrite -(partnC tau (part_gt0 _ _)) (card_Hall sylP). + rewrite (card_Hall hallK) partn_part => [|q]; last exact: leqW. + rewrite (card_Hall hallE) -!partnI; congr (_ * _)%N; apply: eq_partn => q. + by rewrite 4!inE andbC /= 8!inE -leqNgt -eqn_leq eq_sym; case: eqP => // <-. +have nsKP_E: KP <| E. + by rewrite [KP](eq_Hall_pcore _ hallKP) ?pcore_normal ?rank2_ge_pcore_Hall. +have [cKA | not_cKA]:= boolP (A \subset 'C(K)). + pose KA := K <*> A; have defKA: K \x A = KA. + by rewrite dprodEY // coprime_TIg // (coprimegS sAP). + have defA: 'Ohm_1(P) = A by case exceptional_structure. + have{defA} defA: 'Ohm_1('O_p(KP)) = A. + apply/eqP; rewrite -defA eqEsubset OhmS /=; last first. + rewrite pcore_sub_Hall ?(pHall_subl _ _ sylP) ?joing_subr //. + exact: subset_trans (pHall_sub hallKP) sEM. + rewrite -Ohm_id defA OhmS // pcore_max // /normal join_subG. + rewrite (subset_trans sAP) ?joing_subr // cents_norm 1?centsC //=. + by rewrite -defA gFnorm. + have nMsE: E \subset 'N(Ms) by rewrite (subset_trans sEM) ?gFnorm. + have tiMsE: Ms :&: E = 1. + by rewrite coprime_TIg ?(pnat_coprime (pcore_pgroup _ _)). + have <-: Ms * E = M. + apply/eqP; rewrite eqEcard mulG_subG pcore_sub sEM /= TI_cardMg //. + by rewrite (card_Hall hallE) (card_Hall (Msigma_Hall maxM)) ?partnC. + rewrite norm_joinEr -?quotientK ?(subset_trans sAE) //= cosetpre_normal. + rewrite quotient_normal // -defA (char_normal_trans (Ohm_char _ _)) //. + by rewrite (char_normal_trans (pcore_char p _)). +pose q := pdiv #|K : 'C_K(A)|. +have q_pr: prime q by rewrite pdiv_prime // indexg_gt1 subsetI subxx centsC. +have [nKA coKA] := (subset_trans sAP nKP, coprimegS sAP coKP). +have [Q sylQ nQA]: exists2 Q : {group gT}, q.-Sylow(K) Q & A \subset 'N(Q). + by apply: sol_coprime_Sylow_exists => //; exact: (pgroup_sol pA). +have [sQK qQ q'iQK] := and3P sylQ; have [sKE tauK _]:= and3P hallK. +have{q'iQK} not_cQA: ~~ (A \subset 'C(Q)). + apply: contraL q'iQK => cQA; rewrite p'natE // negbK. + rewrite -(Lagrange_index (subsetIl K 'C(A))) ?dvdn_mulr ?pdiv_dvd //. + by rewrite subsetI sQK centsC. +have ntQ: Q :!=: 1 by apply: contraNneq not_cQA => ->; exact: cents1. +have q_dv_K: q %| #|K| := dvdn_trans (pdiv_dvd _) (dvdn_indexg _ _). +have sM'q: q \in (\sigma(M))^' := pgroupP (pgroupS sKE s'E) q q_pr q_dv_K. +have{q_dv_K} tau_q: q \in tau := pgroupP tauK q q_pr q_dv_K. +have sylQ_E: q.-Sylow(E) Q := subHall_Sylow hallK tau_q sylQ. +have sylQ_M: q.-Sylow(M) Q := subHall_Sylow hallE sM'q sylQ_E. +have q'p: p != q by rewrite neq_ltn [p < q]tau_q. +have [regQ | nregQ] := eqVneq 'C_Q(A) 1; last first. + have ncycQ: ~~ cyclic Q. + apply: contra not_cQA => cycQ. + rewrite (coprime_odd_faithful_Ohm1 qQ) ?mFT_odd ?(coprimeSg sQK) //. + rewrite centsC; apply: contraR nregQ => not_sQ1_CA. + rewrite setIC TI_Ohm1 // setIC prime_TIg //. + by rewrite (Ohm1_cyclic_pgroup_prime cycQ qQ ntQ). + have {ncycQ} rQ: 'r_q(Q) = 2. + apply/eqP; rewrite eqn_leq ltnNge -odd_pgroup_rank1_cyclic ?mFT_odd //. + by rewrite -rE -rank_pgroup ?rankS // (pHall_sub sylQ_E). + have [B Eq2B]: exists B, B \in 'E_q^2(Q) by apply/p_rank_geP; rewrite rQ. + have maxB: B \in 'E*_q(G). + apply: subsetP (subsetP (pnElemS q 2 (pHall_sub sylQ_M)) B Eq2B). + by rewrite sigma'_rank2_max // -(p_rank_Sylow sylQ_M). + have CAq: q %| #|'C(A)|. + apply: dvdn_trans (cardSg (subsetIr Q _)). + by have [_ ? _] := pgroup_pdiv (pgroupS (subsetIl Q _) qQ) nregQ. + have [Qstar maxQstar sQ_Qstar] := max_normed_exists qQ nQA. + have [|Qm] := max_normed_2Elem_signaliser q'p _ maxQstar CAq. + by rewrite inE (subsetP (pnElemS p 2 (subsetT M))). + case=> _ sAQm [_ _ cQstarQm]; rewrite (centSS sAQm sQ_Qstar) // in not_cQA. + apply: cQstarQm; apply/implyP=> _; apply/set0Pn; exists B. + have{Eq2B} Eq2B := subsetP (pnElemS q 2 sQ_Qstar) B Eq2B. + rewrite inE Eq2B (subsetP (pmaxElemS q (subsetT _))) // inE maxB inE. + by have [? _ _] := pnElemP Eq2B. +pose Q0 := 'Z(Q); have charQ0: Q0 \char Q := center_char Q. +have nQ0A: A \subset 'N(Q0) := char_norm_trans charQ0 nQA. +have defQ0: [~: A, Q0] = Q0. + rewrite -{2}[Q0](coprime_abelian_cent_dprod nQ0A) ?center_abelian //. + by rewrite setIAC regQ (setIidPl (sub1G _)) dprodg1 commGC. + by rewrite (coprimeSg (subset_trans (center_sub Q) sQK)). +have [_ _ [A1 EpA1 [A2 EpA2 [neqA12 regA1 regA2]]]] := exceptional_structure. +have defA: A1 \x A2 = A by apply/(p2Elem_dprodP Ep2A EpA1 EpA2). +have{defQ0} defQ0: [~: A1, Q0] * [~: A2, Q0] = Q0. + have{defA} [[_ defA cA12 _] [sA2A _ _]] := (dprodP defA, pnElemP EpA2). + by rewrite -commMG ?defA // normsR ?(cents_norm cA12) // (subset_trans sA2A). +have nsQ0M: Q0 <| M. + have sQ0M: Q0 \subset M := subset_trans (center_sub Q) (pHall_sub sylQ_M). + have qQ0: q.-group Q0 := pgroupS (center_sub Q) qQ. + have p'Q0: p^'.-group Q0 by apply: (pi_pnat qQ0); rewrite eq_sym in q'p. + have sM'Q0: \sigma(M)^'.-group Q0 := pi_pnat qQ0 sM'q. + have cQ0Q0: abelian Q0 := center_abelian Q. + have sA_NQ0: A \subset 'N_M(Q0) by rewrite subsetI sAM. + have sEpA_EpN := subsetP (pnElemS p 1 sA_NQ0). + have nsRQ0 := commG_sigma'_1Elem_cyclic maxM sQ0M sM'Q0 sM'p (sEpA_EpN _ _). + rewrite -defQ0 -!(commGC Q0). + by apply: normalM; [case/nsRQ0: EpA1 | case/nsRQ0: EpA2]. +case/exists_inP: sM'q; exists Q => //. +rewrite (subset_trans (char_norms charQ0)) ?(mmax_normal maxM nsQ0M) //= -/Q0. +by apply: contraNneq ntQ; move/(trivg_center_pgroup qQ)->. +Qed. + +End Section11. diff --git a/mathcomp/odd_order/BGsection12.v b/mathcomp/odd_order/BGsection12.v new file mode 100644 index 0000000..b831ebc --- /dev/null +++ b/mathcomp/odd_order/BGsection12.v @@ -0,0 +1,2688 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq choice div fintype. +Require Import path bigop finset prime fingroup morphism perm automorphism. +Require Import quotient action gproduct gfunctor pgroup cyclic commutator. +Require Import center gseries nilpotent sylow abelian maximal hall frobenius. +Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +Require Import BGsection7 BGsection9 BGsection10 BGsection11. + +(******************************************************************************) +(* This file covers B & G, section 12; it defines the prime sets for the *) +(* complements of M`_\sigma in a maximal group M: *) +(* \tau1(M) == the set of p not in \pi(M^`(1)) (thus not in \sigma(M)), *) +(* such that M has p-rank 1. *) +(* \tau2(M) == the set of p not in \sigma(M), such that M has p-rank 2. *) +(* \tau3(M) == the set of p not in \sigma(M), but in \pi(M^`(1)), such *) +(* that M has p-rank 1. *) +(* We also define the following helper predicate, which encapsulates the *) +(* notation conventions defined at the beginning of B & G, Section 12: *) +(* sigma_complement M E E1 E2 E3 <=> *) +(* E is a Hall \sigma(M)^'-subgroup of M, the Ei are Hall *) +(* \tau_i(M)-subgroups of E, and E2 * E1 is a group. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. +Section Definitions. + +Variables (gT : finGroupType) (M : {set gT}). +Local Notation sigma' := \sigma(M)^'. + +Definition tau1 := [pred p in sigma' | 'r_p(M) == 1%N & ~~ (p %| #|M^`(1)|)]. +Definition tau2 := [pred p in sigma' | 'r_p(M) == 2]. +Definition tau3 := [pred p in sigma' | 'r_p(M) == 1%N & p %| #|M^`(1)|]. + +Definition sigma_complement E E1 E2 E3 := + [/\ sigma'.-Hall(M) E, tau1.-Hall(E) E1, tau2.-Hall(E) E2, tau3.-Hall(E) E3 + & group_set (E2 * E1)]. + +End Definitions. + +Notation "\tau1 ( M )" := (tau1 M) + (at level 2, format "\tau1 ( M )") : group_scope. +Notation "\tau2 ( M )" := (tau2 M) + (at level 2, format "\tau2 ( M )") : group_scope. +Notation "\tau3 ( M )" := (tau3 M) + (at level 2, format "\tau3 ( M )") : group_scope. + +Section Section12. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types p q r : nat. +Implicit Types A E H K M Mstar N P Q R S T U V W X Y Z : {group gT}. + +Section Introduction. + +Variables M E : {group gT}. +Hypotheses (maxM : M \in 'M) (hallE : \sigma(M)^'.-Hall(M) E). + +Lemma tau1J x : \tau1(M :^ x) =i \tau1(M). +Proof. by move=> p; rewrite 3!inE sigmaJ p_rankJ derg1 -conjsRg cardJg. Qed. + +Lemma tau2J x : \tau2(M :^ x) =i \tau2(M). +Proof. by move=> p; rewrite 3!inE sigmaJ p_rankJ. Qed. + +Lemma tau3J x : \tau3(M :^ x) =i \tau3(M). +Proof. by move=> p; rewrite 3!inE sigmaJ p_rankJ derg1 -conjsRg cardJg. Qed. + +Lemma tau2'1 : {subset \tau1(M) <= \tau2(M)^'}. +Proof. by move=> p; rewrite !inE; case/and3P=> ->; move/eqP->. Qed. + +Lemma tau3'1 : {subset \tau1(M) <= \tau3(M)^'}. +Proof. by move=> p; rewrite !inE; case/and3P=> -> ->. Qed. + +Lemma tau3'2 : {subset \tau2(M) <= \tau3(M)^'}. +Proof. by move=> p; rewrite !inE; case/andP=> ->; move/eqP->. Qed. + +Lemma ex_sigma_compl : exists F : {group gT}, \sigma(M)^'.-Hall(M) F. +Proof. exact: Hall_exists (mmax_sol maxM). Qed. + +Let s'E : \sigma(M)^'.-group E := pHall_pgroup hallE. +Let sEM : E \subset M := pHall_sub hallE. + +(* For added convenience, this lemma does NOT depend on the maxM assumption. *) +Lemma sigma_compl_sol : solvable E. +Proof. +have [-> | [p p_pr pE]] := trivgVpdiv E; first exact: solvable1. +rewrite (solvableS sEM) // mFT_sol // properT. +apply: contraNneq (pgroupP s'E p p_pr pE) => ->. +have [P sylP] := Sylow_exists p [set: gT]. +by apply/exists_inP; exists P; rewrite ?subsetT. +Qed. +Let solE := sigma_compl_sol. + +Let exHallE pi := exists Ei : {group gT}, pi.-Hall(E) Ei. +Lemma ex_tau13_compl : exHallE \tau1(M) /\ exHallE \tau3(M). +Proof. by split; exact: Hall_exists. Qed. + +Lemma ex_tau2_compl E1 E3 : + \tau1(M).-Hall(E) E1 -> \tau3(M).-Hall(E) E3 -> + exists2 E2 : {group gT}, \tau2(M).-Hall(E) E2 & sigma_complement M E E1 E2 E3. +Proof. +move=> hallE1 hallE3; have [sE1E t1E1 _] := and3P hallE1. +pose tau12 := [predU \tau1(M) & \tau2(M)]. +have t12E1: tau12.-group E1 by apply: sub_pgroup t1E1 => p t1p; apply/orP; left. +have [E21 hallE21 sE1E21] := Hall_superset solE sE1E t12E1. +have [sE21E t12E21 _] := and3P hallE21. +have [E2 hallE2] := Hall_exists \tau2(M) (solvableS sE21E solE). +have [sE2E21 t2E2 _] := and3P hallE2. +have hallE2_E: \tau2(M).-Hall(E) E2. + by apply: subHall_Hall hallE21 _ hallE2 => p t2p; apply/orP; right. +exists E2 => //; split=> //. +suffices ->: E2 * E1 = E21 by apply: groupP. +have coE21: coprime #|E2| #|E1| := sub_pnat_coprime tau2'1 t2E2 t1E1. +apply/eqP; rewrite eqEcard mul_subG ?coprime_cardMg //=. +rewrite -(partnC \tau1(M) (cardG_gt0 E21)) (card_Hall hallE2) mulnC. +rewrite (card_Hall (pHall_subl sE1E21 sE21E hallE1)) leq_pmul2r //. +rewrite dvdn_leq // sub_in_partn // => p t12p t1'p. +by apply: contraLR (pnatPpi t12E21 t12p) => t2'p; apply/norP. +Qed. + +Lemma coprime_sigma_compl : coprime #|M`_\sigma| #|E|. +Proof. exact: pnat_coprime (pcore_pgroup _ _) (pHall_pgroup hallE). Qed. +Let coMsE := coprime_sigma_compl. + +Lemma pi_sigma_compl : \pi(E) =i [predD \pi(M) & \sigma(M)]. +Proof. by move=> p; rewrite /= (card_Hall hallE) pi_of_part // !inE andbC. Qed. + +Lemma sdprod_sigma : M`_\sigma ><| E = M. +Proof. +rewrite sdprodE ?coprime_TIg ?(subset_trans sEM) ?gFnorm //. +apply/eqP; rewrite eqEcard mul_subG ?pcore_sub ?coprime_cardMg //=. +by rewrite (card_Hall (Msigma_Hall maxM)) (card_Hall hallE) partnC. +Qed. + +(* The preliminary remarks in the introduction of B & G, section 12. *) + +Remark der1_sigma_compl : M^`(1) :&: E = E^`(1). +Proof. +have [nsMsM _ defM _ _] := sdprod_context sdprod_sigma. +by rewrite setIC (pprod_focal_coprime defM _ (subxx E)) ?(setIidPr _) ?der_sub. +Qed. + +Remark partition_pi_mmax p : + (p \in \pi(M)) = + [|| p \in \tau1(M), p \in \tau2(M), p \in \tau3(M) | p \in \sigma(M)]. +Proof. +symmetry; rewrite 2!orbA -!andb_orr orbAC -andb_orr orNb andbT. +rewrite orb_andl orNb /= -(orb_idl ((alpha_sub_sigma maxM) p)) orbA orbC -orbA. +rewrite !(eq_sym 'r_p(M)) -!leq_eqVlt p_rank_gt0 orb_idl //. +exact: sigma_sub_pi. +Qed. + +Remark partition_pi_sigma_compl p : + (p \in \pi(E)) = [|| p \in \tau1(M), p \in \tau2(M) | p \in \tau3(M)]. +Proof. +rewrite pi_sigma_compl inE /= partition_pi_mmax !andb_orr /=. +by rewrite andNb orbF !(andbb, andbA) -2!andbA. +Qed. + +Remark tau2E p : (p \in \tau2(M)) = (p \in \pi(E)) && ('r_p(E) == 2). +Proof. +have [P sylP] := Sylow_exists p E. +rewrite -(andb_idl (pnatPpi s'E)) -p_rank_gt0 -andbA; apply: andb_id2l => s'p. +have sylP_M := subHall_Sylow hallE s'p sylP. +by rewrite -(p_rank_Sylow sylP_M) (p_rank_Sylow sylP); case: posnP => // ->. +Qed. + +Remark tau3E p : (p \in \tau3(M)) = (p \in \pi(E^`(1))) && ('r_p(E) == 1%N). +Proof. +have [P sylP] := Sylow_exists p E. +have hallE': \sigma(M)^'.-Hall(M^`(1)) E^`(1). + by rewrite -der1_sigma_compl setIC (Hall_setI_normal _ hallE) ?der_normal. +rewrite 4!inE -(andb_idl (pnatPpi (pHall_pgroup hallE'))) -andbA. +apply: andb_id2l => s'p; have sylP_M := subHall_Sylow hallE s'p sylP. +rewrite -(p_rank_Sylow sylP_M) (p_rank_Sylow sylP) andbC; apply: andb_id2r. +rewrite eqn_leq p_rank_gt0 mem_primes => /and3P[_ p_pr _]. +rewrite (card_Hall hallE') pi_of_part 3?inE ?mem_primes ?cardG_gt0 //=. +by rewrite p_pr inE /= s'p andbT. +Qed. + +Remark tau1E p : + (p \in \tau1(M)) = [&& p \in \pi(E), p \notin \pi(E^`(1)) & 'r_p(E) == 1%N]. +Proof. +rewrite partition_pi_sigma_compl; apply/idP/idP=> [t1p|]. + have [s'p rpM _] := and3P t1p; have [P sylP] := Sylow_exists p E. + have:= tau3'1 t1p; rewrite t1p /= inE /= tau3E -(p_rank_Sylow sylP). + by rewrite (p_rank_Sylow (subHall_Sylow hallE s'p sylP)) rpM !andbT. +rewrite orbC andbC -andbA => /and3P[not_piE'p /eqP rpE]. +by rewrite tau3E tau2E rpE (negPf not_piE'p) andbF. +Qed. + +(* Generate a rank 2 elementary abelian tau2 subgroup in a given complement. *) +Lemma ex_tau2Elem p : + p \in \tau2(M) -> exists2 A, A \in 'E_p^2(E) & A \in 'E_p^2(M). +Proof. +move=> t2p; have [A Ep2A] := p_rank_witness p E. +have <-: 'r_p(E) = 2 by apply/eqP; move: t2p; rewrite tau2E; case/andP. +by exists A; rewrite // (subsetP (pnElemS p _ sEM)). +Qed. + +(* A converse to the above Lemma: if E has an elementary abelian subgroup of *) +(* order p^2, then p must be in tau2. *) +Lemma sigma'2Elem_tau2 p A : A \in 'E_p^2(E) -> p \in \tau2(M). +Proof. +move=> Ep2A; have rE: 'r_p(E) > 1 by apply/p_rank_geP; exists A. +have: p \in \pi(E) by rewrite -p_rank_gt0 ltnW. +rewrite partition_pi_sigma_compl orbCA => /orP[] //. +by rewrite -!andb_orr eqn_leq leqNgt (leq_trans rE) ?andbF ?p_rankS. +Qed. + +(* This is B & G, Lemma 12.1(a). *) +Lemma der1_sigma_compl_nil : nilpotent E^`(1). +Proof. +have sE'E := der_sub 1 E. +have nMaE: E \subset 'N(M`_\alpha) by rewrite (subset_trans sEM) ?gFnorm. +have tiMaE': M`_\alpha :&: E^`(1) = 1. + by apply/trivgP; rewrite -(coprime_TIg coMsE) setISS ?Malpha_sub_Msigma. +rewrite (isog_nil (quotient_isog (subset_trans sE'E nMaE) tiMaE')). +by rewrite (nilpotentS (quotientS _ (dergS 1 sEM))) ?Malpha_quo_nil. +Qed. + +(* This is B & G, Lemma 12.1(g). *) +Lemma tau2_not_beta p : + p \in \tau2(M) -> p \notin \beta(G) /\ {subset 'E_p^2(M) <= 'E*_p(G)}. +Proof. +case/andP=> s'p /eqP rpM; split; first exact: sigma'_rank2_beta' rpM. +by apply/subsetP; exact: sigma'_rank2_max. +Qed. + +End Introduction. + +Implicit Arguments tau2'1 [[M] x]. +Implicit Arguments tau3'1 [[M] x]. +Implicit Arguments tau3'2 [[M] x]. + +(* This is the rest of B & G, Lemma 12.1 (parts b, c, d,e, and f). *) +Lemma sigma_compl_context M E E1 E2 E3 : + M \in 'M -> sigma_complement M E E1 E2 E3 -> + [/\ (*b*) E3 \subset E^`(1) /\ E3 <| E, + (*c*) E2 :==: 1 -> E1 :!=: 1, + (*d*) cyclic E1 /\ cyclic E3, + (*e*) E3 ><| (E2 ><| E1) = E /\ E3 ><| E2 ><| E1 = E + & (*f*) 'C_E3(E) = 1]. +Proof. +move=> maxM [hallE hallE1 hallE2 hallE3 groupE21]. +have [sEM solM] := (pHall_sub hallE, mmax_sol maxM). +have [[sE1E t1E1 _] [sE3E t3E3 _]] := (and3P hallE1, and3P hallE3). +have tiE'E1: E^`(1) :&: E1 = 1. + rewrite coprime_TIg // coprime_pi' ?cardG_gt0 //. + by apply: sub_pgroup t1E1 => p; rewrite (tau1E maxM hallE) => /and3P[]. +have cycE1: cyclic E1. + apply: nil_Zgroup_cyclic. + rewrite odd_rank1_Zgroup ?mFT_odd //; apply: wlog_neg; rewrite -ltnNge. + have [p p_pr ->]:= rank_witness E1; move/ltnW; rewrite p_rank_gt0. + move/(pnatPpi t1E1); rewrite (tau1E maxM hallE) => /and3P[_ _ /eqP <-]. + by rewrite p_rankS. + rewrite abelian_nil // /abelian (sameP commG1P trivgP) -tiE'E1. + by rewrite subsetI (der_sub 1) (dergS 1). +have solE: solvable E := solvableS sEM solM. +have nilE': nilpotent E^`(1) := der1_sigma_compl_nil maxM hallE. +have nsE'piE pi: 'O_pi(E^`(1)) <| E. + exact: char_normal_trans (pcore_char _ _) (der_normal _ _). +have SylowE3 P: Sylow E3 P -> [/\ cyclic P, P \subset E^`(1) & 'C_P(E) = 1]. +- case/SylowP=> p p_pr sylP; have [sPE3 pP _] := and3P sylP. + have [-> | ntP] := eqsVneq P 1. + by rewrite cyclic1 sub1G (setIidPl (sub1G _)). + have t3p: p \in \tau3(M). + rewrite (pnatPpi t3E3) // -p_rank_gt0 -(p_rank_Sylow sylP) -rank_pgroup //. + by rewrite rank_gt0. + have sPE: P \subset E := subset_trans sPE3 sE3E. + have cycP: cyclic P. + rewrite (odd_pgroup_rank1_cyclic pP) ?mFT_odd //. + rewrite (tau3E maxM hallE) in t3p. + by case/andP: t3p => _ /eqP <-; rewrite p_rankS. + have nEp'E: E \subset 'N('O_p^'(E)) by exact: gFnorm. + have nEp'P := subset_trans sPE nEp'E. + have sylP_E := subHall_Sylow hallE3 t3p sylP. + have nsEp'P_E: 'O_p^'(E) <*> P <| E. + rewrite sub_der1_normal ?join_subG ?pcore_sub //=. + rewrite norm_joinEr // -quotientSK //=; last first. + by rewrite (subset_trans (der_sub 1 E)). + have [_ /= <- _ _] := dprodP (nilpotent_pcoreC p nilE'). + rewrite -quotientMidr -mulgA (mulSGid (pcore_max _ _)) ?pcore_pgroup //=. + rewrite quotientMidr quotientS //. + apply: subset_trans (pcore_sub_Hall sylP_E). + by rewrite pcore_max ?pcore_pgroup /=. + have nEP_sol: solvable 'N_E(P) by rewrite (solvableS _ solE) ?subsetIl. + have [K hallK] := Hall_exists p^' nEP_sol; have [sKNEP p'K _] := and3P hallK. + have coPK: coprime #|P| #|K| := pnat_coprime pP p'K. + have sP_NEP: P \subset 'N_E(P) by rewrite subsetI sPE normG. + have mulPK: P * K = 'N_E(P). + apply/eqP; rewrite eqEcard mul_subG //= coprime_cardMg // (card_Hall hallK). + by rewrite (card_Hall (pHall_subl sP_NEP (subsetIl E _) sylP_E)) partnC. + rewrite subsetI in sKNEP; case/andP: sKNEP => sKE nPK. + have nEp'K := subset_trans sKE nEp'E. + have defE: 'O_p^'(E) <*> K * P = E. + have sP_Ep'P: P \subset 'O_p^'(E) <*> P := joing_subr _ _. + have sylP_Ep'P := pHall_subl sP_Ep'P (normal_sub nsEp'P_E) sylP_E. + rewrite -{2}(Frattini_arg nsEp'P_E sylP_Ep'P) /= !norm_joinEr //. + by rewrite -mulgA (normC nPK) -mulPK -{1}(mulGid P) !mulgA. + have ntPE': P :&: E^`(1) != 1. + have sylPE' := Hall_setI_normal (der_normal 1 E) sylP_E. + rewrite -rank_gt0 (rank_Sylow sylPE') p_rank_gt0. + by rewrite (tau3E maxM hallE) in t3p; case/andP: t3p. + have defP := coprime_abelian_cent_dprod nPK coPK (cyclic_abelian cycP). + have{defP} [[PK1 _]|[regKP defP]] := cyclic_pgroup_dprod_trivg pP cycP defP. + have coP_Ep'K: coprime #|P| #|'O_p^'(E) <*> K|. + rewrite (pnat_coprime pP) // -pgroupE norm_joinEr //. + by rewrite pgroupM pcore_pgroup. + rewrite -subG1 -(coprime_TIg coP_Ep'K) setIS ?der1_min // in ntPE'. + rewrite -{1}defE mulG_subG normG normsY // cents_norm //. + exact/commG1P. + by rewrite -{2}defE quotientMidl quotient_abelian ?cyclic_abelian. + split=> //; first by rewrite -defP commgSS. + by apply/trivgP; rewrite -regKP setIS ?centS. +have sE3E': E3 \subset E^`(1). + by rewrite -(Sylow_gen E3) gen_subG; apply/bigcupsP=> P; case/SylowE3. +have cycE3: cyclic E3. + rewrite nil_Zgroup_cyclic ?(nilpotentS sE3E') //. + by apply/forall_inP => P; case/SylowE3. +have regEE3: 'C_E3(E) = 1. + have [// | [p p_pr]] := trivgVpdiv 'C_E3(E). + case/Cauchy=> // x /setIP[]; rewrite -!cycle_subG => sXE3 cEX ox. + have pX: p.-elt x by rewrite /p_elt ox pnat_id. + have [P sylP sXP] := Sylow_superset sXE3 pX. + suffices: <[x]> == 1 by case/idPn; rewrite cycle_eq1 -order_gt1 ox prime_gt1. + rewrite -subG1; case/SylowE3: (p_Sylow sylP) => _ _ <-. + by rewrite subsetI sXP. +have nsE3E: E3 <| E. + have hallE3_E' := pHall_subl sE3E' (der_sub 1 E) hallE3. + by rewrite (nilpotent_Hall_pcore nilE' hallE3_E') /=. +have [sE2E t2E2 _] := and3P hallE2; have [_ nE3E] := andP nsE3E. +have coE21: coprime #|E2| #|E1| := sub_pnat_coprime tau2'1 t2E2 t1E1. +have coE31: coprime #|E3| #|E1| := sub_pnat_coprime tau3'1 t3E3 t1E1. +have coE32: coprime #|E3| #|E2| := sub_pnat_coprime tau3'2 t3E3 t2E2. +have{groupE21} defE: E3 ><| (E2 ><| E1) = E. + have defE21: E2 * E1 = E2 <*> E1 by rewrite -genM_join gen_set_id. + have sE21E: E2 <*> E1 \subset E by rewrite join_subG sE2E. + have nE3E21 := subset_trans sE21E nE3E. + have coE312: coprime #|E3| #|E2 <*> E1|. + by rewrite -defE21 coprime_cardMg // coprime_mulr coE32. + have nE21: E1 \subset 'N(E2). + rewrite (subset_trans (joing_subr E2 E1)) ?sub_der1_norm ?joing_subl //. + rewrite /= -{2}(mulg1 E2) -(setIidPr (der_sub 1 _)) /=. + rewrite -(coprime_mulG_setI_norm defE21) ?gFnorm //. + by rewrite mulgSS ?subsetIl // -tiE'E1 setIC setSI ?dergS. + rewrite (sdprodEY nE21) ?sdprodE ?coprime_TIg //=. + apply/eqP; rewrite eqEcard mul_subG // coprime_cardMg //= -defE21. + rewrite -(partnC \tau3(M) (cardG_gt0 E)) (card_Hall hallE3) leq_mul //. + rewrite coprime_cardMg // (card_Hall hallE1) (card_Hall hallE2). + rewrite -[#|E|`__](partnC \tau2(M)) ?leq_mul ?(partn_part _ tau3'2) //. + rewrite -partnI dvdn_leq // sub_in_partn // => p piEp; apply/implyP. + rewrite inE /= -negb_or /= orbC implyNb orbC. + by rewrite -(partition_pi_sigma_compl maxM hallE). +split=> // [/eqP E2_1|]; last split=> //. + apply: contraTneq (sol_der1_proper solM (subxx _) (mmax_neq1 maxM)) => E1_1. + case/sdprodP: (sdprod_sigma maxM hallE) => _ defM _ _. + rewrite properE der_sub /= negbK -{1}defM mulG_subG Msigma_der1 //. + by rewrite -defE E1_1 E2_1 !sdprodg1 (subset_trans sE3E') ?dergS //. +case/sdprodP: defE => [[_ E21 _ defE21]]; rewrite defE21 => defE nE321 tiE321. +have{defE21} [_ defE21 nE21 tiE21] := sdprodP defE21. +have [nE32 nE31] := (subset_trans sE2E nE3E, subset_trans sE1E nE3E). +rewrite [E3 ><| _]sdprodEY ? sdprodE ?coprime_TIg ?normsY //=. + by rewrite norm_joinEr // -mulgA defE21. +by rewrite norm_joinEr // coprime_cardMg // coprime_mull coE31. +Qed. + +(* This is B & G, Lemma 12.2(a). *) +Lemma prime_class_mmax_norm M p X : + M \in 'M -> p.-group X -> 'N(X) \subset M -> + (p \in \sigma(M)) || (p \in \tau2(M)). +Proof. +move=> maxM pX sNM; rewrite -implyNb; apply/implyP=> sM'p. +by rewrite 3!inE /= sM'p (sigma'_norm_mmax_rank2 _ _ pX). +Qed. + +(* This is B & G, Lemma 12.2(b). *) +Lemma mmax_norm_notJ M Mstar p X : + M \in 'M -> Mstar \in 'M -> + p.-group X -> X \subset M -> 'N(X) \subset Mstar -> + [|| [&& p \in \sigma(M) & M :!=: Mstar], p \in \tau1(M) | p \in \tau3(M)] -> + gval Mstar \notin M :^: G. +Proof. +move: Mstar => H maxM maxH pX sXM sNH; apply: contraL => MG_H. +have [x Gx defH] := imsetP MG_H. +have [sMp | sM'p] := boolP (p \in \sigma(M)); last first. + have:= prime_class_mmax_norm maxH pX sNH. + rewrite defH /= sigmaJ tau2J !negb_or (negPf sM'p) /= => t2Mp. + by rewrite (contraL (@tau2'1 _ p)) // [~~ _]tau3'2. +rewrite 3!inE sMp 3!inE sMp orbF negbK. +have [_ transCX _] := sigma_group_trans maxM sMp pX. +set maxMX := finset _ in transCX. +have maxMX_H: gval H \in maxMX by rewrite inE MG_H (subset_trans (normG X)). +have maxMX_M: gval M \in maxMX by rewrite inE orbit_refl. +have [y cXy ->] := atransP2 transCX maxMX_H maxMX_M. +by rewrite /= conjGid // (subsetP sNH) // (subsetP (cent_sub X)). +Qed. + +(* This is B & G, Lemma 12.3. *) +Lemma nonuniq_p2Elem_cent_sigma M Mstar p A A0 : + M \in 'M -> Mstar \in 'M -> Mstar :!=: M -> A \in 'E_p^2(M) -> + A0 \in 'E_p^1(A) -> 'N(A0) \subset Mstar -> + [/\ (*a*) p \notin \sigma(M) -> A \subset 'C(M`_\sigma :&: Mstar) + & (*b*) p \notin \alpha(M) -> A \subset 'C(M`_\alpha :&: Mstar)]. +Proof. +move: Mstar => H maxM maxH neqMH Ep2A EpA0 sNH. +have p_pr := pnElem_prime Ep2A. +have [sAM abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have [sA0A _ _] := pnElemP EpA0; have pA0 := pgroupS sA0A pA. +have sAH: A \subset H. + by apply: subset_trans (cents_norm _) sNH; exact: subset_trans (centS sA0A). +have nsHsH: H`_\sigma <| H by exact: pcore_normal. +have [sHsH nHsH] := andP nsHsH; have nHsA := subset_trans sAH nHsH. +have nsHsA_H: H`_\sigma <*> A <| H. + have [sHp | sH'p] := boolP (p \in \sigma(H)). + rewrite (joing_idPl _) ?pcore_normal //. + by rewrite (sub_Hall_pcore (Msigma_Hall _)) // (pi_pgroup pA). + have [P sylP sAP] := Sylow_superset sAH pA. + have excH: exceptional_FTmaximal p H A0 A by split=> //; apply/pnElemP. + exact: exceptional_mul_sigma_normal excH sylP sAP. +have cAp' K: + p^'.-group K -> A \subset 'N(K) -> K \subset H -> + [~: K, A] \subset K :&: H`_\sigma. +- move=> p'K nKA sKH; have nHsK := subset_trans sKH nHsH. + rewrite subsetI commg_subl nKA /= -quotient_sub1 ?comm_subG // quotientR //=. + have <-: K / H`_\sigma :&: A / H`_\sigma = 1. + by rewrite setIC coprime_TIg ?coprime_morph ?(pnat_coprime pA p'K). + rewrite subsetI commg_subl commg_subr /= -{2}(quotientYidr nHsA). + by rewrite !quotient_norms //= joingC (subset_trans sKH) ?normal_norm. +have [sMp | sM'p] := boolP (p \in \sigma(M)). + split=> // aM'p; have notMGH: gval H \notin M :^: G. + apply: mmax_norm_notJ maxM maxH pA0 (subset_trans sA0A sAM) sNH _. + by rewrite sMp eq_sym neqMH. + rewrite centsC (sameP commG1P trivgP). + apply: subset_trans (cAp' _ _ _ (subsetIr _ _)) _. + - exact: pi_p'group (pgroupS (subsetIl _ _) (pcore_pgroup _ _)) aM'p. + - by rewrite (normsI _ (normsG sAH)) // (subset_trans sAM) ?gFnorm. + by rewrite setIAC; case/sigma_disjoint: notMGH => // -> _ _; exact: subsetIl. +suffices cMaA: A \subset 'C(M`_\sigma :&: H). + by rewrite !{1}(subset_trans cMaA) ?centS ?setSI // Malpha_sub_Msigma. +have [sHp | sH'p] := boolP (p \in \sigma(H)); last first. + apply/commG1P; apply: contraNeq neqMH => ntA_MsH. + have [P sylP sAP] := Sylow_superset sAH pA. + have excH: exceptional_FTmaximal p H A0 A by split=> //; exact/pnElemP. + have maxAM: M \in 'M(A) by exact/setIdP. + rewrite (exceptional_sigma_uniq maxH excH sylP sAP maxAM) //. + apply: contraNneq ntA_MsH => tiMsHs; rewrite -subG1. + have [sHsA_H nHsA_H] := andP nsHsA_H. + have <-: H`_\sigma <*> A :&: M`_\sigma = 1. + apply/trivgP; rewrite -tiMsHs subsetI subsetIr /=. + rewrite -quotient_sub1 ?subIset ?(subset_trans sHsA_H) //. + rewrite quotientGI ?joing_subl //= joingC quotientYidr //. + rewrite setIC coprime_TIg ?coprime_morph //. + rewrite (pnat_coprime (pcore_pgroup _ _)) // (card_pnElem Ep2A). + by rewrite pnat_exp ?orbF ?pnatE. + rewrite commg_subI // subsetI ?joing_subr ?subsetIl. + by rewrite (subset_trans sAM) ?gFnorm. + by rewrite setIC subIset ?nHsA_H. +have sAHs: A \subset H`_\sigma. + by rewrite (sub_Hall_pcore (Msigma_Hall maxH)) // (pi_pgroup pA). +have [S sylS sAS] := Sylow_superset sAHs pA; have [sSHs pS _] := and3P sylS. +have nsHaH: H`_\alpha <| H := pcore_normal _ _; have [_ nHaH] := andP nsHaH. +have nHaS := subset_trans (subset_trans sSHs sHsH) nHaH. +have nsHaS_H: H`_\alpha <*> S <| H. + rewrite -{2}(quotientGK nsHaH) (norm_joinEr nHaS) -quotientK //. + rewrite cosetpre_normal; apply: char_normal_trans (quotient_normal _ nsHsH). + rewrite /= (nilpotent_Hall_pcore _ (quotient_pHall _ sylS)) ?pcore_char //. + exact: nilpotentS (quotientS _ (Msigma_der1 maxH)) (Malpha_quo_nil maxH). +rewrite (sameP commG1P trivgP). +have <-: H`_\alpha <*> S :&: M`_\sigma = 1. + have: gval M \notin H :^: G. + by apply: contra sM'p; case/imsetP=> x _ ->; rewrite sigmaJ. + case/sigma_disjoint=> // _ ti_aHsM _. + rewrite setIC coprime_TIg ?(pnat_coprime (pcore_pgroup _ _)) //=. + rewrite norm_joinEr // [pnat _ _]pgroupM (pi_pgroup pS) // andbT. + apply: sub_pgroup (pcore_pgroup _ _) => q aHq. + by apply: contraFN (ti_aHsM q) => sMq; rewrite inE /= aHq. +rewrite commg_subI // subsetI ?subsetIl. + by rewrite (subset_trans sAS) ?joing_subr ?(subset_trans sAM) ?gFnorm. +by rewrite setIC subIset 1?normal_norm. +Qed. + +(* This is B & G, Proposition 12.4. *) +Proposition p2Elem_mmax M p A : + M \in 'M -> A \in 'E_p^2(M) -> + (*a*) 'C(A) \subset M + /\ (*b*) ([forall A0 in 'E_p^1(A), 'M('N(A0)) != [set M]] -> + [/\ p \in \sigma(M), M`_\alpha = 1 & nilpotent M`_\sigma]). +Proof. +move=> maxM Ep2A; have p_pr := pnElem_prime Ep2A. +have [sAM abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have [EpAnonuniq |] := altP forall_inP; last first. + rewrite negb_forall_in; case/exists_inP=> A0 EpA0; rewrite negbK. + case/eqP/mem_uniq_mmax=> _ sNA0_M; rewrite (subset_trans _ sNA0_M) //. + by have [sA0A _ _] := pnElemP EpA0; rewrite cents_norm // centS. +have{EpAnonuniq} sCMkApCA y: y \in A^# -> + [/\ 'r('C_M(<[y]>)) <= 2, + p \in \sigma(M)^' -> 'C_(M`_\sigma)[y] \subset 'C_M(A) + & p \in \alpha(M)^' -> 'C_(M`_\alpha)[y] \subset 'C_M(A)]. +- case/setD1P=> nty Ay; pose Y := <[y]>%G. + rewrite -cent_cycle -[<[y]>]/(gval Y). + have EpY: Y \in 'E_p^1(A). + by rewrite p1ElemE // 2!inE cycle_subG Ay -orderE (abelem_order_p abelA) /=. + have [sYA abelY dimY] := pnElemP EpY; have [pY _] := andP abelY. + have [H maxNYH neqHM]: exists2 H, H \in 'M('N(Y)) & H \notin [set M]. + apply/subsetPn; rewrite subset1 negb_or EpAnonuniq //=. + apply/set0Pn; have [|H] := (@mmax_exists _ 'N(Y)); last by exists H. + rewrite mFT_norm_proper ?(mFT_pgroup_proper pY) //. + by rewrite -rank_gt0 (rank_abelem abelY) dimY. + have{maxNYH} [maxH sNYH] := setIdP maxNYH; rewrite inE -val_eqE /= in neqHM. + have ->: 'r('C_M(Y)) <= 2. + apply: contraR neqHM; rewrite -ltnNge => rCMYgt2. + have uniqCMY: 'C_M(Y)%G \in 'U. + by rewrite rank3_Uniqueness ?(sub_mmax_proper maxM) ?subsetIl. + have defU: 'M('C_M(Y)) = [set M] by apply: def_uniq_mmax; rewrite ?subsetIl. + rewrite (eq_uniq_mmax defU maxH) ?subIset //. + by rewrite orbC (subset_trans (cent_sub Y)). + have [cAMs cAMa] := nonuniq_p2Elem_cent_sigma maxM maxH neqHM Ep2A EpY sNYH. + rewrite !{1}subsetI !{1}(subset_trans (subsetIl _ _) (pcore_sub _ _)). + by split=> // [/cAMs | /cAMa]; rewrite centsC; apply: subset_trans; + rewrite setIS ?(subset_trans (cent_sub Y)). +have ntA: A :!=: 1 by rewrite -rank_gt0 (rank_abelem abelA) dimA. +have ncycA: ~~ cyclic A by rewrite (abelem_cyclic abelA) dimA. +have rCMAle2: 'r('C_M(A)) <= 2. + have [y Ay]: exists y, y \in A^# by apply/set0Pn; rewrite setD_eq0 subG1. + have [rCMy _ _] := sCMkApCA y Ay; apply: leq_trans rCMy. + by rewrite rankS // setIS // centS // cycle_subG; case/setIdP: Ay. +have sMp: p \in \sigma(M). + apply: contraFT (ltnn 1) => sM'p; rewrite -dimA -(rank_abelem abelA). + suffices cMsA: A \subset 'C(M`_\sigma). + by rewrite -(setIidPl cMsA) sub'cent_sigma_rank1 // (pi_pgroup pA). + have nMsA: A \subset 'N(M`_\sigma) by rewrite (subset_trans sAM) ?gFnorm. + rewrite centsC /= -(coprime_abelian_gen_cent1 _ _ nMsA) //; last first. + exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat pA _). + rewrite gen_subG; apply/bigcupsP=> y; case/sCMkApCA=> _ sCMsyCA _. + by rewrite (subset_trans (sCMsyCA sM'p)) ?subsetIr. +have [P sylP sAP] := Sylow_superset sAM pA; have [sPM pP _] := and3P sylP. +pose Z := 'Ohm_1('Z(P)). +have sZA: Z \subset A. + have maxA: A \in 'E*_p('C_M(A)). + have sACMA: A \subset 'C_M(A) by rewrite subsetI sAM. + rewrite (subsetP (p_rankElem_max _ _)) // !inE abelA sACMA. + rewrite eqn_leq logn_le_p_rank /=; last by rewrite !inE sACMA abelA. + by rewrite dimA (leq_trans (p_rank_le_rank _ _)). + rewrite [Z](OhmE 1 (pgroupS (center_sub P) pP)) gen_subG. + rewrite -(pmaxElem_LdivP p_pr maxA) -(setIA M) setIid setSI //=. + by rewrite setISS // centS. +have{ntA} ntZ: Z != 1. + by rewrite Ohm1_eq1 (center_nil_eq1 (pgroup_nil pP)) (subG1_contra sAP). +have rPle2: 'r(P) <= 2. + have [z Zz ntz]: exists2 z, z \in Z & z \notin [1]. + by apply/subsetPn; rewrite subG1. + have [|rCMz _ _] := sCMkApCA z; first by rewrite inE ntz (subsetP sZA). + rewrite (leq_trans _ rCMz) ?rankS // subsetI sPM centsC cycle_subG. + by rewrite (subsetP _ z Zz) // (subset_trans (Ohm_sub 1 _)) ?subsetIr. +have aM'p: p \in \alpha(M)^'. + by rewrite !inE -leqNgt -(p_rank_Sylow sylP) -rank_pgroup. +have sMaCMA: M`_\alpha \subset 'C_M(A). +have nMaA: A \subset 'N(M`_\alpha) by rewrite (subset_trans sAM) ?gFnorm. + rewrite -(coprime_abelian_gen_cent1 _ _ nMaA) //; last first. + exact: (pnat_coprime (pcore_pgroup _ _) (pi_pnat pA _)). + rewrite gen_subG; apply/bigcupsP=> y; case/sCMkApCA=> _ _ sCMayCA. + by rewrite (subset_trans (sCMayCA aM'p)) ?subsetIr. +have Ma1: M`_\alpha = 1. + have [q q_pr rMa]:= rank_witness M`_\alpha. + apply: contraTeq rCMAle2; rewrite -ltnNge -rank_gt0 rMa p_rank_gt0 => piMa_q. + have aMq: q \in \alpha(M) := pnatPpi (pcore_pgroup _ _) piMa_q. + apply: leq_trans (rankS sMaCMA); rewrite rMa. + have [Q sylQ] := Sylow_exists q M`_\alpha; rewrite -(p_rank_Sylow sylQ). + by rewrite (p_rank_Sylow (subHall_Sylow (Malpha_Hall maxM) aMq sylQ)). +have nilMs: nilpotent M`_\sigma. + rewrite (nilpotentS (Msigma_der1 maxM)) // (isog_nil (quotient1_isog _)). + by rewrite -Ma1 Malpha_quo_nil. +rewrite (subset_trans (cents_norm (centS sZA))) ?(mmax_normal maxM) //. +apply: char_normal_trans (char_trans (Ohm_char 1 _) (center_char P)) _. +have{sylP} sylP: p.-Sylow(M`_\sigma) P. + apply: pHall_subl _ (pcore_sub _ _) sylP. + by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) // (pi_pgroup pP). +by rewrite (nilpotent_Hall_pcore _ sylP) ?(char_normal_trans (pcore_char _ _)). +Qed. + +(* This is B & G, Theorem 12.5(a) -- this part does not mention a specific *) +(* rank 2 elementary abelian \tau_2(M) subgroup of M. *) + +Theorem tau2_Msigma_nil M p : M \in 'M -> p \in \tau2(M) -> nilpotent M`_\sigma. +Proof. +move=> maxM t2Mp; have [sM'p /eqP rpM] := andP t2Mp. +have [A Ep2A] := p_rank_witness p M; rewrite rpM in Ep2A. +have [_]:= p2Elem_mmax maxM Ep2A; rewrite -negb_exists_in [p \in _](negPf sM'p). +have [[A0 EpA0 /eqP/mem_uniq_mmax[_ sNA0M _]] | _ [] //] := exists_inP. +have{EpA0 sNA0M} excM: exceptional_FTmaximal p M A0 A by []. +have [sAM abelA _] := pnElemP Ep2A; have [pA _] := andP abelA. +have [P sylP sAP] := Sylow_superset sAM pA. +exact: exceptional_sigma_nil maxM excM sylP sAP. +Qed. + +(* This is B & G, Theorem 12.5 (b-f) -- the bulk of the Theorem. *) +Theorem tau2_context M p A (Ms := M`_\sigma) : + M \in 'M -> p \in \tau2(M) -> A \in 'E_p^2(M) -> + [/\ (*b*) forall P, p.-Sylow(M) P -> + abelian P + /\ (A \subset P -> 'Ohm_1(P) = A /\ ~~ ('N(P) \subset M)), + (*c*) Ms <*> A <| M, + (*d*) 'C_Ms(A) = 1, + (*e*) forall Mstar, Mstar \in 'M(A) :\ M -> Ms :&: Mstar = 1 + & (*f*) exists2 A1, A1 \in 'E_p^1(A) & 'C_Ms(A1) = 1]. +Proof. +move=> maxM t2Mp Ep2A; have [sM'p _] := andP t2Mp. +have [_]:= p2Elem_mmax maxM Ep2A; rewrite -negb_exists_in [p \in _](negPf sM'p). +have [[A0 EpA0 /eqP/mem_uniq_mmax[_ sNA0M _]] | _ [] //] := exists_inP. +have{EpA0 sNA0M} excM: exceptional_FTmaximal p M A0 A by []. +have strM := exceptional_structure maxM excM. +have [sAM abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have [P sylP sAP] := Sylow_superset sAM pA. +have nsMsA_M : Ms <*> A <| M := exceptional_mul_sigma_normal maxM excM sylP sAP. +have [_ regA [A1 EpA1 [_ _ [_ regA1 _]]]] := strM P sylP sAP. +split=> // [P1 sylP1 | {P sylP sAP A0 excM}H| ]; last by exists A1. + split=> [|sAP1]; first exact: (exceptional_Sylow_abelian _ excM sylP). + split; first by case/strM: sylP1. + by apply: contra sM'p => sNP1M; apply/exists_inP; exists P1; rewrite // ?inE. +case/setD1P; rewrite -val_eqE /= => neqHM /setIdP[maxH sAH]. +apply/trivgP; rewrite -regA subsetI subsetIl /=. +have Ep2A_H: A \in 'E_p^2(H) by apply/pnElemP. +have [_]:= p2Elem_mmax maxH Ep2A_H; rewrite -negb_exists_in. +have [[A0 EpA0 /eqP/mem_uniq_mmax[_ sNA0H _]]|_ [//|sHp _ nilHs]] := exists_inP. + have [cMSH_A _]:= nonuniq_p2Elem_cent_sigma maxM maxH neqHM Ep2A EpA0 sNA0H. + by rewrite centsC cMSH_A. +have [P sylP sAP] := Sylow_superset sAH pA; have [sPH pP _] := and3P sylP. +have sylP_Hs: p.-Sylow(H`_\sigma) P. + rewrite (pHall_subl _ (pcore_sub _ _) sylP) //. + by rewrite (sub_Hall_pcore (Msigma_Hall maxH)) // (pi_pgroup pP). +have nPH: H \subset 'N(P). + rewrite (nilpotent_Hall_pcore nilHs sylP_Hs). + by rewrite !(char_norm_trans (pcore_char _ _)) ?normG. +have coMsP: coprime #|M`_\sigma| #|P|. + exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat pP _). +rewrite (sameP commG1P trivgP) -(coprime_TIg coMsP) commg_subI ?setIS //. +by rewrite subsetI sAP (subset_trans sAM) ?gFnorm. +Qed. + +(* This is B & G, Corollary 12.6 (a, b, c & f) -- i.e., the assertions that *) +(* do not depend on the decomposition of the complement. *) +Corollary tau2_compl_context M E p A (Ms := M`_\sigma) : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> p \in \tau2(M) -> A \in 'E_p^2(E) -> + [/\ (*a*) A <| E /\ 'E_p^1(E) = 'E_p^1(A), + (*b*) [/\ 'C(A) \subset E, 'N_M(A) = E & ~~ ('N(A) \subset M)], + (*c*) forall X, X \in 'E_p^1(E) -> 'C_Ms(X) != 1 -> 'M('C(X)) = [set M] + & (*f*) forall Mstar, + Mstar \in 'M -> gval Mstar \notin M :^: G -> + Ms :&: Mstar`_\sigma = 1 + /\ [predI \sigma(M) & \sigma(Mstar)] =i pred0]. +Proof. +move=> maxM hallE t2Mp Ep2A; have [sEM sM'E _] := and3P hallE. +have [p_pr [sM'p _]] := (pnElem_prime Ep2A, andP t2Mp). +have [sAE abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have [_ mulMsE nMsE tiMsE] := sdprodP (sdprod_sigma maxM hallE). +have Ep2A_M: A \in 'E_p^2(M) by rewrite (subsetP (pnElemS _ _ sEM)). +have [syl_p_M nsMsAM regA tiMsMA _] := tau2_context maxM t2Mp Ep2A_M. +have nMsA: A \subset 'N(Ms) := subset_trans sAE nMsE. +have nsAE: A <| E. + rewrite /normal sAE -(mul1g A) -tiMsE setIC group_modr // normsI ?normG //. + by rewrite (subset_trans sEM) // -(norm_joinEr nMsA) normal_norm. +have sAsylE P: p.-Sylow(E) P -> 'Ohm_1(P) = A /\ ~~ ('N(P) \subset M). + move=> sylP; have sylP_M: p.-Sylow(M) P by apply: (subHall_Sylow hallE). + have [_] := syl_p_M P sylP_M; apply. + exact: subset_trans (pcore_max pA nsAE) (pcore_sub_Hall sylP). +have not_sNA_M: ~~ ('N(A) \subset M). + have [P sylP] := Sylow_exists p E; have [<-]:= sAsylE P sylP. + exact: contra (subset_trans (char_norms (Ohm_char 1 P))). +have{sAsylE syl_p_M} defEpE: 'E_p^1(E) = 'E_p^1(A). + apply/eqP; rewrite eqEsubset andbC pnElemS //. + apply/subsetP=> X /pnElemP[sXE abelX dimX]; apply/pnElemP; split=> //. + have [pX _ eX] := and3P abelX; have [P sylP sXP] := Sylow_superset sXE pX. + have [<- _]:= sAsylE P sylP; have [_ pP _] := and3P sylP. + by rewrite (OhmE 1 pP) sub_gen // subsetI sXP sub_LdivT. +have defNMA: 'N_M(A) = E. + rewrite -mulMsE setIC -group_modr ?normal_norm //= setIC. + rewrite coprime_norm_cent ?regA ?mul1g //. + exact: (pnat_coprime (pcore_pgroup _ _) (pi_pnat pA _)). +have [sCAM _]: 'C(A) \subset M /\ _ := p2Elem_mmax maxM Ep2A_M. +have sCAE: 'C(A) \subset E by rewrite -defNMA subsetI sCAM cent_sub. +split=> // [X EpX | H maxH not_MGH]; last first. + by case/sigma_disjoint: not_MGH => // _ _; apply; apply: tau2_Msigma_nil t2Mp. +rewrite defEpE in EpX; have [sXA abelX dimX] := pnElemP EpX. +have ntX: X :!=: 1 by rewrite -rank_gt0 (rank_abelem abelX) dimX. +apply: contraNeq => neq_maxCX_M. +have{neq_maxCX_M} [H]: exists2 H, H \in 'M('C(X)) & H \notin [set M]. + apply/subsetPn; rewrite subset1 negb_or neq_maxCX_M. + by have [H maxH]:= mmax_exists (mFT_cent_proper ntX); apply/set0Pn; exists H. +case/setIdP=> maxH sCXH neqHM. +rewrite -subG1 -(tiMsMA H) ?setIS // inE neqHM inE maxH. +exact: subset_trans (sub_abelian_cent cAA sXA) sCXH. +Qed. + +(* This is B & G, Corollary 12.6 (d, e) -- the parts that apply to a *) +(* particular decomposition of the complement. We included an easy consequece *) +(* of part (a), that A is a subgroup of E2, as this is used implicitly later *) +(* in sections 12 and 13. *) +Corollary tau2_regular M E E1 E2 E3 p A (Ms := M`_\sigma) : + M \in 'M -> sigma_complement M E E1 E2 E3 -> + p \in \tau2(M) -> A \in 'E_p^2(E) -> + [/\ (*d*) semiregular Ms E3, + (*e*) semiregular Ms 'C_E1(A) + & A \subset E2]. +Proof. +move=> maxM complEi t2Mp Ep2A; have p_pr := pnElem_prime Ep2A. +have [hallE hallE1 hallE2 hallE3 _] := complEi. +have [sEM sM'E _] := and3P hallE; have [sM'p _] := andP t2Mp. +have [sAE abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have Ep2A_M: A \in 'E_p^2(M) by rewrite (subsetP (pnElemS _ _ sEM)). +have [_ _ _ tiMsMA _] := tau2_context maxM t2Mp Ep2A_M. +have [[nsAE _] _ _ _] := tau2_compl_context maxM hallE t2Mp Ep2A. +have [sCAM _]: 'C(A) \subset M /\ _ := p2Elem_mmax maxM Ep2A_M. +have sAE2: A \subset E2. + exact: normal_sub_max_pgroup (Hall_max hallE2) (pi_pnat pA _) nsAE. +split=> // x /setD1P[ntx]; last first. + case/setIP; rewrite -cent_cycle -!cycle_subG => sXE1 cAX. + pose q := pdiv #[x]; have piXq: q \in \pi(#[x]) by rewrite pi_pdiv order_gt1. + have [Q sylQ] := Sylow_exists q <[x]>; have [sQX qQ _] := and3P sylQ. + have [sE1E t1E1 _] := and3P hallE1; have sQE1 := subset_trans sQX sXE1. + have sQM := subset_trans sQE1 (subset_trans sE1E sEM). + have [H /setIdP[maxH sNQH]]: {H | H \in 'M('N(Q))}. + apply: mmax_exists; rewrite mFT_norm_proper ?(mFT_pgroup_proper qQ) //. + by rewrite -rank_gt0 (rank_pgroup qQ) (p_rank_Sylow sylQ) p_rank_gt0. + apply/trivgP; rewrite -(tiMsMA H) ?setIS //. + by rewrite (subset_trans _ sNQH) ?cents_norm ?centS. + rewrite 3!inE maxH /=; apply/andP; split. + apply: contra_orbit (mmax_norm_notJ maxM maxH qQ sQM sNQH _). + by rewrite (pnatPpi (pgroupS sXE1 t1E1)) ?orbT. + by rewrite (subset_trans _ sNQH) ?cents_norm // centsC (subset_trans sQX). +rewrite -cent_cycle -cycle_subG => sXE3. +pose q := pdiv #[x]; have piXq: q \in \pi(#[x]) by rewrite pi_pdiv order_gt1. +have [Q sylQ] := Sylow_exists q <[x]>; have [sQX qQ _] := and3P sylQ. +have [sE3E t3E3 _] := and3P hallE3; have sQE3 := subset_trans sQX sXE3. +have sQM := subset_trans sQE3 (subset_trans sE3E sEM). +have [H /setIdP[maxH sNQH]]: {H | H \in 'M('N(Q))}. + apply: mmax_exists; rewrite mFT_norm_proper ?(mFT_pgroup_proper qQ) //. + by rewrite -rank_gt0 (rank_pgroup qQ) (p_rank_Sylow sylQ) p_rank_gt0. +apply/trivgP; rewrite -(tiMsMA H) ?setIS //. + by rewrite (subset_trans _ sNQH) ?cents_norm ?centS. +rewrite 3!inE maxH /=; apply/andP; split. + apply: contra_orbit (mmax_norm_notJ maxM maxH qQ sQM sNQH _). + by rewrite (pnatPpi (pgroupS sXE3 t3E3)) ?orbT. +rewrite (subset_trans _ sNQH) ?cents_norm // (subset_trans _ (centS sQE3)) //. +have coE3A: coprime #|E3| #|A|. + by rewrite (pnat_coprime t3E3 (pi_pnat pA _)) ?tau3'2. +rewrite (sameP commG1P trivgP) -(coprime_TIg coE3A) subsetI commg_subl. +have [[_ nsE3E] _ _ _ _] := sigma_compl_context maxM complEi. +by rewrite commg_subr (subset_trans sE3E) ?(subset_trans sAE) ?normal_norm. +Qed. + +(* This is B & G, Theorem 12.7. *) +Theorem nonabelian_tau2 M E p A P0 (Ms := M`_\sigma) (A0 := 'C_A(Ms)%G) : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> p \in \tau2(M) -> A \in 'E_p^2(E) -> + p.-group P0 -> ~~ abelian P0 -> + [/\ (*a*) \tau2(M) =i (p : nat_pred), + (*b*) #|A0| = p /\ Ms \x A0 = 'F(M), + (*c*) forall X, + X \in 'E_p^1(E) :\ A0 -> 'C_Ms(X) = 1 /\ ~~ ('C(X) \subset M) + & (*d*) exists2 E0 : {group gT}, A0 ><| E0 = E + & (*e*) forall x, x \in Ms^# -> {subset \pi('C_E0[x]) <= \tau1(M)}]. +Proof. +rewrite {}/A0 => maxM hallE t2Mp Ep2A pP0 not_cP0P0 /=. +have p_pr := pnElem_prime Ep2A. +have [sEM sM'E _] := and3P hallE; have [sM'p _] := andP t2Mp. +have [sAE abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have Ep2A_M: A \in 'E_p^2(M) by rewrite (subsetP (pnElemS _ _ sEM)). +have [[E1 hallE1] [E3 hallE3]] := ex_tau13_compl hallE. +have [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. +have [regE3 _ sAE2] := tau2_regular maxM complEi t2Mp Ep2A. +have [P sylP sAP] := Sylow_superset sAE2 pA; have [sPE2 pP _] := and3P sylP. +have [S /= sylS sPS] := Sylow_superset (subsetT P) pP. +have pS := pHall_pgroup sylS; have sAS := subset_trans sAP sPS. +have sylP_E: p.-Sylow(E) P := subHall_Sylow hallE2 t2Mp sylP. +have sylP_M: p.-Sylow(M) P := subHall_Sylow hallE sM'p sylP_E. +have [syl_p_M _ regA _ _] := tau2_context maxM t2Mp Ep2A_M. +have{syl_p_M} cPP: abelian P by case: (syl_p_M P). +have{P0 pP0 not_cP0P0} not_cSS: ~~ abelian S. + have [x _ sP0Sx] := Sylow_subJ sylS (subsetT P0) pP0. + by apply: contra not_cP0P0 => cSS; rewrite (abelianS sP0Sx) ?abelianJ. +have [defP | ltPS] := eqVproper sPS; first by rewrite -defP cPP in not_cSS. +have [[nsAE defEp] [sCAE _ _] nregA _] := + tau2_compl_context maxM hallE t2Mp Ep2A. +have defCSA: 'C_S(A) = P. + apply: (sub_pHall sylP_E (pgroupS (subsetIl _ _) pS)). + by rewrite subsetI sPS (sub_abelian_cent2 cPP). + by rewrite subIset // sCAE orbT. +have max2A: A \in 'E_p^2(G) :&: 'E*_p(G). + by rewrite 3!inE subsetT abelA dimA; case: (tau2_not_beta maxM t2Mp) => _ ->. +have def_t2: \tau2(M) =i (p : nat_pred). + move=> q; apply/idP/idP=> [t2Mq |]; last by move/eqnP->. + apply: contraLR (proper_card ltPS); rewrite !inE /= eq_sym -leqNgt => q'p. + apply: wlog_neg => p'q; have [B EqB] := p_rank_witness q E. + have{EqB} Eq2B: B \in 'E_q^2(E). + by move: t2Mq; rewrite (tau2E hallE) => /andP[_ /eqP <-]. + have [sBE abelB dimB]:= pnElemP Eq2B; have [qB _] := andP abelB. + have coBA: coprime #|B| #|A| by exact: pnat_coprime qB (pi_pnat pA _). + have [[nsBE _] [sCBE _ _] _ _] := tau2_compl_context maxM hallE t2Mq Eq2B. + have nBA: A \subset 'N(B) by rewrite (subset_trans sAE) ?normal_norm. + have cAB: B \subset 'C(A). + rewrite (sameP commG1P trivgP) -(coprime_TIg coBA) subsetI commg_subl nBA. + by rewrite commg_subr (subset_trans sBE) ?normal_norm. + have{cAB} qCA: q %| #|'C(A)|. + by apply: dvdn_trans (cardSg cAB); rewrite (card_pnElem Eq2B) dvdn_mull. + have [Q maxQ sBQ] := max_normed_exists qB nBA. + have nnQ: q.-narrow Q. + apply/implyP=> _; apply/set0Pn; exists B. + rewrite 3!inE sBQ abelB dimB (subsetP (pmaxElemS q (subsetT Q))) //=. + rewrite setIC 2!inE sBQ; case: (tau2_not_beta maxM t2Mq) => _ -> //. + by rewrite (subsetP (pnElemS _ _ sEM)). + have [P1 [sylP1 _] [_ _]] := max_normed_2Elem_signaliser q'p max2A maxQ qCA. + move/(_ nnQ)=> cQP1; have sylP1_E: p.-Sylow(E) P1. + apply: pHall_subl (subset_trans _ sCBE) (subsetT E) sylP1. + exact: subset_trans (centS sBQ). + rewrite (card_Hall sylS) -(card_Hall sylP1). + by rewrite (card_Hall sylP_E) -(card_Hall sylP1_E). +have coMsA: coprime #|Ms| #|A|. + by exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat pA _). +have defMs: <<\bigcup_(X in 'E_p^1(A)) 'C_Ms(X)>> = Ms. + have ncycA: ~~ cyclic A by rewrite (abelem_cyclic abelA) dimA. + have [sAM _ _] := pnElemP Ep2A_M. + have{sAM} nMsA: A \subset 'N(Ms) by rewrite (subset_trans sAM) ?gFnorm. + apply/eqP; rewrite eqEsubset andbC gen_subG. + rewrite -{1}[Ms](coprime_abelian_gen_cent1 cAA ncycA nMsA coMsA). + rewrite genS; apply/bigcupsP=> x; rewrite ?subsetIl //; case/setD1P=> ntx Ax. + rewrite /= -cent_cycle (bigcup_max <[x]>%G) // p1ElemE // . + by rewrite 2!inE cycle_subG Ax /= -orderE (abelem_order_p abelA). +have [A0 EpA0 nregA0]: exists2 A0, A0 \in 'E_p^1(A) & 'C_Ms(A0) != 1. + apply/exists_inP; rewrite -negb_forall_in. + apply: contra (Msigma_neq1 maxM); move/forall_inP => regAp. + rewrite -/Ms -defMs -subG1 gen_subG; apply/bigcupsP=> X EpX. + by rewrite subG1 regAp. +have uniqCA0: 'M('C(A0)) = [set M]. + by rewrite nregA // (subsetP (pnElemS _ _ sAE)). +have defSM: S :&: M = P. + apply: sub_pHall (pgroupS (subsetIl S M) pS) _ (subsetIr S M) => //. + by rewrite subsetI sPS (pHall_sub sylP_M). +have{ltPS} not_sSM: ~~ (S \subset M). + by rewrite (sameP setIidPl eqP) defSM proper_neq. +have not_sA0Z: ~~ (A0 \subset 'Z(S)). + apply: contra not_sSM; rewrite subsetI centsC; case/andP=> _ sS_CA0. + by case/mem_uniq_mmax: uniqCA0 => _; exact: subset_trans sS_CA0. +have [EpZ0 dxCSA transNSA] := basic_p2maxElem_structure max2A pS sAS not_cSS. +do [set Z0 := 'Ohm_1('Z(S))%G; set EpA' := _ :\ Z0] in EpZ0 dxCSA transNSA. +have sZ0Z: Z0 \subset 'Z(S) := Ohm_sub 1 _. +have [sA0A _ _] := pnElemP EpA0; have sA0P := subset_trans sA0A sAP. +have EpA'_A0: A0 \in EpA'. + by rewrite 2!inE EpA0 andbT; apply: contraNneq not_sA0Z => ->. +have{transNSA sAP not_sSM defSM} regA0' X: + X \in 'E_p^1(E) :\ A0 -> 'C_Ms(X) = 1 /\ ~~ ('C(X) \subset M). +- case/setD1P=> neqXA0 EpX. + suffices not_sCXM: ~~ ('C(X) \subset M). + split=> //; apply: contraNeq not_sCXM => nregX. + by case/mem_uniq_mmax: (nregA X EpX nregX). + have [sXZ | not_sXZ] := boolP (X \subset 'Z(S)). + apply: contra (subset_trans _) not_sSM. + by rewrite centsC (subset_trans sXZ) ?subsetIr. + have EpA'_X: X \in EpA'. + by rewrite 2!inE -defEp EpX andbT; apply: contraNneq not_sXZ => ->. + have [g NSAg /= defX] := atransP2 transNSA EpA'_A0 EpA'_X. + have{NSAg} [Sg nAg] := setIP NSAg. + have neqMgM: M :^ g != M. + rewrite (sameP eqP normP) (norm_mmax maxM); apply: contra neqXA0 => Mg. + rewrite defX [_ == _](sameP eqP normP) (subsetP (cent_sub A0)) //. + by rewrite (subsetP (centSS (subxx _) sA0P cPP)) // -defSM inE Sg. + apply: contra neqMgM; rewrite defX centJ sub_conjg. + by move/(eq_uniq_mmax uniqCA0) => defM; rewrite -{1}defM ?mmaxJ ?actKV. +have{defMs} defA0: 'C_A(Ms) = A0. + apply/eqP; rewrite eq_sym eqEcard subsetI sA0A (card_pnElem EpA0). + have pCA: p.-group 'C_A(Ms) := pgroupS (subsetIl A _) pA. + rewrite andbC (card_pgroup pCA) leq_exp2l ?prime_gt1 // -ltnS -dimA. + rewrite properG_ltn_log //=; last first. + rewrite properE subsetIl /= subsetI subxx centsC -(subxx Ms) -subsetI. + by rewrite regA subG1 Msigma_neq1. + rewrite centsC -defMs gen_subG (big_setD1 A0) //= subUset subsetIr /=. + by apply/bigcupsP=> X; rewrite -defEp; case/regA0'=> -> _; rewrite sub1G. +rewrite defA0 (group_inj defA0) (card_pnElem EpA0). +have{dxCSA} [Y [cycY sZ0Y]] := dxCSA; move/(_ _ EpA'_A0)=> dxCSA. +have defCP_Ms: 'C_P(Ms) = A0. + move: dxCSA; rewrite defCSA => /dprodP[_ mulA0Y cA0Y tiA0Y]. + rewrite -mulA0Y -group_modl /=; last by rewrite -defA0 subsetIr. + rewrite setIC TI_Ohm1 ?mulg1 // setIC. + have pY: p.-group Y by rewrite (pgroupS _ pP) // -mulA0Y mulG_subr. + have cYY: abelian Y := cyclic_abelian cycY. + have ->: 'Ohm_1(Y) = Z0. + apply/eqP; rewrite eq_sym eqEcard (card_pnElem EpZ0) /= -['Ohm_1(_)]Ohm_id. + rewrite OhmS // (card_pgroup (pgroupS (Ohm_sub 1 Y) pY)). + rewrite leq_exp2l ?prime_gt1 -?p_rank_abelian // -rank_pgroup //. + by rewrite -abelian_rank1_cyclic. + rewrite prime_TIg ?(card_pnElem EpZ0) // centsC (sameP setIidPl eqP) eq_sym. + case: (regA0' Z0) => [|-> _]; last exact: Msigma_neq1. + by rewrite 2!inE defEp EpZ0 andbT; apply: contraNneq not_sA0Z => <-. +have [sPM pA0] := (pHall_sub sylP_M, pgroupS sA0A pA). +have cMsA0: A0 \subset 'C(Ms) by rewrite -defA0 subsetIr. +have nsA0M: A0 <| M. + have [_ defM nMsE _] := sdprodP (sdprod_sigma maxM hallE). + rewrite /normal (subset_trans sA0P) // -defM mulG_subG cents_norm 1?centsC //. + by rewrite -defA0 normsI ?norms_cent // normal_norm. +have defFM: Ms \x A0 = 'F(M). + have nilF := Fitting_nil M. + rewrite dprodE ?(coprime_TIg (coprimegS sA0A coMsA)) //. + have [_ /= defFM cFpp' _] := dprodP (nilpotent_pcoreC p nilF). + have defFp': 'O_p^'('F(M)) = Ms. + apply/eqP; rewrite eqEsubset. + rewrite (sub_Hall_pcore (Msigma_Hall maxM)); last first. + exact: subset_trans (pcore_sub _ _) (Fitting_sub _). + rewrite /pgroup (sub_in_pnat _ (pcore_pgroup _ _)) => [|q piFq]; last first. + have [Q sylQ] := Sylow_exists q 'F(M); have [sQF qQ _] := and3P sylQ. + have ntQ: Q :!=: 1. + rewrite -rank_gt0 (rank_pgroup qQ) (p_rank_Sylow sylQ) p_rank_gt0. + by rewrite (piSg _ piFq) ?pcore_sub. + have sNQM: 'N(Q) \subset M. + rewrite (mmax_normal maxM) // (nilpotent_Hall_pcore nilF sylQ). + by rewrite p_core_Fitting pcore_normal. + apply/implyP; rewrite implyNb /= -def_t2 orbC. + by rewrite (prime_class_mmax_norm maxM qQ). + rewrite pcore_max ?(pi_p'group (pcore_pgroup _ _)) //. + rewrite /normal (subset_trans (Fitting_sub M)) ?gFnorm //. + rewrite Fitting_max ?pcore_normal ?(tau2_Msigma_nil _ t2Mp) //. + rewrite p_core_Fitting defFp' centsC in defFM cFpp'. + rewrite -defFM (centC cFpp'); congr (Ms * _). + apply/eqP; rewrite eqEsubset pcore_max //. + by rewrite -defCP_Ms subsetI cFpp' pcore_sub_Hall. +split=> {defFM}//. +have [[sE1E t1E1 _] t2E2] := (and3P hallE1, pHall_pgroup hallE2). +have defE2: E2 :=: P by rewrite (sub_pHall sylP) // -(eq_pgroup _ def_t2) t2E2. +have [[_ nsE3E] _ _ [defEr _] _] := sigma_compl_context maxM complEi. +have [sE3E nE3E] := andP nsE3E; have{nE3E} nE3E := subset_trans _ nE3E. +have [[_ E21 _ defE21]] := sdprodP defEr; rewrite defE21 => defE nE321 tiE321. +rewrite defE2 in defE21; have{defE21} [_ defPE1 nPE1 tiPE1] := sdprodP defE21. +have [P0 defP nP0E1]: exists2 P0 : {group gT}, A0 \x P0 = P & E1 \subset 'N(P0). + have p'E1b: p^'.-group (E1 / 'Phi(P)). + rewrite quotient_pgroup //; apply: sub_pgroup t1E1 => q /tau2'1. + by rewrite inE /= def_t2. + have defPhiP: 'Phi(P) = 'Phi(Y). + have [_ _ cA0Y tiA0Y] := dprodP dxCSA. + rewrite defCSA dprodEcp // in dxCSA. + have [_ abelA0 _] := pnElemP EpA0; rewrite -trivg_Phi // in abelA0. + by rewrite -(Phi_cprod pP dxCSA) (eqP abelA0) cprod1g. + have abelPb := Phi_quotient_abelem pP; have sA0Pb := quotientS 'Phi(P) sA0P. + have [||P0b] := Maschke_abelem abelPb p'E1b sA0Pb; rewrite ?quotient_norms //. + by rewrite (subset_trans (subset_trans sE1E sEM)) ?normal_norm. + case/dprodP=> _ defPb _ tiAP0b nP0E1b. + have sP0Pb: P0b \subset P / 'Phi(P) by rewrite -defPb mulG_subr. + have [P0 defP0b sPhiP0 sP0P] := inv_quotientS (Phi_normal P) sP0Pb. + exists P0; last first. + rewrite -(quotientSGK (char_norm_trans (Phi_char P) nPE1)); last first. + by rewrite cents_norm ?(sub_abelian_cent2 cPP (Phi_sub P) sP0P). + by rewrite quotient_normG -?defP0b ?(normalS _ _ (Phi_normal P)). + rewrite dprodEY ?(sub_abelian_cent2 cPP) //. + apply/eqP; rewrite eqEsubset join_subG sA0P sP0P /=. + rewrite -(quotientSGK (normal_norm (Phi_normal P))) //=; last first. + by rewrite sub_gen // subsetU // sPhiP0 orbT. + rewrite cent_joinEr ?(sub_abelian_cent2 cPP) //. + rewrite quotientMr //; last by rewrite (subset_trans sP0P) ?gFnorm. + by rewrite -defP0b defPb. + apply/trivgP; case/dprodP: dxCSA => _ _ _ <-. + rewrite subsetI subsetIl (subset_trans _ (Phi_sub Y)) // -defPhiP. + rewrite -quotient_sub1 ?subIset ?(subset_trans sA0P) ?gFnorm //. + by rewrite quotientIG // -defP0b tiAP0b. +have nA0E := subset_trans _ (subset_trans sEM (normal_norm nsA0M)). +have{defP} [_ defAP0 _ tiAP0] := dprodP defP. +have sP0P: P0 \subset P by rewrite -defAP0 mulG_subr. +have sP0E := subset_trans sP0P (pHall_sub sylP_E). +pose E0 := (E3 <*> (P0 <*> E1))%G. +have sP0E1_E: P0 <*> E1 \subset E by rewrite join_subG sP0E. +have sE0E: E0 \subset E by rewrite join_subG sE3E. +have mulA0E0: A0 * E0 = E. + rewrite /= (norm_joinEr (nE3E _ sP0E1_E)) mulgA -(normC (nA0E _ sE3E)). + by rewrite /= -mulgA (norm_joinEr nP0E1) (mulgA A0) defAP0 defPE1. +have tiA0E0: A0 :&: E0 = 1. + rewrite cardMg_TI // mulA0E0 -defE /= (norm_joinEr (nE3E _ sP0E1_E)). + rewrite !TI_cardMg //; last first. + by apply/trivgP; rewrite -tiE321 setIS //= ?norm_joinEr // -defPE1 mulSg. + rewrite mulnCA /= leq_mul // norm_joinEr //= -defPE1. + rewrite !TI_cardMg //; last by apply/trivgP; rewrite -tiPE1 setSI. + by rewrite mulnA -(TI_cardMg tiAP0) defAP0. +have defAE0: A0 ><| E0 = E by rewrite sdprodE ?nA0E. +exists E0 => // x /setD1P[ntx Ms_x] q piCE0x_q. +have: q \in \pi(E) by apply: piSg piCE0x_q; rewrite subIset ?sE0E. +rewrite mem_primes in piCE0x_q; case/and3P: piCE0x_q => q_pr _. +case/Cauchy=> // z /setIP[E0z cxz] oz. +have ntz: z != 1 by rewrite -order_gt1 oz prime_gt1. +have ntCMs_z: 'C_Ms[z] != 1. + apply: contraNneq ntx => reg_z; rewrite (sameP eqP set1gP) -reg_z inE Ms_x. + by rewrite cent1C. +rewrite (partition_pi_sigma_compl maxM hallE) def_t2. +case/or3P => [-> // | pq | t3Mq]. + have EpA0'_z: <[z]>%G \in 'E_p^1(E) :\ A0. + rewrite p1ElemE // !inE -orderE oz (eqnP pq) eqxx cycle_subG. + rewrite (subsetP sE0E) // !andbT; apply: contraNneq ntz => eqA0z. + by rewrite (sameP eqP set1gP) -tiA0E0 inE -eqA0z cycle_id E0z. + have [reg_z _] := regA0' _ EpA0'_z. + by rewrite -cent_cycle reg_z eqxx in ntCMs_z. +rewrite regE3 ?eqxx // !inE ntz /= in ntCMs_z. +by rewrite (mem_normal_Hall hallE3 nsE3E) ?(subsetP sE0E) // /p_elt oz pnatE. +Qed. + +(* This is B & G, Theorem 12.8(c). This part does not use the decompotision *) +(* of the complement, and needs to be proved ahead because it is used with *) +(* different primes in \tau_2(M) in the main argument. We also include an *) +(* auxiliary identity, which is needed in another part of the proof of 12.8. *) +Theorem abelian_tau2_sub_Fitting M E p A S : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> + p \in \tau2(M) -> A \in 'E_p^2(E) -> + p.-Sylow(G) S -> A \subset S -> abelian S -> + [/\ S \subset 'N(S)^`(1), + 'N(S)^`(1) \subset 'F(E), + 'F(E) \subset 'C(S), + 'C(S) \subset E + & 'F('N(A)) = 'F(E)]. +Proof. +move=> maxM hallE t2Mp Ep2A sylS sAS cSS. +have [sAE abelA dimA]:= pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have [sEM sM'E _] := and3P hallE. +have Ep2A_M := subsetP (pnElemS p 2 sEM) A Ep2A. +have eqFC H: A <| H -> 'C(A) \subset H -> 'F(H) = 'F('C(A)). + move=> nsAH sCH; have [_ nAH] := andP nsAH. + apply/eqP; rewrite eqEsubset !Fitting_max ?Fitting_nil //. + by rewrite (char_normal_trans (Fitting_char _)) // /normal sCH norms_cent. + apply: normalS sCH (Fitting_normal H). + have [_ defF cFpFp' _] := dprodP (nilpotent_pcoreC p (Fitting_nil H)). + have sAFp: A \subset 'O_p('F(H)) by rewrite p_core_Fitting pcore_max. + have [x _ sFpSx] := Sylow_subJ sylS (subsetT _) (pcore_pgroup p 'F(H)). + have cFpFp: abelian 'O_p('F(H)) by rewrite (abelianS sFpSx) ?abelianJ. + by rewrite -defF mulG_subG (centSS _ _ cFpFp) // (centSS _ _ cFpFp'). +have [[nsAE _] [sCAE _] _ _ _] := tau2_compl_context maxM hallE t2Mp Ep2A. +have eqFN_FE: 'F('N(A)) = 'F(E) by rewrite (eqFC E) // eqFC ?cent_sub ?normalG. +have sN'FN: 'N(A)^`(1) \subset 'F('N(A)). + rewrite rank2_der1_sub_Fitting ?mFT_odd //. + rewrite ?mFT_sol // mFT_norm_proper ?(mFT_pgroup_proper pA) //. + by rewrite -rank_gt0 (rank_abelem abelA) dimA. + rewrite eqFN_FE (leq_trans (rankS (Fitting_sub E))) //. + have [q q_pr ->]:= rank_witness E; apply: wlog_neg; rewrite -ltnNge => rEgt2. + rewrite (leq_trans (p_rankS q sEM)) // leqNgt. + apply: contra ((alpha_sub_sigma maxM) q) (pnatPpi sM'E _). + by rewrite -p_rank_gt0 (leq_trans _ rEgt2). +have sSE: S \subset E by rewrite (subset_trans _ sCAE) // (centSS _ _ cSS). +have nA_NS: 'N(S) \subset 'N(A). + have [ ] := tau2_context maxM t2Mp Ep2A_M; have sSM := subset_trans sSE sEM. + have sylS_M: p.-Sylow(M) S := pHall_subl sSM (subsetT M) sylS. + by case/(_ S) => // _ [// |<- _] _ _ _ _; exact: char_norms (Ohm_char 1 _). +have sS_NS': S \subset 'N(S)^`(1) := mFT_Sylow_der1 sylS. +have sNS'_FE: 'N(S)^`(1) \subset 'F(E). + by rewrite -eqFN_FE (subset_trans (dergS 1 nA_NS)). +split=> //; last by rewrite (subset_trans (centS sAS)). +have sSFE := subset_trans sS_NS' sNS'_FE; have nilFE := Fitting_nil E. +have sylS_FE := pHall_subl sSFE (subsetT 'F(E)) sylS. +suff sSZF: S \subset 'Z('F(E)) by rewrite centsC (subset_trans sSZF) ?subsetIr. +have [_ <- _ _] := dprodP (center_dprod (nilpotent_pcoreC p nilFE)). +by rewrite -(nilpotent_Hall_pcore nilFE sylS_FE) (center_idP cSS) mulG_subl. +Qed. + +(* This is B & G, Theorem 12.8(a,b,d,e) -- the bulk of the Theorem. We prove *) +(* part (f) separately below, as it does not depend on a decomposition of the *) +(* complement. *) +Theorem abelian_tau2 M E E1 E2 E3 p A S (Ms := M`_\sigma) : + M \in 'M -> sigma_complement M E E1 E2 E3 -> + p \in \tau2(M) -> A \in 'E_p^2(E) -> + p.-Sylow(G) S -> A \subset S -> abelian S -> + [/\ (*a*) E2 <| E /\ abelian E2, + (*b*) \tau2(M).-Hall(G) E2, + (*d*) [/\ 'N(A) = 'N(S), + 'N(S) = 'N(E2), + 'N(E2) = 'N(E3 <*> E2) + & 'N(E3 <*> E2) = 'N('F(E))] + & (*e*) forall X, X \in 'E^1(E1) -> 'C_Ms(X) = 1 -> X \subset 'Z(E)]. +Proof. +move=> maxM complEi t2Mp Ep2A sylS sAS cSS. +have [hallE hallE1 hallE2 hallE3 _] := complEi. +have [sEM sM'E _] := and3P hallE. +have [sE1E t1E1 _] := and3P hallE1. +have [sE2E t2E2 _] := and3P hallE2. +have [sE3E t3E3 _] := and3P hallE3. +have nilF: nilpotent 'F(E) := Fitting_nil E. +have sylE2_sylG_ZFE q Q: + q.-Sylow(E2) Q -> Q :!=: 1 -> q.-Sylow(G) Q /\ Q \subset 'Z('F(E)). +- move=> sylQ ntQ; have [sQE2 qQ _] := and3P sylQ. + have piQq: q \in \pi(Q) by rewrite -p_rank_gt0 -rank_pgroup // rank_gt0. + have t2Mq: q \in \tau2(M) by rewrite (pnatPpi t2E2) // (piSg sQE2). + have sylQ_E: q.-Sylow(E) Q := subHall_Sylow hallE2 t2Mq sylQ. + have rqQ: 'r_q(Q) = 2. + rewrite (tau2E hallE) !inE -(p_rank_Sylow sylQ_E) in t2Mq. + by case/andP: t2Mq => _ /eqP. + have [B Eq2B sBQ]: exists2 B, B \in 'E_q^2(E) & B \subset Q. + have [B Eq2B] := p_rank_witness q Q; have [sBQ abelB rBQ] := pnElemP Eq2B. + exists B; rewrite // !inE rBQ rqQ abelB !andbT. + exact: subset_trans sBQ (pHall_sub sylQ_E). + have [T /= sylT sQT] := Sylow_superset (subsetT Q) qQ. + have qT: q.-group T := pHall_pgroup sylT. + have cTT: abelian T. + apply: wlog_neg => not_cTT. + have [def_t2 _ _ _] := nonabelian_tau2 maxM hallE t2Mq Eq2B qT not_cTT. + rewrite def_t2 !inE in t2Mp; rewrite (eqP t2Mp) in sylS. + by have [x _ ->] := Sylow_trans sylS sylT; rewrite abelianJ. + have sTF: T \subset 'F(E). + have subF := abelian_tau2_sub_Fitting maxM hallE t2Mq Eq2B sylT. + have [sTN' sN'F _ _ _] := subF (subset_trans sBQ sQT) cTT. + exact: subset_trans sTN' sN'F. + have sTE: T \subset E := subset_trans sTF (Fitting_sub E). + have <-: T :=: Q by apply: (sub_pHall sylQ_E). + have sylT_F: q.-Sylow('F(E)) T := pHall_subl sTF (subsetT _) sylT. + have [_ <- _ _] := dprodP (center_dprod (nilpotent_pcoreC q nilF)). + by rewrite -(nilpotent_Hall_pcore nilF sylT_F) (center_idP cTT) mulG_subl. +have hallE2_G: \tau2(M).-Hall(G) E2. + rewrite pHallE subsetT /= -(part_pnat_id t2E2); apply/eqnP. + rewrite !(widen_partn _ (subset_leq_card (subsetT _))). + apply: eq_bigr => q t2q; rewrite -!p_part. + have [Q sylQ] := Sylow_exists q E2; have qQ := pHall_pgroup sylQ. + have sylQ_E: q.-Sylow(E) Q := subHall_Sylow hallE2 t2q sylQ. + have ntQ: Q :!=: 1. + rewrite -rank_gt0 (rank_pgroup qQ) (p_rank_Sylow sylQ_E) p_rank_gt0. + by rewrite (tau2E hallE) in t2q; case/andP: t2q. + have [sylQ_G _] := sylE2_sylG_ZFE q Q sylQ ntQ. + by rewrite -(card_Hall sylQ) -(card_Hall sylQ_G). +have sE2_ZFE: E2 \subset 'Z('F(E)). + rewrite -Sylow_gen gen_subG; apply/bigcupsP=> Q; case/SylowP=> q q_pr sylQ. + have [-> | ntQ] := eqsVneq Q 1; first exact: sub1G. + by have [_ ->] := sylE2_sylG_ZFE q Q sylQ ntQ. +have cE2E2: abelian E2 := abelianS sE2_ZFE (center_abelian _). +have sE2FE: E2 \subset 'F(E) := subset_trans sE2_ZFE (center_sub _). +have nsE2E: E2 <| E. + have hallE2_F := pHall_subl sE2FE (Fitting_sub E) hallE2. + rewrite (nilpotent_Hall_pcore nilF hallE2_F). + exact: char_normal_trans (pcore_char _ _) (Fitting_normal E). +have [_ _ [cycE1 cycE3] [_ defEl] _] := sigma_compl_context maxM complEi. +have [[K _ defK _] _ _ _] := sdprodP defEl; rewrite defK in defEl. +have [nsKE _ mulKE1 nKE1 _] := sdprod_context defEl; have [sKE _] := andP nsKE. +have [nsE3K sE2K _ nE32 tiE32] := sdprod_context defK. +rewrite -sdprodEY // defK. +have{defK} defK: E3 \x E2 = K. + rewrite dprodEsd // (sameP commG1P trivgP) -tiE32 subsetI commg_subr nE32. + by rewrite commg_subl (subset_trans sE3E) ?normal_norm. +have cKK: abelian K. + by have [_ <- cE23 _] := dprodP defK; rewrite abelianM cE2E2 cyclic_abelian. +have [_ sNS'F _ sCS_E defFN] := + abelian_tau2_sub_Fitting maxM hallE t2Mp Ep2A sylS sAS cSS. +have{sCS_E} sSE2: S \subset E2. + rewrite (sub_normal_Hall hallE2 nsE2E (subset_trans cSS sCS_E)). + by rewrite (pi_pgroup (pHall_pgroup sylS)). +have charS: S \char E2. + have sylS_E2: p.-Sylow(E2) S := pHall_subl sSE2 (subsetT E2) sylS. + by rewrite (nilpotent_Hall_pcore (abelian_nil cE2E2) sylS_E2) pcore_char. +have nsSE: S <| E := char_normal_trans charS nsE2E; have [sSE nSE] := andP nsSE. +have charA: A \char S. + have Ep2A_M := subsetP (pnElemS p 2 sEM) A Ep2A. + have sylS_M := pHall_subl (subset_trans sSE sEM) (subsetT M) sylS. + have [] := tau2_context maxM t2Mp Ep2A_M; case/(_ S sylS_M)=> _ [//|<- _]. + by rewrite Ohm_char. +have charE2: E2 \char K. + have hallE2_K := pHall_subl sE2K sKE hallE2. + by rewrite (nilpotent_Hall_pcore (abelian_nil cKK) hallE2_K) pcore_char. +have coKE1: coprime #|K| #|E1|. + rewrite -(dprod_card defK) coprime_mull (sub_pnat_coprime tau3'1 t3E3 t1E1). + by rewrite (sub_pnat_coprime tau2'1 t2E2 t1E1). +have hallK: Hall 'F(E) K. + have hallK: Hall E K. + by rewrite /Hall -divgS sKE //= -(sdprod_card defEl) mulKn. + have sKFE: K \subset 'F(E) by rewrite Fitting_max ?abelian_nil. + exact: pHall_Hall (pHall_subl sKFE (Fitting_sub E) (Hall_pi hallK)). +have charK: K \char 'F(E). + by rewrite (nilpotent_Hall_pcore nilF (Hall_pi hallK)) pcore_char. +have{defFN} [eqNAS eqNSE2 eqNE2K eqNKF]: + [/\ 'N(A) = 'N(S), 'N(S) = 'N(E2), 'N(E2) = 'N(K) & 'N(K) = 'N('F(E))]. + have: #|'N(A)| <= #|'N('F(E))|. + by rewrite subset_leq_card // -defFN gFnorm. + have leCN := subset_leqif_cards (@char_norms gT _ _ _). + have trCN := leqif_trans (leCN _ _ _). + have leq_KtoA := trCN _ _ _ _ charE2 (trCN _ _ _ _ charS (leCN _ _ charA)). + rewrite (geq_leqif (trCN _ _ _ _ charK leq_KtoA)). + by case/and4P; do 4!move/eqP->. +split=> // X E1_1_X regX. +have cK_NK': 'N(K)^`(1) \subset 'C(K). + suffices sKZ: K \subset 'Z('F(E)). + by rewrite -eqNE2K -eqNSE2 (centSS sNS'F sKZ) // centsC subsetIr. + have{hallK} [pi hallK] := HallP hallK. + have [_ <- _ _] := dprodP (center_dprod (nilpotent_pcoreC pi nilF)). + by rewrite -(nilpotent_Hall_pcore nilF hallK) (center_idP cKK) mulG_subl. +have [q EqX] := nElemP E1_1_X; have [sXE1 abelX dimX] := pnElemP EqX. +have sXE := subset_trans sXE1 sE1E. +have nKX := subset_trans sXE (normal_norm nsKE). +have nCSX_NS: 'N(K) \subset 'N('C(K) * X). + rewrite -(quotientGK (cent_normal _)) -quotientK ?norms_cent //. + by rewrite morphpre_norms // sub_abelian_norm ?quotientS ?sub_der1_abelian. +have nKX_NS: 'N(S) \subset 'N([~: K, X]). + have CK_K_1: [~: 'C(K), K] = 1 by apply/commG1P. + rewrite eqNSE2 eqNE2K commGC -[[~: X, K]]mul1g -CK_K_1. + by rewrite -commMG ?CK_K_1 ?norms1 ?normsR. +have not_sNKX_M: ~~ ('N([~: K, X]) \subset M). + have [[sM'p _] sSM] := (andP t2Mp, subset_trans sSE sEM). + apply: contra sM'p => sNKX_M; apply/existsP; exists S. + by rewrite (pHall_subl sSM (subsetT _) sylS) // (subset_trans _ sNKX_M). +have cKX: K \subset 'C(X). + apply: contraR not_sNKX_M; rewrite (sameP commG1P eqP) => ntKX. + rewrite (mmax_normal maxM) //. + have [sKM sM'K] := (subset_trans sKE sEM, pgroupS sKE sM'E). + have piE1q: q \in \pi(E1). + by rewrite -p_rank_gt0 -dimX logn_le_p_rank // inE sXE1. + have sM'q: q \in \sigma(M)^' by rewrite (pnatPpi sM'E) // (piSg sE1E). + have EpX_NK: X \in 'E_q^1('N_M(K)). + by apply: subsetP EqX; rewrite pnElemS // subsetI (subset_trans sE1E). + have q'K: q^'.-group K. + by rewrite p'groupEpi ?coprime_pi' // in coKE1 *; apply: (pnatPpi coKE1). + by have []:= commG_sigma'_1Elem_cyclic maxM sKM sM'K sM'q EpX_NK regX. +rewrite subsetI sXE /= -mulKE1 centM subsetI centsC cKX. +exact: subset_trans sXE1 (cyclic_abelian cycE1). +Qed. + +(* This is B & G, Theorem 12.8(f). *) +Theorem abelian_tau2_norm_Sylow M E p A S : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> p \in \tau2(M) -> A \in 'E_p^2(E) -> + p.-Sylow(G) S -> A \subset S -> abelian S -> + forall X, X \subset 'N(S) -> 'C_S(X) <| 'N(S) /\ [~: S, X] <| 'N(S). +Proof. +move=> maxM hallE t2Mp Ep2A sylS sAS cSS X nSX. +have [_ sNS'F sFCS _ _] := + abelian_tau2_sub_Fitting maxM hallE t2Mp Ep2A sylS sAS cSS. +have{sNS'F sFCS} sNS'CS: 'N(S)^`(1) \subset 'C(S) := subset_trans sNS'F sFCS. +have nCSX_NS: 'N(S) \subset 'N('C(S) * X). + rewrite -quotientK ?norms_cent // -{1}(quotientGK (cent_normal S)). + by rewrite morphpre_norms // sub_abelian_norm ?quotientS ?sub_der1_abelian. +rewrite /normal subIset ?comm_subG ?normG //=; split. + have ->: 'C_S(X) = 'C_S('C(S) * X). + by rewrite centM setIA; congr (_ :&: _); rewrite (setIidPl _) // centsC. + by rewrite normsI ?norms_cent. +have CS_S_1: [~: 'C(S), S] = 1 by exact/commG1P. +by rewrite commGC -[[~: X, S]]mul1g -CS_S_1 -commMG ?CS_S_1 ?norms1 ?normsR. +Qed. + +(* This is B & G, Corollary 12.9. *) +Corollary tau1_act_tau2 M E p A q Q (Ms := M`_\sigma) : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> p \in \tau2(M) -> A \in 'E_p^2(E) -> + q \in \tau1(M) -> Q \in 'E_q^1(E) -> 'C_Ms(Q) = 1 -> [~: A, Q] != 1 -> + let A0 := [~: A, Q]%G in let A1 := ('C_A(Q))%G in + [/\ (*a*) [/\ A0 \in 'E_p^1(A), 'C_A(Ms) = A0 & A0 <| M], + (*b*) gval A0 \notin A1 :^: G + & (*c*) A1 \in 'E_p^1(A) /\ ~~ ('C(A1) \subset M)]. +Proof. +move=> maxM hallE t2Mp Ep2A t1Mq EqQ regQ ntA0 A0 A1. +have [sEM sM'E _] := and3P hallE. +have [sAE abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have [sQE abelQ dimQ] := pnElemP EqQ; have [qQ _ _] := and3P abelQ. +have [p_pr q_pr] := (pnElem_prime Ep2A, pnElem_prime EqQ). +have p_gt1 := prime_gt1 p_pr. +have Ep2A_M := subsetP (pnElemS p 2 sEM) A Ep2A. +have [_ _ regA _ _] := tau2_context maxM t2Mp Ep2A_M. +have [[nsAE _] _ _ _] := tau2_compl_context maxM hallE t2Mp Ep2A. +have [_ nAE] := andP nsAE; have nAQ := subset_trans sQE nAE. +have coAQ: coprime #|A| #|Q|. + exact: sub_pnat_coprime tau2'1 (pi_pnat pA t2Mp) (pi_pnat qQ t1Mq). +have defA: A0 \x A1 = A := coprime_abelian_cent_dprod nAQ coAQ cAA. +have [_ _ _ tiA01] := dprodP defA. +have [sAM sM'A] := (subset_trans sAE sEM, pgroupS sAE sM'E). +have sM'q: q \in \sigma(M)^' by case/andP: t1Mq. +have EqQ_NA: Q \in 'E_q^1('N_M(A)). + by apply: subsetP EqQ; rewrite pnElemS // subsetI sEM. +have q'A: q^'.-group A. + rewrite p'groupEpi ?coprime_pi' // in coAQ *. + by apply: (pnatPpi coAQ); rewrite -p_rank_gt0 (p_rank_abelem abelQ) dimQ. +have [] := commG_sigma'_1Elem_cyclic maxM sAM sM'A sM'q EqQ_NA regQ q'A cAA. +rewrite -[[~: A, Q]]/(gval A0) -/Ms => cMsA0 cycA0 nsA0M. +have sA0A: A0 \subset A by rewrite commg_subl. +have EpA0: A0 \in 'E_p^1(A). + have abelA0 := abelemS sA0A abelA; have [pA0 _] := andP abelA0. + rewrite p1ElemE // !inE sA0A -(Ohm1_id abelA0) /=. + by rewrite (Ohm1_cyclic_pgroup_prime cycA0 pA0). +have defA0: 'C_A(Ms) = A0. + apply/eqP; rewrite eq_sym eqEcard subsetI sA0A cMsA0 (card_pnElem EpA0). + have pCAMs: p.-group 'C_A(Ms) := pgroupS (subsetIl A _) pA. + rewrite (card_pgroup pCAMs) leq_exp2l //= leqNgt. + apply: contra (Msigma_neq1 maxM) => dimCAMs. + rewrite eq_sym -regA (sameP eqP setIidPl) centsC (sameP setIidPl eqP). + by rewrite eqEcard subsetIl (card_pnElem Ep2A) (card_pgroup pCAMs) leq_exp2l. +have EpA1: A1 \in 'E_p^1(A). + rewrite p1ElemE // !inE subsetIl -(eqn_pmul2l (ltnW p_gt1)). + by rewrite -{1}[p](card_pnElem EpA0) (dprod_card defA) (card_pnElem Ep2A) /=. +have defNA0: 'N(A0) = M by apply: mmax_normal. +have not_cA0Q: ~~ (Q \subset 'C(A0)). + apply: contra ntA0 => cA0Q. + by rewrite -subG1 -tiA01 !subsetI subxx sA0A centsC cA0Q. +have rqM: 'r_q(M) = 1%N by apply/eqP; case/and3P: t1Mq. +have q'CA0: q^'.-group 'C(A0). + have [S sylS sQS] := Sylow_superset (subset_trans sQE sEM) qQ. + have qS := pHall_pgroup sylS; rewrite -(p_rank_Sylow sylS) in rqM. + have cycS: cyclic S by rewrite (odd_pgroup_rank1_cyclic qS) ?mFT_odd ?rqM. + have ntS: S :!=: 1 by rewrite -rank_gt0 (rank_pgroup qS) rqM. + have defS1: 'Ohm_1(S) = Q. + apply/eqP; rewrite eq_sym eqEcard -{1}(Ohm1_id abelQ) OhmS //=. + by rewrite (card_pnElem EqQ) (Ohm1_cyclic_pgroup_prime cycS qS). + have sylSC: q.-Sylow('C(A0)) 'C_S(A0). + by rewrite (Hall_setI_normal _ sylS) // -defNA0 cent_normal. + rewrite -partG_eq1 ?cardG_gt0 // -(card_Hall sylSC) -trivg_card1 /=. + by rewrite setIC TI_Ohm1 // defS1 setIC prime_TIg ?(card_pnElem EqQ). +do 2?split=> //. + have: ~~ q^'.-group Q by rewrite /pgroup (card_pnElem EqQ) pnatE ?inE ?negbK. + apply: contra; case/imsetP=> x _ defA01. + rewrite defA01 centJ pgroupJ in q'CA0. + by apply: pgroupS q'CA0; rewrite centsC subsetIr. +have [S sylS sAS] := Sylow_superset (subsetT A) pA. +have [cSS | not_cSS] := boolP (abelian S). + have solE := sigma_compl_sol hallE. + have [E1 hallE1 sQE1] := Hall_superset solE sQE (pi_pnat qQ t1Mq). + have [E3 hallE3] := Hall_exists \tau3(M) solE. + have [E2 _ complEi] := ex_tau2_compl hallE hallE1 hallE3. + have [_ _ _ reg1Z] := abelian_tau2 maxM complEi t2Mp Ep2A sylS sAS cSS. + have E1Q: Q \in 'E^1(E1). + by apply/nElemP; exists q; rewrite // !inE sQE1 abelQ dimQ. + rewrite (subset_trans (reg1Z Q E1Q regQ)) ?subIset // in not_cA0Q. + by rewrite centS ?orbT // (subset_trans sA0A). +have pS := pHall_pgroup sylS. +have [_ _ not_cent_reg _] := nonabelian_tau2 maxM hallE t2Mp Ep2A pS not_cSS. +case: (not_cent_reg A1); rewrite // 2!inE (subsetP (pnElemS p 1 sAE)) // andbT. +by rewrite -val_eqE /= defA0 eq_sym; apply/(TIp1ElemP EpA0 EpA1). +Qed. + +(* This is B & G, Corollary 12.10(a). *) +Corollary sigma'_nil_abelian M N : + M \in 'M -> N \subset M -> \sigma(M)^'.-group N -> nilpotent N -> abelian N. +Proof. +move=> maxM sNM sM'N /nilpotent_Fitting defN. +apply/center_idP; rewrite -{2}defN -{defN}(center_bigdprod defN). +apply: eq_bigr => p _; apply/center_idP; set P := 'O_p(N). +have [-> | ntP] := eqVneq P 1; first exact: abelian1. +have [pP sPN]: p.-group P /\ P \subset N by rewrite pcore_sub pcore_pgroup. +have{sPN sNM sM'N} [sPM sM'P] := (subset_trans sPN sNM, pgroupS sPN sM'N). +have{sPM sM'P} [E hallE sPE] := Hall_superset (mmax_sol maxM) sPM sM'P. +suffices [S sylS cSS]: exists2 S : {group gT}, p.-Sylow(E) S & abelian S. + by have [x _ sPS] := Sylow_subJ sylS sPE pP; rewrite (abelianS sPS) ?abelianJ. +have{N P sPE pP ntP} piEp: p \in \pi(E). + by rewrite (piSg sPE) // -p_rank_gt0 -rank_pgroup // rank_gt0. +rewrite (partition_pi_sigma_compl maxM hallE) orbCA orbC -orbA in piEp. +have [[E1 hallE1] [E3 hallE3]] := ex_tau13_compl hallE. +have [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. +have{complEi} [_ _ [cycE1 cycE3] _ _] := sigma_compl_context maxM complEi. +have{piEp} [t1p | t3p | t2p] := or3P piEp. +- have [S sylS] := Sylow_exists p E1. + exists S; first exact: subHall_Sylow hallE1 t1p sylS. + exact: abelianS (pHall_sub sylS) (cyclic_abelian cycE1). +- have [S sylS] := Sylow_exists p E3. + exists S; first exact: subHall_Sylow hallE3 t3p sylS. + exact: abelianS (pHall_sub sylS) (cyclic_abelian cycE3). +have [s'p rpM] := andP t2p; have [S sylS] := Sylow_exists p E; exists S => //. +have sylS_M := subHall_Sylow hallE s'p sylS. +have [A _ Ep2A] := ex_tau2Elem hallE t2p. +by have [/(_ S sylS_M)[]] := tau2_context maxM t2p Ep2A. +Qed. + +(* This is B & G, Corollary 12.10(b), first assertion. *) +Corollary der_mmax_compl_abelian M E : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> abelian E^`(1). +Proof. +move=> maxM hallE; have [sEM s'E _] := and3P hallE. +have sE'E := der_sub 1 E; have sE'M := subset_trans sE'E sEM. +exact: sigma'_nil_abelian (pgroupS _ s'E) (der1_sigma_compl_nil maxM hallE). +Qed. + +(* This is B & G, Corollary 12.10(b), second assertion. *) +(* Note that we do not require the full decomposition of the complement. *) +Corollary tau2_compl_abelian M E E2 : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> \tau2(M).-Hall(E) E2 -> abelian E2. +Proof. +move: E2 => F2 maxM hallE hallF2; have [sEM s'E _] := and3P hallE. +have [[E1 hallE1] [E3 hallE3]] := ex_tau13_compl hallE. +have [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. +have solE: solvable E := sigma_compl_sol hallE. +have{solE hallF2} [x _ ->{F2}] := Hall_trans solE hallF2 hallE2. +have [-> | ntE] := eqsVneq E2 1; rewrite {x}abelianJ ?abelian1 //. +have [[p p_pr rpE2] [sE2E t2E2 _]] := (rank_witness E2, and3P hallE2). +have piE2p: p \in \pi(E2) by rewrite -p_rank_gt0 -rpE2 rank_gt0. +have t2p := pnatPpi t2E2 piE2p; have [A Ep2A _] := ex_tau2Elem hallE t2p. +have [_ abelA _] := pnElemP Ep2A; have [pA _] := andP abelA. +have [S /= sylS sAS] := Sylow_superset (subsetT A) pA. +have [cSS | not_cSS] := boolP (abelian S). + by have [[_ cE2E2] _ _ _] := abelian_tau2 maxM complEi t2p Ep2A sylS sAS cSS. +have pS := pHall_pgroup sylS. +have [def_t2 _ _ _] := nonabelian_tau2 maxM hallE t2p Ep2A pS not_cSS. +apply: sigma'_nil_abelian (subset_trans _ sEM) (pgroupS _ s'E) _ => //. +by rewrite (eq_pgroup _ def_t2) in t2E2; exact: pgroup_nil t2E2. +Qed. + +(* This is B & G, Corollary 12.10(c). *) +(* We do not really need a full decomposition of the complement in the first *) +(* part, but this reduces the number of assumptions. *) +Corollary tau1_cent_tau2Elem_factor M E p A : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> p \in \tau2(M) -> A \in 'E_p^2(E) -> + [/\ forall E1 E2 E3, sigma_complement M E E1 E2 E3 -> E3 * E2 \subset 'C_E(A), + 'C_E(A) <| E + & \tau1(M).-group (E / 'C_E(A))]. +Proof. +move=> maxM hallE t2p Ep2A; have sEM: E \subset M := pHall_sub hallE. +have nsAE: A <| E by case/(tau2_compl_context maxM): Ep2A => //; case. +have [sAE nAE]: A \subset E /\ E \subset 'N(A) := andP nsAE. +have nsCAE: 'C_E(A) <| E by rewrite norm_normalI ?norms_cent. +have [[E1 hallE1] [E3 hallE3]] := ex_tau13_compl hallE. +have [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. +have{hallE1} [t1E1 sE3E] := (pHall_pgroup hallE1, pHall_sub hallE3). +have{nsAE} sAE2: A \subset E2. + apply: subset_trans (pcore_max _ nsAE) (pcore_sub_Hall hallE2). + by apply: pi_pnat t2p; case/pnElemP: Ep2A => _; case/andP. +have{complEi} defE: (E3 ><| E2) ><| E1 = E. + by case/sigma_compl_context: complEi => // _ _ _ [_ ->]. +have [[K _ defK _] _ _ _] := sdprodP defE; rewrite defK in defE. +have nsKE: K <| E by case/sdprod_context: defE. +have [[sKE nKE] [_ mulE32 nE32 tiE32]] := (andP nsKE, sdprodP defK). +have{sE3E} sK_CEA: K \subset 'C_E(A). + have cE2E2: abelian E2 := tau2_compl_abelian maxM hallE hallE2. + rewrite subsetI sKE -mulE32 mulG_subG (centsS sAE2 cE2E2) (sameP commG1P eqP). + rewrite -subG1 -tiE32 subsetI commg_subl (subset_trans sAE2) //=. + by rewrite (subset_trans _ sAE2) // commg_subr (subset_trans sE3E). +split=> // [_ F2 F3 [_ _ hallF2 hallF3 _] | ]. + have solE: solvable E := solvableS sEM (mmax_sol maxM). + have [x2 Ex2 ->] := Hall_trans solE hallF2 hallE2. + have [x3 Ex3 ->] := Hall_trans solE hallF3 hallE3. + rewrite mulG_subG !sub_conjg !(normsP (normal_norm nsCAE)) ?groupV //. + by rewrite -mulG_subG mulE32. +have [f <-] := homgP (homg_quotientS nKE (normal_norm nsCAE) sK_CEA). +by rewrite morphim_pgroup // /pgroup -divg_normal // -(sdprod_card defE) mulKn. +Qed. + +(* This is B & G, Corollary 12.10(d). *) +Corollary norm_noncyclic_sigma M p P : + M \in 'M -> p \in \sigma(M) -> p.-group P -> P \subset M -> ~~ cyclic P -> + 'N(P) \subset M. +Proof. +move=> maxM sMp pP sPM ncycP. +have [A Ep2A]: exists A, A \in 'E_p^2(P). + by apply/p_rank_geP; rewrite ltnNge -odd_pgroup_rank1_cyclic ?mFT_odd. +have [[sAP _ _] Ep2A_M] := (pnElemP Ep2A, subsetP (pnElemS p 2 sPM) A Ep2A). +have sCAM: 'C(A) \subset M by case/p2Elem_mmax: Ep2A_M. +have [_ _ <- //] := sigma_group_trans maxM sMp pP. +by rewrite mulG_subG subsetIl (subset_trans (centS sAP)). +Qed. + +(* This is B & G, Corollary 12.10(e). *) +Corollary cent1_nreg_sigma_uniq M (Ms := M`_\sigma) x : + M \in 'M -> x \in M^# -> \tau2(M).-elt x -> 'C_Ms[x] != 1 -> + 'M('C[x]) = [set M]. +Proof. +move=> maxM /setD1P[ntx]; rewrite -cycle_subG => sMX t2x. +apply: contraNeq => MCx_neqM. +have{MCx_neqM} [H maxCxH neqHM]: exists2 H, H \in 'M('C[x]) & H \notin [set M]. + apply/subsetPn; have [H MCxH] := mmax_exists (mFT_cent1_proper ntx). + by rewrite eqEcard cards1 (cardD1 H) MCxH andbT in MCx_neqM. +have sCxH: 'C[x] \subset H by case/setIdP: maxCxH. +have s'x: \sigma(M)^'.-elt x by apply: sub_pgroup t2x => p; case/andP. +have [E hallE sXE] := Hall_superset (mmax_sol maxM) sMX s'x. +have [sEM solE] := (pHall_sub hallE, sigma_compl_sol hallE). +have{sXE} [E2 hallE2 sXE2] := Hall_superset solE sXE t2x. +pose p := pdiv #[x]. +have t2p: p \in \tau2(M) by rewrite (pnatPpi t2x) ?pi_pdiv ?order_gt1. +have [A Ep2A sAE2]: exists2 A, A \in 'E_p^2(M) & A \subset E2. + have [A Ep2A_E EpA] := ex_tau2Elem hallE t2p. + have [sAE abelA _] := pnElemP Ep2A_E; have [pA _] := andP abelA. + have [z Ez sAzE2] := Hall_Jsub solE hallE2 sAE (pi_pnat pA t2p). + by exists (A :^ z)%G; rewrite // -(normsP (normsG sEM) z Ez) pnElemJ. +have cE2E2: abelian E2 := tau2_compl_abelian maxM hallE hallE2. +have cxA: A \subset 'C[x] by rewrite -cent_cycle (sub_abelian_cent2 cE2E2). +have maxAH: H \in 'M(A) :\ M by rewrite inE neqHM (subsetP (mmax_ofS cxA)). +have [_ _ _ tiMsMA _] := tau2_context maxM t2p Ep2A. +by rewrite -subG1 -(tiMsMA H maxAH) setIS. +Qed. + +(* This is B & G, Lemma 12.11. *) +Lemma primes_norm_tau2Elem M E p A Mstar : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> p \in \tau2(M) -> A \in 'E_p^2(E) -> + Mstar \in 'M('N(A)) -> + [/\ (*a*) {subset \tau2(M) <= [predD \sigma(Mstar) & \beta(Mstar)]}, + (*b*) [predU \tau1(Mstar) & \tau2(Mstar)].-group (E / 'C_E(A)) + & (*c*) forall q, q \in \pi(E / 'C_E(A)) -> q \in \pi('C_E(A)) -> + [/\ q \in \tau2(Mstar), + exists2 P, P \in 'Syl_p(G) & P <| Mstar + & exists Q, [/\ Q \in 'Syl_q(G), Q \subset Mstar & abelian Q]]]. +Proof. +move=> maxM hallE t2Mp Ep2A; move: Mstar. +have [sAE abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have ntA: A :!=: 1 by exact: (nt_pnElem Ep2A). +have [sEM solE] := (pHall_sub hallE, sigma_compl_sol hallE). +have [_ nsCA_E t1CEAb] := tau1_cent_tau2Elem_factor maxM hallE t2Mp Ep2A. +have [sAM nCA_E] := (subset_trans sAE sEM, normal_norm nsCA_E). +have part_a H: + H \in 'M('N(A)) -> {subset \tau2(M) <= [predD \sigma(H) & \beta(H)]}. +- case/setIdP=> maxH sNA_H q t2Mq. + have sCA_H: 'C(A) \subset H := subset_trans (cent_sub A) sNA_H. + have sAH := subset_trans cAA sCA_H. + have sHp: p \in \sigma(H). + have [// | t2Hp] := orP (prime_class_mmax_norm maxH pA sNA_H). + apply: contraLR sNA_H => sH'p. + have sH'A: \sigma(H)^'.-group A by apply: pi_pnat pA _. + have [F hallF sAF] := Hall_superset (mmax_sol maxH) sAH sH'A. + have Ep2A_F: A \in 'E_p^2(F) by apply/pnElemP. + by have [_ [_ _ ->]]:= tau2_compl_context maxH hallF t2Hp Ep2A_F. + rewrite inE /= -predI_sigma_beta //= negb_and /= orbC. + have [-> /= _] := tau2_not_beta maxM t2Mq. + have [B Eq2B]: exists B, B \in 'E_q^2('C(A)). + have [E2 hallE2 sAE2] := Hall_superset solE sAE (pi_pnat pA t2Mp). + have cE2E2: abelian E2 := tau2_compl_abelian maxM hallE hallE2. + have [Q sylQ] := Sylow_exists q E2; have sQE2 := pHall_sub sylQ. + have sylQ_E := subHall_Sylow hallE2 t2Mq sylQ. + apply/p_rank_geP; apply: leq_trans (p_rankS q (centsS sAE2 cE2E2)). + rewrite -(p_rank_Sylow sylQ) (p_rank_Sylow sylQ_E). + by move: t2Mq; rewrite (tau2E hallE) => /andP[_ /eqP->]. + have [cAB abelB dimB] := pnElemP Eq2B; have sBH := subset_trans cAB sCA_H. + have{Eq2B} Eq2B: B \in 'E_q^2(H) by apply/pnElemP. + have rqHgt1: 'r_q(H) > 1 by apply/p_rank_geP; exists B. + have piHq: q \in \pi(H) by rewrite -p_rank_gt0 ltnW. + rewrite partition_pi_mmax // in piHq. + case/or4P: piHq rqHgt1 => // [|t2Hq _|]; try by case/and3P=> _ /eqP->. + have [_ _ regB _ _] := tau2_context maxH t2Hq Eq2B. + case/negP: ntA; rewrite -subG1 -regB subsetI centsC cAB andbT. + by rewrite (sub_Hall_pcore (Msigma_Hall maxH)) // (pi_pgroup pA). +have part_b H: + H \in 'M('N(A)) -> [predU \tau1(H) & \tau2(H)].-group (E / 'C_E(A)). +- move=> maxNA_H; have [maxH sNA_H] := setIdP maxNA_H. + have [/= bH'p sHp] := andP (part_a H maxNA_H p t2Mp). + have sCA_H: 'C(A) \subset H := subset_trans (cent_sub A) sNA_H. + have sAH := subset_trans cAA sCA_H. + apply/pgroupP=> q q_pr q_dv_CEAb; apply: contraR bH'p => t12H'q. + have [Q sylQ] := Sylow_exists q E; have [sQE qQ _] := and3P sylQ. + have nsAE: A <| E by case/(tau2_compl_context maxM): Ep2A => //; case. + have nAE := normal_norm nsAE; have nAQ := subset_trans sQE nAE. + have sAQ_A: [~: A, Q] \subset A by rewrite commg_subl. + have ntAQ: [~: A, Q] != 1. + apply: contraTneq q_dv_CEAb => /commG1P cAQ. + have nCEA_Q := subset_trans sQE nCA_E. + rewrite -p'natE // -partn_eq1 ?cardg_gt0 //. + rewrite -(card_Hall (quotient_pHall nCEA_Q sylQ)) -trivg_card1 -subG1. + by rewrite quotient_sub1 // subsetI sQE centsC. + have sQH: Q \subset H := subset_trans nAQ sNA_H. + have sHsubH' r X: + r \in \sigma(H) -> X \subset H -> r.-group X -> X \subset H^`(1). + + move=> sHr sXH rX; apply: subset_trans (Msigma_der1 maxH). + by rewrite (sub_Hall_pcore (Msigma_Hall maxH)) // (pi_pgroup rX sHr). + have sAH': A \subset H^`(1) by apply: sHsubH' pA. + have{t12H'q} sQH': Q \subset H^`(1). + have [sHq | sH'q] := boolP (q \in \sigma(H)); first exact: sHsubH' qQ. + have{sH'q} sH'Q: \sigma(H)^'.-group Q by apply: (pi_pnat qQ). + have{sH'Q} [F hallF sQF] := Hall_superset (mmax_sol maxH) sQH sH'Q. + have [-> | ntQ] := eqsVneq Q 1; first exact: sub1G. + have{t12H'q} t3Hq: q \in \tau3(H). + apply: implyP t12H'q; rewrite implyNb -orbA. + rewrite -(partition_pi_sigma_compl maxH hallF) -p_rank_gt0. + by rewrite (leq_trans _ (p_rankS q sQF)) // -rank_pgroup // rank_gt0. + have solF: solvable F := sigma_compl_sol hallF. + have [F3 hallF3 sQF3] := Hall_superset solF sQF (pi_pnat qQ t3Hq). + have [[F1 hallF1] _] := ex_tau13_compl hallF. + have [F2 _ complFi] := ex_tau2_compl hallF hallF1 hallF3. + have [[sF3H' _] _ _ _ _] := sigma_compl_context maxH complFi. + exact: subset_trans sQF3 (subset_trans sF3H' (dergS 1 (pHall_sub hallF))). + have hallHb: \beta(H).-Hall(H) H`_\beta := Mbeta_Hall maxH. + have nilH'b: nilpotent (H^`(1) / H`_\beta) := Mbeta_quo_nil maxH. + have{nilH'b} sAQ_Hb: [~: A, Q] \subset H`_\beta. + rewrite -quotient_cents2 ?(subset_trans _ (gFnorm _ _)) // centsC. + rewrite (sub_nilpotent_cent2 nilH'b) ?quotientS ?coprime_morph //. + rewrite (pnat_coprime (pi_pnat pA t2Mp) (pi_pnat qQ _)) ?tau2'1 //. + by rewrite (pnatPpi t1CEAb) // mem_primes q_pr cardG_gt0. + rewrite (pnatPpi (pHall_pgroup hallHb)) // (piSg sAQ_Hb) // -p_rank_gt0. + by rewrite -(rank_pgroup (pgroupS sAQ_A pA)) rank_gt0. +move=> H maxNA_H; split; last 1 [move=> q piCEAb_q piCEAq] || by auto. +have [_ sHp]: _ /\ p \in \sigma(H) := andP (part_a H maxNA_H p t2Mp). +have{maxNA_H} [maxH sNA_H] := setIdP maxNA_H. +have{sNA_H} sCA_H: 'C(A) \subset H := subset_trans (cent_sub A) sNA_H. +have{piCEAq} [Q0 EqQ0]: exists Q0, Q0 \in 'E_q^1('C_E(A)). + by apply/p_rank_geP; rewrite p_rank_gt0. +have [sQ0_CEA abelQ0 dimQ0]:= pnElemP EqQ0; have [qQ0 cQ0Q0 _] := and3P abelQ0. +have{sQ0_CEA} [sQ0E cAQ0]: Q0 \subset E /\ Q0 \subset 'C(A). + by apply/andP; rewrite -subsetI. +have ntQ0: Q0 :!=: 1 by apply: (nt_pnElem EqQ0). +have{t1CEAb} t1Mq: q \in \tau1(M) := pnatPpi t1CEAb piCEAb_q. +have [Q sylQ sQ0Q] := Sylow_superset sQ0E qQ0; have [sQE qQ _] := and3P sylQ. +have [E1 hallE1 sQE1] := Hall_superset solE sQE (pi_pnat qQ t1Mq). +have rqE: 'r_q(E) = 1%N. + by move: t1Mq; rewrite (tau1E maxM hallE) andbA andbC; case: eqP. +have cycQ: cyclic Q. + by rewrite (odd_pgroup_rank1_cyclic qQ) ?mFT_odd // (p_rank_Sylow sylQ) rqE. +have sCAE: 'C(A) \subset E by case/(tau2_compl_context maxM): Ep2A => // _ []. +have{sCAE} sylCQA: q.-Sylow('C(A)) 'C_Q(A). + by apply: Hall_setI_normal sylQ; rewrite /= -(setIidPr sCAE). +have{sylCQA} defNA: 'C(A) * 'N_('N(A))(Q0) = 'N(A). + apply/eqP; rewrite eqEsubset mulG_subG cent_sub subsetIl /=. + rewrite -{1}(Frattini_arg (cent_normal A) sylCQA) mulgS ?setIS ?char_norms //. + by rewrite (sub_cyclic_char Q0 (cyclicS (subsetIl Q _) cycQ)) subsetI sQ0Q. +have [L maxNQ0_L]: {L | L \in 'M('N(Q0))}. + by apply: mmax_exists; rewrite mFT_norm_proper ?(mFT_pgroup_proper qQ0). +have{maxNQ0_L} [maxL sNQ0_L] := setIdP maxNQ0_L. +have sCQ0_L: 'C(Q0) \subset L := subset_trans (cent_sub Q0) sNQ0_L. +have sAL: A \subset L by rewrite (subset_trans _ sCQ0_L) // centsC. +have sCA_L: 'C(A) \subset L. + by have /p2Elem_mmax[]: A \in 'E_p^2(L) by apply/pnElemP. +have{sCA_L defNA} maxNA_L: L \in 'M('N(A)). + by rewrite inE maxL -defNA setIC mul_subG // subIset ?sNQ0_L. +have t2Lq: q \in \tau2(L). + have /orP[sLq | //] := prime_class_mmax_norm maxL qQ0 sNQ0_L. + by have /orP[/andP[/negP] | ] := pnatPpi (part_b L maxNA_L) piCEAb_q. +have [cQQ [/= sL'q _]] := (cyclic_abelian cycQ, andP t2Lq). +have sQL: Q \subset L := subset_trans (centsS sQ0Q cQQ) sCQ0_L. +have [F hallF sQF] := Hall_superset (mmax_sol maxL) sQL (pi_pnat qQ sL'q). +have [B Eq2B _] := ex_tau2Elem hallF t2Lq. +have [_ sLp]: _ /\ p \in \sigma(L) := andP (part_a L maxNA_L p t2Mp). +have{H sHp maxH sCA_H} <-: L :=: H. + have sLHp: p \in [predI \sigma(L) & \sigma(H)] by apply/andP. + have [_ transCA _] := sigma_group_trans maxH sHp pA. + set S := finset _ in transCA; have sAH := subset_trans cAA sCA_H. + suffices [SH SL]: gval H \in S /\ gval L \in S. + have [c cAc -> /=]:= atransP2 transCA SH SL. + by rewrite conjGid // (subsetP sCA_H). + have [_ _ _ TIsL] := tau2_compl_context maxL hallF t2Lq Eq2B. + apply/andP; rewrite !inE sAH sAL orbit_refl orbit_sym /= andbT. + by apply: contraLR sLHp => /TIsL[] // _ ->. +split=> //. + exists ('O_p(L`_\sigma))%G; last by rewrite /= -pcoreI pcore_normal. + rewrite inE (subHall_Sylow (Msigma_Hall_G maxL) sLp) //. + by rewrite nilpotent_pcore_Hall // (tau2_Msigma_nil maxL t2Lq). +have [Q1 sylQ1 sQQ1] := Sylow_superset (subsetT Q) qQ. +have [sQ0Q1 qQ1] := (subset_trans sQ0Q sQQ1, pHall_pgroup sylQ1). +have [cQ1Q1 | not_cQ1Q1] := boolP (abelian Q1). + by exists Q1; rewrite inE (subset_trans (centsS sQ0Q1 cQ1Q1)). +have [_ _ regB [F0 /=]] := nonabelian_tau2 maxL hallF t2Lq Eq2B qQ1 not_cQ1Q1. +have{regB} ->: 'C_B(L`_\sigma) = Q0; last move=> defF _. + apply: contraTeq sCQ0_L => neqQ0B; case: (regB Q0) => //. + by rewrite 2!inE eq_sym neqQ0B; apply/pnElemP; rewrite (subset_trans sQ0Q). +have{defF} defQ: Q0 \x (F0 :&: Q) = Q. + rewrite dprodEsd ?(centSS (subsetIr F0 Q) sQ0Q) //. + by rewrite (sdprod_modl defF sQ0Q) (setIidPr sQF). +have [[/eqP/idPn//] | [_ eqQ0Q]] := cyclic_pgroup_dprod_trivg qQ cycQ defQ. +have nCEA_Q := subset_trans sQE nCA_E. +case/idPn: piCEAb_q; rewrite -p'natEpi -?partn_eq1 ?cardG_gt0 //. +rewrite -(card_Hall (quotient_pHall nCEA_Q sylQ)) -trivg_card1 -subG1. +by rewrite quotient_sub1 // subsetI sQE -eqQ0Q. +Qed. + +(* This is a generalization of B & G, Theorem 12.12. *) +(* In the B & G text, Theorem 12.12 only establishes the type F structure *) +(* for groups of type I, whereas it is required for the derived groups of *) +(* groups of type II (in the sense of Peterfalvi). Indeed this is exactly *) +(* what is stated in Lemma 15.15(e) and then Theorem B(3). The proof of *) +(* 15.15(c) cites 12.12 in the type I case (K = 1) and then loosely invokes *) +(* a "short and easy argument" inside the proof of 12.12 for the K != 1 case. *) +(* In fact, this involves copying roughly 25% of the proof, whereas the proof *) +(* remains essentially unchanged when Theorem 12.12 is generalized to a *) +(* normal Hall subgroup of E as below. *) +(* Also, we simplify the argument for central tau2 Sylow subgroup S of U by *) +(* by replacing the considerations on the abelian structure of S by a *) +(* comparison of Mho^n-1(S) and Ohm_1(S) (with exponent S = p ^ n), as the *) +(* former is needed anyway to prove regularity when S is not homocyclic. *) +Theorem FTtypeF_complement M (Ms := M`_\sigma) E U : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> Hall E U -> U <| E -> U :!=: 1 -> + {in U^#, forall e, [predU \tau1(M) & \tau3(M)].-elt e -> 'C_Ms[e] = 1} -> + [/\ (*a*) exists A0 : {group gT}, + [/\ A0 <| U, abelian A0 & {in Ms^#, forall x, 'C_U[x] \subset A0}] + & (*b*) exists E0 : {group gT}, + [/\ E0 \subset U, exponent E0 = exponent U + & [Frobenius Ms <*> E0 = Ms ><| E0]]]. +Proof. +set p13 := predU _ _ => maxM hallE /Hall_pi hallU nsUE ntU regU13. +have [[E1 hallE1] [E3 hallE3]] := ex_tau13_compl hallE. +have [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. +have [[sE1E _] [sE2E t2E2 _]] := (andP hallE1, and3P hallE2). +have [[_ nsE3E] _ [cycE1 _] [defE _] _] := sigma_compl_context maxM complEi. +have [[[sE3E t3E3 _][_ nE3E]] [sUE _]] := (and3P hallE3, andP nsE3E, andP nsUE). +have defM: Ms ><| E = M := sdprod_sigma maxM hallE. +have [nsMsM sEM mulMsE nMsE tiMsE] := sdprod_context defM. +have ntMs: Ms != 1 := Msigma_neq1 maxM. +have{defM} defMsU: Ms ><| U = Ms <*> U := sdprod_subr defM sUE. +pose U2 := (E2 :&: U)%G. +have hallU2: \tau2(M).-Hall(U) U2 := Hall_setI_normal nsUE hallE2. +have [U2_1 | ntU2] := eqsVneq U2 1. + have frobMsU: [Frobenius Ms <*> U = Ms ><| U]. + apply/Frobenius_semiregularP=> // e Ue. + apply: regU13 => //; case/setD1P: Ue => _; apply: mem_p_elt. + have: \tau2(M)^'.-group U. + by rewrite -partG_eq1 -(card_Hall hallU2) U2_1 cards1. + apply: sub_in_pnat => p /(piSg sUE). + by rewrite (partition_pi_sigma_compl maxM hallE) orbCA => /orP[] // /idPn. + split; [exists 1%G; rewrite normal1 abelian1 | by exists U]. + by split=> //= x Ux; rewrite (Frobenius_reg_compl frobMsU). +have [[sU2U t2U2 _] [p p_pr rU2]] := (and3P hallU2, rank_witness U2). +have piU2p: p \in \pi(U2) by rewrite -p_rank_gt0 -rU2 rank_gt0. +have t2p: p \in \tau2(M) := pnatPpi t2U2 piU2p. +have [A Ep2A Ep2A_M] := ex_tau2Elem hallE t2p. +have [sAE abelA _] := pnElemP Ep2A; have{abelA} [pA cAA _] := and3P abelA. +have [S sylS sAS] := Sylow_superset (subsetT A) pA. +have [cSS | not_cSS] := boolP (abelian S); last first. + have [_] := nonabelian_tau2 maxM hallE t2p Ep2A (pHall_pgroup sylS) not_cSS. + set A0 := ('C_A(_))%G => [] [oA0 _] _ {defE} [E0 defE regE0]. + have [nsA0E sE0E mulAE0 nAE0 tiAE0] := sdprod_context defE. + have [P sylP] := Sylow_exists p U; have [sPU _] := andP sylP. + have sylP_E := subHall_Sylow hallU (piSg sU2U piU2p) sylP. + have pA0: p.-group A0 by rewrite /pgroup oA0 pnat_id. + have sA0P: A0 \subset P. + by apply: subset_trans (pcore_sub_Hall sylP_E); apply: pcore_max. + have sA0U: A0 \subset U := subset_trans sA0P sPU. + pose U0 := (E0 :&: U)%G. + have defU: A0 ><| U0 = U by rewrite (sdprod_modl defE) // (setIidPr sUE). + have piU0p: p \in \pi(U0). + have:= lognSg p sAE; rewrite (card_pnElem Ep2A) pfactorK //. + rewrite -logn_part -(card_Hall sylP_E) (card_Hall sylP) logn_part. + rewrite -(sdprod_card defU) oA0 lognM // ?prime_gt0 // logn_prime // eqxx. + by rewrite ltnS logn_gt0. + have defM0: Ms ><| U0 = Ms <*> U0 := sdprod_subr defMsU (subsetIr _ _). + have frobM0: [Frobenius Ms <*> U0 = Ms ><| U0]. + apply/Frobenius_semiregularP=> // [|e /setD1P[nte /setIP[E0e Ue]]]. + by rewrite -rank_gt0 (leq_trans _ (p_rank_le_rank p _)) ?p_rank_gt0. + have [ | ] := boolP (p13.-elt e); first by apply: regU13; rewrite !inE nte. + apply: contraNeq => /trivgPn[x /setIP[Ms_x cex] ntx]. + apply/pgroupP=> q q_pr q_dv_x ; rewrite inE /= (regE0 x) ?inE ?ntx //. + rewrite mem_primes q_pr cardG_gt0 (dvdn_trans q_dv_x) ?order_dvdG //. + by rewrite inE E0e cent1C. + have [nsA0U sU0U _ _ _] := sdprod_context defU. + split; [exists A0 | exists U0]. + split=> // [|x Ms_x]; first by rewrite (abelianS (subsetIl A _) cAA). + rewrite -(sdprodW defU) -group_modl ?(Frobenius_reg_compl frobM0) ?mulg1 //. + by rewrite subIset //= orbC -cent_set1 centS // sub1set; case/setD1P: Ms_x. + split=> //; apply/eqP; rewrite eqn_dvd exponentS //=. + rewrite -(partnC p (exponent_gt0 U0)) -(partnC p (exponent_gt0 U)). + apply: dvdn_mul; last first. + rewrite (partn_exponentS sU0U) // -(sdprod_card defU) partnM ?cardG_gt0 //. + by rewrite part_p'nat ?pnatNK // mul1n dvdn_part. + have cPP: abelian P. + have [/(_ P)[] //] := tau2_context maxM t2p Ep2A_M. + by apply: subHall_Sylow hallE _ sylP_E; case/andP: t2p. + have defP: A0 \x (U0 :&: P) = P. + rewrite dprodEsd ?(sub_abelian_cent2 cPP) ?subsetIr //. + by rewrite (sdprod_modl defU) // (setIidPr sPU). + have sylP_U0: p.-Sylow(U0) (U0 :&: P). + rewrite pHallE subsetIl /= -(eqn_pmul2l (cardG_gt0 A0)). + rewrite (dprod_card defP) (card_Hall sylP) -(sdprod_card defU). + by rewrite partnM // part_pnat_id. + rewrite -(exponent_Hall sylP) -(dprod_exponent defP) (exponent_Hall sylP_U0). + rewrite dvdn_lcm (dvdn_trans (exponent_dvdn A0)) //= oA0. + apply: contraLR piU0p; rewrite -p'natE // -partn_eq1 // partn_part //. + by rewrite partn_eq1 ?exponent_gt0 // pnat_exponent -p'groupEpi. +have{t2p Ep2A sylS sAS cSS} [[nsE2E cE2E2] hallE2_G _ _] + := abelian_tau2 maxM complEi t2p Ep2A sylS sAS cSS. +clear p p_pr rU2 piU2p A S Ep2A_M sAE pA cAA. +have nsU2U: U2 <| U by rewrite (normalS sU2U sUE) ?normalI. +have cU2U2: abelian U2 := abelianS (subsetIl _ _) cE2E2. +split. + exists U2; rewrite -set1gE; split=> // x /setDP[Ms_x ntx]. + rewrite (sub_normal_Hall hallU2) ?subsetIl //=. + apply: sub_in_pnat (pgroup_pi _) => q /(piSg (subsetIl U _))/(piSg sUE). + rewrite (partition_pi_sigma_compl maxM) // orbCA => /orP[] // t13q. + rewrite mem_primes => /and3P[q_pr _ /Cauchy[] // y /setIP[Uy cxy] oy]. + case/negP: ntx; rewrite -(regU13 y); first by rewrite inE Ms_x cent1C. + by rewrite !inE -order_gt1 oy prime_gt1. + by rewrite /p_elt oy pnatE. +pose sylU2 S := (S :!=: 1) && Sylow U2 S. +pose cyclicRegular Z S := + [/\ Z <| U, cyclic Z, 'C_Ms('Ohm_1(Z)) = 1 & exponent Z = exponent S]. +suffices /fin_all_exists[Z_ Z_P] S: exists Z, sylU2 S -> cyclicRegular Z S. + pose Z2 := <<\bigcup_(S | sylU2 S) Z_ S>>. + have sZU2: Z2 \subset U2. + rewrite gen_subG; apply/bigcupsP=> S sylS. + have [[/andP[sZE _] _ _ eq_expZS] [_ _ sSU2 _]] := (Z_P S sylS, and4P sylS). + rewrite (sub_normal_Hall hallU2) // -pnat_exponent eq_expZS. + by rewrite pnat_exponent (pgroupS sSU2 t2U2). + have nZ2U: U \subset 'N(Z2). + by rewrite norms_gen ?norms_bigcup //; apply/bigcapsP => S /Z_P[/andP[]]. + have solU: solvable U := solvableS sUE (sigma_compl_sol hallE). + have [U31 hallU31] := Hall_exists \tau2(M)^' solU. + have [[sU31U t2'U31 _] t2Z2] := (and3P hallU31, pgroupS sZU2 t2U2). + pose U0 := (Z2 <*> U31)%G; have /joing_sub[sZ2U0 sU310] := erefl (gval U0). + have sU0U: U0 \subset U by rewrite join_subG (subset_trans sZU2). + have nsZ2U0: Z2 <| U0 by rewrite /normal sZ2U0 (subset_trans sU0U). + have defU0: Z2 * U31 = U0 by rewrite -norm_joinEr // (subset_trans sU31U). + have [hallZ2 hallU31_0] := coprime_mulpG_Hall defU0 t2Z2 t2'U31. + have expU0U: exponent U0 = exponent U. + have exp_t2c U' := partnC \tau2(M) (exponent_gt0 U'). + rewrite -(exp_t2c U) -(exponent_Hall hallU31) -(exponent_Hall hallU2). + rewrite -{}exp_t2c -(exponent_Hall hallU31_0) -(exponent_Hall hallZ2). + congr (_ * _)%N; apply/eqP; rewrite eqn_dvd exponentS //=. + apply/dvdn_partP=> [|p]; first exact: exponent_gt0. + have [S sylS] := Sylow_exists p U2; rewrite -(exponent_Hall sylS). + rewrite pi_of_exponent -p_rank_gt0 -(rank_Sylow sylS) rank_gt0 => ntS. + have{sylS} sylS: sylU2 S by rewrite /sylU2 ntS (p_Sylow sylS). + by have /Z_P[_ _ _ <-] := sylS; rewrite exponentS ?sub_gen ?(bigcup_max S). + exists U0; split=> //. + have ntU0: U0 :!=: 1 by rewrite trivg_exponent expU0U -trivg_exponent. + apply/Frobenius_semiregularP=> //; first by rewrite (sdprod_subr defMsU). + apply: semiregular_sym => x /setD1P[ntx Ms_x]; apply: contraNeq ntx. + rewrite -rank_gt0; have [p p_pr ->] := rank_witness [group of 'C_U0[x]]. + rewrite -in_set1 -set1gE p_rank_gt0 => piCp. + have [e /setIP[U0e cxe] oe]: {e | e \in 'C_U0[x] & #[e] = p}. + by move: piCp; rewrite mem_primes p_pr cardG_gt0; apply: Cauchy. + have nte: e != 1 by rewrite -order_gt1 oe prime_gt1. + have{piCp} piUp: p \in \pi(U). + by rewrite -pi_of_exponent -expU0U pi_of_exponent (piSg _ piCp) ?subsetIl. + have:= piSg sUE piUp; rewrite (partition_pi_sigma_compl maxM) // orbCA orbC. + case/orP=> [t13p | t2p]. + rewrite -(regU13 e) 1?inE ?Ms_x 1?cent1C //. + by rewrite inE nte (subsetP sU0U). + by rewrite /p_elt oe pnatE. + have Z2e: e \in Z2 by rewrite (mem_normal_Hall hallZ2) // /p_elt oe pnatE. + have [S sylS] := Sylow_exists p U2; have [sSU2 pS _] := and3P sylS. + have sylS_U: p.-Sylow(U) S := subHall_Sylow hallU2 t2p sylS. + have ntS: S :!=: 1 by rewrite -rank_gt0 (rank_Sylow sylS_U) p_rank_gt0. + have sylS_U2: sylU2 S by rewrite /sylU2 ntS (p_Sylow sylS). + have [nsZU cycZ regZ1 eqexpZS] := Z_P S sylS_U2. + suffices defZ1: 'Ohm_1(Z_ S) == <[e]>. + by rewrite -regZ1 (eqP defZ1) cent_cycle inE Ms_x cent1C. + have pZ: p.-group (Z_ S) by rewrite -pnat_exponent eqexpZS pnat_exponent. + have sylZ: p.-Sylow(Z2) (Z_ S). + have:= sZU2; rewrite pHallE /Z2 /= -bigprodGE (bigD1 S) //=. + set Z2' := (\prod_(T | _) _)%G => /joing_subP[sZ_U2 sZ2'_U2]. + rewrite joing_subl cent_joinEl ?(sub_abelian_cent2 cU2U2) //=. + suffices p'Z2': p^'.-group Z2'. + rewrite coprime_cardMg ?(pnat_coprime pZ) //. + by rewrite partnM // part_pnat_id // part_p'nat // muln1. + elim/big_ind: Z2' sZ2'_U2 => [_||T /andP[sylT neqTS]]; first exact: pgroup1. + move=> X Y IHX IHY /joing_subP[sXU2 sYU2] /=. + by rewrite cent_joinEl ?(sub_abelian_cent2 cU2U2) // pgroupM IHX ?IHY. + have /Z_P[_ _ _ expYT _] := sylT. + have{sylT} [_ /SylowP[q _ sylT]] := andP sylT. + rewrite -pnat_exponent expYT pnat_exponent. + apply: (pi_pnat (pHall_pgroup sylT)); apply: contraNneq neqTS => eq_qp. + have defOE2 := nilpotent_Hall_pcore (abelian_nil cU2U2). + by rewrite -val_eqE /= (defOE2 _ _ sylS) (defOE2 _ _ sylT) eq_qp. + have nZZ2 := normalS (pHall_sub sylZ) (subset_trans sZU2 sU2U) nsZU. + have Ze: e \in Z_ S by rewrite (mem_normal_Hall sylZ) // /p_elt oe pnat_id. + rewrite (eq_subG_cyclic cycZ) ?Ohm_sub ?cycle_subG // -orderE oe. + by rewrite (Ohm1_cyclic_pgroup_prime _ pZ) //; apply/trivgPn; exists e. +case: (sylU2 S) / andP => [[ntS /SylowP[p p_pr sylS_U2]]|]; last by exists E. +have [sSU2 pS _] := and3P sylS_U2; have [sSE2 sSU] := subsetIP sSU2. +have piSp: p \in \pi(S) by rewrite -p_rank_gt0 -rank_pgroup ?rank_gt0. +have t2p: p \in \tau2(M) := pnatPpi t2U2 (piSg sSU2 piSp). +have sylS_U: p.-Sylow(U) S := subHall_Sylow hallU2 t2p sylS_U2. +have sylS_E: p.-Sylow(E) S := subHall_Sylow hallU (piSg sSU piSp) sylS_U. +have sylS: p.-Sylow(E2) S := pHall_subl sSE2 sE2E sylS_E. +have sylS_G: p.-Sylow(G) S := subHall_Sylow hallE2_G t2p sylS. +have [cSS [/= s'p rpM]] := (abelianS sSU2 cU2U2, andP t2p). +have sylS_M: p.-Sylow(M) S := subHall_Sylow hallE s'p sylS_E. +have rpS: 'r_p(S) = 2 by apply/eqP; rewrite (p_rank_Sylow sylS_M). +have [A] := p_rank_witness p S; rewrite rpS -(setIidPr (pHall_sub sylS_E)). +rewrite pnElemI setIC 2!inE => /andP[sAS Ep2A]. +have [[nsAE defEp] _ nregEp_uniq _] := tau2_compl_context maxM hallE t2p Ep2A. +have [_ sNS'FE _ sCSE _] + := abelian_tau2_sub_Fitting maxM hallE t2p Ep2A sylS_G sAS cSS. +have [_ _ [defNS _ _ _] regE1subZ] + := abelian_tau2 maxM complEi t2p Ep2A sylS_G sAS cSS. +have nSE: E \subset 'N(S) by rewrite -defNS normal_norm. +have [sSE nSU] := (subset_trans sSE2 sE2E, subset_trans sUE nSE). +have n_subNS := abelian_tau2_norm_Sylow maxM hallE t2p Ep2A sylS_G sAS cSS. +have not_sNS_M: ~~ ('N(S) \subset M). + by apply: contra s'p => sNS_M; apply/exists_inP; exists S; rewrite // inE. +have regNNS Z (Z1 := 'Ohm_1(Z)%G): + Z \subset S -> cyclic Z -> Z :!=: 1 -> 'N(S) \subset 'N(Z1) -> 'C_Ms(Z1) = 1. +- move=> sZS cycZ ntZ nZ1_NS; apply: contraNeq not_sNS_M => nregZ1. + have sZ1S: Z1 \subset S := subset_trans (Ohm_sub 1 Z) sZS. + have EpZ1: Z1 \in 'E_p^1(E). + rewrite p1ElemE // !inE (subset_trans sZ1S) //=. + by rewrite (Ohm1_cyclic_pgroup_prime _ (pgroupS sZS pS)). + have /= uCZ1 := nregEp_uniq _ EpZ1 nregZ1. + apply: (subset_trans nZ1_NS); apply: (sub_uniq_mmax uCZ1 (cent_sub _)). + by rewrite mFT_norm_proper ?(mFT_pgroup_proper (pgroupS sZ1S pS)) ?Ohm1_eq1. +have [_ nsCEA t1CEAb] := tau1_cent_tau2Elem_factor maxM hallE t2p Ep2A. +have [cSU | not_cSU] := boolP (U \subset 'C(S)). + pose n := logn p (exponent S); pose Sn := 'Mho^n.-1(S)%G. + have n_gt0: 0 < n by rewrite -pi_of_exponent -logn_gt0 in piSp. + have expS: (exponent S = p ^ n.-1 * p)%N. + rewrite -expnSr prednK -?p_part //. + by rewrite part_pnat_id ?pnat_exponent ?expg_exponent. + have sSnS1: Sn \subset 'Ohm_1(S). + rewrite (OhmE 1 pS) /= (MhoE _ pS); apply/genS/subsetP=> _ /imsetP[x Sx ->]. + by rewrite !inE groupX //= -expgM -expS expg_exponent. + have sSZ: S \subset 'Z(U) by rewrite subsetI sSU centsC. + have{sSZ} nsZU z: z \in S -> <[z]> <| U. + by move/(subsetP sSZ)=> ZUz; rewrite sub_center_normal ?cycle_subG. + have [homoS | ltSnS1] := eqVproper sSnS1. + have Ep2A_M := subsetP (pnElemS p 2 sEM) A Ep2A. + have [_ _ _ _ [A1 EpA1 regA1]] := tau2_context maxM t2p Ep2A_M. + have [sA1A _ oA1] := pnElemPcard EpA1. + have /cyclicP[zn defA1]: cyclic A1 by rewrite prime_cyclic ?oA1. + have [z Sz def_zn]: exists2 z, z \in S & zn = z ^+ (p ^ n.-1). + apply/imsetP; rewrite -(MhoEabelian _ pS cSS) homoS (OhmE 1 pS). + rewrite mem_gen // !inE -cycle_subG -defA1 (subset_trans sA1A) //=. + by rewrite -oA1 defA1 expg_order. + have oz: #[z] = exponent S. + by rewrite expS; apply: orderXpfactor; rewrite // -def_zn orderE -defA1. + exists <[z]>%G; split; rewrite ?cycle_cyclic ?exponent_cycle ?nsZU //. + by rewrite (Ohm_p_cycle _ (mem_p_elt pS Sz)) subn1 oz -def_zn -defA1. + have [z Sz /esym oz] := exponent_witness (abelian_nil cSS). + exists <[z]>%G; split; rewrite ?cycle_cyclic ?exponent_cycle ?nsZU //. + have ntz: <[z]> != 1 by rewrite trivg_card1 -orderE oz -dvdn1 -trivg_exponent. + rewrite regNNS ?cycle_cyclic ?cycle_subG //=. + suffices /eqP->: 'Ohm_1(<[z]>) == Sn by apply: char_norms; apply: gFchar. + have [p_z pS1] := (mem_p_elt pS Sz, pgroupS (Ohm_sub 1 S) pS). + rewrite eqEcard (Ohm1_cyclic_pgroup_prime _ p_z) ?cycle_cyclic //. + rewrite (Ohm_p_cycle _ p_z) oz -/n subn1 cycle_subG Mho_p_elt //=. + rewrite (card_pgroup (pgroupS sSnS1 pS1)) (leq_exp2l _ 1) ?prime_gt1 //. + by rewrite -ltnS -rpS p_rank_abelian ?properG_ltn_log. +have{not_cSU} [q q_pr piUq]: {q | prime q & q \in \pi(U / 'C(S))}. + have [q q_pr rCE] := rank_witness (U / 'C(S)); exists q => //. + by rewrite -p_rank_gt0 -rCE rank_gt0 -subG1 quotient_sub1 ?norms_cent. +have{piUq} [piCESbq piUq]: q \in \pi(E / 'C_E(S)) /\ q \in \pi(U). + rewrite /= setIC (card_isog (second_isog (norms_cent nSE))). + by rewrite (piSg _ piUq) ?quotientS // (pi_of_dvd _ _ piUq) ?dvdn_quotient. +have [Q1 sylQ1_U] := Sylow_exists q U; have [sQ1U qQ1 _] := and3P sylQ1_U. +have sylQ1: q.-Sylow(E) Q1 := subHall_Sylow hallU piUq sylQ1_U. +have sQ1E := subset_trans sQ1U sUE; have nSQ1 := subset_trans sQ1E nSE. +have [Q sylQ sQ1Q] := Sylow_superset nSQ1 qQ1; have [nSQ qQ _] := and3P sylQ. +have Ep2A_M := subsetP (pnElemS p 2 sEM) A Ep2A. +have ltCQ1_S: 'C_S(Q1) \proper S. + rewrite properE subsetIl /= subsetI subxx centsC -sQ1E -subsetI. + have nCES_Q1: Q1 \subset 'N('C_E(S)) by rewrite (setIidPr sCSE) norms_cent. + rewrite -quotient_sub1 // subG1 -rank_gt0. + by rewrite (rank_Sylow (quotient_pHall nCES_Q1 sylQ1)) p_rank_gt0. +have coSQ: coprime #|S| #|Q|. + suffices p'q: q != p by rewrite (pnat_coprime pS) // (pi_pnat qQ). + apply: contraNneq (proper_subn ltCQ1_S) => eq_qp; rewrite subsetI subxx. + rewrite (sub_abelian_cent2 cE2E2) // (sub_normal_Hall hallE2) //. + by rewrite (pi_pgroup qQ1) ?eq_qp. +have not_sQ1CEA: ~~ (Q1 \subset 'C_E(A)). + rewrite subsetI sQ1E; apply: contra (proper_subn ltCQ1_S) => /= cAQ1. + rewrite subsetIidl centsC coprime_abelian_faithful_Ohm1 ?(coprimegS sQ1Q) //. + by case: (tau2_context maxM t2p Ep2A_M); case/(_ S sylS_M)=> _ [|->] //. +have t1q: q \in \tau1(M). + rewrite (pnatPpi t1CEAb) // -p_rank_gt0. + have nCEA_Q1: Q1 \subset 'N('C_E(A)) := subset_trans sQ1E (normal_norm nsCEA). + rewrite -(rank_Sylow (quotient_pHall nCEA_Q1 sylQ1)) rank_gt0. + by rewrite -subG1 quotient_sub1. +have cycQ1: cyclic Q1. + have [x _ sQ1E1x] := Hall_psubJ hallE1 t1q sQ1E qQ1. + by rewrite (cyclicS sQ1E1x) ?cyclicJ. +have defQ1: Q :&: E = Q1. + apply: (sub_pHall sylQ1) (subsetIr Q E); last by rewrite subsetI sQ1Q. + by rewrite (pgroupS (subsetIl Q _)). +pose Q0 := 'C_Q(S); have nsQ0Q: Q0 <| Q by rewrite norm_normalI ?norms_cent. +have [sQ0Q nQ0Q] := andP nsQ0Q; have nQ01 := subset_trans sQ1Q nQ0Q. +have coSQ0: coprime #|S| #|Q0| := coprimegS sQ0Q coSQ. +have ltQ01: Q0 \proper Q1. + rewrite /proper -{1}defQ1 setIS //. + apply: contra (proper_subn ltCQ1_S) => sQ10. + by rewrite subsetIidl (centsS sQ10) // centsC subsetIr. +have [X]: exists2 X, X \in subgroups Q & ('C_S(X) != 1) && ([~: S, X] != 1). + apply/exists_inP; apply: contraFT (ltnn 1); rewrite negb_exists_in => irrS. + have [sQ01 not_sQ10] := andP ltQ01. + have qQb: q.-group (Q / Q0) by exact: quotient_pgroup. + have ntQ1b: Q1 / Q0 != 1 by rewrite -subG1 quotient_sub1. + have ntQb: Q / Q0 != 1 := subG1_contra (quotientS _ sQ1Q) ntQ1b. + have{irrS} regQ: semiregular (S / Q0) (Q / Q0). + move=> Q0x; rewrite 2!inE -cycle_subG -cycle_eq1 -cent_cycle andbC. + case/andP; case/(inv_quotientS nsQ0Q)=> X /= -> {Q0x} sQ0X sXQ ntXb. + have [nSX nQ0X] := (subset_trans sXQ nSQ, subset_trans sXQ nQ0Q). + rewrite -quotient_TI_subcent ?(coprime_TIg coSQ0) //. + apply: contraTeq (forallP irrS X) => ntCSXb; rewrite inE sXQ negbK. + apply/andP; split. + by apply: contraNneq ntCSXb => ->; rewrite quotient1. + apply: contraNneq ntXb; move/commG1P => cXS. + by rewrite quotientS1 // subsetI sXQ centsC. + have{regQ} cycQb: cyclic (Q / Q0). + have nSQb: Q / Q0 \subset 'N(S / Q0) by exact: quotient_norms. + apply: odd_regular_pgroup_cyclic qQb (mFT_quo_odd _ _) _ nSQb regQ. + rewrite -(isog_eq1 (quotient_isog _ _)) ?coprime_TIg 1?coprime_sym //. + by rewrite cents_norm // centsC subsetIr. + have rQ1: 'r(Q1) = 1%N. + apply/eqP; rewrite (rank_Sylow sylQ1). + by rewrite (tau1E maxM hallE) in t1q; case/and3P: t1q. + have: 'r(Q) <= 1; last apply: leq_trans. + have nQ0_Ohm1Q := subset_trans (Ohm_sub 1 Q) nQ0Q. + rewrite -rQ1 -rank_Ohm1 rankS // -(quotientSGK _ sQ01) //. + rewrite (subset_trans (morphim_Ohm _ _ nQ0Q)) //= -quotientE -/Q0. + rewrite -(cardSg_cyclic cycQb) ?Ohm_sub ?quotientS //. + have [_ q_dv_Q1b _] := pgroup_pdiv (pgroupS (quotientS _ sQ1Q) qQb) ntQ1b. + by rewrite (Ohm1_cyclic_pgroup_prime cycQb qQb ntQb). + have ltNA_G: 'N(A) \proper G. + by rewrite defNS mFT_norm_proper // (mFT_pgroup_proper pS). + have [H maxNA_H] := mmax_exists ltNA_G. + have nCEA_Q1 := subset_trans sQ1E (normal_norm nsCEA). + have [_ _] := primes_norm_tau2Elem maxM hallE t2p Ep2A maxNA_H. + case/(_ q)=> [||t2Hq [S2 sylS2 nsS2H] _]. + - rewrite -p_rank_gt0 -(rank_Sylow (quotient_pHall _ sylQ1)) //. + by rewrite rank_gt0 -subG1 quotient_sub1. + - rewrite -p_rank_gt0 -rQ1 (rank_pgroup qQ1) -p_rank_Ohm1 p_rankS //. + have: 'Z(E) \subset 'C_E(A); last apply: subset_trans. + by rewrite setIS ?centS // normal_sub. + have [x Ex sQ1xE1] := Hall_pJsub hallE1 t1q sQ1E qQ1. + rewrite -(conjSg _ _ x) (normsP (normal_norm (center_normal E))) //. + set Q11x := _ :^ x; have oQ11x: #|Q11x| = q. + by rewrite cardJg (Ohm1_cyclic_pgroup_prime _ qQ1) // -rank_gt0 rQ1. + apply: regE1subZ. + apply/nElemP; exists q; rewrite p1ElemE // !inE oQ11x. + by rewrite (subset_trans _ sQ1xE1) //= conjSg Ohm_sub. + have: cyclic Q11x by rewrite prime_cyclic ?oQ11x. + case/cyclicP=> y defQ11x; rewrite /= -/Q11x defQ11x cent_cycle regU13 //. + rewrite !inE -order_gt1 -cycle_subG /order -defQ11x oQ11x prime_gt1 //. + rewrite sub_conjg (subset_trans (Ohm_sub 1 Q1)) //. + by rewrite (normsP (normal_norm nsUE)) ?groupV. + by rewrite /p_elt /order -defQ11x oQ11x pnatE //; apply/orP; left. + rewrite inE in sylS2; have [sS2H _]:= andP nsS2H. + have sylS2_H := pHall_subl sS2H (subsetT H) sylS2. + have [maxH sNS_H] := setIdP maxNA_H; rewrite /= defNS in sNS_H. + have sylS_H := pHall_subl (subset_trans (normG S) sNS_H) (subsetT H) sylS_G. + have defS2: S :=: S2 := uniq_normal_Hall sylS2_H nsS2H (Hall_max sylS_H). + have sylQ_H: q.-Sylow(H) Q by rewrite -(mmax_normal maxH nsS2H) -defS2. + by rewrite (rank_Sylow sylQ_H); case/andP: t2Hq => _ /eqP->. +rewrite inE => sXQ /=; have nSX := subset_trans sXQ nSQ. +set S1 := [~: S, X]; set S2 := 'C_S(X) => /andP[ntS2 ntS1]. +have defS12: S1 \x S2 = S. + by apply: coprime_abelian_cent_dprod; rewrite ?(coprimegS sXQ). +have /mulG_sub[sS1S sS2S] := dprodW defS12. +have [cycS1 cycS2]: cyclic S1 /\ cyclic S2. + apply/andP; rewrite !(abelian_rank1_cyclic (abelianS _ cSS)) //. + rewrite -(leqif_add (leqif_geq _) (leqif_geq _)) ?rank_gt0 // addn1 -rpS. + rewrite !(rank_pgroup (pgroupS _ pS)) ?(p_rank_abelian p (abelianS _ cSS)) //. + by rewrite -lognM ?cardG_gt0 // (dprod_card (Ohm_dprod 1 defS12)). +have [nsS2NS nsS1NS]: S2 <| 'N(S) /\ S1 <| 'N(S) := n_subNS X nSX. +pose Z := if #|S1| < #|S2| then [group of S2] else [group of S1]. +have [ntZ sZS nsZN cycZ]: [/\ Z :!=: 1, Z \subset S, Z <| 'N(S) & cyclic Z]. + by rewrite /Z; case: ifP. +have nsZU: Z <| U := normalS (subset_trans sZS sSU) nSU nsZN. +exists Z; split=> //=. + by rewrite regNNS // (char_norm_trans (Ohm_char 1 Z)) // normal_norm. +rewrite -(dprod_exponent defS12) /= (fun_if val) fun_if !exponent_cyclic //=. +rewrite (card_pgroup (pgroupS sS1S pS)) (card_pgroup (pgroupS sS2S pS)) //. +by rewrite /= -/S1 -/S2 ltn_exp2l ?prime_gt1 // -fun_if expn_max. +Qed. + +(* This is B & G, Theorem 12.13. *) +Theorem nonabelian_Uniqueness p P : p.-group P -> ~~ abelian P -> P \in 'U. +Proof. +move=> pP not_cPP; have [M maxP_M] := mmax_exists (mFT_pgroup_proper pP). +have sigma_p L: L \in 'M(P) -> p \in \sigma(L). + case/setIdP=> maxL sPL; apply: contraR not_cPP => sL'p. + exact: sigma'_nil_abelian maxL sPL (pi_pnat pP _) (pgroup_nil pP). +have{maxP_M} [[maxM sPM] sMp] := (setIdP maxP_M, sigma_p M maxP_M). +rewrite (uniq_mmax_subset1 maxM sPM); apply/subsetP=> H maxP_H; rewrite inE. +have{sigma_p maxP_H} [[maxH sPH] sHp] := (setIdP maxP_H, sigma_p H maxP_H). +without loss{pP sPH sPM} sylP: P not_cPP / p.-Sylow(M :&: H) P. + move=> IH; have sP_MH: P \subset M :&: H by rewrite subsetI sPM. + have [S sylS sPS] := Sylow_superset sP_MH pP. + exact: IH (contra (abelianS sPS) not_cPP) sylS. +have [sP_MH pP _] := and3P sylP; have [sPM sPH] := subsetIP sP_MH. +have ncycP := contra (@cyclic_abelian _ _) not_cPP. +have{sHp} sNMH: 'N(P) \subset M :&: H. + by rewrite subsetI !(@norm_noncyclic_sigma _ p). +have{sylP} sylP_M: p.-Sylow(M) P. + have [S sylS sPS] := Sylow_superset sPM pP; have pS := pHall_pgroup sylS. + have [-> // | ltPS] := eqVproper sPS. + have /andP[sNP] := nilpotent_proper_norm (pgroup_nil pS) ltPS. + rewrite (sub_pHall sylP _ sNP) ?subxx ?(pgroupS (subsetIl _ _)) //. + by rewrite subIset // orbC sNMH. +have{sMp} sylP_G: p.-Sylow(G) P := sigma_Sylow_G maxM sMp sylP_M. +have sylP_H: p.-Sylow(H) P := pHall_subl sPH (subsetT H) sylP_G. +have [rPgt2 | rPle2] := ltnP 2 'r(P). + have uniqP: P \in 'U by rewrite rank3_Uniqueness ?(mFT_pgroup_proper pP). + have defMP: 'M(P) = [set M] := def_uniq_mmax uniqP maxM sPM. + by rewrite -val_eqE /= (eq_uniq_mmax defMP maxH). +have rpP: 'r_p(P) = 2. + apply/eqP; rewrite eqn_leq -{1}rank_pgroup // rPle2 ltnNge. + by rewrite -odd_pgroup_rank1_cyclic ?mFT_odd. +have:= mFT_rank2_Sylow_cprod sylP_G rPle2 not_cPP. +case=> Q [not_cQQ dimQ eQ] [R cycR [defP defR1]]. +have sQP: Q \subset P by have /mulG_sub[] := cprodW defP. +have pQ: p.-group Q := pgroupS sQP pP. +have oQ: #|Q| = (p ^ 3)%N by rewrite (card_pgroup pQ) dimQ. +have esQ: extraspecial Q by apply: (p3group_extraspecial pQ); rewrite ?dimQ. +have p_pr := extraspecial_prime pQ esQ; have p_gt1 := prime_gt1 p_pr. +pose Z := 'Z(Q)%G; have oZ: #|Z| = p := card_center_extraspecial pQ esQ. +have nsZQ: Z <| Q := center_normal Q; have [sZQ nZQ] := andP nsZQ. +have [[defPhiQ defQ'] _]: ('Phi(Q) = Z /\ Q^`(1) = Z) /\ _ := esQ. +have defZ: 'Ohm_1 ('Z(P)) = Z. + have [_ <- _] := cprodP (center_cprod defP). + by rewrite (center_idP (cyclic_abelian cycR)) -defR1 mulSGid ?Ohm_sub. +have ncycQ: ~~ cyclic Q := contra (@cyclic_abelian _ Q) not_cQQ. +have rQgt1: 'r_p(Q) > 1. + by rewrite ltnNge -(odd_pgroup_rank1_cyclic pQ) ?mFT_odd. +have [A Ep2A]: exists A, A \in 'E_p^2(Q) by exact/p_rank_geP. +wlog uniqNEpA: M H maxM maxH sP_MH sNMH sPM sPH sylP_M sylP_H / + ~~ [exists A0 in 'E_p^1(A) :\ Z, 'M('N(A0)) == [set M]]. +- move=> IH; case: exists_inP (IH M H) => [[A0 EpA0 defMA0] _ | _ -> //]. + case: exists_inP {IH}(IH H M) => [[A1 EpA1 defMA1] _ | _]; last first. + by rewrite setIC eq_sym => ->. + have [sAQ abelA dimA] := pnElemP Ep2A; have sAP := subset_trans sAQ sQP. + have transNP: [transitive 'N_P(A), on 'E_p^1(A) :\ Z | 'JG]. + have [|_ _] := basic_p2maxElem_structure _ pP sAP not_cPP. + have Ep2A_G := subsetP (pnElemS p 2 (subsetT Q)) A Ep2A. + rewrite inE Ep2A_G (subsetP (p_rankElem_max p G)) //. + by rewrite -(p_rank_Sylow sylP_G) rpP. + by rewrite (group_inj defZ). + have [x NPx defA1] := atransP2 transNP EpA0 EpA1. + have Mx: x \in M by rewrite (subsetP sPM) //; case/setIP: NPx. + rewrite eq_sym -in_set1 -(group_inj (conjGid Mx)). + by rewrite -(eqP defMA1) defA1 /= normJ mmax_ofJ (eqP defMA0) set11. +apply: contraR uniqNEpA => neqHM; have sQM := subset_trans sQP sPM. +suffices{A Ep2A} [ntMa nonuniqNZ]: M`_\alpha != 1 /\ 'M('N(Z)) != [set M]. + have [A0 EpA0 defMNA0]: exists2 A0, A0 \in 'E_p^1(A) & 'M('N(A0)) == [set M]. + apply/exists_inP; apply: contraR ntMa; rewrite negb_exists_in => uniqNA1. + have{Ep2A} Ep2A: A \in 'E_p^2(M) := subsetP (pnElemS p 2 sQM) A Ep2A. + by have [_ [//|_ ->]] := p2Elem_mmax maxM Ep2A. + apply/exists_inP; exists A0; rewrite // 2!inE EpA0 andbT. + by apply: contraNneq nonuniqNZ => <-. +have coMaQ: coprime #|M`_\alpha| #|Q|. + apply: pnat_coprime (pcore_pgroup _ _) (pi_pnat pQ _). + by rewrite !inE -(p_rank_Sylow sylP_M) rpP. +have nMaQ: Q \subset 'N(M`_\alpha) by rewrite (subset_trans sQM) ?gFnorm. +have [coMaZ nMaZ] := (coprimegS sZQ coMaQ, subset_trans sZQ nMaQ). +pose K := 'N_(M`_\alpha)(Z). +have defKC: 'C_(M`_\alpha)(Z) = K by rewrite -coprime_norm_cent. +have coKQ: coprime #|K| #|Q| := coprimeSg (subsetIl _ _) coMaQ. +have nKQ: Q \subset 'N(K) by rewrite normsI ?norms_norm. +have [coKZ nKZ] := (coprimegS sZQ coKQ, subset_trans sZQ nKQ). +have sKH: K \subset H. + have sZH := subset_trans sZQ (subset_trans sQP sPH). + rewrite -(quotientSGK (subsetIr _ _) sZH) /= -/Z -/K. + have cQQb: abelian (Q / Z) by rewrite -defQ' sub_der1_abelian. + rewrite -(coprime_abelian_gen_cent cQQb) ?coprime_morph ?quotient_norms //. + rewrite gen_subG /= -/K -/Z; apply/bigcupsP=> Ab; rewrite andbC; case/andP. + case/(inv_quotientN nsZQ)=> A -> sZA nsAQ; have sAQ := normal_sub nsAQ. + rewrite (isog_cyclic (third_isog _ _ _)) // -/Z => cycQA. + have pA: p.-group A := pgroupS sAQ pQ. + have rAgt1: 'r_p(A) > 1. + have [-> // | ltAQ] := eqVproper sAQ. + rewrite ltnNge -(odd_pgroup_rank1_cyclic pA) ?mFT_odd //. + apply: contraL cycQA => cycA /=; have cAA := cyclic_abelian cycA. + suff <-: Z :=: A by rewrite -defPhiQ (contra (@Phi_quotient_cyclic _ Q)). + apply/eqP; rewrite eqEcard sZA /= oZ (card_pgroup pA) (leq_exp2l _ 1) //. + by rewrite -abelem_cyclic // /abelem pA cAA (dvdn_trans (exponentS sAQ)). + have [A1 EpA1] := p_rank_geP rAgt1. + rewrite -(setIidPr (subset_trans sAQ (subset_trans sQP sPH))) pnElemI in EpA1. + have{EpA1} [Ep2A1 sA1A]:= setIdP EpA1; rewrite inE in sA1A. + have [sCA1_H _]: 'C(A1) \subset H /\ _ := p2Elem_mmax maxH Ep2A1. + rewrite -quotient_TI_subcent ?(subset_trans sAQ) ?(coprime_TIg coKZ) //= -/K. + by rewrite quotientS // subIset // orbC (subset_trans (centS sA1A)). +have defM: M`_\alpha * (M :&: H) = M. + rewrite setIC in sNMH *. + have [defM eq_aM_bM] := nonuniq_norm_Sylow_pprod maxM maxH neqHM sylP_G sNMH. + by rewrite [M`_\alpha](eq_pcore M eq_aM_bM). +do [split; apply: contraNneq neqHM] => [Ma1 | uniqNZ]. + by rewrite -val_eqE /= (eq_mmax maxM maxH) // -defM Ma1 mul1g subsetIr. +have [_ sNZM]: _ /\ 'N(Z) \subset M := mem_uniq_mmax uniqNZ. +rewrite -val_eqE /= (eq_uniq_mmax uniqNZ maxH) //= -(setIidPr sNZM). +have sZ_MH: Z \subset M :&: H := subset_trans sZQ (subset_trans sQP sP_MH). +rewrite -(pprod_norm_coprime_prod defM) ?pcore_normal ?mmax_sol //. +by rewrite mulG_subG /= defKC sKH setIAC subsetIr. +Qed. + +(* This is B & G, Corollary 12.14. We have removed the redundant assumption *) +(* p \in \sigma(M), and restricted the quantification over P to the part of *) +(* the conclusion where it is mentioned. *) +(* Usage note: it might be more convenient to state that P is a Sylow *) +(* subgroup of M rather than M`_\sigma; check later use. *) +Corollary cent_der_sigma_uniq M p X (Ms := M`_\sigma) : + M \in 'M -> X \in 'E_p^1(M) -> (p \in \beta(M)) || (X \subset Ms^`(1)) -> + 'M('C(X)) = [set M] /\ (forall P, p.-Sylow(Ms) P -> 'M(P) = [set M]). +Proof. +move=> maxM EpX bMp_sXMs'; have p_pr := pnElem_prime EpX. +have [sXM abelX oX] := pnElemPcard EpX; have [pX _] := andP abelX. +have ntX: X :!=: 1 := nt_pnElem EpX isT; have ltCXG := mFT_cent_proper ntX. +have sMp: p \in \sigma(M). + have [bMp | sXMs'] := orP bMp_sXMs'; first by rewrite beta_sub_sigma. + rewrite -pnatE // -[p]oX; apply: pgroupS (subset_trans sXMs' (der_sub 1 _)) _. + exact: pcore_pgroup. +have hallMs: \sigma(M).-Hall(M) Ms by exact: Msigma_Hall. +have sXMs: X \subset Ms by rewrite (sub_Hall_pcore hallMs) // /pgroup oX pnatE. +have [P sylP sXP]:= Sylow_superset sXMs pX. +have sylP_M: p.-Sylow(M) P := subHall_Sylow hallMs sMp sylP. +have sylP_G := sigma_Sylow_G maxM sMp sylP_M. +have [sPM pP _] := and3P sylP_M; have ltPG := mFT_pgroup_proper pP. +suffices [-> uniqP]: 'M('C(X)) = [set M] /\ 'M(P) = [set M]. + split=> // Py sylPy; have [y Ms_y ->] := Sylow_trans sylP sylPy. + rewrite (def_uniq_mmaxJ _ uniqP) (group_inj (conjGid _)) //. + exact: subsetP (pcore_sub _ _) y Ms_y. +have [rCPXgt2 | rCPXle2] := ltnP 2 'r_p('C_P(X)). + have [sCPX_P sCPX_CX] := subsetIP (subxx 'C_P(X)). + have [ltP ltCX] := (mFT_pgroup_proper pP, mFT_cent_proper ntX). + have sCPX_M := subset_trans sCPX_P sPM. + have ltCPX_G := sub_proper_trans sCPX_P ltPG. + suff uniqCPX: 'M('C_P(X)) = [set M] by rewrite !(def_uniq_mmaxS _ _ uniqCPX). + apply: (def_uniq_mmax (rank3_Uniqueness _ _)) => //. + exact: leq_trans (p_rank_le_rank p _). +have nnP: p.-narrow P. + apply: wlog_neg; rewrite negb_imply; case/andP=> rP _. + by apply/narrow_centP; rewrite ?mFT_odd //; exists X. +have{bMp_sXMs'} [bM'p sXMs']: p \notin \beta(M) /\ X \subset Ms^`(1). + move: bMp_sXMs'; rewrite !inE -negb_exists_in. + by case: exists_inP => // [[]]; exists P. +have defMs: 'O_p^'(Ms) ><| P = Ms. + by have [_ hallMp' _] := beta_max_pdiv maxM bM'p; exact/sdprod_Hall_p'coreP. +have{defMs} sXP': X \subset P^`(1). + have{defMs} [_ defMs nMp'P tiMp'P] := sdprodP defMs. + have [injMp'P imMp'P] := isomP (quotient_isom nMp'P tiMp'P). + rewrite -(injmSK injMp'P) // morphim_der // {injMp'P}imMp'P morphim_restrm. + rewrite (setIidPr sXP) /= -quotientMidl defMs -quotient_der ?quotientS //. + by rewrite -defMs mul_subG ?normG. +have [rPgt2 | rPle2] := ltnP 2 'r_p(P). + case/eqP: ntX; rewrite -(setIidPl sXP'). + by case/(narrow_cent_dprod pP (mFT_odd P)): rCPXle2. +have not_cPP: ~~ abelian P. + by rewrite (sameP derG1P eqP) (subG1_contra sXP') ?ntX. +have sXZ: X \subset 'Z(P). + rewrite -rank_pgroup // in rPle2. + have := mFT_rank2_Sylow_cprod sylP_G rPle2 not_cPP. + case=> Q [not_cQQ dimQ _] [R]; move/cyclic_abelian=> cRR [defP _]. + have [_ mulQR _] := cprodP defP; have [sQP _] := mulG_sub mulQR. + rewrite (subset_trans sXP') // -(der_cprod 1 defP) (derG1P cRR) cprodg1. + have{dimQ} dimQ: logn p #|Q| <= 3 by rewrite dimQ. + have [[_ ->] _] := p3group_extraspecial (pgroupS sQP pP) not_cQQ dimQ. + by case/cprodP: (center_cprod defP) => _ <- _; exact: mulG_subl. +have uniqP: 'M(P) = [set M]. + exact: def_uniq_mmax (nonabelian_Uniqueness pP not_cPP) maxM sPM. +rewrite (def_uniq_mmaxS _ ltCXG uniqP) //. +by rewrite centsC (subset_trans sXZ) // subsetIr. +Qed. + +(* This is B & G, Proposition 12.15. *) +Proposition sigma_subgroup_embedding M q X Mstar : + M \in 'M -> q \in \sigma(M) -> X \subset M -> q.-group X -> X :!=: 1 -> + Mstar \in 'M('N(X)) :\ M -> + [/\ (*a*) gval Mstar \notin M :^: G, + forall S, q.-Sylow(M :&: Mstar) S -> X \subset S -> + (*b*) 'N(S) \subset M + /\ (*c*) q.-Sylow(Mstar) S + & if q \in \sigma(Mstar) + (*d*) then + [/\ (*1*) Mstar`_\beta * (M :&: Mstar) = Mstar, + (*2*) {subset \tau1(Mstar) <= [predU \tau1(M) & \alpha(M)]} + & (*3*) M`_\beta = M`_\alpha /\ M`_\alpha != 1] + (*e*) else + [/\ (*1*) q \in \tau2(Mstar), + (*2*) {subset [predI \pi(M) & \sigma(Mstar)] <= \beta(Mstar)} + & (*3*) \sigma(Mstar)^'.-Hall(Mstar) (M :&: Mstar)]]. +Proof. +move: Mstar => H maxM sMq sXM qX ntX /setD1P[neqHM maxNX_H]. +have [q_pr _ _] := pgroup_pdiv qX ntX; have [maxH sNX_H] := setIdP maxNX_H. +have sXH := subset_trans (normG X) sNX_H. +have sX_MH: X \subset M :&: H by apply/subsetIP. +have parts_bc S: + q.-Sylow(M :&: H) S -> X \subset S -> 'N(S) \subset M /\ q.-Sylow(H) S. +- move=> sylS sXS; have [sS_MH qS _] := and3P sylS. + have [sSM sSH] := subsetIP sS_MH. + have sNS_M: 'N(S) \subset M. + have [cycS|] := boolP (cyclic S); last exact: norm_noncyclic_sigma qS _. + have [T sylT sST] := Sylow_superset sSM qS; have [sTM qT _] := and3P sylT. + rewrite -(nilpotent_sub_norm (pgroup_nil qT) sST). + exact: norm_sigma_Sylow sylT. + rewrite (sub_pHall sylS (pgroupS (subsetIl T _) qT)) //. + by rewrite subsetI sST normG. + by rewrite setISS // (subset_trans (char_norms _) sNX_H) // sub_cyclic_char. + split=> //; have [T sylT sST] := Sylow_superset sSH qS. + have [sTH qT _] := and3P sylT. + rewrite -(nilpotent_sub_norm (pgroup_nil qT) sST) //. + rewrite (sub_pHall sylS (pgroupS (subsetIl T _) qT)) //=. + by rewrite subsetI sST normG. + by rewrite /= setIC setISS. +have [S sylS sXS] := Sylow_superset sX_MH qX; have [sS_MH qS _] := and3P sylS. +have [sSM sSH] := subsetIP sS_MH; have [sNS_M sylS_H] := parts_bc S sylS sXS. +have notMGH: gval H \notin M :^: G. + by apply: mmax_norm_notJ maxM maxH qX sXM sNX_H _; rewrite sMq eq_sym neqHM. +have /orP[sHq | t2Hq] := prime_class_mmax_norm maxH qX sNX_H; last first. + have [/= sH'q rqH] := andP t2Hq; rewrite [q \in _](negPf sH'q); split=> //. + have [A Eq2A] := p_rank_witness q S; have [sAS abelA dimA] := pnElemP Eq2A. + rewrite (p_rank_Sylow sylS_H) (eqP rqH) in dimA; have [qA _] := andP abelA. + have [sAH sAM] := (subset_trans sAS sSH, subset_trans sAS sSM). + have [F hallF sAF] := Hall_superset (mmax_sol maxH) sAH (pi_pnat qA sH'q). + have tiHsM: H`_\sigma :&: M = 1. + have{Eq2A} Eq2A: A \in 'E_q^2(H) by apply/pnElemP. + have [_ _ _ -> //] := tau2_context maxH t2Hq Eq2A. + by rewrite 3!inE eq_sym neqHM maxM. + have{Eq2A} Eq2A_F: A \in 'E_q^2(F) by apply/pnElemP. + have [[nsAF _] [sCA_F _ _] _ TIsH] + := tau2_compl_context maxH hallF t2Hq Eq2A_F. + have sNA_M: 'N(A) \subset M. + apply: norm_noncyclic_sigma maxM sMq qA sAM _. + by rewrite (abelem_cyclic abelA) dimA. + have ->: M :&: H = F. + have [[_ <- _ _] [_ nAF]] := (sdprodP (sdprod_sigma maxH hallF), andP nsAF). + by rewrite -(group_modr _ (subset_trans nAF sNA_M)) setIC tiHsM mul1g. + split=> // p /andP[/= piMp sHp]; apply: wlog_neg => bH'p. + have bM'q: q \notin \beta(M). + by rewrite -predI_sigma_beta // inE /= sMq; case/tau2_not_beta: t2Hq. + have sM'p: p \notin \sigma(M). + rewrite orbit_sym in notMGH; have [_ TIsHM] := TIsH M maxM notMGH. + by have:= TIsHM p; rewrite inE /= sHp /= => ->. + have p'CA: p^'.-group 'C(A). + by rewrite (pgroupS sCA_F) // (pi'_p'group (pHall_pgroup hallF)). + have p_pr: prime p by rewrite mem_primes in piMp; case/andP: piMp. + have [lt_pq | lt_qp | eq_pq] := ltngtP p q; last 1 first. + - by rewrite eq_pq sMq in sM'p. + - have bH'q: q \notin \beta(H) by apply: contra sH'q; apply: beta_sub_sigma. + have [|[P sylP cPA] _ _] := beta'_cent_Sylow maxH bH'p bH'q qA. + by rewrite lt_pq sAH orbT. + have sylP_H := subHall_Sylow (Msigma_Hall maxH) sHp sylP. + have piPp: p \in \pi(P). + by rewrite -p_rank_gt0 (p_rank_Sylow sylP_H) p_rank_gt0 sigma_sub_pi. + by rewrite centsC in cPA; case/eqnP: (pnatPpi (pgroupS cPA p'CA) piPp). + have bM'p: p \notin \beta(M) by apply: contra sM'p; apply: beta_sub_sigma. + have [P sylP] := Sylow_exists p M; have [sMP pP _] := and3P sylP. + have [|[Q1 sylQ1 cQ1P] _ _] := beta'_cent_Sylow maxM bM'q bM'p pP. + by rewrite lt_qp sMP orbT. + have sylQ1_M := subHall_Sylow (Msigma_Hall maxM) sMq sylQ1. + have [x Mx sAQ1x] := Sylow_subJ sylQ1_M sAM qA. + have sPxCA: P :^ x \subset 'C(A) by rewrite (centsS sAQ1x) // centJ conjSg. + have piPx_p: p \in \pi(P :^ x). + by rewrite /= cardJg -p_rank_gt0 (p_rank_Sylow sylP) p_rank_gt0. + by case/eqnP: (pnatPpi (pgroupS sPxCA p'CA) piPx_p). +rewrite sHq; split=> //. +have sNS_HM: 'N(S) \subset H :&: M by rewrite subsetI (norm_sigma_Sylow sHq). +have sylS_G: q.-Sylow(G) S := sigma_Sylow_G maxH sHq sylS_H. +have [defM eq_abM] := nonuniq_norm_Sylow_pprod maxM maxH neqHM sylS_G sNS_HM. +rewrite setIC eq_sym in sNS_HM neqHM defM. +have [defH eq_abH] := nonuniq_norm_Sylow_pprod maxH maxM neqHM sylS_G sNS_HM. +rewrite [M`_\alpha](eq_pcore M eq_abM) -/M`_\beta. +split=> // [r t1Hr|]; last first. + split=> //; apply: contraNneq neqHM => Mb1. + by rewrite -val_eqE /= (eq_mmax maxM maxH) // -defM Mb1 mul1g subsetIr. +have [R sylR] := Sylow_exists r (M :&: H); have [sR_MH rR _] := and3P sylR. +have [sRM sRH] := subsetIP sR_MH; have [sH'r rrH not_rH'] := and3P t1Hr. +have bH'r: r \notin \beta(H). + by apply: contra sH'r; rewrite -eq_abH; apply: alpha_sub_sigma. +have sylR_H: r.-Sylow(H) R. + rewrite pHallE sRH -defH -LagrangeMr partnM ?cardG_gt0 //. + rewrite -(card_Hall sylR) part_p'nat ?mul1n ?(pnat_dvd (dvdn_indexg _ _)) //=. + by rewrite (pi_p'nat (pcore_pgroup _ _)). +rewrite inE /= orbC -implyNb eq_abM; apply/implyP=> bM'r. +have sylR_M: r.-Sylow(M) R. + rewrite pHallE sRM -defM -LagrangeMr partnM ?cardG_gt0 //. + rewrite -(card_Hall sylR) part_p'nat ?mul1n ?(pnat_dvd (dvdn_indexg _ _)) //=. + by rewrite (pi_p'nat (pcore_pgroup _ _)). +have rrR: 'r_r(R) = 1%N by rewrite (p_rank_Sylow sylR_H) (eqP rrH). +have piRr: r \in \pi(R) by rewrite -p_rank_gt0 rrR. +suffices not_piM'r: r \notin \pi(M^`(1)). + rewrite inE /= -(p_rank_Sylow sylR_M) rrR /= -negb_or /=. + apply: contra not_piM'r; case/orP=> [sMr | rM']. + have sRMs: R \subset M`_\sigma. + by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) ?(pi_pgroup rR). + by rewrite (piSg (Msigma_der1 maxM)) // (piSg sRMs). + by move: piRr; rewrite !mem_primes !cardG_gt0; case/andP=> ->. +have coMbR: coprime #|M`_\beta| #|R|. + exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat rR _). +have sylRM': r.-Sylow(M^`(1)) _ := Hall_setI_normal (der_normal 1 M) sylR_M. +rewrite -p'groupEpi -partG_eq1 -(card_Hall sylRM') -trivg_card1 /=. +rewrite (pprod_focal_coprime defM (pcore_normal _ _)) //. +rewrite coprime_TIg ?(pnat_coprime rR (pgroupS (dergS 1 (subsetIr _ _)) _)) //. +by rewrite p'groupEpi mem_primes (negPf not_rH') !andbF. +Qed. + +(* This is B & G, Corollary 12.16. *) +Corollary sigma_Jsub M Y : + M \in 'M -> \sigma(M).-group Y -> Y :!=: 1 -> + [/\ exists x, Y :^ x \subset M`_\sigma + & forall E p H, + \sigma(M)^'.-Hall(M) E -> p \in \pi(E) -> p \notin \beta(G) -> + H \in 'M(Y) -> gval H \notin M :^: G -> + [/\ (*a*) 'r_p('N_H(Y)) <= 1 + & (*b*) p \in \tau1(M) -> p \notin \pi(('N_H(Y))^`(1))]]. +Proof. +move=> maxM sM_Y ntY. +have ltYG: Y \proper G. + have ltMsG: M`_\sigma \proper G. + exact: sub_proper_trans (pcore_sub _ _) (mmax_proper maxM). + rewrite properEcard subsetT (leq_ltn_trans _ (proper_card ltMsG)) //. + rewrite dvdn_leq ?cardG_gt0 // (card_Hall (Msigma_Hall_G maxM)). + by rewrite -(part_pnat_id sM_Y) partn_dvd // cardSg ?subsetT. +have [q q_pr rFY] := rank_witness 'F(Y). +have [X [ntX qX charX]]: exists X, [/\ gval X :!=: 1, q.-group X & X \char Y]. + exists ('O_q(Y))%G; rewrite pcore_pgroup pcore_char //. + rewrite -rank_gt0 /= -p_core_Fitting. + rewrite (rank_Sylow (nilpotent_pcore_Hall q (Fitting_nil Y))) -rFY. + by rewrite rank_gt0 (trivg_Fitting (mFT_sol ltYG)). +have sXY: X \subset Y := char_sub charX. +have sMq: q \in \sigma(M). + apply: (pnatPpi (pgroupS sXY sM_Y)). + by rewrite -p_rank_gt0 -(rank_pgroup qX) rank_gt0. +without loss sXMs: M maxM sM_Y sMq / X \subset M`_\sigma. + move=> IH; have [Q sylQ] := Sylow_exists q M`_\sigma. + have sQMs := pHall_sub sylQ. + have sylQ_G := subHall_Sylow (Msigma_Hall_G maxM) sMq sylQ. + have [x Gx sXQx] := Sylow_subJ sylQ_G (subsetT X) qX. + have: X \subset M`_\sigma :^ x by rewrite (subset_trans sXQx) ?conjSg. + rewrite -MsigmaJ => /IH; rewrite sigmaJ mmaxJ (eq_pgroup _ (sigmaJ _ _)). + case => // [[y sYyMx] parts_ab]. + split=> [|E p H hallE piEp bG'p maxY_H notMGH]. + by exists (y * x^-1); rewrite conjsgM sub_conjgV -MsigmaJ. + have:= parts_ab (E :^ x)%G p H; rewrite tau1J /= cardJg pHallJ2. + rewrite (eq_pHall _ _ (eq_negn (sigmaJ _ _))). + by rewrite 2!orbit_sym (orbit_transl (mem_orbit _ _ _)) //; apply. +have pre_part_a E p H: + \sigma(M)^'.-Hall(M) E -> p \in \pi(E) -> + H \in 'M(Y) -> gval H \notin M :^: G -> 'r_p(H :&: M) <= 1. +- move=> hallE piEp /setIdP[maxH sYH] notMGH; rewrite leqNgt. + apply: contra ntX => /p_rank_geP[A /pnElemP[/subsetIP[sAH sAM] abelA dimA]]. + have{abelA dimA} Ep2A: A \in 'E_p^2(M) by apply/pnElemP. + have rpMgt1: 'r_p(M) > 1 by apply/p_rank_geP; exists A. + have t2Mp: p \in \tau2(M). + move: piEp; rewrite (partition_pi_sigma_compl maxM hallE) orbCA orbC. + by rewrite -2!andb_orr orNb eqn_leq leqNgt rpMgt1 !andbF. + have sM'p := pnatPpi (pHall_pgroup hallE) piEp. + have [_ _ _ tiMsH _] := tau2_context maxM t2Mp Ep2A. + rewrite -subG1 -(tiMsH H); first by rewrite subsetI sXMs (subset_trans sXY). + by rewrite 3!inE maxH (contra_orbit _ _ notMGH). +have [sNX_M | not_sNX_M] := boolP ('N(X) \subset M). + have sNY_M: 'N(Y) \subset M := subset_trans (char_norms charX) sNX_M. + split=> [|E p H hallE piEp bG'p maxY_H notMGH]; last split. + - exists 1; rewrite act1 (sub_Hall_pcore (Msigma_Hall maxM)) //. + exact: subset_trans (normG Y) sNY_M. + - rewrite (leq_trans (p_rankS p (setIS H sNY_M))) ?(pre_part_a E) //. + case/and3P=> _ _; apply: contra; rewrite mem_primes => /and3P[_ _ pM']. + by apply: dvdn_trans pM' (cardSg (dergS 1 _)); rewrite subIset ?sNY_M ?orbT. +have [L maxNX_L] := mmax_exists (mFT_norm_proper ntX (mFT_pgroup_proper qX)). +have [maxL sNX_L] := setIdP maxNX_L. +have{maxNX_L} maxNX_L: L \in 'M('N(X)) :\ M. + by rewrite 2!inE maxNX_L andbT; apply: contraNneq not_sNX_M => <-. +have sXM := subset_trans sXMs (pcore_sub _ M). +have [notMGL _ embedL] := sigma_subgroup_embedding maxM sMq sXM qX ntX maxNX_L. +pose K := (if q \in \sigma(L) then L`_\beta else L`_\sigma)%G. +have sM'K: \sigma(M)^'.-group K. + rewrite orbit_sym in notMGL. + rewrite /K; case: (boolP (q \in _)) embedL => [sLq _ | sL'q [t2Lq _ _]]. + have [_ TIaLsM _] := sigma_disjoint maxL maxM notMGL. + apply: sub_pgroup (pcore_pgroup _ _) => p bLp. + by apply: contraFN (TIaLsM p) => /= sMp; rewrite inE /= beta_sub_alpha. + have [F hallF] := ex_sigma_compl maxL. + have [A Ep2A _] := ex_tau2Elem hallF t2Lq. + have [_ _ _ TIsLs] := tau2_compl_context maxL hallF t2Lq Ep2A. + have{TIsLs} [_ TIsLsM] := TIsLs M maxM notMGL. + apply: sub_pgroup (pcore_pgroup _ _) => p sLp. + by apply: contraFN (TIsLsM p) => /= sMp; rewrite inE /= sLp. +have defL: K * (M :&: L) = L. + rewrite /K; case: (q \in _) embedL => [] [] // _ _. + by move/(sdprod_Hall_pcoreP (Msigma_Hall maxL)); case/sdprodP. +have sYL := subset_trans (char_norm charX) sNX_L. +have [x sYxMs]: exists x, Y :^ x \subset M`_\sigma. + have solML := solvableS (subsetIl M L) (mmax_sol maxM). + have [H hallH] := Hall_exists \sigma(M) solML. + have [sHM sHL] := subsetIP (pHall_sub hallH). + have hallH_L: \sigma(M).-Hall(L) H. + rewrite pHallE sHL -defL -LagrangeMr partnM ?cardG_gt0 //. + rewrite -(card_Hall hallH) part_p'nat ?mul1n //=. + exact: pnat_dvd (dvdn_indexg _ _) sM'K. + have [x _ sYxH]:= Hall_Jsub (mmax_sol maxL) hallH_L sYL sM_Y. + exists x; rewrite (sub_Hall_pcore (Msigma_Hall maxM)) ?pgroupJ //. + exact: subset_trans sYxH sHM. +split=> [|E p H hallE piEp bG'p maxY_H notMGH]; first by exists x. +have p'K: p^'.-group K. + have bL'p: p \notin \beta(L). + by rewrite -predI_sigma_beta // negb_and bG'p orbT. + rewrite /K; case: (q \in _) embedL => [_ | [_ bLp _]]. + by rewrite (pi_p'group (pcore_pgroup _ _)). + rewrite (pi_p'group (pcore_pgroup _ _)) //; apply: contra bL'p => /= sLp. + by rewrite bLp // inE /= (piSg (pHall_sub hallE)). +have sNHY_L: 'N_H(Y) \subset L. + by rewrite subIset ?(subset_trans (char_norms charX)) ?orbT. +rewrite (leq_trans (p_rankS p sNHY_L)); last first. + have [P sylP] := Sylow_exists p (M :&: L). + have [_ sPL] := subsetIP (pHall_sub sylP). + have{sPL} sylP_L: p.-Sylow(L) P. + rewrite pHallE sPL -defL -LagrangeMr partnM ?cardG_gt0 //. + rewrite -(card_Hall sylP) part_p'nat ?mul1n //=. + exact: pnat_dvd (dvdn_indexg _ _) p'K. + rewrite -(p_rank_Sylow sylP_L) {P sylP sylP_L}(p_rank_Sylow sylP). + by rewrite /= setIC (pre_part_a E) // inE maxL. +split=> // t1Mp; rewrite (contra ((piSg (dergS 1 sNHY_L)) p)) // -p'groupEpi. +have nsKL: K <| L by rewrite /K; case: ifP => _; apply: pcore_normal. +have [sKL nKL] := andP nsKL; have nKML := subset_trans (subsetIr M L) nKL. +suffices: p^'.-group (K * (M :&: L)^`(1)). + have sder := subset_trans (der_sub 1 _). + rewrite -norm_joinEr ?sder //; apply: pgroupS => /=. + rewrite norm_joinEr -?quotientSK ?sder //= !quotient_der //. + by rewrite -{1}defL quotientMidl. +rewrite pgroupM p'K (pgroupS (dergS 1 (subsetIl M L))) // p'groupEpi. +by rewrite mem_primes andbA andbC negb_and; case/and3P: t1Mp => _ _ ->. +Qed. + +(* This is B & G, Lemma 12.17. *) +Lemma sigma_compl_embedding M E (Ms := M`_\sigma) : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> + [/\ 'C_Ms(E) \subset Ms^`(1), [~: Ms, E] = Ms + & forall g (MsMg := Ms :&: M :^ g), g \notin M -> + [/\ cyclic MsMg, \beta(M)^'.-group MsMg & MsMg :&: Ms^`(1) = 1]]. +Proof. +move=> maxM hallE; have [sEM s'E _] := and3P hallE. +have solMs: solvable Ms := solvableS (pcore_sub _ _) (mmax_sol maxM). +have defM := coprime_der1_sdprod (sdprod_sigma maxM hallE). +have{s'E} coMsE: coprime #|Ms| #|E| := pnat_coprime (pcore_pgroup _ _) s'E. +have{defM coMsE} [-> ->] := defM coMsE solMs (Msigma_der1 maxM). +split=> // g MsMg notMg. +have sMsMg: \sigma(M).-group MsMg := pgroupS (subsetIl _ _) (pcore_pgroup _ _). +have EpMsMg p n X: X \in 'E_p^n(MsMg) -> n > 0 -> + n = 1%N /\ ~~ ((p \in \beta(M)) || (X \subset Ms^`(1))). +- move=> EpX n_gt0; have [sXMsMg abelX dimX] := pnElemP EpX. + have [[sXMs sXMg] [pX _]] := (subsetIP sXMsMg, andP abelX). + have sXM := subset_trans sXMs (pcore_sub _ _). + have piXp: p \in \pi(X) by rewrite -p_rank_gt0 p_rank_abelem ?dimX. + have sMp: p \in \sigma(M) := pnatPpi (pgroupS sXMs (pcore_pgroup _ _)) piXp. + have not_sCX_M: ~~ ('C(X) \subset M). + apply: contra notMg => sCX_M; rewrite -groupV. + have [transCX _ _] := sigma_group_trans maxM sMp pX. + have [|c CXc [m Mm ->]] := transCX g^-1 sXM; rewrite ?sub_conjgV //. + by rewrite groupM // (subsetP sCX_M). + have cycX: cyclic X. + apply: contraR not_sCX_M => ncycX; apply: subset_trans (cent_sub _) _. + exact: norm_noncyclic_sigma maxM sMp pX sXM ncycX. + have n1: n = 1%N by apply/eqP; rewrite eqn_leq -{1}dimX -abelem_cyclic ?cycX. + rewrite n1 in dimX *; split=> //; apply: contra not_sCX_M. + by case/cent_der_sigma_uniq=> //; [apply/pnElemP | case/mem_uniq_mmax]. +have tiMsMg_Ms': MsMg :&: Ms^`(1) = 1. + apply/eqP/idPn; rewrite -rank_gt0 => /rank_geP[X /nElemP[p]]. + case/pnElemP=> /subsetIP[sXMsMg sXMs'] abelX dimX. + by case: (EpMsMg p 1%N X) => //; [apply/pnElemP | rewrite sXMs' orbT]. +split=> //; last first. + apply: sub_in_pnat sMsMg => p. + by rewrite -p_rank_gt0 => /p_rank_geP[X /EpMsMg[] // _ /norP[]]. +rewrite abelian_rank1_cyclic. + by rewrite leqNgt; apply/rank_geP=> [[X /nElemP[p /EpMsMg[]]]]. +by rewrite (sameP derG1P trivgP) -tiMsMg_Ms' subsetI der_sub dergS ?subsetIl. +Qed. + +(* This is B & G, Lemma 12.18. *) +(* We corrected an omission in the text, which fails to quote Lemma 10.3 to *) +(* justify the two p-rank inequalities (12.5) and (12.6), and indeed *) +(* erroneously refers to 12.2(a) for (12.5). Note also that the loosely *) +(* justified equalities of Ohm_1 subgroups are in fact unnecessary. *) +Lemma cent_Malpha_reg_tau1 M p q P Q (Ma := M`_\alpha) : + M \in 'M -> p \in \tau1(M) -> q \in p^' -> P \in 'E_p^1(M) -> Q :!=: 1 -> + P \subset 'N(Q) -> 'C_Q(P) = 1 -> 'M('N(Q)) != [set M] -> + [/\ (*a*) Ma != 1 -> q \notin \alpha(M) -> q.-group Q -> Q \subset M -> + 'C_Ma(P) != 1 /\ 'C_Ma(Q <*> P) = 1 + & (*b*) q.-Sylow(M) Q -> + [/\ \alpha(M) =i \beta(M), Ma != 1, q \notin \alpha(M), + 'C_Ma(P) != 1 & 'C_Ma(Q <*> P) = 1]]. +Proof. +move=> maxM t1p p'q EpP ntQ nQP regPQ nonuniqNQ. +set QP := Q <*> P; set CaQP := 'C_Ma(QP); set part_a := _ -> _. +have ssolM := solvableS _ (mmax_sol maxM). +have [sPM abelP oP] := pnElemPcard EpP; have{abelP} [pP _] := andP abelP. +have p_pr := pnElem_prime EpP; have [s'p _] := andP t1p. +have a'p: p \in \alpha(M)^' by apply: contra s'p; apply: alpha_sub_sigma. +have{a'p} [a'P t2'p] := (pi_pgroup pP a'p, tau2'1 t1p). +have uniqCMX: 'M('C_M(_)) = [set M] := def_uniq_mmax _ maxM (subsetIl _ _). +have nQ_CMQ: 'C_M(Q) \subset 'N(Q) by rewrite setIC subIset ?cent_sub. +have part_a_holds: part_a. + move=> ntMa a'q qQ sQM; have{p'q} p'Q := pi_pgroup qQ p'q. + have{p'Q} coQP: coprime #|Q| #|P| by rewrite coprime_sym (pnat_coprime pP). + have{a'q} a'Q: \alpha(M)^'.-group Q by rewrite (pi_pgroup qQ). + have rCMaQle1: 'r('C_Ma(Q)) <= 1. + rewrite leqNgt; apply: contra nonuniqNQ => rCMaQgt1; apply/eqP. + apply: def_uniq_mmaxS (uniqCMX Q _) => //; last exact: cent_alpha'_uniq. + exact: mFT_norm_proper (mFT_pgroup_proper qQ). + have rCMaPle1: 'r('C_Ma(P)) <= 1. + have: ~~ ('N(P) \subset M). + by apply: contra (prime_class_mmax_norm maxM pP) _; apply/norP. + rewrite leqNgt; apply: contra => rCMaPgt1. + apply: (sub_uniq_mmax (uniqCMX P _)); first exact: cent_alpha'_uniq. + by rewrite /= setIC subIset ?cent_sub. + exact: mFT_norm_proper (nt_pnElem EpP _) (mFT_pgroup_proper pP). + have [sMaM nMaM] := andP (pcore_normal _ M : Ma <| M). + have aMa: \alpha(M).-group Ma by rewrite pcore_pgroup. + have nMaQP: QP \subset 'N(Ma) by rewrite join_subG !(subset_trans _ nMaM). + have{nMaM} coMaQP: coprime #|Ma| #|QP|. + by rewrite (pnat_coprime aMa) ?[QP]norm_joinEr // [pnat _ _]pgroupM ?a'Q. + pose r := pdiv #|if CaQP == 1 then Ma else CaQP|. + have{ntMa} piMar: r \in \pi(Ma). + rewrite /r; case: ifPn => [_| ntCaQP]; first by rewrite pi_pdiv cardG_gt1. + by rewrite (piSg (subsetIl Ma 'C(QP))) // pi_pdiv cardG_gt1. + have{aMa} a_r: r \in \alpha(M) := pnatPpi aMa piMar. + have [r'Q r'P] : r^'.-group Q /\ r^'.-group P by rewrite !(pi'_p'group _ a_r). + have [Rc /= sylRc] := Sylow_exists r [group of CaQP]. + have [sRcCaQP rRc _] := and3P sylRc; have [sRcMa cQPRc] := subsetIP sRcCaQP. + have nRcQP: QP \subset 'N(Rc) by rewrite cents_norm // centsC. + have{nMaQP rRc coMaQP sRcCaQP sRcMa nRcQP} [R [sylR nR_QP sRcR]] := + coprime_Hall_subset nMaQP coMaQP (ssolM _ sMaM) sRcMa rRc nRcQP. + have{nR_QP} [[sRMa rR _] [nRQ nRP]] := (and3P sylR, joing_subP nR_QP). + have{piMar} ntR: R :!=: 1 by rewrite -rank_gt0 (rank_Sylow sylR) p_rank_gt0. + have [r_pr _ _] := pgroup_pdiv rR ntR. + have sylR_M := subHall_Sylow (Malpha_Hall maxM) a_r sylR. + have{rCMaQle1 a_r} not_cRQ: ~~ (Q \subset 'C(R)). + apply: contraL rCMaQle1; rewrite centsC => cQR; rewrite -ltnNge ltnW //. + by rewrite (leq_trans a_r) // -(rank_Sylow sylR_M) rankS // subsetI sRMa. + have [R1 [charR1 _ _ expR1 rCR1_AutR]] := critical_odd rR (mFT_odd R) ntR. + have [sR1R nR1R] := andP (char_normal charR1); have rR1 := pgroupS sR1R rR. + have [nR1P nR1Q] := (char_norm_trans charR1 nRP, char_norm_trans charR1 nRQ). + have [coR1Q coR1P] := (pnat_coprime rR1 r'Q, pnat_coprime rR1 r'P). + have {rCR1_AutR not_cRQ} not_cR1Q: ~~ (Q \subset 'C(R1)). + apply: contra not_cRQ => cR1Q; rewrite -subsetIidl. + rewrite -quotient_sub1 ?normsI ?normG ?norms_cent // subG1 trivg_card1. + rewrite (pnat_1 _ (quotient_pgroup _ r'Q)) //= -ker_conj_aut. + rewrite (card_isog (first_isog_loc _ _)) //; apply: pgroupS rCR1_AutR. + apply/subsetP=> za; case/morphimP=> z nRz Qz ->; rewrite inE Aut_aut inE. + apply/subsetP=> x R1x; rewrite inE [_ x _]norm_conj_autE ?(subsetP sR1R) //. + by rewrite /conjg -(centsP cR1Q z) ?mulKg. + pose R0 := 'C_R1(Q); have sR0R1: R0 \subset R1 := subsetIl R1 'C(Q). + have nR0P: P \subset 'N(R0) by rewrite normsI ?norms_cent. + have nR0Q: Q \subset 'N(R0) by rewrite normsI ?norms_cent ?normG. + pose R1Q := R1 <*> Q; have defR1Q: R1 * Q = R1Q by rewrite -norm_joinEr. + have [[sR1_R1Q sQ_R1Q] tiR1Q] := (joing_sub (erefl R1Q), coprime_TIg coR1Q). + have not_nilR1Q: ~~ nilpotent R1Q. + by apply: contra not_cR1Q => /sub_nilpotent_cent2; apply. + have not_nilR1Qb: ~~ nilpotent (R1Q / R0). + apply: contra not_cR1Q => nilR1Qb. + have [nilR1 solR1] := (pgroup_nil rR1, pgroup_sol rR1). + rewrite centsC -subsetIidl -(nilpotent_sub_norm nilR1 sR0R1) //= -/R0. + rewrite -(quotientSGK (subsetIr R1 _)) ?coprime_quotient_cent //= -/R0. + rewrite quotientInorm subsetIidl /= centsC -/R0. + by rewrite (sub_nilpotent_cent2 nilR1Qb) ?quotientS ?coprime_morph. + have coR1QP: coprime #|R1Q| #|P|. + by rewrite -defR1Q TI_cardMg // coprime_mull coR1P. + have defR1QP: R1Q ><| P = R1Q <*> P. + by rewrite sdprodEY ?normsY ?coprime_TIg. + have sR1Ma := subset_trans sR1R sRMa; have sR1M := subset_trans sR1Ma sMaM. + have solR1Q: solvable R1Q by rewrite ssolM // !join_subG sR1M. + have solR1QP: solvable (R1Q <*> P) by rewrite ssolM // !join_subG sR1M sQM. + have defCR1QP: 'C_R1Q(P) = 'C_R1(P). + by rewrite -defR1Q -subcent_TImulg ?regPQ ?mulg1 //; apply/subsetIP. + have ntCR1P: 'C_R1(P) != 1. + apply: contraNneq not_nilR1Q => regPR1. + by rewrite (prime_Frobenius_sol_kernel_nil defR1QP) ?oP ?defCR1QP. + split; first exact: subG1_contra (setSI _ sR1Ma) ntCR1P. + have{rCMaPle1} cycCRP: cyclic 'C_R(P). + have rCRP: r.-group 'C_R(P) := pgroupS (subsetIl R _) rR. + rewrite (odd_pgroup_rank1_cyclic rCRP) ?mFT_odd -?rank_pgroup {rCRP}//. + by rewrite (leq_trans (rankS _) rCMaPle1) ?setSI. + have{ntCR1P} oCR1P: #|'C_R1(P)| = r. + have cycCR1P: cyclic 'C_R1(P) by rewrite (cyclicS _ cycCRP) ?setSI. + apply: cyclic_abelem_prime ntCR1P => //. + by rewrite abelemE ?cyclic_abelian // -expR1 exponentS ?subsetIl. + apply: contraNeq not_nilR1Qb => ntCaQP. + have{Rc sRcR sylRc cQPRc ntCaQP} ntCRQP: 'C_R(QP) != 1. + suffices: Rc :!=: 1 by apply: subG1_contra; apply/subsetIP. + rewrite -rank_gt0 (rank_Sylow sylRc) p_rank_gt0. + by rewrite /r (negPf ntCaQP) pi_pdiv cardG_gt1. + have defR1QPb: (R1Q / R0) ><| (P / R0) = R1Q <*> P / R0. + have [_ <- nR1QP _] := sdprodP defR1QP; rewrite quotientMr //. + by rewrite sdprodE ?quotient_norms // coprime_TIg ?coprime_morph. + have tiPR0: R0 :&: P = 1 by rewrite coprime_TIg // (coprimeSg sR0R1). + have prPb: prime #|P / R0| by rewrite -(card_isog (quotient_isog _ _)) ?oP. + rewrite (prime_Frobenius_sol_kernel_nil defR1QPb) ?quotient_sol //. + rewrite -coprime_quotient_cent ?(subset_trans sR0R1) // quotientS1 //=. + rewrite defCR1QP -{2}(setIidPl sR1R) -setIA subsetI subsetIl. + apply: subset_trans (setIS R (centS (joing_subl Q P))). + rewrite -(cardSg_cyclic cycCRP) ?setIS ?setSI ?centS ?joing_subr // oCR1P. + by have [_ -> _] := pgroup_pdiv (pgroupS (subsetIl R _) rR) ntCRQP. +split=> // sylQ; have [sQM qQ _] := and3P sylQ. +have ltQG := mFT_pgroup_proper qQ; have ltNQG := mFT_norm_proper ntQ ltQG. +have{p'q} p'Q := pi_pgroup qQ p'q. +have{p'Q} coQP: coprime #|Q| #|P| by rewrite coprime_sym (pnat_coprime pP). +have sQM': Q \subset M^`(1). + by rewrite -(coprime_cent_prod nQP) ?ssolM // regPQ mulg1 commgSS. +have ntMa: Ma != 1. + apply: contraNneq nonuniqNQ => Ma1. + rewrite (mmax_normal maxM _ ntQ) ?mmax_sup_id //. + apply: char_normal_trans (der_normal 1 M). + have sylQ_M': q.-Sylow(M^`(1)) Q := pHall_subl sQM' (der_sub 1 M) sylQ. + rewrite (nilpotent_Hall_pcore _ sylQ_M') ?pcore_char //. + by rewrite (isog_nil (quotient1_isog _)) -Ma1 Malpha_quo_nil. +have a'q: q \notin \alpha(M). + apply: contra nonuniqNQ => a_q. + have uniqQ: Q \in 'U by rewrite rank3_Uniqueness ?(rank_Sylow sylQ). + by rewrite (def_uniq_mmaxS _ _ (def_uniq_mmax _ _ sQM)) ?normG. +have b'q := contra (@beta_sub_alpha _ M _) a'q. +case: part_a_holds => // ntCaP regQP; split=> {ntCaP regQP}// r. +apply/idP/idP=> [a_r | ]; last exact: beta_sub_alpha. +apply: contraR nonuniqNQ => b'r; apply/eqP. +apply: def_uniq_mmaxS (uniqCMX Q _) => //. +have q'r: r != q by apply: contraNneq a'q => <-. +by have [|_ -> //] := beta'_cent_Sylow maxM b'r b'q qQ; rewrite q'r sQM'. +Qed. + +(* This is B & G, Lemma 12.19. *) +(* We have used lemmas 10.8(b) and 10.2(c) rather than 10.9(a) as suggested *) +(* in the text; this avoids a quantifier inversion! *) +Lemma der_compl_cent_beta' M E : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> + exists2 H : {group gT}, \beta(M)^'.-Hall(M`_\sigma) H & E^`(1) \subset 'C(H). +Proof. +move=> maxM hallE; have [sEM s'E _] := and3P hallE. +have s'E': \sigma(M)^'.-group E^`(1) := pgroupS (der_sub 1 E) s'E. +have b'E': \beta(M)^'.-group E^`(1). + by apply: sub_pgroup s'E' => p; apply: contra; apply: beta_sub_sigma. +have solM': solvable M^`(1) := solvableS (der_sub 1 M) (mmax_sol maxM). +have [K hallK sE'K] := Hall_superset solM' (dergS 1 sEM) b'E'. +exists (K :&: M`_\sigma)%G. + apply: Hall_setI_normal hallK. + exact: normalS (Msigma_der1 maxM) (der_sub 1 M) (pcore_normal _ M). +have nilK: nilpotent K. + by have [sKM' b'K _] := and3P hallK; apply: beta'_der1_nil sKM'. +rewrite (sub_nilpotent_cent2 nilK) ?subsetIl ?(coprimeSg (subsetIr _ _)) //. +exact: pnat_coprime (pcore_pgroup _ _) s'E'. +Qed. + +End Section12. + +Implicit Arguments tau2'1 [[gT] [M] x]. +Implicit Arguments tau3'1 [[gT] [M] x]. +Implicit Arguments tau3'2 [[gT] [M] x]. + diff --git a/mathcomp/odd_order/BGsection13.v b/mathcomp/odd_order/BGsection13.v new file mode 100644 index 0000000..be68f9d --- /dev/null +++ b/mathcomp/odd_order/BGsection13.v @@ -0,0 +1,1116 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import bigop finset prime fingroup morphism perm automorphism quotient. +Require Import action gproduct gfunctor pgroup cyclic center commutator. +Require Import gseries nilpotent sylow abelian maximal hall frobenius. +Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +Require Import BGsection7 BGsection9 BGsection10 BGsection12. + +(******************************************************************************) +(* This file covers B & G, section 13; the title subject of the section, *) +(* prime and regular actions, was defined in the frobenius.v file. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Section Section13. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types p q q_star r : nat. +Implicit Types A E H K L M Mstar N P Q Qstar R S T U V W X Y Z : {group gT}. + +Section OneComplement. + +Variables M E : {group gT}. +Hypotheses (maxM : M \in 'M) (hallE : \sigma(M)^'.-Hall(M) E). + +Let sEM : E \subset M := pHall_sub hallE. +Let sM'E : \sigma(M)^'.-group E := pHall_pgroup hallE. + +(* This is B & G, Lemma 13.1. *) +Lemma Msigma_setI_mmax_central p H : + H \in 'M -> p \in \pi(E) -> p \in \pi(H) -> p \notin \tau1(H) -> + [~: M`_\sigma :&: H, M :&: H] != 1 -> gval H \notin M :^: G -> + [/\ (*a*) forall P, P \subset M :&: H -> p.-group P -> + P \subset 'C(M`_\sigma :&: H), + (*b*) p \notin \tau2(H) + & (*c*) p \in \tau1(M) -> p \in \beta(G)]. +Proof. +move=> maxH piEp piHp t1H'p; set R := [~: _, _] => ntR notMGH. +have [q sMq piH'q]: exists2 q, q \in \sigma(M) & q \in \pi(H^`(1)). + pose q := pdiv #|R|; have q_pr: prime q by rewrite pdiv_prime ?cardG_gt1. + have q_dv : q %| _ := dvdn_trans (pdiv_dvd _) (cardSg _). + exists q; last by rewrite mem_primes q_pr cardG_gt0 q_dv ?commgSS ?subsetIr. + rewrite (pgroupP (pcore_pgroup _ M)) ?q_dv //. + have sR_MsM: R \subset [~: M`_\sigma, M] by rewrite commgSS ?subsetIl. + by rewrite (subset_trans sR_MsM) // commg_subl gFnorm. +have [Y sylY] := Sylow_exists q H^`(1); have [sYH' qY _] := and3P sylY. +have nsHbH: H`_\beta <| H := pcore_normal _ _; have [_ nHbH] := andP nsHbH. +have sYH := subset_trans sYH' (der_sub 1 H); have nHbY := subset_trans sYH nHbH. +have nsHbY_H: H`_\beta <*> Y <| H. + rewrite -{2}(quotientGK nsHbH) -quotientYK ?cosetpre_normal //. + rewrite (char_normal_trans _ (der_normal 1 _)) //= -quotient_der //. + rewrite (nilpotent_Hall_pcore _ (quotient_pHall nHbY sylY)) ?pcore_char //. + exact: Mbeta_quo_nil. +have sYNY: Y \subset 'N_H(Y) by rewrite subsetI sYH normG. +have{nsHbY_H} defH: H`_\beta * 'N_H(Y) = H. + rewrite -(mulSGid sYNY) mulgA -(norm_joinEr nHbY). + rewrite (Frattini_arg _ (pHall_subl _ _ sylY)) ?joing_subr //. + by rewrite join_subG Mbeta_der1. +have ntY: Y :!=: 1 by rewrite -cardG_gt1 (card_Hall sylY) p_part_gt1. +have{ntY} [_] := sigma_Jsub maxM (pi_pgroup qY sMq) ntY. +have maxY_H: H \in 'M(Y) by apply/setIdP. +move/(_ E p H hallE piEp _ maxY_H notMGH) => b'p_t3Hp. +case t2Hp: (p \in \tau2(H)). + have b'p: p \notin \beta(G) by case/tau2_not_beta: t2Hp. + have rpH: 'r_p(H) = 2 by apply/eqP; case/andP: t2Hp. + have p'Hb: p^'.-group H`_\beta. + rewrite (pi_p'group (pcore_pgroup _ H)) // inE /=. + by rewrite -predI_sigma_beta // negb_and b'p orbT. + case: b'p_t3Hp; rewrite // -(p_rank_p'quotient p'Hb) ?subIset ?nHbH //=. + by rewrite -quotientMidl defH p_rank_p'quotient ?rpH. +have [S sylS] := Sylow_exists p H; have [sSH pS _] := and3P sylS. +have sSH': S \subset H^`(1). + have [sHp | sH'p] := boolP (p \in \sigma(H)). + apply: subset_trans (Msigma_der1 maxH). + by rewrite (sub_Hall_pcore (Msigma_Hall _)) // (pi_pgroup pS). + have sH'_S: \sigma(H)^'.-group S by rewrite (pi_pgroup pS). + have [F hallF sSF] := Hall_superset (mmax_sol maxH) sSH sH'_S. + have t3Hp: p \in \tau3(H). + have:= partition_pi_sigma_compl maxH hallF p. + by rewrite (pi_sigma_compl hallF) inE /= sH'p piHp (negPf t1H'p) t2Hp. + have [[F1 hallF1] [F3 hallF3]] := ex_tau13_compl hallF. + have [F2 _ complFi] := ex_tau2_compl hallF hallF1 hallF3. + have [[sF3F' nsF3F] _ _ _ _] := sigma_compl_context maxH complFi. + apply: subset_trans (subset_trans sF3F' (dergS 1 (pHall_sub hallF))). + by rewrite (sub_normal_Hall hallF3) ?(pi_pgroup pS). +have sylS_H' := pHall_subl sSH' (der_sub 1 H) sylS. +split=> // [P sPMH pP | t1Mp]; last first. + apply/idPn=> b'p; have [_ /(_ t1Mp)/negP[]] := b'p_t3Hp b'p. + have p'Hb: p^'.-group H`_\beta. + rewrite (pi_p'group (pcore_pgroup _ H)) // inE /=. + by rewrite -predI_sigma_beta // negb_and b'p orbT. + rewrite -p_rank_gt0 -(p_rank_p'quotient p'Hb) ?comm_subG ?subIset ?nHbH //=. + rewrite quotient_der ?subIset ?nHbH // -quotientMidl defH -quotient_der //=. + rewrite p_rank_p'quotient ?comm_subG // -(rank_Sylow sylS_H'). + by rewrite (rank_Sylow sylS) p_rank_gt0. +have nsHaH: H`_\alpha <| H := pcore_normal _ _; have [_ nHaH] := andP nsHaH. +have [sPM sPH] := subsetIP sPMH; have nHaS := subset_trans sSH nHaH. +have nsHaS_H: H`_\alpha <*> S <| H. + rewrite -{2}(quotientGK nsHaH) -quotientYK ?cosetpre_normal //. + rewrite (char_normal_trans _ (der_normal 1 _)) //= -quotient_der //. + rewrite (nilpotent_Hall_pcore _ (quotient_pHall nHaS sylS_H')) ?pcore_char //. + exact: Malpha_quo_nil. +have [sHaS_H nHaS_H] := andP nsHaS_H. +have sP_HaS: P \subset H`_\alpha <*> S. + have [x Hx sPSx] := Sylow_subJ sylS sPH pP; apply: subset_trans sPSx _. + by rewrite sub_conjg (normsP nHaS_H) ?groupV ?joing_subr. +have coHaS_Ms: coprime #|H`_\alpha <*> S| #|M`_\sigma|. + rewrite (p'nat_coprime _ (pcore_pgroup _ _)) // -pgroupE norm_joinEr //. + rewrite pgroupM andbC (pi_pgroup pS) ?(pnatPpi (pHall_pgroup hallE)) //=. + apply: sub_pgroup (pcore_pgroup _ _) => r aHr. + have [|_ ti_aH_sM _] := sigma_disjoint maxH maxM; first by rewrite orbit_sym. + by apply: contraFN (ti_aH_sM r) => sMr; apply/andP. +rewrite (sameP commG1P trivgP) -(coprime_TIg coHaS_Ms) commg_subI ?setIS //. +by rewrite subsetI sP_HaS (subset_trans sPM) ?gFnorm. +Qed. + +(* This is B & G, Corollary 13.2. *) +Corollary cent_norm_tau13_mmax p P H : + (p \in \tau1(M)) || (p \in \tau3(M)) -> + P \subset M -> p.-group P -> H \in 'M('N(P)) -> + [/\ (*a*) forall P1, P1 \subset M :&: H -> p.-group P1 -> + P1 \subset 'C(M`_\sigma :&: H), + (*b*) forall X, X \subset E :&: H -> \tau1(H)^'.-group X -> + X \subset 'C(M`_\sigma :&: H) + & (*c*) [~: M`_\sigma :&: H, M :&: H] != 1 -> + p \in \sigma(H) /\ (p \in \tau1(M) -> p \in \beta(H))]. +Proof. +move=> t13Mp sPM pP /setIdP[maxH sNP_H]. +have ntP: P :!=: 1. + by apply: contraTneq sNP_H => ->; rewrite norm1 proper_subn ?mmax_proper. +have st2Hp: (p \in \sigma(H)) || (p \in \tau2(H)). + exact: (prime_class_mmax_norm maxH pP sNP_H). +have not_MGH: gval H \notin M :^: G. + apply: contraL st2Hp => /imsetP[x _ ->]; rewrite sigmaJ tau2J negb_or. + by have:= t13Mp; rewrite -2!andb_orr !inE => /and3P[-> /eqP->]. +set R := [~: _, _]; have [/commG1P | ntR] := altP (R =P 1). + rewrite centsC => cMH; split=> // X sX_EH _; apply: subset_trans cMH => //. + by rewrite (subset_trans sX_EH) ?setSI. +have piEp: p \in \pi(E). + by rewrite (partition_pi_sigma_compl maxM) // orbCA t13Mp orbT. +have piHp: p \in \pi(H). + by rewrite (partition_pi_mmax maxH) orbCA orbC -!orbA st2Hp !orbT. +have t1H'p: p \notin \tau1(H). + by apply: contraL st2Hp; rewrite negb_or !inE => /and3P[-> /eqP->]. +case: (Msigma_setI_mmax_central maxH piEp) => // cMsH t2H'p b_p. +split=> // [X sX_EH t1'X | _]. + have [sXE sXH] := subsetIP sX_EH. + rewrite -(Sylow_gen X) gen_subG; apply/bigcupsP=> Q; case/SylowP=> q _ sylQ. + have [-> | ntQ] := eqsVneq Q 1; first exact: sub1G. + have piXq: q \in \pi(X) by rewrite -p_rank_gt0 -(rank_Sylow sylQ) rank_gt0. + have [[piEq piHq] t1H'q] := (piSg sXE piXq, piSg sXH piXq, pnatPpi t1'X piXq). + have [sQX qQ _] := and3P sylQ; have sXM := subset_trans sXE sEM. + case: (Msigma_setI_mmax_central maxH piEq) => // -> //. + by rewrite subsetI !(subset_trans sQX). +rewrite (negPf t2H'p) orbF in st2Hp. +by rewrite -predI_sigma_beta // {3}/in_mem /= st2Hp. +Qed. + +(* This is B & G, Corollary 13.3(a). *) +Lemma cyclic_primact_Msigma p P : + p.-Sylow(E) P -> cyclic P -> semiprime M`_\sigma P. +Proof. +move=> sylP cycP x /setD1P[]; rewrite -cycle_eq1 -cycle_subG => ntX sXP. +have [sPE pP _] := and3P sylP; rewrite -cent_cycle. +have sPM := subset_trans sPE sEM; have sXM := subset_trans sXP sPM. +have pX := pgroupS sXP pP; have ltXG := mFT_pgroup_proper pX. +have t13p: (p \in \tau1(M)) || (p \in \tau3(M)). + rewrite (tau1E maxM hallE) (tau3E maxM hallE) -p_rank_gt0 -(rank_Sylow sylP). + rewrite eqn_leq rank_gt0 (subG1_contra sXP) //= andbT -andb_orl orNb. + by rewrite -abelian_rank1_cyclic ?cyclic_abelian. +have [H maxNH] := mmax_exists (mFT_norm_proper ntX ltXG). +have [cMsX _ _] := cent_norm_tau13_mmax t13p sXM pX maxNH. +have [_ sNH] := setIdP maxNH. +apply/eqP; rewrite eqEsubset andbC setIS ?centS // subsetI subsetIl /= centsC. +apply: subset_trans (cMsX P _ pP) (centS _). + rewrite subsetI sPM (subset_trans (cents_norm _) sNH) //. + by rewrite sub_abelian_cent // cyclic_abelian. +by rewrite setIS // (subset_trans (cent_sub _) sNH). +Qed. + +(* This is B & G, Corollary 13.3(b). *) +Corollary tau3_primact_Msigma E3 : + \tau3(M).-Hall(E) E3 -> semiprime M`_\sigma E3. +Proof. +move=> hallE3 x /setD1P[]; rewrite -cycle_eq1 -cycle_subG => ntX sXE3. +have [sE3E t3E3 _] := and3P hallE3; rewrite -cent_cycle. +have [[E1 hallE1] _] := ex_tau13_compl hallE. +have [E2 _ complEi] := ex_tau2_compl hallE hallE1 hallE3. +have [[sE3E' nsE3E] _ [_ cycE3] _ _] := sigma_compl_context maxM complEi. +apply/eqP; rewrite eqEsubset andbC setIS ?centS // subsetI subsetIl /= centsC. +pose p := pdiv #[x]; have p_pr: prime p by rewrite pdiv_prime ?cardG_gt1. +have t3p: p \in \tau3(M) by rewrite (pgroupP (pgroupS sXE3 t3E3)) ?pdiv_dvd. +have t13p: [|| p \in \tau1(M) | p \in \tau3(M)] by rewrite t3p orbT. +have [y Xy oy]:= Cauchy p_pr (pdiv_dvd _). +have ntY: <[y]> != 1 by rewrite -cardG_gt1 -orderE oy prime_gt1. +have pY: p.-group <[y]> by rewrite /pgroup -orderE oy pnat_id. +have [H maxNH] := mmax_exists (mFT_norm_proper ntY (mFT_pgroup_proper pY)). +have sYE3: <[y]> \subset E3 by rewrite cycle_subG (subsetP sXE3). +have sYE := subset_trans sYE3 sE3E; have sYM := subset_trans sYE sEM. +have [_ cMsY _] := cent_norm_tau13_mmax t13p sYM pY maxNH. +have [_ sNH] := setIdP maxNH. +have sE3H': E3 \subset H^`(1). + rewrite (subset_trans sE3E') ?dergS // (subset_trans _ sNH) ?normal_norm //. + by rewrite (char_normal_trans _ nsE3E) // sub_cyclic_char. +apply: subset_trans (cMsY E3 _ _) (centS _). +- rewrite subsetI sE3E (subset_trans (cents_norm _) sNH) //. + by rewrite sub_abelian_cent ?cyclic_abelian. +- rewrite (pgroupS sE3H') //; apply/pgroupP=> q _ q_dv_H'. + by rewrite !inE q_dv_H' !andbF. +by rewrite setIS // (subset_trans _ sNH) // cents_norm ?centS ?cycle_subG. +Qed. + +(* This is B & G, Theorem 13.4. *) +(* Note how the non-structural steps in the proof (top of p. 99, where it is *) +(* deduced that C_M_alpha(P) <= C_M_alpha(R) from q \notin \alpha, and then *) +(* C_M_alpha(P) = C_M_alpha(R) from r \in tau_1(M) !!), are handled cleanly *) +(* on lines 5-12 of the proof by a conditional expression for the former, and *) +(* a without loss tactic for the latter. *) +(* Also note that the references to 10.12 and 12.2 are garbled (some are *) +(* missing, and some are exchanged!). *) +Theorem cent_tau1Elem_Msigma p r P R (Ms := M`_\sigma) : + p \in \tau1(M) -> P \in 'E_p^1(E) -> R \in 'E_r^1('C_E(P)) -> + 'C_Ms(P) \subset 'C_Ms(R). +Proof. +have /andP[sMsM nMsM]: Ms <| M := pcore_normal _ M. +have coMsE: coprime #|Ms| #|E| := coprime_sigma_compl hallE. +pose Ma := M`_\alpha; have sMaMs: Ma \subset Ms := Malpha_sub_Msigma maxM. +rewrite pnElemI -setIdE => t1Mp EpP /setIdP[ErR cPR]. +without loss symPR: p r P R EpP ErR cPR t1Mp / + r \in \tau1(M) -> 'C_Ma(P) \subset 'C_Ma(R) -> 'C_Ma(P) = 'C_Ma(R). +- move=> IH; apply: (IH p r) => // t1Mr sCaPR; apply/eqP; rewrite eqEsubset. + rewrite sCaPR -(setIidPl sMaMs) -!setIA setIS ?(IH r p) 1?centsC // => _. + by case/eqVproper; rewrite // /proper sCaPR andbF. +do [rewrite !subsetI !subsetIl /=; set cRCaP := _ \subset _] in symPR *. +pose Mz := 'O_(if cRCaP then \sigma(M) else \alpha(M))(M); pose C := 'C_Mz(P). +suffices: C \subset 'C(R) by rewrite /C /Mz /cRCaP; case: ifP => // ->. +have sMzMs: Mz \subset Ms by rewrite /Mz; case: ifP => // _. +have sCMs: C \subset Ms by rewrite subIset ?sMzMs. +have [[sPE abelP dimP] [sRE abelR dimR]] := (pnElemP EpP, pnElemP ErR). +have [sPM sRM] := (subset_trans sPE sEM, subset_trans sRE sEM). +have [[pP cPP _] [rR _]] := (and3P abelP, andP abelR). +have coCR: coprime #|C| #|R| := coprimeSg sCMs (coprimegS sRE coMsE). +have ntP: P :!=: 1 by exact: nt_pnElem EpP _. +pose ST := [set S | Sylow C (gval S) & R \subset 'N(S)]. +have sST_CP S: S \in ST -> S \subset C by case/setIdP=> /SylowP[q _ /andP[]]. +rewrite -{sST_CP}[C](Sylow_transversal_gen sST_CP) => [|q _]; last first. + have nMzR: R \subset 'N(Mz) by rewrite (subset_trans sRM) // gFnorm. + have{nMzR} nCR: R \subset 'N(C) by rewrite normsI // norms_cent // cents_norm. + have solC := solvableS (subset_trans sCMs sMsM) (mmax_sol maxM). + have [S sylS nSR] := coprime_Hall_exists q nCR coCR solC. + by exists S; rewrite // inE (p_Sylow sylS) nSR. +rewrite gen_subG; apply/bigcupsP=> S /setIdP[/SylowP[q _ sylS] nSR] {ST}. +have [sSC qS _] := and3P sylS. +have [sSMs [sSMz cPS]] := (subset_trans sSC sCMs, subsetIP sSC). +rewrite (sameP commG1P eqP) /=; set Q := [~: S, R]; apply/idPn => ntQ. +have sQS: Q \subset S by [rewrite commg_subl]; have qQ := pgroupS sQS qS. +have piQq: q \in \pi(Q) by rewrite -p_rank_gt0 -(rank_pgroup qQ) rank_gt0. +have{piQq} [nQR piSq] := (commg_normr R S : R \subset 'N(Q), piSg sQS piQq). +have [H maxNH] := mmax_exists (mFT_norm_proper ntP (mFT_pgroup_proper pP)). +have [maxH sNH] := setIdP maxNH; have sCPH := subset_trans (cent_sub _) sNH. +have [sPH sRH] := (subset_trans cPP sCPH, subset_trans cPR sCPH). +have [sSM sSH] := (subset_trans sSMs sMsM, subset_trans cPS sCPH). +have [sQM sQH] := (subset_trans sQS sSM, subset_trans sQS sSH). +have ntMsH_R: [~: Ms :&: H, R] != 1. + by rewrite (subG1_contra _ ntQ) ?commSg // subsetI sSMs. +have sR_EH: R \subset E :&: H by apply/subsetIP. +have ntMsH_MH: [~: Ms :&: H, M :&: H] != 1. + by rewrite (subG1_contra _ ntMsH_R) ?commgS // (subset_trans sR_EH) ?setSI. +have t13Mp: p \in [predU \tau1(M) & \tau3(M)] by apply/orP; left. +have [_ cMsH_t1H' [//|sHp bHp]] := cent_norm_tau13_mmax t13Mp sPM pP maxNH. +have{cMsH_t1H'} t1Hr: r \in \tau1(H). + apply: contraR ntMsH_R => t1H'r; rewrite (sameP eqP commG1P) centsC. + by rewrite cMsH_t1H' // (pi_pgroup rR). +have ntCHaRQ: 'C_(H`_\alpha)(R <*> Q) != 1. + rewrite centY (subG1_contra _ ntP) ?subsetI //= centsC cPR centsC. + rewrite (subset_trans sQS cPS) (subset_trans _ (Mbeta_sub_Malpha H)) //. + by rewrite (sub_Hall_pcore (Mbeta_Hall maxH)) // (pi_pgroup pP) ?bHp. +have not_MGH: gval H \notin M :^: G. + by apply: contraL sHp => /imsetP[x _ ->]; rewrite sigmaJ; case/andP: t1Mp. +have neqHM: H :!=: M := contra_orbit _ _ not_MGH. +have cSS: abelian S. + apply: contraR neqHM => /(nonabelian_Uniqueness qS)uniqS. + by rewrite (eq_uniq_mmax (def_uniq_mmax uniqS maxM sSM) maxH sSH). +have tiQcR: 'C_Q(R) = 1 by rewrite coprime_abel_cent_TI // (coprimeSg sSC). +have sMq: q \in \sigma(M) := pnatPpi (pcore_pgroup _ M) (piSg sSMs piSq). +have aH'q: q \notin \alpha(H). + have [|_ tiHaMs _] := sigma_disjoint maxH maxM; first by rewrite orbit_sym. + by apply: contraFN (tiHaMs q) => aHq; apply/andP. +have piRr: r \in \pi(R) by rewrite -p_rank_gt0 p_rank_abelem ?dimR. +have ErH_R: R \in 'E_r^1(H) by rewrite !inE sRH abelR dimR. +have{piRr} sM'r: r \in \sigma(M)^' := pnatPpi (pgroupS sRE sM'E) piRr. +have r'q: q \in r^' by apply: contraTneq sMq => ->. +have ntHa: H`_\alpha != 1 by rewrite (subG1_contra _ ntCHaRQ) ?subsetIl. +have uniqNQ: 'M('N(Q)) = [set H]. + apply: contraNeq ntCHaRQ; rewrite joingC. + by case/(cent_Malpha_reg_tau1 _ _ r'q ErH_R) => //; case=> //= _ -> _. +have maxNQ_H: H \in 'M('N(Q)) :\ M by rewrite uniqNQ !inE neqHM /=. +have{maxNQ_H} [_ _] := sigma_subgroup_embedding maxM sMq sQM qQ ntQ maxNQ_H. +have [sHq [_ st1HM [_ ntMa]] | _ [_ _ sM'MH]] := ifP; last first. + have piPp: p \in \pi(P) by rewrite -p_rank_gt0 p_rank_abelem ?dimP. + have sPMH: P \subset M :&: H by apply/subsetIP. + by have/negP[] := pnatPpi (pgroupS sPMH (pHall_pgroup sM'MH)) piPp. +have{st1HM} t1Mr: r \in \tau1(M). + by case/orP: (st1HM r t1Hr); rewrite //= (contraNF ((alpha_sub_sigma _) r)). +have aM'q: q \notin \alpha(M). + have [_ tiMaHs _] := sigma_disjoint maxM maxH not_MGH. + by apply: contraFN (tiMaHs q) => aMq; apply/andP. +have ErM_R: R \in 'E_r^1(M) by rewrite !inE sRM abelR dimR. +have: 'M('N(Q)) != [set M] by rewrite uniqNQ (inj_eq (@set1_inj _)). +case/(cent_Malpha_reg_tau1 _ _ r'q ErM_R) => //. +case=> //= ntCMaP tiCMaPQ _; case/negP: ntCMaP. +rewrite -subG1 -{}tiCMaPQ centY setICA subsetIidr /= -/Ma -/Q centsC. +apply/commG1P/three_subgroup; apply/commG1P. + by rewrite commGC (commG1P _) ?sub1G ?subsetIr. +apply: subset_trans (subsetIr Ma _); rewrite /= -symPR //. + rewrite commg_subl normsI //; last by rewrite norms_cent // cents_norm. + by rewrite (subset_trans sSM) ?gFnorm. +apply: contraR aM'q => not_cRCaP; apply: pnatPpi (pgroupS sSMz _) piSq. +by rewrite (negPf not_cRCaP) pcore_pgroup. +Qed. + +(* This is B & G, Theorem 13.5. *) +Theorem tau1_primact_Msigma E1 : \tau1(M).-Hall(E) E1 -> semiprime M`_\sigma E1. +Proof. +move=> hallE1 x /setD1P[]; rewrite -cycle_eq1 -cycle_subG => ntX sXE1. +rewrite -cent_cycle; have [sE1E t1E1 _] := and3P hallE1. +have [_ [E3 hallE3]] := ex_tau13_compl hallE. +have{hallE3} [E2 _ complEi] := ex_tau2_compl hallE hallE1 hallE3. +have [_ _ [cycE1 _] _ _ {E2 E3 complEi}] := sigma_compl_context maxM complEi. +apply/eqP; rewrite eqEsubset andbC setIS ?centS //= subsetI subsetIl /=. +have [p _ rX] := rank_witness <[x]>; rewrite -rank_gt0 {}rX in ntX. +have t1p: p \in \tau1(M) by rewrite (pnatPpi t1E1) // (piSg sXE1) -?p_rank_gt0. +have{ntX} [P EpP] := p_rank_geP ntX; have{EpP} [sPX abelP dimP] := pnElemP EpP. +have{sXE1} sPE1 := subset_trans sPX sXE1. +have{dimP} EpP: P \in 'E_p^1(E) by rewrite !inE abelP dimP (subset_trans sPE1). +apply: {x sPX abelP} subset_trans (setIS _ (centS sPX)) _; rewrite centsC. +rewrite -(Sylow_gen E1) gen_subG; apply/bigcupsP=> S; case/SylowP=> r _ sylS. +have [[sSE1 rS _] [-> | ntS]] := (and3P sylS, eqsVneq S 1); first exact: sub1G. +have [cycS sSE] := (cyclicS sSE1 cycE1, subset_trans sSE1 sE1E). +have /p_rank_geP[R ErR]: 0 < 'r_r(S) by rewrite -rank_pgroup ?rank_gt0. +have{ErR} [sRS abelR dimR] := pnElemP ErR; have sRE1 := subset_trans sRS sSE1. +have{abelR dimR} ErR: R \in 'E_r^1('C_E(P)). + rewrite !inE abelR dimR (subset_trans sRE1) // subsetI sE1E. + by rewrite sub_abelian_cent ?cyclic_abelian. +rewrite centsC (subset_trans (cent_tau1Elem_Msigma t1p EpP ErR)) //. +have [y defR]: exists y, R :=: <[y]> by apply/cyclicP; exact: cyclicS cycS. +have sylS_E: r.-Sylow(E) S. + apply: subHall_Sylow hallE1 (pnatPpi t1E1 _) (sylS). + by rewrite -p_rank_gt0 -(rank_Sylow sylS) rank_gt0. +rewrite defR cent_cycle (cyclic_primact_Msigma sylS_E cycS) ?subsetIr //. +by rewrite !inE -cycle_subG -cycle_eq1 -defR (nt_pnElem ErR). +Qed. + +(* This is B & G, Lemma 13.6. *) +(* The wording at the top of the proof is misleading: it should say: by *) +(* Corollary 12.14, it suffices to show that we can't have both q \in beta(M) *) +(* and X \notsubset M_sigma^(1). Also, the reference to 12.13 should be 12.19 *) +(* or 10.9 (we've used 12.19 here). *) +Lemma cent_cent_Msigma_tau1_uniq E1 P q X (Ms := M`_\sigma) : + \tau1(M).-Hall(E) E1 -> P \subset E1 -> P :!=: 1 -> + X \in 'E_q^1('C_Ms(P)) -> + 'M('C(X)) = [set M] /\ (forall S, q.-Sylow(Ms) S -> 'M(S) = [set M]). +Proof. +move=> hallE1 sPE1 ntP EqX; have [sE1E t1E1 _] := and3P hallE1. +rewrite (cent_semiprime (tau1_primact_Msigma hallE1)) //= -/Ms in EqX. +have{P ntP sPE1} ntE1 := subG1_contra sPE1 ntP. +have /andP[sMsM nMsM]: Ms <| M := pcore_normal _ M. +have coMsE: coprime #|Ms| #|E| := coprime_sigma_compl hallE. +have [solMs nMsE] := (solvableS sMsM (mmax_sol maxM), subset_trans sEM nMsM). +apply: cent_der_sigma_uniq => //. + by move: EqX; rewrite -(setIidPr sMsM) -setIA pnElemI => /setIP[]. +have{EqX} [[sXcMsE1 abelX _] ntX] := (pnElemP EqX, nt_pnElem EqX isT). +apply: contraR ntX => /norP[b'q not_sXMs']; rewrite -subG1. +have [S sylS nSE] := coprime_Hall_exists q nMsE coMsE solMs. +have{abelX} [[sSMs qS _] [qX _]] := (and3P sylS, andP abelX). +have sScMsE': S \subset 'C_Ms(E^`(1)). + have [H hallH cHE'] := der_compl_cent_beta' maxM hallE. + have [Q sylQ] := Sylow_exists q H; have [sQH qQ _] := and3P sylQ. + have{cHE' sQH} cQE' := centsS sQH cHE'; have sE'E := der_sub 1 E. + have [nMsE' coMsE'] := (coprimegS sE'E coMsE, subset_trans sE'E nMsE). + have{H hallH sylQ} sylQ := subHall_Sylow hallH b'q sylQ. + have nSE' := subset_trans sE'E nSE; have nQE' := cents_norm cQE'. + have [x cE'x ->] := coprime_Hall_trans coMsE' nMsE' solMs sylS nSE' sylQ nQE'. + by rewrite conj_subG // subsetI (pHall_sub sylQ) centsC. +without loss{qX} sXS: X sXcMsE1 not_sXMs' / X \subset S. + have [nMsE1 coMsE1 IH] := (subset_trans sE1E nMsE, coprimegS sE1E coMsE). + have [nSE1 [sXMs cE1X]] := (subset_trans sE1E nSE, subsetIP sXcMsE1). + have [|Q [sylQ nQE1 sXQ]] := coprime_Hall_subset nMsE1 coMsE1 solMs sXMs qX. + by rewrite cents_norm // centsC. + have [//|x cE1x defS] := coprime_Hall_trans nMsE1 _ solMs sylS nSE1 sylQ nQE1. + have Ms_x: x \in Ms by case/setIP: cE1x. + rewrite -(conjs1g x^-1) -sub_conjg IH //; last by rewrite defS conjSg. + by rewrite -(conjGid cE1x) conjSg. + by rewrite -(normsP (der_norm 1 _) x) ?conjSg. +have [sXMs cE1X] := subsetIP sXcMsE1. +have [_ [E3 hallE3]] := ex_tau13_compl hallE. +have{hallE3} [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. +have{not_sXMs' E3 complEi} ntE2: E2 :!=: 1. + apply: contraNneq not_sXMs' => E2_1. + have [[sE3E' _] _ _ [defE _] _] := sigma_compl_context maxM complEi. + have [sCMsE_Ms' _ _] := sigma_compl_embedding maxM hallE. + have{defE} [_ defE _ _] := sdprodP defE; rewrite E2_1 sdprod1g in defE. + rewrite (subset_trans _ sCMsE_Ms') // -defE centM setIA subsetI. + by rewrite (subset_trans (subset_trans sXS sScMsE')) ?setIS ?centS. +have{ntE2 E2 hallE2} [p p_pr t2p]: exists2 p, prime p & p \in \tau2(M). + have [[p p_pr rE2] [_ t2E2 _]] := (rank_witness E2, and3P hallE2). + by exists p; rewrite ?(pnatPpi t2E2) // -p_rank_gt0 -rE2 rank_gt0. +have [A Ep2A Ep2A_M] := ex_tau2Elem hallE t2p. +have [_ _ tiCMsA _ _] := tau2_context maxM t2p Ep2A_M. +have [[nsAE _] _ _ _] := tau2_compl_context maxM hallE t2p Ep2A. +have [sAE abelA _] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. +have cCAE1_X: X \subset 'C('C_A(E1)). + rewrite centsC; apply/subsetP=> y; case/setIP=> Ay cE1y. + have [-> | nty] := eqVneq y 1; first exact: group1. + have oY: #[y] = p := abelem_order_p abelA Ay nty. + have [r _ rE1] := rank_witness E1. + have{rE1} rE1: 'r_r(E1) > 0 by rewrite -rE1 rank_gt0. + have [R ErR] := p_rank_geP rE1; have{ErR} [sRE1 abelR dimR] := pnElemP ErR. + have t1r: r \in \tau1(M) by rewrite (pnatPpi t1E1) -?p_rank_gt0. + have ErR: R \in 'E_r^1(E) by rewrite !inE abelR dimR (subset_trans sRE1). + have EpY: <[y]>%G \in 'E_p^1('C_E(R)). + rewrite p1ElemE // !inE -oY eqxx subsetI (centsS sRE1) cycle_subG //=. + by rewrite (subsetP sAE). + rewrite -sub_cent1 -cent_cycle (subset_trans sXcMsE1) //. + apply: subset_trans (setIS _ (centS sRE1)) _. + rewrite -subsetIidl setIAC subsetI subsetIr andbT. + exact: cent_tau1Elem_Msigma t1r ErR EpY. +have nAE1 := subset_trans sE1E (normal_norm nsAE). +have coAE1: coprime #|A| #|E1|. + by apply: pnat_coprime pA (pi_p'group t1E1 (contraL _ t2p)); apply: tau2'1. +rewrite -tiCMsA -(coprime_cent_prod nAE1 coAE1) ?abelian_sol // centM setIA. +rewrite subsetI cCAE1_X (subset_trans (subset_trans sXS sScMsE')) ?setIS //. +by rewrite centS ?commgSS. +Qed. + +End OneComplement. + +(* This is B & G, Lemma 13.7. *) +(* We've had to plug a gap in this proof: on p. 100, l. 6-7 it is asserted *) +(* the conclusion (E1 * E3 acts in a prime manner on M_sigma) follows from *) +(* the fact that E1 and E3 have coprime orders and act in a prime manner with *) +(* the same set of fixed points. This seems to imply the following argument: *) +(* For any x \in M_sigma, *) +(* C_(E1 * E3)[x] = C_E1[x] * C_E3[x] is either E1 * E3 or 1, *) +(* i.e., E1 * E3 acts in a prime manner on M_sigma. *) +(* This is improper because the distributivity of centralisers over coprime *) +(* products only hold under normality conditions that do not hold in this *) +(* instance. The correct argument, which involves using the prime action *) +(* assumption a second time, only relies on the fact that E1 and E3 are Hall *) +(* subgroups of the group E1 * E3. The fact that E3 <| E (Lemma 12.1(a)), *) +(* implicitly needed to justify that E1 * E3 is a group, can also be used to *) +(* simplify the argument, and we do so. *) +Lemma tau13_primact_Msigma M E E1 E2 E3 : + M \in 'M -> sigma_complement M E E1 E2 E3 -> ~ semiregular E3 E1 -> + semiprime M`_\sigma (E3 <*> E1). +Proof. +move=> maxM complEi not_regE13; set Ms := M`_\sigma. +have [hallE hallE1 hallE2 hallE3 _] := complEi. +have [[sE1E t1E1 _] [sE3E t3E3 _]] := (and3P hallE1, and3P hallE3). +have [[sEM _ _] [_ t2E2 _]] := (and3P hallE, and3P hallE2). +have [[_ nsE3E] _ [_ cycE3] [defE _] tiE3cE]:= sigma_compl_context maxM complEi. +have [[_ nE3E] [sMsM nMsM]] := (andP nsE3E, andP (pcore_normal _ M : Ms <| M)). +have [P]: exists2 P, P \in 'E^1(E1) & 'C_E3(P) != 1. + apply/exists_inP; rewrite -negb_forall_in; apply/forall_inP=> regE13. + apply: not_regE13 => x /setD1P[]; rewrite -cycle_subG -cycle_eq1 -rank_gt0. + case/rank_geP=> P E1xP sXE1; apply/trivgP; move: E1xP. + rewrite /= -(setIidPr sXE1) nElemI -setIdE => /setIdP[E1_P sPX]. + by rewrite -(eqP (regE13 P E1_P)) -cent_cycle setIS ?centS. +rewrite -{1}(setIidPr sE1E) nElemI -setIdE => /setIdP[/nElemP[p EpP] sPE1]. +rewrite -{1}(setIidPl sE3E) -setIA setIC -rank_gt0 => /rank_geP[R]. +rewrite nElemI -setIdE => /setIdP[/nElemP[r ErR] sRE3]. +have t1p: p \in \tau1(M). + rewrite (pnatPpi (pgroupS sPE1 t1E1)) //= (card_p1Elem EpP). + by rewrite pi_of_prime ?inE ?(pnElem_prime EpP) //. +have prE1 := tau1_primact_Msigma maxM hallE hallE1. +have prE3 := tau3_primact_Msigma maxM hallE hallE3. +have sCsPR: 'C_Ms(P) \subset 'C_Ms(R) by apply: cent_tau1Elem_Msigma EpP ErR. +have [eqCsPR | ltCsPR] := eqVproper sCsPR. + move=> x; case/setD1P; rewrite -cycle_eq1 -cycle_subG => ntX sXE31. + apply/eqP; rewrite -cent_cycle eqEsubset andbC setIS ?centS //=. + have eqCsE13: 'C_Ms(E1) = 'C_Ms(E3). + rewrite -(cent_semiprime prE1 sPE1) ?(nt_pnElem EpP) //. + by rewrite -(cent_semiprime prE3 sRE3) ?(nt_pnElem ErR). + rewrite centY setICA eqCsE13 -setICA setIid. + have sE31E: E3 <*> E1 \subset E by apply/joing_subP. + have nE3E1 := subset_trans sE1E nE3E. + pose y := x.`_\tau1(M); have sYX: <[y]> \subset <[x]> := cycleX x _. + have sXE := subset_trans sXE31 sE31E; have sYE := subset_trans sYX sXE. + have [t1'x | not_t1'x] := boolP (\tau1(M)^'.-elt x). + rewrite (cent_semiprime prE3 _ ntX) // (sub_normal_Hall hallE3) //. + apply: pnat_dvd t3E3; rewrite -(Gauss_dvdr _ (p'nat_coprime t1'x t1E1)). + by rewrite mulnC (dvdn_trans _ (dvdn_cardMg _ _)) -?norm_joinEr ?cardSg. + have{not_t1'x} ntY: #[y] != 1%N by rewrite order_constt partn_eq1. + apply: subset_trans (setIS _ (centS sYX)) _. + have [solE nMsE] := (sigma_compl_sol hallE, subset_trans sEM nMsM). + have [u Eu sYuE1] := Hall_Jsub solE hallE1 sYE (p_elt_constt _ _). + rewrite -(conjSg _ _ u) !conjIg -!centJ (normsP nMsE) ?(normsP nE3E) //=. + by rewrite -eqCsE13 (cent_semiprime prE1 sYuE1) // trivg_card1 cardJg. +have ntCsR: 'C_Ms(R) != 1. + by rewrite -proper1G (sub_proper_trans _ ltCsPR) ?sub1G. +have ntR: R :!=: 1 by rewrite (nt_pnElem ErR). +have [cEPR abelR dimR] := pnElemP ErR; have [rR _ _] := and3P abelR. +have{cEPR} [sRE cPR] := subsetIP cEPR; have sRM := subset_trans sRE sEM. +have E2_1: E2 :=: 1. + have [x defR] := cyclicP (cyclicS sRE3 cycE3). + apply: contraNeq ntCsR; rewrite -rank_gt0; have [q _ ->] := rank_witness E2. + rewrite p_rank_gt0 defR cent_cycle. move/(pnatPpi t2E2) => t2q. + have [A Eq2A _] := ex_tau2Elem hallE t2q. + have [-> //] := tau2_regular maxM complEi t2q Eq2A. + by rewrite !inE -cycle_subG -cycle_eq1 -defR sRE3 (nt_pnElem ErR). +have nRE: E \subset 'N(R) by rewrite (char_norm_trans _ nE3E) ?sub_cyclic_char. +have [H maxNH] := mmax_exists (mFT_norm_proper ntR (mFT_pgroup_proper rR)). +have [maxH sNH] := setIdP maxNH; have sEH := subset_trans nRE sNH. +have ntCsR_P: [~: 'C_Ms(R), P] != 1. + rewrite (sameP eqP commG1P); apply: contra (proper_subn ltCsPR). + by rewrite subsetI subsetIl. +have sCsR_MsH: 'C_Ms(R) \subset Ms :&: H. + exact: setIS Ms (subset_trans (cent_sub R) sNH). +have ntMsH_P: [~: Ms :&: H, P] != 1 by rewrite (subG1_contra _ ntCsR_P) ?commSg. +have tiE1cMsH: 'C_E1(Ms :&: H) = 1. + apply: contraNeq (proper_subn ltCsPR) => ntCE1MsH. + rewrite (cent_semiprime prE1 sPE1) ?(nt_pnElem EpP) //. + rewrite -(cent_semiprime prE1 (subsetIl _ _) ntCE1MsH) /=. + by rewrite subsetI subsetIl centsC subIset // orbC centS. +have t3r: r \in \tau3(M). + by rewrite (pnatPpi t3E3) ?(piSg sRE3) // -p_rank_gt0 p_rank_abelem ?dimR. +have t13r: (r \in \tau1(M)) || (r \in \tau3(M)) by rewrite t3r orbT. +have [sE1H sRH] := (subset_trans sE1E sEH, subset_trans sRE sEH). +have [_ ct1H'R [|sHr _]] := cent_norm_tau13_mmax maxM hallE t13r sRM rR maxNH. + rewrite (subG1_contra _ ntMsH_P) // commgS // (subset_trans sPE1) //. + by rewrite subsetI (subset_trans sE1E). +have t1H_E1: \tau1(H).-group E1. + apply/pgroupP=> q q_pr /Cauchy[] // x E1x ox. + apply: contraLR (prime_gt1 q_pr) => t1H'q; rewrite -ox cardG_gt1 negbK. + rewrite -subG1 -tiE1cMsH subsetI cycle_subG E1x /= ct1H'R //. + by rewrite (setIidPl sEH) cycle_subG (subsetP sE1E). + by rewrite /pgroup -orderE ox pnatE. +have sH'_E1: \sigma(H)^'.-group E1 by apply: sub_pgroup t1H_E1 => q /andP[]. +have [F hallF sE1F] := Hall_superset (mmax_sol maxH) sE1H sH'_E1. +have [F1 hallF1 sE1F1] := Hall_superset (sigma_compl_sol hallF) sE1F t1H_E1. +have{sHr} sRHs: R \subset H`_\sigma. + by rewrite (sub_Hall_pcore (Msigma_Hall maxH)) ?(pi_pgroup rR). +have cRE1: E1 \subset 'C(R). + rewrite centsC (centsS sE1F1) // -subsetIidl subsetI subxx -sRHs -subsetI. + have prF1 := tau1_primact_Msigma maxH hallF hallF1. + rewrite -(cent_semiprime prF1 (subset_trans sPE1 sE1F1)); last first. + exact: nt_pnElem EpP _. + by rewrite subsetI sRHs. +case/negP: ntR; rewrite -subG1 -tiE3cE subsetI sRE3 centsC -(sdprodW defE). +by rewrite E2_1 sdprod1g mul_subG // sub_abelian_cent ?cyclic_abelian. +Qed. + +(* This is B & G, Lemma 13.8. *) +(* We had to plug a significant hole in the proof text: in the sixth *) +(* paragraph of the proof, it is asserted that because M = N_M(Q)M_alpha and *) +(* r is in pi(C_M(P)), P centralises some non-trivial r-subgroup of N_M(Q). *) +(* This does not seem to follow directly, even taking into account that r is *) +(* not in alpha(M): while it is true that N_M(Q) contains a Sylow r-subgroup *) +(* of M, this subgroup does not need to contain an r-group centralized by P. *) +(* We can only establish the required result by making use of the fact that M *) +(* has a normal p-complement K (because p is in tau_1(M)), as then N_K(Q) *) +(* will contain a p-invariant Sylow r-subgroup S of K and M (coprime action) *) +(* and then any r-subgroup of M centralised by P will be in K, and hence *) +(* conjugate in C_K(P) to a subgroup of S (coprime action again). *) +Lemma tau1_mmaxI_asymmetry M Mstar p P q Q q_star Qstar : + (*1*) [/\ M \in 'M, Mstar \in 'M & gval Mstar \notin M :^: G] -> + (*2*) [/\ p \in \tau1(M), p \in \tau1(Mstar) & P \in 'E_p^1(M :&: Mstar)] -> + (*3*) [/\ q.-Sylow(M :&: Mstar) Q, q_star.-Sylow(M :&: Mstar) Qstar, + P \subset 'N(Q) & P \subset 'N(Qstar)] -> + (*4*) 'C_Q(P) = 1 /\ 'C_Qstar(P) = 1 -> + (*5*) 'N(Q) \subset Mstar /\ 'N(Qstar) \subset M -> + False. +Proof. +move: Mstar q_star Qstar => L u U. (* Abbreviate Mstar by L, Qstar by U. *) +move=> [maxM maxL notMGL] [t1Mp t1Lp EpP] [sylQ sylU nQP nUP]. +move=> [regPQ regPU] [sNQL sNUM]; rewrite setIC in sylU. (* for symmetry *) +have notLGM: gval M \notin L :^: G by rewrite orbit_sym. (* for symmetry *) +have{EpP} [ntP [sPML abelP dimP]] := (nt_pnElem EpP isT, pnElemP EpP). +have{sPML} [[sPM sPL] [pP _ _]] := (subsetIP sPML, and3P abelP). +have solCP: solvable 'C(P) by rewrite mFT_sol ?mFT_cent_proper. +pose Qprops M q Q := [&& q.-Sylow(M) Q, q != p, q \notin \beta(M), + 'C_(M`_\beta)(P) != 1 & 'C_(M`_\beta)(P <*> Q) == 1]. +have{sylQ sylU} [hypQ hypU]: Qprops M q Q /\ Qprops L u U. + apply/andP; apply/nandP=> not_hypQ. + without loss{not_hypQ}: L u U M q Q maxM maxL notMGL notLGM t1Mp t1Lp sPM sPL + sylQ sylU nQP nUP regPQ regPU sNQL sNUM / ~~ Qprops M q Q. + - by move=> IH; case: not_hypQ; [apply: (IH L u U) | apply: (IH M q Q)]. + case/and5P; have [_ qQ _] := and3P sylQ. + have{sylQ} sylQ: q.-Sylow(M) Q. + by rewrite -Sylow_subnorm -(setIidPr sNQL) setIA Sylow_subnorm. + have ntQ: Q :!=: 1. + by apply: contraTneq sNQL => ->; rewrite norm1 proper_subn ?mmax_proper. + have p'q: q != p. + apply: contraNneq ntQ => def_q; rewrite (trivg_center_pgroup qQ) //. + apply/trivgP; rewrite -regPQ setIS // centS //. + by rewrite (norm_sub_max_pgroup (Hall_max sylQ)) ?def_q. + have EpP: P \in 'E_p^1(M) by apply/pnElemP. + have [|_ [// | abM]] := cent_Malpha_reg_tau1 maxM t1Mp p'q EpP ntQ nQP regPQ. + apply: contraNneq notMGL => uniqNQ. + by rewrite (eq_uniq_mmax uniqNQ) ?orbit_refl. + by rewrite joingC /alpha_core abM !(eq_pcore _ abM) => _ -> -> ->. +have bML_CMbP: forall M L, [predU \beta(M) & \beta(L)].-group 'C_(M`_\beta)(P). + move=> ? ?; apply: pgroupS (subsetIl _ _) (sub_pgroup _ (pcore_pgroup _ _)). + by move=> ?; rewrite !inE => ->. +have [H hallH sCMbP_H] := Hall_superset solCP (subsetIr _ _) (bML_CMbP M L). +have [[s _ rFH] [cPH bH _]] := (rank_witness 'F(H), and3P hallH). +have{sCMbP_H rFH cPH} piFHs: s \in \pi('F(H)). + rewrite -p_rank_gt0 -rFH rank_gt0 (trivg_Fitting (solvableS cPH solCP)). + by rewrite (subG1_contra sCMbP_H) //; case/and5P: hypQ. +without loss{bH} bMs: L u U M q Q maxM maxL notMGL notLGM t1Mp t1Lp sPM sPL + nQP nUP regPQ regPU sNQL sNUM hypQ hypU hallH / s \in \beta(M). +- move=> IH; have:= pnatPpi bH (piSg (Fitting_sub H) piFHs). + case/orP; [apply: IH hypQ hypU hallH | apply: IH hypU hypQ _] => //. + by apply: etrans (eq_pHall _ _ _) hallH => ?; apply: orbC. +without loss{bML_CMbP} sCMbP_H: H hallH piFHs / 'C_(M`_\beta)(P) \subset H. + have [x cPx sCMbP_Hx] := Hall_subJ solCP hallH (subsetIr _ _) (bML_CMbP M L). + by move=> IH; apply: IH sCMbP_Hx; rewrite ?pHallJ //= FittingJ cardJg. +pose X := 'O_s(H); have sylX := nilpotent_pcore_Hall s (Fitting_nil H). +have{piFHs sylX} ntX: X != 1. + by rewrite -rank_gt0 /= -p_core_Fitting (rank_Sylow sylX) p_rank_gt0. +have [[cPH bH _] [sXH nXH]] := (and3P hallH, andP (pcore_normal s H : X <| H)). +have [cPX sX] := (subset_trans sXH cPH, pcore_pgroup s H : s.-group X). +have{hypQ} [sylQ p'q bM'q ntCMbP] := and5P hypQ; apply: negP. +apply: subG1_contra (ntX); rewrite /= centY !subsetI (subset_trans _ cPH) //=. +have nsMbM : M`_\beta <| M := pcore_normal _ M; have [sMbM nMbM] := andP nsMbM. +have hallMb := Mbeta_Hall maxM; have [_ bMb _] := and3P hallMb. +have{ntX} sHM: H \subset M. + have [g sXMg]: exists g, X \subset M :^ g. + have [S sylS] := Sylow_exists s M`_\beta; have [sSMb _ _] := and3P sylS. + have sylS_G := subHall_Sylow (Mbeta_Hall_G maxM) bMs sylS. + have [g _ sXSg] := Sylow_subJ sylS_G (subsetT X) sX. + by exists g; rewrite (subset_trans sXSg) // conjSg (subset_trans sSMb). + have [t _ rFC] := rank_witness 'F('C_(M`_\beta)(P)). + pose Y := 'O_t('C_(M`_\beta)(P)). + have [sYC tY] := (pcore_sub t _ : Y \subset _, pcore_pgroup t _ : t.-group Y). + have sYMb := subset_trans sYC (subsetIl _ _); have bMY := pgroupS sYMb bMb. + have{rFC} ntY: Y != 1. + rewrite -rank_gt0 /= -p_core_Fitting. + rewrite (rank_Sylow (nilpotent_pcore_Hall t (Fitting_nil _))) -rFC. + by rewrite rank_gt0 (trivg_Fitting (solvableS (subsetIr _ _) solCP)). + have bMt: t \in \beta(M). + by rewrite (pnatPpi bMY) // -p_rank_gt0 -rank_pgroup ?rank_gt0. + have sHMg: H \subset M :^ g. + rewrite (subset_trans nXH) // beta_norm_sub_mmax ?mmaxJ // /psubgroup sXMg. + by rewrite (pi_pgroup sX) 1?betaJ. + have sYMg: Y \subset M :^ g := subset_trans (subset_trans sYC sCMbP_H) sHMg. + have sNY_M: 'N(Y) \subset M. + by rewrite beta_norm_sub_mmax // /psubgroup (subset_trans sYMb). + have [_ trCY _] := sigma_group_trans maxM (beta_sub_sigma maxM bMt) tY. + have [|| h cYh /= defMg] := (atransP2 trCY) M (M :^ g). + - by rewrite inE orbit_refl (subset_trans (normG _) sNY_M). + - by rewrite inE mem_orbit ?in_setT. + by rewrite defMg conjGid // (subsetP sNY_M) ?(subsetP (cent_sub _)) in sHMg. +have sXMb: X \subset M`_\beta. + by rewrite (sub_Hall_pcore hallMb) ?(subset_trans sXH sHM) ?(pi_pgroup sX). +rewrite sXMb (sameP commG1P eqP) /= -/X -subG1. +have [sQL [sQM qQ _]] := (subset_trans (normG Q) sNQL, and3P sylQ). +have nsLbL : L`_\beta <| L := pcore_normal _ L; have [sLbL nLbL] := andP nsLbL. +have nLbQ := subset_trans sQL nLbL. +have [<- ti_aLsM _] := sigma_disjoint maxL maxM notLGM. +rewrite subsetI (subset_trans _ (Mbeta_sub_Msigma maxM)) ?andbT; last first. + by rewrite (subset_trans (commgSS sXMb sQM)) // commg_subl nMbM. +have defQ: [~: Q, P] = Q. + rewrite -{2}(coprime_cent_prod nQP) ?(pgroup_sol qQ) ?regPQ ?mulg1 //. + by rewrite (p'nat_coprime (pi_pnat qQ _) pP). +suffices sXL: X \subset L. + apply: subset_trans (Mbeta_sub_Malpha L). + rewrite -quotient_sub1 ?comm_subG ?quotientR ?(subset_trans _ nLbL) //. + have <-: (M`_\beta :&: L) / L`_\beta :&: 'O_q(L^`(1) / L`_\beta) = 1. + rewrite coprime_TIg // coprime_morphl // (coprimeSg (subsetIl _ _)) //. + exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat (pcore_pgroup _ _) _). + rewrite commg_subI // subsetI. + rewrite quotientS; last by rewrite subsetI sXMb. + rewrite (char_norm_trans (pcore_char _ _)) ?quotient_norms //. + by rewrite (subset_trans sXL) ?der_norm. + rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ (Mbeta_quo_nil _))) //. + rewrite quotient_pgroup ?quotient_norms //. + by rewrite normsI ?(subset_trans sQM nMbM) ?normsG. + by rewrite quotientS // -defQ commgSS // (subset_trans nQP). +have{hypU} [r bLr piHr]: exists2 r, r \in \beta(L) & r \in \pi(H). + have [_ _ _] := and5P hypU; rewrite -rank_gt0. + have [r _ ->] := rank_witness 'C_(L`_\beta)(P); rewrite p_rank_gt0 => piCr _. + have [piLb_r piCPr] := (piSg (subsetIl _ _) piCr, piSg (subsetIr _ _) piCr). + have bLr: r \in \beta(L) := pnatPpi (pcore_pgroup _ L) piLb_r. + exists r; rewrite //= (card_Hall hallH) pi_of_part // inE /= piCPr. + by rewrite inE /= bLr orbT. +have sM'r: r \notin \sigma(M). + by apply: contraFN (ti_aLsM r) => sMr; rewrite inE /= beta_sub_alpha. +have defM: M`_\beta * 'N_M(Q) = M. + have nMbQ := subset_trans sQM nMbM. + have nsMbQ_M: M`_\beta <*> Q <| M. + rewrite -{2}(quotientGK nsMbM) -quotientYK ?cosetpre_normal //. + rewrite (eq_Hall_pcore (nilpotent_pcore_Hall q (Mbeta_quo_nil _))) //. + apply: char_normal_trans (pcore_char _ _) (quotient_normal _ _). + exact: der_normal. + rewrite quotient_pHall // (pHall_subl _ (der_sub 1 M) sylQ) //. + by rewrite -defQ commgSS // (subset_trans nUP). + have sylQ_MbQ := pHall_subl (joing_subr _ Q) (normal_sub nsMbQ_M) sylQ. + rewrite -{3}(Frattini_arg nsMbQ_M sylQ_MbQ) /= norm_joinEr // -mulgA. + by congr (_ * _); rewrite mulSGid // subsetI sQM normG. +have [[sM'p _ not_pM'] [sL'p _]] := (and3P t1Mp, andP t1Lp). +have{not_pM'} [R ErR nQR]: exists2 R, R \in 'E_r^1('C_M(P)) & R \subset 'N(Q). + have p'r: r \in p^' by apply: contraNneq sL'p => <-; apply: beta_sub_sigma. + have p'M': p^'.-group M^`(1). + by rewrite p'groupEpi mem_primes (negPf not_pM') !andbF. + pose K := 'O_p^'(M); have [sKM nKM] := andP (pcore_normal _ M : K <| M). + have hallK: p^'.-Hall(M) K. + rewrite -(pquotient_pHall _ (der_normal 1 M)) ?quotient_pgroup //. + rewrite -pquotient_pcore ?der_normal // nilpotent_pcore_Hall //. + by rewrite abelian_nil ?der_abelian. + by rewrite (normalS _ sKM) ?pcore_max ?der_normal. + have sMbK: M`_\beta \subset K. + by rewrite (subset_trans (Mbeta_der1 maxM)) // pcore_max ?der_normal. + have coKP: coprime #|K| #|P| := p'nat_coprime (pcore_pgroup _ M) pP. + have [solK sNK] := (solvableS sKM (mmax_sol maxM), subsetIl K 'N(Q)). + have [nKP coNP] := (subset_trans sPM nKM, coprimeSg sNK coKP). + have nNP: P \subset 'N('N_K(Q)) by rewrite normsI // norms_norm. + have [S sylS nSP] := coprime_Hall_exists r nNP coNP (solvableS sNK solK). + have /subsetIP[sSK nQS]: S \subset 'N_K(Q) := pHall_sub sylS. + have sylS_K: r.-Sylow(K) S. + rewrite pHallE sSK /= -/K -(setIidPr sKM) -defM -group_modl // setIAC. + rewrite (setIidPr sKM) -LagrangeMr partnM // -(card_Hall sylS). + rewrite part_p'nat ?mul1n 1?(pnat_dvd (dvdn_indexg _ _)) //. + by apply: (pi_p'nat bMb); apply: contra sM'r; exact: beta_sub_sigma. + have rC: 'r_r('C_M(P)) > 0 by rewrite p_rank_gt0 (piSg _ piHr) // subsetI sHM. + have{rC} [R ErR] := p_rank_geP rC; have [sRcMP abelR _] := pnElemP ErR. + have{sRcMP abelR} [[sRM cPR] [rR _]] := (subsetIP sRcMP, andP abelR). + have nRP: P \subset 'N(R) by rewrite cents_norm 1?centsC. + have sRK: R \subset K by rewrite sub_Hall_pcore ?(pi_pgroup rR). + have [T [sylT nTP sRT]] := coprime_Hall_subset nKP coKP solK sRK rR nRP. + have [x cKPx defS] := coprime_Hall_trans nKP coKP solK sylS_K nSP sylT nTP. + rewrite -(conjGid (subsetP (setSI _ sKM) x cKPx)). + by exists (R :^ x)%G; rewrite ?pnElemJ ?(subset_trans _ nQS) // defS conjSg. +have [sRcMP abelR _] := pnElemP ErR; have ntR := nt_pnElem ErR isT. +have{sRcMP abelR} [[sRM cPR] [rR _]] := (subsetIP sRcMP, andP abelR). +have sNR_L: 'N(R) \subset L. + by rewrite beta_norm_sub_mmax /psubgroup ?(subset_trans nQR) ?(pi_pgroup rR). +have sPR_M: P <*> R \subset M by rewrite join_subG (subset_trans nUP). +have sM'_PR: \sigma(M)^'.-group (P <*> R). + by rewrite cent_joinEr // pgroupM (pi_pgroup rR) // (pi_pgroup pP). +have [E hallE sPRE] := Hall_superset (mmax_sol maxM) sPR_M sM'_PR. +have{sPRE} [sPE sRE] := joing_subP sPRE. +have EpP: P \in 'E_p^1(E) by apply/pnElemP. +have{ErR} ErR: R \in 'E_r^1('C_E(P)). + by rewrite -(setIidPr (pHall_sub hallE)) setIAC pnElemI inE ErR inE. +apply: subset_trans (cents_norm (subset_trans _ (subsetIr M`_\sigma _))) sNR_L. +apply: subset_trans (cent_tau1Elem_Msigma maxM hallE t1Mp EpP ErR). +by rewrite subsetI cPX (subset_trans sXMb) ?Mbeta_sub_Msigma. +Qed. + +(* This is B & G, Theorem 13.9. *) +Theorem sigma_partition M Mstar : + M \in 'M -> Mstar \in 'M -> gval Mstar \notin M :^: G -> + [predI \sigma(M) & \sigma(Mstar)] =i pred0. +Proof. +move: Mstar => L maxM maxL notMGL q; apply/andP=> [[/= sMq sLq]]. +have [E hallE] := ex_sigma_compl maxM; have [sEM sM'E _] := and3P hallE. +have [_ _ nMsE _] := sdprodP (sdprod_sigma maxM hallE). +have coMsE: coprime #|M`_\sigma| #|E| := pnat_coprime (pcore_pgroup _ _) sM'E. +have [|S sylS nSE] := coprime_Hall_exists q nMsE coMsE. + exact: solvableS (pcore_sub _ _) (mmax_sol maxM). +have [sSMs qS _] := and3P sylS. +have sylS_M := subHall_Sylow (Msigma_Hall maxM) sMq sylS. +have ntS: S :!=: 1. + by rewrite -rank_gt0 (rank_Sylow sylS_M) p_rank_gt0 sigma_sub_pi. +without loss sylS_L: L maxL sLq notMGL / q.-Sylow(L) S. + have sylS_G := sigma_Sylow_G maxM sMq sylS_M. + have [T sylT] := Sylow_exists q L; have sylT_G := sigma_Sylow_G maxL sLq sylT. + have [x Gx ->] := Sylow_trans sylT_G (sigma_Sylow_G maxM sMq sylS_M). + case/(_ (L :^ x)%G); rewrite ?mmaxJ ?sigmaJ ?pHallJ2 //. + by rewrite (orbit_transr _ (mem_orbit 'Js L Gx)). +have [[sSL _] [[E1 hallE1] [E3 hallE3]]] := (andP sylS_L, ex_tau13_compl hallE). +have [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. +have E2_1: E2 :==: 1. + apply: contraR ntS; rewrite -rank_gt0; have [p _ ->] := rank_witness E2. + rewrite p_rank_gt0 => /(pnatPpi (pHall_pgroup hallE2))t2p. + have [A Ep2A _] := ex_tau2Elem hallE t2p. + have [_ _ _ ti_sM] := tau2_compl_context maxM hallE t2p Ep2A. + rewrite -subG1; have [<- _] := ti_sM L maxL notMGL; rewrite subsetI sSMs /=. + by rewrite (sub_Hall_pcore (Msigma_Hall maxL) sSL) (pi_pgroup qS). +have: E1 :!=: 1 by have [_ -> //] := sigma_compl_context maxM complEi. +rewrite -rank_gt0; have [p _ ->] := rank_witness E1; case/p_rank_geP=> P EpP. +have [[sPE1 abelP dimP] [sE1E t1E1 _]] := (pnElemP EpP, and3P hallE1). +have ntP: P :!=: 1 by rewrite (nt_pnElem EpP). +have piPp: p \in \pi(P) by rewrite -p_rank_gt0 ?p_rank_abelem ?dimP. +have t1Mp: p \in \tau1(M) by rewrite (pnatPpi _ (piSg sPE1 _)). +have sPE := subset_trans sPE1 sE1E; have sPM := subset_trans sPE sEM. +have [sNM sNL] := (norm_sigma_Sylow sMq sylS_M, norm_sigma_Sylow sLq sylS_L). +have nSP := subset_trans sPE nSE; have sPL := subset_trans nSP sNL. +have regPS: 'C_S(P) = 1. + apply: contraNeq (contra_orbit _ _ notMGL); rewrite -rank_gt0. + rewrite (rank_pgroup (pgroupS _ qS)) ?subsetIl //; case/p_rank_geP=> Q /=. + rewrite -(setIidPr sSMs) setIAC pnElemI; case/setIdP=> EqQ _. + have [_ uniqSq] := cent_cent_Msigma_tau1_uniq maxM hallE hallE1 sPE1 ntP EqQ. + by rewrite (eq_uniq_mmax (uniqSq S sylS) maxL sSL). +have t1Lp: p \in \tau1(L). + have not_cMsL_P: ~~ (P \subset 'C(M`_\sigma :&: L)). + apply: contra ntS => cMsL_P; rewrite -subG1 -regPS subsetIidl centsC. + by rewrite (subset_trans cMsL_P) ?centS // subsetI sSMs. + apply: contraR (not_cMsL_P) => t1L'p. + have [piEp piLp] := (piSg sPE piPp, piSg sPL piPp). + have [] := Msigma_setI_mmax_central maxM hallE maxL piEp piLp t1L'p _ notMGL. + apply: contraNneq not_cMsL_P; move/commG1P; rewrite centsC. + by apply: subset_trans; rewrite subsetI sPM. + by move->; rewrite ?(abelem_pgroup abelP) // subsetI sPM. +case: (@tau1_mmaxI_asymmetry M L p P q S q S) => //. + by rewrite !inE subsetI sPM sPL abelP dimP. +by rewrite (pHall_subl _ (subsetIl M L) sylS_M) // subsetI (pHall_sub sylS_M). +Qed. + +(* This is B & G, Theorem 13.10. *) +Theorem tau13_regular M E E1 E2 E3 p P : + M \in 'M -> sigma_complement M E E1 E2 E3 -> + P \in 'E_p^1(E1) -> ~~ (P \subset 'C(E3)) -> + [/\ (*a*) semiregular E3 E1, + (*b*) semiregular M`_\sigma E3 + & (*c*) 'C_(M`_\sigma)(P) != 1]. +Proof. +move=> maxM complEi EpP not_cE3P. +have nsMsM: M`_\sigma <| M := pcore_normal _ M; have [sMsM nMsM] := andP nsMsM. +have [hallMs sMaMs] := (Msigma_Hall maxM, Malpha_sub_Msigma maxM). +have [[sE3E' nsE3E] _ [_ cycE3] _ _] := sigma_compl_context maxM complEi. +have [hallE hallE1 _ hallE3] := complEi. +have [[_ sM_Ms _] [sEM sM'E _]] := (and3P hallMs, and3P hallE). +have [[sE1E t1E1 _] [sE3E t3E3 _] _] := (and3P hallE1, and3P hallE3). +have [sPE1 abelP dimP] := pnElemP EpP; have [pP _ _] := and3P abelP. +have [ntP t1MP] := (nt_pnElem EpP isT, pgroupS sPE1 t1E1). +have sPE := subset_trans sPE1 sE1E; have sPM := subset_trans sPE sEM. +have piPp: p \in \pi(P) by rewrite -p_rank_gt0 p_rank_abelem ?dimP. +have t1Mp: p \in \tau1(M) := pnatPpi t1MP piPp. +have [Q sylQ not_cPQ]: exists2 Q, Sylow E3 (gval Q) & ~~ (P \subset 'C(Q)). + apply/exists_inP; rewrite -negb_forall_in; apply: contra not_cE3P. + move/forall_inP=> cPE3; rewrite centsC -(Sylow_gen E3) gen_subG. + by apply/bigcupsP=> Q sylQ; rewrite centsC cPE3. +have{sylQ} [q q_pr sylQ] := SylowP _ _ sylQ; have [sQE3 qQ _] := and3P sylQ. +have ntQ: Q :!=: 1 by apply: contraNneq not_cPQ => ->; apply: cents1. +have t3Mq: q \in \tau3(M). + by rewrite (pnatPpi t3E3) // -p_rank_gt0 -(rank_Sylow sylQ) rank_gt0. +have nQP: P \subset 'N(Q). + rewrite (subset_trans sPE) ?normal_norm //. + by rewrite (char_normal_trans _ nsE3E) ?sub_cyclic_char. +have regPQ: 'C_Q(P) = 1. + apply: contraNeq not_cPQ; rewrite setIC => /meet_Ohm1. + rewrite setIC => /prime_meetG/=/implyP. + rewrite (Ohm1_cyclic_pgroup_prime (cyclicS sQE3 cycE3) qQ) // q_pr centsC. + apply: (coprime_odd_faithful_Ohm1 qQ) nQP _ (mFT_odd _). + exact: sub_pnat_coprime tau3'1 (pgroupS sQE3 t3E3) t1MP. +have sQE' := subset_trans sQE3 sE3E'. +have sQM := subset_trans (subset_trans sQE3 sE3E) sEM. +have [L maxNL] := mmax_exists (mFT_norm_proper ntQ (mFT_pgroup_proper qQ)). +have [maxL sNQL] := setIdP maxNL; have sQL := subset_trans (normG Q) sNQL. +have notMGL: gval L \notin M :^: G. + by apply: mmax_norm_notJ maxM maxL qQ sQM sNQL _; rewrite t3Mq !orbT. +have [ntCMaP tiCMaQP]: 'C_(M`_\alpha)(P) != 1 /\ 'C_(M`_\alpha)(Q <*> P) = 1. + have EpMP: P \in 'E_p^1(M) by apply/pnElemP. + have p'q: q != p by apply: contraNneq (tau3'1 t1Mp) => <-. + have [|_ [] //] := cent_Malpha_reg_tau1 maxM t1Mp p'q EpMP ntQ nQP regPQ. + by apply: contraTneq maxNL => ->; rewrite inE (contra_orbit _ _ notMGL). + have sM'q: q \in \sigma(M)^' by case/andP: t3Mq. + exact: subHall_Sylow hallE sM'q (subHall_Sylow hallE3 t3Mq sylQ). +split=> [x E1x | x E3x |]; last exact: subG1_contra (setSI _ sMaMs) ntCMaP. + apply: contraNeq ntCMaP => ntCE3X. + have prE31: semiprime M`_\sigma (E3 <*> E1). + apply: tau13_primact_Msigma maxM complEi _ => regE13. + by rewrite regE13 ?eqxx in ntCE3X. + rewrite -subG1 -tiCMaQP /= -(setIidPl sMaMs) -!setIA setIS //. + rewrite (cent_semiprime prE31 _ ntP) ?setIS ?centS //=. + by rewrite -!genM_join genS ?mulgSS. + by rewrite sub_gen // subsetU // sPE1 orbT. +have prE3 := tau3_primact_Msigma maxM hallE hallE3. +apply/eqP; apply/idPn; rewrite prE3 {x E3x}// -rank_gt0. +have [u _ -> ruC] := rank_witness 'C_(M`_\sigma)(E3). +have sCMs := subsetIl M`_\sigma 'C(E3). +have sMu: u \in \sigma(M) by rewrite (pnatPpi (pgroupS sCMs _)) -?p_rank_gt0. +have nCE: E \subset 'N('C_(M`_\sigma)(E3)). + by rewrite normsI ?norms_cent ?(normal_norm nsE3E) // (subset_trans sEM). +have coCE := coprimeSg sCMs (pnat_coprime (pcore_pgroup _ _) sM'E). +have solC := solvableS (subset_trans sCMs sMsM) (mmax_sol maxM). +have{nCE coCE solC} [U sylU nUE] := coprime_Hall_exists u nCE coCE solC. +have ntU: U :!=: 1 by rewrite -rank_gt0 (rank_Sylow sylU). +have cMsL_Q: Q \subset 'C(M`_\sigma :&: L). + have t13q: (q \in \tau1(M)) || (q \in \tau3(M)) by rewrite t3Mq orbT. + have [-> //] := cent_norm_tau13_mmax maxM hallE t13q sQM qQ maxNL. + by rewrite subsetI sQM. +rewrite /= -(cent_semiprime prE3 sQE3 ntQ) in sylU. +have [sUMs cQU] := subsetIP (pHall_sub sylU). +have{cMsL_Q} sylU_MsL: u.-Sylow(M`_\sigma :&: L) U. + apply: pHall_subl sylU; last by rewrite subsetI subsetIl centsC. + by rewrite subsetI sUMs (subset_trans (cents_norm _) sNQL). +have sylU_ML: u.-Sylow(M :&: L) U. + apply: subHall_Sylow sMu sylU_MsL. + by rewrite /= -(setIidPl sMsM) -setIA (setI_normal_Hall nsMsM) ?subsetIl. +have [sUML uU _] := and3P sylU_ML; have{sUML} [sUM sUL] := subsetIP sUML. +have [sNUM regPU]: 'N(U) \subset M /\ 'C_U(P) = 1. + have [bMu | bM'u] := boolP (u \in \beta(M)). + have [bM_U sMbMa] := (pi_pgroup uU bMu, Mbeta_sub_Malpha M). + split; first by rewrite beta_norm_sub_mmax /psubgroup ?sUM. + apply/trivgP; rewrite -tiCMaQP centY setIA setSI // subsetI cQU. + by rewrite (subset_trans _ sMbMa) // (sub_Hall_pcore (Mbeta_Hall maxM)). + have sylU_Ms: u.-Sylow(M`_\sigma) U. + have [H hallH cHE'] := der_compl_cent_beta' maxM hallE. + rewrite pHallE sUMs (card_Hall sylU) eqn_dvd. + rewrite partn_dvd ?cardG_gt0 ?cardSg ?subsetIl //=. + rewrite -(@partn_part u \beta(M)^') => [|v /eqnP-> //]. + rewrite -(card_Hall hallH) partn_dvd ?cardG_gt0 ?cardSg //. + by rewrite subsetI (pHall_sub hallH) centsC (subset_trans sQE'). + split; first exact: norm_sigma_Sylow (subHall_Sylow hallMs sMu sylU_Ms). + apply: contraNeq (contra_orbit _ _ notMGL); rewrite -rank_gt0. + rewrite (rank_pgroup (pgroupS _ uU)) ?subsetIl // => /p_rank_geP[X] /=. + rewrite -(setIidPr sUMs) setIAC pnElemI -setIdE => /setIdP[EuX sXU]. + have [_ uniqU] := cent_cent_Msigma_tau1_uniq maxM hallE hallE1 sPE1 ntP EuX. + by rewrite (eq_uniq_mmax (uniqU U sylU_Ms) maxL). +have sPL := subset_trans nQP sNQL. +have sPML: P \subset M :&: L by apply/subsetIP. +have t1Lp: p \in \tau1(L). + have not_cMsL_P: ~~ (P \subset 'C(M`_\sigma :&: L)). + apply: contra ntU => cMsL_P; rewrite -subG1 -regPU subsetIidl. + by rewrite centsC (centsS (pHall_sub sylU_MsL)). + apply: contraR (not_cMsL_P) => t1L'p. + have [piEp piLp] := (piSg sPE piPp, piSg sPL piPp). + case: (Msigma_setI_mmax_central maxM hallE maxL piEp piLp) => // [|->] //. + apply: contraNneq not_cMsL_P => /commG1P. + by rewrite centsC; apply: subset_trans sPML. +have EpMLP: P \in 'E_p^1(M :&: L) by apply/pnElemP. +case: (@tau1_mmaxI_asymmetry M L p P q Q u U) => //. +have [sylQ_E [sM'q _]] := (subHall_Sylow hallE3 t3Mq sylQ, andP t3Mq). +have sylQ_M := subHall_Sylow hallE sM'q sylQ_E. +have sQML: Q \subset M :&: L by apply/subsetIP. +by rewrite (subset_trans sPE nUE) (pHall_subl sQML _ sylQ_M) ?subsetIl. +Qed. + +(* This is B & G, Corollary 13.11. *) +Corollary tau13_nonregular M E E1 E2 E3 : + M \in 'M -> sigma_complement M E E1 E2 E3 -> ~ semiregular M`_\sigma E3 -> + [/\ (*a*) E1 :!=: 1, + (*b*) E3 ><| E1 = E, + (*c*) semiprime M`_\sigma E + & (*d*) forall X, X \in 'E^1(E) -> X <| E]. +Proof. +move=> maxM complEi not_regE3Ms. +have [hallE hallE1 hallE2 hallE3 _] := complEi. +have [[sE1E t1E1 _] [sE3E t3E3 _]] := (and3P hallE1, and3P hallE3). +have{hallE2} E2_1: E2 :==: 1. + apply/idPn; rewrite -rank_gt0; have [p _ ->] := rank_witness E2. + rewrite p_rank_gt0 => /(pnatPpi (pHall_pgroup hallE2))t2p. + have [A Ep2A _] := ex_tau2Elem hallE t2p. + by apply: not_regE3Ms; case: (tau2_regular maxM complEi t2p Ep2A). +have [_ ntE1 [cycE1 cycE3] [defE _] _] := sigma_compl_context maxM complEi. +rewrite (eqP E2_1) sdprod1g in defE; have{ntE1} ntE1 := ntE1 E2_1. +have [nsE3E _ mulE31 nE31 _] := sdprod_context defE. +have cE3E1 P: P \in 'E^1(E1) -> P \subset 'C(E3). + by case/nElemP=> p EpP; apply/idPn => /(tau13_regular maxM complEi EpP)[]. +split=> // [|X EpX]. + rewrite -mulE31 -norm_joinEr //. + have [-> | ntE3] := eqsVneq E3 1. + by rewrite joing1G; apply: (tau1_primact_Msigma maxM hallE hallE1). + apply: tau13_primact_Msigma maxM complEi _ => regE13. + have:= ntE1; rewrite -rank_gt0; case/rank_geP=> P EpP. + have cPE3: E3 \subset 'C(P) by rewrite centsC cE3E1. + have [p Ep1P] := nElemP EpP; have [sPE1 _ _] := pnElemP Ep1P. + have ntP: P :!=: 1 by apply: (nt_pnElem Ep1P). + by case/negP: ntE3; rewrite -(setIidPl cPE3) (cent_semiregular regE13 _ ntP). +have [p Ep1X] := nElemP EpX; have [sXE abelX oX] := pnElemPcard Ep1X. +have [p_pr ntX] := (pnElem_prime Ep1X, nt_pnElem Ep1X isT). +have tau31p: p \in [predU \tau3(M) & \tau1(M)]. + rewrite (pgroupP (pgroupS sXE _)) ?oX // -mulE31 pgroupM. + rewrite (sub_pgroup _ t3E3) => [|q t3q]; last by rewrite inE /= t3q. + by rewrite (sub_pgroup _ t1E1) // => q t1q; rewrite inE /= t1q orbT. +have [/= t3p | t1p] := orP tau31p. + rewrite (char_normal_trans _ nsE3E) ?sub_cyclic_char //. + by rewrite (sub_normal_Hall hallE3) // (pi_pgroup (abelem_pgroup abelX)). +have t1X := pi_pgroup (abelem_pgroup abelX) t1p. +have [e Ee sXeE1] := Hall_Jsub (sigma_compl_sol hallE) hallE1 sXE t1X. +rewrite /normal sXE -(conjSg _ _ e) conjGid //= -normJ -mulE31 mulG_subG. +rewrite andbC sub_abelian_norm ?cyclic_abelian ?cents_norm // centsC cE3E1 //. +rewrite -(setIidPr sE1E) nElemI !inE sXeE1 andbT. +by rewrite -(pnElemJ e) conjGid // def_pnElem in Ep1X; case/setIP: Ep1X. +Qed. + +(* This is B & G, Lemma 13.12. *) +Lemma tau12_regular M E p q P A : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> + p \in \tau1(M) -> P \in 'E_p^1(E) -> q \in \tau2(M) -> A \in 'E_q^2(E) -> + 'C_A(P) != 1 -> + 'C_(M`_\sigma)(P) = 1. +Proof. +move=> maxM hallE t1p EpP t2q Eq2A ntCAP; apply: contraNeq (ntCAP) => ntCMsP. +have [[nsAE _] _ uniq_cMs _] := tau2_compl_context maxM hallE t2q Eq2A. +have [sPE abelP dimP] := pnElemP EpP; have [pP _] := andP abelP. +have ntP: P :!=: 1 by apply: nt_pnElem EpP _. +have [solE t1P] := (sigma_compl_sol hallE, pi_pgroup pP t1p). +have [E1 hallE1 sPE1] := Hall_superset solE sPE t1P. +have [_ [E3 hallE3]] := ex_tau13_compl hallE. +have [E2 _ complEi] := ex_tau2_compl hallE hallE1 hallE3. +have not_cAP: ~~ (P \subset 'C(A)). + have [_ regCE1A _] := tau2_regular maxM complEi t2q Eq2A. + apply: contra ntCMsP => cAP; rewrite (cent_semiregular regCE1A _ ntP) //. + exact/subsetIP. +have [sAE abelA dimA] := pnElemP Eq2A; have [qA _] := andP abelA. +pose Y := 'C_A(P)%G; have{abelA dimA} EqY: Y \in 'E_q^1('C_E(P)). + have sYA: Y \subset A := subsetIl A _; have abelY := abelemS sYA abelA. + rewrite !inE setSI // abelY eqn_leq -{2}rank_abelem // rank_gt0 -ltnS -dimA. + by rewrite properG_ltn_log //= /proper subsetIl subsetIidl centsC. +have ntCMsY: 'C_(M`_\sigma)(Y) != 1. + by apply: subG1_contra ntCMsP; apply: cent_tau1Elem_Msigma t1p EpP EqY. +have EqEY: Y \in 'E_q^1(E) by rewrite pnElemI in EqY; case/setIP: EqY. +have uniqCY := uniq_cMs _ EqEY ntCMsY. +have [ntA nAE] := (nt_pnElem Eq2A isT, normal_norm nsAE). +have [L maxNL] := mmax_exists (mFT_norm_proper ntA (mFT_pgroup_proper qA)). +have [sLq t12Lp]: q \in \sigma(L) /\ (p \in \tau1(L)) || (p \in \tau2(L)). + have [sLt2 t12cA' _] := primes_norm_tau2Elem maxM hallE t2q Eq2A maxNL. + split; first by have /andP[] := sLt2 q t2q. + apply: pnatPpi (pgroupS (quotientS _ sPE) t12cA') _. + rewrite -p_rank_gt0 -rank_pgroup ?quotient_pgroup // rank_gt0. + rewrite -subG1 quotient_sub1 ?subsetI ?sPE // (subset_trans sPE) //. + by rewrite normsI ?normG ?norms_cent. +have [maxL sNL] := setIdP maxNL; have sEL := subset_trans nAE sNL. +have sL'p: p \in \sigma(L)^' by move: t12Lp; rewrite -andb_orr => /andP[]. +have [sPL sL'P] := (subset_trans sPE sEL, pi_pgroup pP sL'p). +have{sL'P} [F hallF sPF] := Hall_superset (mmax_sol maxL) sPL sL'P. +have solF := sigma_compl_sol hallF. +have [sAL sL_A] := (subset_trans (normG A) sNL, pi_pgroup qA sLq). +have sALs: A \subset L`_\sigma by rewrite (sub_Hall_pcore (Msigma_Hall maxL)). +have neqLM: L != M by apply: contraTneq sLq => ->; case/andP: t2q. +have{t12Lp} [t1Lp | t2Lp] := orP t12Lp. + have [F1 hallF1 sPF1] := Hall_superset solF sPF (pi_pgroup pP t1Lp). + have EqLsY: Y \in 'E_q^1('C_(L`_\sigma)(P)). + by rewrite !inE setSI //; case/pnElemP: EqY => _ -> ->. + have [defL _] := cent_cent_Msigma_tau1_uniq maxL hallF hallF1 sPF1 ntP EqLsY. + by rewrite -in_set1 -uniqCY defL set11 in neqLM. +have sCPL: 'C(P) \subset L. + have [B Ep2B _] := ex_tau2Elem hallF t2Lp. + have EpFP: P \in 'E_p^1(F) by apply/pnElemP. + have [_ _ uniq_cLs _] := tau2_compl_context maxL hallF t2Lp Ep2B. + by case/mem_uniq_mmax: (uniq_cLs _ EpFP (subG1_contra (setSI _ sALs) ntCAP)). +have Eq2MA: A \in 'E_q^2(M). + by move: Eq2A; rewrite -(setIidPr (pHall_sub hallE)) pnElemI => /setIP[]. +have [_ _ _ tiMsL _] := tau2_context maxM t2q Eq2MA. +by case/negP: ntCMsP; rewrite -subG1 -(tiMsL L) ?setIS // 3!inE neqLM maxL. +Qed. + +(* This is B & G, Lemma 13.13. *) +Lemma tau13_nonregular_sigma M E p P : + M \in 'M -> \sigma(M)^'.-Hall(M) E -> + P \in 'E_p^1(E) -> (p \in \tau1(M)) || (p \in \tau3(M)) -> + 'C_(M`_\sigma)(P) != 1 -> + {in 'M('N(P)), forall Mstar, p \in \sigma(Mstar)}. +Proof. +move=> maxM hallE EpP t13Mp ntCMsP L maxNL /=. +have [maxL sNL] := setIdP maxNL. +have [sPE abelP dimP] := pnElemP EpP; have [pP _] := andP abelP. +have [solE ntP] := (sigma_compl_sol hallE, nt_pnElem EpP isT). +have /orP[// | t2Lp] := prime_class_mmax_norm maxL pP sNL. +have:= ntCMsP; rewrite -rank_gt0 => /rank_geP[Q /nElemP[q EqQ]]. +have [sQcMsP abelQ dimQ] := pnElemP EqQ; have [sQMs cPQ] := subsetIP sQcMsP. +have piQq: q \in \pi(Q) by rewrite -p_rank_gt0 p_rank_abelem ?dimQ. +have sMq: q \in \sigma(M) := pnatPpi (pgroupS sQMs (pcore_pgroup _ M)) piQq. +have rpM: 'r_p(M) = 1%N by move: t13Mp; rewrite -2!andb_orr andbCA; case: eqP. +have sL'q: q \notin \sigma(L). + have notMGL: gval L \notin M :^: G. + by apply: contraL t2Lp => /imsetP[x _ ->]; rewrite tau2J 2!inE rpM andbF. + by apply: contraFN (sigma_partition maxM maxL notMGL q) => sLq; apply/andP. +have [[sL'p _] [qQ _]] := (andP t2Lp, andP abelQ). +have sL'PQ: \sigma(L)^'.-group (P <*> Q). + by rewrite cent_joinEr // pgroupM (pi_pgroup pP) // (pi_pgroup qQ). +have sPQ_L: P <*> Q \subset L. + by rewrite (subset_trans _ sNL) // join_subG normG cents_norm. +have{sPQ_L sL'PQ} [F hallF sPQF] := Hall_superset (mmax_sol maxL) sPQ_L sL'PQ. +have{sPQF} [sPF sQF] := joing_subP sPQF. +have [A Ep2A _] := ex_tau2Elem hallF t2Lp. +have [[nsAF defA1] _ _ _] := tau2_compl_context maxL hallF t2Lp Ep2A. +have EpAP: P \in 'E_p^1(A) by rewrite -defA1; apply/pnElemP. +have{EpAP} sPA: P \subset A by case/pnElemP: EpAP. +have sCQM: 'C(Q) \subset M. + suffices: 'M('C(Q)) = [set M] by case/mem_uniq_mmax. + have [t1Mp | t3Mp] := orP t13Mp. + have [E1 hallE1 sPE1] := Hall_superset solE sPE (pi_pgroup pP t1Mp). + by have [] := cent_cent_Msigma_tau1_uniq maxM hallE hallE1 sPE1 ntP EqQ. + have [E3 hallE3 sPE3] := Hall_superset solE sPE (pi_pgroup pP t3Mp). + have [[E1 hallE1] _] := ex_tau13_compl hallE; have [sE1E _] := andP hallE1. + have [E2 _ complEi] := ex_tau2_compl hallE hallE1 hallE3. + have [regE3 | ntE1 _ prE _] := tau13_nonregular maxM complEi. + by rewrite (cent_semiregular regE3 sPE3 ntP) eqxx in ntCMsP. + rewrite /= (cent_semiprime prE) // -(cent_semiprime prE sE1E ntE1) in EqQ. + by have [] := cent_cent_Msigma_tau1_uniq maxM hallE hallE1 _ ntE1 EqQ. +have not_cQA: ~~ (A \subset 'C(Q)). + have [_ abelA dimA] := pnElemP Ep2A; apply: contraFN (ltnn 1) => cQA. + by rewrite -dimA -p_rank_abelem // -rpM p_rankS ?(subset_trans cQA sCQM). +have t1Lq: q \in \tau1(L). + have [_ nsCF t1Fb] := tau1_cent_tau2Elem_factor maxL hallF t2Lp Ep2A. + rewrite (pnatPpi (pgroupS (quotientS _ sQF) t1Fb)) //. + rewrite -p_rank_gt0 -rank_pgroup ?quotient_pgroup // rank_gt0. + rewrite -subG1 quotient_sub1 ?(subset_trans _ (normal_norm nsCF)) //. + by rewrite subsetI sQF centsC. +have defP: 'C_A(Q) = P. + apply/eqP; rewrite eq_sym eqEcard subsetI sPA centsC cPQ /=. + have [_ abelA dimA] := pnElemP Ep2A; have [pA _] := andP abelA. + rewrite (card_pgroup (pgroupS _ pA)) ?subsetIl // (card_pgroup pP) dimP. + rewrite leq_exp2l ?prime_gt1 ?(pnElem_prime EpP) //. + by rewrite -ltnS -dimA properG_ltn_log // /proper subsetIl subsetIidl. +have EqFQ: Q \in 'E_q^1(F) by exact/pnElemP. +have regQLs: 'C_(L`_\sigma)(Q) = 1. + by rewrite (tau12_regular maxL hallF t1Lq EqFQ t2Lp Ep2A) // defP. +have ntAQ: [~: A, Q] != 1 by rewrite (sameP eqP commG1P). +have [_ _ [_]] := tau1_act_tau2 maxL hallF t2Lp Ep2A t1Lq EqFQ regQLs ntAQ. +by case/negP; rewrite /= defP (subset_trans (cent_sub P)). +Qed. + +End Section13. + diff --git a/mathcomp/odd_order/BGsection14.v b/mathcomp/odd_order/BGsection14.v new file mode 100644 index 0000000..493f634 --- /dev/null +++ b/mathcomp/odd_order/BGsection14.v @@ -0,0 +1,2513 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import bigop finset prime fingroup morphism perm automorphism quotient. +Require Import action gproduct gfunctor pgroup cyclic center commutator. +Require Import gseries nilpotent sylow abelian maximal hall frobenius. +Require Import ssralg ssrnum ssrint rat. +Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +Require Import BGsection7 BGsection9 BGsection10 BGsection12 BGsection13. + +(******************************************************************************) +(* This file covers B & G, section 14, starting with the definition of the *) +(* sigma-decomposition of elements, sigma-supergroups, and basic categories *) +(* of maximal subgroups: *) +(* sigma_decomposition x == the set of nontrivial constituents x.`_\sigma(M), *) +(* with M ranging over maximal sugroups of G. *) +(* (x is the product of these). *) +(* \ell_\sigma[x] == #|sigma_decomposition x|. *) +(* 'M_\sigma(X) == the set of maximal subgroups M such that X is a *) +(* a subset of M`_\sigma. *) +(* 'M_\sigma[x] := 'M_\sigma(<[x]>) *) +(* \kappa(M) == the set of primes p in \tau1(M) or \tau3(M), such *) +(* that 'C_(M`_\sigma)(P) != 1 for some subgroup of *) +(* M of order p, i.e., the primes for which M fails *) +(* to be a Frobenius group. *) +(* kappa_complement M U K <-> U and K are respectively {kappa, sigma}'- and *) +(* kappa-Hall subgroups of M, whose product is a *) +(* sigma-complement of M. This corresponds to the *) +(* notation introduced at the start of section 15 in *) +(* B & G, but is needed here to capture the use of *) +(* bound variables of 14.2(a) in the statement of *) +(* Lemma 14.12. *) +(* 'M_'F == the set of maximal groups M for which \kappa(M) *) +(* is empty, i.e., the maximal groups of Frobenius *) +(* type (in the final classification, this becomes *) +(* Type I). *) +(* 'M_'P == the complement to 'M_'F in 'M, i.e., the set of M *) +(* for which at least E1 has a proper prime action *) +(* on M`_\sigma. *) +(* 'M_'P1 == the set of maximal subgroups M such that \pi(M) *) +(* is the disjoint union of \sigma(M) and \kappa(M), *) +(* i.e., for which the entire complement acts in a *) +(* prime manner (this troublesome subset of 'M_'P is *) +(* ultimately refined into Types III-V in the final *) +(* classification). *) +(* 'M_'P2 == the complement to 'M_'P1 in 'M_'P; this becomes *) +(* Type II in the final classification. *) +(* 'N[x] == if x != 1 and 'M_\sigma[x] > 1, the unique group *) +(* in 'M('C[x]) (see B & G, Theorem 14.4), and the *) +(* trivial group otherwise. *) +(* 'R[x] := 'C_('N[x]`_\sigma)[x]; if \ell_\sigma[x] == 1, *) +(* this is the normal Hall subgroup of 'C[x] that *) +(* acts sharply transitively by conjugagtion on *) +(* 'M`_\sigma[x] (again, by Theorem 14.4). *) +(* M^~~ == the union of all the cosets x *: 'R[x], with x *) +(* ranging over (M`_\sigma)^#. This will become the *) +(* support set for the Dade isometry for M in the *) +(* character theory part of the proof. *) +(* It seems 'R[x] and 'N[x]`_\sigma play a somewhat the role of a signalizer *) +(* functor in the FT proof; in particular 'R[x] will be used to construct the *) +(* Dade isometry in the character theory part of the proof. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope ring_scope. +Local Open Scope nat_scope. +Import GRing.Theory Num.Theory GroupScope. + +Section Definitons. + +Variable gT : minSimpleOddGroupType. +Implicit Type x : gT. +Implicit Type M X : {set gT}. + +Definition sigma_decomposition x := + [set x.`_\sigma(M) | M : {group gT} in 'M]^#. +Definition sigma_length x := #|sigma_decomposition x|. + +Definition sigma_mmax_of X := [set M in 'M | X \subset M`_\sigma]. + +Definition FT_signalizer_base x := + if #|sigma_mmax_of <[x]>| > 1 then odflt 1%G (pick (mem 'M('C[x]))) else 1%G. + +Definition FT_signalizer x := 'C_((FT_signalizer_base x)`_\sigma)[x]. + +Definition sigma_cover M := \bigcup_(x in (M`_\sigma)^#) x *: FT_signalizer x. + +Definition tau13 M := [predU \tau1(M) & \tau3(M)]. + +Fact kappa_key : unit. Proof. by []. Qed. +Definition kappa_def M : nat_pred := + [pred p in tau13 M | [exists P in 'E_p^1(M), 'C_(M`_\sigma)(P) != 1]]. +Definition kappa := locked_with kappa_key kappa_def. +Canonical kappa_unlockable := [unlockable fun kappa]. + +Definition sigma_kappa M := [predU \sigma(M) & kappa M]. + +Definition kappa_complement (M U K : {set gT}) := + [/\ (sigma_kappa M)^'.-Hall(M) U, (kappa M).-Hall(M) K & group_set (U * K)]. + +Definition TypeF_maxgroups := [set M in 'M | (kappa M)^'.-group M]. + +Definition TypeP_maxgroups := 'M :\: TypeF_maxgroups. + +Definition TypeP1_maxgroups := + [set M in TypeP_maxgroups | (sigma_kappa M).-group M]. + +Definition TypeP2_maxgroups := TypeP_maxgroups :\: TypeP1_maxgroups. + +End Definitons. + +Notation "\ell_ \sigma ( x )" := (sigma_length x) + (at level 8, format "\ell_ \sigma ( x )") : group_scope. + +Notation "''M_' \sigma ( X )" := (sigma_mmax_of X) + (at level 8, format "''M_' \sigma ( X )") : group_scope. + +Notation "''M_' \sigma [ x ]" := (sigma_mmax_of <[x]>) + (at level 8, format "''M_' \sigma [ x ]") : group_scope. + +Notation "''N' [ x ]" := (FT_signalizer_base x) + (at level 8, format "''N' [ x ]") : group_scope. + +Notation "''R' [ x ]" := (FT_signalizer x) + (at level 8, format "''R' [ x ]") : group_scope. + +Notation "M ^~~" := (sigma_cover M) + (at level 2, format "M ^~~") : group_scope. + +Notation "\tau13 ( M )" := (tau13 M) + (at level 8, format "\tau13 ( M )") : group_scope. + +Notation "\kappa ( M )" := (kappa M) + (at level 8, format "\kappa ( M )") : group_scope. + +Notation "\sigma_kappa ( M )" := (sigma_kappa M) + (at level 8, format "\sigma_kappa ( M )") : group_scope. + +Notation "''M_' ''F'" := (TypeF_maxgroups _) + (at level 2, format "''M_' ''F'") : group_scope. + +Notation "''M_' ''P'" := (TypeP_maxgroups _) + (at level 2, format "''M_' ''P'") : group_scope. + +Notation "''M_' ''P1'" := (TypeP1_maxgroups _) + (at level 2, format "''M_' ''P1'") : group_scope. + +Notation "''M_' ''P2'" := (TypeP2_maxgroups _) + (at level 2, format "''M_' ''P2'") : group_scope. + +Section Section14. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types p q q_star r : nat. +Implicit Types x y z : gT. +Implicit Types A E H K L M Mstar N P Q Qstar R S T U V W X Y Z : {group gT}. + +(* Basic properties of the sigma decomposition. *) +Lemma mem_sigma_decomposition x M (xM := x.`_\sigma(M)) : + M \in 'M -> xM != 1 -> xM \in sigma_decomposition x. +Proof. by move=> maxM nt_xM; rewrite !inE nt_xM; apply: mem_imset. Qed. + +Lemma sigma_decompositionJ x z : + sigma_decomposition (x ^ z) = sigma_decomposition x :^ z. +Proof. +rewrite conjD1g -[_ :^ z]imset_comp; congr _^#. +by apply: eq_in_imset => M maxM; rewrite /= consttJ. +Qed. + +Lemma ell_sigmaJ x z : \ell_\sigma(x ^ z) = \ell_\sigma(x). +Proof. by rewrite /sigma_length sigma_decompositionJ cardJg. Qed. + +Lemma sigma_mmaxJ M (X : {set gT}) z : + ((M :^ z)%G \in 'M_\sigma(X :^ z)) = (M \in 'M_\sigma(X)). +Proof. by rewrite inE mmaxJ MsigmaJ conjSg !inE. Qed. + +Lemma card_sigma_mmaxJ (X : {set gT}) z : + #|'M_\sigma(X :^ z)| = #|'M_\sigma(X)|. +Proof. +rewrite -(card_setact 'JG _ z^-1) setactVin ?inE //. +by apply: eq_card => M; rewrite inE sigma_mmaxJ. +Qed. + +Lemma sigma_decomposition_constt' x M (sM := \sigma(M)) : + M \in 'M -> sigma_decomposition x.`_sM^' = sigma_decomposition x :\ x.`_sM. +Proof. +move=> maxM; apply/setP=> y; rewrite !inE andbCA; apply: andb_id2l => nty. +apply/imsetP/andP=> [ | [neq_y_xM /imsetP]] [L maxL def_y]. + have not_sMy: ~~ sM.-elt y. + apply: contra nty => sMy; rewrite -order_eq1 (pnat_1 sMy) // def_y. + by apply: p_eltX; apply: p_elt_constt. + split; first by apply: contraNneq not_sMy => ->; apply: p_elt_constt. + have notMGL: gval L \notin M :^: G. + apply: contra not_sMy; rewrite def_y; case/imsetP=> z _ ->. + by rewrite (eq_constt _ (sigmaJ M z)) p_elt_constt. + apply/imsetP; exists L; rewrite // def_y sub_in_constt // => p _ sLp. + by apply: contraFN (sigma_partition maxM maxL notMGL p) => sMp; apply/andP. +exists L; rewrite ?sub_in_constt // => p _ sLp. +suffices notMGL: gval L \notin M :^: G. + by apply: contraFN (sigma_partition maxM maxL notMGL p) => sMp; apply/andP. +apply: contra neq_y_xM; rewrite def_y => /imsetP[z _ ->]. +by rewrite (eq_constt _ (sigmaJ M z)). +Qed. + +(* General remarks about the sigma-decomposition, p. 105 of B & G. *) +Remark sigma_mmax_exists p : + p \in \pi(G) -> {M : {group gT} | M \in 'M & p \in \sigma(M)}. +Proof. +move=> piGp; have [P sylP] := Sylow_exists p [set: gT]. +have ntP: P :!=: 1 by rewrite -rank_gt0 (rank_Sylow sylP) p_rank_gt0. +have ltPG: P \proper G := mFT_pgroup_proper (pHall_pgroup sylP). +have [M maxNM] := mmax_exists (mFT_norm_proper ntP ltPG). +have{maxNM} [maxM sNM] := setIdP maxNM; have sPM := subset_trans (normG P) sNM. +have{sylP} sylP := pHall_subl sPM (subsetT M) sylP. +by exists M => //; apply/exists_inP; exists P. +Qed. + +Lemma ell_sigma0P x : reflect (x = 1) (\ell_\sigma(x) == 0). +Proof. +rewrite cards_eq0 setD_eq0. +apply: (iffP idP) => [x1 | ->]; last first. + by apply/subsetP=> _ /imsetP[M _ ->]; rewrite constt1 inE. +rewrite -(prod_constt x) big1_seq //= => p _; apply: contraTeq x1 => nt_xp. +have piXp: p \in \pi(#[x]) by rewrite -p_part_gt1 -order_constt order_gt1. +have [M maxM sMp] := sigma_mmax_exists (piSg (subsetT _) piXp). +apply/subsetPn; exists (x.`_(\sigma(M))); first exact: mem_imset. +by rewrite (sameP set1P constt1P); apply: contraL sMp => /pnatPpi; apply. +Qed. + +Remark sigma_decomposition_subG x H : + x \in H -> sigma_decomposition x \subset H. +Proof. +by move=> Hx; apply/subsetP=> _ /setD1P[_ /imsetP[M _ ->]]; apply: groupX. +Qed. + +Remark prod_sigma_decomposition x : + \prod_(x_sM in sigma_decomposition x) x_sM = x. +Proof. +rewrite -big_filter filter_index_enum; set e := enum _. +have: uniq e := enum_uniq _; have: e =i sigma_decomposition x := mem_enum _. +elim: {x}e (x) => [|y e IHe] x def_e /= Ue. + by rewrite big_nil (ell_sigma0P x _) //; apply/pred0P; apply: fsym. +have{Ue} [not_e_y Ue] := andP Ue. +have [nty] := setD1P (etrans (fsym def_e y) (mem_head _ _)). +case/imsetP=> M maxM def_y; rewrite big_cons -(consttC \sigma(M) x) -def_y. +congr (y * _); apply: IHe Ue => z; rewrite sigma_decomposition_constt' //. +rewrite -def_y inE -def_e !inE andb_orr andNb andb_idl //. +by apply: contraTneq => ->. +Qed. + +Lemma ell1_decomposition x : + \ell_\sigma(x) == 1%N -> sigma_decomposition x = [set x]. +Proof. +case/cards1P=> y sdx_y. +by rewrite -{2}[x]prod_sigma_decomposition sdx_y big_set1. +Qed. + +Lemma Msigma_ell1 M x : + M \in 'M -> x \in (M`_\sigma)^# -> \ell_\sigma(x) == 1%N. +Proof. +move=> maxM /setD1P[ntx Ms_x]; apply/cards1P. +have sMx: \sigma(M).-elt x := mem_p_elt (pcore_pgroup _ _) Ms_x. +have def_xM: x.`_(\sigma(M)) = x := constt_p_elt sMx. +exists x; apply/eqP; rewrite eqEsubset sub1set !inE ntx -setD_eq0 /=. +rewrite -{2 3}def_xM -sigma_decomposition_constt' // (constt1P _) ?p_eltNK //. +by rewrite -cards_eq0 (sameP (ell_sigma0P 1) eqP) eqxx; apply: mem_imset. +Qed. + +Remark ell_sigma1P x : + reflect (x != 1 /\ 'M_\sigma[x] != set0) (\ell_\sigma(x) == 1%N). +Proof. +apply: (iffP idP) => [ell1x | [ntx]]; last first. + case/set0Pn=> M /setIdP[maxM]; rewrite cycle_subG => Ms_x. + by rewrite (Msigma_ell1 maxM) // in_setD1 ntx. +have sdx_x: x \in sigma_decomposition x by rewrite ell1_decomposition ?set11. +have{sdx_x} [ntx sdx_x] := setD1P sdx_x; split=> //; apply/set0Pn. +have{sdx_x} [M maxM def_x] := imsetP sdx_x; rewrite // -cycle_eq1 in ntx. +have sMx: \sigma(M).-elt x by rewrite def_x p_elt_constt. +have [[z sXzMs] _] := sigma_Jsub maxM sMx ntx. +by exists (M :^ z^-1)%G; rewrite inE mmaxJ maxM MsigmaJ -sub_conjg. +Qed. + +Remark ell_sigma_le1 x :(\ell_\sigma(x) <= 1) = ('M_\sigma[x] != set0). +Proof. +rewrite -[_ <= 1](mem_iota 0 2) !inE (sameP (ell_sigma0P x) eqP). +rewrite (sameP (ell_sigma1P x) andP); case: eqP => //= ->; symmetry. +have [M max1M] := mmax_exists (mFT_pgroup_proper (@pgroup1 gT 2)). +have [maxM _] := setIdP max1M; apply/set0Pn; exists M. +by rewrite inE maxM cycle1 sub1G. +Qed. + +(* Basic properties of \kappa and the maximal group subclasses. *) +Lemma kappaJ M x : \kappa(M :^ x) =i \kappa(M). +Proof. +move=> p; rewrite unlock 3!{1}inE /= tau1J tau3J; apply: andb_id2l => _. +apply/exists_inP/exists_inP=> [] [P EpP ntCMsP]. + rewrite -(conjsgK x M); exists (P :^ x^-1)%G; first by rewrite pnElemJ. + by rewrite MsigmaJ centJ -conjIg -subG1 sub_conjg conjs1g subG1. +exists (P :^ x)%G; first by rewrite pnElemJ. +by rewrite MsigmaJ centJ -conjIg -subG1 sub_conjg conjs1g subG1. +Qed. + +Lemma kappa_tau13 M p : p \in \kappa(M) -> (p \in \tau1(M)) || (p \in \tau3(M)). +Proof. by rewrite unlock => /andP[]. Qed. + +Lemma kappa_sigma' M : {subset \kappa(M) <= \sigma(M)^'}. +Proof. by move=> p /kappa_tau13/orP[] /andP[]. Qed. + +Remark rank_kappa M p : p \in \kappa(M) -> 'r_p(M) = 1%N. +Proof. by case/kappa_tau13/orP=> /and3P[_ /eqP]. Qed. + +Lemma kappa_pi M : {subset \kappa(M) <= \pi(M)}. +Proof. by move=> p kMp; rewrite -p_rank_gt0 rank_kappa. Qed. + +Remark kappa_nonregular M p P : + p \in \kappa(M) -> P \in 'E_p^1(M) -> 'C_(M`_\sigma)(P) != 1. +Proof. +move=> kMp EpP; have rpM := rank_kappa kMp. +have [sPM abelP oP] := pnElemPcard EpP; have [pP _] := andP abelP. +have [Q EpQ nregQ]: exists2 Q, Q \in 'E_p^1(M) & 'C_(M`_\sigma)(Q) != 1. + by apply/exists_inP; rewrite unlock in kMp; case/andP: kMp. +have [sQM abelQ oQ] := pnElemPcard EpQ; have [pQ _] := andP abelQ. +have [S sylS sQS] := Sylow_superset sQM pQ; have [_ pS _] := and3P sylS. +have [x Mx sPxS] := Sylow_Jsub sylS sPM pP. +rewrite -(inj_eq (@conjsg_inj _ x)) conjs1g conjIg -centJ. +rewrite (normsP (normal_norm (pcore_normal _ _))) // (subG1_contra _ nregQ) //. +rewrite setIS ?centS // -(cardSg_cyclic _ sPxS sQS) ?cardJg ?oP ?oQ //. +by rewrite (odd_pgroup_rank1_cyclic pS) ?mFT_odd // (p_rank_Sylow sylS) rpM. +Qed. + +Lemma ex_kappa_compl M K : + M \in 'M -> \kappa(M).-Hall(M) K -> + exists U : {group gT}, kappa_complement M U K. +Proof. +move=> maxM hallK; have [sKM kK _] := and3P hallK. +have s'K: \sigma(M)^'.-group K := sub_pgroup (@kappa_sigma' M) kK. +have [E hallE sKE] := Hall_superset (mmax_sol maxM) sKM s'K. +pose sk' := \sigma_kappa(M)^'. +have [U hallU] := Hall_exists sk' (sigma_compl_sol hallE). +exists U; split=> //. + by apply: subHall_Hall hallE _ hallU => p; case/norP. +suffices ->: U * K = E by apply: groupP. +have [[sUE sk'U _] [sEM s'E _]] := (and3P hallU, and3P hallE). +apply/eqP; rewrite eqEcard mulG_subG sUE sKE /= coprime_cardMg; last first. + by apply: p'nat_coprime (sub_pgroup _ sk'U) kK => p; case/norP. +rewrite -(partnC \kappa(M) (cardG_gt0 E)) -{2}(part_pnat_id s'E) mulnC. +rewrite -(card_Hall (pHall_subl sKE sEM hallK)) leq_mul // -partnI. +by rewrite -(@eq_partn sk') -?(card_Hall hallU) // => p; apply: negb_or. +Qed. + +Lemma FtypeP M : reflect (M \in 'M /\ \kappa(M) =i pred0) (M \in 'M_'F). +Proof. +do [apply: (iffP setIdP) => [] [maxM k'M]; split] => // [p|]. + by apply/idP=> /= kMp; case/negP: (pnatPpi k'M (kappa_pi kMp)). +by apply/pgroupP=> p _ _; rewrite inE /= k'M. +Qed. + +Lemma PtypeP M : reflect (M \in 'M /\ exists p, p \in \kappa(M)) (M \in 'M_'P). +Proof. +apply: (iffP setDP) => [[maxM kM] | [maxM [p kMp]]]; split => //. + rewrite inE maxM !negb_and cardG_gt0 /= all_predC negbK in kM. + by have [p _ kMp] := hasP kM; exists p. +by apply/FtypeP=> [[_ kM0]]; rewrite kM0 in kMp. +Qed. + +Lemma trivg_kappa M K : + M \in 'M -> \kappa(M).-Hall(M) K -> (K :==: 1) = (M \in 'M_'F). +Proof. +by move=> maxM hallK; rewrite inE maxM trivg_card1 (card_Hall hallK) partG_eq1. +Qed. + +(* Could go in Section 10. *) +Lemma not_sigma_mmax M : M \in 'M -> ~~ \sigma(M).-group M. +Proof. +move=> maxM; apply: contraL (mmax_sol maxM); rewrite negb_forall_in => sM. +apply/existsP; exists M; rewrite mmax_neq1 // subsetIidl andbT. +apply: subset_trans (Msigma_der1 maxM). +by rewrite (sub_Hall_pcore (Msigma_Hall maxM)). +Qed. + +Lemma trivg_kappa_compl M U K : + M \in 'M -> kappa_complement M U K -> (U :==: 1) = (M \in 'M_'P1). +Proof. +move=> maxM [hallU _ _]; symmetry. +rewrite 3!inE maxM /= trivg_card1 (card_Hall hallU) partG_eq1 pgroupNK andbT. +apply: andb_idl => skM; apply: contra (not_sigma_mmax maxM). +by apply: sub_in_pnat => p /(pnatPpi skM)/orP[] // kMp /negP. +Qed. + +Lemma FtypeJ M x : ((M :^ x)%G \in 'M_'F) = (M \in 'M_'F). +Proof. by rewrite inE mmaxJ pgroupJ (eq_p'group _ (kappaJ M x)) !inE. Qed. + +Lemma PtypeJ M x : ((M :^ x)%G \in 'M_'P) = (M \in 'M_'P). +Proof. by rewrite !in_setD mmaxJ FtypeJ. Qed. + +Lemma P1typeJ M x : ((M :^ x)%G \in 'M_'P1) = (M \in 'M_'P1). +Proof. +rewrite inE PtypeJ pgroupJ [M \in 'M_'P1]inE; congr (_ && _). +by apply: eq_pgroup => p; rewrite inE /= kappaJ sigmaJ. +Qed. + +Lemma P2typeJ M x : ((M :^ x)%G \in 'M_'P2) = (M \in 'M_'P2). +Proof. by rewrite in_setD PtypeJ P1typeJ -in_setD. Qed. + +(* This is B & G, Lemma 14.1. *) +Lemma sigma'_kappa'_facts M p S (A := 'Ohm_1(S)) (Ms := M`_\sigma) : + M \in 'M -> p.-Sylow(M) S -> + [&& p \in \pi(M), p \notin \sigma(M) & p \notin \kappa(M)] -> + [/\ M \in 'M_'F :|: 'M_'P2, logn p #|A| <= 2, 'C_Ms(A) = 1 & nilpotent Ms]. +Proof. +move=> maxM sylS /and3P[piMp sM'p kM'p]; have [sSM pS _] := and3P sylS. +rewrite 8!(maxM, inE) /= !andbT negb_and orb_andr orbN andbT negbK. +rewrite (contra (fun skM => pnatPpi skM piMp)) ?orbT; last exact/norP. +rewrite partition_pi_mmax // (negPf sM'p) orbF orbCA in piMp. +have{piMp} [t2p | t13p] := orP piMp. + rewrite (tau2_Msigma_nil maxM t2p); have [_ rpM] := andP t2p. + have{rpM} rS: 2 <= 'r_p(S) by rewrite (p_rank_Sylow sylS) (eqP rpM). + have [B EpB] := p_rank_geP rS; have{EpB} [sBS abelB dimB] := pnElemP EpB. + have EpB: B \in 'E_p^2(M) by rewrite !inE abelB dimB (subset_trans sBS). + have [defB _ regB _ _] := tau2_context maxM t2p EpB. + by rewrite /A -dimB; have [_ [|->]] := defB S sylS. +have [ntS cycS]: S :!=: 1 /\ cyclic S. + rewrite (odd_pgroup_rank1_cyclic pS) ?mFT_odd // (p_rank_Sylow sylS). + apply/andP; rewrite -rank_gt0 (rank_Sylow sylS) -eqn_leq eq_sym. + by rewrite -2!andb_orr orNb andbT inE /= sM'p in t13p. +have [p_pr _ _] := pgroup_pdiv pS ntS. +have oA: #|A| = p by rewrite (Ohm1_cyclic_pgroup_prime cycS pS). +have sAM: A \subset M := subset_trans (Ohm_sub 1 S) sSM. +have regA: 'C_Ms(A) = 1. + apply: contraNeq kM'p => nregA; rewrite unlock; apply/andP; split=> //. + by apply/exists_inP; exists [group of A]; rewrite ?p1ElemE // !inE sAM oA /=. +have defMsA: Ms ><| A = Ms <*> A. + rewrite sdprodEY ?coprime_TIg ?(subset_trans sAM) ?gFnorm // oA. + by rewrite (pnat_coprime (pcore_pgroup _ _)) ?pnatE. +rewrite (prime_Frobenius_sol_kernel_nil defMsA) ?oA ?(pfactorK 1) //. +by rewrite (solvableS _ (mmax_sol maxM)) // join_subG pcore_sub. +Qed. + +Lemma notP1type_Msigma_nil M : + (M \in 'M_'F) || (M \in 'M_'P2) -> nilpotent M`_\sigma. +Proof. +move=> notP1maxM; suffices [maxM]: M \in 'M /\ ~~ \sigma_kappa(M).-group M. + rewrite negb_and cardG_gt0 => /allPn[p piMp /norP[s'p k'p]]. + by have [S /sigma'_kappa'_facts[] //] := Sylow_exists p M; apply/and3P. +have [/setIdP[maxM k'M] | /setDP[PmaxM]] := orP notP1maxM; last first. + by rewrite inE PmaxM; case/setDP: PmaxM. +split=> //; apply: contra (not_sigma_mmax maxM). +by apply: sub_in_pnat => p piMp /orP[] // /idPn[]; exact: (pnatPpi k'M). +Qed. + +(* This is B & G, Proposition 14.2. *) +Proposition Ptype_structure M K (Ms := M`_\sigma) (Kstar := 'C_Ms(K)) : + M \in 'M_'P -> \kappa(M).-Hall(M) K -> + [/\ (*a*) exists2 U : {group gT}, + kappa_complement M U K /\ Ms ><| (U ><| K) = M + & [/\ abelian U, semiprime Ms K & semiregular U K], + (*b*) (*1.2*) K \x Kstar = 'N_M(K) + /\ {in 'E^1(K), forall X, + (*1.1*) 'N_M(X) = 'N_M(K) + /\ (*2*) {in 'M('N(X)), forall Mstar, X \subset Mstar`_\sigma}}, + (*c*) Kstar != 1 /\ {in 'E^1(Kstar), forall X, 'M('C(X)) = [set M]}, + [/\ (*d*) {in ~: M, forall g, Kstar :&: M :^ g = 1} + /\ {in M :\: 'N_M(K), forall g, K :&: K :^ g = 1}, + (*e*) {in \pi(Kstar), forall p S, + p.-Sylow(M) S -> 'M(S) = [set M] /\ ~~ (S \subset Kstar)} + & (*f*) forall Y, \sigma(M).-group Y -> Y :&: Kstar != 1 -> Y \subset Ms] + & (*g*) M \in 'M_'P2 -> + [/\ \sigma(M) =i \beta(M), prime #|K|, nilpotent Ms + & normedTI Ms^# G M]]. +Proof. +move: @Kstar => Ks PmaxM hallK; have [maxM notFmaxM] := setDP PmaxM. +have sMs: \sigma(M).-group M`_\sigma := pcore_pgroup _ M. +have{notFmaxM} ntK: K :!=: 1 by rewrite (trivg_kappa maxM). +have [sKM kK _] := and3P hallK; have s'K := sub_pgroup (@kappa_sigma' M) kK. +have solM := mmax_sol maxM; have [E hallE sKE] := Hall_superset solM sKM s'K. +have [[sEM s'E _] [_ [E3 hallE3]]] := (and3P hallE, ex_tau13_compl hallE). +have [F1 hallF1] := Hall_exists \tau1(M) (solvableS sKM solM). +have [[sF1K t1F1 _] solE] := (and3P hallF1, sigma_compl_sol hallE). +have [E1 hallE1 sFE1] := Hall_superset solE (subset_trans sF1K sKE) t1F1. +have [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. +have [[_ nsE3E] _ [cycE1 _] [defEl defE] _] := sigma_compl_context maxM complEi. +have [sE1E t1E1 _] := and3P hallE1; have sE1M := subset_trans sE1E sEM. +have [sE3E t3E3 _] := and3P hallE3; have sE3M := subset_trans sE3E sEM. +set part_a := exists2 U, _ & _; pose b1_hyp := {in 'E^1(K), forall X, X <| K}. +have [have_a nK1K ntE1 sE1K]: [/\ part_a, b1_hyp, E1 :!=: 1 & E1 \subset K]. + have [t1K | not_t1K] := boolP (\tau1(M).-group K). + have sKE1: K \subset E1 by rewrite (sub_pHall hallF1 t1K). + have prE1 := tau1_primact_Msigma maxM hallE hallE1. + have st1k: {subset \tau1(M) <= \kappa(M)}. + move=> p t1p; rewrite unlock 3!inE /= t1p /=. + have [X]: exists X, X \in 'E_p^1(E1). + apply/p_rank_geP; rewrite p_rank_gt0 /= (card_Hall hallE1). + by rewrite pi_of_part // inE /= (partition_pi_sigma_compl maxM) ?t1p. + rewrite -(setIidPr sE1M) pnElemI -setIdE => /setIdP[EpX sXE1]. + pose q := pdiv #|K|; have piKq: q \in \pi(K) by rewrite pi_pdiv cardG_gt1. + have /p_rank_geP[Y]: 0 < 'r_q(K) by rewrite p_rank_gt0. + rewrite -(setIidPr sKM) pnElemI -setIdE => /setIdP[EqY sYK]. + have ntCMsY := kappa_nonregular (pnatPpi kK piKq) EqY. + have [ntY sYE1] := (nt_pnElem EqY isT, subset_trans sYK sKE1). + apply/exists_inP; exists X; rewrite ?(subG1_contra _ ntCMsY) //=. + by rewrite (cent_semiprime prE1 sYE1 ntY) ?setIS ?centS. + have defK := sub_pHall hallK (sub_pgroup st1k t1E1) sKE1 sE1M. + split=> [|X||]; rewrite ?defK //; last first. + rewrite -defK; case/nElemP=> p /pnElemP[sXE1 _ _]. + by rewrite char_normal // sub_cyclic_char. + have [[U _ defU _] _ _ _] := sdprodP defE; rewrite defU in defE. + have [nsUE _ mulUE1 nUE1 _] := sdprod_context defE. + have [[_ V _ defV] _] := sdprodP defEl; rewrite defV. + have [_ <- nE21 _] := sdprodP defV => /mulGsubP[nE32 nE31] _. + have regUK: semiregular U K. + move=> y /setD1P[]; rewrite -cycle_subG -cent_cycle -order_gt1. + rewrite -pi_pdiv; move: (pdiv _) => p pi_y_p Ky. + have piKp := piSg Ky pi_y_p; have t1p := pnatPpi t1K piKp. + move: pi_y_p; rewrite -p_rank_gt0 => /p_rank_geP[Y] /=. + rewrite -{1}(setIidPr (subset_trans Ky sKE)) pnElemI -setIdE. + case/setIdP=> EpY sYy; have EpMY := subsetP (pnElemS _ _ sEM) Y EpY. + apply: contraNeq (kappa_nonregular (pnatPpi kK piKp) EpMY). + move/(subG1_contra (setIS U (centS sYy))). + have{y sYy Ky} sYE1 := subset_trans sYy (subset_trans Ky sKE1). + have ntY: Y :!=: 1 by apply: (nt_pnElem EpY). + rewrite -subG1 /=; have [_ <- _ tiE32] := sdprodP defU. + rewrite -subcent_TImulg ?subsetI ?(subset_trans sYE1) // mulG_subG. + rewrite !subG1 /= => /nandP[nregE3Y | nregE2Y]. + have pr13 := cent_semiprime (tau13_primact_Msigma maxM complEi _). + rewrite pr13 ?(subset_trans sYE1) ?joing_subr //; last first. + by move/cent_semiregular=> regE31; rewrite regE31 ?eqxx in nregE3Y. + pose q := pdiv #|'C_E3(Y)|. + have piE3q: q \in \pi(E3). + by rewrite (piSg (subsetIl E3 'C(Y))) // pi_pdiv cardG_gt1. + have /p_rank_geP[X]: 0 < 'r_q(M :&: E3). + by rewrite (setIidPr sE3M) p_rank_gt0. + rewrite pnElemI -setIdE => /setIdP[EqX sXE3]. + rewrite -subG1 -(_ : 'C_Ms(X) = 1) ?setIS ?centS //. + by rewrite (subset_trans sXE3) ?joing_subl. + apply: contraTeq (pnatPpi t3E3 piE3q) => nregMsX; apply: tau3'1. + suffices kq: q \in \kappa(M). + rewrite (pnatPpi t1K) //= (card_Hall hallK) pi_of_part //. + by rewrite inE /= kappa_pi. + rewrite unlock 3!inE /= (pnatPpi t3E3 piE3q) orbT /=. + by apply/exists_inP; exists X. + pose q := pdiv #|'C_E2(Y)|; have [sE2E t2E2 _] := and3P hallE2. + have piCE2Yq: q \in \pi('C_E2(Y)) by rewrite pi_pdiv cardG_gt1. + have [X]: exists X, X \in 'E_q^1(E :&: 'C_E2(Y)). + by apply/p_rank_geP; rewrite /= setIA (setIidPr sE2E) p_rank_gt0. + rewrite pnElemI -setIdE => /setIdP[EqX sXcE2Y]. + have t2q:= pnatPpi t2E2 (piSg (subsetIl _ _) piCE2Yq). + have [A Eq2A _] := ex_tau2Elem hallE t2q. + have [[_ defEq1] _ _ _] := tau2_compl_context maxM hallE t2q Eq2A. + rewrite (tau12_regular maxM hallE t1p EpY t2q Eq2A) //. + rewrite (subG1_contra _ (nt_pnElem EqX _)) // subsetI. + rewrite defEq1 in EqX; case/pnElemP: EqX => -> _ _. + by rewrite (subset_trans sXcE2Y) ?subsetIr. + have [defM [sUE _]] := (sdprod_sigma maxM hallE, andP nsUE). + have hallU: \sigma_kappa(M)^'.-Hall(M) U. + suffices: [predI \sigma(M)^' & \kappa(M)^'].-Hall(M) U. + by apply: etrans; apply: eq_pHall=> p; rewrite inE /= negb_or. + apply: subHall_Hall hallE _ _ => [p|]; first by case/andP. + rewrite pHallE partnI (part_pnat_id s'E) -pHallE. + have hallK_E: \kappa(M).-Hall(E) K := pHall_subl sKE sEM hallK. + by apply/(sdprod_normal_pHallP nsUE hallK_E); rewrite -defK. + exists U; [rewrite -{2}defK defE | rewrite -{1}defK]; split=> //. + by split; rewrite // -defK mulUE1 groupP. + apply: abelianS (der_mmax_compl_abelian maxM hallE). + rewrite -(coprime_cent_prod nUE1) ?(solvableS sUE) //. + by rewrite {2}defK (cent_semiregular regUK) // mulg1 commgSS. + by rewrite (coprime_sdprod_Hall_r defE); apply: pHall_Hall hallE1. + move: not_t1K; rewrite negb_and cardG_gt0 => /allPn[p piKp t1'p]. + have kp := pnatPpi kK piKp; have t3p := kappa_tau13 kp. + rewrite [p \in _](negPf t1'p) /= in t3p. + have [X]: exists X, X \in 'E_p^1(K) by apply/p_rank_geP; rewrite p_rank_gt0. + rewrite -{1}(setIidPr sKM) pnElemI -setIdE => /setIdP[EpX sXK]. + have sXE3: X \subset E3. + rewrite (sub_normal_Hall hallE3) ?(subset_trans sXK) ?(pi_pgroup _ t3p) //. + by case/pnElemP: EpX => _ /andP[]. + have [nregX ntX] := (kappa_nonregular kp EpX, nt_pnElem EpX isT). + have [regE3|ntE1 {defE}defE prE nE1_E] := tau13_nonregular maxM complEi. + by case/eqP: nregX; rewrite (cent_semiregular regE3). + have defK: E :=: K. + apply: (sub_pHall hallK _ sKE sEM); apply/pgroupP=> q q_pr q_dv_E. + have{q_dv_E} piEq: q \in \pi(E) by rewrite mem_primes q_pr cardG_gt0. + rewrite unlock; apply/andP; split=> /=. + apply: pnatPpi piEq; rewrite -pgroupE -(sdprodW defE). + rewrite pgroupM (sub_pgroup _ t3E3) => [|r t3r]; last by apply/orP; right. + by rewrite (sub_pgroup _ t1E1) // => r t1r; apply/orP; left. + have:= piEq; rewrite -p_rank_gt0 => /p_rank_geP[Y]. + rewrite -{1}(setIidPr sEM) pnElemI -setIdE => /setIdP[EqY sYE]. + rewrite (cent_semiprime prE) ?(subset_trans sXK) // in nregX. + apply/exists_inP; exists Y => //; apply: subG1_contra nregX. + by rewrite setIS ?centS. + have defM := sdprod_sigma maxM hallE. + rewrite /b1_hyp -defK; split=> //; exists 1%G; last first. + by split; [exact: abelian1 | rewrite -defK | exact: semiregular1l]. + rewrite sdprod1g; do 2?split=> //; rewrite ?mul1g ?groupP -?defK //. + rewrite pHallE sub1G cards1 eq_sym partG_eq1 pgroupNK /=. + have{defM} [_ defM _ _] := sdprodP defM; rewrite -{2}defM defK pgroupM. + rewrite (sub_pgroup _ sMs) => [|r sr]; last by apply/orP; left. + by rewrite (sub_pgroup _ kK) // => r kr; apply/orP; right. +set part_b := _ /\ _; set part_c := _ /\ _; set part_d := _ /\ _. +have [U [complUK defM] [cUU prMsK regUK]] := have_a. +have [hallU _ _] := complUK; have [sUM sk'U _] := and3P hallU. +have have_b: part_b. + have coMsU: coprime #|Ms| #|U|. + by rewrite (pnat_coprime sMs (sub_pgroup _ sk'U)) // => p; case/norP. + have{defM} [[_ F _ defF]] := sdprodP defM; rewrite defF. + have [_ <- nUK _] := sdprodP defF; rewrite mulgA mulG_subG => defM. + case/andP=> nMsU nMsK _. + have coMsU_K: coprime #|Ms <*> U| #|K|. + rewrite norm_joinEr // (p'nat_coprime _ kK) // -pgroupE. + rewrite pgroupM // (sub_pgroup _ sMs) => [|r]; last first. + by apply: contraL; apply: kappa_sigma'. + by apply: sub_pgroup _ sk'U => r /norP[]. + have defNK X: X <| K -> X :!=: 1 -> 'N_M(X) = Ks * K. + case/andP=> sXK nXK ntX; rewrite -defM -(norm_joinEr nMsU). + rewrite setIC -group_modr // setIC. + rewrite coprime_norm_cent ?(subset_trans sXK) ?normsY //; last first. + by rewrite (coprimegS sXK). + rewrite /= norm_joinEr -?subcent_TImulg ?(coprime_TIg coMsU) //; last first. + by rewrite subsetI !(subset_trans sXK). + by rewrite (cent_semiregular regUK) // mulg1 (cent_semiprime prMsK). + rewrite /part_b dprodE ?subsetIr //; last first. + rewrite setICA setIA (coprime_TIg (coprimeSg _ coMsU_K)) ?setI1g //. + by rewrite joing_subl. + rewrite -centC ?subsetIr // defNK //; split=> // X Eq1X. + have [q EqX] := nElemP Eq1X; have ntX := nt_pnElem EqX isT. + have:= EqX; rewrite -{1}(setIidPr sKE) pnElemI -setIdE. + case/setIdP=> EqEX sXK; split; first by rewrite defNK ?nK1K. + have [_ abelX dimX] := pnElemP EqX; have [qX _] := andP abelX. + have kq: q \in \kappa(M). + by rewrite (pnatPpi kK (piSg sXK _)) // -p_rank_gt0 p_rank_abelem ?dimX. + have nregX := kappa_nonregular kq (subsetP (pnElemS _ _ sEM) _ EqEX). + have sq := tau13_nonregular_sigma maxM hallE EqEX (kappa_tau13 kq) nregX. + move=> H maxNH; have [maxH sNXH] := setIdP maxNH. + rewrite (sub_Hall_pcore (Msigma_Hall maxH)) ?(subset_trans (normG X)) //. + exact: pi_pgroup qX (sq H maxNH). +have have_c: part_c. + pose p := pdiv #|E1|; have piE1p: p \in \pi(E1) by rewrite pi_pdiv cardG_gt1. + have:= piE1p; rewrite -p_rank_gt0 => /p_rank_geP[Y]. + rewrite -(setIidPr sE1M) pnElemI -setIdE => /setIdP[EpY sYE1]. + have [sYK ntY] := (subset_trans sYE1 sE1K, nt_pnElem EpY isT). + split=> [|X /nElemP[q]]. + rewrite /Ks -(cent_semiprime prMsK sYK) //. + exact: kappa_nonregular (pnatPpi kK (piSg sE1K piE1p)) EpY. + rewrite /= -(cent_semiprime prMsK sYK) // => EqX. + by have [] := cent_cent_Msigma_tau1_uniq maxM hallE hallE1 sYE1 ntY EqX. +have [[defNK defK1] [_ uniqKs]] := (have_b, have_c). +have have_d: part_d. + split=> g. + rewrite inE; apply: contraNeq; rewrite -rank_gt0. + case/rank_geP=> X; rewrite nElemI -setIdE -groupV => /setIdP[EpX sXMg]. + have [_ sCXMs] := mem_uniq_mmax (uniqKs _ EpX). + case/nElemP: EpX => p /pnElemP[/subsetIP[sXMs _] abelX dimX]. + have [[pX _] sXM] := (andP abelX, subset_trans sXMs (pcore_sub _ _)). + have piXp: p \in \pi(X) by rewrite -p_rank_gt0 p_rank_abelem ?dimX. + have sp := pnatPpi sMs (piSg sXMs piXp). + have [def_g _ _] := sigma_group_trans maxM sp pX. + have [|c cXc [m Mm ->]] := def_g g^-1 sXM; first by rewrite sub_conjgV. + by rewrite groupM // (subsetP sCXMs). + case/setDP=> Mg; apply: contraNeq; rewrite -rank_gt0 /=. + case/rank_geP=> X; rewrite nElemI -setIdE => /setIdP[EpX sXKg]. + have [<- _] := defK1 X EpX; rewrite 2!inE Mg /=. + have{EpX} [p EpX] := nElemP EpX; have [_ abelX dimX] := pnElemP EpX. + have defKp1: {in 'E_p^1(K), forall Y, 'Ohm_1('O_p(K)) = Y}. + move=> Y EpY; have E1K_Y: Y \in 'E^1(K) by apply/nElemP; exists p. + have piKp: p \in \pi(K) by rewrite -p_rank_gt0; apply/p_rank_geP; exists Y. + have [pKp sKpK] := (pcore_pgroup p K, pcore_sub p K). + have cycKp: cyclic 'O_p(K). + rewrite (odd_pgroup_rank1_cyclic pKp) ?mFT_odd //. + by rewrite -(rank_kappa (pnatPpi kK piKp)) p_rankS ?(subset_trans sKpK). + have [sYK abelY oY] := pnElemPcard EpY; have [pY _] := andP abelY. + have sYKp: Y \subset 'O_p(K) by rewrite pcore_max ?nK1K. + apply/eqP; rewrite eq_sym eqEcard -{1}(Ohm1_id abelY) OhmS //= oY. + rewrite (Ohm1_cyclic_pgroup_prime cycKp pKp) ?(subG1_contra sYKp) //=. + exact: nt_pnElem EpY _. + rewrite sub_conjg -[X :^ _]defKp1 ?(defKp1 X) //. + by rewrite !inE sub_conjgV sXKg abelemJ abelX cardJg dimX. +split=> {part_a part_b part_c have_a have_b have_c}//; first split=> //. +- move=> q; rewrite /Ks -(cent_semiprime prMsK sE1K ntE1) => picMsE1q. + have sq := pnatPpi (pcore_pgroup _ M) (piSg (subsetIl _ _) picMsE1q). + move: picMsE1q; rewrite -p_rank_gt0; case/p_rank_geP=> y EqY S sylS. + have [sSM qS _] := and3P sylS. + have sSMs: S \subset M`_\sigma. + by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) ?(pi_pgroup qS). + have sylS_Ms: q.-Sylow(M`_\sigma) S := pHall_subl sSMs (pcore_sub _ M) sylS. + have [_]:= cent_cent_Msigma_tau1_uniq maxM hallE hallE1 (subxx _) ntE1 EqY. + move/(_ S sylS_Ms) => uniqS; split=> //; rewrite subsetI sSMs /=. + pose p := pdiv #|E1|; have piE1p: p \in \pi(E1) by rewrite pi_pdiv cardG_gt1. + have [s'p _] := andP (pnatPpi t1E1 piE1p). + have [P sylP] := Sylow_exists p E1; have [sPE1 pP _] := and3P sylP. + apply: contra (s'p) => cE1S; apply/exists_inP; exists P. + exact: subHall_Sylow s'p (subHall_Sylow hallE1 (pnatPpi t1E1 piE1p) sylP). + apply: (sub_uniq_mmax uniqS); first by rewrite cents_norm // (centsS sPE1). + apply: mFT_norm_proper (mFT_pgroup_proper pP). + by rewrite -rank_gt0 (rank_Sylow sylP) p_rank_gt0. +- move=> Y sY ntYKs; have ntY: Y :!=:1 := subG1_contra (subsetIl _ _) ntYKs. + have [[x sYxMs] _] := sigma_Jsub maxM sY ntY; rewrite sub_conjg in sYxMs. + suffices Mx': x^-1 \in M by rewrite (normsP _ _ Mx') ?gFnorm in sYxMs. + rewrite -(setCK M) inE; apply: contra ntYKs => M'x'. + rewrite setIC -(setIidPr sYxMs) -/Ms -[Ms](setIidPr (pcore_sub _ _)). + by rewrite conjIg !setIA; have [-> // _] := have_d; rewrite !setI1g. +rewrite inE PmaxM andbT -(trivg_kappa_compl maxM complUK) => ntU. +have [regMsU nilMs]: 'C_Ms(U) = 1 /\ nilpotent Ms. + pose p := pdiv #|U|; have piUp: p \in \pi(U) by rewrite pi_pdiv cardG_gt1. + have sk'p := pnatPpi sk'U piUp. + have [S sylS] := Sylow_exists p U; have [sSU _] := andP sylS. + have sylS_M := subHall_Sylow hallU sk'p sylS. + rewrite -(setIidPr (centS (subset_trans (Ohm_sub 1 S) sSU))) setIA. + have [|_ _ -> ->] := sigma'_kappa'_facts maxM sylS_M; last by rewrite setI1g. + by rewrite -negb_or (piSg sUM). +have [[_ F _ defF] _ _ _] := sdprodP defM; rewrite defF in defM. +have hallMs: \sigma(M).-Hall(M) Ms by apply: Msigma_Hall. +have hallF: \sigma(M)^'.-Hall(M) F by apply/(sdprod_Hall_pcoreP hallMs). +have frF: [Frobenius F = U ><| K] by apply/Frobenius_semiregularP. +have ntMs: Ms != 1 by apply: Msigma_neq1. +have prK: prime #|K|. + have [solF [_ _ nMsF _]] := (sigma_compl_sol hallF, sdprodP defM). + have solMs: solvable Ms := solvableS (pcore_sub _ _) (mmax_sol maxM). + have coMsF: coprime #|Ms| #|F|. + by rewrite (coprime_sdprod_Hall_r defM) (pHall_Hall hallF). + by have [] := Frobenius_primact frF solF nMsF solMs ntMs coMsF prMsK. +have eq_sb: \sigma(M) =i \beta(M). + suffices bMs: \beta(M).-group Ms. + move=> p; apply/idP/idP=> [sp|]; last exact: beta_sub_sigma. + rewrite (pnatPpi bMs) //= (card_Hall (Msigma_Hall maxM)) pi_of_part //. + by rewrite inE /= sigma_sub_pi. + have [H hallH cHF'] := der_compl_cent_beta' maxM hallF. + rewrite -pgroupNK -partG_eq1 -(card_Hall hallH) -trivg_card1 -subG1. + rewrite -regMsU subsetI (pHall_sub hallH) centsC (subset_trans _ cHF') //. + have [solU [_ mulUK nUK _]] := (abelian_sol cUU, sdprodP defF). + have coUK: coprime #|U| #|K|. + rewrite (p'nat_coprime (sub_pgroup _ (pHall_pgroup hallU)) kK) // => p. + by case/norP. + rewrite -(coprime_cent_prod nUK) // (cent_semiregular regUK) // mulg1. + by rewrite -mulUK commgSS ?mulG_subl ?mulG_subr. +split=> //; apply/normedTI_P; rewrite setD_eq0 subG1 setTI normD1 gFnorm. +split=> // g _; rewrite -setI_eq0 conjD1g -setDIl setD_eq0 subG1 /= -/Ms. +have [_ _ b'MsMg] := sigma_compl_embedding maxM hallE. +apply: contraR => notMg; have{b'MsMg notMg} [_ b'MsMg _] := b'MsMg g notMg. +rewrite -{2}(setIidPr (pHall_sub hallMs)) conjIg setIA coprime_TIg // cardJg. +by apply: p'nat_coprime b'MsMg _; rewrite -(eq_pnat _ eq_sb). +Qed. + +(* This is essentially the skolemized form of 14.2(a). *) +Lemma kappa_compl_context M U K : + M \in 'M -> kappa_complement M U K -> + [/\ \sigma(M)^'.-Hall(M) (U <*> K), + M`_\sigma ><| (U ><| K) = M, + semiprime M`_\sigma K, + semiregular U K + & K :!=: 1 -> abelian U]. +Proof. +move=> maxM [hallU hallK gsetUK]; set E := U <*> K. +have mulUK: U * K = E by rewrite -(gen_set_id gsetUK) genM_join. +have [[sKM kK _] [sUM sk'U _]] := (and3P hallK, and3P hallU). +have tiUK: U :&: K = 1. + by apply: coprime_TIg (p'nat_coprime (sub_pgroup _ sk'U) kK) => p; case/norP. +have hallE: \sigma(M)^'.-Hall(M) E. + rewrite pHallE /= -/E -mulUK mul_subG //= TI_cardMg //. + rewrite -(partnC \kappa(M) (part_gt0 _ _)) (partn_part _ (@kappa_sigma' M)). + apply/eqP; rewrite -partnI -(card_Hall hallK) mulnC; congr (_ * _)%N. + by rewrite (card_Hall hallU); apply: eq_partn => p; apply: negb_or. +have [K1 | ntK] := altP (K :=P: 1). + rewrite K1 sdprodg1 -{1}(mulg1 U) -{1}K1 mulUK sdprod_sigma //. + by split=> //; first apply: semiregular_prime; apply: semiregular1r. +have PmaxM: M \in 'M_'P by rewrite inE maxM -(trivg_kappa maxM hallK) andbT. +have [[V [complV defM] [cVV prK regK]] _ _ _ _] := Ptype_structure PmaxM hallK. +have [[_ F _ defF] _ _ _] := sdprodP defM; rewrite defF in defM. +have hallF: \sigma(M)^'.-Hall(M) F. + exact/(sdprod_Hall_pcoreP (Msigma_Hall maxM)). +have [a Ma /= defFa] := Hall_trans (mmax_sol maxM) hallE hallF. +have [hallV _ _] := complV; set sk' := \sigma_kappa(M)^' in hallU hallV sk'U. +have [nsVF sKF _ _ _] := sdprod_context defF. +have [[[sVF _] [sFM _]] [sEM _]] := (andP nsVF, andP hallF, andP hallE). +have hallV_F: sk'.-Hall(F) V := pHall_subl sVF sFM hallV. +have hallU_E: sk'.-Hall(E) U := pHall_subl (joing_subl _ _) sEM hallU. +have defV: 'O_sk'(F) = V := normal_Hall_pcore hallV_F nsVF. +have hallEsk': sk'.-Hall(E) 'O_sk'(E). + by rewrite [E]defFa pcoreJ pHallJ2 /= defV. +have defU: 'O_sk'(E) = U by rewrite (eq_Hall_pcore hallEsk' hallU_E). +have nUE: E \subset 'N(U) by rewrite -defU gFnorm. +have hallK_E: \kappa(M).-Hall(E) K := pHall_subl (joing_subr _ _) sEM hallK. +have hallK_F: \kappa(M).-Hall(F) K := pHall_subl sKF sFM hallK. +have hallKa_E: \kappa(M).-Hall(E) (K :^ a) by rewrite [E]defFa pHallJ2. +have [b Eb /= defKab] := Hall_trans (sigma_compl_sol hallE) hallK_E hallKa_E. +have defVa: V :^ a = U by rewrite -defV -pcoreJ -defFa defU. +split=> // [| x Kx | _]; last by rewrite -defVa abelianJ. + by rewrite [U ><| K]sdprodEY ?sdprod_sigma //; case/joing_subP: nUE. +rewrite -(conjgKV (a * b) x) -(normsP nUE b Eb) -defVa -conjsgM. +rewrite -cent_cycle cycleJ centJ -conjIg cent_cycle regK ?conjs1g //. +by rewrite -mem_conjg conjD1g conjsgM -defKab. +Qed. + +(* This is B & G, Corollary 14.3. *) +Corollary pi_of_cent_sigma M x x' : + M \in 'M -> x \in (M`_\sigma)^# -> + x' \in ('C_M[x])^# -> \sigma(M)^'.-elt x' -> + (*1*) \kappa(M).-elt x' /\ 'C[x] \subset M + \/ (*2*) [/\ \tau2(M).-elt x', \ell_\sigma(x') == 1%N & 'M('C[x']) = [set M]]. +Proof. +move: x' => y maxM /setD1P[ntx Ms_x] /setD1P[nty cMxy] s'y. +have [My cxy] := setIP cMxy. +have [t2y | not_t2y] := boolP (\tau2(M).-elt y); [right | left]. + have uniqCy: 'M('C[y]) = [set M]; last split=> //. + apply: cent1_nreg_sigma_uniq; rewrite // ?inE ?nty //. + by apply/trivgPn; exists x; rewrite // inE Ms_x cent1C. + pose p := pdiv #[y]; have piYp: p \in \pi(#[y]) by rewrite pi_pdiv order_gt1. + have t2p := pnatPpi t2y piYp; have [E hallE] := ex_sigma_compl maxM. + have [A Ep2A _] := ex_tau2Elem hallE t2p. + have pA: p.-group A by case/pnElemP: Ep2A => _ /andP[]. + have ntA: A :!=: 1 by rewrite (nt_pnElem Ep2A). + have [H maxNH] := mmax_exists (mFT_norm_proper ntA (mFT_pgroup_proper pA)). + have [st2MsH _ _] := primes_norm_tau2Elem maxM hallE t2p Ep2A maxNH. + have [maxH _] := setIdP maxNH. + have sHy: \sigma(H).-elt y by apply: sub_p_elt t2y => q /st2MsH/andP[]. + rewrite /sigma_length (cardsD1 y.`_(\sigma(H))). + rewrite mem_sigma_decomposition //; last by rewrite constt_p_elt. + rewrite eqSS -sigma_decomposition_constt' //. + by apply/ell_sigma0P; rewrite (constt1P _) ?p_eltNK. +have{not_t2y} [p piYp t2'p]: exists2 p, p \in \pi(#[y]) & p \notin \tau2(M). + by apply/allPn; rewrite negb_and cardG_gt0 in not_t2y. +have sYM: <[y]> \subset M by rewrite cycle_subG. +have piMp: p \in \pi(M) := piSg sYM piYp. +have t13p: p \in [predU \tau1(M) & \tau3(M)]. + move: piMp; rewrite partition_pi_mmax // (negPf t2'p) /= orbA. + by case/orP=> // sMy; case/negP: (pnatPpi s'y piYp). +have [X]: exists X, X \in 'E_p^1(<[y]>) by apply/p_rank_geP; rewrite p_rank_gt0. +rewrite -(setIidPr sYM) pnElemI -setIdE => /setIdP[EpX sXy]. +have kp: p \in \kappa(M). + rewrite unlock; apply/andP; split=> //; apply/exists_inP; exists X => //. + apply/trivgPn; exists x; rewrite // inE Ms_x (subsetP (centS sXy)) //. + by rewrite cent_cycle cent1C. +have [sXM abelX dimX] := pnElemP EpX; have [pX _] := andP abelX. +have [K hallK sXK] := Hall_superset (mmax_sol maxM) sXM (pi_pgroup pX kp). +have PmaxM: M \in 'M_'P. + by rewrite 2!inE maxM andbT; apply: contraL kp => k'M; exact: (pnatPpi k'M). +have [_ [defNK defNX] [_ uniqCKs] _ _] := Ptype_structure PmaxM hallK. +have{defNX} sCMy_nMK: 'C_M[y] \subset 'N_M(K). + have [|<- _] := defNX X. + by apply/nElemP; exists p; rewrite !inE sXK abelX dimX. + by rewrite setIS // cents_norm // -cent_cycle centS. +have [[sMK kK _] [_ mulKKs cKKs _]] := (and3P hallK, dprodP defNK). +have s'K: \sigma(M)^'.-group K := sub_pgroup (@kappa_sigma' M) kK. +have sMs: \sigma(M).-group M`_\sigma := pcore_pgroup _ M. +have sKs: \sigma(M).-group 'C_(M`_\sigma)(K) := pgroupS (subsetIl _ _) sMs. +have{s'K sKs} [hallK_N hallKs] := coprime_mulGp_Hall mulKKs s'K sKs. +split. + rewrite (mem_p_elt kK) // (mem_normal_Hall hallK_N) ?normal_subnorm //. + by rewrite (subsetP sCMy_nMK) // inE My cent1id. +have Mx: x \in M := subsetP (pcore_sub _ _) x Ms_x. +have sxKs: <[x]> \subset 'C_(M`_\sigma)(K). + rewrite cycle_subG (mem_normal_Hall hallKs) ?(mem_p_elt sMs) //=. + by rewrite -mulKKs /normal mulG_subr mulG_subG normG cents_norm // centsC. + by rewrite (subsetP sCMy_nMK) // inE Mx cent1C. +have /rank_geP[Z]: 0 < 'r(<[x]>) by rewrite rank_gt0 cycle_eq1. +rewrite /= -(setIidPr sxKs) nElemI -setIdE => /setIdP[E1KsZ sZx]. +have [_ sCZM] := mem_uniq_mmax (uniqCKs Z E1KsZ). +by rewrite (subset_trans _ sCZM) // -cent_cycle centS. +Qed. + +(* This is B & G, Theorem 14.4. *) +(* We are omitting the first half of part (a), since we have taken it as our *) +(* definition of 'R[x]. *) +Theorem FT_signalizer_context x (N := 'N[x]) (R := 'R[x]) : + \ell_\sigma(x) == 1%N -> + [/\ [/\ [transitive R, on 'M_\sigma[x] | 'JG], + #|R| = #|'M_\sigma[x]|, + R <| 'C[x] + & Hall 'C[x] R] + & #|'M_\sigma[x]| > 1 -> + [/\ 'M('C[x]) = [set N], + (*a*) R :!=: 1, + (*c1*) \tau2(N).-elt x, + (*f*) N \in 'M_'F :|: 'M_'P2 + & {in 'M_\sigma[x], forall M, + [/\ (*b*) R ><| 'C_(M :&: N)[x] = 'C[x], + (*c2*) {subset \tau2(N) <= \sigma(M)}, + (*d*) {subset [predI \pi(M) & \sigma(N)] <= \beta(N)} + & (*e*) \sigma(N)^'.-Hall(N) (M :&: N)]}]]. +Proof. +rewrite {}/N {}/R => ell1x; have [ntx ntMSx] := ell_sigma1P x ell1x. +have [M MSxM] := set0Pn _ ntMSx; have [maxM Ms_x] := setIdP MSxM. +rewrite cycle_subG in Ms_x; have sMx := mem_p_elt (pcore_pgroup _ M) Ms_x. +have Mx: x \in M := subsetP (pcore_sub _ _) x Ms_x. +have [MSx_le1 | MSx_gt1] := leqP _ 1. + rewrite /'R[x] /'N[x] ltnNge MSx_le1 (trivgP (pcore_sub _ _)) setI1g normal1. + have <-: [set M] = 'M_\sigma[x]. + by apply/eqP; rewrite eqEcard sub1set MSxM cards1. + by rewrite /Hall atrans_acts_card ?imset_set1 ?cards1 ?sub1G ?coprime1n. +have [q pi_x_q]: exists q, q \in \pi(#[x]). + by exists (pdiv #[x]); rewrite pi_pdiv order_gt1. +have{sMx} sMq: q \in \sigma(M) := pnatPpi sMx pi_x_q. +have [X EqX]: exists X, X \in 'E_q^1(<[x]>). + by apply/p_rank_geP; rewrite p_rank_gt0. +have [sXx abelX dimX] := pnElemP EqX; have [qX cXX _] := and3P abelX. +have ntX: X :!=: 1 := nt_pnElem EqX isT. +have sXM: X \subset M by rewrite (subset_trans sXx) ?cycle_subG. +have [N maxNX_N] := mmax_exists (mFT_norm_proper ntX (mFT_pgroup_proper qX)). +have [maxN sNX_N] := setIdP maxNX_N; pose R := 'C_(N`_\sigma)[x]%G. +have sCX_N: 'C(X) \subset N := subset_trans (cent_sub X) sNX_N. +have sCx_N: 'C[x] \subset N by rewrite -cent_cycle (subset_trans (centS sXx)). +have sMSx_MSX: 'M_\sigma[x] \subset 'M_\sigma(X). + apply/subsetP=> M1 /setIdP[maxM1 sxM1]. + by rewrite inE maxM1 (subset_trans sXx). +have nsRCx: R <| 'C[x] by rewrite /= setIC (normalGI sCx_N) ?pcore_normal. +have hallR: \sigma(N).-Hall('C[x]) R. + exact: setI_normal_Hall (pcore_normal _ _) (Msigma_Hall maxN) sCx_N. +have transCX: [transitive 'C(X), on 'M_\sigma(X) | 'JG]. + have [_ trCX _ ] := sigma_group_trans maxM sMq qX. + case/imsetP: trCX => _ /setIdP[/imsetP[y _ ->] sXMy] trCX. + have maxMy: (M :^ y)%G \in 'M by rewrite mmaxJ. + have sXMys: X \subset (M :^ y)`_\sigma. + by rewrite (sub_Hall_pcore (Msigma_Hall _)) // (pi_pgroup qX) ?sigmaJ. + apply/imsetP; exists (M :^ y)%G; first exact/setIdP. + apply/setP=> Mz; apply/setIdP/imsetP=> [[maxMz sXMzs] | [z cXz -> /=]]. + suffices: gval Mz \in orbit 'Js 'C(X) (M :^ y). + by case/imsetP=> z CXz /group_inj->; exists z. + rewrite -trCX inE andbC (subset_trans sXMzs) ?pcore_sub //=. + apply/idPn => /(sigma_partition maxM maxMz)/=/(_ q)/idP[]. + rewrite inE /= sMq (pnatPpi (pgroupS sXMzs (pcore_pgroup _ _))) //. + by rewrite -p_rank_gt0 p_rank_abelem ?dimX. + by rewrite mmaxJ -(normP (subsetP (cent_sub X) z cXz)) MsigmaJ conjSg. +have MSX_M: M \in 'M_\sigma(X) := subsetP sMSx_MSX M MSxM. +have not_sCX_M: ~~ ('C(X) \subset M). + apply: contraL MSx_gt1 => sCX_M. + rewrite -leqNgt (leq_trans (subset_leq_card sMSx_MSX)) //. + rewrite -(atransP transCX _ MSX_M) card_orbit astab1JG. + by rewrite (setIidPl (normsG sCX_M)) indexgg. +have neqNM: N :!=: M by apply: contraNneq not_sCX_M => <-. +have maxNX'_N: N \in 'M('N(X)) :\ M by rewrite 2!inE neqNM. +have [notMGN _] := sigma_subgroup_embedding maxM sMq sXM qX ntX maxNX'_N. +have sN'q: q \notin \sigma(N). + by apply: contraFN (sigma_partition maxM maxN notMGN q) => sNq; exact/andP. +rewrite (negPf sN'q) => [[t2Nq s_piM_bN hallMN]]. +have defN: N`_\sigma ><| (M :&: N) = N. + exact/(sdprod_Hall_pcoreP (Msigma_Hall maxN)). +have Nx: x \in N by rewrite (subsetP sCx_N) ?cent1id. +have MNx: x \in M :&: N by rewrite inE Mx. +have sN'x: \sigma(N)^'.-elt x by rewrite (mem_p_elt (pHall_pgroup hallMN)). +have /andP[sNsN nNsN]: N`_\sigma <| N := pcore_normal _ _. +have nNs_x: x \in 'N(N`_\sigma) := subsetP nNsN x Nx. +have defCx: R ><| 'C_(M :&: N)[x] = 'C[x]. + rewrite -{2}(setIidPr sCx_N) /= -cent_cycle (subcent_sdprod defN) //. + by rewrite subsetI andbC normsG ?cycle_subG. +have transR: [transitive R, on 'M_\sigma[x] | 'JG]. + apply/imsetP; exists M => //; apply/setP=> L. + apply/idP/imsetP=> [MSxL | [u Ru ->{L}]]; last first. + have [_ cxu] := setIP Ru; rewrite /= -cent_cycle in cxu. + by rewrite -(normsP (cent_sub _) u cxu) sigma_mmaxJ. + have [u cXu defL] := atransP2 transCX MSX_M (subsetP sMSx_MSX _ MSxL). + have [_ mulMN nNsMN tiNsMN] := sdprodP defN. + have:= subsetP sCX_N u cXu; rewrite -mulMN -normC //. + case/imset2P=> v w /setIP[Mv _] Ns_w def_u; exists w => /=; last first. + by apply: group_inj; rewrite defL /= def_u conjsgM (conjGid Mv). + rewrite inE Ns_w -groupV (sameP cent1P commgP) -in_set1 -set1gE -tiNsMN. + rewrite setICA setIC (setIidPl sNsN) inE groupMl ?groupV //. + rewrite memJ_norm // groupV Ns_w /= -(norm_mmax maxM) inE sub_conjg. + rewrite invg_comm -(conjSg _ _ w) -{2}(conjGid Mx) -!conjsgM -conjg_Rmul. + rewrite -conjgC conjsgM -(conjGid Mv) -(conjsgM M) -def_u. + rewrite -[M :^ u](congr_group defL) conjGid // -cycle_subG. + by have [_ Ls_x] := setIdP MSxL; rewrite (subset_trans Ls_x) ?pcore_sub. +have oR: #|R| = #|'M_\sigma[x]|. + rewrite -(atransP transR _ MSxM) card_orbit astab1JG (norm_mmax maxM). + rewrite -setIAC /= -{3}(setIidPl sNsN) -(setIA _ N) -(setIC M). + by have [_ _ _ ->] := sdprodP defN; rewrite setI1g indexg1. +have ntR: R :!=: 1 by rewrite -cardG_gt1 oR. +have [y Ns_y CNy_x]: exists2 y, y \in (N`_\sigma)^# & x \in ('C_N[y])^#. + have [y Ry nty] := trivgPn _ ntR; have [Ns_y cxy] := setIP Ry. + by exists y; rewrite 2!inE ?nty // inE Nx cent1C ntx. +have kN'q: q \notin \kappa(N). + rewrite (contra (@kappa_tau13 N q)) // negb_or (contraL (@tau2'1 _ _ _)) //. + exact: tau3'2. +have [[kNx _] | [t2Nx _ uniqN]] := pi_of_cent_sigma maxN Ns_y CNy_x sN'x. + by case/idPn: (pnatPpi kNx pi_x_q). +have defNx: 'N[x] = N. + apply/set1P; rewrite -uniqN /'N[x] MSx_gt1. + by case: pickP => // /(_ N); rewrite uniqN /= set11. +rewrite /'R[x] {}defNx -(erefl (gval R)) (pHall_Hall hallR). +split=> // _; split=> // [|L MSxL]. + rewrite !(maxN, inE) /=; case: (pgroup _ _) => //=; rewrite andbT. + apply: contra kN'q => skN_N; have:= pnatPpi (mem_p_elt skN_N Nx) pi_x_q. + by case/orP=> //=; rewrite (negPf sN'q). +have [u Ru ->{L MSxL}] := atransP2 transR MSxM MSxL; rewrite /= cardJg. +have [Ns_u cxu] := setIP Ru; have Nu := subsetP sNsN u Ns_u. +rewrite -{1}(conjGid Ru) -(conjGid cxu) -{1 6 7}(conjGid Nu) -!conjIg pHallJ2. +split=> // [|p t2Np]. + rewrite /= -(setTI 'C[x]) -!(setICA setT) -!morphim_conj. + exact: injm_sdprod (subsetT _) (injm_conj _ _) defCx. +have [A Ep2A _] := ex_tau2Elem hallMN t2Np. +have [[nsAMN _] _ _ _] := tau2_compl_context maxN hallMN t2Np Ep2A. +have{Ep2A} Ep2A: A \in 'E_p^2(M) by move: Ep2A; rewrite pnElemI; case/setIP. +have rpM: 'r_p(M) > 1 by apply/p_rank_geP; exists A. +have: p \in \pi(M) by rewrite -p_rank_gt0 ltnW. +rewrite sigmaJ partition_pi_mmax // !orbA; case/orP=> //. +rewrite orbAC -2!andb_orr -(subnKC rpM) andbF /= => t2Mp. +case/negP: ntX; rewrite -subG1 (subset_trans sXx) //. +have [_ _ <- _ _] := tau2_context maxM t2Mp Ep2A. +have [[sAM abelA _] [_ nAMN]] := (pnElemP Ep2A, andP nsAMN). +rewrite -coprime_norm_cent ?(subset_trans sAM) ?gFnorm //. + by rewrite cycle_subG inE Ms_x (subsetP nAMN). +have [[sM'p _] [pA _]] := (andP t2Mp, andP abelA). +exact: pnat_coprime (pcore_pgroup _ _) (pi_pgroup pA sM'p). +Qed. + +(* A useful supplement to Theorem 14.4. *) +Lemma cent1_sub_uniq_sigma_mmax x M : + #|'M_\sigma[x]| == 1%N -> M \in 'M_\sigma[x] -> 'C[x] \subset M. +Proof. +move: M => M0 /cards1P[M defMSx] MS_M0; move: MS_M0 (MS_M0). +rewrite {1}defMSx => /set1P->{M0} MSxM; have [maxM _] := setIdP MSxM. +rewrite -(norm_mmax maxM); apply/normsP=> y cxy; apply: congr_group. +by apply/set1P; rewrite -defMSx -(mulKg y x) (cent1P cxy) cycleJ sigma_mmaxJ. +Qed. + +Remark cent_FT_signalizer x : x \in 'C('R[x]). +Proof. by rewrite -sub_cent1 subsetIr. Qed. + +(* Because the definition of 'N[x] uses choice, we can only prove it commutes *) +(* with conjugation now that we have established that the choice is unique. *) +Lemma FT_signalizer_baseJ x z : 'N[x ^ z] :=: 'N[x] :^ z. +Proof. +case MSx_gt1: (#|'M_\sigma[x]| > 1); last first. + by rewrite /'N[x] /'N[_] cycleJ card_sigma_mmaxJ MSx_gt1 conjs1g. +have [x1 | ntx] := eqVneq x 1. + rewrite x1 conj1g /'N[1] /= norm1. + case: pickP => [M maxTM | _]; last by rewrite if_same conjs1g. + by have [maxM] := setIdP maxTM; case/idPn; rewrite proper_subn ?mmax_proper. +apply: congr_group; apply/eqP; rewrite eq_sym -in_set1. +have ell1xz: \ell_\sigma(x ^ z) == 1%N. + by rewrite ell_sigmaJ; apply/ell_sigma1P; rewrite -cards_eq0 -lt0n ltnW. +have [_ [|<- _ _ _ _]] := FT_signalizer_context ell1xz. + by rewrite cycleJ card_sigma_mmaxJ. +rewrite -conjg_set1 normJ mmax_ofJ; rewrite ell_sigmaJ in ell1xz. +by have [_ [//|-> _ _ _ _]] := FT_signalizer_context ell1xz; apply: set11. +Qed. + +Lemma FT_signalizerJ x z : 'R[x ^ z] :=: 'R[x] :^ z. +Proof. +by rewrite /'R[x] /'R[_] FT_signalizer_baseJ MsigmaJ -conjg_set1 normJ conjIg. +Qed. + +Lemma sigma_coverJ x z : x ^ z *: 'R[x ^ z] = (x *: 'R[x]) :^ z. +Proof. by rewrite FT_signalizerJ conjsMg conjg_set1. Qed. + +Lemma sigma_supportJ M z : (M :^ z)^~~ = M^~~ :^ z. +Proof. +rewrite -bigcupJ /_^~~ MsigmaJ -conjD1g (big_imset _ (in2W (act_inj 'J z))) /=. +by apply: eq_bigr => x _; rewrite sigma_coverJ. +Qed. + +(* This is the remark imediately above B & G, Lemma 14.5; note the adjustment *) +(* allowing for the case x' = 1. *) +Remark sigma_cover_decomposition x x' : + \ell_\sigma(x) == 1%N -> x' \in 'R[x] -> + sigma_decomposition (x * x') = x |: [set x']^#. +Proof. +move=> ell1x Rx'; have [-> | ntx'] := eqVneq x' 1. + by rewrite mulg1 setDv setU0 ell1_decomposition. +rewrite setDE (setIidPl _) ?sub1set ?inE // setUC. +have ntR: #|'R[x]| > 1 by rewrite cardG_gt1; apply/trivgPn; exists x'. +have [Ns_x' cxx'] := setIP Rx'; move/cent1P in cxx'. +have [[_ <- _ _] [//| maxN _ t2Nx _ _]] := FT_signalizer_context ell1x. +have{maxN} [maxN _] := mem_uniq_mmax maxN. +have sNx' := mem_p_elt (pcore_pgroup _ _) Ns_x'. +have sN'x: \sigma('N[x])^'.-elt x by apply: sub_p_elt t2Nx => p /andP[]. +have defx': (x * x').`_\sigma('N[x]) = x'. + by rewrite consttM // (constt1P sN'x) mul1g constt_p_elt. +have sd_xx'_x': x' \in sigma_decomposition (x * x'). + by rewrite 2!inE ntx' -{1}defx'; apply: mem_imset. +rewrite -(setD1K sd_xx'_x') -{3}defx' -sigma_decomposition_constt' ?consttM //. +by rewrite constt_p_elt // (constt1P _) ?p_eltNK ?mulg1 // ell1_decomposition. +Qed. + +(* This is the simplified form of remark imediately above B & G, Lemma 14.5. *) +Remark nt_sigma_cover_decomposition x x' : + \ell_\sigma(x) == 1%N -> x' \in 'R[x]^# -> + sigma_decomposition (x * x') = [set x; x']. +Proof. +move=> ell1x /setD1P[ntx' Rx']; rewrite sigma_cover_decomposition //. +by rewrite setDE (setIidPl _) ?sub1set ?inE // setUC. +Qed. + +Remark mem_sigma_cover_decomposition x g : + \ell_\sigma(x) == 1%N -> g \in x *: 'R[x] -> x \in sigma_decomposition g. +Proof. +by move=> ell1x /lcosetP[x' Rx' ->]; rewrite sigma_cover_decomposition ?setU11. +Qed. + +Remark ell_sigma_cover x g : + \ell_\sigma(x) == 1%N -> g \in x *: 'R[x] -> \ell_\sigma(g) <= 2. +Proof. +move=> ell1x /lcosetP[x' Rx' ->]. +rewrite /(\ell_\sigma(_)) sigma_cover_decomposition // cardsU1. +by rewrite (leq_add (leq_b1 _)) // -(cards1 x') subset_leq_card ?subsetDl. +Qed. + +Remark ell_sigma_support M g : M \in 'M -> g \in M^~~ -> \ell_\sigma(g) <= 2. +Proof. +by move=> maxM /bigcupP[x Msx]; apply: ell_sigma_cover; apply: Msigma_ell1 Msx. +Qed. + +(* This is B & G, Lemma 14.5(a). *) +Lemma sigma_cover_disjoint x y : + \ell_\sigma(x) == 1%N -> \ell_\sigma(y) == 1%N -> x != y -> + [disjoint x *: 'R[x] & y *: 'R[y]]. +Proof. +move=> ell1x ell1y neq_xy; apply/pred0P=> g /=. +have [[ntx _] [nty _]] := (ell_sigma1P x ell1x, ell_sigma1P y ell1y). +apply: contraNF (ntx) => /andP[/lcosetP[x' Rxx' ->{g}] /= yRy_xx']. +have def_y: y = x'. + apply: contraTeq (mem_sigma_cover_decomposition ell1y yRy_xx') => neq_yx'. + by rewrite sigma_cover_decomposition // !inE negb_or nty eq_sym neq_xy. +have [[_ <- _ _] [|uniqCx _ _ _ _]] := FT_signalizer_context ell1x. + by rewrite cardG_gt1; apply/trivgPn; exists x'; rewrite // -def_y. +have{uniqCx} [maxNx sCxNx] := mem_uniq_mmax uniqCx. +have Rx_y: y \in 'R[x] by [rewrite def_y]; have [Nxs_y cxy] := setIP Rx_y. +have Ry_x: x \in 'R[y]. + by rewrite -def_y -(cent1P cxy) mem_lcoset mulKg in yRy_xx'. +have MSyNx: 'N[x] \in 'M_\sigma[y] by rewrite inE maxNx cycle_subG. +have [[_ <- _ _] [|uniqCy _ _ _]] := FT_signalizer_context ell1y. + by rewrite cardG_gt1; apply/trivgPn; exists x. +have{uniqCy} [_ sCyNy] := mem_uniq_mmax uniqCy. +case/(_ 'N[x] MSyNx)=> /sdprodP[_ _ _ tiRyNx] _ _ _. +rewrite -in_set1 -set1gE -tiRyNx -setIA (setIidPr sCyNy) inE Ry_x /=. +by rewrite inE cent1C cxy (subsetP sCxNx) ?cent1id. +Qed. + +(* This is B & G, Lemma 14.5(b). *) +Lemma sigma_support_disjoint M1 M2 : + M1 \in 'M -> M2 \in 'M -> gval M2 \notin M1 :^: G -> [disjoint M1^~~ & M2^~~]. +Proof. +move=> maxM1 maxM2 notM1GM2; rewrite -setI_eq0 -subset0 big_distrl. +apply/bigcupsP=> x M1s_x; rewrite big_distrr; apply/bigcupsP=> y M2s_y /=. +have [ell1x ell1y] := (Msigma_ell1 maxM1 M1s_x, Msigma_ell1 maxM2 M2s_y). +rewrite subset0 setI_eq0 sigma_cover_disjoint //. +have{M1s_x M2s_y}[[ntx M1s_x] [_ M2s_y]] := (setD1P M1s_x, setD1P M2s_y). +pose p := pdiv #[x]; have pixp: p \in \pi(#[x]) by rewrite pi_pdiv order_gt1. +apply: contraFN (sigma_partition maxM1 maxM2 notM1GM2 p) => eq_xy. +rewrite inE /= (pnatPpi (mem_p_elt (pcore_pgroup _ _) M1s_x)) //=. +by rewrite (pnatPpi (mem_p_elt (pcore_pgroup _ _) M2s_y)) -?(eqP eq_xy). +Qed. + +(* This is B & G, Lemma 14.5(c). *) +Lemma card_class_support_sigma M : + M \in 'M -> #|class_support M^~~ G| = (#|M`_\sigma|.-1 * #|G : M|)%N. +Proof. +move=> maxM; rewrite [#|M`_\sigma|](cardsD1 1) group1 /=. +set MsG := class_support (M`_\sigma)^# G; pose P := [set x *: 'R[x] | x in MsG]. +have ellMsG x: x \in MsG -> \ell_\sigma(x) == 1%N. + by case/imset2P=> y z My _ ->; rewrite ell_sigmaJ (Msigma_ell1 maxM). +have tiP: trivIset P. + apply/trivIsetP=> _ _ /imsetP[x MsGx ->] /imsetP[y MsGy ->] neq_xRyR. + by rewrite sigma_cover_disjoint ?ellMsG //; apply: contraNneq neq_xRyR => ->. +have->: class_support M^~~ G = cover P. + apply/setP=> az; apply/imset2P/bigcupP=> [[a z] | [xRz]]. + case/bigcupP=> x Ms_x xRa Gz ->; exists (x ^ z *: 'R[x ^ z]). + by apply: mem_imset; exact: mem_imset2. + by rewrite sigma_coverJ memJ_conjg. + case/imsetP=> _ /imset2P[x z Ms_x Gz ->] ->; rewrite sigma_coverJ. + by case/imsetP=> a xRa ->; exists a z => //; apply/bigcupP; exists x. +rewrite -(eqnP tiP) big_imset /= => [|x y MsGx MsGy eq_xyR]; last first. + have: x *: 'R[x] != set0 by rewrite -cards_eq0 -lt0n card_lcoset cardG_gt0. + rewrite -[x *: _]setIid {2}eq_xyR setI_eq0. + by apply: contraNeq => neq_xy; rewrite sigma_cover_disjoint ?ellMsG. +rewrite -{2}(norm_mmax maxM) -astab1JG -indexgI -card_orbit. +set MG := orbit _ G M; rewrite mulnC -sum_nat_const. +transitivity (\sum_(Mz in MG) \sum_(x in (Mz`_\sigma)^#) 1); last first. + apply: eq_bigr => _ /imsetP[z _ ->]; rewrite sum1_card MsigmaJ. + by rewrite -conjD1g cardJg. +rewrite (exchange_big_dep (mem MsG)) /= => [|Mz xz]; last first. + case/imsetP=> z Gz ->; rewrite MsigmaJ -conjD1g => /imsetP[x Ms_x ->{xz}]. + exact: mem_imset2. +apply: eq_bigr => x MsGx; rewrite card_lcoset sum1dep_card. +have ell1x := ellMsG x MsGx; have [ntx _] := ell_sigma1P x ell1x. +have [[transRx -> _ _] _] := FT_signalizer_context ell1x. +apply: eq_card => Mz; rewrite 2!inE cycle_subG in_setD1 ntx /=. +apply: andb_id2r => Mzs_x. +apply/idP/imsetP=> [maxMz | [z _ ->]]; last by rewrite mmaxJ. +have [y t Ms_y _ def_x] := imset2P MsGx; have{Ms_y} [_ Ms_y] := setD1P Ms_y. +have [MSxMz MSxMt]: Mz \in 'M_\sigma[x] /\ (M :^ t)%G \in 'M_\sigma[x]. + by rewrite {2}def_x cycleJ sigma_mmaxJ inE maxMz inE maxM !cycle_subG. +have [z _ ->] := atransP2 transRx MSxMt MSxMz. +by exists (t * z); rewrite ?inE ?actM. +Qed. + +(* This is B & G, Lemma 14.6. *) +Lemma sigma_decomposition_dichotomy (g : gT) : + g != 1 -> + [exists (x | \ell_\sigma(x) == 1%N), x^-1 * g \in 'R[x]] + (+) [exists (y | \ell_\sigma(y) == 1%N), + let y' := y^-1 * g in + [exists M in 'M_\sigma[y], (y' \in ('C_M[y])^#) && \kappa(M).-elt y']]. +Proof. +move=> ntg; have [[x ell1x Rx'] | ] := altP exists_inP. + rewrite /= negb_exists_in; apply/forall_inP=> y ell1y. + set y' := y^-1 * g; set x' := x^-1 * g in Rx'. + apply/existsP=> -[M /and3P[MSyM CMy_y' kMy']]. + have [maxM Ms_y] := setIdP MSyM; rewrite cycle_subG in Ms_y. + have [nty'] := setD1P CMy_y'; case/setIP=> My'; move/cent1P=> cyy'. + have [[nty _] sMy]:= (ell_sigma1P y ell1y, mem_p_elt (pcore_pgroup _ _) Ms_y). + have sM'y': \sigma(M)^'.-elt y' := sub_p_elt (@kappa_sigma' M) kMy'. + have t2M'y': \tau2(M)^'.-elt y'. + apply: sub_p_elt kMy' => p; move/kappa_tau13. + by case/orP; [apply: tau2'1 | apply: contraL; apply: tau3'2]. + have xx'_y: y \in pred2 x x'. + suffices: y \in x |: [set x']^# by rewrite !inE nty. + rewrite -sigma_cover_decomposition // mulKVg 2!inE nty /=. + apply/imsetP; exists M => //; rewrite -(mulKVg y g) -/y' consttM //. + by rewrite (constt_p_elt sMy) (constt1P sM'y') mulg1. + have nt_x': x' != 1 by case/pred2P: xx'_y; rewrite /x' => <-. + have maxCY_M: M \in 'M('C[y]). + have Ms1_y: y \in (M`_\sigma)^# by apply/setD1P. + rewrite inE maxM; case/pi_of_cent_sigma: CMy_y' => // [[//] | [t2y']]. + by rewrite -order_eq1 (pnat_1 t2y' t2M'y') in nty'. + have [[_ <- _ _] [|uniqNx _ t2Nx _ _]] := FT_signalizer_context ell1x. + by rewrite cardG_gt1; apply/trivgPn; exists x'. + rewrite -order_gt1 (pnat_1 sMy _) // -/(_.-elt _) in nty. + have{xx'_y} [eq_yx | eq_yx']: y = x \/ y = x' := pred2P xx'_y. + rewrite eq_yx uniqNx in maxCY_M *; rewrite (set1P maxCY_M). + by apply: sub_p_elt t2Nx => p; case/andP. + have eq_xy': x = y' by apply: (mulIg y); rewrite cyy' {1}eq_yx' !mulKVg. + have [[z _ defM] | notMGNx] := altP (@orbitP _ _ _ 'Js G 'N[x] M). + rewrite -order_eq1 (pnat_1 _ t2M'y') // in nty'. + by rewrite -defM (eq_pnat _ (tau2J _ _)) -eq_xy'. + have Ns_y: y \in 'N[x]`_\sigma by rewrite eq_yx'; case/setIP: Rx'. + apply: sub_p_elt (mem_p_elt (pcore_pgroup _ _) Ns_y) => p sNp. + have [maxN _] := mem_uniq_mmax uniqNx. + by apply: contraFN (sigma_partition _ _ notMGNx p) => // sMp; apply/andP. +rewrite negb_exists_in => /forall_inP not_sign_g. +apply: wlog_neg; rewrite negb_exists_in => /forall_inP not_kappa_g. +have s'g M: M \in 'M -> g \in M -> g.`_\sigma(M) = 1. + move=> maxM; set x := g.`_\sigma(M); pose x' := g.`_(\sigma(M))^'. + have def_x': x^-1 * g = x' by rewrite -(consttC \sigma(M) g) mulKg. + apply: contraTeq => ntx. + have ell1x: \ell_\sigma(x) == 1%N. + rewrite /sigma_length (cardsD1 x.`_\sigma(M)). + rewrite -sigma_decomposition_constt' // mem_sigma_decomposition //. + by apply/ell_sigma0P; apply/constt1P; rewrite p_eltNK p_elt_constt. + by rewrite sub_in_constt // => ?. + apply: contra (not_sign_g _ ell1x) => Mg; rewrite def_x'. + have [-> | ntx'] := eqVneq x' 1; first exact: group1. + have cxx': x \in 'C[x'] by apply/cent1P; apply: commuteX2. + have cMx_x': x' \in ('C_M[x])^# by rewrite 3!inE ntx' cent1C cxx' groupX. + have Ms_x: x \in M`_\sigma. + by rewrite (mem_Hall_pcore (Msigma_Hall maxM)) ?p_elt_constt ?groupX. + have Ms1x: x \in (M`_\sigma)^# by apply/setD1P. + have sM'x': (\sigma(M))^'.-elt x' := p_elt_constt _ _. + have [[kMx' _] | [_ ell1x' uniqM]] := pi_of_cent_sigma maxM Ms1x cMx_x' sM'x'. + case/existsP: (not_kappa_g _ ell1x); exists M; rewrite def_x' cMx_x' /=. + by rewrite inE maxM cycle_subG Ms_x. + have MSx'_gt1: #|'M_\sigma[x']| > 1. + have [_ ntMSx'] := ell_sigma1P _ ell1x'. + rewrite ltn_neqAle lt0n cards_eq0 ntMSx' andbT eq_sym. + apply: contra ntx' => MSx'_eq1; rewrite -order_eq1 (pnat_1 _ sM'x') //. + have [N MSx'N] := set0Pn _ ntMSx'; have [maxN Ns_x'] := setIdP MSx'N. + rewrite -(eq_uniq_mmax uniqM maxN) ?cent1_sub_uniq_sigma_mmax //. + exact: pgroupS Ns_x' (pcore_pgroup _ _). + have defNx': 'N[x'] = M. + by apply: set1_inj; case/FT_signalizer_context: ell1x' => _ [|<-]. + case/negP: (not_sign_g _ ell1x'). + by rewrite -(consttC \sigma(M)^' g) mulKg consttNK inE defNx' Ms_x. +have [x sg_x]: exists x, x \in sigma_decomposition g. + by apply/set0Pn; rewrite -cards_eq0 (sameP (ell_sigma0P g) eqP). +have{sg_x} [ntx /imsetP[M maxM def_x]] := setD1P sg_x. +wlog MSxM: M maxM def_x / M \in 'M_\sigma[x]. + have sMx: \sigma(M).-elt x by rewrite def_x p_elt_constt. + have [|[z Ms_xz] _] := sigma_Jsub maxM sMx; first by rewrite cycle_eq1. + move/(_ (M :^ z^-1)%G)->; rewrite ?mmaxJ ?(eq_constt _ (sigmaJ M _)) //. + by rewrite inE mmaxJ maxM MsigmaJ -sub_conjg. +have ell1x: \ell_\sigma(x) == 1%N. + by apply/ell_sigma1P; split=> //; apply/set0Pn; exists M. +have notMg: g \notin M by apply: contra ntx; rewrite def_x; move/s'g->. +have cxg: g \in 'C[x] by rewrite cent1C def_x groupX ?cent1id. +have MSx_gt1: #|'M_\sigma[x]| > 1. + rewrite ltnNge; apply: contra notMg => MSx_le1; apply: subsetP cxg. + have [_ ntMSx] := ell_sigma1P _ ell1x. + by rewrite cent1_sub_uniq_sigma_mmax // eqn_leq MSx_le1 lt0n cards_eq0. +have [_ [//|defNx _ _ _]] := FT_signalizer_context ell1x. +case/(_ M)=> // _ _ _ hallMN; have [maxN sCxN] := mem_uniq_mmax defNx. +have Ng: <[g]> \subset 'N[x] by rewrite cycle_subG (subsetP sCxN). +have sN'g: \sigma('N[x])^'.-elt g by apply/constt1P; rewrite s'g // -cycle_subG. +have [z _ MNgz] := Hall_subJ (mmax_sol maxN) hallMN Ng sN'g. +case/eqP: ntx; rewrite def_x -(eq_constt _ (sigmaJ M z)) s'g ?mmaxJ //. +by move: MNgz; rewrite conjIg cycle_subG => /setIP[]. +Qed. + +Section PTypeEmbedding. +Implicit Types Mi Mj : {group gT}. +Implicit Type Ks : {set gT}. + +(* This is B & G, Theorem 14.7. *) +(* This theorem provides the basis for the maximal subgroup classification, *) +(* the main output of the local analysis. Note that we handle differently the *) +(* two separate instances of non-structural proof (by analogy) that occur in *) +(* the textbook, p. 112, l. 7 and p. 113, l. 22. For the latter we simply use *) +(* global induction on the size of the class support of the TI-set \hat{Z} *) +(* (for this reason we have kept the assertion that this is greater than half *) +(* of the size of G, even though this is not used later in the proof; we did *) +(* drop the more precise lower bound). For the former we prove a preliminary *) +(* lemma that summarizes the four results of the beginning of the proof that *) +(* used after p. 112, l. 7 -- note that this also gets rid of a third non *) +(* structural argument (on p. 112, l. 5). *) +(* Also, note that the direct product decomposition of Z and the K_i, and *) +(* its direct relation with the sigma-decomposition of elements of Z (p. 112, *) +(* l. 13-19) is NOT materially used in the rest of the argument, though it *) +(* does obviously help a human reader forge a mental picture of the situation *) +(* at hand. Only the first remark, l. 13, is used to prove the alternative *) +(* definition of T implicit in the remarks l. 22-23. Accordingly, we have *) +(* suppressed most of these intermediate results: we have only kept the proof *) +(* that Z is the direct product of the K_i^*, though we discard this result *) +(* immediately (its 24-line proof just nudges the whole proof size slightyly *) +(* over the 600-line bar). *) +Theorem Ptype_embedding M K : + M \in 'M_'P -> \kappa(M).-Hall(M) K -> + exists2 Mstar, Mstar \in 'M_'P /\ gval Mstar \notin M :^: G + & let Kstar := 'C_(M`_\sigma)(K) in + let Z := K <*> Kstar in let Zhat := Z :\: (K :|: Kstar) in + [/\ (*a*) {in 'E^1(K), forall X, 'M('C(X)) = [set Mstar]}, + (*b*) \kappa(Mstar).-Hall(Mstar) Kstar /\ \sigma(M).-Hall(Mstar) Kstar, + (*c*) 'C_(Mstar`_\sigma)(Kstar) = K /\ \kappa(M) =i \tau1(M), + (*d*) [/\ cyclic Z, M :&: Mstar = Z, + {in K^#, forall x, 'C_M[x] = Z}, + {in Kstar^#, forall y, 'C_Mstar[y] = Z} + & {in K^# & Kstar^#, forall x y, 'C[x * y] = Z}] +& [/\ (*e*) [/\ normedTI Zhat G Z, {in ~: M, forall g, [disjoint Zhat & M :^ g]} + & (#|G|%:R / 2%:R < #|class_support Zhat G|%:R :> rat)%R ], + (*f*) M \in 'M_'P2 /\ prime #|K| \/ Mstar \in 'M_'P2 /\ prime #|Kstar|, + (*g*) {in 'M_'P, forall H, gval H \in M :^: G :|: Mstar :^: G} + & (*h*) M^`(1) ><| K = M]]. +Proof. +pose isKi Ks M K := [&& M \in 'M_'P, \kappa(M).-Hall(M) K & Ks \subset K]. +move: M K; have Pmax_sym M K X (Ks := 'C_(M`_\sigma)(K)) (Z := K <*> Ks) Mi : + M \in 'M_'P -> \kappa(M).-Hall(M) K -> X \in 'E^1(K) -> Mi \in 'M('N(X)) -> + [/\ Z \subset Mi, gval Mi \notin M :^: G, exists Ki, isKi Ks Mi Ki + & {in 'E^1(Ks), forall Xs, Z \subset 'N_Mi(gval Xs)}]. +- move=> PmaxM hallK E1X maxNMi. + have [[_ maxM] [maxMi sNXMi]] := (setIdP PmaxM, setIdP maxNMi). + have [_ [defNK defNX] [ntKs uniqCKs] _ _] := Ptype_structure PmaxM hallK. + rewrite -/Ks in defNK ntKs uniqCKs; have [_ mulKKs cKKs _] := dprodP defNK. + have{mulKKs} defZ: 'N_M(K) = Z by rewrite -mulKKs -cent_joinEr. + have sZMi: Z \subset Mi. + by rewrite -defZ; have [<- _] := defNX X E1X; rewrite setIC subIset ?sNXMi. + have [sKMi sKsMi] := joing_subP sZMi. + have sXMis: X \subset Mi`_\sigma by have [_ ->] := defNX X E1X. + have sMiX: \sigma(Mi).-group X := pgroupS sXMis (pcore_pgroup _ _). + have [q EqX] := nElemP E1X; have [sXK abelX dimX] := pnElemP EqX. + have piXq: q \in \pi(X) by rewrite -p_rank_gt0 p_rank_abelem ?dimX. + have notMGMi: gval Mi \notin M :^: G. + apply: contraL (pnatPpi sMiX piXq); case/imsetP=> a _ ->; rewrite sigmaJ. + exact: kappa_sigma' (pnatPpi (pHall_pgroup hallK) (piSg sXK piXq)). + have kMiKs: \kappa(Mi).-group Ks. + apply/pgroupP=> p p_pr /Cauchy[] // xs Ks_xs oxs. + pose Xs := <[xs]>%G; have sXsKs: Xs \subset Ks by rewrite cycle_subG. + have EpXs: Xs \in 'E_p^1(Ks) by rewrite p1ElemE // !inE sXsKs -oxs /=. + have sMi'Xs: \sigma(Mi)^'.-group Xs. + rewrite /pgroup /= -orderE oxs pnatE //=. + apply: contraFN (sigma_partition maxM maxMi notMGMi p) => /= sMi_p. + rewrite inE /= sMi_p -pnatE // -oxs andbT. + exact: pgroupS sXsKs (pgroupS (subsetIl _ _) (pcore_pgroup _ _)). + have uniqM: 'M('C(Xs)) = [set M] by apply: uniqCKs; apply/nElemP; exists p. + have [x Xx ntx] := trivgPn _ (nt_pnElem EqX isT). + have Mis_x: x \in (Mi`_\sigma)^# by rewrite !inE ntx (subsetP sXMis). + have CMix_xs: xs \in ('C_Mi[x])^#. + rewrite 2!inE -order_gt1 oxs prime_gt1 // inE -!cycle_subG. + rewrite (subset_trans sXsKs) //= sub_cent1 (subsetP _ x Xx) //. + by rewrite centsC (centSS sXsKs sXK). + have{sMi'Xs} [|[_ _]] := pi_of_cent_sigma maxMi Mis_x CMix_xs sMi'Xs. + by case; rewrite /p_elt oxs pnatE. + case/mem_uniq_mmax=> _ sCxsMi; case/negP: notMGMi. + by rewrite -(eq_uniq_mmax uniqM maxMi) ?orbit_refl //= cent_cycle. + have{kMiKs} [Ki hallKi sKsKi] := Hall_superset (mmax_sol maxMi) sKsMi kMiKs. + have{ntKs} PmaxMi: Mi \in 'M_'P. + rewrite !(maxMi, inE) andbT /= -partG_eq1 -(card_Hall hallKi) -trivg_card1. + exact: subG1_contra sKsKi ntKs. + have [_ [defNKi defNXs] _ _ _] := Ptype_structure PmaxMi hallKi. + split=> //= [|Xs]; first by exists Ki; apply/and3P. + rewrite -{1}[Ks](setIidPr sKsKi) nElemI -setIdE => /setIdP[E1Xs sXsKs]. + have{defNXs} [defNXs _] := defNXs _ E1Xs; rewrite join_subG /= {2}defNXs. + by rewrite !subsetI sKMi sKsMi cents_norm ?normsG ?(centsS sXsKs) // centsC. +move=> M K PmaxM hallK /=; set Ks := 'C_(M`_\sigma)(K); set Z := K <*> Ks. +move: {2}_.+1 (ltnSn #|class_support (Z :\: (K :|: Ks)) G|) => nTG. +elim: nTG => // nTG IHn in M K PmaxM hallK Ks Z *; rewrite ltnS => leTGn. +have [maxM notFmaxM]: M \in 'M /\ M \notin 'M_'F := setDP PmaxM. +have{notFmaxM} ntK: K :!=: 1 by rewrite (trivg_kappa maxM). +have [_ [defNK defNX] [ntKs uniqCKs] _ _] := Ptype_structure PmaxM hallK. +rewrite -/Ks in defNK ntKs uniqCKs; have [_ mulKKs cKKs _] := dprodP defNK. +have{mulKKs} defZ: 'N_M(K) = Z by rewrite -mulKKs -cent_joinEr. +pose MNX := \bigcup_(X in 'E^1(K)) 'M('N(X)); pose MX := M |: MNX. +have notMG_MNX: {in MNX, forall Mi, gval Mi \notin M :^: G}. + by move=> Mi /bigcupP[X E1X /(Pmax_sym M K)[]]. +have MX0: M \in MX := setU11 M MNX. +have notMNX0: M \notin MNX by apply/negP=> /notMG_MNX; rewrite orbit_refl. +pose K_ Mi := odflt K [pick Ki | isKi Ks Mi Ki]. +pose Ks_ Mi := 'C_(Mi`_\sigma)(K_ Mi). +have K0: K_ M = K. + rewrite /K_; case: pickP => // K1 /and3P[_ /and3P[_ kK1 _] sKsK1]. + have sM_Ks: \sigma(M).-group Ks := pgroupS (subsetIl _ _) (pcore_pgroup _ _). + rewrite -(setIid Ks) coprime_TIg ?eqxx ?(pnat_coprime sM_Ks) // in ntKs. + exact: sub_pgroup (@kappa_sigma' M) (pgroupS sKsK1 kK1). +have Ks0: Ks_ M = Ks by rewrite /Ks_ K0. +have K_spec: {in MNX, forall Mi, isKi Ks Mi (K_ Mi)}. + move=> Mi /bigcupP[X _ /(Pmax_sym M K)[] // _ _ [Ki Ki_ok] _]. + by rewrite /K_; case: pickP => // /(_ Ki)/idP. +have PmaxMX: {in MX, forall Mi, Mi \in 'M_'P /\ \kappa(Mi).-Hall(Mi)(K_ Mi)}. + by move=> Mi /setU1P[-> | /K_spec/and3P[]//]; rewrite K0. +have ntKsX: {in MX, forall Mi, Ks_ Mi != 1}. + by move=> Mi /PmaxMX[MX_Mi /Ptype_structure[] // _ _ []]. +pose co_sHallK Mi Zi := + let sMi := \sigma(Mi) in sMi^'.-Hall(Zi) (K_ Mi) /\ sMi.-Hall(Zi) (Ks_ Mi). +have hallK_Zi: {in MX, forall Mi, co_sHallK Mi (K_ Mi \x Ks_ Mi)}. + move=> Mi MXi; have [PmaxMi hallKi] := PmaxMX _ MXi. + have [_ [defNKs _] _ _ _] := Ptype_structure PmaxMi hallKi. + have [_ mulKKs _ _] := dprodP defNKs; rewrite defNKs. + have sMi_Kis: _.-group (Ks_ Mi) := pgroupS (subsetIl _ _) (pcore_pgroup _ _). + have sMi'Ki := sub_pgroup (@kappa_sigma' _) (pHall_pgroup hallKi). + exact: coprime_mulGp_Hall mulKKs sMi'Ki sMi_Kis. +have{K_spec} defZX: {in MX, forall Mi, K_ Mi \x Ks_ Mi = Z}. + move=> Mi MXi; have [-> | MNXi] := setU1P MXi; first by rewrite K0 Ks0 defNK. + have /and3P[PmaxMi hallKi sKsKi] := K_spec _ MNXi. + have [X E1X maxNMi] := bigcupP MNXi. + have{defNX} [defNX /(_ Mi maxNMi) sXMis] := defNX X E1X. + have /rank_geP[Xs E1Xs]: 0 < 'r(Ks) by rewrite rank_gt0. + have [_ [defNi defNXi] _ _ _] := Ptype_structure PmaxMi hallKi. + have [defNXs _] := defNXi _ (subsetP (nElemS 1 sKsKi) _ E1Xs). + have [_ hallKis] := hallK_Zi _ MXi; rewrite defNi in hallKis. + have sZNXs: Z \subset 'N_Mi(Xs) by case/(Pmax_sym M K): maxNMi => // _ _ _ ->. + apply/eqP; rewrite eqEsubset andbC {1}defNi -defNXs sZNXs. + have [_ _ cKiKis tiKiKis] := dprodP defNi; rewrite dprodEY // -defZ -defNX. + have E1KiXs: Xs \in 'E^1(K_ Mi) := subsetP (nElemS 1 sKsKi) Xs E1Xs. + have [|_ _ _ -> //] := Pmax_sym Mi _ Xs M PmaxMi hallKi E1KiXs. + have [p EpXs] := nElemP E1Xs; have [_] := pnElemP EpXs; case/andP=> pXs _ _. + rewrite inE maxM (sub_uniq_mmax (uniqCKs _ E1Xs)) ?cent_sub //=. + exact: mFT_norm_proper (nt_pnElem EpXs isT) (mFT_pgroup_proper pXs). + have [q /pnElemP[sXK abelX dimX]] := nElemP E1X. + apply/nElemP; exists q; apply/pnElemP; split=> //. + have nKisZi: Ks_ Mi <| 'N_Mi(K_ Mi) by case/dprod_normal2: defNi. + rewrite (sub_normal_Hall hallKis) ?(pgroupS sXMis (pcore_pgroup _ _)) //=. + by rewrite -defNXs (subset_trans sXK) // (subset_trans (joing_subl _ Ks)). +have{hallK_Zi} hallK_Z: {in MX, forall Mi, co_sHallK Mi Z}. + by move=> Mi MXi; rewrite -(defZX _ MXi); apply: hallK_Zi. +have nsK_Z: {in MX, forall Mi, K_ Mi <| Z /\ Ks_ Mi <| Z}. + by move=> Mi /defZX; apply: dprod_normal2. +have tiKs: {in MX &, forall Mi Mj, gval Mi != gval Mj -> Ks_ Mi :&: Ks_ Mj = 1}. + move=> Mi Mj MXi MXj; apply: contraNeq; rewrite -rank_gt0. + case/rank_geP=> X E1X; move: E1X (E1X); rewrite /= {1}setIC {1}nElemI. + case/setIP=> E1jX _; rewrite nElemI => /setIP[E1iX _]. + have [[maxKi hallKi] [maxKj hallKj]] := (PmaxMX _ MXi, PmaxMX _ MXj). + have [_ _ [_ uniqMi] _ _] := Ptype_structure maxKi hallKi. + have [_ _ [_ uniqMj] _ _] := Ptype_structure maxKj hallKj. + by rewrite val_eqE -in_set1 -(uniqMj _ E1jX) (uniqMi _ E1iX) set11. +have sKsKX: {in MX &, forall Mi Mj, Mj != Mi -> Ks_ Mj \subset K_ Mi}. + move=> Mi Mj MXi MXj /= neqMji; have [hallKi hallKsi] := hallK_Z _ MXi. + have [[_ nsKsjZ] [nsKiZ _]] := (nsK_Z _ MXj, nsK_Z _ MXi). + rewrite (sub_normal_Hall hallKi) ?(normal_sub nsKsjZ) // -partG_eq1. + by rewrite -(card_Hall (Hall_setI_normal _ hallKsi)) //= setIC tiKs ?cards1. +have exMNX X: X \in 'E^1(K) -> exists2 Mi, Mi \in MNX & X \subset Mi`_\sigma. + move=> E1X; have [p EpX] := nElemP E1X; have [_ abelX _] := pnElemP EpX. + have ltXG: X \proper G := mFT_pgroup_proper (abelem_pgroup abelX). + have [Mi maxNMi] := mmax_exists (mFT_norm_proper (nt_pnElem EpX isT) ltXG). + have MNXi: Mi \in MNX by apply/bigcupP; exists X. + by exists Mi => //; have [_ ->] := defNX X E1X. +have dprodKs_eqZ: \big[dprod/1]_(Mi in MX) Ks_ Mi = Z; last clear dprodKs_eqZ. + have sYKs_KX Mi: + Mi \in MX -> <<\bigcup_(Mj in MX | Mj != Mi) Ks_ Mj>> \subset K_ Mi. + - move=> MXi; rewrite gen_subG. + by apply/bigcupsP=> Mj /= /andP[]; apply: sKsKX. + transitivity <<\bigcup_(Mi in MX) Ks_ Mi>>; apply/eqP. + rewrite -bigprodGE; apply/bigdprodYP => Mi MXi; rewrite bigprodGE. + apply: subset_trans (sYKs_KX _ MXi) _; apply/dprodYP. + have [_ defZi cKiKs tiKiKs] := dprodP (defZX _ MXi). + by rewrite dprodC joingC dprodEY. + rewrite eqEsubset {1}(bigD1 M) //= Ks0 setUC -joingE -joing_idl. + rewrite genS ?setSU ?big_setU1 //=; last by rewrite -K0 sYKs_KX. + rewrite setUC -joingE -joing_idl Ks0 genS ?setSU // -(Sylow_gen K) gen_subG. + apply/bigcupsP=> P; case/SylowP=> p p_pr /=; case/and3P=> sPK pP _. + have [-> | ] := eqsVneq P 1; first exact: sub1G. + rewrite -rank_gt0 (rank_pgroup pP); case/p_rank_geP=> X EpX. + have EpKX: X \in 'E_p^1(K) := subsetP (pnElemS p 1 sPK) X EpX. + have{EpKX} E1X: X \in 'E^1(K) by apply/nElemP; exists p. + have [Mi MNXi sXMis] := exMNX X E1X; have MXi: Mi \in MX by rewrite setU1r. + have [[_ nsKsi] [_ hallKsi]] := (nsK_Z _ MXi, hallK_Z _ MXi). + have sPZ: P \subset Z := subset_trans sPK (joing_subl _ _). + rewrite sub_gen ?(bigcup_max Mi) // (sub_normal_Hall hallKsi) //. + rewrite (pi_pgroup pP) // (pnatPpi (pcore_pgroup _ _) (piSg sXMis _)) //. + by have [_ ? dimX] := pnElemP EpX; rewrite -p_rank_gt0 p_rank_abelem ?dimX. +pose PZ := [set (Ks_ Mi)^# | Mi in MX]; pose T := Z^# :\: cover PZ. +have defT: \bigcup_(Mi in MX) (Ks_ Mi)^# * (K_ Mi)^# = T. + apply/setP=> x; apply/bigcupP/setDP=> [[Mi MXi] | [Zx notZXx]]. + case/mulsgP=> y y' /setD1P[nty Ks_y] /setD1P[nty' Ky'] defx. + have [_ defZi cKsKi tiKsKi] := dprodP (defZX _ MXi). + rewrite 2!inE -[Z]defZi -(centC cKsKi) andbC {1}defx mem_mulg //=. + have notKx: x \notin K_ Mi. + by rewrite -in_set1 -set1gE -tiKsKi inE Ks_y andbT defx groupMr in nty *. + split; first exact: group1_contra notKx. + rewrite cover_imset; apply/bigcupP=> [[Mj MXj /setD1P[_ Ksj_x]]]. + rewrite (subsetP (sKsKX Mi Mj _ _ _)) // in notKx. + apply: contraNneq nty' => eqMji; rewrite -in_set1 -set1gE -tiKsKi inE Ky'. + by rewrite -(groupMl _ Ks_y) -defx -eqMji. + have{Zx} [ntx Zx] := setD1P Zx. + have [Mi MXi notKi_x]: exists2 Mi, Mi \in MX & x \notin K_ Mi. + have [Kx | notKx] := boolP (x \in K); last by exists M; rewrite ?K0. + pose p := pdiv #[x]; have xp: p \in \pi(#[x]) by rewrite pi_pdiv order_gt1. + have /p_rank_geP[X EpX]: 0 < 'r_p(<[x]>) by rewrite p_rank_gt0. + have [sXx abelX dimX] := pnElemP EpX. + have piXp: p \in \pi(X) by rewrite -p_rank_gt0 p_rank_abelem ?dimX. + have sXK: X \subset K by rewrite (subset_trans sXx) ?cycle_subG. + have E1X: X \in 'E^1(K) by apply/nElemP; exists p; apply/pnElemP. + have [Mi MNXi sXMis] := exMNX X E1X; have MXi: Mi \in MX := setU1r M MNXi. + have sXZ: X \subset Z := subset_trans sXK (joing_subl _ _). + have sMip: p \in \sigma(Mi) := pnatPpi (pcore_pgroup _ _) (piSg sXMis piXp). + have [hallKi _] := hallK_Z _ MXi. + exists Mi => //; apply: contraL sMip => Ki_x. + exact: pnatPpi (mem_p_elt (pHall_pgroup hallKi) Ki_x) xp. + have [_ defZi cKisKi _] := dprodP (defZX _ MXi). + rewrite -[Z]defZi -(centC cKisKi) in Zx. + have [y y' Kis_y Ki_y' defx] := mulsgP Zx. + have Kis1y: y \in (Ks_ Mi)^#. + rewrite 2!inE Kis_y andbT; apply: contraNneq notKi_x => y1. + by rewrite defx y1 mul1g. + exists Mi; rewrite // defx mem_mulg // 2!inE Ki_y' andbT. + apply: contraNneq notZXx => y'1; rewrite cover_imset. + by apply/bigcupP; exists Mi; rewrite // defx y'1 mulg1. +have oT: #|T| = #|Z| + #|MNX| - (\sum_(Mi in MX) #|Ks_ Mi|). + have tiTPZ Kis: Kis \in PZ -> [disjoint T & Kis]. + move=> Z_Kis; rewrite -setI_eq0 setIDAC setD_eq0. + by rewrite (bigcup_max Kis) ?subsetIr. + have notPZset0: set0 \notin PZ. + apply/imsetP=> [[Mi MXi]]; apply/eqP; rewrite /= eq_sym setD_eq0 subG1. + exact: ntKsX. + have [| tiPZ injKs] := trivIimset _ notPZset0. + move=> Mi Mj MXi MXj /= neqMji. + by rewrite -setI_eq0 -setDIl setD_eq0 setIC tiKs. + have{tiPZ} [tiPZ notPZ_T] := trivIsetU1 tiTPZ tiPZ notPZset0. + rewrite (eq_bigr (fun Mi : {group gT} => 1 + #|(Ks_ Mi)^#|)%N); last first. + by move=> Mi _; rewrite (cardsD1 1) group1. + rewrite big_split sum1_card cardsU1 notMNX0 (cardsD1 1 Z) group1 /=. + have ->: Z^# = cover (T |: PZ). + rewrite -(setID Z^# (cover PZ)) setUC (setIidPr _) /cover ?big_setU1 //=. + apply/bigcupsP=> _ /imsetP[Mi MXi ->]; apply: setSD. + by case/nsK_Z: MXi => _ /andP[]. + by rewrite addnAC subnDl -(eqnP tiPZ) big_setU1 // big_imset //= addnK. +have tiTscov: {in 'M, forall H, [disjoint T & H^~~]}. + move=> H maxH; apply/pred0P=> t; apply/andP=> [[/= Tt scovHt]]. + have ntt: t != 1 by have [/setD1P[]] := setDP Tt. + have [x Hs_x xR_y] := bigcupP scovHt; have ell1x := Msigma_ell1 maxH Hs_x. + have:= sigma_decomposition_dichotomy ntt. + rewrite (introT existsP) /=; last by exists x; rewrite ell1x -mem_lcoset. + rewrite -defT in Tt; have [Mi MXi Zi_t] := bigcupP Tt. + case/mulsgP: Zi_t => y y' /setD1P[nty Ks_y] /setD1P[nty' Ky'] ->. + case/existsP; exists y; rewrite mulKg. + have [[Mis_y cKy] [PmaxMi hallKi]] := (setIP Ks_y, PmaxMX _ MXi). + have [[maxMi _] [sKiMi kMiKi _]] := (setDP PmaxMi, and3P hallKi). + rewrite (Msigma_ell1 maxMi) ?inE ?nty //=; apply/existsP; exists Mi. + rewrite inE maxMi cycle_subG Mis_y 3!inE nty' (subsetP sKiMi) //=. + by rewrite (subsetP _ _ Ky') ?sub_cent1 // (mem_p_elt kMiKi). +have nzT: T != set0. + have [[y Ksy nty] [y' Ky' nty']] := (trivgPn _ ntKs, trivgPn _ ntK). + apply/set0Pn; exists (y * y'); rewrite -defT; apply/bigcupP. + by exists M; rewrite ?MX0 // K0 Ks0 mem_mulg 2?inE ?nty ?nty'. +have ntiT: normedTI T G Z. + have sTZ: {subset T <= Z} by apply/subsetP; rewrite 2!subDset setUA subsetUr. + have nTZ: Z \subset 'N(T). + rewrite normsD ?norms_bigcup ?normD1 ?normG //. + apply/bigcapsP=> _ /imsetP[Mi MXi ->]; rewrite normD1. + by case/nsK_Z: MXi => _ /andP[]. + apply/normedTI_P; rewrite setTI /= -/Z. + split=> // a _ /pred0Pn[t /andP[/= Tt]]; rewrite mem_conjg => Tta. + have{Tta} [Zt Zta] := (sTZ t Tt, sTZ _ Tta). + move: Tt; rewrite -defT => /bigcupP[Mi MXi]. + case/mulsgP=> y y' /setD1P[nty Kisy] /setD1P[nty' Kiy'] def_yy'. + have [[hallKi hallKis] [nsKiZ _]] := (hallK_Z _ MXi, nsK_Z _ MXi). + have [[PmaxMi hallKiMi] defZi] := (PmaxMX _ MXi, defZX _ MXi). + have [_ [defNKi _] _ [[]]] := Ptype_structure PmaxMi hallKiMi. + rewrite -defNKi defZi -/(Ks_ _) => tiKsi tiKi _ _ _. + have [defy defy']: y = t.`_\sigma(Mi) /\ y' = t.`_\sigma(Mi)^'. + have [_ cKiy] := setIP Kisy; have cy'y := centP cKiy _ Kiy'. + have sMi_y := mem_p_elt (pHall_pgroup hallKis) Kisy. + have sMi'y' := mem_p_elt (pHall_pgroup hallKi) Kiy'. + rewrite def_yy' !consttM // constt_p_elt // 2?(constt1P _) ?p_eltNK //. + by rewrite mulg1 mul1g constt_p_elt. + have: a \in Mi. + apply: contraR nty; rewrite -in_setC -in_set1 -set1gE; move/tiKsi <-. + rewrite inE Kisy mem_conjg defy -consttJ groupX ?(subsetP _ _ Zta) //. + by rewrite -defZi defNKi subsetIl. + apply/implyP; apply: contraR nty'; rewrite negb_imply andbC -in_setD. + rewrite -in_set1 -set1gE => /tiKi <-; rewrite inE Kiy' defy' mem_conjg. + by rewrite -consttJ (mem_normal_Hall hallKi nsKiZ) ?p_elt_constt ?groupX. +have [_ tiT /eqP defNT] := and3P ntiT; rewrite setTI in defNT. +pose n : rat := #|MNX|%:R; pose g : rat := #|G|%:R. +pose z : rat := #|Z|%:R; have nz_z: z != 0%R := natrG_neq0 _ _. +pose k_ Mi : rat := #|K_ Mi|%:R. +have nz_ks: #|Ks_ _|%:R != 0%R :> rat := natrG_neq0 _ _. +pose TG := class_support T G. +have oTG: (#|TG|%:R = (1 + n / z - \sum_(Mi in MX) (k_ Mi)^-1) * g)%R. + rewrite /TG class_supportEr -cover_imset -(eqnP tiT). + rewrite (eq_bigr (fun _ => #|T|)) => [|_ /imsetP[x _ ->]]; last first. + by rewrite cardJg. + rewrite sum_nat_const card_conjugates setTI defNT. + rewrite natrM natf_indexg ?subsetT //= -/z -mulrA mulrC; congr (_ * _)%R. + rewrite oT natrB; last by rewrite ltnW // -subn_gt0 lt0n -oT cards_eq0. + rewrite mulrC natrD -/n -/z natr_sum /=. + rewrite mulrBl mulrDl big_distrl divff //=; congr (_ - _)%R. + apply: eq_bigr => Mi MXi; have defZi := defZX _ MXi. + by rewrite /z -(dprod_card defZi) natrM invfM mulrC divfK. +have neMNX: MNX != set0. + move: ntK; rewrite -rank_gt0 => /rank_geP[X /exMNX[Mi MNXi _]]. + by apply/set0Pn; exists Mi. +have [Mi MXi P2maxMi]: exists2 Mi, Mi \in MX & Mi \in 'M_'P2. + apply/exists_inP; apply: negbNE; rewrite negb_exists_in. + apply/forall_inP=> allP1; pose ssup Mi := class_support (gval Mi)^~~ G. + have{allP1} min_ssupMX Mi: + Mi \in MX -> (#|ssup Mi|%:R >= ((k_ Mi)^-1 - (z *+ 2)^-1) * g)%R. + - move=> MXi; have [PmaxMi hallKi] := PmaxMX _ MXi. + have [[U [complU defMi] _]] := Ptype_structure PmaxMi hallKi. + case=> defZi _ _ _ _; have [maxMi _] := setDP PmaxMi. + have{complU} U1: U :==: 1; last rewrite {U U1}(eqP U1) sdprod1g in defMi. + rewrite (trivg_kappa_compl maxMi complU). + by apply: contraR (allP1 _ MXi) => ?; apply/setDP. + rewrite card_class_support_sigma // natrM natf_indexg ?subsetT // -/g. + rewrite mulrCA mulrC ler_wpmul2r ?ler0n // -subn1 natrB ?cardG_gt0 //. + rewrite mulr1n mulrBl -{1}(sdprod_card defMi) natrM invfM. + rewrite mulVKf ?natrG_neq0 // ler_add2l ler_opp2 -(mulr_natr _ 2) invfM. + rewrite ler_pdivr_mulr ?natrG_gt0 // mulrC mulrA. + have sZM: Z \subset M by rewrite -defZ subsetIl. + have sZMi: Z \subset Mi by rewrite -(defZX _ MXi) defZi subsetIl. + rewrite -natf_indexg //= -/Z ler_pdivl_mulr ?(ltr0Sn _ 1) // mul1r ler_nat. + rewrite indexg_gt1 /= -/Z subEproper /proper sZMi andbF orbF. + apply: contraNneq notMNX0 => defMiZ; have [Mj MNXj] := set0Pn _ neMNX. + have maxZ: [group of Z] \in 'M by rewrite !inE defMiZ in maxMi *. + have eqZ := group_inj (eq_mmax maxZ _ _); rewrite -(eqZ M) //. + have [Xj E1Xj maxNMj] := bigcupP MNXj; have [maxMj _] := setIdP maxNMj. + by rewrite (eqZ Mj) //; case/(Pmax_sym M K): maxNMj. + pose MXsup := [set ssup Mi | Mi in MX]. + have notMXsup0: set0 \notin MXsup. + apply/imsetP=> [[Mi /PmaxMX[/setDP[maxMi _] _] /esym/eqP/set0Pn[]]]. + have [x Mis_x ntx] := trivgPn _ (Msigma_neq1 maxMi). + exists (x ^ 1); apply: mem_imset2; rewrite ?inE //. + by apply/bigcupP; exists x; rewrite ?inE ?ntx // lcoset_refl. + have [Mi Mj MXi MXj /= neqMij | tiMXsup inj_ssup] := trivIimset _ notMXsup0. + apply/pred0Pn=> [[_ /andP[/imset2P[x y1 signMi_x _ ->]]]] /=. + rewrite /ssup class_supportEr /= => /bigcupP[y2 _]. + rewrite -mem_conjgV -conjsgM -sigma_supportJ; set H := Mj :^ _ => Hx. + suffices: [disjoint Mi^~~ & H^~~]. + by case/pred0Pn; exists x; rewrite /= {1}signMi_x Hx. + have [[PmaxMi _] [PmaxMj _]] := (PmaxMX _ MXi, PmaxMX _ MXj). + have [[maxMi _] [maxMj _]] := (setDP PmaxMi, setDP PmaxMj). + apply: sigma_support_disjoint; rewrite ?mmaxJ //. + rewrite (orbit_transr _ (mem_orbit _ _ _)) ?inE //=. + apply: contra (ntKsX _ MXi); case/imsetP=> y _ /= defMj; rewrite -/(Ks_ _). + have sKisKj: Ks_ Mi \subset K_ Mj by rewrite sKsKX // eq_sym. + rewrite -(setIidPl sKisKj) coprime_TIg //. + have [[_ hallKis] [hallKj _]] := (hallK_Z _ MXi, hallK_Z _ MXj). + apply: pnat_coprime (pHall_pgroup hallKj). + by rewrite defMj -pgroupE (eq_pgroup _ (sigmaJ _ _)) (pHall_pgroup hallKis). + have [|tiPG notMXsupTG]: _ /\ TG \notin _ := trivIsetU1 _ tiMXsup notMXsup0. + move=> _ /imsetP[Mi /PmaxMX[/setDP[maxMi _] _] ->]. + apply/pred0Pn=> [[_ /andP[/imset2P[x y1 Tx _ ->]]]] /=. + rewrite /ssup class_supportEr => /bigcupP[y2 _]. + rewrite -mem_conjgV -conjsgM -sigma_supportJ; set H := Mi :^ _ => Hx. + have maxH: [group of H] \in 'M by rewrite mmaxJ. + by case/andP: (pred0P (tiTscov _ maxH) x). + suffices: (g <= #|cover (TG |: MXsup)|%:R)%R. + rewrite ler_nat (cardsD1 1 G) group1 ltnNge subset_leq_card //. + apply/bigcupsP=> _ /setU1P[|/imsetP[Mi /PmaxMX[/setDP[maxMi _] _]]] ->. + rewrite /TG class_supportEr; apply/bigcupsP=> x _. + rewrite sub_conjg (normP _) ?normD1 ?(subsetP (normG _)) ?inE //. + by rewrite subDset setUC subsetU // setSD ?subsetT. + rewrite /ssup class_supportEr; apply/bigcupsP=> x _. + rewrite subsetD1 subsetT mem_conjg conj1g {x}/=. + move/ell_sigma0P: (@erefl gT 1); rewrite cards_eq0. + apply: contraL => /bigcupP[x Mis_x xR1]; apply/set0Pn; exists x. + exact: mem_sigma_cover_decomposition (Msigma_ell1 maxMi Mis_x) xR1. + rewrite -(eqnP tiPG) big_setU1 ?big_imset //= natrD natr_sum. + suffices: (g <= #|TG|%:R + \sum_(i in MX) ((k_ i)^-1 - (z *+ 2)^-1) * g)%R. + by move/ler_trans->; rewrite // ler_add2l ler_sum. + rewrite -big_distrl /= oTG -/g -mulrDl big_split /= sumr_const. + rewrite addrA subrK -(mulr_natl _ 2) -[_ *+ _]mulr_natl invfM mulrN. + rewrite mulrA -addrA -mulrBl -{1}(mul1r g) ler_wpmul2r ?ler0n //. + rewrite ler_addl -(mul0r z^-1)%R ler_wpmul2r ?invr_ge0 ?ler0n //. + rewrite subr_ge0 ler_pdivr_mulr ?(ltr0Sn _ 1) // -natrM ler_nat. + by rewrite muln2 -addnn cardsU1 leq_add2r notMNX0 lt0n cards_eq0. +have [prKi nilMis]: prime #|K_ Mi| /\ nilpotent Mi`_\sigma. + by have [PmaxMi /Ptype_structure[] // _ _ _ _ []] := PmaxMX _ MXi. +have [Mj MXj neqMji]: exists2 Mj, Mj \in MX & Mj :!=: Mi. + have [Mj |] := pickP (mem ((MX) :\ Mi)); first by case/setD1P; exists Mj. + move/eq_card0/eqP; rewrite -(eqn_add2l true) -{1}MXi -cardsD1 cardsU1. + by rewrite notMNX0 eqSS cards_eq0 (negPf neMNX). +have defKjs: Ks_ Mj = K_ Mi. + have sKjsKi: Ks_ Mj \subset K_ Mi by rewrite sKsKX. + apply/eqP; rewrite eqEcard sKjsKi (prime_nt_dvdP _ _ (cardSg sKjsKi)) //=. + by rewrite -trivg_card1 ntKsX. +have defMXij: MX = [set Mi; Mj]. + symmetry; rewrite -(setD1K MXi); congr (_ |: _); apply/eqP. + rewrite eqEcard sub1set cards1 (cardsD1 Mj) 2!inE neqMji MXj /= ltnS leqn0. + apply/pred0Pn=> [[Mk /setD1P[neMkj /setD1P[neMki MXk]]]]. + have sKskKsj: Ks_ Mk \subset Ks_ Mj by rewrite defKjs sKsKX. + by case/negP: (ntKsX _ MXk); rewrite -(setIidPl sKskKsj) tiKs. +have defKsi: Ks_ Mi = K_ Mj. + apply/eqP; rewrite eqEcard sKsKX 1?eq_sym //=. + rewrite -(@leq_pmul2r #|Ks_ Mj|) ?cardG_gt0 // (dprod_card (defZX _ MXj)). + by rewrite defKjs mulnC (dprod_card (defZX _ MXi)). +have{nilMis} cycZ: cyclic Z. + have cycKi := prime_cyclic prKi. + apply: nil_Zgroup_cyclic. + apply/forall_inP=> S /SylowP[p _ /and3P[sSZ pS _]]. + have [[hallKi hallKis] [nsKi nsKis]] := (hallK_Z _ MXi, nsK_Z _ MXi). + have [sMi_p | sMi'p] := boolP (p \in \sigma(Mi)); last first. + by rewrite (cyclicS _ cycKi) // (sub_normal_Hall hallKi) ?(pi_pgroup pS). + have sSKj: S \subset K_ Mj. + by rewrite -defKsi (sub_normal_Hall hallKis) ?(pi_pgroup pS). + rewrite (odd_pgroup_rank1_cyclic pS) ?mFT_odd //. + apply: wlog_neg; rewrite -ltnNge ltn_neqAle p_rank_gt0 => /andP[_ piSp]. + have [_ /and3P[sKjMj kKj _]] := PmaxMX _ MXj. + rewrite -(rank_kappa (pnatPpi kKj (piSg sSKj piSp))) p_rankS //. + exact: subset_trans sSKj sKjMj. + rewrite (dprod_nil (defZX _ MXi)) abelian_nil ?cyclic_abelian //=. + exact: (nilpotentS (subsetIl _ _)) nilMis. +have cycK: cyclic K := cyclicS (joing_subl _ _) cycZ. +have defM: M^`(1) ><| K = M. + have [U complU] := ex_kappa_compl maxM hallK; have [hallU _ _] := complU. + have [_ defM _ regUK _] := kappa_compl_context maxM complU. + have{hallU} [[sUM _] [sKM kK _]] := (andP hallU, and3P hallK). + case/sdprodP: defM => [[_ E _ defE]]; rewrite defE. + case/sdprodP: defE => _ <-{E} nUK _ defM /mulGsubP[nMsU nMsK] tiMsUK. + pose MsU := M`_\sigma <*> U; have nMsUK: K \subset 'N(MsU) by rewrite normsY. + have defMl: MsU * K = M by rewrite [MsU]norm_joinEr // -mulgA. + have coUK := regular_norm_coprime nUK regUK. + have ->: M^`(1) = MsU. + apply/eqP; rewrite eqEsubset; apply/andP; split; last first. + have solU := solvableS sUM (mmax_sol maxM). + rewrite join_subG Msigma_der1 //= -(coprime_cent_prod nUK coUK solU). + by rewrite (cent_semiregular regUK) // mulg1 commgSS. + apply: der1_min; first by rewrite -{1}defMl mulG_subG normG. + by rewrite -{2}defMl quotientMidl quotient_abelian ?cyclic_abelian. + rewrite sdprodE ?coprime_TIg //= norm_joinEr //. + rewrite (coprime_dvdl (dvdn_cardMg _ _)) // coprime_mull coUK. + rewrite (pnat_coprime (pcore_pgroup _ _) (sub_pgroup _ kK)) //. + exact: kappa_sigma'. +have{neMNX} [Mstar MNX'star] := set0Pn _ neMNX. +have defMNX: MNX = [set Mstar]. + apply/eqP; rewrite eq_sym eqEcard sub1set MNX'star /= -(leq_add2l true). + by rewrite -{1}notMNX0 -cardsU1 -/MX defMXij setUC cards2 neqMji !cards1. +have MXstar: Mstar \in MX by rewrite setU1r. +have [[PmaxMstar hallKstar] defZstar] := (PmaxMX _ MXstar, defZX _ MXstar). +have [maxMstar _] := setDP PmaxMstar. +have notMGMstar := notMG_MNX _ MNX'star; exists Mstar => //. +have [defKs defKs_star]: Ks = K_ Mstar /\ Ks_ Mstar = K. + rewrite /Ks /Ks_ -K0; rewrite /MX defMNX 3!inE val_eqE in neqMji MXj MXi. + by case/set2P: MXi (negPf neqMji) MXj => <- ->; rewrite ?orbF /= => /eqP <-. +have hallKs: \sigma(M).-Hall(Mstar) Ks. + have sKsMstar: Ks \subset Mstar by rewrite defKs (pHall_sub hallKstar). + have sM_Ks: \sigma(M).-group Ks := pgroupS (subsetIl _ _) (pcore_pgroup _ _). + have [Y hallY sKsY] := Hall_superset (mmax_sol maxMstar) sKsMstar sM_Ks. + have [sYMstar sM_Y _] := and3P hallY; apply: etrans hallY; congr pHall. + have sYMs: Y \subset M`_\sigma. + case/Ptype_structure: hallK => // _ _ _ [_ _ -> //]. + by rewrite (setIidPr sKsY). + apply/eqP; rewrite eqEsubset sKsY subsetI sYMs (sameP commG1P trivgP) /=. + have <-: M`_\sigma :&: Mstar`_\sigma = 1. + rewrite coprime_TIg // (pnat_coprime (pcore_pgroup _ _)) //. + apply: sub_pgroup (pcore_pgroup _ _) => q sM1q. + apply: contraFN (sigma_partition maxM maxMstar notMGMstar q) => sMq. + exact/andP. + rewrite commg_subI //. + by rewrite subsetI sYMs (subset_trans sYMstar) ?gFnorm. + rewrite subsetI -{1}defKs_star subsetIl. + by rewrite (subset_trans (pHall_sub hallK)) ?gFnorm. +have oTGgt_g2: (g / 2%:R < #|TG|%:R)%R. + rewrite oTG big_setU1 //= /n defMNX big_set1 cards1 mulrC mul1r. + rewrite ltr_pmul2r ?(ltr_nat _ 0) ?cardG_gt0 // /k_ K0 -defKs. + rewrite /z -defZ -(dprod_card defNK) natrM invfM opprD. + pose hm u : rat := (1 - u%:R^-1)%R; set lhs := (_^-1)%R. + suffices: (lhs < hm #|K| * hm #|Ks|)%R. + by rewrite mulrBl !mulrBr !mul1r mulr1 opprB addrAC !addrA. + have hm_inc u v: 0 < u <= v -> (hm u <= hm v)%R. + case/andP=> u_gt0 le_uv; rewrite ler_add2l ler_opp2. + have v_gt0 := leq_trans u_gt0 le_uv. + rewrite -(mul1r _^-1)%R ler_pdivr_mulr ?ltr0n //. + by rewrite ler_pdivl_mull ?ltr0n // mulr1 ler_nat. + have le_pdiv H: 0 < pdiv #|H| <= #|H| by rewrite pdiv_gt0 dvdn_leq ?pdiv_dvd. + have{le_pdiv} hm_pdiv := hm_inc _ _ (le_pdiv _). + have hm_ge0 u: (0 <= hm u)%R. + by case: u => // u; rewrite subr_ge0 invf_le1 ?ltr0Sn ?(ler_nat _ 1). + do 2![rewrite mulrC (ltr_le_trans _ (ler_wpmul2r (hm_ge0 _) (hm_pdiv _))) //]. + set p := pdiv #|K|; set q := pdiv #|Ks|. + have [odd_p odd_q]: odd p /\ odd q. + by split; apply: dvdn_odd (pdiv_dvd _) (mFT_odd _). + without loss [lt1p ltpq]: p q odd_p odd_q / 1 < p /\ p < q. + have [p_pr q_pr]: prime p /\ prime q by rewrite !pdiv_prime ?cardG_gt1. + have [ltpq | ltqp | eqpq] := ltngtP p q. + - by apply; rewrite ?prime_gt1. + - by rewrite mulrC; apply; rewrite ?prime_gt1. + have [] := hallK_Z _ MX0. + rewrite K0 Ks0 => /and3P[_ sM'K _] /and3P[_ sMKs _]. + case/negP: (pgroupP sM'K _ p_pr (pdiv_dvd _)); rewrite eqpq. + exact: pgroupP sMKs _ q_pr (pdiv_dvd _). + have p_gt2: 2 < p by rewrite odd_geq. + apply: ltr_le_trans (isT : lhs < hm 3 * hm 5)%R _. + by rewrite ler_pmul ?hm_inc ?hm_ge0 //= odd_geq ?(leq_trans _ ltpq). +have defZhat: Z :\: (K :|: Ks) = T. + rewrite /T cover_imset big_setU1 //= defMNX big_set1 defKs_star Ks0. + by rewrite -setDUl setDDl setUC setD1K // inE group1. +rewrite defZhat {1}defKs; split; first 2 [by split]. +- by rewrite -defKs_star; case/Ptype_structure: hallKstar => // _ _ []. +- split=> [|p]; first by rewrite -defKs_star defKs. + apply/idP/idP=> [kMp | t1p]. + have /orP[// | /and3P[_ _ p_dv_M']] := kappa_tau13 kMp. + have hallM': \kappa(M)^'.-Hall(M) M^`(1). + exact/(sdprod_normal_pHallP (der_normal 1 M) hallK). + have piMp: p \in \pi(M) by rewrite kappa_pi. + case/idPn: kMp; apply: (pnatPpi (pHall_pgroup hallM')). + by move: piMp; rewrite !mem_primes !cardG_gt0 /= => /andP[->]. + apply: (pnatPpi (pHall_pgroup hallK)); have [_ _ not_p_dv_M'] := and3P t1p. + have: p \in \pi(M) by rewrite (partition_pi_mmax maxM) t1p ?orbT. + rewrite !mem_primes !cardG_gt0 /= => /andP[p_pr]. + by rewrite p_pr -(sdprod_card defM) Euclid_dvdM // (negPf not_p_dv_M'). +- split=> // [| x | y | x y K1_x Ks1_y]. + + have defMsMstar: M`_\sigma :&: Mstar = Ks. + apply: sub_pHall hallKs _ _ (subsetIr _ _). + exact: pgroupS (subsetIl _ _) (pcore_pgroup _ _). + by rewrite subsetI subsetIl /= -/Ks defKs (pHall_sub hallKstar). + have nKsMMstar: M :&: Mstar \subset 'N(Ks). + by rewrite -defMsMstar normsIG ?gFnorm. + have [_ [defNKs _] _ _ _] := Ptype_structure PmaxMstar hallKstar. + rewrite -(setIidPl nKsMMstar) -setIA defKs -defNKs defZstar. + by rewrite -defZ setIA setIid. + + case/setD1P; rewrite -cycle_eq1 -cycle_subG -cent_cycle => ntx sxK. + apply/eqP; rewrite eqEsubset andbC subsetI -{1}defZ subsetIl. + rewrite sub_abelian_cent ?cyclic_abelian //=; last first. + by rewrite (subset_trans sxK) ?joing_subl. + move: ntx; rewrite -rank_gt0 /= -{1}(setIidPr sxK) => /rank_geP[X]. + rewrite nElemI -setIdE -defZ => /setIdP[E1X sXx]. + by have [<- _] := defNX _ E1X; rewrite setIS ?cents_norm ?centS. + + case/setD1P; rewrite -cycle_eq1 -cycle_subG -cent_cycle => nty syKs. + have [_ [defNKs defNY] _ _ _] := Ptype_structure PmaxMstar hallKstar. + rewrite defZstar -defKs in defNKs defNY. + apply/eqP; rewrite eqEsubset andbC subsetI {1}defNKs subsetIl. + rewrite sub_abelian_cent ?cyclic_abelian //=; last first. + by rewrite (subset_trans syKs) ?joing_subr. + move: nty; rewrite -rank_gt0 /= -{1}(setIidPr syKs) => /rank_geP[Y]. + rewrite nElemI -setIdE defNKs => /setIdP[E1Y sYy]. + by have [<- _] := defNY _ E1Y; rewrite setIS ?cents_norm ?centS. + have [[_ K_x] [_ Ks_y]] := (setD1P K1_x, setD1P Ks1_y). + apply/eqP; rewrite eqEsubset sub_cent1 -(centsP cKKs) //. + have Tyx: y * x \in T by rewrite -defT big_setU1 //= inE Ks0 K0 mem_mulg. + rewrite (subset_trans _ (cent1_normedTI ntiT Tyx)) ?setTI //. + rewrite (subsetP _ _ Tyx) // -defZhat setDE subIset //. + by rewrite -abelianE cyclic_abelian. +split=> // [||H PmaxH]. +- split=> // a notMa. + have{tiKs} [_ _ _ [[tiKs _] _ _] _] := Ptype_structure PmaxM hallK. + rewrite -defT big_setU1 //= defMNX big_set1 -defKs defKs_star Ks0 K0. + rewrite centC ?(centSS _ _ cKKs) ?subsetDl // setUid. + apply/pred0Pn=> [[_ /andP[/mulsgP[x y K1_x Ks1_y ->] /= Ma_xy]]]. + have [[_ K_x] [nty Ks_y]] := (setD1P K1_x, setD1P Ks1_y); case/negP: nty. + rewrite -in_set1 -set1gE -(tiKs a notMa) inE Ks_y. + suffices ->: y = (x * y).`_\sigma(M) by rewrite groupX. + rewrite consttM; last by red; rewrite -(centsP cKKs). + have sM'K := sub_pgroup (@kappa_sigma' M) (pHall_pgroup hallK). + rewrite (constt1P (mem_p_elt sM'K K_x)) mul1g constt_p_elt //. + exact: mem_p_elt (pHall_pgroup hallKs) Ks_y. +- have:= set21 Mi Mj; rewrite -defMXij /MX defMNX defKs -K0. + by case/set2P=> <-; [left | right]. +have [maxH _] := setDP PmaxH. +have{maxH}[L hallL] := Hall_exists \kappa(H) (mmax_sol maxH). +pose Ls := 'C_(H`_\sigma)(L); pose S := (L <*> Ls) :\: (L :|: Ls). +have{IHn} oSGgt_g2: (g / 2%:R < #|class_support S G|%:R)%R. + have [|nTG_leS] := ltnP #|class_support S G| nTG. + by case/IHn=> // Sstar _ [_ _ _ _ [[_ _ -> //]]]. + apply: ltr_le_trans oTGgt_g2 _; rewrite ler_nat /TG -defZhat. + exact: leq_trans leTGn nTG_leS. +have{oSGgt_g2 oTGgt_g2} meetST: ~~ [disjoint TG & class_support S G]. + rewrite -leq_card_setU; apply: contraTneq (leqnn #|G|) => tiTGS. + rewrite -ltnNge -(ltr_nat [realFieldType of rat]) -/g. + rewrite -{1}[g](@divfK _ 2%:R) // mulr_natr. + apply: ltr_le_trans (ltr_add oTGgt_g2 oSGgt_g2) _. + by rewrite -natrD -tiTGS ler_nat cardsT max_card. +have{meetST} [x Tx [a Sx]]: exists2 x, x \in T & exists a, x \in S :^ a. + have [_ /andP[/imset2P[x a1 Tx _ ->]]] := pred0Pn meetST. + rewrite class_supportEr => /bigcupP[a2 _ Sa2_xa1]. + by exists x => //; exists (a2 * a1^-1); rewrite conjsgM mem_conjgV. +rewrite {}/S {}/Ls in Sx; without loss a1: a H L PmaxH hallL Sx / a = 1. + move/(_ 1 (H :^ a)%G (L :^ a)%G); rewrite conjsg1 PtypeJ PmaxH pHallJ2. + rewrite (eq_pHall _ _ (kappaJ H a)) hallL MsigmaJ centJ. + rewrite -conjIg -conjYg -conjUg -conjDg Sx !inE. + by rewrite !(orbit_transr _ (mem_orbit _ _ _)) ?inE //; exact. +have [_ [defNL _] [_ uniqH] _ _] := Ptype_structure PmaxH hallL. +do [rewrite {a}a1 conjsg1; set Ls := 'C_(_)(L)] in Sx defNL. +have{x Sx Tx} [Mk MXk ntLsMks]: exists2 Mk, Mk \in MX & Ls :&: Ks_ Mk != 1. + have [_ _ cLLs tiLLs] := dprodP defNL. + pose W := L <*> Ls; pose y := x.`_\sigma(H); pose ys := y.`_\sigma(Mi). + have Zy: y \in Z by apply: groupX; case/setDP: Tx; case/setD1P=> _ ->. + have{hallL} [hallL hallLs]: \sigma(H)^'.-Hall(W) L /\ \sigma(H).-Hall(W) Ls. + apply: coprime_mulGp_Hall; first by rewrite /= cent_joinEr. + exact: sub_pgroup (@kappa_sigma' H) (pHall_pgroup hallL). + exact: pgroupS (subsetIl _ _) (pcore_pgroup _ _). + have [nsLW nsLsW]: L <| W /\ Ls <| W := cprod_normal2 (cprodEY cLLs). + have{Sx} [Ls_y nty]: y \in Ls /\ y != 1. + move: Sx; rewrite 2!inE negb_or -andbA -/W; case/and3P=> notLx _ Wx. + split; first by rewrite (mem_normal_Hall hallLs) ?p_elt_constt ?groupX. + by rewrite (sameP eqP constt1P) -(mem_normal_Hall hallL). + have [[hallKi hallKis] [nsKi nsKis]] := (hallK_Z _ MXi, nsK_Z _ MXi). + have [/constt1P sM'y | ntys] := altP (ys =P 1). + exists Mj; rewrite // defKjs. + by apply/trivgPn; exists y; rewrite // inE Ls_y (mem_normal_Hall hallKi). + exists Mi => //; apply/trivgPn; exists ys; rewrite // inE groupX //=. + by rewrite (mem_normal_Hall hallKis) ?p_elt_constt // groupX. +suffices ->: H = Mk. + by move: MXk; rewrite /MX defMNX => /set2P[]->; rewrite inE orbit_refl ?orbT. +move: ntLsMks; rewrite -rank_gt0 => /rank_geP[Y E1Y]. +have:= E1Y; rewrite nElemI => /setIP[E1LsY _]. +apply: set1_inj; rewrite -(uniqH _ E1LsY). +have [PmaxMk hallKk] := PmaxMX _ MXk. +have [_ _ [_ -> //]] := Ptype_structure PmaxMk hallKk. +by rewrite /= setIC nElemI in E1Y; case/setIP: E1Y. +Qed. + +End PTypeEmbedding. + +(* This is the first part of B & G, Corollary 14.8. *) +Corollary P1type_trans : {in 'M_'P1 &, forall M H, gval H \in M :^: G}. +Proof. +move=> M H P1maxM P1maxH; have [PmaxM _] := setIdP P1maxM. +have [[maxM _] [PmaxH _]] := (setDP PmaxM, setIdP P1maxH). +have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). +have [Mstar _ [_ _ _ _ [_ [|]]]] := Ptype_embedding PmaxM hallK. + by case; rewrite inE P1maxM. +case=> /setDP[_ /negP notP1maxMstar] _. +case/(_ H PmaxH)/setUP=> // /imsetP[a _ /group_inj defH]. +by rewrite defH P1typeJ in P1maxH. +Qed. + +(* This is the second part of B & G, Corollary 14.8. *) +Corollary Ptype_trans : {in 'M_'P, forall M, + exists2 Mstar, Mstar \in 'M_'P /\ gval Mstar \notin M :^: G + & {in 'M_'P, forall H, gval H \in M :^: G :|: Mstar :^: G}}. +Proof. +move=> M PmaxM; have [maxM _] := setDP PmaxM. +have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). +have [Mstar PmaxMstar [_ _ _ _ [_ _ inMMs _]]] := Ptype_embedding PmaxM hallK. +by exists Mstar. +Qed. + +(* This is B & G, Corollary 14.9. *) +Corollary mFT_partition : + let Pcover := [set class_support M^~~ G | M : {group gT} in 'M] in + [/\ (*1*) 'M_'P == set0 :> {set {group gT}} -> partition Pcover G^# + & (*2*) forall M K, M \in 'M_'P -> \kappa(M).-Hall(M) K -> + let Ks := 'C_(M `_\sigma)(K) in let Z := K <*> Ks in + let Zhat := Z :\: (K :|: Ks) in + let ClZhat := class_support Zhat G in + partition (ClZhat |: Pcover) G^# /\ ClZhat \notin Pcover]. +Proof. +move=> Pcover; have notPcover0: set0 \notin Pcover. + apply/imsetP=> [[M maxM]]; apply/eqP; rewrite eq_sym; apply/set0Pn. + have [x Ms_x ntx] := trivgPn _ (Msigma_neq1 maxM); exists x. + rewrite class_supportEl; apply/bigcupP; exists x; last exact: class_refl. + by apply/bigcupP; exists x; [apply/setD1P | apply: lcoset_refl]. +have tiPcover: trivIset Pcover. + apply/trivIsetP=> _ _ /imsetP[M maxM ->] /imsetP[H maxH ->] notMGH. + rewrite -setI_eq0 !{1}class_supportEr big_distrr big1 //= => a Ga. + rewrite big_distrl big1 //= => b Gb; apply/eqP. + rewrite -!{1}sigma_supportJ setI_eq0 sigma_support_disjoint ?mmaxJ //. + apply: contra notMGH; rewrite {a Ga}(orbit_transr _ (mem_orbit _ _ Ga)). + rewrite {b Gb}(orbit_transl (mem_orbit _ _ Gb))=> /imsetP[c Gc ->] /=. + by rewrite sigma_supportJ class_supportGidl. +have ntPcover: cover Pcover \subset G^#. + apply/bigcupsP=> _ /imsetP[M maxM ->]; rewrite class_supportEr. + apply/bigcupsP=> a _; rewrite subsetD1 subsetT mem_conjg conj1g {a}//=. + move/ell_sigma0P: (@erefl gT 1); rewrite cards_eq0; apply: contraL. + case/bigcupP=> x Ms_x xR1; apply/set0Pn; exists x. + exact: mem_sigma_cover_decomposition (Msigma_ell1 maxM Ms_x) xR1. +split=> [MP0 | M K PmaxM hallK Ks Z Zhat ClZhat]. + rewrite /partition eqEsubset ntPcover tiPcover notPcover0 !andbT /=. + apply/subsetP=> x; rewrite !inE andbT => ntx. + have:= sigma_decomposition_dichotomy ntx. + have [[y ell1y yRx] _ | _] := exists_inP. + have [nty /set0Pn[M /setIdP[maxM Ms_y]]] := ell_sigma1P _ ell1y. + apply/bigcupP; exists (class_support M^~~ G); first exact: mem_imset. + rewrite -(conjg1 x) mem_imset2 ?inE //. + apply/bigcupP; exists y; last by rewrite mem_lcoset. + by rewrite !inE nty -cycle_subG. + case/exists_inP=> y _; move: (_ * x) => y' /existsP[M]. + case/and3P => /setIdP[maxM _] /setD1P[nty' /setIP[My' _]] kMy' {y}. + case/set0Pn: MP0; exists M; rewrite 2!inE maxM andbT. + apply: contra nty' => kM'M; rewrite -order_eq1 (pnat_1 kMy') //. + exact: mem_p_elt kM'M My'. +have [_ [defNK _] [ntKs _] _ _] := Ptype_structure PmaxM hallK. +have [Mst [PmaxMst _] [_ [hallKst _] [defK _]]] := Ptype_embedding PmaxM hallK. +rewrite -/Ks -/Z -/Zhat in ntKs hallKst * => _ [_ _ conjMMst _]. +have [_ _ [ntK _] _ _] := Ptype_structure PmaxMst hallKst. +have [maxM _] := setDP PmaxM; rewrite defK in ntK. +have [|//|tiZPcover notPcovZ]: _ /\ ClZhat \notin _ := trivIsetU1 _ tiPcover _. + move=> HcovG; case/imsetP=> H maxH ->{HcovG}. + rewrite -setI_eq0 /ClZhat !class_supportEr big_distrr big1 //= => a _. + rewrite big_distrl big1 //= => b _; apply/eqP; rewrite -cards_eq0. + rewrite -(cardJg _ b^-1) conjIg conjsgK -conjsgM -sigma_supportJ cards_eq0. + wlog ->: a b H maxH / H :^ (a * b^-1) = H. + by move/(_ a a (H :^ (a * b^-1))%G); rewrite mmaxJ mulgV act1 => ->. + rewrite setIC big_distrl big1 //= => y Hs_y; apply/setP=> x; rewrite in_set0. + rewrite 3!inE mem_lcoset negb_or -andbA; apply/and4P=> [[yRx notKx notKs_x]]. + rewrite /Z cent_joinEr ?subsetIr //; case/mulsgP=> z z' Kz Ks_z' defx. + have:= sigma_decomposition_dichotomy (group1_contra notKx). + rewrite (introT exists_inP) /=; last first. + by exists y; rewrite // (Msigma_ell1 maxH). + have [Ms_z' cKz'] := setIP Ks_z'; case/exists_inP; exists z'. + rewrite (Msigma_ell1 maxM) ?inE // Ms_z' andbT. + by apply: contraNneq notKx => z'1; rewrite defx z'1 mulg1. + apply/existsP; exists M; rewrite inE maxM cycle_subG Ms_z'. + rewrite defx -(centP cKz') // mulKg (mem_p_elt (pHall_pgroup hallK)) //=. + rewrite 3!inE (subsetP (pHall_sub hallK)) //= cent1C !andbT. + rewrite andbC cent1C (subsetP _ _ Kz) ?sub_cent1 //=. + by apply: contraNneq notKs_x => z1; rewrite defx z1 mul1g. +split=> //; rewrite /partition eqEsubset 2!inE {}tiZPcover negb_or notPcover0. +rewrite /cover big_setU1 {notPcovZ}//= subUset ntPcover subsetD1 subsetT. +rewrite {}/ClZhat {}/Zhat !andbT /= andbC; apply/and3P; split. +- have [[y Ks_y nty] [y' Ky' nty']] := (trivgPn _ ntKs, trivgPn _ ntK). + rewrite eq_sym; apply/set0Pn; exists ((y' * y) ^ 1). + apply: mem_imset2; rewrite 2?inE // groupMl // groupMr // -/Ks negb_or. + have [_ _ _ tiKKs] := dprodP defNK. + rewrite -[Z]genM_join ?mem_gen ?mem_mulg //= andbT; apply/andP; split. + by apply: contra nty => Ky; rewrite -in_set1 -set1gE -tiKKs inE Ky. + by apply: contra nty' => Ks_y'; rewrite -in_set1 -set1gE -tiKKs inE Ky'. +- rewrite class_supportEr; apply/bigcupP=> [[a _]]. + by rewrite mem_conjg conj1g 2!inE !group1. +apply/subsetP=> x; case/setD1P=> ntx _; apply/setUP. +case: exists_inP (sigma_decomposition_dichotomy ntx) => [[y ell1y yRx] _ | _]. + have [nty] := ell_sigma1P _ ell1y; case/set0Pn=> H; case/setIdP=> maxH Hs_y. + right; apply/bigcupP; exists (class_support H^~~ G); first exact: mem_imset. + rewrite -[x]conjg1 mem_imset2 ?inE //; apply/bigcupP. + by exists y; rewrite ?mem_lcoset // !inE nty -cycle_subG. +case/exists_inP=> y ell1y /existsP[H]; set y' := y^-1 * x. +case/and3P=> /setIdP[maxH Hs_y] /setD1P[nty' /setIP[Hy' cyy']] kHy'. +rewrite {ntK ntKs maxM defNK}/Z /Ks; left. +wlog{Ks Mst PmaxMst hallKst conjMMst defK maxH} defH: M K PmaxM hallK / H :=: M. + move=> IH; have PmaxH: H \in 'M_'P. + apply/PtypeP; split=> //; exists (pdiv #[y']). + by rewrite (pnatPpi kHy') // pi_pdiv order_gt1. + have [|] := setUP (conjMMst H PmaxH); case/imsetP=> a Ga defH. + have:= IH _ (K :^ a)%G _ _ defH. + rewrite (eq_pHall _ _ (kappaJ _ _)) pHallJ2 PtypeJ MsigmaJ centJ. + by rewrite -conjIg -conjUg -conjYg -conjDg class_supportGidl //; apply. + have:= IH _ [group of Ks :^ a] _ _ defH. + rewrite (eq_pHall _ _ (kappaJ _ _)) pHallJ2 PtypeJ MsigmaJ centJ. + rewrite -conjIg -conjUg -conjYg -conjDg setUC joingC defK. + by rewrite class_supportGidl //; apply. +have /andP[sMsM nMsM]: M`_\sigma <| M := pcore_normal _ M. +have{Hs_y} Ms_y: y \in M`_\sigma by rewrite -defH -cycle_subG. +wlog{H defH Hy' kHy'} Ky': K hallK / y' \in K. + have [maxM _] := setDP PmaxM; rewrite -cycle_subG defH in Hy' kHy'. + have [a Ma Ka_y'] := Hall_subJ (mmax_sol maxM) hallK Hy' kHy'. + move/(_ (K :^ a)%G); rewrite pHallJ // -cycle_subG. + rewrite -{1 2}(normsP nMsM a Ma) centJ -conjIg -conjYg -conjUg -conjDg. + by rewrite class_supportGidl ?inE //; apply. +rewrite -[x]conjg1 mem_imset2 ?group1 //. +have [Mst _ [_ _ _ [cycZ _ defZ _ _] _]] := Ptype_embedding PmaxM hallK. +rewrite -(mulKVg y x) -/y' 2!inE negb_or andbC. +do [set Ks := 'C_(_)(K); set Z := K <*> _] in cycZ defZ *. +have Ks_y: y \in Ks. + have cKZ := sub_abelian_cent (cyclic_abelian cycZ) (joing_subl K Ks). + rewrite inE Ms_y (subsetP cKZ) // -(defZ y'); last by rewrite !inE nty'. + by rewrite inE cent1C (subsetP sMsM). +have [_ [defNK _] _ _ _] := Ptype_structure PmaxM hallK. +have{defNK} [_ _ cKKs tiKKs] := dprodP defNK. +rewrite [Z]joingC cent_joinEl // mem_mulg // groupMr // groupMl //= -/Ks. +apply/andP; split. + have [nty _] := ell_sigma1P _ ell1y. + by apply: contra nty => Ky; rewrite -in_set1 -set1gE -tiKKs inE Ky. +by apply: contra nty' => Ks_y'; rewrite -in_set1 -set1gE -tiKKs inE Ky'. +Qed. + +(* This is B & G, Corollary 14.10. *) +Corollary ell_sigma_leq_2 x : \ell_\sigma(x) <= 2. +Proof. +have [/ell_sigma0P/eqP-> // | ntx] := eqVneq x 1. +case sigma_x: (x \in cover [set class_support M^~~ G | M : {group gT} in 'M]). + case/bigcupP: sigma_x => _ /imsetP[M maxM ->]. + case/imset2P=> x0 a /bigcupP[y Ms_y yRx0] _ ->; rewrite ell_sigmaJ. + exact: ell_sigma_cover (Msigma_ell1 maxM Ms_y) yRx0. +have G1x: x \in G^# by rewrite !inE andbT. +have [FpartG1 PpartG1] := mFT_partition. +have [/eqP/FpartG1 partG1 | [M PmaxM]] := set_0Vmem ('M_'P : {set {group gT}}). + by rewrite -(cover_partition partG1) sigma_x in G1x. +have [maxM _] := setDP PmaxM. +have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). +have{PpartG1} [/cover_partition defG1 notZsigma] := PpartG1 M K PmaxM hallK. +rewrite -{}defG1 /cover big_setU1 {notZsigma}// inE {}sigma_x orbF in G1x. +case/imset2P: G1x => x0 a /setDP[]. +have [Mst [PmaxMst _] [_ _ [defK _] _ _]] := Ptype_embedding PmaxM hallK. +rewrite cent_joinEr ?subsetIr // => /mulsgP[y' y Ky' /= Ks_y ->]. +rewrite inE; have [-> | nty] := eqVneq y 1; first by rewrite mulg1 Ky'. +have [-> | nty' _ _ ->] := eqVneq y' 1; first by rewrite mul1g Ks_y orbT. +have [Ms_y cKy] := setIP Ks_y; set Ks := 'C_(_)(_) in Ks_y defK. +have Msts_y': y' \in Mst`_\sigma by move: Ky'; rewrite -defK => /setIP[]. +have kMy': \kappa(M).-elt y' := mem_p_elt (pHall_pgroup hallK) Ky'. +have{kMy'} sM'y': \sigma(M)^'.-elt y' := sub_pgroup (@kappa_sigma' _) kMy'. +rewrite ell_sigmaJ /sigma_length (cardsD1 (y' * y).`_\sigma(M)). +rewrite (leq_add (leq_b1 _)) // -sigma_decomposition_constt' //. +rewrite consttM /commute ?(centP cKy) // constt_p_elt //. +rewrite (constt1P _) ?p_eltNK ?(mem_p_elt (pcore_pgroup _ _) Ms_y) // mulg1. +have [maxMst _] := setDP PmaxMst; rewrite leq_eqVlt (Msigma_ell1 maxMst) //. +by rewrite !inE nty' Msts_y'. +Qed. + +(* This is B & G, Lemma 14.11. *) +Lemma primes_non_Fitting_Ftype M E q Q : + M \in 'M_'F -> \sigma(M)^'.-Hall(M) E -> + Q \in 'E_q^1(E) -> ~~ (Q \subset 'F(E)) -> + exists2 Mstar, Mstar \in 'M & + [\/ (*1*) q \in \tau2(Mstar) /\ 'M('C(Q)) = [set Mstar] + | (*2*) q \in \kappa(Mstar) /\ Mstar \in 'M_'P1 ]. +Proof. +move=> FmaxM hallE EqQ notsFE_Q; have [maxM k'M] := FtypeP _ FmaxM. +have [sQE abelQ dimQ] := pnElemP EqQ; have [qQ _] := andP abelQ. +have [q_pr oQ] := (pnElem_prime EqQ, card_pnElem EqQ : #|Q| = q). +have t1Mq: q \in \tau1(M). + have: q \in \pi(E) by rewrite -p_rank_gt0; apply/p_rank_geP; exists Q. + rewrite (partition_pi_sigma_compl maxM hallE) => /or3P[// | t2q | t3q]. + have [A EqA _] := ex_tau2Elem hallE t2q. + have [[nsAE defA1] _ _ _] := tau2_compl_context maxM hallE t2q EqA. + have sQA: Q \subset A by move: EqQ; rewrite defA1 => /pnElemP[]. + rewrite (subset_trans sQA) ?Fitting_max // ?abelian_nil // in notsFE_Q. + by have [_ abelA _] := pnElemP EqA; apply: abelem_abelian abelA. + have [[E1 hallE1] [E3 hallE3]] := ex_tau13_compl hallE. + have [E2 _ complEi] := ex_tau2_compl hallE hallE1 hallE3. + have [[_ nsE3E] _ [_ cycE3] _ _] := sigma_compl_context maxM complEi. + have sQE3: Q \subset E3 by rewrite (sub_normal_Hall hallE3) ?(pi_pgroup qQ). + rewrite (subset_trans sQE3) ?Fitting_max ?abelian_nil // in notsFE_Q. + exact: cyclic_abelian cycE3. +have q'FE: q^'.-group 'F(E). + have [R sylR sQR] := Sylow_superset sQE qQ; have [sRE qR _] := and3P sylR. + have cycR: cyclic R. + rewrite (odd_pgroup_rank1_cyclic qR) ?mFT_odd // (p_rank_Sylow sylR) //. + by move: t1Mq; rewrite (tau1E maxM hallE) eqn_leq; case/and4P. + rewrite -partG_eq1 -(card_Hall (Hall_setI_normal (Fitting_normal E) sylR)). + have sRFER: R :&: 'F(E) \subset R := subsetIl _ _. + apply: contraR notsFE_Q; rewrite -trivg_card1 => ntRFE. + rewrite (subset_trans _ (subsetIr R _)) // -(cardSg_cyclic cycR) // oQ. + by have [] := pgroup_pdiv (pgroupS sRFER qR) ntRFE. +have cE'E': abelian E^`(1) := der_mmax_compl_abelian maxM hallE. +pose K := [~: E, Q]; have cKK: abelian K by rewrite (abelianS (commgS E sQE)). +have nsKE: K <| E by rewrite /normal commg_norml comm_subG. +have q'K: q^'.-group K by rewrite (pgroupS _ q'FE) // Fitting_max ?abelian_nil. +have [sKE nKE] := andP nsKE; have nKQ := subset_trans sQE nKE. +have defKQ: [~: K, Q] = K. + have nsKQ_E: K <*> Q <| E. + rewrite -(quotientGK nsKE) -(quotientYK nKQ) cosetpre_normal /= -/K. + by rewrite /normal quotientS // cents_norm // quotient_cents2r. + have [_ sylQ] := coprime_mulGp_Hall (esym (norm_joinEr nKQ)) q'K qQ. + have defE: K * 'N_E(Q) = E. + rewrite -{2}(Frattini_arg nsKQ_E sylQ) /= norm_joinEr //= -/K -mulgA. + by congr (K * _); rewrite mulSGid // subsetI sQE normG. + have cQ_NEQ: [~: 'N_E(Q), Q] = 1. + apply/trivgP; rewrite -(coprime_TIg (pnat_coprime qQ q'K)) subsetI. + by rewrite commg_subr subsetIr commSg ?subsetIl. + by rewrite {2}/K -defE commMG ?cQ_NEQ ?mulg1 1?normsR ?subsetIr ?subIset ?nKE. +have [sEM s'E _] := and3P hallE; have sQM := subset_trans sQE sEM. +have [sKM s'K] := (subset_trans sKE sEM, pgroupS sKE s'E). +have regQ: 'C_(M`_\sigma)(Q) = 1. + apply/eqP; apply: contraFT (k'M q) => nregQ. + have EqQ_M: Q \in 'E_q^1(M) by apply/pnElemP. + by rewrite unlock 3!inE /= t1Mq; apply/exists_inP; exists Q. +have nsKM: K <| M. + have [s'q _] := andP t1Mq. + have EqQ_NK: Q \in 'E_q^1('N_M(K)) by apply/pnElemP; rewrite subsetI sQM. + have:= commG_sigma'_1Elem_cyclic maxM sKM s'K s'q EqQ_NK regQ q'K cKK. + by rewrite defKQ; case. +have ntK: K != 1. + apply: contraNneq notsFE_Q => /commG1P cQE. + by rewrite Fitting_max ?(pgroup_nil qQ) // /normal sQE cents_norm. +pose p := pdiv #|K|; have p_pr: prime p by rewrite pdiv_prime ?cardG_gt1. +have piKp: p \in \pi(K) by rewrite pi_pdiv cardG_gt1. +have t2Mp: p \in \tau2(M). + have s'p := pnatPpi s'K piKp. + have sylKp: p.-Sylow(K) 'O_p(K) := nilpotent_pcore_Hall p (abelian_nil cKK). + rewrite inE /= s'p ?(sigma'_norm_mmax_rank2 maxM s'p (pHall_pgroup sylKp)) //. + rewrite (mmax_normal maxM) ?(char_normal_trans (pcore_char _ _)) //. + by rewrite -rank_gt0 (rank_Sylow sylKp) p_rank_gt0. +have [A EpA _] := ex_tau2Elem hallE t2Mp. +have [sAE] := pnElemP EpA; case/andP=> pA _ dimA. +have [[nsAE _] _ _ _] := tau2_compl_context maxM hallE t2Mp EpA. +have nAQ := subset_trans sQE (normal_norm nsAE). +have [S sylS sAS]:= Sylow_superset (subsetT A) pA. +have not_cSS: ~~ abelian S. + apply: contra notsFE_Q => cSS; rewrite Fitting_max ?(pgroup_nil qQ) //. + have solE := sigma_compl_sol hallE. + have [E1 hallE1 sQE1] := Hall_superset solE sQE (pi_pgroup qQ t1Mq). + have [_ [E3 hallE3]] := ex_tau13_compl hallE. + have [E2 _ complEi] := ex_tau2_compl hallE hallE1 hallE3. + have [_ _ _ sQZ] := abelian_tau2 maxM complEi t2Mp EpA sylS sAS cSS. + by rewrite sub_center_normal ?{}sQZ //; apply/nElemP; exists q; apply/pnElemP. +have [] := nonabelian_tau2 maxM hallE t2Mp EpA (pHall_pgroup sylS) not_cSS. +set A0 := 'C_A(M`_\sigma)%G => _ [oA0 defFM] _ _. +have defA0: A0 :=: K. + have sA0E: A0 \subset E by rewrite subIset ?sAE. + have sKA0: K \subset A0. + have [_ _ _ tiMsE] := sdprodP (sdprod_sigma maxM hallE). + rewrite -(mul1g A0) -tiMsE setIC group_modr // subsetI sKE. + by have [_ -> _ _] := dprodP defFM; rewrite Fitting_max ?abelian_nil. + by apply/eqP; rewrite eqEsubset prime_meetG ?(setIidPr sKA0) ?oA0. +have ntA: A :!=: 1 := nt_pnElem EpA isT. +have [H maxNH] := mmax_exists (mFT_norm_proper ntA (mFT_pgroup_proper pA)). +have [maxH sNH] := setIdP maxNH; have sQH := subset_trans nAQ sNH. +exists H => //. +have: p \in [predD \sigma(H) & \beta(H)] /\ q \in [predU \tau1(H) & \tau2(H)]. + have [-> // piAb _] := primes_norm_tau2Elem maxM hallE t2Mp EpA maxNH. + rewrite (pnatPpi piAb) // (piSg (quotientS _ sQE)) //=. + have piQq: q \in \pi(Q) by rewrite -p_rank_gt0 p_rank_abelem ?dimQ. + rewrite /= card_quotient ?normsI ?norms_cent // ?normsG //. + rewrite -indexgI setIA (setIidPl sQE) prime_TIg ?indexg1 // ?oQ //. + rewrite (sameP commG1P eqP) (subG1_contra _ ntK) //= -/K -defKQ commGC. + by rewrite -defA0 commgS ?subsetIl. +case=> /andP[/= b'Hp sHP] t12Hq. +have nregQHs: 'C_(H`_\sigma)(Q) != 1. + apply: subG1_contra (setSI _ _) (_ : 'C_A(Q) != 1). + rewrite (sub_Hall_pcore (Msigma_Hall maxH)) ?(pi_pgroup pA) //. + exact: subset_trans (normG A) sNH. + apply: contraTneq (leqnn 1) => regQA; rewrite -ltnNge -dimA. + rewrite -(leq_exp2l _ _ (prime_gt1 p_pr)) -card_pgroup // -oA0 defA0. + have coAQ := pnat_coprime (pi_pnat pA t2Mp) (pi_pnat qQ (tau2'1 t1Mq)). + rewrite subset_leq_card // -(coprime_cent_prod nAQ) ?(pgroup_sol pA) //. + by rewrite regQA mulg1 commSg. +have{t12Hq} [/= t1Hq | /= t2Hq] := orP t12Hq. + have EqQ_H: Q \in 'E_q^1(H) by apply/pnElemP. + have kHq: q \in \kappa(H). + by rewrite unlock 3!inE /= t1Hq; apply/exists_inP; exists Q. + right; split=> //; apply: contraR b'Hp => notP1maxH. + have PmaxH: H \in 'M_'P by apply/PtypeP; split=> //; exists q. + have [L hallL] := Hall_exists \kappa(H) (mmax_sol maxH). + by have [_ _ _ _ [|<- //]] := Ptype_structure PmaxH hallL; apply/setDP. +left; split=> //. +have [x defQ]: exists x, Q :=: <[x]> by apply/cyclicP; rewrite prime_cyclic ?oQ. +rewrite defQ cent_cycle in nregQHs *; rewrite (cent1_nreg_sigma_uniq maxH) //. + by rewrite 2!inE -cycle_eq1 -cycle_subG -defQ (nt_pnElem EqQ). +by rewrite /p_elt /order -defQ oQ pnatE. +Qed. + +(* This is B & G, Lemma 14.12. *) +(* Note that the assumption M \in 'M_'P2 could be weakened to M \in 'M_'P, *) +(* since the assumption H \in 'M('N(R)) implies H != 1, and hence U != 1. *) +Lemma P2type_signalizer M Mstar U K r R H : + M \in 'M_'P2 -> kappa_complement M U K -> Mstar \in 'M('C(K)) -> + r.-Sylow(U) R -> H \in 'M('N(R)) -> + [/\ H \in 'M_'F, U \subset H`_\sigma, U <*> K = M :&: H + & [/\ ~~ ('N_H(U) \subset M), K \subset 'F(H :&: Mstar) + & \sigma(H)^'.-Hall(H) (H :&: Mstar)]]. +Proof. +move=> P2maxM complU maxCMstar sylR maxNH; have [hallU hallK _] := complU. +have [PmaxM notP1maxM] := setDP P2maxM; have [maxM notFmaxM] := setDP PmaxM. +have [[sUM sk'M_U _] [sKM kK _]] := (and3P hallU, and3P hallK). +have{notFmaxM} ntK: K :!=: 1 by rewrite (trivg_kappa maxM). +have [hallE defM _ regK /(_ ntK)cUU] := kappa_compl_context maxM complU. +case/sdprodP: defM => [[_ E _ defE] _ _ _]. +have [nsUE sKE mulUK nUK tiUK] := sdprod_context defE. +rewrite (norm_joinEr nUK) mulUK in hallE *. +have [Mst [PmaxMst notMGMst] [uniqMst []]] := Ptype_embedding PmaxM hallK. +set Ks := 'C_(_)(K) => hallKs; case/and3P=> sKsMst sM_Ks _ [defK _]. +case=> cycZ ziMMst _ _ _ [_ _ defPmax _]. +have [_ [defNK _] [ntKs _] _ [//|_ q_pr _ _]] := Ptype_structure PmaxM hallK. +set q := #|K| in q_pr. +have{uniqMst} uniqMst: 'M('C(K)) = [set Mst]. + by apply: uniqMst; apply/nElemP; exists q; rewrite p1ElemE // !inE subxx /=. +have{maxCMstar} ->: Mstar = Mst by [apply/set1P; rewrite -uniqMst] => {Mstar}. +have [maxH sNRH] := setIdP maxNH. +have ntR: R :!=: 1. + by apply: contraTneq sNRH => ->; rewrite norm1 proper_subn ?mmax_proper. +have piUr: r \in \pi(U) by rewrite -p_rank_gt0 -(rank_Sylow sylR) rank_gt0. +have r_pr: prime r by move: piUr; rewrite mem_primes; case/andP. +have sylR_M := subHall_Sylow hallU (pnatPpi sk'M_U piUr) sylR. +have [/= s'Mr k'Mr] := norP (pnatPpi sk'M_U piUr). +have [sRH [sRM rR _]] := (subset_trans (normG R) sNRH, and3P sylR_M). +have notMGH: gval H \notin M :^: G. + apply: contra s'Mr; case/imsetP=> a _ defH. + rewrite -(sigmaJ _ a) -defH; apply/exists_inP; exists R => //. + by rewrite pHallE sRH /= (card_Hall sylR_M) defH cardJg. +have sK_uniqMst a: K \subset Mst :^ a -> a \in Mst. + move=> sKMa; apply: contraR ntK; rewrite -in_setC => Mst'a. + have [_ _ _ [[tiK_MstG _] _ _] _] := Ptype_structure PmaxMst hallKs. + by rewrite -(tiK_MstG a) // defK (setIidPl sKMa). +have [_ _] := dprodP defNK; rewrite -/Ks => cKKs tiKKs. +have snK_sMst L: K <|<| L -> L \subset Mst. + elim: {L}_.+1 {-2}L (ltnSn #|L|) => // n IHn A leAn. + case/subnormalEr=> [<- | [L [snKL nsLA ltLA]]]. + by rewrite -defK subIset // pcore_sub. + have [sKL sLMst]: K \subset L /\ L \subset Mst. + by rewrite subnormal_sub // IHn // (leq_trans (proper_card ltLA)). + apply/subsetP=> a Aa; rewrite -groupV sK_uniqMst // (subset_trans sKL) //. + by rewrite -sub_conjg (normsP (normal_norm nsLA)). +have sEH: E \subset H. + apply: subset_trans (char_norm_trans _ (normal_norm nsUE)) sNRH. + by rewrite (nilpotent_Hall_pcore (abelian_nil cUU) sylR) pcore_char. +have [sUH sKH]: U \subset H /\ K \subset H by apply/mulGsubP; rewrite mulUK. +have notMstGH: gval H \notin Mst :^: G. + apply: contra ntR => /imsetP[a _ defH]. + have{a defH} defH: H :=: Mst by rewrite -(conjGid (sK_uniqMst a _)) -?defH. + rewrite -(setIidPl sRH) -(setIidPl sRM) -setIA defH ziMMst coprime_TIg //=. + rewrite cent_joinEr // TI_cardMg //= coprime_mulr -/Ks. + rewrite (p'nat_coprime (pi_pnat rR _) kK) //=. + exact: p'nat_coprime (pi_pnat rR _) sM_Ks. +have FmaxH: H \in 'M_'F. + suffices: H \notin 'M_'P by rewrite inE maxH andbT negbK. + by apply: (contra (defPmax H)); rewrite inE; apply/norP. +have sKMsts: K \subset Mst`_\sigma by rewrite -defK subsetIl. +have s'H_K: \sigma(H)^'.-group K. + apply/pgroupP=> p p_pr p_dv_K; have [maxMst _] := setDP PmaxMst. + apply: contraFN (sigma_partition maxMst maxH notMstGH p) => /= sHp. + by rewrite inE /= (pgroupP (pgroupS sKMsts _)) ?pcore_pgroup. +have [D hallD sKD] := Hall_superset (mmax_sol maxH) sKH s'H_K. +have piKq: q \in \pi(K) by rewrite pi_of_prime ?inE. +have sK_FD: K \subset 'F(D). + have EqK: K \in 'E_q^1(D) by rewrite p1ElemE // !inE sKD /=. + have sMst_q: q \in \sigma(Mst). + by rewrite (pnatPpi (pcore_pgroup _ _) (piSg sKMsts _)). + apply: contraR notP1maxM => not_sKFD. + have [L _ ] := primes_non_Fitting_Ftype FmaxH hallD EqK not_sKFD. + case=> [[t2Lq ]|[kLq P1maxL]]. + rewrite uniqMst => /set1_inj defL. + by rewrite -defL 3!inE sMst_q in t2Lq. + have [PmaxL _] := setIdP P1maxL. + case/setUP: (defPmax L PmaxL) => /imsetP[a _ defL]. + by rewrite (group_inj defL) P1typeJ in P1maxL. + move: kLq; rewrite defL kappaJ unlock 3!inE /=. + by rewrite -andb_orr inE /= sMst_q. +have sDMst: D \subset Mst. + apply: snK_sMst (subnormal_trans _ (normal_subnormal (Fitting_normal D))). + exact: nilpotent_subnormal (Fitting_nil D) sK_FD. +have defUK: [~: U, K] = U. + rewrite -{2}(coprime_cent_prod nUK) ?abelian_sol //; last first. + by apply: p'nat_coprime (sub_pgroup _ sk'M_U) kK => ? /norP[]. + by rewrite (cent_semiregular regK) ?mulg1. +have qK: q.-group K := pnat_id q_pr. +have sUHs: U \subset H`_\sigma. + have [nsHsH _ mulHsD nHsD _] := sdprod_context (sdprod_sigma maxH hallD). + have nHsDq := subset_trans (pcore_sub q D) nHsD. + pose HsDq := H`_\sigma <*> 'O_q(D). + have defHsDq: H`_\sigma * 'O_q(D) = HsDq by rewrite -norm_joinEr. + have hallHs_HsDq: q^'.-Hall(HsDq) H`_\sigma. + have [|//] := coprime_mulGp_Hall defHsDq _ (pcore_pgroup _ _). + rewrite p'groupEpi; apply: (contra (pnatPpi (pcore_pgroup _ _))). + exact: pnatPpi s'H_K piKq. + have sK_HsDq: K \subset HsDq. + rewrite sub_gen ?subsetU // orbC -p_core_Fitting. + by rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ (Fitting_nil _))) ?qK. + have [|sHsDq_H nHsDq_H] := andP (_ : HsDq <| H). + rewrite -(quotientGK nsHsH) -[HsDq]quotientYK //= cosetpre_normal //. + by rewrite -{3}mulHsD quotientMidl quotient_normal // pcore_normal. + have sU_HsDq: U \subset HsDq. + by rewrite -defUK (subset_trans (commgSS sUH sK_HsDq)) // commg_subr. + rewrite (sub_normal_Hall hallHs_HsDq) //. + rewrite p'groupEpi; apply: (contraL (pnatPpi sk'M_U)) => /=. + by rewrite inE /= orbC (pnatPpi kK). + exact: normalS (joing_subl _ _) _ (pcore_normal _ _). +have defNMU: 'N_M(U) = E. + have [_ mulHsE nHsE _] := sdprodP (sdprod_sigma maxM hallE). + have [sUE nUE] := andP nsUE; rewrite -mulHsE -normC // -group_modl //=. + rewrite coprime_norm_cent ?(subset_trans sUE) //; last first. + exact: coprimegS sUE (coprime_sigma_compl hallE). + have sR1U: 'Ohm_1(R) \subset U := subset_trans (Ohm_sub 1 R) (pHall_sub sylR). + rewrite (trivgP (subset_trans (setIS _ (centS sR1U)) _)) ?mulg1 //. + have [|_ _ -> //] := sigma'_kappa'_facts maxM sylR_M. + by rewrite s'Mr (piSg sUM). +have sHsFH: H`_\sigma \subset 'F(H). + rewrite Fitting_max ?pcore_normal //. + have [S] := Sylow_exists q H; case/sigma'_kappa'_facts=> {S}//. + have [_ k'H] := setIdP FmaxH. + rewrite [~~ _](pnatPpi (pHall_pgroup hallD) (piSg sKD _)) //=. + by rewrite [~~ _](pnatPpi k'H) (piSg sKH). +suffices ->: H :&: Mst = D. + set sk' := _^' in sk'M_U hallU; pose Fu := 'O_sk'('F(H)). + have [sUFH nilFH] := (subset_trans sUHs sHsFH, Fitting_nil H). + have hallFu: sk'.-Hall('F(H)) Fu := nilpotent_pcore_Hall sk' nilFH. + have sUFu: U \subset Fu by rewrite (sub_Hall_pcore hallFu). + have nsFuH: Fu <| H := char_normal_trans (pcore_char _ _) (Fitting_normal _). + have [[sFuFH sk'Fu _] [sFuH nFuH]] := (and3P hallFu, andP nsFuH). + have defU: M :&: Fu = U. + have sk'MFu: sk'.-group(M :&: Fu) := pgroupS (subsetIr M _) sk'Fu. + by rewrite (sub_pHall hallU sk'MFu) ?subsetIl // subsetI sUM. + do 2?split=> //. + apply/eqP; rewrite eqEsubset subsetI (pHall_sub hallE) sEH /=. + by rewrite -defNMU subsetI subsetIl -defU normsGI. + apply: contra (contra_orbit _ _ notMGH) => sNHU_M. + rewrite (eq_mmax maxH maxM (subset_trans _ sNHU_M)) // subsetIidl. + rewrite -(nilpotent_sub_norm (nilpotentS sFuFH nilFH) sUFu) //= -/Fu. + by rewrite -{2}defU subsetI subsetIl (subset_trans (setSI _ sFuH)). +have [maxMst _] := setDP PmaxMst. +have [_ <- _ _] := sdprodP (sdprod_sigma maxH hallD). +rewrite -{2}(mul1g D) setIC -group_modr // setIC; congr (_ * _). +apply/eqP; apply: wlog_neg => ntHsMst. +have nregHsK: 'C_(H`_\sigma)(K) != 1. + rewrite (subG1_contra _ ntHsMst) // subsetI subsetIl (sameP commG1P trivgP). + have <-: H`_\sigma :&: Mst`_\sigma = 1. + apply: card_le1_trivg; rewrite leqNgt -pi_pdiv; set p := pdiv _. + apply: contraFN (sigma_partition maxMst maxH notMstGH p) => piIp. + rewrite inE /= (pnatPpi (pcore_pgroup _ _) (piSg (subsetIl _ _) piIp)). + by rewrite (pnatPpi (pcore_pgroup _ _) (piSg (subsetIr _ _) piIp)). + rewrite commg_subI ?setIS ?gFnorm // subsetI sKMsts. + by rewrite (subset_trans sKH) ?gFnorm. +have t2Hq: q \in \tau2(H). + have: q \in \pi(D) := piSg sKD piKq. + rewrite (partition_pi_sigma_compl maxH hallD) orbCA; case/orP=> // t13Hq. + case/FtypeP: FmaxH => _ /(_ q)/idP[]; rewrite unlock 3!inE /= t13Hq. + by apply/exists_inP; exists K => //; rewrite p1ElemE // !inE sKH /=. +have [A EqA_D EqA] := ex_tau2Elem hallD t2Hq. +have [_ _ _ -> //] := tau2_context maxH t2Hq EqA. +rewrite 3!inE -val_eqE /= eq_sym (contra_orbit _ _ notMstGH) maxMst. +by have [sAD _ _] := pnElemP EqA_D; apply: subset_trans sAD sDMst. +Qed. + +(* This is B & G, Lemma 14.13(a). *) +(* Part (b) is not needed for the Peterfalvi revision of the character theory *) +(* part of the proof. *) +Lemma non_disjoint_signalizer_Frobenius x M : + \ell_\sigma(x) == 1%N -> #|'M_\sigma[x]| > 1 -> + M \in 'M_\sigma[x] -> ~~ (\sigma('N[x])^'.-group M) -> + M \in 'M_'F /\ \tau2(M)^'.-group M. +Proof. +move=> ell1x ntR SMxM; have [maxM Ms_x] := setIdP SMxM. +rewrite negb_and cardG_gt0 all_predC negbK => /hasP[q /= piMq sNq]. +have [Q EqQ]: exists Q, Q \in 'E_q^1(M) by apply/p_rank_geP; rewrite p_rank_gt0. +have [ntQ [sQM abelQ dimQ]] := (nt_pnElem EqQ isT, pnElemP EqQ). +have [[qQ _] q_pr] := (andP abelQ, pnElem_prime EqQ). +have [_ [//| uniqN _ t2Nx _]] := FT_signalizer_context ell1x. +case/(_ M SMxM)=> _ st2NsM spM_sbN _; have [maxN sCxN] := mem_uniq_mmax uniqN. +have bNq: q \in \beta('N[x]) by rewrite spM_sbN //= 4!inE /= piMq. +have bGq: q \in \beta(G) by move: bNq; rewrite -predI_sigma_beta // inE /= sNq. +set p := pdiv #[x]; have pi_p: p \in \pi(#[x]). + by rewrite pi_pdiv order_gt1 (sameP eqP (ell_sigma0P _)) (eqP ell1x). +have sMp: p \in \sigma(M) := pnatPpi (pcore_pgroup _ _) (piSg Ms_x pi_p). +have t2Np: p \in \tau2('N[x]) := pnatPpi t2Nx pi_p. +have notMGN: gval 'N[x] \notin M :^: G. + apply: contraL t2Np => /imsetP[a _ ->]. + by rewrite negb_and negbK /= sigmaJ sMp. +have sM'q: q \in \sigma(M)^'. + by apply: contraFN (sigma_partition maxM maxN notMGN q) => sMq; apply/andP. +have [g sQNg]: exists g, Q \subset 'N[x] :^ g. + have [Q1 sylQ1] := Sylow_exists q 'N[x]. + have [g _ sQQ1g] := Sylow_subJ (sigma_Sylow_G maxN sNq sylQ1) (subsetT Q) qQ. + by exists g; rewrite (subset_trans sQQ1g) // conjSg (pHall_sub sylQ1). +have EqNQ: Q \in 'E_q^1('N[x] :^ g) by apply/pnElemP. +have uniqNg: 'M('C(Q)) = [set 'N[x] :^ g]%G. + by case/cent_der_sigma_uniq: EqNQ; rewrite ?mmaxJ 1?betaJ ?bNq. +have b'Mp: p \notin \beta(M). + by rewrite -predI_sigma_beta // inE /= sMp /=; case/tau2_not_beta: t2Np. +have{p pi_p sMp t2Np b'Mp} FmaxM: M \in 'M_'F. + have [P1maxM | notP1maxM] := boolP (M \in 'M_'P1); last first. + have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). + apply: contraR b'Mp => notFmaxM; have PmaxM: M \in 'M_'P by apply/setDP. + by have [_ _ _ _ [|<- //]] := Ptype_structure PmaxM hallK; apply/setDP. + have [PmaxM skM] := setIdP P1maxM. + have kMq: q \in \kappa(M). + by case/orP: (pnatPpi skM piMq) => //= sMq; case/negP: sM'q. + have [K hallK sQK] := Hall_superset (mmax_sol maxM) sQM (pi_pnat qQ kMq). + have EqKQ: Q \in 'E_q^1(K) by apply/pnElemP. + have [L _ [uniqL [kLhallKs sMhallKs] _ _ _]] := Ptype_embedding PmaxM hallK. + set Ks := 'C_(_)(K) in kLhallKs sMhallKs. + have{uniqL} defL: 'N[x] :^ g = L. + apply: congr_group; apply: set1_inj; rewrite -uniqNg uniqL //. + by apply/nElemP; exists q. + have rpL: 'r_p(L) = 2. + by apply/eqP; case/andP: t2Np => _; rewrite -defL p_rankJ. + suffices piKs_p: p \in \pi(Ks). + by rewrite rank_kappa // (pnatPpi (pHall_pgroup kLhallKs)) in rpL. + have [P sylP] := Sylow_exists p [group of Ks]. + have sylP_L: p.-Sylow(L) P := subHall_Sylow sMhallKs sMp sylP. + by rewrite -p_rank_gt0 -(rank_Sylow sylP) (rank_Sylow sylP_L) ?rpL. +split=> //; apply: sub_pgroup (pgroup_pi _) => p piMp; apply/negP=> /= t2Mp. +have rpN: 'r_p('N[x]) <= 1. + have: p \notin \beta('N[x]). + rewrite -(predI_sigma_beta maxN) negb_and /= orbC. + by have [-> _] := tau2_not_beta maxM t2Mp. + apply: contraR; rewrite -ltnNge => rpN; rewrite spM_sbN // inE /= piMp. + have: p \in \pi('N[x]) by rewrite -p_rank_gt0 ltnW. + rewrite (partition_pi_mmax maxN) orbCA => /orP[t2Np | ]. + by case/andP: t2Mp => /negP[]; apply: st2NsM. + by rewrite orbA -!andb_orr eqn_leq leqNgt rpN andbF. +have [E hallE sQE] := Hall_superset (mmax_sol maxM) sQM (pi_pgroup qQ sM'q). +have [A Ep2A _] := ex_tau2Elem hallE t2Mp; have [_ abelA dimA] := pnElemP Ep2A. +pose A0 := [~: A, Q]%G; pose A1 := 'C_A(Q)%G. +have sCQNg: 'C(Q) \subset 'N[x] :^ g by have [] := mem_uniq_mmax uniqNg. +have ntA0: A0 :!=: 1. + rewrite (sameP eqP commG1P); apply: contraL rpN => cQA. + rewrite -ltnNge -(p_rankJ p _ g); apply/p_rank_geP. + by exists A; apply/pnElemP; rewrite (subset_trans cQA). +have t1Mq: q \in \tau1(M). + have [_ nsCEA_E t1Eb] := tau1_cent_tau2Elem_factor maxM hallE t2Mp Ep2A. + rewrite (pnatPpi t1Eb) // (piSg (quotientS _ sQE)) // -p_rank_gt0. + rewrite -rank_pgroup ?quotient_pgroup // rank_gt0 -subG1. + rewrite quotient_sub1 ?(subset_trans _ (normal_norm nsCEA_E)) //. + by rewrite subsetI sQE centsC (sameP commG1P eqP). +have EqEQ: Q \in 'E_q^1(E) by apply/pnElemP. +have regMsQ: 'C_(M`_\sigma)(Q) = 1. + apply: contraTeq FmaxM => nregMsQ; apply/FtypeP=> [[_]]; move/(_ q). + by rewrite unlock 3!inE /= t1Mq; case/exists_inP; exists Q. +have [[]] := tau1_act_tau2 maxM hallE t2Mp Ep2A t1Mq EqEQ regMsQ ntA0. +rewrite -/A0 -/A1 => EpA0 cMsA0 _ notA1GA0 [EpA1 _]. +have [sA0A abelA0 oA0] := pnElemPcard EpA0; have [pA0 _] := andP abelA0. +have [sA1A abelA1 oA1] := pnElemPcard EpA1; have [pA1 _] := andP abelA1. +have sA0N: A0 \subset 'N[x]. + rewrite -cMsA0 (subset_trans _ sCxN) //= -cent_cycle (centsS Ms_x) //. + exact: subsetIr. +have [P sylP sA0P] := Sylow_superset sA0N pA0; have [_ pP _] := and3P sylP. +have cycP: cyclic P. + by rewrite (odd_pgroup_rank1_cyclic pP) ?mFT_odd ?(p_rank_Sylow sylP). +have sA1gN: A1 :^ g^-1 \subset 'N[x] by rewrite sub_conjgV subIset ?sCQNg ?orbT. +have [|z _ sA1gzP] := Sylow_Jsub sylP sA1gN; first by rewrite pgroupJ. +case/imsetP: notA1GA0; exists (g^-1 * z); rewrite ?inE // conjsgM. +by apply/eqP; rewrite (eq_subG_cyclic cycP) // !cardJg oA0 oA1. +Qed. + +End Section14. + + diff --git a/mathcomp/odd_order/BGsection15.v b/mathcomp/odd_order/BGsection15.v new file mode 100644 index 0000000..2238534 --- /dev/null +++ b/mathcomp/odd_order/BGsection15.v @@ -0,0 +1,1509 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq choice div fintype. +Require Import path bigop finset prime fingroup morphism perm automorphism. +Require Import quotient action gproduct gfunctor pgroup cyclic commutator. +Require Import center gseries nilpotent sylow abelian maximal hall frobenius. +Require Import BGsection1 BGsection2 BGsection3 BGsection4 BGsection5. +Require Import BGsection6 BGsection7 BGsection9 BGsection10 BGsection12. +Require Import BGsection13 BGsection14. + +(******************************************************************************) +(* This file covers B & G, section 15; it fills in the picture of maximal *) +(* subgroups that was sketched out in section14, providing an intrinsic *) +(* characterization of M`_\sigma and establishing the TI property for the *) +(* "kernels" of maximal groups. We introduce only one new definition: *) +(* M`_\F == the (direct) product of all the normal Sylow subgroups of M; *) +(* equivalently, the largest normal nilpotent Hall subgroup of M *) +(* We will refer to M`_\F as the Fitting core or F-core of M. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Section Definitions. + +Variables (gT : finGroupType) (M : {set gT}). + +Definition Fitting_core := + <<\bigcup_(P : {group gT} | Sylow M P && (P <| M)) P>>. +Canonical Structure Fitting_core_group := [group of Fitting_core]. + +End Definitions. + +Notation "M `_ \F" := (Fitting_core M) + (at level 3, format "M `_ \F") : group_scope. +Notation "M `_ \F" := (Fitting_core_group M) : Group_scope. + +Section FittingCore. + +Variable (gT : finGroupType) (M : {group gT}). +Implicit Types H P : {group gT}. +Implicit Type p : nat. + +Lemma Fcore_normal : M`_\F <| M. +Proof. +rewrite -[M`_\F]bigprodGE. +elim/big_ind: _ => [|P Q nsP nsG|P /andP[] //]; first exact: normal1. +by rewrite /normal normsY ?normal_norm // join_subG ?normal_sub. +Qed. +Hint Resolve Fcore_normal. + +Lemma Fcore_sub : M`_\F \subset M. +Proof. by case/andP: Fcore_normal. Qed. + +Lemma Fcore_sub_Fitting : M`_\F \subset 'F(M). +Proof. +rewrite gen_subG; apply/bigcupsP=> P /andP[/SylowP[p _ /and3P[_ pP _]] nsP]. +by rewrite Fitting_max // (pgroup_nil pP). +Qed. + +Lemma Fcore_nil : nilpotent M`_\F. +Proof. exact: nilpotentS Fcore_sub_Fitting (Fitting_nil M). Qed. + +Lemma Fcore_max pi H : + pi.-Hall(M) H -> H <| M -> nilpotent H -> H \subset M`_\F. +Proof. +move=> hallH nsHM nilH; have [sHM pi_H _] := and3P hallH. +rewrite -(nilpotent_Fitting nilH) FittingEgen genS //. +apply/bigcupsP=> [[p /= _] piHp]; rewrite (bigcup_max 'O_p(H)%G) //. +have sylHp := nilpotent_pcore_Hall p nilH. +have sylHp_M := subHall_Sylow hallH (pnatPpi pi_H piHp) sylHp. +by rewrite (p_Sylow sylHp_M) (char_normal_trans (pcore_char _ _)). +Qed. + +Lemma Fcore_dprod : \big[dprod/1]_(P | Sylow M (gval P) && (P <| M)) P = M`_\F. +Proof. +rewrite -[M`_\F]bigprodGE. +apply/eqP/bigdprodYP=> P /andP[/SylowP[p p_pr sylP] nsPM]. +have defP := normal_Hall_pcore sylP nsPM. +have /dprodP[_ _ cFpFp' tiFpFp'] := nilpotent_pcoreC p (Fitting_nil M). +have /dprodYP := dprodEY cFpFp' tiFpFp'; rewrite /= p_core_Fitting defP. +apply: subset_trans; rewrite bigprodGE gen_subG. +apply/bigcupsP=> Q => /andP[/andP[/SylowP[q _ sylQ] nsQM] neqQP]. +have defQ := normal_Hall_pcore sylQ nsQM; rewrite -defQ -p_core_Fitting. +apply: sub_pcore => q' /eqnP->; apply: contraNneq neqQP => eq_qp. +by rewrite -val_eqE /= -defP -defQ eq_qp. +Qed. + +Lemma Fcore_pcore_Sylow p : p \in \pi(M`_\F) -> p.-Sylow(M) 'O_p(M). +Proof. +rewrite /= -(bigdprod_card Fcore_dprod) mem_primes => /and3P[p_pr _]. +have not_p_dv_1: ~ p %| 1 by rewrite gtnNdvd ?prime_gt1. +elim/big_ind: _ => // [p1 p2 IH1 IH2|P /andP[/SylowP[q q_pr sylP] nsPM p_dv_P]]. + by rewrite Euclid_dvdM // => /orP[/IH1 | /IH2]. +have qP := pHall_pgroup sylP. +by rewrite (eqnP (pgroupP qP p p_pr p_dv_P)) (normal_Hall_pcore sylP). +Qed. + +Lemma p_core_Fcore p : p \in \pi(M`_\F) -> 'O_p(M`_\F) = 'O_p(M). +Proof. +move=> piMFp /=; rewrite -(pcore_setI_normal p Fcore_normal). +apply/setIidPl; rewrite sub_gen // (bigcup_max 'O_p(M)%G) //= pcore_normal. +by rewrite (p_Sylow (Fcore_pcore_Sylow piMFp)). +Qed. + +Lemma Fcore_Hall : \pi(M`_\F).-Hall(M) M`_\F. +Proof. +rewrite Hall_pi // /Hall Fcore_sub coprime_pi' ?cardG_gt0 //=. +apply/pnatP=> // p p_pr; apply: contraL => /= piMFp; rewrite -p'natE //. +rewrite -partn_eq1 // -(eqn_pmul2l (part_gt0 p #|M`_\F|)) muln1. +rewrite -partnM ?cardG_gt0 // Lagrange ?Fcore_sub //. +rewrite -(card_Hall (nilpotent_pcore_Hall p Fcore_nil)) /=. +by rewrite p_core_Fcore // (card_Hall (Fcore_pcore_Sylow piMFp)). +Qed. + +Lemma pcore_Fcore pi : {subset pi <= \pi(M`_\F)} -> 'O_pi(M`_\F) = 'O_pi(M). +Proof. +move=> s_pi_MF; rewrite -(pcore_setI_normal pi Fcore_normal). +apply/setIidPl; rewrite (sub_normal_Hall Fcore_Hall) ?pcore_sub //. +exact: sub_pgroup s_pi_MF (pcore_pgroup pi M). +Qed. + +Lemma Fcore_pcore_Hall pi : {subset pi <= \pi(M`_\F)} -> pi.-Hall(M) 'O_pi(M). +Proof. +move=> s_pi_MF; apply: (subHall_Hall Fcore_Hall s_pi_MF). +by rewrite /= -pcore_Fcore // (nilpotent_pcore_Hall pi Fcore_nil). +Qed. + +End FittingCore. + +Lemma morphim_Fcore : GFunctor.pcontinuous Fitting_core. +Proof. +move=> gT rT G D f; have nsGF_G := Fcore_normal G. +suffices hall_fGF: \pi(G`_\F).-Hall(f @* (D :&: G)) (f @* (D :&: G`_\F)). + rewrite !morphimIdom in hall_fGF. + by rewrite (Fcore_max hall_fGF) ?morphim_normal // morphim_nil ?Fcore_nil. +rewrite morphim_pHall ?subsetIl //= -{2}(setIidPr (Fcore_sub G)) setIA. +by rewrite !(setIC (D :&: G)) (setI_normal_Hall nsGF_G) ?subsetIr ?Fcore_Hall. +Qed. + +Canonical Structure Fcore_igFun := [igFun by Fcore_sub & morphim_Fcore]. +Canonical Structure Fcore_gFun := [gFun by morphim_Fcore]. +Canonical Structure Fcore_pgFun := [pgFun by morphim_Fcore]. + +Section MoreFittingCore. + +Variables (gT rT : finGroupType) (D : {group gT}) (f : {morphism D >-> rT}). +Implicit Types (M H : {group gT}) (R : {group rT}). + +Lemma Fcore_char M : M`_\F \char M. Proof. exact: gFchar. Qed. + +Lemma FcoreJ M x : (M :^ x)`_\F = M`_\F :^ x. +Proof. +rewrite -{1}(setTI M) -morphim_conj. +by rewrite -injmF ?injm_conj ?subsetT // morphim_conj setTI. +Qed. + +Lemma injm_Fcore M : 'injm f -> M \subset D -> f @* M`_\F = (f @* M)`_\F. +Proof. by move=> injf sMD; rewrite injmF. Qed. + +Lemma isom_Fcore M R : isom M R f -> M \subset D -> isom M`_\F R`_\F f. +Proof. by move=> isoMR sMD; apply: gFisom. Qed. + +Lemma isog_Fcore M R : M \isog R -> M`_\F \isog R`_\F. +Proof. by move=> isoMR; apply: gFisog. Qed. + +End MoreFittingCore. + +Section Section15. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types p q q_star r : nat. +Implicit Types x y z : gT. +Implicit Types A E H K L M Mstar N P Q Qstar R S T U V W X Y Z : {group gT}. + +Lemma Fcore_sub_Msigma M : M \in 'M -> M`_\F \subset M`_\sigma. +Proof. +move=> maxM; rewrite gen_subG. +apply/bigcupsP=> P /andP[/SylowP[p _ sylP] nsPM]; have [sPM pP _] := and3P sylP. +have [-> | ntP] := eqsVneq P 1; first exact: sub1G. +rewrite (sub_Hall_pcore (Msigma_Hall maxM)) // (pi_pgroup pP) //. +by apply/exists_inP; exists P; rewrite ?(mmax_normal maxM). +Qed. + +Lemma Fcore_eq_Msigma M : + M \in 'M -> reflect (M`_\F = M`_\sigma) (nilpotent M`_\sigma). +Proof. +move=> maxM; apply: (iffP idP) => [nilMs | <-]; last exact: Fcore_nil. +apply/eqP; rewrite eqEsubset Fcore_sub_Msigma //. +by rewrite (Fcore_max (Msigma_Hall maxM)) ?pcore_normal. +Qed. + +(* This is B & G, Lemma 15.1. *) +(* We have made all semidirect products explicits, and omitted the assertion *) +(* M`_\sigma \subset M^`(1), which is exactly covered by Msigma_der1. *) +(* Some refactoring is definitely needed here, to avoid the mindless cut *) +(* and paste of a large fragment of the proof of Lemma 12.12. *) +Lemma kappa_structure M U K (Ms := M`_\sigma) : + M \in 'M -> kappa_complement M U K -> + [/\ (*a*) [/\ (Ms ><| U) ><| K = M, cyclic K & abelian (M^`(1) / Ms)], + (*b*) K :!=: 1 -> Ms ><| U = M^`(1) /\ abelian U, + (*c*) forall X, X \subset U -> X :!=: 1 -> 'C_Ms(X) != 1 -> + [/\ 'M('C(X)) = [set M], cyclic X & \tau2(M).-group X], + (*d*) abelian <<\bigcup_(x in Ms^#) 'C_U[x]>> + & (*e*) U :!=: 1 -> exists U0, + [/\ gval U0 \subset U, exponent (gval U0) = exponent U + & [Frobenius Ms <*> U0 = Ms ><| U0]]]. +Proof. +move=> maxM complU; have [hallU hallK _] := complU. +have [hallE defM _ regUK cUU] := kappa_compl_context maxM complU. +have [[_ E _ defE]] := sdprodP defM. +have [nsUE sKE mulUK nUK tiUK] := sdprod_context defE. +rewrite defE -{1 2}mulUK mulgA => defMl /mulGsubP[nMsU nMsK] tiMsE. +have [/andP[sMsM nMsM] [sUE nUE]] := (pcore_normal _ M : Ms <| M, andP nsUE). +rewrite norm_joinEr // mulUK in hallE. +have [[sEM s'M_E _] [sUM sk'U _]] := (and3P hallE, and3P hallU). +have defMsU: Ms ><| U = Ms <*> U. + by apply: sdprodEY nMsU (trivgP _); rewrite -tiMsE -mulUK setIS ?mulG_subl. +have{defM} defM: Ms <*> U ><| K = M. + rewrite sdprodE ?normsY ?coprime_TIg //=; first by rewrite norm_joinEr. + rewrite -(sdprod_card defMsU) coprime_mull andbC regular_norm_coprime //=. + by rewrite (coprimegS sKE) ?(pnat_coprime (pcore_pgroup _ _)). +rewrite defMsU quotient_der //= -/Ms -{2}defMl -mulgA mulUK. +rewrite quotientMidl -quotient_der ?(subset_trans sEM) //. +rewrite quotient_abelian ?(der_mmax_compl_abelian maxM hallE) //. +set part_c := forall U, _; have c_holds: part_c. + move=> X sXU ntX nregMsX; have sXE := subset_trans sXU sUE. + have [x /setIP[Ms_x cXx] ntx] := trivgPn _ nregMsX. + have Ms1x: x \in Ms^# by apply/setD1P. + have piCx_hyp: {in X^#, forall x', x' \in ('C_M[x])^# /\ \sigma(M)^'.-elt x'}. + move=> x' /setD1P[ntx' Xx']; have Ex' := subsetP sXE x' Xx'. + rewrite 3!inE ntx' (subsetP sEM) ?(mem_p_elt s'M_E) //=. + by rewrite (subsetP _ _ Xx') ?sub_cent1. + have piCx x' X1x' := (* GG -- ssreflect evar generalization fails in trunk *) + let: conj c e := piCx_hyp x' X1x' in pi_of_cent_sigma maxM Ms1x c e. + have t2X: \tau2(M).-group X. + apply/pgroupP=> p p_pr /Cauchy[] // x' Xx' ox'. + have X1x': x' \in X^# by rewrite !inE Xx' -order_gt1 ox' prime_gt1. + have [[]|[]] := piCx _ X1x'; last by rewrite /p_elt ox' pnatE. + case/idPn; have:= mem_p_elt (pgroupS sXU sk'U) Xx'. + by rewrite /p_elt ox' !pnatE // => /norP[]. + suffices cycX: cyclic X. + split=> //; have [x' defX] := cyclicP cycX. + have X1x': x' \in X^# by rewrite !inE -cycle_eq1 -cycle_subG -defX ntX /=. + have [[kX _]|[_ _]] := piCx _ X1x'; last by rewrite defX cent_cycle. + rewrite -(setIid X) coprime_TIg ?eqxx // {2}defX in ntX. + rewrite (pnat_coprime t2X (sub_pgroup _ kX)) // => p kp. + by rewrite inE /= negb_and rank_kappa ?orbT. + have [E2 hallE2 sXE2] := Hall_superset (sigma_compl_sol hallE) sXE t2X. + rewrite abelian_rank1_cyclic; last first. + exact: abelianS sXE2 (tau2_compl_abelian maxM hallE hallE2). + have [p _ ->] := rank_witness X; rewrite leqNgt; apply: contra nregMsX => rpX. + have t2p: p \in \tau2(M) by rewrite (pnatPpi t2X) // -p_rank_gt0 ltnW. + rewrite -(setIidPr (subset_trans sXE sEM)) in rpX. + case/p_rank_geP: rpX => A; rewrite pnElemI -setIdE; case/setIdP=> Ep2A sAX. + rewrite -subG1; have [_ _ <- _ _] := tau2_context maxM t2p Ep2A. + by rewrite setIS ?centS. +have hallU_E: Hall E U := pHall_Hall (pHall_subl sUE sEM hallU). +have UtypeF := FTtypeF_complement maxM hallE hallU_E nsUE. +set k'U13 := ({in _, _}) in UtypeF. +have/UtypeF{UtypeF k'U13}UtypeF: k'U13. + move=> x /setD1P[]; rewrite -order_gt1 -pi_pdiv. + set p := pdiv _ => pi_x_p Ux t13x. + apply: contraNeq (pnatPpi (mem_p_elt sk'U Ux) pi_x_p) => nreg_x. + apply/orP; right; rewrite unlock /= inE /= (pnatPpi t13x) //=. + have sxM: <[x]> \subset M by rewrite cycle_subG (subsetP sUM). + move: pi_x_p; rewrite -p_rank_gt0 /= -(setIidPr sxM) => /p_rank_geP[P]. + rewrite pnElemI -setIdE => /setIdP[EpP sPx]; apply/exists_inP; exists P => //. + by rewrite (subG1_contra _ nreg_x) //= -cent_cycle setIS ?centS. +have [K1 | ntK] := altP (K :=P: 1). + rewrite {2}K1 cyclic1; rewrite K1 mulg1 in mulUK; rewrite -mulUK in hallE. + have ltM'M := sol_der1_proper (mmax_sol maxM) (subxx _) (mmax_neq1 maxM). + suffices /UtypeF[[A0 [_ abA0 genA0]] frobM]: U :!=: 1. + by split => //; apply: abelianS abA0; rewrite gen_subG; apply/bigcupsP. + apply: contraNneq (proper_subn ltM'M); rewrite -{1}defMl => ->. + by rewrite K1 !mulg1 Msigma_der1. +have PmaxM: M \in 'M_'P by rewrite inE maxM -(trivg_kappa maxM hallK) andbT. +have [_ _ [_ _ _ [cycZ _ _ _ _] [_ _ _ defM']]] := Ptype_embedding PmaxM hallK. +have{cycZ cUU} [cycK cUU] := (cyclicS (joing_subl _ _) cycZ, cUU ntK). +split=> // [_||/UtypeF[] //]; first split=> //. + apply/eqP; rewrite eq_sym eqEcard -(leq_pmul2r (cardG_gt0 K)). + have [nsMsU_M _ mulMsU _ _] := sdprod_context defM. + rewrite (sdprod_card defM) (sdprod_card defM') der1_min ?normal_norm //=. + by rewrite -(isog_abelian (sdprod_isog defM)) cyclic_abelian. +by apply: abelianS cUU; rewrite gen_subG -big_distrr subsetIl. +Qed. + +(* This is B & G, Theorem 15.2. *) +(* It is this theorem that implies that the non-functorial definition of *) +(* M`_\sigma used in B & G is equivalent to the original definition in FT *) +(* (also used in Peterfalvi). *) +(* Proof notes: this proof contained two non-structural arguments: taking D *) +(* to be K-invariant, and reusing the nilpotent Frobenius kernel argument for *) +(* Q1 (bottom of p. 118). We handled the first with a "without loss", but for *) +(* the second we had to spell out explicitly the assumptions and conclusions *) +(* of the nilpotent kernel argument that were spread throughout the last *) +(* paragraph p. 118. *) +(* We also had to make a few additions to the argument at the top of p. 119; *) +(* while the statement of the Theorem states that F(M) = C_M(Qbar), the text *) +(* only shows that F(M) = C_Msigma(Qbar), and we need to show that K acts *) +(* regularly on Qbar to complete the proof; this follows from the values of *) +(* orders of K, Kstar and Qbar. In addition we need to show much earlier *) +(* that K acts faithfully on Q, to show that C_M(Q) is included in Ms, and *) +(* this requires a use of 14.2(e) not mentioned in the text; in addition, the *) +(* reference to coprime action (Proposition 1.5) on p. 119 l. 1 is somewhat *) +(* misleading, since we actually need to use the coprime stabilizer Lemma 1.9 *) +(* to show that C_D(Qbar) = C_D(Q) = 1 (unless we splice in the proof of that *) +(* lemma). *) +Theorem Fcore_structure M (Ms := M`_\sigma) : + M \in 'M -> + [/\ M`_\F != 1, M`_\F \subset Ms, Ms \subset M^`(1) & M^`(1) \proper M] + /\ (forall K D : {group gT}, + \kappa(M).-Hall(M) K -> M`_\F != M`_\sigma -> + let p := #|K| in let Ks := 'C_Ms(K) in + let q := #|Ks| in let Q := 'O_q(M) in + let Q0 := 'C_Q(D) in let Qbar := Q / Q0 in + q^'.-Hall(M`_\sigma) D -> + [/\ (*a*) [/\ M \in 'M_'P1, Ms ><| K = M & Ms = M ^`(1)], + (*b*) [/\ prime p, prime q, q \in \pi(M`_\F) & q \in \beta(M)], + [/\ (*c*) q.-Sylow(M) Q, + (*d*) nilpotent D + & (*e*) Q0 <| M], + (*f*) [/\ minnormal Qbar (M / Q0), q.-abelem Qbar & #|Qbar| = (q ^ p)%N] + & (*g*) [/\ Ms^`(1) = M^`(2), + M^`(2) \subset 'F(M), + [/\ Q <*> 'C_M(Q) = 'F(M), + 'C_M(Qbar | 'Q) = 'F(M) + & 'C_Ms (Ks / Q0 | 'Q) = 'F(M)] + & 'F(M) \proper Ms]]). +Proof. +move=> maxM; set M' := M^`(1); set M'' := M^`(2). +have nsMsM: Ms <| M := pcore_normal _ M; have [sMsM nMsM] := andP nsMsM. +have solM := mmax_sol maxM; have solMs: solvable Ms := solvableS sMsM solM. +have sMF_Ms: M`_\F \subset Ms := Fcore_sub_Msigma maxM. +have ltM'M: M' \proper M by rewrite (sol_der1_proper solM) ?mmax_neq1. +have sMsM': Ms \subset M' := Msigma_der1 maxM. +have [-> | ltMF_Ms] := eqVproper sMF_Ms; first by rewrite eqxx Msigma_neq1. +set KDpart := (X in _ /\ X); suffices KD_holds: KDpart. + do 2!split=> //; have [K hallK] := Hall_exists \kappa(M) solM. + pose q := #|'C_(M`_\sigma)(K)|; have [D hallD] := Hall_exists q^' solMs. + have [_ [_ _ piMFq _] _ _ _] := KD_holds K D hallK (proper_neq ltMF_Ms) hallD. + by rewrite -rank_gt0 (leq_trans _ (p_rank_le_rank q _)) ?p_rank_gt0. +move=> {KDpart} K D hallK neMF_Ms p Ks q Q /= hallD. +have not_nilMs: ~~ nilpotent Ms by rewrite (sameP (Fcore_eq_Msigma maxM) eqP). +have P1maxM: M \in 'M_'P1; last have [PmaxM _] := setIdP P1maxM. + apply: contraR not_nilMs => notP1maxM; apply: notP1type_Msigma_nil. + by rewrite orbC inE notP1maxM inE maxM andbT orNb. +have ntK: K :!=: 1 by rewrite inE maxM andbT -(trivg_kappa maxM hallK) in PmaxM. +have [defMs defM]: Ms = M' /\ Ms ><| K = M. + have [U complU] := ex_kappa_compl maxM hallK. + have U1: U :=: 1 by apply/eqP; rewrite (trivg_kappa_compl maxM complU). + have [[defM _ _] [//| defM' _] _ _ _] := kappa_structure maxM complU. + by rewrite U1 sdprodg1 in defM defM'. +have [_ mulMsK nMsK _] := sdprodP defM; rewrite /= -/Ms in mulMsK nMsK. +have [sKM kK _] := and3P hallK; have s'K := sub_pgroup (@kappa_sigma' _ _) kK. +have coMsK: coprime #|Ms| #|K| := pnat_coprime (pcore_pgroup _ _) s'K. +have q_pr: prime q. + have [L _ [_ _ _ _ [_]]] := Ptype_embedding PmaxM hallK. + by rewrite inE P1maxM => [[] []]. +have hallMs: \sigma(M).-Hall(M) Ms := Msigma_Hall maxM. +have sMq: q \in \sigma(M). + by rewrite -pnatE // -pgroupE (pgroupS (subsetIl _ _) (pcore_pgroup _ _)). +have{s'K kK} q'K: q^'.-group K := pi'_p'group s'K sMq. +have nsQM: Q <| M := pcore_normal q M; have [sQM nQM] := andP nsQM. +have qQ: q.-group Q := pcore_pgroup _ _. +have sQMs: Q \subset Ms by rewrite (sub_Hall_pcore hallMs) ?(pi_pgroup qQ). +have [K1 prK1 sK1K]: exists2 K1, prime #|gval K1| & K1 \subset K. + have:= ntK; rewrite -rank_gt0; have [r r_pr ->] := rank_witness K. + by case/p_rank_geP=> K1 /pnElemPcard[? _ oK1]; exists K1; rewrite ?oK1. +have coMsK1 := coprimegS sK1K coMsK; have coQK1 := coprimeSg sQMs coMsK1. +have prMsK: semiprime Ms K by have [[? _ []] ] := Ptype_structure PmaxM hallK. +have defCMsK1: 'C_Ms(K1) = Ks. + by rewrite (cent_semiprime prMsK) // -cardG_gt1 prime_gt1. +have sK1M := subset_trans sK1K sKM; have nQK1 := subset_trans sK1M nQM. +have{sMsM'} sKsQ: Ks \subset Q. + have defMsK: [~: Ms, K] = Ms by case/coprime_der1_sdprod: defM. + have hallQ := nilpotent_pcore_Hall q (Fitting_nil M). + rewrite -[Q]p_core_Fitting (sub_Hall_pcore hallQ) //; first exact: pnat_id. + apply: prime_meetG => //; apply: contraNneq not_nilMs => tiKsFM. + suffices <-: 'F(Ms) = Ms by apply: Fitting_nil. + apply/eqP; rewrite eqEsubset Fitting_sub /= -{1}defMsK. + rewrite (odd_sdprod_primact_commg_sub_Fitting defM) ?mFT_odd //. + apply/trivgP; rewrite -tiKsFM setIAC setSI //= -/Ms subsetI Fitting_sub /=. + by rewrite Fitting_max ?Fitting_nil // (char_normal_trans (Fitting_char _)). +have nilMs_Q: nilpotent (Ms / Q). + have [nMsK1 tiQK1] := (subset_trans sK1K nMsK, coprime_TIg coQK1). + have prK1b: prime #|K1 / Q| by rewrite -(card_isog (quotient_isog _ _)). + have defMsK1: (Ms / Q) ><| (K1 / Q) = (Ms / Q) <*> (K1 / Q). + by rewrite sdprodEY ?quotient_norms // coprime_TIg ?coprime_morph. + apply: (prime_Frobenius_sol_kernel_nil defMsK1) => //. + by rewrite (solvableS _ (quotient_sol _ solM)) ?join_subG ?quotientS. + by rewrite -coprime_quotient_cent ?quotientS1 /= ?defCMsK1. +have defQ: 'O_q(Ms) = Q by rewrite -(setIidPl sQMs) pcore_setI_normal. +have sylQ: q.-Sylow(Ms) Q. + have nsQMs: Q <| Ms by rewrite -defQ pcore_normal. + rewrite -(pquotient_pHall qQ) // /= -/Q -{3}defQ. + by rewrite -(pquotient_pcore _ qQ) ?nilpotent_pcore_Hall. +have{sMq hallMs} sylQ_M := subHall_Sylow hallMs sMq sylQ. +have sQ_MF: Q \subset M`_\F. + by rewrite sub_gen ?(bigcup_max [group of Q]) ?(p_Sylow sylQ_M) ?pcore_normal. +have{sQ_MF} piMFq: q \in \pi(M`_\F). + by rewrite (piSg sQ_MF) // (piSg sKsQ) // pi_of_prime ?inE /=. +without loss nDK: D hallD / K \subset 'N(D). + have [E hallE nEK] := coprime_Hall_exists q^' nMsK coMsK solMs. + have [x Ms_x ->] := Hall_trans solMs hallD hallE. + set Q0 := 'C__(_)%G; rewrite -(isog_nil (conj_isog _ _)) -['C_Q(_)]/(gval Q0). + move/(_ E hallE nEK)=> IH; suffices ->: Q0 = [group of 'C_Q(E)] by []. + apply: group_inj => /=; have Mx: x \in M := subsetP (pcore_sub _ _) x Ms_x. + rewrite /= -/Q -{1}(normsP nQM x Mx) centJ -conjIg (normsP _ x Mx) //. + by case: IH => _ _ [_ _]; case/andP. +set Q0 := 'C_Q(D); set Qb := Q / Q0. +have defQD: Q ><| D = Ms by rewrite -defQ in sylQ *; apply/sdprod_Hall_pcoreP. +have [_ mulQD nQD tiQD] := sdprodP defQD; rewrite /= -/Q -/Ms in mulQD nQD tiQD. +have nilD: nilpotent D. + by rewrite (isog_nil (quotient_isog nQD tiQD)) /= -quotientMidl mulQD. +have [sDMs q'D _] := and3P hallD; have sDM := subset_trans sDMs sMsM. +have sDKM: D <*> K \subset M by rewrite join_subG sDM. +have q'DK: q^'.-group (D <*> K) by rewrite norm_joinEr // pgroupM q'D. +have{K1 sK1M sK1K coMsK1 coQK1 prK1 defCMsK1 nQK1 solMs} Qi_rec Qi: + Qi \in |/|_Q(D <*> K; q) -> Q0 \subset Qi -> Qi \proper Q -> + exists L, [/\ L \in |/|_Q(D <*> K; q), Qi <| L, minnormal (L / Qi) (M / Qi) + & ~~ ((Ks \subset L) ==> (Ks \subset Qi))]. +- case/setIdP=> /andP[sQiQ qQi] nQiDK sQ0i ltQiQ. + have ltQiN := nilpotent_proper_norm (pgroup_nil qQ) ltQiQ. + have [Lb minLb sLbQ]: {Lb | minnormal (gval Lb) (M / Qi) & Lb \subset Q / Qi}. + apply: mingroup_exists; rewrite quotient_norms //= andbT -quotientInorm. + by rewrite -subG1 quotient_sub1 ?subsetIr // proper_subn. + have [ntLb nLbM] := andP (mingroupp minLb). + have nsQiN: Qi <| 'N_M(Qi) by rewrite normal_subnorm (subset_trans sQiQ). + have: Lb <| 'N_M(Qi) / Qi. + by rewrite quotientInorm /normal (subset_trans sLbQ) ?quotientS. + case/(inv_quotientN nsQiN) => L defLb sQij /=; case/andP. + case/subsetIP=> sLM nQij nLN; exists L. + have{sLbQ} sLQ: L \subset Q by rewrite -(quotientSGK nQij sQiQ) -defLb. + rewrite inE /psubgroup /normal sLQ sQij nQij (pgroupS sLQ qQ) -defLb. + have nLDK: D <*> K \subset 'N(L) by apply: subset_trans nLN; exact/subsetIP. + have sLD_Ms: L <*> D \subset Ms by rewrite join_subG (subset_trans sLQ). + have coLD_K1: coprime #|L <*> D| #|K1| := coprimeSg sLD_Ms coMsK1. + have [[nQiD nQiK] [nLD nLK]] := (joing_subP nQiDK, joing_subP nLDK). + have [nQiK1 nLK1] := (subset_trans sK1K nQiK, subset_trans sK1K nLK). + split=> //; apply: contra ntLb => regLK. + have [sLLD sDLD] := joing_subP (subxx (L <*> D)). + suffices nilLDbar: nilpotent (L <*> D / Qi). + rewrite defLb -subG1 -(quotientS1 sQ0i) /= -/Q. + rewrite coprime_quotient_cent ?(pgroup_sol qQ) ?(pnat_coprime qQ) //=. + rewrite subsetI quotientS //= (sub_nilpotent_cent2 nilLDbar) ?quotientS //. + by rewrite coprime_morph ?(p'nat_coprime q'D (pgroupS sLQ qQ)). + have defLK1b: (L <*> D / Qi) ><| (K1 / Qi) = (L <*> D / Qi) <*> (K1 / Qi). + rewrite sdprodEY ?coprime_TIg ?quotient_norms //=. + by rewrite (subset_trans sK1K) // normsY. + by rewrite coprime_morph // (coprimeSg sLD_Ms). + have [sQiLD sLD_M] := (subset_trans sQij sLLD, subset_trans sLD_Ms sMsM). + have{regLK}: 'C_(L <*> D / Qi)(K1 / Qi) = 1. + rewrite -coprime_quotient_cent ?(subset_trans sK1K) ?(solvableS sLD_M) //=. + rewrite -(setIidPr sLD_Ms) setIAC defCMsK1 quotientS1 //= -/Ks joingC. + rewrite norm_joinEl // -(setIidPl sKsQ) -setIA -group_modr // tiQD mul1g. + have [-> | ntLKs] := eqVneq (Ks :&: L) 1; first exact: sub1G. + by rewrite subIset ?(implyP regLK) // prime_meetG. + apply: (prime_Frobenius_sol_kernel_nil defLK1b). + by apply: solvableS (quotient_sol _ solM); rewrite join_subG !quotientS. + by rewrite -(card_isog (quotient_isog _ _)) ?coprime_TIg // (coprimeSg sQiQ). +have ltQ0Q: Q0 \proper Q. + rewrite properEneq subsetIl andbT; apply: contraNneq not_nilMs => defQ0. + rewrite -dprodEsd // in defQD; last by rewrite centsC -defQ0 subsetIr. + by rewrite (dprod_nil defQD) (pgroup_nil qQ). +have [nQK coQK] := (subset_trans sKM nQM, pnat_coprime qQ q'K). +have solQ := pgroup_sol qQ. (* must come late: Coq diverges on solQ <> solMs *) +have [coDK coQD] := (coprimeSg sDMs coMsK, pnat_coprime qQ q'D). +have nQ0K: K \subset 'N(Q0) by rewrite normsI ?norms_cent. +have nQ0D: D \subset 'N(Q0) by rewrite cents_norm // centsC subsetIr. +have nQ0DK: D <*> K \subset 'N(Q0) by apply/joing_subP. +have [|Q1 [DKinvQ1 nsQ01 minQ1b nregQ1b]] := Qi_rec _ _ (subxx _) ltQ0Q. + by rewrite inE /psubgroup (pgroupS _ qQ) ?subsetIl. +have{Qi_rec nregQ1b DKinvQ1} [tiQ0Ks defQ1]: Q0 :&: Ks = 1 /\ Q1 :=: Q. + move: nregQ1b; rewrite negb_imply; case/andP=> sKsQ1 not_sKsQ0. + split=> //; first by rewrite setIC prime_TIg. + have [] := setIdP DKinvQ1; case/andP; case/eqVproper=> // ltQ1Q _ _. + have [Q2 [_ _ _]] := Qi_rec Q1 DKinvQ1 (normal_sub nsQ01) ltQ1Q. + by rewrite sKsQ1 implybT. +have [nsQ0Q minQb]: Q0 <| Q /\ minnormal Qb (M / Q0) by rewrite /Qb -defQ1. +have{Q1 defQ1 minQ1b nsQ01} abelQb: q.-abelem Qb. + have qQb: q.-group Qb := quotient_pgroup _ qQ; have solQb := pgroup_sol qQb. + by rewrite -is_abelem_pgroup // (minnormal_solvable_abelem minQb). +have [cQbQb [sQ0Q nQ0Q]] := (abelem_abelian abelQb, andP nsQ0Q). +have nQ0M: M \subset 'N(Q0) by rewrite -mulMsK -mulQD -mulgA !mul_subG. +have nsQ0M: Q0 <| M by rewrite /normal subIset ?sQM. +have sFM_QCQ: 'F(M) \subset Q <*> 'C_M(Q). + have [_ /= mulQQ' cQQ' _] := dprodP (nilpotent_pcoreC q (Fitting_nil M)). + rewrite -{3}mulQQ' p_core_Fitting cent_joinEr ?subsetIr //= -/Q in cQQ' *. + by rewrite mulgS // subsetI (subset_trans (pcore_sub _ _) (Fitting_sub M)). +have sQCQ_CMsQb: Q <*> 'C_M(Q) \subset 'C_Ms(Qb | 'Q). + rewrite join_subG !(subsetI _ Ms) sQMs /= !sub_astabQ nQ0Q /= -/Q0 -/Qb. + rewrite -abelianE cQbQb quotient_cents ?subsetIr //= andbC subIset ?nQ0M //=. + rewrite -(coprime_mulG_setI_norm mulMsK) ?norms_cent //= -/Ms. + suffices ->: 'C_K(Q) = 1 by rewrite mulg1 subsetIl. + have: ~~ (Q \subset Ks); last apply: contraNeq => ntCKQ. + have [_ _ _ [_]] := Ptype_structure PmaxM hallK. + by move/(_ q); rewrite pi_of_prime //; case/(_ (eqxx q) _ sylQ_M). + rewrite -[Ks](cent_semiprime prMsK _ ntCKQ) ?subsetIl //. + by rewrite subsetI sQMs centsC subsetIr. +have nCQb: M \subset 'N('C(Qb | 'Q)). + by rewrite (subset_trans _ (astab_norm _ _)) ?actsQ. +have defFM: 'C_Ms(Qb | 'Q) = 'F(M). + apply/eqP; rewrite eqEsubset andbC (subset_trans sFM_QCQ) //=. + rewrite Fitting_max //=; first by rewrite /normal subIset ?sMsM //= normsI. + rewrite -(coprime_mulG_setI_norm mulQD) ?(subset_trans sMsM) //= -/Q. + rewrite mulg_nil ?(nilpotentS (subsetIl _ _)) ?(pgroup_nil qQ) //= -/Q. + rewrite (centsS (subsetIl _ _)) //= -/Q. + have cDQ0: 'C_D(Qb | 'Q) \subset 'C(Q0) by rewrite subIset // centsC subsetIr. + rewrite (stable_factor_cent cDQ0) ?(coprimegS (subsetIl _ _) coQD) //. + by rewrite /stable_factor commGC -sub_astabQR ?subsetIr // subIset ?nQ0D. +have{sFM_QCQ sQCQ_CMsQb} ->: Q <*> 'C_M(Q) = 'F(M). + by apply/eqP; rewrite eqEsubset sFM_QCQ andbT -defFM. +have ltFM_Ms: 'F(M) \proper Ms. + rewrite properEneq -{2}defFM subsetIl andbT. + by apply: contraNneq not_nilMs => <-; apply: Fitting_nil. +have sQFM: Q \subset 'F(M) by rewrite -[Q]p_core_Fitting pcore_sub. +have not_cDQb: ~~ (D / Q0 \subset 'C(Qb)). + apply: contra (proper_subn ltFM_Ms) => cDQb; rewrite -mulQD mulG_subG sQFM /=. + by rewrite -defFM subsetI sDMs sub_astabQ nQ0D. +have [_ minQbP] := mingroupP minQb. +have regQbDb: 'C_Qb(D / Q0) = 1. + apply: contraNeq not_cDQb => ntCQDb; rewrite centsC; apply/setIidPl. + apply: minQbP (subsetIl _ _); rewrite ntCQDb /= -/Qb -(setIidPl cQbQb) -setIA. + by rewrite -centM -quotientMl //= mulQD normsI ?norms_cent ?quotient_norms. +have tiQ0 R: q^'.-group R -> Q0 :&: R = 1. + by move/(pnat_coprime (pgroupS sQ0Q qQ))/coprime_TIg. +have oKb: #|K / Q0| = p by rewrite -(card_isog (quotient_isog _ (tiQ0 _ _))). +have oKsb: #|Ks / Q0| = q. + by rewrite -(card_isog (quotient_isog _ tiQ0Ks)) ?(subset_trans sKsQ). +have regDK: 'C_D(K) = 1. + by rewrite -(setIidPl sDMs) -setIA setIC coprime_TIg ?(coprimeSg sKsQ). +have{tiQ0} frobDKb: [Frobenius D <*> K / Q0 = (D / Q0) ><| (K / Q0)]. + have ntDb: D / Q0 != 1 by apply: contraNneq not_cDQb => ->; apply: sub1G. + have ntKb: K / Q0 != 1 by rewrite -(isog_eq1 (quotient_isog _ (tiQ0 _ _))). + apply/Frobenius_semiregularP => // [|xb]. + by apply: quotient_coprime_sdprod; rewrite ?sdprodEY ?coprime_TIg. + have [f [_ ker_f _ im_f]] := restrmP (coset_morphism Q0) nQ0DK. + have{ker_f} inj_f: 'injm f by rewrite ker_f ker_coset setIC tiQ0. + rewrite /= /quotient -!im_f ?joing_subl ?joing_subr //. + rewrite 2!inE andbC => /andP[/morphimP[x DKx Kx ->{xb}]]. + rewrite morph_injm_eq1 // -injm_subcent1 ?joing_subl // => ntx. + rewrite -{3}(setIidPl sDMs) -setIA prMsK ?inE ?ntx //. + by rewrite setICA regDK setIg1 morphim1. +have defKsb: 'C_Qb(K / Q0) = Ks / Q0. + rewrite /Ks -mulQD coprime_cent_mulG // regDK mulg1. + by rewrite coprime_quotient_cent ?subsetIl. +have{frobDKb regQbDb} [p_pr oQb cQbD']: + [/\ prime p, #|Qb| = (q ^ p)%N & (D / Q0)^`(1) \subset 'C(Qb)]. +- have ntQb: Qb != 1 by rewrite -subG1 quotient_sub1 ?proper_subn. + have prQbK: semiprime Qb (K / Q0). + move=> xb; rewrite 2!inE andbC; case/andP; case/morphimP=> x nQ0x Kx -> ntx. + have{ntx} ntx: x != 1 by apply: contraNneq ntx => ->; rewrite morph1. + transitivity ('C_Q[x] / Q0); last first. + rewrite -(coprime_quotient_cent (subsetIl Q _) nQ0K coQK solQ) /= -/Q0. + by rewrite -/Q -(setIidPl sQMs) -!setIA prMsK // !inE ntx. + rewrite -!cent_cycle -quotient_cycle //; rewrite -!cycle_subG in Kx nQ0x. + by rewrite coprime_quotient_cent ?(coprimegS Kx). + have:= Frobenius_primact frobDKb _ _ _ ntQb _ prQbK regQbDb. + have [nQDK solDK] := (subset_trans sDKM nQM, solvableS sDKM solM). + rewrite !quotient_sol ?quotient_norms // coprime_morph ?(pnat_coprime qQ) //=. + rewrite -/Qb oKb defKsb prime_cyclic oKsb // subsetI der_sub /=. + by case=> //= -> -> ->. +have{cQbD'} sM''FM: M'' \subset 'F(M). + have nQMs := subset_trans sMsM nQM. + rewrite [M'']dergSn -/M' -defMs -(quotientSGK _ sQFM) ?comm_subG //. + rewrite (quotient_der 1) //= -/Ms -mulQD quotientMidl -quotient_der //= -/Q. + by rewrite quotientS // -defFM subsetI sub_astabQ !comm_subG ?quotient_der. +have sQ0Ms := subset_trans sQ0Q sQMs. +have ->: 'C_Ms(Ks / Q0 | 'Q) = 'F(M). + have sFMcKsb: 'F(M) \subset 'C_Ms(Ks / Q0 | 'Q). + by rewrite -defFM setIS ?astabS ?quotientS. + have nCMsKsbM: M \subset 'N('C_Ms(Ks / Q0 | 'Q)). + rewrite -{1}mulMsK mulG_subG sub_der1_norm ?subsetIl //= -/Ms; last first. + by rewrite {1}defMs (subset_trans sM''FM sFMcKsb). + rewrite normsI // (subset_trans _ (astab_norm _ _)) ?actsQ //. + by rewrite cents_norm // centsC subsetIr. + apply/eqP; rewrite eqEsubset sFMcKsb -defFM subsetI subsetIl. + rewrite sub_astabQ /= -/Q0 subIset ?(subset_trans sMsM) //= andbT centsC. + apply/setIidPl; apply: minQbP (subsetIl _ _). + rewrite andbC normsI ?norms_cent ?quotient_norms //= -/Qb. + have: Ks / Q0 != 1 by rewrite -cardG_gt1 oKsb prime_gt1. + apply: subG1_contra; rewrite (quotientGI _ sQ0Ms) quotient_astabQ /= -/Q0. + by rewrite subsetI quotientS // centsC subsetIr. +have ->: 'C_M(Qb | 'Q) = 'F(M). + apply/eqP; rewrite eqEsubset -{2}defFM setSI //= andbT. + rewrite -(coprime_mulG_setI_norm mulMsK) //= -defFM mulG_subG subxx /=. + rewrite subsetI subsetIr -(quotientSGK _ sQ0Ms) 1?subIset ?nQ0K //= -/Q0. + rewrite quotientIG; last by rewrite sub_astabQ normG trivg_quotient sub1G. + rewrite quotient_astabQ /= -/Q0 prime_TIg ?sub1G ?oKb //. + rewrite centsC -subsetIidl defKsb; apply: contra (@subset_leq_card _ _ _) _. + by rewrite -ltnNge oQb oKsb (ltn_exp2l 1) prime_gt1. +suffices ->: q \in \beta(M) by do 2!split=> //; last rewrite {1}defMs. +apply: contraR not_cDQb; rewrite negb_forall_in; case/exists_inP=> Q1 sylQ1. +rewrite negbK {Q1 sylQ1}(eq_Hall_pcore sylQ_M sylQ1) -/Q => nnQ. +have sD_DK': D \subset (D <*> K)^`(1). + rewrite -{1}(coprime_cent_prod nDK) ?nilpotent_sol // regDK mulg1. + by rewrite commgSS ?joing_subl ?joing_subr. +rewrite quotient_cents // (subset_trans sD_DK') //. +have nQDK := subset_trans sDKM nQM; have nCQDK := norms_cent nQDK. +rewrite der1_min // -(isog_abelian (second_isog nCQDK)) setIC /=. +rewrite -ker_conj_aut (isog_abelian (first_isog_loc _ _)) //=; set A := _ @* _. +have norm_q := normal_norm (pcore_normal q _). +rewrite {norm_q}(isog_abelian (quotient_isog (norm_q _ _) _)) /=; last first. + by rewrite coprime_TIg ?coprime_morphr //= (pnat_coprime (pcore_pgroup q _)). +have AutA: A \subset [Aut Q] := Aut_conj_aut _ _. +have [|//] := Aut_narrow qQ (mFT_odd _) nnQ _ AutA (morphim_odd _ (mFT_odd _)). +exact: morphim_sol (solvableS sDKM solM). +Qed. + +(* This is B & G, Corollary 15.3(a). *) +Corollary cent_Hall_sigma_sdprod M H pi : + M \in 'M -> pi.-Hall(M`_\sigma) H -> H :!=: 1 -> + exists X, [/\ gval X \subset M, cyclic X, \tau2(M).-group X + & 'C_(M`_\sigma)(H) ><| X = 'C_M(H)]. +Proof. +move=> maxM hallH ntH; have hallMs := Msigma_Hall maxM. +have nsMsM: M`_\sigma <| M := pcore_normal _ M; have [sMsM nMsM] := andP nsMsM. +have sMs := pHall_pgroup hallMs; have [sHMs piH _] := and3P hallH. +have k'CH: \kappa(M)^'.-group 'C_M(H). + apply/idPn; rewrite negb_and cardG_gt0 all_predC negbK => /hasP[p piCHp kMp]. + have PmaxM: M \in 'M_'P by apply/PtypeP; split; last exists p. + have [X]: exists X, X \in 'E_p^1('C_M(H)). + by apply/p_rank_geP; rewrite p_rank_gt0. + case/pnElemP; case/subsetIP=> sXM cHX abelX dimX; have [pX _] := andP abelX. + have [K hallK sXK] := Hall_superset (mmax_sol maxM) sXM (pi_pgroup pX kMp). + have E1X: X \in 'E^1(K) by apply/nElemP; exists p; apply/pnElemP. + have [q q_pr rqH] := rank_witness H; have [S sylS] := Sylow_exists q H. + have piSq: q \in \pi(S). + by rewrite -p_rank_gt0 (p_rank_Sylow sylS) -rqH rank_gt0. + have [_ [defNK defNX] _ [_ not_sylCP _] _] := Ptype_structure PmaxM hallK. + have{defNX} [defNX _] := defNX X E1X. + have [[_ nsKs] [_ mulKKs _ _]] := (dprod_normal2 defNK, dprodP defNK). + have s'K := sub_pgroup (@kappa_sigma' _ _) (pHall_pgroup hallK). + have [_ hallKs] := coprime_mulGp_Hall mulKKs s'K (pgroupS (subsetIl _ _) sMs). + have [sSH _] := andP sylS. + have sylS_Ms := subHall_Sylow hallH (pnatPpi piH (piSg sSH piSq)) sylS. + have [sSMs _] := andP sylS_Ms; have sS := pgroupS sSMs sMs. + have sylS_M := subHall_Sylow hallMs (pnatPpi sS piSq) sylS_Ms. + have sSKs: S \subset 'C_(M`_\sigma)(K). + rewrite (sub_normal_Hall hallKs) //= -defNX subsetI (pHall_sub sylS_M) /=. + by rewrite cents_norm // centsC (centsS sSH). + by have [_ /negP] := not_sylCP q (piSg sSKs piSq) S sylS_M. +have solCH := solvableS (subsetIl M 'C(H)) (mmax_sol maxM). +have [X hallX] := Hall_exists \sigma(M)^' solCH; exists X. +have nsCsH: 'C_(M`_\sigma)(H) <| 'C_M(H) by rewrite /normal setSI // normsIG. +have hallCs: \sigma(M).-Hall('C_M(H)) 'C_(M`_\sigma)(H). + by rewrite -(setIidPl sMsM) -setIA (setI_normal_Hall nsMsM) ?subsetIl. +rewrite (sdprod_normal_p'HallP nsCsH hallX hallCs). +have [-> | ntX] := eqsVneq X 1; first by rewrite sub1G cyclic1 pgroup1. +have [sXCH s'X _] := and3P hallX; have [sXM cHX] := subsetIP sXCH. +have sk'X: \sigma_kappa(M)^'.-group X. + apply/pgroupP=> p p_pr p_dv_X; apply/norP=> /=. + by split; [apply: (pgroupP s'X) | apply: (pgroupP (pgroupS sXCH k'CH))]. +have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). +have [U complU] := ex_kappa_compl maxM hallK; have [hallU _ _] := complU. +have [a Ma sXaU] := Hall_Jsub (mmax_sol maxM) hallU sXM sk'X. +have [_ _ cycX _ _] := kappa_structure maxM complU. +rewrite -(cyclicJ _ a) -(pgroupJ _ _ a); have [||//] := cycX _ sXaU. + by rewrite -(isog_eq1 (conj_isog X a)). +rewrite -(normsP nMsM a Ma) centJ -conjIg -(isog_eq1 (conj_isog _ a)). +by rewrite (subG1_contra _ ntH) // subsetI sHMs centsC. +Qed. + +(* This is B & G, Corollary 15.3(b). *) +Corollary sigma_Hall_tame M H pi x a : + M \in 'M -> pi.-Hall(M`_\sigma) H -> x \in H -> x ^ a \in H -> + exists2 b, b \in 'N_M(H) & x ^ a = x ^ b. +Proof. +move=> maxM hallH Hx Hxa; have [sHMs piH _] := and3P hallH. +have SMxM: M \in 'M_\sigma[x] by rewrite inE maxM cycle_subG (subsetP sHMs). +have SMxMa: (M :^ a^-1)%G \in 'M_\sigma[x]. + by rewrite inE mmaxJ maxM cycle_subG /= MsigmaJ mem_conjgV (subsetP sHMs). +have [-> | ntx] := eqVneq x 1; first by exists 1; rewrite ?group1 ?conj1g. +have ell1x: \ell_\sigma(x) == 1%N. + by apply/ell_sigma1P; split=> //; apply/set0Pn; exists M. +have ntH: H :!=: 1 by apply/trivgPn; exists x. +have [[transR _ _ _] _] := FT_signalizer_context ell1x. +have [c Rc defMc] := atransP2 transR SMxM SMxMa. +pose b := c * a; have def_xa: x ^ a = x ^ b. + by have [_ cxc] := setIP Rc; rewrite conjgM {3}/conjg -(cent1P cxc) mulKg. +have Mb: b \in M. + by rewrite -(norm_mmax maxM) inE conjsgM -(congr_group defMc) actKV. +have nsMsM: M`_\sigma <| M := pcore_normal _ _; have [sMsM _] := andP nsMsM. +have [nsHM | not_nsHM] := boolP (H <| M). + by exists b; rewrite // (mmax_normal maxM) ?setIid. +have neqMFs: M`_\F != M`_\sigma. + apply: contraNneq not_nsHM => eq_MF_Ms. + have nilMs: nilpotent M`_\sigma by apply/Fcore_eq_Msigma. + rewrite (eq_Hall_pcore (nilpotent_pcore_Hall _ nilMs) hallH). + exact: char_normal_trans (pcore_char _ _) nsMsM. +have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). +pose q := #|'C_(M`_\sigma)(K)|. +have [D hallD] := Hall_exists q^' (solvableS sMsM (mmax_sol maxM)). +have [[_ sMFs _ _]] := Fcore_structure maxM; case/(_ K D) => //; rewrite -/q. +set Q := 'O_q(M) => _ [_ q_pr piMFq _] [sylQ nilD _] _ _. +have sQMs: Q \subset M`_\sigma. + by apply: subset_trans sMFs; rewrite -[Q](p_core_Fcore piMFq) pcore_sub. +have sylQ_Ms: q.-Hall(M`_\sigma) Q := pHall_subl sQMs sMsM sylQ. +have nsQM: Q <| M := pcore_normal q M; have [_ qQ _] := and3P sylQ. +have nsQ_Ms: Q <| M`_\sigma := normalS sQMs sMsM nsQM. +have defMs: Q ><| D = M`_\sigma := sdprod_normal_p'HallP nsQ_Ms hallD sylQ_Ms. +have [_ mulQD nQD tiQD] := sdprodP defMs; rewrite -/Q in mulQD nQD tiQD. +have nQH := subset_trans sHMs (normal_norm nsQ_Ms). +have nsQHM: Q <*> H <| M. + rewrite -(quotientGK nsQM) -quotientYK // cosetpre_normal //= -/Q. + have nilMsQ: nilpotent (M`_\sigma / Q). + by rewrite -mulQD quotientMidl -(isog_nil (quotient_isog _ _)). + have hallMsQpi := nilpotent_pcore_Hall pi nilMsQ. + rewrite (eq_Hall_pcore hallMsQpi (quotient_pHall nQH hallH)). + by rewrite (char_normal_trans (pcore_char _ _)) ?quotient_normal. +have tiQH: Q :&: H = 1. + apply: coprime_TIg (p'nat_coprime (pi_pgroup qQ _) piH). + apply: contra not_nsHM => pi_q; rewrite (joing_idPr _) // in nsQHM. + by rewrite (normal_sub_max_pgroup (Hall_max hallH)) ?(pi_pgroup qQ). +have defM: Q * 'N_M(H) = M. + have hallH_QH: pi.-Hall(Q <*> H) H. + by rewrite (pHall_subl (joing_subr _ _) _ hallH) // join_subG sQMs. + have solQH := solvableS (normal_sub nsQHM) (mmax_sol maxM). + rewrite -{2}(Hall_Frattini_arg solQH nsQHM hallH_QH) /= norm_joinEr //. + by rewrite -mulgA [H * _]mulSGid // subsetI (subset_trans sHMs sMsM) normG. +move: Mb; rewrite -{1}defM; case/mulsgP=> z n Qz Nn defb; exists n => //. +rewrite def_xa defb conjgM [x ^ z](conjg_fixP _) // -in_set1 -set1gE -tiQH. +rewrite inE {1}commgEr groupMr // -mem_conjgV groupV /=. +rewrite (normsP (normal_norm nsQM)) ?Qz; last first. + by rewrite groupV (subsetP sMsM) // (subsetP sHMs). +rewrite commgEl groupMl ?groupV //= -(conjsgK n H) mem_conjgV -conjgM -defb. +by have [_ nHn] := setIP Nn; rewrite (normP nHn) -def_xa. +Qed. + +(* This is B & G, Corollary 15.4. *) +(* This result does not appear to be needed for the actual proof. *) +Corollary nilpotent_Hall_sigma H : + nilpotent H -> Hall G H -> exists2 M, M \in 'M & H \subset M`_\sigma. +Proof. +move=> nilH /Hall_pi/=hallH; have [_ piH _] := and3P hallH. +have [-> | ntH] := eqsVneq H 1. + by have [M maxM] := any_mmax gT; exists M; rewrite ?sub1G. +pose p := pdiv #|H|; have piHp: p \in \pi(H) by rewrite pi_pdiv cardG_gt1. +pose S := 'O_p(H)%G; have sylS: p.-Sylow(H) S := nilpotent_pcore_Hall p nilH. +have [sSH pS _] := and3P sylS. +have ntS: S :!=: 1 by rewrite -rank_gt0 (rank_Sylow sylS) p_rank_gt0. +have [M maxNM] := mmax_exists (mFT_norm_proper ntS (mFT_pgroup_proper pS)). +have [maxM sNM] := setIdP maxNM; exists M => //. +have sSM: S \subset M := subset_trans (normG _) sNM. +have sylS_G := subHall_Sylow hallH (pnatPpi piH piHp) sylS. +have sylS_M := pHall_subl sSM (subsetT M) sylS_G. +have sMp: p \in \sigma(M) by apply/exists_inP; exists S. +have sSMs: S \subset M`_\sigma. + by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) ?(pi_pgroup pS). +rewrite -(nilpotent_Fitting nilH) FittingEgen gen_subG. +apply/bigcupsP=> [[q /= _] piHq]; have [-> // | p'q] := eqVneq q p. +have sylS_Ms := pHall_subl sSMs (pcore_sub _ _) sylS_M. +have [X [_ cycX t2X defCS]] := cent_Hall_sigma_sdprod maxM sylS_Ms ntS. +have{defCS} [nCMsCS _ defCS _ _] := sdprod_context defCS. +have t2'CMs: \tau2(M)^'.-group 'C_(M`_\sigma)(S). + apply: pgroupS (subsetIl _ _) (sub_pgroup _ (pcore_pgroup _ _)) => r. + by apply: contraL => /andP[]. +have [hallCMs hallX] := coprime_mulGp_Hall defCS t2'CMs t2X. +have sHqCS: 'O_q(H) \subset 'C_M(S). + rewrite (setIidPr (subset_trans (cent_sub _) sNM)). + rewrite (sub_nilpotent_cent2 nilH) ?pcore_sub //. + exact: pnat_coprime pS (pi_pgroup (pcore_pgroup _ _) _). +have [t2q | t2'q] := boolP (q \in \tau2(M)); last first. + apply: subset_trans (subsetIl _ 'C(S)). + by rewrite (sub_normal_Hall hallCMs) // (pi_pgroup (pcore_pgroup _ _)). +have sylHq := nilpotent_pcore_Hall q nilH. +have sylHq_G := subHall_Sylow hallH (pnatPpi piH piHq) sylHq. +have sylHq_CS := pHall_subl sHqCS (subsetT _) sylHq_G. +have [Q sylQ] := Sylow_exists q X; have [sQX _] := andP sylQ. +have sylQ_CS := subHall_Sylow hallX t2q sylQ. +case/andP: t2q => _. +rewrite eqn_leq andbC ltnNge (leq_trans (p_rankS q (subsetT _))) //. +rewrite -(rank_Sylow sylHq_G) (rank_Sylow sylHq_CS) -(rank_Sylow sylQ_CS). +by rewrite (leq_trans (rankS sQX)) // -abelian_rank1_cyclic ?cyclic_abelian. +Qed. + +(* This is B & G, Corollary 15.5. *) +(* We have changed the condition K != 1 to the equivalent M \in 'M_'P, as *) +(* avoids a spurrious quantification on K. *) +(* The text is quite misleading here, since Corollary 15.3(a) is needed for *) +(* bith the nilpotent and the non-nilpotent case -- indeed, it is not really *) +(* needed in the non-nilpotent case! *) +Corollary Fitting_structure M (H := M`_\F) (Y := 'O_\sigma(M)^'('F(M))) : + M \in 'M -> + [/\ (*a*) cyclic Y /\ \tau2(M).-group Y, + (*b*) [/\ M^`(2) \subset 'F(M), + H \* 'C_M(H) = 'F(M) + & 'F(M`_\sigma) \x Y = 'F(M)], + (*c*) H \subset M^`(1) /\ nilpotent (M^`(1) / H) + & (*d*) M \in 'M_'P -> 'F(M) \subset M^`(1)]. +Proof. +move=> maxM; have nilF := Fitting_nil M. +have sHF: H \subset 'F(M) := Fcore_sub_Fitting M. +have nsMsM: M`_\sigma <| M := pcore_normal _ _; have [sMsM nMsM] := andP nsMsM. +have sMs: \sigma(M).-group M`_\sigma := pcore_pgroup _ _. +have nsFM: 'F(M) <| M := Fitting_normal M; have [sFM nFM] := andP nsFM. +have sYF: Y \subset 'F(M) := pcore_sub _ _; have sYM := subset_trans sYF sFM. +have defF: 'F(M`_\sigma) \x Y = 'F(M). + rewrite -(nilpotent_pcoreC \sigma(M) nilF); congr (_ \x _). + apply/eqP; rewrite eqEsubset pcore_max ?(pgroupS (Fitting_sub _)) //=. + rewrite Fitting_max ?(nilpotentS (pcore_sub _ _)) //=. + by rewrite -(pcore_setI_normal _ nsFM) norm_normalI ?(subset_trans sMsM). + rewrite /normal (char_norm_trans (Fitting_char _)) ?(subset_trans sFM) //. + by rewrite Fitting_max ?Fitting_nil // (char_normal_trans (Fitting_char _)). +have [[ntH sHMs sMsM' _] nnil_struct] := Fcore_structure maxM. +have hallH: \pi(H).-Hall(M`_\sigma) H := pHall_subl sHMs sMsM (Fcore_Hall M). +have [X [_ cycX t2X defCH]] := cent_Hall_sigma_sdprod maxM hallH ntH. +have hallX: \sigma(M)^'.-Hall('C_M(H)) X. + have [_ mulCsH_X _ _] := sdprodP defCH. + have [|//] := coprime_mulpG_Hall mulCsH_X (pgroupS (subsetIl _ _) sMs). + by apply: sub_pgroup t2X => p /andP[]. +have sYX: Y \subset X. + rewrite (normal_sub_max_pgroup (Hall_max hallX)) ?pcore_pgroup //. + rewrite /normal (char_norm_trans (pcore_char _ _)) ?subIset ?nFM //= -/Y. + have [_ _ cFsY _] := dprodP defF. + rewrite subsetI sYM (sub_nilpotent_cent2 nilF) //= -/Y -/H. + exact: pnat_coprime (pgroupS sHMs sMs) (pcore_pgroup _ _). +have [cycY t2Y]: cyclic Y /\ \tau2(M).-group Y. + by rewrite (cyclicS sYX cycX) (pgroupS sYX t2X). +have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). +have [U complU] := ex_kappa_compl maxM hallK. +have [[defM _ cM'M'b] defM' _ _ _] := kappa_structure maxM complU. +have d_holds: M \in 'M_'P -> 'F(M) \subset M^`(1). + rewrite inE maxM andbT -(trivg_kappa maxM hallK) => ntK. + rewrite -(dprodW defF) mulG_subG (subset_trans (Fitting_sub _)) //= -/Y. + have{defM'} [[defM' _] nsM'M] := (defM' ntK, der_normal 1 M). + have hallM': \kappa(M)^'.-Hall(M) M^`(1). + by apply/(sdprod_normal_pHallP nsM'M hallK); rewrite /= -defM'. + rewrite (sub_normal_Hall hallM') ?(sub_pgroup _ t2Y) // => p. + by case/andP=> _; apply: contraL => /rank_kappa->. +have defF_H: 'C_M(H) \subset 'F(M) -> H \* 'C_M(H) = 'F(M). + move=> sCHF; apply/eqP; rewrite cprodE ?subsetIr // eqEsubset ?mul_subG //=. + have hallH_F := pHall_subl sHF sFM (Fcore_Hall M). + have nsHF := normalS sHF sFM (Fcore_normal M). + have /dprodP[_] := nilpotent_pcoreC \pi(H) nilF. + rewrite (normal_Hall_pcore hallH_F nsHF) /= -/H => defF_H cHFH' _. + by rewrite -defF_H mulgS // subsetI (subset_trans (pcore_sub _ _)). +have [eqHMs | neqHMs] := eqVneq H M`_\sigma. + split=> //; [split=> // | by rewrite eqHMs abelian_nil]. + by rewrite (subset_trans _ sHF) //= eqHMs der1_min ?comm_subG. + rewrite defF_H // -(sdprodW defCH) -eqHMs mulG_subG subIset ?sHF //=. + rewrite Fitting_max ?abelian_nil ?cyclic_abelian //. + rewrite -(normal_Hall_pcore hallX) ?(char_normal_trans (pcore_char _ _)) //. + by rewrite norm_normalI // eqHMs norms_cent. + move: defCH; rewrite -dprodEsd; first by case/dprod_normal2. + by rewrite -eqHMs (centsS (subsetIl _ _)); case/subsetIP: (pHall_sub hallX). +pose q := #|'C_(M`_\sigma)(K)|; pose Q := 'O_q(M). +have [D hallD] := Hall_exists q^' (solvableS sMsM (mmax_sol maxM)). +case/(_ K D): nnil_struct => //=; rewrite -/H -/q -/Q. +move=> [_ _ defMs] [_ _ piHq _] [sylQ nilD _] _ [_ -> [defF_Q _ _] _]. +have sQH: Q \subset H by rewrite -[Q](p_core_Fcore piHq) pcore_sub. +split=> //; rewrite -?{}defMs; split=> //. + by rewrite defF_H // -defF_Q joingC sub_gen // subsetU ?setIS ?centS. +have sQMs := subset_trans sQH sHMs; have sylQ_Ms := pHall_subl sQMs sMsM sylQ. +have nsQ_Ms: Q <| M`_\sigma := normalS sQMs sMsM (pcore_normal _ _). +have defMs: Q ><| D = M`_\sigma := sdprod_normal_p'HallP nsQ_Ms hallD sylQ_Ms. +have [_ <- _ _] := sdprodP defMs; rewrite -quotientMidl mulgA (mulGSid sQH). +by rewrite quotientMidl quotient_nil. +Qed. + +(* This is B & G, Corollary 15.6. *) +(* Note that the proof of the F-core noncyclicity given in the text only *) +(* applies to the nilpotent case, and we need to use a different assertion of *) +(* 15.2 if Msigma is not nilpotent. *) +Corollary Ptype_cyclics M K (Ks := 'C_(M`_\sigma)(K)) : + M \in 'M_'P -> \kappa(M).-Hall(M) K -> + [/\ Ks != 1, cyclic Ks, Ks \subset M^`(2), Ks \subset M`_\F + & ~~ cyclic M`_\F]. +Proof. +move=> PmaxM hallK; have [ntK maxM] := setIdP PmaxM. +rewrite -(trivg_kappa maxM hallK) in ntK. +have [_ _ [ntKs _] _ _] := Ptype_structure PmaxM hallK. +have [_ _ [_ _ _ [cycZ _ _ _ _] [_ _ _ defM]]] := Ptype_embedding PmaxM hallK. +have{cycZ} cycKs: cyclic Ks := cyclicS (joing_subr _ _) cycZ. +have solM': solvable M^`(1) := solvableS (der_sub 1 M) (mmax_sol maxM). +have sMsM' := Msigma_der1 maxM. +have{defM} sKsM'': Ks \subset M^`(2). + have coM'K: coprime #|M^`(1)| #|K|. + by rewrite (coprime_sdprod_Hall_r defM) (pHall_Hall hallK). + have [_] := coprime_der1_sdprod defM coM'K solM' (subxx _). + exact: subset_trans (setSI _ sMsM'). +have [eqMFs | neqMFs] := eqVneq M`_\F M`_\sigma. + split=> //; rewrite eqMFs ?subsetIl //; apply: contra ntKs => cycMs. + rewrite -subG1 (subset_trans sKsM'') // (sameP trivgP derG1P) /= -derg1. + have cycF: cyclic 'F(M). + have [[cycY _] [_ _ defF] _ _] := Fitting_structure maxM. + have [[x defMs] [y defY]] := (cyclicP cycMs, cyclicP cycY). + rewrite defMs (nilpotent_Fitting (abelian_nil (cycle_abelian _))) in defF. + have [_ mulXY cxy _] := dprodP defF. + rewrite -mulXY defY -cycleM ?cycle_cyclic //. + by apply/cent1P; rewrite -cycle_subG sub_cent1 -cycle_subG -defY. + by rewrite /order -defMs -defY coprime_pcoreC. + apply: abelianS (cyclic_abelian cycF). + apply: subset_trans (cent_sub_Fitting (mmax_sol maxM)). + rewrite der1_min ?normsI ?normG ?norms_cent ?gFnorm //=. + rewrite -ker_conj_aut (isog_abelian (first_isog_loc _ _)) ?gFnorm //=. + exact: abelianS (Aut_conj_aut _ _) (Aut_cyclic_abelian cycF). +have [D hallD] := Hall_exists #|Ks|^' (solvableS sMsM' solM'). +have [_ /(_ K D)[]//=] := Fcore_structure maxM; rewrite -/Ks. +set q := #|Ks|; set Q := 'O_q(M) => _ [_ q_pr piMFq bMq] [sylQ _ _] _ _. +have sQMF: Q \subset M`_\F by rewrite -[Q]p_core_Fcore ?pcore_sub. +have qKs: q.-group Ks := pnat_id q_pr. +have sKsM := subset_trans sKsM'' (der_sub 2 M). +split=> //; first by apply: subset_trans sQMF; rewrite (sub_Hall_pcore sylQ). +apply: contraL (beta_sub_alpha bMq) => /(cyclicS sQMF)cycQ; rewrite -leqNgt. +by rewrite leqW // -(rank_Sylow sylQ) -abelian_rank1_cyclic ?cyclic_abelian. +Qed. + +(* This is B & G, Theorem 15.7. *) +(* We had to change the statement of the Theorem, because the first equality *) +(* of part (c) does not appear to be valid: if M is of type F, we know very *) +(* little of the action E1 on the Sylow subgroups of E2, and so E2 might have *) +(* a Sylow subgroup that meets F(M) but is also centralised by E1 and hence *) +(* intesects M' trivially; luckily, only the inclusion M' \subset F(M) seems *) +(* to be needed in the sequel. *) +(* We have also restricted the quantification on the Ei to part (c), and *) +(* factored and simplified some of the assertions of parts (e2) and (e3); it *) +(* appears they could perhaps be simplified futher, since the assertions on *) +(* Op(H) and Op'(H) do not appear to be used in the Peterfalvi revision of *) +(* the character theory part of the proof. *) +(* Proof notes: we had to correct/complete several arguments of the B & G *) +(* text. The use of 12.6(d) for parts (c) and (d), p. 120, l. 5 from the *) +(* bottom, is inapropriate as tau2 could be empty. The assertion X1 != Z0 on *) +(* l. 5, p. 121 needs to be strengthened to ~~ (X1 \subset Z0) in order to *) +(* prove that #|Z0| is prime -- only then are the assertions equivalent. The *) +(* use of Lemma 10.13(b), l. 7, p. 121, requires that B be maximal in G, not *) +(* only in P as is stated l. 6; proving the stronger assertion requires using *) +(* Corollary 15.3(b), which the text does not mention. The regularity *) +(* property stated l. 11-13 is too weak to be used in the type P1 case -- the *) +(* regularity assumption need to be restricted to a subgroup of prime order. *) +(* Finally, the cyclicity of Z(H) is not actually needed in the proof. *) +Theorem nonTI_Fitting_structure M g (H := (M`_\F)%G) : + let X := ('F(M) :&: 'F(M) :^ g)%G in + M \in 'M -> g \notin M -> X :!=: 1 -> + [/\ (*a*) M \in 'M_'F :|: 'M_'P1 /\ H :=: M`_\sigma, + (*b*) X \subset H /\ cyclic X, + (*c*) M^`(1) \subset 'F(M) /\ M`_\sigma \x 'O_\sigma(M)^'('F(M)) = 'F(M), + (*d*) forall E E1 E2 E3, sigma_complement M E E1 E2 E3 -> + [/\ E3 :=: 1, E2 <| E, E / E2 \isog E1 & cyclic (E / E2)] + & (*e*) (*1*) [/\ M \in 'M_'F, abelian H & 'r(H) = 2] + \/ let p := #|X| in [/\ prime p, ~~ abelian 'O_p(H), cyclic 'O_p^'(H) + & (*2*) {in \pi(H), forall q, exponent (M / H) %| q.-1} + \/ (*3*) [/\ #|'O_p(H)| = (p ^ 3)%N, M \in 'M_'P1 & #|M / H| %| p.+1] + ]]. +Proof. +move=> X maxM notMg ntX; have nilH: nilpotent H := Fcore_nil M. +have /andP[sHM nHM]: H <| M := Fcore_normal M. +have [[cycY t2Y] [_ _ defF] _ _] := Fitting_structure maxM. +set Y := 'O_\sigma(M)^'('F(M)) in cycY t2Y defF *. +have not_cycMp: {in \pi(X), forall p, ~~ cyclic 'O_p(M)}. + move=> p; rewrite mem_primes => /and3P[p_pr _ p_dv_X]. + apply: contra notMg => cycMp. + have hallMp := nilpotent_pcore_Hall p (Fitting_nil M). + have{cycMp} defNx1: {in 'F(M), forall x1, #[x1] = p -> 'N(<[x1]>) = M}. + move=> x1; rewrite /order -cycle_subG => sX1F oX1. + rewrite (mmax_normal maxM) //; last by rewrite -cardG_gt1 oX1 prime_gt1. + rewrite (char_normal_trans _ (pcore_normal p M)) ?sub_cyclic_char //=. + by rewrite -p_core_Fitting (sub_Hall_pcore hallMp) // /pgroup oX1 pnat_id. + have [x1 Xx1 ox1] := Cauchy p_pr p_dv_X; have [Fx1 Fgx1] := setIP Xx1. + rewrite -(norm_mmax maxM) inE -{1}(defNx1 (x1 ^ g^-1)) -?mem_conjg ?orderJ //. + by rewrite cycleJ normJ actKV -(defNx1 x1). +have{cycY} sX: \sigma(M).-group X. + apply: sub_pgroup (pgroup_pi X) => p piXp. + apply: contraR (not_cycMp p piXp) => s'p; rewrite -p_core_Fitting. + by apply: cyclicS (sub_pcore _ _) cycY => p1; move/eqnP->. +have sXMs: X \subset M`_\sigma. + by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) // subIset ?Fitting_sub. +have E1X_facts p X1 (C1 := 'C_H(X1)%G): + X1 \in 'E_p^1(X) -> [/\ C1 \notin 'U, 'r(C1) <= 2 & abelian C1]. +- move=> EpX1; have [sX1X /andP[pX1 _] _] := pnElemP EpX1. + have piXp: p \in \pi(X) by rewrite -p_rank_gt0; apply/p_rank_geP; exists X1. + have not_sCX1M: ~~ ('C(X1) \subset M). + have [[sX1F sX1Fg] sFM] := (subsetIP sX1X, Fitting_sub M). + apply: contra notMg => sCX1M; rewrite -groupV. + have [trCX1 _ _] := sigma_group_trans maxM (pnatPpi sX piXp) pX1. + have [||c cX1c] := trCX1 g^-1; rewrite ?(subset_trans _ sFM) ?sub_conjgV //. + by case=> m Mm ->; rewrite groupM // (subsetP sCX1M). + have ltCX1_G: 'C(X1) \proper G := mFT_cent_proper (nt_pnElem EpX1 isT). + have ltC1G: C1 \proper G := sub_proper_trans (subsetIr H _) ltCX1_G. + have{ltCX1_G} nonuniqC1: C1 \notin 'U. + have sC1M: C1 \subset M by rewrite subIset ?Fcore_sub. + apply: contra not_sCX1M => uniqC1. + by rewrite (sub_uniq_mmax (def_uniq_mmax uniqC1 maxM sC1M)) ?subsetIr. + split=> //; first by rewrite leqNgt (contra (rank3_Uniqueness _)). + have sC1H: C1 \subset H := subsetIl _ _. + have dprodC1: 'F(C1) = C1 := nilpotent_Fitting (nilpotentS sC1H nilH). + apply/center_idP; rewrite -{2}dprodC1 -(center_bigdprod dprodC1). + apply: eq_bigr => r _; apply/center_idP; apply: contraR nonuniqC1. + move/(nonabelian_Uniqueness (pcore_pgroup _ _)). + exact: uniq_mmaxS (pcore_sub _ _) ltC1G. +pose p := pdiv #|X|; have piXp: p \in \pi(X) by rewrite pi_pdiv cardG_gt1. +have /p_rank_geP[X1 EpX1]: 0 < 'r_p(X) by rewrite p_rank_gt0. +have [sMp ntX1] := (pnatPpi sX piXp, nt_pnElem EpX1 isT). +have [p_pr oX1] := (pnElem_prime EpX1, card_pnElem EpX1 : #|X1| = p). +have [sX1X abelX1 dimX1] := pnElemP EpX1; have [pX1 _] := andP abelX1. +have [nonuniqC1 rC1 cC1C1] := E1X_facts p X1 EpX1. +have [cycX b'p]: cyclic X /\ p \in \beta(M)^'. + have [E hallE] := ex_sigma_compl maxM. + have [_ _] := sigma_compl_embedding maxM hallE. + case/(_ g notMg); set X2 := _ :&: _ => cycX2 b'X2 _. + have sXMg: X \subset M :^ g by rewrite subIset // conjSg Fitting_sub orbT. + have{sXMg} sXX2: X \subset X2 by rewrite subsetI sXMs. + by rewrite (pnatPpi b'X2 (piSg sXX2 piXp)) (cyclicS sXX2). +have b'H: \beta(M)^'.-group H. + apply: sub_pgroup (pgroup_pi _) => r piHr; have [-> // | p'r] := eqVneq r p. + apply/existsP; exists 'O_r(M)%G; rewrite /= Fcore_pcore_Sylow // negbK. + apply/implyP; rewrite ltnNge -rank_pgroup ?pcore_pgroup ?(leq_trans _ rC1) //. + rewrite rankS // subsetI /= -{1}(p_core_Fcore piHr) pcore_sub //. + rewrite -p_core_Fitting (sub_nilpotent_cent2 (Fitting_nil M)) ?pcore_sub //. + exact: subset_trans sX1X (subsetIl _ _). + exact: pnat_coprime pX1 (pi_pgroup (pcore_pgroup r _) _). +have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). +have [sKM kK _] := and3P hallK; have s'K := sub_pgroup (@kappa_sigma' _ M) kK. +have [U complU] := ex_kappa_compl maxM hallK. +have [[defM cycK _] defM' _ _ exU0] := kappa_structure maxM complU. +have{b'p} FP1maxM: M \in 'M_'F :|: 'M_'P1. + apply: contraR b'p; rewrite inE /=; case/norP=> notFmaxF notP1maxF. + have PmaxM: M \in 'M_'P by apply/setDP. + by have [_ _ _ _ [| <- //]] := Ptype_structure PmaxM hallK; apply/setDP. +have defH: H :=: M`_\sigma. + apply/eqP; apply/idPn=> neqHMs; pose q := #|'C_(M`_\sigma)(K)|. + have solMs: solvable M`_\sigma := solvableS (pcore_sub _ _) (mmax_sol maxM). + have [D hallD] := Hall_exists q^' solMs. + have [_ /(_ K D)[] // _ [_ _ piHq /idPn[]]] := Fcore_structure maxM. + exact: pnatPpi b'H piHq. +have{sXMs} sXH: X \subset H by rewrite defH. +have{b'H} sM'F: M^`(1) \subset 'F(M). + rewrite Fitting_max ?der_normal // (isog_nil (quotient1_isog _)). + suffices <-: M`_\beta = 1 by apply: Mbeta_quo_nil. + apply/eqP; rewrite trivg_card1 (card_Hall (Mbeta_Hall maxM)). + rewrite -(partn_part _ (beta_sub_sigma maxM)) -(card_Hall (Msigma_Hall maxM)). + by rewrite /= -defH partG_eq1. +have{defF} defF: M`_\sigma \x Y = 'F(M). + by rewrite -defF -defH nilpotent_Fitting. +split=> // [E E1 E2 E3 complEi | {Y t2Y defF sM'F}]. + have [[sE3E' _] _ [cycE1 _] [_ defE] _]:= sigma_compl_context maxM complEi. + have [hallE _ _ hallE3 _] := complEi; have [sE3E t3E3 _] := and3P hallE3. + have E3_1: E3 :=: 1. + have [sEM s'E _] := and3P hallE; have sE'M' := dergS 1 sEM. + have sE3F: E3 \subset 'F(M) := subset_trans sE3E' (subset_trans sE'M' sM'F). + rewrite -(setIidPr sE3F) coprime_TIg // -(dprod_card defF) coprime_mull. + rewrite (pnat_coprime (pcore_pgroup _ _) (pgroupS sE3E s'E)). + exact: p'nat_coprime (sub_pgroup (@tau3'2 _ M) t2Y) t3E3. + have{defE} defE: E2 ><| E1 = E by rewrite -defE E3_1 sdprod1g. + have [-> _ mulE21 nE21 tiE21] := sdprod_context defE. + by rewrite -mulE21 quotientMidl quotient_cyclic // isog_sym quotient_isog. +have{defM'} defM_P1: M \in 'M_'P1 -> H ><| K = M /\ M^`(1) = H. + move=> P1maxM; have [PmaxM _] := setIdP P1maxM. + have U1: U :=: 1 by apply/eqP; rewrite (trivg_kappa_compl maxM complU). + have ntK: K :!=: 1 by rewrite (trivg_kappa maxM hallK); case/setDP: PmaxM. + by have [<- _] := defM' ntK; rewrite -{1}defM U1 sdprodg1 -defH. +pose P := 'O_p(H); have sylP: p.-Sylow(H) P := nilpotent_pcore_Hall p nilH. +have [sPH pP _] := and3P sylP. +have [cHH | {not_cycMp} not_cHH] := boolP (abelian H); [left | right]. + have [-> | P1maxM] := setUP FP1maxM; last first. + have [PmaxM _] := setIdP P1maxM. + have [ntKs _ sKsM'' _ _] := Ptype_cyclics PmaxM hallK. + case/eqP: (subG1_contra sKsM'' ntKs); apply/derG1P. + by rewrite /= -derg1; have [_ ->] := defM_P1 P1maxM. + split=> //; apply/eqP; rewrite eqn_leq (leq_trans _ rC1) //=; last first. + by rewrite rankS // subsetIidl (centsS sX1X) // (sub_abelian_cent cHH). + apply: leq_trans (rankS (pcore_sub p _)). + rewrite ltnNge -abelian_rank1_cyclic ?(abelianS sPH) //=. + by rewrite p_core_Fcore ?(piSg sXH) ?not_cycMp. +have sX1P: X1 \subset P by rewrite (sub_Hall_pcore sylP) ?(subset_trans sX1X). +have [_ mulPHp' cPHp' _] := dprodP (nilpotent_pcoreC p nilH : P \x _ = H). +have cHp'Hp': abelian 'O_p^'(H). + by rewrite (abelianS _ cC1C1) // subsetI pcore_sub (centsS sX1P). +have not_cPP: ~~ abelian P. + by apply: contra not_cHH => cPP; rewrite -mulPHp' abelianM cPP cHp'Hp'. +have{E1X_facts} pX: p.-group X. + apply: sub_pgroup (pgroup_pi _) => q; rewrite -p_rank_gt0. + case/p_rank_geP=> X2 EqX2; have [_ _ cC2C2] := E1X_facts _ _ EqX2. + case/pnElemP: EqX2 => sX2X /andP[qX2 _] _; have sX2H := subset_trans sX2X sXH. + apply: contraR not_cPP => q'p; rewrite (abelianS _ cC2C2) // subsetI sPH. + by rewrite (sub_nilpotent_cent2 nilH) ?(p'nat_coprime (pi_pgroup qX2 _) pP). +have sXP: X \subset P by rewrite (sub_Hall_pcore sylP). +pose Z0 := 'Ohm_1('Z(P)); have sZ0ZP: Z0 \subset 'Z(P) := Ohm_sub 1 _. +have{sZ0ZP} [sZ0P cPZ0] := subsetIP sZ0ZP. +have not_sX1Z0: ~~ (X1 \subset Z0). + apply: contra not_cPP => sX1Z0; apply: abelianS cC1C1. + by rewrite subsetI sPH (centsS sX1Z0) // centsC. +pose B := X1 <*> Z0; have sBP: B \subset P by rewrite join_subG sX1P. +have defB: X1 \x Z0 = B by rewrite dprodEY ?prime_TIg ?oX1 ?(centsS sX1P). +have pZP: p.-group 'Z(P) := pgroupS (center_sub _) pP. +have abelZ0: p.-abelem Z0 by rewrite Ohm1_abelem ?center_abelian. +have{abelZ0} abelB: p.-abelem B by rewrite (dprod_abelem _ defB) abelX1. +have sylP_Ms: p.-Sylow(M`_\sigma) P by rewrite -defH. +have sylP_G: p.-Sylow(G) P := subHall_Sylow (Msigma_Hall_G maxM) sMp sylP_Ms. +have max_rB A: p.-abelem A -> B \subset A -> 'r_p(A) <= 2. + move=> abelA /joing_subP[sX1A _]; have [pA cAA _] := and3P abelA. + suffices [a [nX1a sAaP]]: exists a, a \in 'N(X1) /\ A :^ a \subset P. + rewrite -rank_pgroup // -(rankJ _ a) (leq_trans _ rC1) ?rankS //= subsetI. + by rewrite -(normP nX1a) centJ conjSg (subset_trans sAaP) ?(centsS sX1A). + have [a _ sAaP] := Sylow_Jsub sylP_G (subsetT A) pA. + have [x1 defX1]: exists x1, X1 :=: <[x1]>. + by apply/cyclicP; rewrite prime_cyclic ?oX1. + have Px1: x1 \in P by rewrite -cycle_subG -defX1. + have Px1a: x1 ^ a \in P. + by rewrite (subsetP sAaP) // memJ_conjg -cycle_subG -defX1. + have [b nPb def_xb] := sigma_Hall_tame maxM sylP_Ms Px1 Px1a. + exists (a * b^-1); rewrite !inE !actM !sub_conjgV defX1 /= -!cycleJ def_xb. + by have{nPb} [_ nPb] := setIP nPb; rewrite (normP nPb). +have rpB: 'r_p(B) = 2. + apply/eqP; rewrite eqn_leq max_rB // -(p_rank_dprod p defB). + rewrite p_rank_abelem ?dimX1 // ltnS p_rank_Ohm1 -rank_pgroup // rank_gt0. + by rewrite center_nil_eq1 ?(pgroup_nil pP) ?(subG1_contra sXP). +have oZ0: #|Z0| = p. + apply/eqP; rewrite -(eqn_pmul2l (cardG_gt0 X1)) (dprod_card defB) oX1. + by rewrite (card_pgroup (abelem_pgroup abelB)) -p_rank_abelem ?rpB. +have p2maxElemB: [group of B] \in 'E_p^2(G) :&: 'E*_p(G). + rewrite !inE subsetT abelB -p_rank_abelem // rpB /=. + apply/maxgroupP; rewrite !inE subsetT /= -/B; split=> // A. + case/pElemP=> _ abelA sBA; have [pA _] := andP abelA. + apply/eqP; rewrite eq_sym eqEcard sBA (card_pgroup pA). + rewrite (card_pgroup (abelem_pgroup abelB)) -!p_rank_abelem // rpB. + by rewrite leq_exp2l ?prime_gt1 ?max_rB. +have{not_sX1Z0} defX: X :=: X1. + have sX_CPB: X \subset 'C_P(B). + rewrite centY !subsetI sXP sub_abelian_cent ?cyclic_abelian //=. + by rewrite centsC (centsS sXP). + have [C defCPB]: exists C, X1 \x C = 'C_P(B). + have [_ [C]] := basic_p2maxElem_structure p2maxElemB pP sBP not_cPP. + case=> _ _ defCPB _; exists C; rewrite defCPB // !inE joing_subl abelX1. + by rewrite -val_eqE eqEsubset negb_and not_sX1Z0 /= dimX1. + have defX: X1 \x (C :&: X) = X by rewrite (dprod_modl defCPB) // (setIidPr _). + by move/eqP: ntX1; case/(cyclic_pgroup_dprod_trivg pX cycX): defX; case. +have cycHp': cyclic 'O_p^'(H). + rewrite abelian_rank1_cyclic // leqNgt; apply: contra nonuniqC1 => rHp'. + apply: (uniq_mmaxS (setIS H (centS sX1P))). + by rewrite mFT_sol_proper nilpotent_sol // (nilpotentS (subsetIl _ _)). + apply: cent_uniq_Uniqueness (subsetIr _ _) (leq_trans rHp' (rankS _)). + exact: nonabelian_Uniqueness pP not_cPP. + by rewrite subsetI pcore_sub. +rewrite {1}defX oX1 /= -[M`_\F]/(gval H) -/P; split=> //. +pose Z q := 'Ohm_1('Z('O_q(H)))%G. +have charZ q: Z q \char H. + have:= char_trans (center_char _) (pcore_char q H). + exact: char_trans (Ohm_char 1 _). +have{cycHp'} oZ: {in \pi(H), forall q, #|Z q| = q}. + move=> q piHp; have [-> // | p'q] := eqVneq q p. + have qHq: q.-group 'O_q(H) := pcore_pgroup q H. + have sHqHp': 'O_q(H) \subset 'O_p^'(H) by apply: sub_pcore => r /eqnP->. + rewrite /= (center_idP (abelianS sHqHp' cHp'Hp')). + apply: Ohm1_cyclic_pgroup_prime (cyclicS sHqHp' cycHp') qHq _. + by rewrite -rank_gt0 (rank_Sylow (nilpotent_pcore_Hall q nilH)) p_rank_gt0. +have regZq_dv_q1 A q: + A \subset 'N(H) -> q \in \pi(H) -> semiregular (Z q) A -> #|A| %| q.-1. +- move=> nHA piHq regA. + by rewrite -(oZ q piHq) regular_norm_dvd_pred // (char_norm_trans (charZ q)). +have [FmaxM | {U complU defM exU0}P1maxM] := setUP FP1maxM. + left=> q piHq; have K1: K :=: 1 by apply/eqP; rewrite (trivg_kappa maxM). + have ntU: U :!=: 1 by rewrite (trivg_kappa_compl maxM complU) 2!inE FmaxM. + rewrite K1 sdprodg1 -defH in defM; have [_ mulHU nHU tiHU] := sdprodP defM. + rewrite -mulHU quotientMidl -(exponent_isog (quotient_isog nHU tiHU)). + have [U0 [sU0U <- frobMsU0]] := exU0 ntU; have nHU0 := subset_trans sU0U nHU. + apply: dvdn_trans (exponent_dvdn U0) _; apply: regZq_dv_q1 => // x U0x. + apply/trivgP; rewrite -(Frobenius_reg_ker frobMsU0 U0x) setSI //= -defH. + exact: (char_sub (charZ _)). +have{defM_P1} [[defM defM'] [PmaxM _]] := (defM_P1 P1maxM, setIdP P1maxM). +have [_ mulHK nHK tiHK] := sdprodP defM; have p'K := pi'_p'group s'K sMp. +have coHK: coprime #|H| #|K| by rewrite defH (pnat_coprime (pcore_pgroup _ _)). +have{coHK} coPK: coprime #|P| #|K| := coprimeSg sPH coHK. +have oMH: #|M / H| = #|K|. + by rewrite -mulHK quotientMidl -(card_isog (quotient_isog nHK tiHK)). +pose Ks := 'C_H(K); have prKs: prime #|Ks|. + have [Ms _ [_ _ _ _ [_]]] := Ptype_embedding PmaxM hallK. + by rewrite inE P1maxM -defH; do 2!case. +have sKsP: Ks \subset P. + have sKsM'': Ks \subset M^`(2) by rewrite /Ks defH; case/Ptype_cyclics: hallK. + rewrite (subset_trans sKsM'') 1?der1_min //= -derg1 defM' ?gFnorm //. + by rewrite -mulPHp' quotientMidl quotient_abelian. +have oKs: #|Ks| = p. + apply/eqP; apply: pnatPpi pP (piSg sKsP _). + by rewrite mem_primes prKs cardG_gt0 dvdnn. +have [prHK ntKs]: semiprime H K /\ Ks != 1. + by rewrite /Ks defH; case/Ptype_structure: hallK => // [[_ _ [_ ? _]] _ []]. +have [K_dv_p1 | {regZq_dv_q1}] := altP (@implyP (Ks :==: Z0) (#|K| %| p.-1)). + left=> q piHq; rewrite (dvdn_trans (exponent_dvdn _)) // oMH. + have [eqZqKs | neqZqKs] := eqVneq Ks (Z q). + have def_q: q = p by rewrite -(oZ q piHq) -eqZqKs. + by rewrite def_q K_dv_p1 // eqZqKs def_q. + apply: regZq_dv_q1 => // x Kx; rewrite -(setIidPl (char_sub (charZ q))). + rewrite -setIA prHK {x Kx}// setIC (prime_TIg prKs) //. + have q_pr: prime q by rewrite mem_primes in piHq; case/and3P: piHq. + apply: contra neqZqKs => sKsZq; rewrite eqEsubset sKsZq /=. + by rewrite prime_meetG ?oZ // (setIidPr sKsZq). +rewrite {Z oZ charZ}negb_imply; case/andP; move/eqP=> defKs not_K_dv_p1. +have nPK: K \subset 'N(P) := char_norm_trans (pcore_char p H) nHK. +have defZP: 'Z(P) = Ks. + apply/eqP; rewrite eqEsubset andbC {1}defKs Ohm_sub subsetI subIset ?sPH //. + have nZPK: K \subset 'N('Z(P)) := char_norm_trans (center_char _) nPK. + have coZPK: coprime #|'Z(P)| #|K| := coprimeSg (center_sub _) coPK. + rewrite centsC (coprime_odd_faithful_Ohm1 pZP) ?mFT_odd //. + by rewrite /= -/Z0 -defKs centsC subsetIr. +have rPle2: 'r(P) <= 2. + case/setIP: p2maxElemB; case/pnElemP=> _ _ dimB pmaxB. + have Ep2B: [group of B] \in 'E_p^2(P) by apply/pnElemP. + rewrite leqNgt; apply: contra not_K_dv_p1 => rPgt2. + have tiKcP: 'C_K(P) = 1. + apply/trivgP/subsetP=> x => /setIP[Kx cPx]. + apply: contraR ntX1 => ntx; rewrite -subG1. + have [_ _ _ <-] := dprodP defB; rewrite subsetIidl -defKs. + rewrite -[Ks](prHK x) 1?inE ?ntx // (subset_trans sX1P) //=. + by rewrite subsetI sPH sub_cent1. + rewrite (card_isog (quotient1_isog _)) -tiKcP -ker_conj_aut. + rewrite (card_isog (first_isog_loc _ nPK)) /=. + set A := _ @* _; have AutA: A \subset Aut P := Aut_conj_aut _ _. + have solA: solvable A by rewrite morphim_sol ?abelian_sol ?cyclic_abelian. + have oddA: odd #|A| by rewrite morphim_odd ?mFT_odd. + have nnP: p.-narrow P. + apply/implyP=> _; apply/set0Pn; exists [group of B]. + by rewrite inE Ep2B (subsetP (pmaxElemS p (subsetT P))) // inE pmaxB inE. + have [x defK] := cyclicP cycK; have Kx: x \in K by rewrite defK cycle_id. + have nPx := subsetP nPK x Kx; rewrite /A defK morphim_cycle //. + have Axj: conj_aut [group of P] x \in A by exact: mem_morphim. + have [_ _ -> //] := Aut_narrow pP (mFT_odd _) nnP solA AutA oddA. + by rewrite morph_p_elt // (mem_p_elt p'K). +have{rPle2} dimP: logn p #|P| = 3. + have [S [_ <- _] [C cycC]] := mFT_rank2_Sylow_cprod sylP_G rPle2 not_cPP. + case=> defP defZS; congr (logn p #|(_ : {set _})|). + suffices defC: 'Ohm_1(C) = C by rewrite -defC defZS cprod_center_id in defP. + suffices <-: 'Z(P) = C by rewrite defZP (@Ohm1_id _ p) // prime_abelem. + have [_ <- _] := cprodP (center_cprod defP). + by rewrite -defZS (center_idP (cyclic_abelian cycC)) mulSGid ?Ohm_sub. +have oP: #|P| = (p ^ 1.*2.+1)%N by rewrite (card_pgroup pP) dimP. +right; split; rewrite // {}oMH. +have esP: extraspecial P by apply: (p3group_extraspecial pP); rewrite ?dimP. +have defPK: P ><| K = P <*> K by rewrite sdprodEY // coprime_TIg. +have copK: coprime p #|K| by rewrite -oX1 (coprimeSg sX1P). +have [x|] := repr_extraspecial_prime_sdprod_cycle pP esP defPK cycK oP copK. + move/prHK=> defCHx /=; rewrite /= -/P -{1}(setIidPl sPH) -setIA defCHx -/Ks. + by rewrite -defZP setIA setIid. +by rewrite addn1 subn1 (negPf not_K_dv_p1) orbF. +Qed. + +(* A subset of the above, that does not require the non-TI witness. *) +Lemma nonTI_Fitting_facts M : + M \in 'M -> ~~ normedTI 'F(M)^# G M -> + [/\ M \in 'M_'F :|: 'M_'P1, M`_\F = M`_\sigma & M^`(1) \subset 'F(M)]. +Proof. +move=> maxM nonTI; suff: [exists (y | y \notin M), 'F(M) :&: 'F(M) :^ y != 1]. + by case/exists_inP=> y notMy /nonTI_Fitting_structure[] // [[-> ?] _ []]. +rewrite -negb_forall_in; apply: contra nonTI => /forall_inP tiF. +apply/normedTI_P; rewrite normD1 setTI gFnorm setD_eq0 subG1. +split=> // [|g _]; first by rewrite (trivg_Fitting (mmax_sol maxM)) mmax_neq1. +by apply: contraR => /tiF; rewrite -setI_eq0 conjD1g -setDIl setD_eq0 subG1. +Qed. + +(* This is B & G, Theorem 15.8, due to Feit and Thompson (1991). *) +(* We handle the non-structural step on l. 5, p. 122 by choosing A not to be *) +(* a q-group, if possible, so that when it turns out to be we know q is the *) +(* only prime in tau2(H). Since this uniqueness does not appear to be used *) +(* later we could also eliminate this complication. *) +(* We also avoid the circularity in proving that A is a q-group and that Q *) +(* is non-abelian by deriving that Q is in U from 14.2(e) rather than 12.13. *) +Theorem tau2_P2type_signalizer M Mstar U K r R H (q := #|K|) : + M \in 'M_'P2 -> kappa_complement M U K -> Mstar \in 'M('C(K)) -> + r.-Sylow(U) R -> H \in 'M('N(R)) -> ~~ \tau2(H)^'.-group H -> + [/\ prime q, \tau2(H) =i (q : nat_pred) & \tau2(M)^'.-group M]. +Proof. +move: Mstar => L P2maxM complU maxCK_L sylR maxNR_H not_t2'H. +have [[PmaxM notP1maxM] [hallU hallK _]] := (setDP P2maxM, complU). +have q_pr: prime q by have [_ _ _ _ []] := Ptype_structure PmaxM hallK. +have [[maxH _] [maxM _]] := (setIdP maxNR_H, setDP PmaxM). +have [maxL sCKL] := setIdP maxCK_L; have hallLs := Msigma_Hall maxL. +have [_ sUHs] := P2type_signalizer P2maxM complU maxCK_L sylR maxNR_H. +set D := H :&: L => defUK [_ sKFD hallD] {r R sylR maxNR_H}. +set uniq_q := _ =i _. +have{not_t2'H} [q1 t2Hq neq_q]: exists2 q1, q1 \in \tau2(H) & q1 = q -> uniq_q. + move: not_t2'H; rewrite negb_and cardG_gt0 all_predC negbK /= has_filter. + set s := filter _ _ => nts. + have mem_s: s =i \tau2(H). + move=> q1; rewrite mem_filter; apply: andb_idr => /= t2q1. + by rewrite (partition_pi_mmax maxH) t2q1 !orbT. + have [all_q | ] := altP (@allP _ (pred1 q) s); last first. + by case/allPn=> q1; rewrite mem_s=> t2q1; move/eqnP=> ne_q1q; exists q1. + have s_q1: head q s \in s by case: (s) nts => // q1 s' _; exact: predU1l. + exists (head q s) => [|def_q q1]; rewrite -mem_s //. + by apply/idP/idP; [exact: all_q | move/eqnP->; rewrite -def_q]. +have [A /= Eq2A Eq2A_H] := ex_tau2Elem hallD t2Hq; rewrite -/D in Eq2A. +have [b'q qmaxA]: q1 \notin \beta(G) /\ A \in 'E*_q1(G). + by have [-> ->] := tau2_not_beta maxH t2Hq. +have [sDH s'HD _] := and3P hallD. +have [sAH abelA dimA] := pnElemP Eq2A_H; have [qA _] := andP abelA. +have [[nsAD _] _ _ _] := tau2_compl_context maxH hallD t2Hq Eq2A. +have cKA: A \subset 'C(K). + have sFD: 'F(D) \subset D := Fitting_sub _; have sFH := subset_trans sFD sDH. + have cFF: abelian 'F(D). + exact: sigma'_nil_abelian maxH sFH (pgroupS sFD s'HD) (Fitting_nil _). + exact: sub_abelian_cent2 cFF (Fitting_max nsAD (pgroup_nil qA)) sKFD. +have sAL: A \subset L := subset_trans cKA sCKL. +pose Ks := 'C_(M`_\sigma)(K). +have [PmaxL hallKs defK]: + [/\ L \in 'M_'P, \kappa(L).-Hall(L) Ks & 'C_(L`_\sigma)(Ks) = K]. +- have [L1 [? _] [defL1 [? _] [? _] _ _]] := Ptype_embedding PmaxM hallK. + suffices ->: L = L1 by []; apply/set1P; rewrite defL1 // in maxCK_L. + by apply/nElemP; exists q; rewrite p1ElemE // !inE subxx eqxx. +have sKLs: K \subset L`_\sigma by rewrite -defK subsetIl. +have sLq: q \in \sigma(L). + by rewrite -pnatE // -pgroupE (pgroupS sKLs) ?pcore_pgroup. +have sLq1: q1 \in \sigma(L). + apply: contraLR sLq => s'Lq1; rewrite -negnK negbK /= -pnatE // -pgroupE. + have s'LA: \sigma(L)^'.-group A by exact: pi_pgroup qA _. + have [E hallE sAE] := Hall_superset (mmax_sol maxL) sAL s'LA. + have EqA_E: A \in 'E_q1^2(E) by exact/pnElemP. + have [[sEL s'E _] t2Lq1] := (and3P hallE, sigma'2Elem_tau2 maxL hallE EqA_E). + have [_ [sCAE _ _] _ _] := tau2_compl_context maxL hallE t2Lq1 EqA_E. + by apply: pgroupS (subset_trans _ sCAE) s'E; rewrite centsC. +have sALs: A \subset L`_\sigma by rewrite sub_Hall_pcore ?(pi_pgroup qA). +have solL: solvable L`_\sigma := solvableS (pcore_sub _ _) (mmax_sol maxL). +pose Q := 'O_q(L)%G; have{solL} [Ds hallDs] := Hall_exists q^' solL. +have sQFL: Q \subset 'F(L) by rewrite -[gval Q]p_core_Fitting pcore_sub. +have [sAFL sylQ]: A \subset 'F(L) /\ q.-Sylow(L) Q. + have [defLF | neqLF] := eqVneq L`_\F L`_\sigma. + split; first by rewrite (subset_trans sALs) // -defLF Fcore_sub_Fitting. + by rewrite Fcore_pcore_Sylow // defLF mem_primes q_pr cardG_gt0 cardSg. + have [_ /(_ _ Ds hallKs neqLF)] := Fcore_structure maxL. + rewrite /= defK -/q -/Q; case=> // _ _ [-> _ nsQ0L] _ [_ _ [_ _ <-] _]. + rewrite subsetI sALs sub_astabQ quotient_cents // (subset_trans sAL) //. + exact: normal_norm nsQ0L. +have{sLq1} neqHL: H :!=: L. + by apply: contraTneq t2Hq => ->; rewrite negb_and negbK /= sLq1. +have def_q1: q1 = q. + have uniqQ: Q \in 'U. + have [_ _ _ [_ uniqQ _] _] := Ptype_structure PmaxL hallKs. + apply/uniq_mmaxP; exists L; case/uniqQ: sylQ => //=; rewrite defK. + by rewrite pi_of_prime ?inE. + apply: contraNeq neqHL => q'q1. + rewrite (eq_uniq_mmax (def_uniq_mmax _ maxL sAL) maxH sAH) //. + rewrite (cent_uniq_Uniqueness uniqQ) ?(rank_abelem abelA) ?dimA //. + rewrite (sub_nilpotent_cent2 (Fitting_nil L)) //. + exact: pnat_coprime (pcore_pgroup _ _) (pi_pgroup qA _). +split=> //; first exact: neq_q. +rewrite {q1 neq_q}def_q1 in qA Eq2A Eq2A_H t2Hq abelA dimA qmaxA b'q. +have{b'q} b'q: q \notin \beta(L) by rewrite -predI_sigma_beta // inE /= sLq. +have P1maxL: L \in 'M_'P1. + apply: contraR b'q => notP1maxL. + by have [_ _ _ _ [|<- //]] := Ptype_structure PmaxL hallKs; apply/setDP. +have nilLs: nilpotent L`_\sigma. + rewrite (sameP (Fcore_eq_Msigma maxL) eqP); apply: contraR b'q => neqLF. + have [_ /(_ _ Ds hallKs neqLF)] := Fcore_structure maxL. + by rewrite /= defK -/q -/Q; case=> // _ [_ _ _ ->] _ _ _. +have defL': L^`(1) = L`_\sigma. + have [Us complUs] := ex_kappa_compl maxL hallKs. + have [_ [|<- _ _ _ _]] := kappa_structure maxL complUs. + by rewrite (trivg_kappa maxL hallKs) //; case/setDP: PmaxL. + suffices ->: Us :=: 1 by rewrite sdprodg1. + by apply/eqP; rewrite (trivg_kappa_compl maxL complUs). +have [ntK sKLs']: K :!=: 1 /\ K \subset L`_\sigma^`(1). + by rewrite -defL' -defK; case/Ptype_cyclics: hallKs. +have [sQL qQ _] := and3P sylQ. +have not_cQQ: ~~ abelian Q. + have sKL: K \subset L := subset_trans sKLs (pcore_sub _ _). + have sKQ: K \subset Q by rewrite (sub_Hall_pcore sylQ) /pgroup ?pnat_id. + have sQLs: Q \subset L`_\sigma by rewrite sub_Hall_pcore ?(pi_pgroup qQ). + have defLs: 'O_q^'(L`_\sigma) * Q = L`_\sigma. + rewrite -(setIidPl sQLs) pcore_setI_normal ?pcore_normal //. + by have [_] := dprodP (nilpotent_pcoreC q^' nilLs); rewrite pcoreNK. + apply: contra ntK => cQQ; rewrite -subG1 /= -(derG1P cQQ) -subsetIidl. + rewrite -(pprod_focal_coprime defLs) ?subsetIidl ?pcore_normal //. + by rewrite coprime_sym (coprimeSg sKQ) ?coprime_pcoreC. +pose X := 'C_A(H`_\sigma)%G. +have [sXA cHsX]: X \subset A /\ X \subset 'C(H`_\sigma) by apply/subsetIP. +have{not_cQQ} oX: #|X| = q. + by have [_ []] := nonabelian_tau2 maxH hallD t2Hq Eq2A qQ not_cQQ. +have neqXK: X :!=: K. + apply: contraNneq neqHL => defX; rewrite (eq_mmax maxH maxL) //. + have [_ <- _ _] := sdprodP (sdprod_sigma maxH hallD). + by rewrite mulG_subG subsetIr (subset_trans _ sCKL) // centsC -defX. +have{neqXK sXA} not_sXM: ~~ (X \subset M). + apply: contra neqXK => sXM; rewrite eqEcard oX leqnn andbT; apply/joing_idPl. + have [[sKM kK _] cKX] := (and3P hallK, subset_trans sXA cKA). + apply: sub_pHall hallK _ (joing_subl _ _) _; last by rewrite join_subG sKM. + by rewrite /= (cent_joinEr cKX) pgroupM {2}/pgroup oX andbb. +have{not_sXM} not_sCUM: ~~ ('C(U) \subset M). + exact: contra (subset_trans (centsS sUHs cHsX)) not_sXM. +apply/pgroupP=> r r_pr _; apply: contra not_sCUM => /= t2Mr. +have [hallUK _ _ _ _] := kappa_compl_context maxM complU. +have [[B Er2B _] [sUKM _]] := (ex_tau2Elem hallUK t2Mr, andP hallUK). +have [[nsBUK _] [sCBUK _ _] _ _ ] := tau2_compl_context maxM hallUK t2Mr Er2B. +apply: subset_trans (centS _) (subset_trans sCBUK sUKM). +have [sBUK /andP[rB _] _] := pnElemP Er2B. +have maxU_UK := Hall_max (pHall_subl (joing_subl _ _) sUKM hallU). +rewrite (normal_sub_max_pgroup maxU_UK) // (pi_pgroup rB) //. +apply: contraL t2Mr; rewrite negb_and negbK /= inE /=. +by case: (r \in _) => //=; move/rank_kappa->. +Qed. + +(* This is B & G, Theorem 15.9, parts (a) and (b), due to D. Sibley and Feit *) +(* and Thompson, respectively. *) +(* We have dropped part (c) because it is not used later, and localizing the *) +(* quantification on r would complicate the proof needlessly. *) +Theorem nonFtype_signalizer_base M x : + M \in 'M -> x \in M`_\sigma^# -> + ~~ ('C[x] \subset M) -> 'N[x] \notin 'M_'F -> + [/\ (*a*) M \in 'M_'F, 'N[x] \in 'M_'P2 + & (*b*) exists2 E : {group gT}, \sigma(M)^'.-Hall(M) E + & cyclic (gval E) /\ [Frobenius M = M`_\sigma ><| E]]. +Proof. +move=> maxM Ms1x not_sCxM notFmaxN; have ell1x := Msigma_ell1 maxM Ms1x. +have [[ntx Ms_x] [y cxy notMy]] := (setD1P Ms1x, subsetPn not_sCxM). +have SMxM: M \in 'M_\sigma[x] by rewrite inE maxM cycle_subG. +have SMxMy: (M :^ y)%G \in 'M_\sigma[x]. + by rewrite inE mmaxJ maxM gen_subG -(normP cxy) /= MsigmaJ conjSg sub1set. +have neq_MyM: M :^ y != M by rewrite (sameP eqP normP) norm_mmax. +have SMx_gt1: #|'M_\sigma[x]| > 1. + by rewrite (cardD1 M) SMxM (cardD1 (M :^ y)%G) inE /= SMxMy neq_MyM. +have [_ [//|uniqN _ t2Nx]] := FT_signalizer_context ell1x. +rewrite inE (negPf notFmaxN) /= => P2maxN /(_ M SMxM)[_ st2NsM _ hallMN]. +pose r := pdiv #[x]; have pixr: r \in \pi(#[x]) by rewrite pi_pdiv order_gt1. +have t2Nr := pnatPpi t2Nx pixr; have sMr := st2NsM r t2Nr. +have [[PmaxN _] [_ s'N_MN _]] := (setDP P2maxN, and3P hallMN). +have [K hallK] := Hall_exists \kappa('N[x]) (sigma_compl_sol hallMN). +have [U hallU] := Hall_exists \kappa('N[x])^' (sigma_compl_sol hallMN). +have hallK_N := subHall_Hall hallMN (@kappa_sigma' _ _) hallK. +have [[sKMN kK _] [sUMN k'U _]] := (and3P hallK, and3P hallU). +have mulUK: U * K = M :&: 'N[x]. + apply/eqP; rewrite eqEcard mulG_subG sUMN sKMN. + rewrite coprime_cardMg ?(p'nat_coprime k'U) //= mulnC. + by rewrite (card_Hall hallU) (card_Hall hallK) partnC ?cardG_gt0. +have complU: kappa_complement 'N[x] U K. + split=> //; last by rewrite mulUK groupP. + apply: (subHall_Hall hallMN) => [p|]; first by case/norP. + rewrite pHallE sUMN /= (card_Hall hallU) eq_sym; apply/eqP. + apply: eq_in_partn => p piMNp; rewrite inE /= negb_or /=. + by rewrite [~~ _](pnatPpi s'N_MN). +have prK: prime #|K| by case/Ptype_structure: hallK_N => // _ _ _ _ []. +have ntK: K :!=: 1 by rewrite -cardG_gt1 prime_gt1. +have [maxN _] := setDP PmaxN. +have [defUK cUU regUK]: [/\ U ><| K = M :&: 'N[x], abelian U & 'C_U(K) = 1]. + have [_ defM _ regUK -> //] := kappa_compl_context maxN complU. + have [[_ UK _ defUK] _ _ _] := sdprodP defM. + by rewrite (cent_semiregular regUK) // defUK; case/sdprodP: defUK => _ <-. +have [[R sylR] [s'Nr rrN]] := (Sylow_exists r (M :&: 'N[x]), andP t2Nr). +have [[sRMN rR _] sylR_N] := (and3P sylR, subHall_Sylow hallMN s'Nr sylR). +have [nsUMN _ _ nUK _] := sdprod_context defUK. +have [[sRM sRN] [sKM _]] := (subsetIP sRMN, subsetIP sKMN). +have sRU: R \subset U. + rewrite (sub_normal_Hall hallU nsUMN sRMN) (pi_pgroup rR) //. + by apply: contraL rrN; move/rank_kappa->. +have sNRM: 'N(R) \subset M. + apply: (norm_noncyclic_sigma maxM sMr rR sRM). + rewrite (odd_pgroup_rank1_cyclic rR) ?mFT_odd //. + by rewrite (p_rank_Sylow sylR_N) (eqnP rrN). +have sylR_U := pHall_subl sRU sUMN sylR. +have maxNRM: M \in 'M('N(R)) by rewrite inE maxM. +have [L maxCK_L] := mmax_exists (mFT_cent_proper ntK). +have FmaxM: M \in 'M_'F; last split=> //. + by have [] := P2type_signalizer P2maxN complU maxCK_L sylR_U maxNRM. +have nilMs: nilpotent M`_\sigma by rewrite notP1type_Msigma_nil ?FmaxM. +have sMsF: M`_\sigma \subset 'F(M) by rewrite Fitting_max ?pcore_normal. +have defR: R :=: 'O_r(U) := nilpotent_Hall_pcore (abelian_nil cUU) sylR_U. +have nRK: K \subset 'N(R) by rewrite defR (char_norm_trans (pcore_char r U)). +have ntR: R :!=: 1. + rewrite -rank_gt0 (rank_Sylow sylR_N) p_rank_gt0. + by rewrite (partition_pi_mmax maxN) t2Nr !orbT. +have not_nilRK: ~~ nilpotent (R <*> K). + apply: contra ntR => nilRK; rewrite -subG1 -regUK subsetI sRU. + rewrite (sub_nilpotent_cent2 nilRK) ?joing_subl ?joing_subr //. + by rewrite (coprimegS sRU) ?(pnat_coprime kK). +have{not_nilRK} not_sKMs: ~~ (K \subset M`_\sigma). + apply: contra not_nilRK => sKMs; apply: nilpotentS nilMs. + by rewrite join_subG (sub_Hall_pcore (Msigma_Hall maxM)) // (pi_pgroup rR). +have s'MK: \sigma(M)^'.-group K. + rewrite /pgroup pnatE //; apply: contra not_sKMs; rewrite /= -pnatE // => sK. + by rewrite (sub_Hall_pcore (Msigma_Hall maxM)). +have [E hallE sKE] := Hall_superset (mmax_sol maxM) sKM s'MK. +have [[E1 hallE1] [E3 hallE3]] := ex_tau13_compl hallE. +have [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. +have [[_ t1E1 _] [sEM _]] := (and3P hallE1, andP hallE). +have E2_1: E2 :=: 1. + have [sE2E t2E2 _] := and3P hallE2. + rewrite -(setIidPl sE2E) coprime_TIg ?(pnat_coprime t2E2 (pgroupS sEM _)) //. + apply: contraR ntR => not_t2'M. + have:= tau2_P2type_signalizer P2maxN complU maxCK_L sylR_U maxNRM not_t2'M. + case=> _ _ t2'N; rewrite -(setIidPl sRN) coprime_TIg //. + by rewrite (pnat_coprime (pi_pgroup rR t2Nr)). +have E3_1: E3 :=: 1. + have ntX: 'F(M) :&: 'F(M) :^ y != 1. + apply/trivgPn; exists x; rewrite // inE mem_conjg !(subsetP sMsF) //. + by rewrite /conjg invgK mulgA (cent1P cxy) mulgK. + have [_ _ _ defE _] := nonTI_Fitting_structure maxM notMy ntX. + by case/defE: complEi. +have [cycE defE]: cyclic E /\ E :=: E1. + have [_ _ [cycE1 _] [<- _] _] := sigma_compl_context maxM complEi. + by rewrite E2_1 E3_1 !sdprod1g. +have [ntMs ntE] := (Msigma_neq1 maxM, subG1_contra sKE ntK). +have defM := sdprod_sigma maxM hallE. +exists E => //; split=> //; apply/Frobenius_semiregularP=> // z Ez. +have{Ez} [ntz szE1] := setD1P Ez; rewrite defE -cycle_subG in szE1. +pose q := pdiv #[z]; have pizq: q \in \pi(#[z]) by rewrite pi_pdiv order_gt1. +have szM: <[z]> \subset M by rewrite (subset_trans _ sEM) ?defE. +have [_ k'M] := setIdP FmaxM; have k'q := pnatPpi k'M (piSg szM pizq). +have t1q := pnatPpi t1E1 (piSg szE1 pizq). +move: pizq; rewrite -p_rank_gt0 => /p_rank_geP[Z]. +rewrite /= -(setIidPr szM) pnElemI -setIdE => /setIdP[EqZ sZz]. +apply: contraNeq k'q => ntCMsx /=. +rewrite unlock 3!inE /= t1q; apply/exists_inP; exists Z => //. +by rewrite (subG1_contra _ ntCMsx) ?setIS //= -cent_cycle centS. +Qed. + +End Section15. + + diff --git a/mathcomp/odd_order/BGsection16.v b/mathcomp/odd_order/BGsection16.v new file mode 100644 index 0000000..a37edba --- /dev/null +++ b/mathcomp/odd_order/BGsection16.v @@ -0,0 +1,1359 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import bigop finset prime fingroup morphism perm automorphism quotient. +Require Import action gproduct gfunctor pgroup cyclic center commutator. +Require Import gseries nilpotent sylow abelian maximal hall frobenius. +Require Import BGsection1 BGsection2 BGsection3 BGsection4 BGsection5. +Require Import BGsection6 BGsection7 BGsection9 BGsection10 BGsection12. +Require Import BGsection13 BGsection14 BGsection15. + +(******************************************************************************) +(* This file covers B & G, section 16; it summarises all the results of the *) +(* local analysis. Some of the definitions of B & G have been adapted to fit *) +(* in beter with Peterfalvi, section 8, dropping unused properties and adding *) +(* a few missing ones. This file also defines the following: *) +(* of_typeF M U <-> M = M`_\F ><| U is of type F, in the sense of *) +(* Petervalvi (8.1) rather than B & G section 14. *) +(* is_typeF_complement M U U0 <-> U0 is a subgroup of U with the same *) +(* exponent as U, such that M`_\F ><| U0 is a Frobenius *) +(* group; this corresponds to Peterfalvi (8.1)(c). *) +(* is_typeF_inertia M U U1 <-> U1 <| U is abelian and contains 'C_U[x] for *) +(* all x in M`_\F^#, and thus the inertia groups of all *) +(* nonprincipal irreducible characters of M`_\F; this *) +(* corresponds to Peterfalvi (8.1)(b). *) +(* of_typeI M U <-> M = M`_\F ><| U is of type I, in the sense of *) +(* Peterfalvi (8.3); this definition is almost identical *) +(* to B & G conditions (Ii) - (Iv), except that (Iiv) is *) +(* dropped, as is the condition p \in \pi* in (Iv)(c). *) +(* Also, the condition 'O_p^'(M) cyclic, present in both *) +(* B & G and Peterfalvi, is weakened to 'O_p^'(M`_\F) *) +(* cyclic, because B & G, Theorem 15.7 only proves the *) +(* weaker statement, and we did not manage to improve it. *) +(* This appears to be a typo in B & G that was copied *) +(* over to Petrfalvi (8.3). It is probably no consequence *) +(* because (8.3) is only used in (12.6) and (12.10) and *) +(* neither use the assumption that 'O_p^'(M) is cyclic. *) +(* For defW : W1 \x W2 = W we also define: *) +(* of_typeP M U defW <-> M = M`_\F ><| U ><| W1 is of type P, in the sense of *) +(* Peterfalvi (8.4) rather than B & G section 14, where *) +(* (8.4)(d,e) hold for W and W2 (i.e., W2 = C_M^(1)(W1)). *) +(* of_typeII_IV M U defW <-> M = M`_\F ><| U ><| W1 is of type II, III, or IV *) +(* in the sense of Peterfalvi (8.6)(a). This is almost *) +(* exactly the contents of B & G, (T1)-(T7), except that *) +(* (T6) is dropped, and 'C_(M`_\F)(W1) \subset M^`(2) is *) +(* added (PF, (8.4)(d) and B & G, Theorem C(3)). *) +(* of_typeII M U defW <-> M = M`_\F ><| U ><| W1 is of type II in the sense *) +(* of Peterfalvi (8.6); this differs from B & G by *) +(* dropping the rank 2 clause in IIiii and replacing IIv *) +(* by B(2)(3) (note that IIv is stated incorrectly: M' *) +(* should be M'^#). *) +(* of_typeIII M U defW <-> M = M`_\F ><| U ><| W1 is of type III in the sense *) +(* of Peterfalvi (8.6). *) +(* of_typeIV M U defW <-> M = M`_\F ><| U ><| W1 is of type IV in the sense *) +(* of Peterfalvi (8.6). *) +(* of_typeV M U defW <-> U = 1 and M = M`_\F ><| W1 is of type V in the *) +(* sense of Peterfalvi (8.7); this differs from B & G (V) *) +(* dropping the p \in \pi* condition in clauses (V)(b) *) +(* and (V)(c). *) +(* exists_typeP spec <-> spec U defW holds for some groups U, W, W1 and W2 *) +(* such that defW : W1 \x W2 = W; spec will be one of *) +(* (of_typeP M), (of_typeII_IV M), (of_typeII M), etc. *) +(* FTtype_spec i M <-> M is of type i, for 0 < i <= 5, in the sense of the *) +(* predicates above, for the appropriate complements to *) +(* M`_\F and M^`(1). *) +(* FTtype M == the type of M, in the sense above, when M is a maximal *) +(* subgroup of G (this is an integer i between 1 and 5). *) +(* M`_\s == an alternative, combinatorial definition of M`_\sigma *) +(* := M`_\F if M is of type I or II, else M^`(1) *) +(* 'A1(M) == the "inner Dade support" of a maximal group M, as *) +(* defined in Peterfalvi (8.10). *) +(* := (M`_\s)^# *) +(* 'A(M) == the "Dade support" of M as defined in Peterfalvi (8.10) *) +(* (this differs from B & G by excluding 1). *) +(* 'A0(M) == the "outer Dade support" of M as defined in Peterfalvi *) +(* (8.10) (this differs from B & G by excluding 1). *) +(* 'M^G == a transversal of the conjugacy classes of 'M. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Section GeneralDefinitions. + +Variable gT : finGroupType. +Implicit Types V W X : {set gT}. + +End GeneralDefinitions. + +Section Definitions. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). + +Implicit Types M U V W X : {set gT}. + +Definition is_typeF_inertia M U (H := M`_\F) U1 := + [/\ U1 <| U, abelian U1 & {in H^#, forall x, 'C_U[x] \subset U1}]. + +Definition is_typeF_complement M U (H := M`_\F) U0 := + [/\ U0 \subset U, exponent U0 = exponent U & [Frobenius H <*> U0 = H ><| U0]]. + +Definition of_typeF M U (H := M`_\F) := + [/\ (*a*) [/\ H != 1, U :!=: 1 & H ><| U = M], + (*b*) exists U1 : {group gT}, is_typeF_inertia M U U1 + & (*c*) exists U0 : {group gT}, is_typeF_complement M U U0]. + +Definition of_typeI M (H := M`_\F) U := + of_typeF M U + /\ [\/ (*a*) normedTI H^# G M, + (*b*) abelian H /\ 'r(H) = 2 + | (*c*) {in \pi(H), forall p, exponent U %| p.-1} + /\ (exists2 p, p \in \pi(H) & cyclic 'O_p^'(H))]. + +Section Ptypes. + +Variables M U W W1 W2 : {set gT}. +Let H := M`_\F. +Let M' := M^`(1). +Implicit Type defW : W1 \x W2 = W. + +Definition of_typeP defW := + [/\ (*a*) [/\ cyclic W1, Hall M W1, W1 != 1 & M' ><| W1 = M], + (*b*) [/\ nilpotent U, U \subset M', W1 \subset 'N(U) & H ><| U = M'], + (*c*) [/\ ~~ cyclic H, M^`(2) \subset 'F(M), H * 'C_M(H) = 'F(M) + & 'F(M) \subset M'], + (*d*) [/\ cyclic W2, W2 != 1, W2 \subset H, W2 \subset M^`(2) + & {in W1^#, forall x, 'C_M'[x] = W2}] + & (*e*) normedTI (W :\: (W1 :|: W2)) G W]. + +Definition of_typeII_IV defW := + [/\ of_typeP defW, U != 1, prime #|W1| & normedTI 'F(M)^# G M]. + +Definition of_typeII defW := + [/\ of_typeII_IV defW, abelian U, ~~ ('N(U) \subset M), + of_typeF M' U & M'`_\F = H]. + +Definition of_typeIII defW := + [/\ of_typeII_IV defW, abelian U & 'N(U) \subset M]. + +Definition of_typeIV defW := + [/\ of_typeII_IV defW, ~~ abelian U & 'N(U) \subset M]. + +Definition of_typeV defW := + [/\ of_typeP defW /\ U = 1 + & [\/ (*a*) normedTI H^# G M, + (*b*) exists2 p, p \in \pi(H) & #|W1| %| p.-1 /\ cyclic 'O_p^'(H) + | (*c*) exists2 p, p \in \pi(H) + & [/\ #|'O_p(H)| = (p ^ 3)%N, #|W1| %| p.+1 & cyclic 'O_p^'(H)]]]. + +End Ptypes. + +CoInductive exists_typeP (spec : forall U W W1 W2, W1 \x W2 = W -> Prop) : Prop + := FTtypeP_Spec (U W W1 W2 : {group gT}) defW of spec U W W1 W2 defW. + +Definition FTtype_spec i M := + match i with + | 1%N => exists U : {group gT}, of_typeI M U + | 2 => exists_typeP (of_typeII M) + | 3 => exists_typeP (of_typeIII M) + | 4 => exists_typeP (of_typeIV M) + | 5 => exists_typeP (of_typeV M) + | _ => False + end. + +Definition FTtype M := + if \kappa(M)^'.-group M then 1%N else + if M`_\sigma != M^`(1) then 2 else + if M`_\F == M`_\sigma then 5 else + if abelian (M`_\sigma / M`_\F) then 3 else 4. + +Lemma FTtype_range M : 0 < FTtype M <= 5. +Proof. by rewrite /FTtype; do 4!case: ifP => // _. Qed. + +Definition FTcore M := if 0 < FTtype M <= 2 then M`_\F else M^`(1). +Fact FTcore_is_group M : group_set (FTcore M). +Proof. rewrite /FTcore; case: ifP => _; exact: groupP. Qed. +Canonical Structure FTcore_group M := Group (FTcore_is_group M). + +Definition FTsupport1 M := (FTcore M)^#. + +Let FTder M := M^`(FTtype M != 1%N). + +Definition FTsupport M := \bigcup_(x in FTsupport1 M) 'C_(FTder M)[x]^#. + +Definition FTsupport0 M (pi := \pi(FTder M)) := + FTsupport M :|: [set x in M | ~~ pi.-elt x & ~~ pi^'.-elt x]. + +Definition mmax_transversal U := orbit_transversal 'JG U 'M. + +End Definitions. + +Notation "M `_ \s" := (FTcore M) (at level 3, format "M `_ \s") : group_scope. +Notation "M `_ \s" := (FTcore_group M) : Group_scope. + +Notation "''A1' ( M )" := (FTsupport1 M) + (at level 8, format "''A1' ( M )") : group_scope. + +Notation "''A' ( M )" := (FTsupport M) + (at level 8, format "''A' ( M )") : group_scope. + +Notation "''A0' ( M )" := (FTsupport0 M) + (at level 8, format "''A0' ( M )") : group_scope. + +Notation "''M^' G" := (mmax_transversal G) + (at level 3, format "''M^' G") : group_scope. + +Section Section16. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types p q q_star r : nat. +Implicit Types x y z : gT. +Implicit Types A E H K L M Mstar N P Q Qstar R S T U V W X Y Z : {group gT}. + +(* Structural properties of the M`_\s definition. *) +Lemma FTcore_char M : M`_\s \char M. +Proof. by rewrite /FTcore; case: ifP; rewrite gFchar. Qed. + +Lemma FTcore_normal M : M`_\s <| M. +Proof. by rewrite char_normal ?FTcore_char. Qed. + +Lemma FTcore_norm M : M \subset 'N(M`_\s). +Proof. by rewrite char_norm ?FTcore_char. Qed. + +Lemma FTcore_sub M : M`_\s \subset M. +Proof. by rewrite char_sub ?FTcore_char. Qed. + +Lemma FTcore_type1 M : FTtype M == 1%N -> M`_\s = M`_\F. +Proof. by rewrite /M`_\s => /eqP->. Qed. + +Lemma FTcore_type2 M : FTtype M == 2 -> M`_\s = M`_\F. +Proof. by rewrite /M`_\s => /eqP->. Qed. + +Lemma FTcore_type_gt2 M : FTtype M > 2 -> M`_\s = M^`(1). +Proof. by rewrite /M`_\s => /subnKC <-. Qed. + +Lemma FTsupp1_type1 M : FTtype M == 1%N -> 'A1(M) = M`_\F^#. +Proof. by move/FTcore_type1 <-. Qed. + +Lemma FTsupp1_type2 M : FTtype M == 2 -> 'A1(M) = M`_\F^#. +Proof. by move/FTcore_type2 <-. Qed. + +Lemma FTsupp1_type_gt2 M : FTtype M > 2 -> 'A1(M) = M^`(1)^#. +Proof. by move/FTcore_type_gt2 <-. Qed. + +(* This section covers the characterization of the F, P, P1 and P2 types of *) +(* maximal subgroups summarized at the top of p. 125. in B & G. *) +Section KappaClassification. + +Variables M U K : {group gT}. +Hypotheses (maxM : M \in 'M) (complU : kappa_complement M U K). + +Remark trivgFmax : (M \in 'M_'F) = (K :==: 1). +Proof. by rewrite (trivg_kappa maxM); case: complU. Qed. + +Remark trivgPmax : (M \in 'M_'P) = (K :!=: 1). +Proof. by rewrite inE trivgFmax maxM andbT. Qed. + +Remark FmaxP : reflect (K :==: 1 /\ U :!=: 1) (M \in 'M_'F). +Proof. +rewrite (trivg_kappa_compl maxM complU) 2!inE. +have [_ hallK _] := complU; rewrite (trivg_kappa maxM hallK). +by apply: (iffP idP) => [-> | []]. +Qed. + +Remark P1maxP : reflect (K :!=: 1 /\ U :==: 1) (M \in 'M_'P1). +Proof. +rewrite (trivg_kappa_compl maxM complU) inE -trivgPmax. +by apply: (iffP idP) => [|[] //]; case/andP=> ->. +Qed. + +Remark P2maxP : reflect (K :!=: 1 /\ U :!=: 1) (M \in 'M_'P2). +Proof. +rewrite (trivg_kappa_compl maxM complU) -trivgPmax. +by apply: (iffP setDP) => [] []. +Qed. + +End KappaClassification. + +(* This section covers the combinatorial statements of B & G, Lemma 16.1. It *) +(* needs to appear before summary theorems A-E because we are following *) +(* Peterfalvi in anticipating the classification in the definition of the *) +(* kernel sets A1(M), A(M) and A0(M). The actual correspondence between the *) +(* combinatorial classification and the mathematical description, i.e., the *) +(* of_typeXX predicates, will be given later. *) +Section FTtypeClassification. + +Variable M : {group gT}. +Hypothesis maxM : M \in 'M. + +Lemma kappa_witness : + exists UK : {group gT} * {group gT}, kappa_complement M UK.1 UK.2. +Proof. +have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). +by have [U complU] := ex_kappa_compl maxM hallK; exists (U, K). +Qed. + +(* This is B & G, Lemma 16.1(a). *) +Lemma FTtype_Fmax : (M \in 'M_'F) = (FTtype M == 1%N). +Proof. +by rewrite inE maxM /FTtype; case: (_.-group M) => //; do 3!case: ifP => // _. +Qed. + +Lemma FTtype_Pmax : (M \in 'M_'P) = (FTtype M != 1%N). +Proof. by rewrite inE maxM andbT FTtype_Fmax. Qed. + +(* This is B & G, Lemma 16.1(b). *) +Lemma FTtype_P2max : (M \in 'M_'P2) = (FTtype M == 2). +Proof. +have [[U K] /= complU] := kappa_witness. +rewrite (sameP (P2maxP maxM complU) andP) -(trivgFmax maxM complU) FTtype_Fmax. +have [-> // | notMtype1] := altP eqP. +have ntK: K :!=: 1 by rewrite -(trivgFmax maxM complU) FTtype_Fmax. +have [_ [//|defM' _] _ _ _] := kappa_structure maxM complU. +do [rewrite /FTtype; case: ifP => // _] in notMtype1 *. +rewrite -cardG_gt1 eqEcard Msigma_der1 //= -(sdprod_card defM') -ltnNge. +rewrite -(@ltn_pmul2l #|M`_\sigma|) ?cardG_gt0 //= muln1. +by case: leqP => // _; do 2!case: ifP=> //. +Qed. + +(* This covers the P1 part of B & G, Lemma 16.1 (c) and (d). *) +Lemma FTtype_P1max : (M \in 'M_'P1) = (2 < FTtype M <= 5). +Proof. +rewrite 2!ltn_neqAle -!andbA FTtype_range andbT eq_sym -FTtype_P2max. +rewrite eq_sym -FTtype_Pmax in_setD inE. +by case: (M \in _); rewrite ?andbT ?andbF ?negbK. +Qed. + +Lemma Msigma_eq_der1 : M \in 'M_'P1 -> M`_\sigma = M^`(1). +Proof. +have [[U K] /= complU] := kappa_witness. +case/(P1maxP maxM complU)=> ntK; move/eqP=> U1. +by have [_ [//|<- _] _ _ _] := kappa_structure maxM complU; rewrite U1 sdprodg1. +Qed. + +Lemma def_FTcore : M`_\s = M`_\sigma. +Proof. +rewrite /FTcore -mem_iota 2!inE -FTtype_Fmax -FTtype_P2max. +have [notP1maxM |] := ifPn. + by apply/Fcore_eq_Msigma; rewrite ?notP1type_Msigma_nil. +case/norP=> notFmaxM; rewrite inE andbC inE maxM notFmaxM negbK => P1maxM. +by rewrite Msigma_eq_der1. +Qed. + +(* Other relations between the 'core' groups. *) + +Lemma FTcore_sub_der1 : M`_\s \subset M^`(1)%g. +Proof. by rewrite def_FTcore Msigma_der1. Qed. + +Lemma Fcore_sub_FTcore : M`_\F \subset M`_\s. +Proof. by rewrite def_FTcore Fcore_sub_Msigma. Qed. + +Lemma mmax_Fcore_neq1 : M`_\F != 1. +Proof. by have [[]] := Fcore_structure maxM. Qed. + +Lemma mmax_Fitting_neq1 : 'F(M) != 1. +Proof. exact: subG1_contra (Fcore_sub_Fitting M) mmax_Fcore_neq1. Qed. + +Lemma FTcore_neq1 : M`_\s != 1. +Proof. exact: subG1_contra Fcore_sub_FTcore mmax_Fcore_neq1. Qed. + +Lemma norm_mmax_Fcore : 'N(M`_\F) = M. +Proof. exact: mmax_normal (gFnormal _ _) mmax_Fcore_neq1. Qed. + +Lemma norm_FTcore : 'N(M`_\s) = M. +Proof. exact: mmax_normal (FTcore_normal _) FTcore_neq1. Qed. + +Lemma norm_mmax_Fitting : 'N('F(M)) = M. +Proof. exact: mmax_normal (gFnormal _ _) mmax_Fitting_neq1. Qed. + +(* This is B & G, Lemma 16.1(f). *) +Lemma Fcore_eq_FTcore : reflect (M`_\F = M`_\s) (FTtype M \in pred3 1%N 2 5). +Proof. +rewrite /FTcore -mem_iota 3!inE orbA; case type12M: (_ || _); first by left. +move: type12M FTtype_P1max; rewrite /FTtype; do 2![case: ifP => // _] => _. +rewrite !(fun_if (leq^~ 5)) !(fun_if (leq 3)) !if_same /= => P1maxM. +rewrite Msigma_eq_der1 // !(fun_if (eq_op^~ 5)) if_same. +by rewrite [if _ then _ else _]andbT; apply: eqP. +Qed. + +(* This is the second part of B & G, Lemma 16.1(c). *) +Lemma Fcore_neq_FTcore : (M`_\F != M`_\s) = (FTtype M \in pred2 3 4). +Proof. +have:= FTtype_range M; rewrite -mem_iota (sameP eqP Fcore_eq_FTcore). +by do 5!case/predU1P=> [-> //|]. +Qed. + +Lemma FTcore_eq_der1 : FTtype M > 2 -> M`_\s = M^`(1). +Proof. +move=> FTtype_gt2; rewrite def_FTcore Msigma_eq_der1 // FTtype_P1max. +by rewrite FTtype_gt2; case/andP: (FTtype_range M). +Qed. + +End FTtypeClassification. + +(* Internal automorphism. *) +Lemma FTtypeJ M x : FTtype (M :^ x) = FTtype M. +Proof. +rewrite /FTtype (eq_p'group _ (kappaJ _ _)) pgroupJ MsigmaJ FcoreJ derJ. +rewrite !(can_eq (conjsgK x)); do 4!congr (if _ then _ else _). +rewrite -quotientInorm normJ -conjIg /= setIC -{1 3}(setIidPr (normG M`_\F)). +rewrite -!morphim_conj -morphim_quotm ?normalG //= => nsMFN. +by rewrite injm_abelian /= ?im_quotient // injm_quotm ?injm_conj. +Qed. + +Lemma FTcoreJ M x : (M :^ x)`_\s = M`_\s :^ x. +Proof. by rewrite /FTcore FTtypeJ FcoreJ derJ; case: ifP. Qed. + +Lemma FTsupp1J M x : 'A1(M :^ x) = 'A1(M) :^ x. +Proof. by rewrite conjD1g -FTcoreJ. Qed. + +Lemma FTsuppJ M x : 'A(M :^ x) = 'A(M) :^ x. +Proof. +rewrite -bigcupJ /'A(_) FTsupp1J big_imset /=; last exact: in2W (conjg_inj x). +by apply: eq_bigr => y _; rewrite FTtypeJ derJ cent1J -conjIg conjD1g. +Qed. + +Lemma FTsupp0J M x : 'A0(M :^ x) = 'A0(M) :^ x. +Proof. +apply/setP=> y; rewrite mem_conjg !inE FTsuppJ !mem_conjg; congr (_ || _ && _). +by rewrite FTtypeJ !p_eltJ derJ /= cardJg. +Qed. + +(* Inclusion/normality of class function supports. *) + +Lemma FTsupp_sub0 M : 'A(M) \subset 'A0(M). +Proof. exact: subsetUl. Qed. + +Lemma FTsupp0_sub M : 'A0(M) \subset M^#. +Proof. +rewrite subUset andbC subsetD1 setIdE subsetIl !inE p_elt1 andbF /=. +by apply/bigcupsP=> x _; rewrite setSD ?subIset ?der_sub. +Qed. + +Lemma FTsupp_sub M : 'A(M) \subset M^#. +Proof. exact: subset_trans (FTsupp_sub0 M) (FTsupp0_sub M). Qed. + +Lemma FTsupp1_norm M : M \subset 'N('A1(M)). +Proof. by rewrite normD1 (normal_norm (FTcore_normal M)). Qed. + +Lemma FTsupp_norm M : M \subset 'N('A(M)). +Proof. +apply/subsetP=> y My; rewrite inE -bigcupJ; apply/bigcupsP=> x A1x. +rewrite (bigcup_max (x ^ y)) ?memJ_norm ?(subsetP (FTsupp1_norm M)) //. +by rewrite conjD1g conjIg cent1J (normsP _ y My) ?gFnorm. +Qed. + +Lemma FTsupp0_norm M : M \subset 'N('A0(M)). +Proof. +rewrite normsU ?FTsupp_norm // setIdE normsI //. +by apply/normsP=> x _; apply/setP=> y; rewrite mem_conjg !inE !p_eltJ. +Qed. + +Section MmaxFTsupp. +(* Support inclusions that depend on the maximality of M. *) + +Variable M : {group gT}. +Hypothesis maxM : M \in 'M. + +Lemma FTsupp1_sub : 'A1(M) \subset 'A(M). +Proof. +apply/subsetP=> x A1x; apply/bigcupP; exists x => //. +have [ntx Ms_x] := setD1P A1x; rewrite 3!inE ntx cent1id (subsetP _ x Ms_x) //. +by case: (~~ _); rewrite ?FTcore_sub_der1 ?FTcore_sub. +Qed. + +Lemma FTsupp1_sub0 : 'A1(M) \subset 'A0(M). +Proof. exact: subset_trans FTsupp1_sub (FTsupp_sub0 M). Qed. + +Lemma FTsupp1_neq0 : 'A1(M) != set0. +Proof. by rewrite setD_eq0 subG1 FTcore_neq1. Qed. + +Lemma FTsupp_neq0 : 'A(M) != set0. +Proof. +by apply: contraNneq FTsupp1_neq0 => AM_0; rewrite -subset0 -AM_0 FTsupp1_sub. +Qed. + +Lemma FTsupp0_neq0 : 'A0(M) != set0. +Proof. +by apply: contraNneq FTsupp_neq0 => A0M_0; rewrite -subset0 -A0M_0 FTsupp_sub0. +Qed. + +Lemma Fcore_sub_FTsupp1 : M`_\F^# \subset 'A1(M). +Proof. exact: setSD (Fcore_sub_FTcore maxM). Qed. + +Lemma Fcore_sub_FTsupp : M`_\F^# \subset 'A(M). +Proof. exact: subset_trans Fcore_sub_FTsupp1 FTsupp1_sub. Qed. + +Lemma Fcore_sub_FTsupp0 : M`_\F^# \subset 'A0(M). +Proof. exact: subset_trans Fcore_sub_FTsupp1 FTsupp1_sub0. Qed. + +Lemma Fitting_sub_FTsupp : 'F(M)^# \subset 'A(M). +Proof. +pose pi := \pi(M`_\F); have nilF := Fitting_nil M. +have [U defF]: {U : {group gT} | M`_\F \x U = 'F(M)}. + have hallH := pHall_subl (Fcore_sub_Fitting M) (gFsub _ _) (Fcore_Hall M). + exists 'O_pi^'('F(M))%G; rewrite (nilpotent_Hall_pcore nilF hallH). + exact: nilpotent_pcoreC. +apply/subsetP=> xy /setD1P[ntxy Fxy]; apply/bigcupP. +have [x [y [Hx Vy Dxy _]]] := mem_dprod defF Fxy. +have [z [ntz Hz czxy]]: exists z, [/\ z != 1%g, z \in M`_\F & x \in 'C[z]]. + have [-> | ntx] := eqVneq x 1%g; last by exists x; rewrite ?cent1id. + by have /trivgPn[z ntz] := mmax_Fcore_neq1 maxM; exists z; rewrite ?group1. +exists z; first by rewrite !inE ntz (subsetP (Fcore_sub_FTcore maxM)). +rewrite 3!inE ntxy {2}Dxy groupMl //= andbC (subsetP _ y Vy) //=; last first. + by rewrite sub_cent1 (subsetP _ _ Hz) // centsC; have [] := dprodP defF. +rewrite -FTtype_Pmax // (subsetP _ xy Fxy) //. +case MtypeP: (M \in _); last exact: gFsub. +by have [_ _ _ ->] := Fitting_structure maxM. +Qed. + +Lemma Fitting_sub_FTsupp0 : 'F(M)^# \subset 'A0(M). +Proof. exact: subset_trans Fitting_sub_FTsupp (FTsupp_sub0 M). Qed. + +Lemma FTsupp_eq1 : (2 < FTtype M)%N -> 'A(M) = 'A1(M). +Proof. +move=> typeMgt2; rewrite /'A(M) -(subnKC typeMgt2) /= -FTcore_eq_der1 //. +apply/setP=> y; apply/bigcupP/idP=> [[x A1x /setD1P[nty /setIP[Ms_y _]]] | A1y]. + exact/setD1P. +by exists y; rewrite // inE in_setI cent1id andbT -in_setD. +Qed. + +End MmaxFTsupp. + +Section SingleGroupSummaries. + +Variables M U K : {group gT}. +Hypotheses (maxM : M \in 'M) (complU : kappa_complement M U K). + +Let Kstar := 'C_(M`_\sigma)(K). + +Theorem BGsummaryA : + [/\ (*1*) [/\ M`_\sigma <| M, \sigma(M).-Hall(M) M`_\sigma & + \sigma(M).-Hall(G) M`_\sigma], + (*2*) \kappa(M).-Hall(M) K /\ cyclic K, + (*3*) [/\ U \in [complements to M`_\sigma <*> K in M], + K \subset 'N(U), + M`_\sigma <*> U <| M, + U <| U <*> K + & M`_\sigma * U * K = M], + (*4*) {in K^#, forall k, 'C_U[k] = 1} + & + [/\ (*5*) Kstar != 1 /\ {in K^#, forall k, K \x Kstar = 'C_M[k]}, + (*6*) [/\ M`_\F != 1, M`_\F \subset M`_\sigma, M`_\sigma \subset M^`(1), + M^`(1) \proper M & nilpotent (M^`(1) / M`_\F)], + (*7*) [/\ M^`(2) \subset 'F(M), M`_\F * 'C_M(M`_\F) = 'F(M) + & K :!=: 1 -> 'F(M) \subset M^`(1)] + & (*8*) M`_\F != M`_\sigma -> + [/\ U :=: 1, normedTI 'F(M)^# G M & prime #|K| ]]]. +Proof. +have [hallU hallK _] := complU; split. +- by rewrite pcore_normal Msigma_Hall // Msigma_Hall_G. +- by have [[]] := kappa_structure maxM complU. +- have [_ defM _ _ _] := kappa_compl_context maxM complU. + have [[_ UK _ defUK]] := sdprodP defM; rewrite defUK. + have [nsKUK _ mulUK nUK _] := sdprod_context defUK. + rewrite -mulUK mulG_subG mulgA => mulMsUK /andP[nMsU nMsK] _. + rewrite (norm_joinEr nUK) mulUK; split=> //. + rewrite inE coprime_TIg /= norm_joinEr //=. + by rewrite -mulgA (normC nUK) mulgA mulMsUK !eqxx. + rewrite (pnat_coprime _ (pHall_pgroup hallU)) // -pgroupE pgroupM. + rewrite (sub_pgroup _ (pHall_pgroup hallK)) => [|p k_p]; last first. + by apply/orP; right. + by rewrite (sub_pgroup _ (pcore_pgroup _ _)) => // p s_p; apply/orP; left. + have{defM} [[defM _ _] _ _ _ _] := kappa_structure maxM complU. + have [[MsU _ defMsU] _ _ _ _] := sdprodP defM; rewrite defMsU in defM. + have [_ mulMsU _ _] := sdprodP defMsU. + by rewrite norm_joinEr // mulMsU; case/sdprod_context: defM. +- by have [] := kappa_compl_context maxM complU. +split. +- have [K1 | ntK] := eqsVneq K 1. + rewrite /Kstar K1 cent1T setIT Msigma_neq1 // setDv. + by split=> // k; rewrite inE. + have PmaxM: M \in 'M_'P by rewrite inE -(trivg_kappa maxM hallK) ntK. + have [_ [defNK _] [-> _] _ _] := Ptype_structure PmaxM hallK. + have [_ _ cKKs tiKKs] := dprodP defNK; rewrite dprodEY //; split=> // k Kk. + by have [_ _ [_ _ _ [_ _ -> // _ _] _]] := Ptype_embedding PmaxM hallK. +- have [_ _ [_ ->] _] := Fitting_structure maxM. + by have [[]] := Fcore_structure maxM. +- have [_ [-> defF _] _ sFM'] := Fitting_structure maxM. + have [_ -> _] := cprodP defF; split=> // ntK. + by rewrite sFM' // inE -(trivg_kappa maxM hallK) ntK. +move=> not_nilMs; pose q := #|Kstar|. +have solMs: solvable M`_\sigma := solvableS (pcore_sub _ _) (mmax_sol maxM). +have [D hallD] := Hall_exists q^' solMs. +have [_] := Fcore_structure maxM; case/(_ K D)=> //. +case=> P1maxM _ _ [-> _ _ _] _ _ _; split=> //. + by apply/eqP; rewrite (trivg_kappa_compl maxM complU). +by apply: contraR not_nilMs; case/nonTI_Fitting_facts=> // _ -> _. +Qed. + +(* This is a variant of B & G, Lemma 16.1(e) that better fits the Peterfalvi *) +(* definitions. *) +Lemma sdprod_FTder : M`_\sigma ><| U = M^`(FTtype M != 1%N). +Proof. +rewrite -FTtype_Fmax // (trivgFmax maxM complU). +have [[defM _ _] defM' _ _ _] := kappa_structure maxM complU. +by case: (altP eqP) defM' defM => [-> _ | _ [] //]; rewrite sdprodg1. +Qed. + +Theorem BGsummaryB : + [/\ (*1*) forall p S, p.-Sylow(U) S -> abelian S /\ 'r(S) <= 2, + (*2*) abelian <<U :&: 'A(M)>>, + (*3*) exists U0 : {group gT}, + [/\ U0 \subset U, exponent U0 = exponent U & [disjoint U0 & 'A(M)]] + & (*4*) (forall X, X \subset U -> X :!=: 1 -> 'C_(M`_\sigma)(X) != 1 -> + 'M('C(X)) = [set M]) + /\ (*5*) ('A(M) != 'A1(M) -> normedTI ('A(M) :\: 'A1(M)) G M)]. +Proof. +split. +- move=> p S sylS; have [hallU _ _] := complU; have [sUM sk'U _] := and3P hallU. + have [-> | ntS] := eqsVneq S 1; first by rewrite abelian1 rank1. + have sk'p: p \in \sigma_kappa(M)^'. + by rewrite (pnatPpi sk'U) // -p_rank_gt0 -(rank_Sylow sylS) rank_gt0. + have{sylS} sylS := subHall_Sylow hallU sk'p sylS. + have [[sSM pS _] [/= s'p _]] := (and3P sylS, norP sk'p). + rewrite (sigma'_nil_abelian maxM) ?(pi_pgroup pS) ?(pgroup_nil pS) //. + rewrite (rank_Sylow sylS) leqNgt (contra _ s'p) //; exact: alpha_sub_sigma. +- have [_ _ _ cUA_UA _] := kappa_structure maxM complU. + apply: abelianS cUA_UA; rewrite genS // -big_distrr ?setIS -?def_FTcore //=. + by apply/bigcupsP=> x A1x; rewrite (bigcup_max x) // setDE setIAC subsetIr. +- have [-> | ntU] := eqsVneq U 1. + exists 1%G; split; rewrite // disjoint_sym disjoints_subset. + by apply/bigcupsP=> x _; rewrite setDE subsetIr. + have [_ _ _ _ [// | U0 [sU0U expU0 frobU0]]] := kappa_structure maxM complU. + exists U0; split; rewrite // -setI_eq0 big_distrr /= /'A1(M) def_FTcore //. + rewrite big1 // => x A1x; apply/eqP; rewrite setIDA setD_eq0 setICA. + by rewrite (Frobenius_reg_compl frobU0) ?subsetIr. +set part4 := forall X, _; have part4holds: part4. + move=> X sXU ntX nregX. + by have [_ _] := kappa_structure maxM complU; case/(_ X). +have [[nsMsM _ _] _ [_ _ nsMsUM _ _] _ _] := BGsummaryA. +have{nsMsM nsMsUM}[[_ nMsM] [_ nMsUM]] := (andP nsMsM, andP nsMsUM). +rewrite eqEsubset FTsupp1_sub // -setD_eq0 andbT; set B := _ :\: _. +have nBM: M \subset 'N(B) by rewrite normsD ?FTsupp_norm ?FTsupp1_norm. +have uniqB y (u := y.`_\sigma(M)^'): y \in B -> 'M('C[u]) = [set M]. + case/setDP=> /bigcupP[x /setD1P[ntx Ms_x] /setD1P[nty /setIP[M'y cxy]]]. + rewrite !inE nty def_FTcore //= in Ms_x * => notMs_y; set M' := M^`(_) in M'y. + have [nsMsM' _ _ _ _] := sdprod_context sdprod_FTder. + have [[sMsM' nMsM'] sM'M]:= (andP nsMsM', der_sub _ M : M' \subset M). + have hallMs := pHall_subl sMsM' sM'M (Msigma_Hall maxM). + have hallU: \sigma(M)^'.-Hall(M') U. + by rewrite -(compl_pHall _ hallMs) sdprod_compl ?sdprod_FTder. + have suM': <[u]> \subset M' by rewrite cycle_subG groupX. + have solM': solvable M' := solvableS sM'M (mmax_sol maxM). + have [z M'z suzU] := Hall_Jsub solM' hallU suM' (p_elt_constt _ _). + have Mz': z^-1 \in M by rewrite groupV (subsetP sM'M). + rewrite -(conjgK z u) -(group_inj (conjGid Mz')) -cent_cycle. + rewrite !cycleJ centJ; apply: def_uniq_mmaxJ (part4holds _ suzU _ _). + rewrite /= -cycleJ cycle_eq1 -consttJ; apply: contraNneq notMs_y. + move/constt1P; rewrite p_eltNK p_eltJ => sMy. + by rewrite (mem_normal_Hall hallMs). + rewrite -(normsP nMsM' z M'z) centJ -conjIg -(isog_eq1 (conj_isog _ _)). + by apply/trivgPn; exists x; rewrite //= inE Ms_x cent_cycle cent1C groupX. +split=> // nzB; apply/normedTI_P; rewrite setTI; split=> // a _. +case/pred0Pn=> x /andP[/= Bx]; rewrite mem_conjg => /uniqB/(def_uniq_mmaxJ a). +rewrite consttJ -normJ conjg_set1 conjgKV uniqB // => /set1_inj defM. +by rewrite -(norm_mmax maxM) inE {2}defM. +Qed. + +Let Z := K <*> Kstar. +Let Zhat := Z :\: (K :|: Kstar). + +(* We strengthened the uniqueness condition in part (4) to *) +(* 'M_\sigma(K) = [set Mstar]. *) +Theorem BGsummaryC : K :!=: 1 -> + [/\ + [/\ (*1*) abelian U /\ ~~ ('N(U) \subset M), + (*2*) [/\ cyclic Kstar, Kstar != 1, Kstar \subset M`_\F & ~~ cyclic M`_\F] + & (*3*) M`_\sigma ><| U = M^`(1) /\ Kstar \subset M^`(2)], + exists Mstar, + [/\ (*4*) [/\ Mstar \in 'M_'P, 'C_(Mstar`_\sigma)(Kstar) = K, + \kappa(Mstar).-Hall(Mstar) Kstar + & 'M_\sigma(K) = [set Mstar]], (* uniqueness *) + (*5*) {in 'E^1(Kstar), forall X, 'M('C(X)) = [set M]} + /\ {in 'E^1(K), forall Y, 'M('C(Y)) = [set Mstar]}, + (*6*) [/\ M :&: Mstar = Z, K \x Kstar = Z & cyclic Z] + & (*7*) (M \in 'M_'P2 \/ Mstar \in 'M_'P2) + /\ {in 'M_'P, forall H, gval H \in M :^: G :|: Mstar :^: G}] +& [/\ (*8*) normedTI Zhat G Z, + (*9*) let B := 'A0(M) :\: 'A(M) in + B = class_support Zhat M /\ normedTI B G M, + (*10*) U :!=: 1 -> + [/\ prime #|K|, normedTI 'F(M)^# G M & M`_\sigma \subset 'F(M)] + & (*11*) U :==: 1 -> prime #|Kstar| ]]. +Proof. +move=> ntK; have [_ hallK _] := complU. +have PmaxM: M \in 'M_'P by rewrite inE -(trivg_kappa maxM hallK) ntK. +split. +- have [_ [//|-> ->] _ _ _] := kappa_structure maxM complU. + have [-> -> -> -> ->] := Ptype_cyclics PmaxM hallK; do 2!split=> //. + have [L maxCK_L] := mmax_exists (mFT_cent_proper ntK). + have [-> | ntU] := eqsVneq U 1. + by rewrite norm1 proper_subn // mmax_proper. + have P2maxM: M \in 'M_'P2 by rewrite inE -(trivg_kappa_compl maxM complU) ntU. + have [r _ rU] := rank_witness U; have [R sylR] := Sylow_exists r U. + have ntR: R :!=: 1 by rewrite -rank_gt0 (rank_Sylow sylR) -rU rank_gt0. + have ltRG: R \proper G := mFT_pgroup_proper (pHall_pgroup sylR). + have [H maxNR_H] := mmax_exists (mFT_norm_proper ntR ltRG). + apply: contra (subset_trans (subsetIr H _)) _. + by have [_ _ _ [->]] := P2type_signalizer P2maxM complU maxCK_L sylR maxNR_H. +- have [L [PmaxL _] [uniqL []]] := Ptype_embedding PmaxM hallK. + rewrite -/Kstar -/Z -/Zhat => hallKstar _ [defK _] [cycZ defML _ _ _]. + case=> _ P2_MorL Pmax_conjMorL _; exists L. + suffices uniqMSK: 'M_\sigma(K) = [set L]. + have [_ [defNK _] [_ uniqM] _ _] := Ptype_structure PmaxM hallK. + do 2!split=> //; last by case: P2_MorL => [] [-> _]; [left | right]. + by have [_ _ cKKs tiKKs] := dprodP defNK; rewrite dprodEY. + have sKLs: K \subset L`_\sigma by rewrite -defK subsetIl. + have [X E1X]: exists X, X \in 'E^1(K) by apply/rank_geP; rewrite rank_gt0. + have sXK: X \subset K by case/nElemP: E1X => ? /pnElemP[]. + have [maxL sCXL] := mem_uniq_mmax (uniqL X E1X). + have [x defKx] := cyclicP (cyclicS (joing_subl _ _) cycZ). + have SMxL: L \in 'M_\sigma[x] by rewrite -defKx inE maxL. + have ell1x: \ell_\sigma(x) == 1%N. + by rewrite (Msigma_ell1 maxL) // !inE -cycle_eq1 -cycle_subG -defKx ntK. + apply/eqP; rewrite eq_sym eqEcard defKx sub1set SMxL cards1 leqNgt. + apply/negP=> ntSMx; have [_ [//|_ ntR _ _]] := FT_signalizer_context ell1x. + case/(_ L)=> // /sdprodP[_ _ _ tiRL]; case/negP: ntR. + rewrite -subG1 -tiRL subsetIidl -setIA setICA setISS ?pcore_sub //. + by rewrite subsetIidr /= -cent_cycle -defKx (subset_trans (centS sXK) sCXL). +split; last 1 first. +- rewrite (trivg_kappa_compl maxM complU) => P1maxM. + have [L _ [_ _ _ _ [_ [] [] //]]] := Ptype_embedding PmaxM hallK. + by rewrite inE P1maxM. +- by have [L _ [_ _ _ _ [[]]]] := Ptype_embedding PmaxM hallK. +- have [L _ [_ _ _]] := Ptype_embedding PmaxM hallK; rewrite -/Zhat -/Z. + case=> cycZ defML defCK defCKs defCZhat [[tiZhat tiZhatM _] _ _ defM] B. + have sZM: Z \subset M by rewrite -[Z]defML subsetIl. + have sZhM: Zhat \subset M by rewrite subDset setUC subsetU ?sZM. + suffices defB: B = class_support Zhat M. + split=> //; apply/normedTI_P. + rewrite setTI normsD ?FTsupp_norm ?FTsupp0_norm //; split=> // [|g _]. + case/andP: tiZhat => /set0Pn[z Zz] _; apply/set0Pn; exists z. + by rewrite defB mem_class_support. + rewrite defB => /pred0Pn[_ /andP[/imset2P[z x Zz Mx ->] /= Bg_zx]]. + apply/idPn; rewrite -(groupMr g (groupVr Mx)) -in_setC. + case/tiZhatM/pred0Pn; exists z; rewrite /= Zz conjsgM mem_conjgV. + by apply: subsetP Bg_zx; rewrite conjSg class_support_subG. + rewrite /B /'A0(M); set M' := M^`(_); set su := \pi(M'). + have defM': M' = M^`(1) by rewrite /M' -FTtype_Pmax ?PmaxM. + have{hallK} hallM': su.-Hall(M) M'. + by rewrite Hall_pi //= -/M' defM' (sdprod_Hall defM) (pHall_Hall hallK). + have{hallM'} hallK: su^'.-Hall(M) K. + by rewrite -(compl_pHall _ hallM') /= -/M' defM' sdprod_compl. + have su'K: su^'.-group K := pHall_pgroup hallK. + have suKs: su.-group Kstar. + by rewrite (pgroupS _ (pgroup_pi _)) ///= -/M' defM' subIset ?Msigma_der1. + apply/setP=> x; rewrite !inE; apply/andP/imset2P=> [[]| [y a]]; last first. + case/setDP=> Zy; rewrite inE => /norP[not_Ky notKs_y] Ma ->. + have My := subsetP sZM y Zy; have Mya := groupJ My Ma. + have [not_suy not_su'y]: ~~ su.-elt y /\ ~~ su^'.-elt y. + have defZ: K * Kstar = Z by rewrite -cent_joinEr ?subsetIr. + have [hallK_Z hallKs] := coprime_mulGp_Hall defZ su'K suKs. + have ns_Z := sub_abelian_normal _ (cyclic_abelian cycZ). + rewrite -(mem_normal_Hall hallKs) -?ns_Z ?joing_subr // notKs_y. + by rewrite -(mem_normal_Hall hallK_Z) -?ns_Z ?joing_subl. + rewrite Mya !p_eltJ not_suy not_su'y orbT; split=> //. + apply: contra not_suy => /bigcupP[_ _ /setD1P[_ /setIP[M'ya _]]]. + by rewrite -(p_eltJ _ y a) (mem_p_elt (pgroup_pi _)). + move/negPf=> -> /and3P[Mx not_sux not_su'x]; set y := x.`_su^'. + have syM: <[y]> \subset M by rewrite cycle_subG groupX. + have [a Ma Kya] := Hall_Jsub (mmax_sol maxM) hallK syM (p_elt_constt _ _). + have{Kya} K1ya: y ^ a \in K^#. + rewrite !inE -cycle_subG cycleJ Kya andbT -consttJ. + by apply: contraNneq not_sux; move/constt1P; rewrite p_eltNK p_eltJ. + exists (x ^ a) a^-1; rewrite ?groupV ?conjgK // 2!inE andbC negb_or. + rewrite -[Z](defCK _ K1ya) inE groupJ // cent1C -consttJ groupX ?cent1id //. + by rewrite (contra (mem_p_elt su'K)) ?(contra (mem_p_elt suKs)) ?p_eltJ. +rewrite (trivg_kappa_compl maxM complU) => notP1maxM. +have P2maxM: M \in 'M_'P2 by exact/setDP. +split; first by have [_ _ _ _ []] := Ptype_structure PmaxM hallK. + apply: contraR notP1maxM; case/nonTI_Fitting_facts=> //. + by case/setUP=> //; case/idPn; case/setDP: PmaxM. +have [<- | neqMF_Ms] := eqVneq M`_\F M`_\sigma; first exact: Fcore_sub_Fitting. +have solMs: solvable M`_\sigma := solvableS (pcore_sub _ _) (mmax_sol maxM). +have [D hallD] := Hall_exists #|Kstar|^' solMs. +by case: (Fcore_structure maxM) notP1maxM => _ /(_ K D)[] // [[->]]. +Qed. + +End SingleGroupSummaries. + +Theorem BGsummaryD M : M \in 'M -> + [/\ (*1*) {in M`_\sigma &, forall x y, y \in x ^: G -> y \in x ^: M}, + (*2*) forall g (Ms := M`_\sigma), g \notin M -> + Ms:&: M :^ g = Ms :&: Ms :^ g /\ cyclic (Ms :&: M :^ g), + (*3*) {in M`_\sigma^#, forall x, + [/\ Hall 'C[x] 'C_M[x], 'R[x] ><| 'C_M[x] = 'C[x] + & let MGx := [set Mg in M :^: G | x \in Mg] in + [transitive 'R[x], on MGx | 'Js] /\ #|'R[x]| = #|MGx| ]} + & (*4*) {in M`_\sigma^#, forall x (N := 'N[x]), ~~ ('C[x] \subset M) -> + [/\ 'M('C[x]) = [set N] /\ N`_\F = N`_\sigma, + x \in 'A(N) :\: 'A1(N) /\ N \in 'M_'F :|: 'M_'P2, + \sigma(N)^'.-Hall(N) (M :&: N) + & N \in 'M_'P2 -> + [/\ M \in 'M_'F, + exists2 E, [Frobenius M = M`_\sigma ><| gval E] & cyclic E + & ~~ normedTI (M`_\F)^# G M]]}]. +Proof. +move=> maxM; have [[U K] /= complU] := kappa_witness maxM. +have defSM: {in M`_\sigma^#, + forall x, [set Mg in M :^: G | x \in Mg] = val @: 'M_\sigma[x]}. +- move=> x /setD1P[ntx Ms_x]. + have SMxM: M \in 'M_\sigma[x] by rewrite inE maxM cycle_subG. + apply/setP=> /= Mg; apply/setIdP/imsetP=> [[] | [H]]. + case/imsetP=> g _ -> Mg_x; exists (M :^ g)%G => //=. + rewrite inE cycle_subG (mem_Hall_pcore (Msigma_Hall _)) ?mmaxJ // maxM. + by rewrite (eq_p_elt _ (sigmaJ _ _)) (mem_p_elt (pcore_pgroup _ M)). + case/setIdP=> maxH; rewrite cycle_subG => Hs_x ->. + split; last exact: subsetP (pcore_sub _ _) x Hs_x. + pose p := pdiv #[x]; have pixp: p \in \pi(#[x]) by rewrite pi_pdiv order_gt1. + apply/idPn=> /(sigma_partition maxM maxH)/(_ p). + rewrite inE /= (pnatPpi (mem_p_elt (pcore_pgroup _ _) Ms_x)) //. + by rewrite (pnatPpi (mem_p_elt (pcore_pgroup _ _) Hs_x)). +split. +- have hallMs := pHall_subl (subxx _) (subsetT _) (Msigma_Hall_G maxM). + move=> x y Ms_x Ms_y /=/imsetP[a _ def_y]; rewrite def_y in Ms_y *. + have [b /setIP[Mb _ ->]] := sigma_Hall_tame maxM hallMs Ms_x Ms_y. + exact: mem_imset. +- move=> g notMg; split. + apply/eqP; rewrite eqEsubset andbC setIS ?conjSg ?pcore_sub //=. + rewrite subsetI subsetIl -MsigmaJ. + rewrite (sub_Hall_pcore (Msigma_Hall _)) ?mmaxJ ?subsetIr //. + rewrite (eq_pgroup _ (sigmaJ _ _)). + exact: pgroupS (subsetIl _ _) (pcore_pgroup _ _). + have [E hallE] := ex_sigma_compl maxM. + by have [_ _] := sigma_compl_embedding maxM hallE; case/(_ g). +- move=> x Ms1x /=. + have [[ntx Ms_x] ell1x] := (setD1P Ms1x, Msigma_ell1 maxM Ms1x). + have [[trR oR nsRC hallR] defC] := FT_signalizer_context ell1x. + have SMxM: M \in 'M_\sigma[x] by rewrite inE maxM cycle_subG. + suffices defCx: 'R[x] ><| 'C_M[x] = 'C[x]. + split=> //; first by rewrite -(sdprod_Hall defCx). + rewrite defSM //; split; last by rewrite (card_imset _ val_inj). + apply/imsetP; exists (gval M); first exact: mem_imset. + by rewrite -(atransP trR _ SMxM) -imset_comp. + have [| SMgt1] := leqP #|'M_\sigma[x]| 1. + rewrite leq_eqVlt {2}(cardD1 M) SMxM orbF => eqSMxM. + have ->: 'R[x] = 1 by apply/eqP; rewrite trivg_card1 oR. + by rewrite sdprod1g (setIidPr _) ?cent1_sub_uniq_sigma_mmax. + have [uniqN _ _ _ defCx] := defC SMgt1. + have{defCx} [[defCx _ _ _] [_ sCxN]] := (defCx M SMxM, mem_uniq_mmax uniqN). + by rewrite -setIA (setIidPr sCxN) in defCx. +move=> x Ms1x /= not_sCM. +have [[ntx Ms_x] ell1x] := (setD1P Ms1x, Msigma_ell1 maxM Ms1x). +have SMxM: M \in 'M_\sigma[x] by rewrite inE maxM cycle_subG. +have SMgt1: #|'M_\sigma[x]| > 1. + apply: contraR not_sCM; rewrite -leqNgt leq_eqVlt {2}(cardD1 M) SMxM orbF. + by move/cent1_sub_uniq_sigma_mmax->. +have [_ [//|uniqN ntR t2Nx notP1maxN]] := FT_signalizer_context ell1x. +have [maxN sCxN] := mem_uniq_mmax uniqN. +case/(_ M SMxM)=> _ st2NsM _ ->; split=> //. +- by rewrite (Fcore_eq_Msigma maxN (notP1type_Msigma_nil _)) // -in_setU. +- split=> //; apply/setDP; split. + have [y Ry nty] := trivgPn _ ntR; have [Nsy cxy] := setIP Ry. + apply/bigcupP; exists y; first by rewrite 2!inE def_FTcore ?nty. + rewrite 3!inE ntx cent1C cxy -FTtype_Pmax //= andbT. + have Nx: x \in 'N[x] by rewrite (subsetP sCxN) ?cent1id. + case PmaxN: ('N[x] \in 'M_'P) => //. + have [KN hallKN] := Hall_exists \kappa('N[x]) (mmax_sol maxN). + have [_ _ [_ _ _ _ [_ _ _ defN]]] := Ptype_embedding PmaxN hallKN. + have hallN': \kappa('N[x])^'.-Hall('N[x]) 'N[x]^`(1). + exact/(sdprod_normal_pHallP (der_normal 1 _) hallKN). + rewrite (mem_normal_Hall hallN') ?der_normal // (sub_p_elt _ t2Nx) // => p. + by case/andP=> _; apply: contraL => /rank_kappa->. + rewrite 2!inE ntx def_FTcore //=; apply: contra ntx => Ns_x. + rewrite -(constt_p_elt (mem_p_elt (pcore_pgroup _ _) Ns_x)). + by rewrite (constt1P (sub_p_elt _ t2Nx)) // => p; case/andP. +move=> P2maxN; have [PmaxN _] := setDP P2maxN; have [_ notFmaxN] := setDP PmaxN. +have [FmaxM _ [E _]] := nonFtype_signalizer_base maxM Ms1x not_sCM notFmaxN. +case=> cycE frobM; split=> //; first by exists E. +move: SMgt1; rewrite (cardsD1 M) SMxM ltnS lt0n => /pred0Pn[My /setD1P[neqMyM]]. +move/(mem_imset val); rewrite -defSM //= => /setIdP[/imsetP[y _ defMy] My_x]. +rewrite (Fcore_eq_Msigma maxM (notP1type_Msigma_nil _)) ?FmaxM //. +apply/normedTI_P=> [[_ _ /(_ y (in_setT y))/contraR/implyP/idPn[]]]. +rewrite -{1}(norm_mmax maxM) (sameP normP eqP) -defMy neqMyM. +apply/pred0Pn; exists x; rewrite /= conjD1g !inE ntx Ms_x /= -MsigmaJ. +rewrite (mem_Hall_pcore (Msigma_Hall _)) ?mmaxJ /= -?defMy //. +by rewrite defMy (eq_p_elt _ (sigmaJ _ _)) (mem_p_elt (pcore_pgroup _ _) Ms_x). +Qed. + +Lemma mmax_transversalP : + [/\ 'M^G \subset 'M, is_transversal 'M^G (orbit 'JG G @: 'M) 'M, + {in 'M^G &, injective (fun M => M :^: G)} + & {in 'M, forall M, exists x, (M :^ x)%G \in 'M^G}]. +Proof. +have: [acts G, on 'M | 'JG] by apply/actsP=> x _ M; rewrite mmaxJ. +case/orbit_transversalP; rewrite -/mmax_transversal => -> -> injMX memMX. +split=> // [M H MX_M MX_H /= eqMH | M /memMX[x _]]; last by exists x. +have /orbitP[x Gx defH]: val H \in M :^: G by rewrite eqMH orbit_refl. +by apply/eqP; rewrite -injMX // -(group_inj defH) (mem_orbit 'JG). +Qed. + +(* We are conforming to the statement of B & G, but we defer the introduction *) +(* of 'M^G to Peterfalvi (8.17), which requires several other changes. *) +Theorem BGsummaryE : + [/\ (*1*) forall M, M \in 'M -> + #|class_support M^~~ G| = (#|M`_\sigma|.-1 * #|G : M|)%N, + (*2*) {in \pi(G), forall p, + exists2 M : {group gT}, M \in 'M & p \in \sigma(M)} + /\ {in 'M &, forall M H, + gval H \notin M :^: G -> [predI \sigma(M) & \sigma(H)] =i pred0} + & (*3*) let PG := [set class_support M^~~ G | M : {group gT} in 'M] in + [/\ partition PG (cover PG), + 'M_'P = set0 :> {set {group gT}} -> cover PG = G^# + & forall M K, M \in 'M_'P -> \kappa(M).-Hall(M) K -> + let Kstar := 'C_(M`_\sigma)(K) in + let Zhat := (K <*> Kstar) :\: (K :|: Kstar) in + partition [set class_support Zhat G; cover PG] G^#]]. +Proof. +split=> [||PG]; first exact: card_class_support_sigma. + by split=> [p /sigma_mmax_exists[M]|]; [exists M | apply: sigma_partition]. +have [noPmax | ntPmax] := eqVneq 'M_'P (set0 : {set {group gT}}). + rewrite noPmax; move/eqP in noPmax; have [partPG _] := mFT_partition gT. + have /and3P[/eqP-> _ _] := partPG noPmax; rewrite partPG //. + by split=> // M K; rewrite inE. +have [_ partZPG] := mFT_partition gT. +have partPG: partition PG (cover PG). + have [M PmaxM] := set0Pn _ ntPmax; have [maxM _] := setDP PmaxM. + have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). + have{partZPG} [/and3P[_ tiPG]] := partZPG M K PmaxM hallK. + rewrite inE => /norP[_ notPGset0] _; apply/and3P; split=> //; apply/trivIsetP. + by apply: sub_in2 (trivIsetP tiPG) => C; apply: setU1r. +split=> // [noPmax | M K PmaxM hallK]; first by case/eqP: ntPmax. +have [/=] := partZPG M K PmaxM hallK; rewrite -/PG; set Z := class_support _ G. +case/and3P=> /eqP defG1 tiZPG; rewrite 2!inE => /norP[nzZ _] notPGZ. +have [_ tiPG nzPG] := and3P partPG; have [maxM _] := setDP PmaxM. +rewrite /cover big_setU1 //= -/(cover PG) in defG1. +rewrite /trivIset /cover !big_setU1 //= (eqnP tiPG) -/(cover PG) in tiZPG. +have tiZ_PG: Z :&: cover PG = set0. + by apply/eqP; rewrite setI_eq0 -leq_card_setU eq_sym. +have notUPGZ: Z \notin [set cover PG]. + by rewrite inE; apply: contraNneq nzZ => defZ; rewrite -tiZ_PG -defZ setIid. +rewrite /partition /trivIset /(cover _) !big_setU1 // !big_set1 /= -defG1. +rewrite eqxx tiZPG !inE negb_or nzZ /= eq_sym; apply: contraNneq nzPG => PG0. +apply/imsetP; exists M => //; apply/eqP; rewrite eq_sym -subset0 -PG0. +by rewrite (bigcup_max (class_support M^~~ G)) //; apply: mem_imset. +Qed. + +Let typePfacts M (H := M`_\F) U W1 W2 W (defW : W1 \x W2 = W) : + M \in 'M -> of_typeP M U defW -> + [/\ M \in 'M_'P, \kappa(M).-Hall(M) W1, 'C_H(W1) = W2, + (M \in 'M_'P1) = (U :==: 1) || ('N(U) \subset M) + & let Ms := M`_\sigma in + Ms = M^`(1) -> (H == Ms) = (U :==: 1) /\ abelian (Ms / H) = abelian U]. +Proof. +move=> maxM []{defW}; move: W1 W2 => K Ks [cycK hallK ntK defM] /=. +have [[_ /= sHMs sMsM' _] _] := Fcore_structure maxM. +rewrite -/H in sHMs * => [] [nilU sUM' nUK defM'] _ [_ ntKs sKsH _ prKsK _]. +have [_ sKM mulM'K _ tiM'K] := sdprod_context defM. +have defKs: 'C_H(K) = Ks. + have [[x defK] sHM'] := (cyclicP cycK, subset_trans sHMs sMsM'). + have K1x: x \in K^# by rewrite !inE -cycle_eq1 -cycle_subG -defK subxx andbT. + by rewrite -(setIidPl sHM') -setIA defK cent_cycle prKsK // (setIidPr _). +have{hallK} kK: \kappa(M).-group K. + apply: sub_pgroup (pgroup_pi K) => p piKp. + rewrite unlock 4!inE -!andb_orr orNb andbT -andbA. + have [X EpX]: exists X, X \in 'E_p^1(K). + by apply/p_rank_geP; rewrite p_rank_gt0. + have [sXK abelX dimX] := pnElemP EpX; have [pX _] := andP abelX. + have sXM := subset_trans sXK sKM. + have ->: p \in \sigma(M)^'. + apply: contra (nt_pnElem EpX isT) => sp. + rewrite -subG1 -tiM'K subsetI (subset_trans _ sMsM') //. + by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) ?(pi_pgroup pX). + have ->: 'r_p(M) == 1%N. + rewrite -(p_rank_Hall (Hall_pi hallK)) // eqn_leq p_rank_gt0 piKp andbT. + apply: leq_trans (p_rank_le_rank p K) _. + by rewrite -abelian_rank1_cyclic ?cyclic_abelian. + apply/existsP; exists X; rewrite 2!inE sXM abelX dimX /=. + by rewrite (subG1_contra _ ntKs) // -defKs setISS ?centS. +have PmaxM : M \in 'M_'P. + apply/PtypeP; split=> //; exists (pdiv #|K|). + by rewrite (pnatPpi kK) // pi_pdiv cardG_gt1. +have hallK: \kappa(M).-Hall(M) K. + rewrite pHallE sKM -(eqn_pmul2l (cardG_gt0 M^`(1))) (sdprod_card defM). + have [K1 hallK1] := Hall_exists \kappa(M) (mmax_sol maxM). + have [_ _ [_ _ _ _ [_ _ _ defM1]]] := Ptype_embedding PmaxM hallK1. + by rewrite -(card_Hall hallK1) /= (sdprod_card defM1). +split=> // [|->]; first set Ms := M`_\sigma; last first. + rewrite trivg_card_le1 -(@leq_pmul2l #|H|) ?cardG_gt0 // muln1. + split; first by rewrite (sdprod_card defM') eqEcard (subset_trans sHMs). + have [_ mulHU nHU tiHU] := sdprodP defM'. + by rewrite -mulHU quotientMidl (isog_abelian (quotient_isog nHU tiHU)). +have [U1 | /= ntU] := altP eqP. + rewrite inE PmaxM -{2}mulM'K /= -defM' U1 sdprodg1 pgroupM. + have sH: \sigma(M).-group H := pgroupS sHMs (pcore_pgroup _ _). + rewrite (sub_pgroup _ sH) => [|p sMp]; last by rewrite inE /= sMp. + by rewrite (sub_pgroup _ kK) // => p kMp; rewrite inE /= kMp orbT. +have [P1maxM | notP1maxM] := boolP (M \in _). + have defMs: Ms = M^`(1). + have [U1 complU1] := ex_kappa_compl maxM hallK. + have [_ [//|<- _] _ _ _] := kappa_structure maxM complU1. + by case: (P1maxP maxM complU1 P1maxM) => _; move/eqP->; rewrite sdprodg1. + pose p := pdiv #|U|; have piUp: p \in \pi(U) by rewrite pi_pdiv cardG_gt1. + have hallU: \pi(H)^'.-Hall(M^`(1)) U. + have sHM': H \subset M^`(1) by rewrite -defMs. + have hallH := pHall_subl sHM' (der_sub 1 M) (Fcore_Hall M). + by rewrite -(compl_pHall _ hallH) ?sdprod_compl. + have piMs_p: p \in \pi(Ms) by rewrite defMs (piSg sUM'). + have{piMs_p} sMp: p \in \sigma(M) := pnatPpi (pcore_pgroup _ _) piMs_p. + have sylP: p.-Sylow(M^`(1)) 'O_p(U). + apply: (subHall_Sylow hallU (pnatPpi (pHall_pgroup hallU) piUp)). + exact: nilpotent_pcore_Hall nilU. + rewrite (subset_trans (char_norms (pcore_char p U))) //. + rewrite (norm_sigma_Sylow sMp) //. + by rewrite (subHall_Sylow (Msigma_Hall maxM)) //= -/Ms defMs. +suffices complU: kappa_complement M U K. + by symmetry; apply/idPn; have [[[]]] := BGsummaryC maxM complU ntK. +split=> //; last by rewrite -norm_joinEr ?groupP. +rewrite pHallE (subset_trans _ (der_sub 1 M)) //=. +rewrite -(@eqn_pmul2l #|H|) ?cardG_gt0 // (sdprod_card defM'). +have [U1 complU1] := ex_kappa_compl maxM hallK. +have [hallU1 _ _] := complU1; rewrite -(card_Hall hallU1). +have [_ [// | defM'1 _] _ _ _] := kappa_structure maxM complU1. +rewrite [H](Fcore_eq_Msigma maxM _) ?(sdprod_card defM'1) //. +by rewrite notP1type_Msigma_nil // in_setD notP1maxM PmaxM orbT. +Qed. + +(* This is B & G, Lemma 16.1. *) +Lemma FTtypeP i M : M \in 'M -> reflect (FTtype_spec i M) (FTtype M == i). +Proof. +move=> maxM; pose Ms := M`_\sigma; pose M' := M^`(1); pose H := M`_\F. +have [[ntH sHMs sMsM' _] _] := Fcore_structure maxM. +apply: (iffP eqP) => [<- | ]; last first. + case: i => [// | [[U [[[_ _ defM] _ [U0 [sU0U expU0 frobM]]] _]] | ]]. + apply/eqP; rewrite -FTtype_Fmax //; apply: wlog_neg => notFmaxM. + have PmaxM: M \in 'M_'P by apply/setDP. + apply/FtypeP; split=> // p; apply/idP=> kp. + have [X EpX]: exists X, X \in 'E_p^1(U0). + apply/p_rank_geP; rewrite p_rank_gt0 -pi_of_exponent expU0 pi_of_exponent. + have: p \in \pi(M) by rewrite kappa_pi. + rewrite /= -(sdprod_card defM) pi_ofM ?cardG_gt0 //; case/orP=> // Fk. + have [[_ sMFMs _ _] _] := Fcore_structure maxM. + case/negP: (kappa_sigma' kp). + exact: pnatPpi (pcore_pgroup _ _) (piSg sMFMs Fk). + have [[sXU0 abelX _] ntX] := (pnElemP EpX, nt_pnElem EpX isT). + have kX := pi_pgroup (abelem_pgroup abelX) kp. + have [_ sUM _ _ _] := sdprod_context defM. + have sXM := subset_trans sXU0 (subset_trans sU0U sUM). + have [K hallK sXK] := Hall_superset (mmax_sol maxM) sXM kX. + have [ntKs _ _ sKsMF _] := Ptype_cyclics PmaxM hallK; case/negP: ntKs. + rewrite -subG1 -(cent_semiregular (Frobenius_reg_ker frobM) sXU0 ntX). + by rewrite subsetI sKsMF subIset // centS ?orbT. + case=> [[U W K Ks defW [[PtypeM ntU _ _] _ not_sNUM _ _]] | ]. + apply/eqP; rewrite -FTtype_P2max // inE andbC. + by have [-> _ _ -> _] := typePfacts maxM PtypeM; rewrite negb_or ntU. + case=> [[U W K Ks defW [[PtypeM ntU _ _] cUU sNUM]] | ]. + have [_ _ _] := typePfacts maxM PtypeM. + rewrite (negPf ntU) sNUM FTtype_P1max // cUU /FTtype -/Ms -/M' -/H. + by case: ifP => // _; case: (Ms =P M') => // -> _ [//|-> ->]. + case=> [[U W K Ks defW [[PtypeM ntU _ _] not_cUU sNUM]] | ]. + have [_ _ _] := typePfacts maxM PtypeM. + rewrite (negPf ntU) (negPf not_cUU) sNUM FTtype_P1max // /FTtype -/Ms -/M'. + by case: ifP => // _; case: (Ms =P M') => // -> _ [//|-> ->]. + case=> // [[U K Ks W defW [[PtypeM U_1] _]]]. + have [_ _ _] := typePfacts maxM PtypeM. + rewrite U_1 eqxx FTtype_P1max //= /FTtype -/Ms -/M' -/H. + by case: ifP => // _; case: (Ms =P M') => // -> _ [//|-> _]. +have [[U K] /= complU] := kappa_witness maxM; have [hallU hallK _] := complU. +have [K1 | ntK] := eqsVneq K 1. + have FmaxM: M \in 'M_'F by rewrite -(trivg_kappa maxM hallK) K1. + have ->: FTtype M = 1%N by apply/eqP; rewrite -FTtype_Fmax. + have ntU: U :!=: 1 by case/(FmaxP maxM complU): FmaxM. + have defH: H = Ms. + by apply/Fcore_eq_Msigma; rewrite // notP1type_Msigma_nil ?FmaxM. + have defM: H ><| U = M. + by have [_] := kappa_compl_context maxM complU; rewrite defH K1 sdprodg1. + exists U; split. + have [_ _ _ cU1U1 exU0] := kappa_structure maxM complU. + split=> //; last by rewrite -/Ms -defH in exU0; exact: exU0. + exists [group of <<\bigcup_(x in (M`_\sigma)^#) 'C_U[x]>>]. + split=> //= [|x Hx]; last by rewrite sub_gen //= -/Ms -defH (bigcup_max x). + rewrite -big_distrr /= /normal gen_subG subsetIl. + rewrite norms_gen ?normsI ?normG //; apply/subsetP=> u Uu. + rewrite inE sub_conjg; apply/bigcupsP=> x Msx. + rewrite -sub_conjg -normJ conjg_set1 (bigcup_max (x ^ u)) ?memJ_norm //. + by rewrite normD1 (subsetP (gFnorm _ _)) // (subsetP (pHall_sub hallU)). + have [|] := boolP [forall (y | y \notin M), 'F(M) :&: 'F(M) :^ y == 1]. + move/forall_inP=> TI_F; constructor 1; apply/normedTI_P. + rewrite setD_eq0 subG1 mmax_Fcore_neq1 // setTI normD1 gFnorm. + split=> // x _; apply: contraR => /TI_F/eqP tiFx. + rewrite -setI_eq0 conjD1g -setDIl setD_eq0 -set1gE -tiFx. + by rewrite setISS ?conjSg ?Fcore_sub_Fitting. + rewrite negb_forall_in => /exists_inP[y notMy ntX]. + have [_ _ _ _] := nonTI_Fitting_structure maxM notMy ntX. + case=> [[] | [_]]; first by constructor 2. + move: #|_| => p; set P := 'O_p(H); rewrite /= -/H => not_cPP cycHp'. + case=> [expU | [_]]; [constructor 3 | by rewrite 2!inE FmaxM]. + split=> [q /expU | ]. + have [_ <- nHU tiHU] := sdprodP defM. + by rewrite quotientMidl -(exponent_isog (quotient_isog _ _)). + have sylP: p.-Sylow(H) P := nilpotent_pcore_Hall _ (Fcore_nil M). + have ntP: P != 1 by apply: contraNneq not_cPP => ->; exact: abelian1. + by exists p; rewrite // -p_rank_gt0 -(rank_Sylow sylP) rank_gt0. +have PmaxM: M \in 'M_'P by rewrite inE -(trivg_kappa maxM hallK) ntK. +have [Mstar _ [_ _ _ [cycW _ _ _ _]]] := Ptype_embedding PmaxM hallK. +case=> [[tiV _ _] _ _ defM {Mstar}]. +have [_ [_ cycK] [_ nUK _ _ _] _] := BGsummaryA maxM complU; rewrite -/H. +case=> [[ntKs defCMK] [_ _ _ _ nilM'H] [sM''F defF /(_ ntK)sFM'] types34]. +have hallK_M := pHall_Hall hallK. +have [/= [[cUU not_sNUM]]] := BGsummaryC maxM complU ntK; rewrite -/H -/M' -/Ms. +case=> cycKs _ sKsH not_cycH [defM' sKsM''] _ [_ _ type2 _]. +pose Ks := 'C_H(K); pose W := K <*> Ks; pose V := W :\: (K :|: Ks). +have defKs: 'C_Ms(K) = Ks by rewrite -(setIidPr sKsH) setIA (setIidPl sHMs). +rewrite {}defKs -/W -/V in ntKs tiV cycW cycKs sKsM'' sKsH defCMK. +have{defCMK} prM'K: {in K^#, forall k, 'C_M'[k] = Ks}. + have sKsM': Ks \subset M' := subset_trans sKsM'' (der_sub 1 _). + move=> k; move/defCMK=> defW; have:= dprod_modr defW sKsM'. + have [_ _ _ ->] := sdprodP defM; rewrite dprod1g. + by rewrite setIA (setIidPl (der_sub 1 M)). +have [sHM' nsM'M] := (subset_trans sHMs sMsM', der_normal 1 M : M' <| M). +have hallM': \kappa(M)^'.-Hall(M) M' by apply/(sdprod_normal_pHallP _ hallK). +have [sM'M k'M' _] := and3P hallM'. +have hallH_M': \pi(H).-Hall(M') H := pHall_subl sHM' sM'M (Fcore_Hall M). +have nsHM' := normalS sHM' sM'M (Fcore_normal M). +have defW: K \x Ks = W. + rewrite dprodEY ?subsetIr //= setIC; apply/trivgP. + by have [_ _ _ <-] := sdprodP defM; rewrite setSI ?subIset ?sHM'. +have [Ueq1 | ntU] := eqsVneq U 1; last first. + have P2maxM: M \in 'M_'P2 by rewrite inE -(trivg_kappa_compl maxM complU) ntU. + have ->: FTtype M = 2 by apply/eqP; rewrite -FTtype_P2max. + have defH: H = Ms. + by apply/Fcore_eq_Msigma; rewrite // notP1type_Msigma_nil ?P2maxM ?orbT. + have [//|pr_K tiFM _] := type2; rewrite -defH in defM'. + have [_ sUM' _ _ _] := sdprod_context defM'. + have MtypeP: of_typeP M U defW by split; rewrite // abelian_nil. + have defM'F: M'`_\F = H. + apply/eqP; rewrite eqEsubset (Fcore_max hallH_M') ?Fcore_nil // andbT. + rewrite (Fcore_max (subHall_Hall hallM' _ (Fcore_Hall _))) ?Fcore_nil //. + by move=> p piM'Fp; apply: pnatPpi k'M' (piSg (Fcore_sub _) piM'Fp). + exact: char_normal_trans (Fcore_char _) nsM'M. + exists U _ K _ defW; split=> //; split; first by rewrite defM'F. + by exists U; split=> // x _; apply: subsetIl. + have [_ _ _ _ /(_ ntU)] := kappa_structure maxM complU. + by rewrite -/Ms -defH -defM'F. +have P1maxM: M \in 'M_'P1 by rewrite -(trivg_kappa_compl maxM complU) Ueq1. +have: 2 < FTtype M <= 5 by rewrite -FTtype_P1max. +rewrite /FTtype -/H -/Ms; case: ifP => // _; case: eqP => //= defMs _. +have [Y hallY nYK]: exists2 Y, \pi(H)^'.-Hall(M') (gval Y) & K \subset 'N(Y). + apply: coprime_Hall_exists; first by case/sdprodP: defM. + by rewrite (coprime_sdprod_Hall_l defM) (pHall_Hall hallM'). + exact: solvableS sM'M (mmax_sol maxM). +have{defM'} defM': H ><| Y = M' by apply/(sdprod_normal_p'HallP _ hallY). +have MtypeP: of_typeP M Y defW. + have [_ sYM' mulHY nHY tiHY] := sdprod_context defM'. + do 2!split=> //; rewrite (isog_nil (quotient_isog nHY tiHY)). + by rewrite /= -quotientMidl mulHY. +have [_ _ _ sNYG [//| defY1 ->]] := typePfacts maxM MtypeP. +rewrite defY1; have [Y1 | ntY] := altP (Y :=P: 1); last first. + move/esym: sNYG; rewrite (negPf ntY) P1maxM /= => sNYG. + have [|_ tiFM prK] := types34; first by rewrite defY1. + by case: ifPn; exists Y _ K _ defW. +exists Y _ K _ defW; split=> //=. +have [|] := boolP [forall (y | y \notin M), 'F(M) :&: 'F(M) :^ y == 1]. + move/forall_inP=> TI_F; constructor 1; apply/normedTI_P. + rewrite setD_eq0 subG1 mmax_Fcore_neq1 // setTI normD1 gFnorm. + split=> // x _; apply: contraR => /TI_F/eqP tiFx. + rewrite -setI_eq0 conjD1g -setDIl setD_eq0 -set1gE -tiFx. + by rewrite setISS ?conjSg ?Fcore_sub_Fitting. +rewrite negb_forall_in => /exists_inP[y notMy ntX]. +have [_ _ _ _] := nonTI_Fitting_structure maxM notMy ntX. +case=> [[] | [_]]; first by case/idPn; case/setDP: PmaxM. +move: #|_| => p; set P := 'O_p(H); rewrite /= -/H => not_cPP cycHp'. +have sylP: p.-Sylow(H) P := nilpotent_pcore_Hall _ (Fcore_nil M). +have ntP: P != 1 by apply: contraNneq not_cPP => ->; exact: abelian1. +have piHp: p \in \pi(H) by rewrite -p_rank_gt0 -(rank_Sylow sylP) rank_gt0. +have defH: H = Ms by apply/eqP; rewrite defY1 Y1. +rewrite -defMs -defH in defM; have [_ <- nHU tiHU] := sdprodP defM. +rewrite quotientMidl -(card_isog (quotient_isog _ _)) //. +rewrite -(exponent_isog (quotient_isog _ _)) // exponent_cyclic //=. +case=> [K_dv_H1 | []]; [constructor 2 | constructor 3]; exists p => //. +by rewrite K_dv_H1. +Qed. + +(* This is B & G, Theorem I. *) +(* Note that the first assertion is not used in the Perterfalvi revision of *) +(* the character theory part of the proof. *) +Theorem BGsummaryI : + [/\ forall H x a, Hall G H -> nilpotent H -> x \in H -> x ^ a \in H -> + exists2 y, y \in 'N(H) & x ^ a = x ^ y + & {in 'M, forall M, FTtype M == 1%N} + \/ exists ST : {group gT} * {group gT}, let (S, T) := ST in + [/\ S \in 'M /\ T \in 'M, + exists Wi : {group gT} * {group gT}, let (W1, W2) := Wi in + let W := W1 <*> W2 in let V := W :\: (W1 :|: W2) in + (*a*) [/\ cyclic W, normedTI V G W & W1 :!=: 1 /\ W2 :!=: 1] /\ + (*b*) [/\ S^`(1) ><| W1 = S, T^`(1) ><| W2 = T & S :&: T = W], + (*c*) {in 'M, forall M, FTtype M != 1%N -> + exists x, S :^ x = M \/ T :^ x = M}, + (*d*) FTtype S == 2 \/ FTtype T == 2 + & (*e*) 1 < FTtype S <= 5 /\ 1 < FTtype T <= 5]]. +Proof. +split=> [H x a hallH nilH Hx|]. + have [M maxM sHMs] := nilpotent_Hall_sigma nilH hallH. + have{hallH} hallH := pHall_subl sHMs (subsetT _) (Hall_pi hallH). + by case/(sigma_Hall_tame maxM hallH Hx) => // y; case/setIP; exists y. +have [allFM | ] := boolP (('M : {set {group gT}}) \subset 'M_'F). + by left=> M maxM; rewrite -FTtype_Fmax // (subsetP allFM). +case/subsetPn=> S maxS notFmaxS; right. +have PmaxS: S \in 'M_'P by exact/setDP. +have [[U W1] /= complU] := kappa_witness maxS; have [_ hallW1 _] := complU. +have ntW1: W1 :!=: 1 by rewrite (trivg_kappa maxS). +have [[_ [_]]] := BGsummaryC maxS complU ntW1; set W2 := 'C_(_)(W1) => ntW2 _. +set W := W1 <*> W2; set V := W :\: _ => _ _ [T [[PmaxT defW1 hallW2 _] _]]. +case=> defST _ cycW [P2maxST PmaxST] [tiV _ _] _. +have [maxT _] := setDP PmaxT. +have [_ _ [_ _ _ _ [_ _ _ defS]]] := Ptype_embedding PmaxS hallW1. +have [_ _ [_ _ _ _ [_ _ _ defT]]] := Ptype_embedding PmaxT hallW2. +exists (S, T); split=> //; first by exists (W1, [group of W2]). +- move=> M maxM; rewrite /= -FTtype_Pmax //. + by case/PmaxST/setUP => /imsetP[x _ ->]; exists x; by [left | right]. +- by rewrite -!{1}FTtype_P2max. +rewrite !{1}(ltn_neqAle 1) -!{1}andbA !{1}FTtype_range // !{1}andbT. +by rewrite !{1}(eq_sym 1%N) -!{1}FTtype_Pmax. +Qed. + +Lemma FTsupp0_type1 M : FTtype M == 1%N -> 'A0(M) = 'A(M). +Proof. +move=> typeM; apply/setUidPl/subsetP=> x; rewrite typeM !inE => /and3P[Mx]. +by rewrite (mem_p_elt (pgroup_pi M)). +Qed. + +Lemma FTsupp0_typeP M (H := M`_\F) U W1 W2 W (defW : W1 \x W2 = W) : + M \in 'M -> of_typeP M U defW -> + let V := W :\: (W1 :|: W2) in 'A0(M) :\: 'A(M) = class_support V M. +Proof. +move: W1 W2 => K Ks in defW * => maxM MtypeP /=. +have [[_ _ ntK _] _ _ _ _] := MtypeP. +have [PmaxM hallK defKs _ _] := typePfacts maxM MtypeP. +have [[_ sHMs _ _] _] := Fcore_structure maxM. +have [V complV] := ex_kappa_compl maxM hallK. +have [[_ [_ _ sKsH _] _] _ [_ [-> _] _ _]] := BGsummaryC maxM complV ntK. +by rewrite -(setIidPr sKsH) setIA (setIidPl sHMs) defKs -(dprodWY defW). +Qed. + +(* This is the part of B & G, Theorem II that is relevant to the proof of *) +(* Peterfalvi (8.7). We drop the considerations on the set of supporting *) +(* groups, in particular (Tii)(a), but do include additional information on D *) +(* namely the fact that D is included in 'A1(M), not just 'A(M). *) +Theorem BGsummaryII M (X : {set gT}) : + M \in 'M -> X \in pred2 'A(M) 'A0(M) -> + let D := [set x in X | ~~ ('C[x] \subset M)] in + [/\ D \subset 'A1(M), (* was 'A(M) in B & G *) + (*i*) {in X, forall x a, x ^ a \in X -> exists2 y, y \in M & x ^ a = x ^ y} + & {in D, forall x (L := 'N[x]), + [/\ (*ii*) 'M('C[x]) = [set L], FTtype L \in pred2 1%N 2, + [/\ (*b*) L`_\F ><| (M :&: L) = L, + (*c*) {in X, forall y, coprime #|L`_\F| #|'C_M[y]| }, + (*d*) x \in 'A(L) :\: 'A1(L) + & (*e*) 'C_(L`_\F)[x] ><| 'C_M[x] = 'C[x]] + & (*iii*) FTtype L == 2 -> + exists2 E, [Frobenius M = M`_\F ><| gval E] & cyclic E]}]. +Proof. +move=> maxM defX. +have sA0M: 'A0(M) \subset M := subset_trans (FTsupp0_sub M) (subsetDl M 1). +have sAA0: 'A(M) \subset 'A0(M) := FTsupp_sub0 M. +have sAM: 'A(M) \subset M := subset_trans sAA0 sA0M. +without loss {defX} ->: X / X = 'A0(M). + case/pred2P: defX => ->; move/(_ _ (erefl _))=> //. + set D0 := finset _ => [[sD0A1 tameA0 signD0]] D. + have sDD0: D \subset D0 by rewrite /D /D0 !setIdE setSI. + split=> [|x Ax a Axa|x Dx]; first exact: subset_trans sDD0 sD0A1. + by apply: tameA0; exact: (subsetP sAA0). + have [/= -> -> [-> coA0L -> -> frobL]] := signD0 x (subsetP sDD0 x Dx). + by do 2![split=> //] => y Ay; rewrite coA0L // (subsetP sAA0). +move=> {X} D; pose Ms := M`_\sigma. +have tiA0A x a: x \in 'A0(M) :\: 'A(M) -> x ^ a \notin 'A(M). + rewrite 3!inE; case: (x \in _) => //= /and3P[_ notM'x _]. + apply: contra notM'x => /bigcupP[y _ /setD1P[_ /setIP[Mx _]]]. + by rewrite -(p_eltJ _ _ a) (mem_p_elt (pgroup_pi _)). +have tiA0 x a: x \in 'A0(M) :\: 'A1(M) -> x ^ a \in 'A0(M) -> a \in M. + case/setDP=> A0x notA1x A0xa. + have [Mx Mxa] := (subsetP sA0M x A0x, subsetP sA0M _ A0xa). + have [[U K] /= complU] := kappa_witness maxM. + have [Ax | notAx] := boolP (x \in 'A(M)). + have [_ _ _ [_]] := BGsummaryB maxM complU; set B := _ :\: _ => tiB. + have Bx: x \in B by apply/setDP. + have /tiB/normedTI_memJ_P: 'A(M) != 'A1(M) by apply: contraTneq Ax => ->. + case=> _ _ /(_ x) <- //; rewrite 3?inE // conjg_eq1; apply/andP; split. + apply: contra notA1x; rewrite !inE def_FTcore // => /andP[->]. + by rewrite !(mem_Hall_pcore (Msigma_Hall maxM)) // p_eltJ. + by apply: contraLR Ax => notAxa; rewrite -(conjgK a x) tiA0A // inE notAxa. + have ntK: K :!=: 1. + rewrite -(trivgFmax maxM complU) FTtype_Fmax //. + by apply: contra notAx => /FTsupp0_type1 <-. + have [_ _ [_ [_ /normedTI_memJ_P[_ _ tiB]] _ _]]:= BGsummaryC maxM complU ntK. + by rewrite -(tiB x) inE ?tiA0A ?notAx // inE notAx. +have sDA1: D \subset 'A1(M). + apply/subsetPn=> [[x /setIdP[A0x not_sCxM] notA1x]]. + case/subsetP: not_sCxM => a cxa. + by apply: (tiA0 x); [exact/setDP | rewrite /conjg -(cent1P cxa) mulKg]. +have sDMs1: D \subset Ms^# by rewrite /Ms -def_FTcore. +have [tameMs _ signM PsignM] := BGsummaryD maxM. +split=> // [x A0x a A0xa|x Dx]. + have [A1x | notA1x] := boolP (x \in 'A1(M)); last first. + by exists a; rewrite // (tiA0 x) // inE notA1x. + case/setD1P: A1x => _; rewrite def_FTcore // => Ms_x. + apply/imsetP; rewrite tameMs ?mem_imset ?inE //. + rewrite (mem_Hall_pcore (Msigma_Hall maxM)) ?(subsetP sA0M) //. + by rewrite p_eltJ (mem_p_elt (pcore_pgroup _ _) Ms_x). +have [Ms1x [_ not_sCxM]] := (subsetP sDMs1 x Dx, setIdP Dx). +have [[uniqN defNF] [ANx typeN hallMN] type2] := PsignM x Ms1x not_sCxM. +have [maxN _] := mem_uniq_mmax uniqN. +split=> //; last 1 first. +- rewrite -FTtype_P2max // => /type2[FmaxM]. + by rewrite (Fcore_eq_Msigma maxM _) // notP1type_Msigma_nil ?FmaxM. +- by rewrite !inE -FTtype_Fmax // -FTtype_P2max // -in_setU. +split=> // [|y A0y|]; rewrite defNF ?sdprod_sigma //=; last by case/signM: Ms1x. +rewrite coprime_pi' ?cardG_gt0 // -pgroupE. +rewrite (eq_p'group _ (pi_Msigma maxN)); apply: wlog_neg => not_sNx'CMy. +have ell1x := Msigma_ell1 maxM Ms1x. +have SMxM: M \in 'M_\sigma[x] by rewrite inE maxM cycle_subG; case/setD1P: Ms1x. +have MSx_gt1: #|'M_\sigma[x]| > 1. + rewrite ltn_neqAle lt0n {2}(cardD1 M) SMxM andbT eq_sym. + by apply: contra not_sCxM; move/cent1_sub_uniq_sigma_mmax->. +have [FmaxM t2'M]: M \in 'M_'F /\ \tau2(M)^'.-group M. + apply: (non_disjoint_signalizer_Frobenius ell1x MSx_gt1 SMxM). + by apply: contra not_sNx'CMy; exact: pgroupS (subsetIl _ _). +have defA0: 'A0(M) = Ms^#. + rewrite FTsupp0_type1; last by rewrite -FTtype_Fmax. + rewrite /'A(M) /'A1(M) -FTtype_Fmax // FmaxM def_FTcore //= -/Ms. + apply/setP => z; apply/bigcupP/idP=> [[t Ms1t] | Ms1z]; last first. + have [ntz Ms_z] := setD1P Ms1z. + by exists z; rewrite // 3!inE ntz cent1id (subsetP (pcore_sub _ _) z Ms_z). + case/setD1P=> ntz; case/setIP=> Mz ctz. + rewrite 2!inE ntz (mem_Hall_pcore (Msigma_Hall maxM)) //. + apply: sub_in_pnat (pnat_pi (order_gt0 z)) => p _ pi_z_p. + have szM: <[z]> \subset M by rewrite cycle_subG. + have [piMp [_ k'M]] := (piSg szM pi_z_p, setIdP FmaxM). + apply: contraR (pnatPpi k'M piMp) => s'p /=. + rewrite unlock; apply/andP; split. + have:= piMp; rewrite (partition_pi_mmax maxM) (negPf s'p) orbF. + by rewrite orbCA [p \in _](negPf (pnatPpi t2'M piMp)). + move: pi_z_p; rewrite -p_rank_gt0 /= -(setIidPr szM). + case/p_rank_geP=> P; rewrite pnElemI -setIdE => /setIdP[EpP sPz]. + apply/exists_inP; exists P => //; apply/trivgPn. + have [ntt Ms_t] := setD1P Ms1t; exists t => //. + by rewrite inE Ms_t (subsetP (centS sPz)) // cent_cycle cent1C. +move: A0y; rewrite defA0 => /setD1P[nty Ms_y]. +have sCyMs: 'C_M[y] \subset Ms. + rewrite -[Ms](setD1K (group1 _)) -subDset /= -defA0 subsetU //. + rewrite (bigcup_max y) //; first by rewrite 2!inE nty def_FTcore. + by rewrite -FTtype_Fmax ?FmaxM. +have notMGN: gval 'N[x] \notin M :^: G. + have [_ [//|_ _ t2Nx _ _]] := FT_signalizer_context ell1x. + have [ntx Ms_x] := setD1P Ms1x; have sMx := mem_p_elt (pcore_pgroup _ _) Ms_x. + apply: contra ntx => /imsetP[a _ defN]. + rewrite -order_eq1 (pnat_1 sMx (sub_p_elt _ t2Nx)) // => p. + by rewrite defN tau2J // => /andP[]. +apply: sub_pgroup (pgroupS sCyMs (pcore_pgroup _ _)) => p sMp. +by apply: contraFN (sigma_partition maxM maxN notMGN p) => sNp; apply/andP. +Qed. + +End Section16. + + diff --git a/mathcomp/odd_order/BGsection2.v b/mathcomp/odd_order/BGsection2.v new file mode 100644 index 0000000..fc5f489 --- /dev/null +++ b/mathcomp/odd_order/BGsection2.v @@ -0,0 +1,1153 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div fintype. +Require Import bigop prime binomial finset fingroup morphism perm automorphism. +Require Import quotient action gfunctor commutator gproduct. +Require Import ssralg finalg zmodp cyclic center pgroup gseries nilpotent. +Require Import sylow abelian maximal hall. +Require poly ssrint. +Require Import matrix mxalgebra mxrepresentation mxabelem. +Require Import BGsection1. + +(******************************************************************************) +(* This file covers the useful material in B & G, Section 2. This excludes *) +(* part (c) of Proposition 2.1 and part (b) of Proposition 2.2, which are not *) +(* actually used in the rest of the proof; also the rest of Proposition 2.1 *) +(* is already covered by material in file mxrepresentation.v. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Section BGsection2. + +Import GroupScope GRing.Theory FinRing.Theory poly.UnityRootTheory ssrint.IntDist. +Local Open Scope ring_scope. + +Implicit Types (F : fieldType) (gT : finGroupType) (p : nat). + +(* File mxrepresentation.v covers B & G, Proposition 2.1 as follows: *) +(* - Proposition 2.1 (a) is covered by Lemmas mx_abs_irr_cent_scalar and *) +(* cent_mx_scalar_abs_irr. *) +(* - Proposition 2.2 (b) is our definition of "absolutely irreducible", and *) +(* is thus covered by cent_mx_scalar_abs_irr and mx_abs_irrP. *) +(* - Proposition 2.2 (c) is partly covered by the construction in submodule *) +(* MatrixGenField, which extends the base field with a single element a of *) +(* K = Hom_FG(M, M), rather than all of K, thus avoiding the use of the *) +(* Wedderburn theorem on finite division rings (by the primitive element *) +(* theorem this is actually equivalent). The corresponding representation *) +(* is built by MatrixGenField.gen_repr. In B & G, Proposition 2.1(c) is *) +(* only used in case II of the proof of Theorem 3.10, which we greatly *) +(* simplify by making use of the Wielandt fixpoint formula, following *) +(* Peterfalvi (Theorem 9.1). In this formalization the limited form of *) +(* 2.1(c) is used to streamline the proof that groups of odd order are *) +(* p-stable (B & G, Appendix A.5(c)). *) + +(* This is B & G, Proposition 2.2(a), using internal isomorphims (mx_iso). *) +Proposition mx_irr_prime_index F gT (G H : {group gT}) n M (nsHG : H <| G) + (rG : mx_representation F G n) (rH := subg_repr rG (normal_sub nsHG)) : + group_closure_field F gT -> mx_irreducible rG -> cyclic (G / H)%g -> + mxsimple rH M -> {in G, forall x, mx_iso rH M (M *m rG x)} -> + mx_irreducible rH. +Proof. +move=> closedF irrG /cyclicP[Hx defGH] simM isoM; have [sHG nHG] := andP nsHG. +have [modM nzM _] := simM; pose E_H := enveloping_algebra_mx rH. +have absM f: (M *m f <= M)%MS -> {a | (a \in E_H)%MS & M *m a = M *m f}. + move=> sMf; set rM := submod_repr modM; set E_M := enveloping_algebra_mx rM. + pose u := mxvec (in_submod M (val_submod 1%:M *m f)) *m pinvmx E_M. + have EHu: (gring_mx rH u \in E_H)%MS := gring_mxP rH u. + exists (gring_mx rH u) => //; rewrite -(in_submodK sMf). + rewrite -(in_submodK (mxmodule_envelop modM EHu _)) //; congr (val_submod _). + transitivity (in_submod M M *m gring_mx rM u). + rewrite /gring_mx /= !mulmx_sum_row !linear_sum; apply: eq_bigr => i /= _. + by rewrite !linearZ /= !rowK !mxvecK -in_submodJ. + rewrite /gring_mx /= mulmxKpV ?submx_full ?mxvecK //; last first. + suffices: mx_absolutely_irreducible rM by case/andP. + by apply: closedF; exact/submod_mx_irr. + rewrite {1}[in_submod]lock in_submodE -mulmxA mulmxA -val_submodE -lock. + by rewrite mulmxA -in_submodE in_submodK. +have /morphimP[x nHx Gx defHx]: Hx \in (G / H)%g by rewrite defGH cycle_id. +have{Hx defGH defHx} defG : G :=: <[x]> <*> H. + rewrite -(quotientGK nsHG) defGH defHx -quotient_cycle //. + by rewrite joingC quotientK ?norm_joinEr // cycle_subG. +have [e def1]: exists e, 1%:M = \sum_(z in G) e z *m (M *m rG z). + apply/sub_sumsmxP; have [X sXG [<- _]] := Clifford_basis irrG simM. + by apply/sumsmx_subP=> z Xz; rewrite (sumsmx_sup z) ?(subsetP sXG). +have [phi inj_phi hom_phi defMx] := isoM x Gx. +have Mtau: (M *m (phi *m rG x^-1%g) <= M)%MS. + by rewrite mulmxA (eqmxMr _ defMx) repr_mxK. +have Mtau': (M *m (rG x *m invmx phi) <= M)%MS. + by rewrite mulmxA -(eqmxMr _ defMx) mulmxK. +have [[tau Htau defMtau] [tau' Htau' defMtau']] := (absM _ Mtau, absM _ Mtau'). +have tau'K: tau' *m tau = 1%:M. + rewrite -[tau']mul1mx def1 !mulmx_suml; apply: eq_bigr => z Gz. + have [f _ hom_f] := isoM z Gz; move/eqmxP; case/andP=> _; case/submxP=> v ->. + rewrite (mulmxA _ v) -2!mulmxA; congr (_ *m _). + rewrite -(hom_envelop_mxC hom_f) ?envelop_mxM //; congr (_ *m _). + rewrite mulmxA defMtau' -(mulmxKpV Mtau') -mulmxA defMtau (mulmxA _ M). + by rewrite mulmxKpV // !mulmxA mulmxKV // repr_mxK. +have cHtau_x: centgmx rH (tau *m rG x). + apply/centgmxP=> y Hy; have [u defMy] := submxP (mxmoduleP modM y Hy). + have Gy := subsetP sHG y Hy. + rewrite mulmxA; apply: (canRL (repr_mxKV rG Gx)). + rewrite -!mulmxA /= -!repr_mxM ?groupM ?groupV // (conjgC y) mulKVg. + rewrite -[rG y]mul1mx -{1}[tau]mul1mx def1 !mulmx_suml. + apply: eq_bigr => z Gz; have [f _ hom_f] := isoM z Gz. + move/eqmxP; case/andP=> _; case/submxP=> v ->; rewrite -!mulmxA. + congr (_ *m (_ *m _)); rewrite {v} !(mulmxA M). + rewrite -!(hom_envelop_mxC hom_f) ?envelop_mxM ?(envelop_mx_id rH) //. + congr (_ *m f); rewrite !mulmxA defMy -(mulmxA u) defMtau (mulmxA u) -defMy. + rewrite !mulmxA (hom_mxP hom_phi) // -!mulmxA; congr (M *m (_ *m _)). + by rewrite /= -!repr_mxM ?groupM ?groupV // -conjgC. + by rewrite -mem_conjg (normsP nHG). +have{cHtau_x} cGtau_x: centgmx rG (tau *m rG x). + rewrite /centgmx {1}defG join_subG cycle_subG !inE Gx /= andbC. + rewrite (subset_trans cHtau_x); last by rewrite rcent_subg subsetIr. + apply/eqP; rewrite -{2 3}[rG x]mul1mx -tau'K !mulmxA; congr (_ *m _ *m _). + case/envelop_mxP: Htau' => u ->. + rewrite !(mulmx_suml, mulmx_sumr); apply: eq_bigr => y Hy. + by rewrite -!(scalemxAl, scalemxAr) (centgmxP cHtau_x) ?mulmxA. +have{cGtau_x} [a def_tau_x]: exists a, tau *m rG x = a%:M. + by apply/is_scalar_mxP; apply: mx_abs_irr_cent_scalar cGtau_x; exact: closedF. +apply: mx_iso_simple (eqmx_iso _ _) simM; apply/eqmxP; rewrite submx1 sub1mx. +case/mx_irrP: (irrG) => _ -> //; rewrite /mxmodule {1}defG join_subG /=. +rewrite cycle_subG inE Gx andbC (subset_trans modM) ?rstabs_subg ?subsetIr //=. +rewrite -{1}[M]mulmx1 -tau'K mulmxA -mulmxA def_tau_x mul_mx_scalar. +by rewrite scalemx_sub ?(mxmodule_envelop modM Htau'). +Qed. + +(* This is B & G, Lemma 2.3. Note that this is not used in the FT proof. *) +Lemma rank_abs_irr_dvd_solvable F gT (G : {group gT}) n rG : + @mx_absolutely_irreducible F _ G n rG -> solvable G -> n %| #|G|. +Proof. +move=> absG solG. +without loss closF: F rG absG / group_closure_field F gT. + move=> IH; apply: (@group_closure_field_exists gT F) => [[F' f closF']]. + by apply: IH (map_repr f rG) _ closF'; rewrite map_mx_abs_irr. +elim: {G}_.+1 {-2}G (ltnSn #|G|) => // m IHm G leGm in n rG absG solG *. +have [G1 | ntG] := eqsVneq G 1%g. + by rewrite abelian_abs_irr ?G1 ?abelian1 // in absG; rewrite (eqP absG) dvd1n. +have [H nsHG p_pr] := sol_prime_factor_exists solG ntG. +set p := #|G : H| in p_pr. +pose sHG := normal_sub nsHG; pose rH := subg_repr rG sHG. +have irrG := mx_abs_irrW absG. +wlog [L simL _]: / exists2 L, mxsimple rH L & (L <= 1%:M)%MS. + by apply: mxsimple_exists; rewrite ?mxmodule1 //; case: irrG. +have ltHG: H \proper G. + by rewrite properEcard sHG -(Lagrange sHG) ltn_Pmulr // prime_gt1. +have dvLH: \rank L %| #|H|. + have absL: mx_absolutely_irreducible (submod_repr (mxsimple_module simL)). + by apply: closF; exact/submod_mx_irr. + apply: IHm absL (solvableS (normal_sub nsHG) solG). + by rewrite (leq_trans (proper_card ltHG)). +have [_ [x Gx H'x]] := properP ltHG. +have prGH: prime #|G / H|%g by rewrite card_quotient ?normal_norm. +wlog sH: / socleType rH by exact: socle_exists. +pose W := PackSocle (component_socle sH simL). +have card_sH: #|sH| = #|G : 'C_G[W | 'Cl]|. + rewrite -cardsT; have ->: setT = orbit 'Cl G W. + apply/eqP; rewrite eqEsubset subsetT. + have /imsetP[W' _ defW'] := Clifford_atrans irrG sH. + have WW': W' \in orbit 'Cl G W by rewrite orbit_in_sym // -defW' inE. + by rewrite defW' andbT; apply/subsetP=> W''; exact: orbit_in_trans. + rewrite orbit_stabilizer // card_in_imset //. + exact: can_in_inj (act_reprK _). +have sHcW: H \subset 'C_G[W | 'Cl]. + apply: subset_trans (subset_trans (joing_subl _ _) (Clifford_astab sH)) _. + apply/subsetP=> z; rewrite !inE => /andP[->]; apply: subset_trans. + exact: subsetT. +have [|] := prime_subgroupVti ('C_G[W | 'Cl] / H)%G prGH. + rewrite quotientSGK ?normal_norm // => cClG. + have def_sH: setT = [set W]. + apply/eqP; rewrite eq_sym eqEcard subsetT cards1 cardsT card_sH. + by rewrite -indexgI (setIidPl cClG) indexgg. + suffices L1: (L :=: 1%:M)%MS. + by rewrite L1 mxrank1 in dvLH; exact: dvdn_trans (cardSg sHG). + apply/eqmxP; rewrite submx1. + have cycH: cyclic (G / H)%g by rewrite prime_cyclic. + have [y Gy|_ _] := mx_irr_prime_index closF irrG cycH simL; last first. + by apply; rewrite ?submx1 //; case simL. + have simLy: mxsimple rH (L *m rG y) by exact: Clifford_simple. + pose Wy := PackSocle (component_socle sH simLy). + have: (L *m rG y <= Wy)%MS by rewrite PackSocleK component_mx_id. + have ->: Wy = W by apply/set1P; rewrite -def_sH inE. + by rewrite PackSocleK; exact: component_mx_iso. +rewrite (setIidPl _) ?quotientS ?subsetIl // => /trivgP. +rewrite quotient_sub1 //; last by rewrite subIset // normal_norm. +move/setIidPl; rewrite (setIidPr sHcW) /= => defH. +rewrite -(Lagrange sHG) -(Clifford_rank_components irrG W) card_sH -defH. +rewrite mulnC dvdn_pmul2r // (_ : W :=: L)%MS //; apply/eqmxP. +have sLW: (L <= W)%MS by rewrite PackSocleK component_mx_id. +rewrite andbC sLW; have [modL nzL _] := simL. +have [_ _] := (Clifford_rstabs_simple irrG W); apply=> //. +rewrite /mxmodule rstabs_subg /= -Clifford_astab1 -astabIdom -defH. +by rewrite -(rstabs_subg rG sHG). +Qed. + +(* This section covers the many parts B & G, Proposition 2.4; only the last *) +(* part (k) in used in the rest of the proof, and then only for Theorem 2.5. *) +Section QuasiRegularCyclic. + +Variables (F : fieldType) (q' h : nat). + +Local Notation q := q'.+1. +Local Notation V := 'rV[F]_q. +Local Notation E := 'M[F]_q. + +Variables (g : E) (eps : F). + +Hypothesis gh1 : g ^+ h = 1. +Hypothesis prim_eps : h.-primitive_root eps. + +Let h_gt0 := prim_order_gt0 prim_eps. +Let eps_h := prim_expr_order prim_eps. +Let eps_mod_h m := expr_mod m eps_h. +Let inj_eps : injective (fun i : 'I_h => eps ^+ i). +Proof. +move=> i j eq_ij; apply/eqP; move/eqP: eq_ij. +by rewrite (eq_prim_root_expr prim_eps) !modn_small. +Qed. + +Let inhP m : m %% h < h. Proof. by rewrite ltn_mod. Qed. +Let inh m := Ordinal (inhP m). + +Let V_ i := eigenspace g (eps ^+ i). +Let n_ i := \rank (V_ i). +Let E_ i := eigenspace (lin_mx (mulmx g^-1 \o mulmxr g)) (eps ^+ i). +Let E2_ i t := + (kermx (lin_mx (mulmxr (cokermx (V_ t)) \o mulmx (V_ i))) + :&: kermx (lin_mx (mulmx (\sum_(j < h | j != i %[mod h]) V_ j)%MS)))%MS. + +Local Notation "''V_' i" := (V_ i) (at level 8, i at level 2, format "''V_' i"). +Local Notation "''n_' i" := (n_ i) (at level 8, i at level 2, format "''n_' i"). +Local Notation "''E_' i" := (E_ i) (at level 8, i at level 2, format "''E_' i"). +Local Notation "'E_ ( i )" := (E_ i) (at level 8, only parsing). +Local Notation "e ^g" := (g^-1 *m (e *m g)) + (at level 8, format "e ^g") : ring_scope. +Local Notation "'E_ ( i , t )" := (E2_ i t) + (at level 8, format "''E_' ( i , t )"). + +Let inj_g : g \in GRing.unit. +Proof. by rewrite -(unitrX_pos _ h_gt0) gh1 unitr1. Qed. + +Let Vi_mod i : 'V_(i %% h) = 'V_i. +Proof. by rewrite /V_ eps_mod_h. Qed. + +Let g_mod i := expr_mod i gh1. + +Let EiP i e : reflect (e^g = eps ^+ i *: e) (e \in 'E_i)%MS. +Proof. +rewrite (sameP eigenspaceP eqP) mul_vec_lin -linearZ /=. +by rewrite (can_eq mxvecK); exact: eqP. +Qed. + +Let E2iP i t e : + reflect ('V_i *m e <= 'V_t /\ forall j, j != i %[mod h] -> 'V_j *m e = 0)%MS + (e \in 'E_(i, t))%MS. +Proof. +rewrite sub_capmx submxE !(sameP sub_kermxP eqP) /=. +rewrite !mul_vec_lin !mxvec_eq0 /= -submxE -submx0 sumsmxMr. +apply: (iffP andP) => [[->] | [-> Ve0]]; last first. + by split=> //; apply/sumsmx_subP=> j ne_ji; rewrite Ve0. +move/sumsmx_subP=> Ve0; split=> // j ne_ji; apply/eqP. +by rewrite -submx0 -Vi_mod (Ve0 (inh j)) //= modn_mod. +Qed. + +Let sumV := (\sum_(i < h) 'V_i)%MS. + +(* This is B & G, Proposition 2.4(a). *) +Proposition mxdirect_sum_eigenspace_cycle : (sumV :=: 1%:M)%MS /\ mxdirect sumV. +Proof. +have splitF: group_splitting_field F (Zp_group h). + move: prim_eps (abelianS (subsetT (Zp h)) (Zp_abelian _)). + by rewrite -{1}(card_Zp h_gt0); exact: primitive_root_splitting_abelian. +have F'Zh: [char F]^'.-group (Zp h). + apply/pgroupP=> p p_pr; rewrite card_Zp // => /dvdnP[d def_h]. + apply/negP=> /= charFp. + have d_gt0: d > 0 by move: h_gt0; rewrite def_h; case d. + have: eps ^+ d == 1. + rewrite -(inj_eq (fmorph_inj [rmorphism of Frobenius_aut charFp])). + by rewrite rmorph1 /= Frobenius_autE -exprM -def_h eps_h. + by rewrite -(prim_order_dvd prim_eps) gtnNdvd // def_h ltn_Pmulr // prime_gt1. +case: (ltngtP h 1) => [|h_gt1|h1]; last first; last by rewrite ltnNge h_gt0. + rewrite /sumV mxdirectE /= h1 !big_ord1; split=> //. + apply/eqmxP; rewrite submx1; apply/eigenspaceP. + by rewrite mul1mx scale1r idmxE -gh1 h1. +pose mxZ (i : 'Z_h) := g ^+ i. +have mxZ_repr: mx_repr (Zp h) mxZ. + by split=> // i j _ _; rewrite /mxZ /= {3}Zp_cast // expr_mod // exprD. +pose rZ := MxRepresentation mxZ_repr. +have ZhT: Zp h = setT by rewrite /Zp h_gt1. +have memZh: _ \in Zp h by move=> i; rewrite ZhT inE. +have def_g: g = rZ Zp1 by []. +have lin_rZ m (U : 'M_(m, q)) a: + U *m g = a *: U -> forall i, U *m rZ i%:R = (a ^+ i) *: U. +- move=> defUg i; rewrite repr_mxX //. + elim: i => [|i IHi]; first by rewrite mulmx1 scale1r. + by rewrite !exprS -scalerA mulmxA defUg -IHi scalemxAl. +rewrite mxdirect_sum_eigenspace => [|j k _ _]; last exact: inj_eps. +split=> //; apply/eqmxP; rewrite submx1. +wlog [I M /= simM <- _]: / mxsemisimple rZ 1. + exact: mx_reducible_semisimple (mxmodule1 _) (mx_Maschke rZ F'Zh) _. +apply/sumsmx_subP=> i _; have simMi := simM i; have [modMi _ _] := simMi. +set v := nz_row (M i); have nz_v: v != 0 by exact: nz_row_mxsimple simMi. +have rankMi: \rank (M i) = 1%N. + by rewrite (mxsimple_abelian_linear splitF _ simMi) //= ZhT Zp_abelian. +have defMi: (M i :=: v)%MS. + apply/eqmxP; rewrite andbC -(geq_leqif (mxrank_leqif_eq _)) ?nz_row_sub //. + by rewrite rankMi lt0n mxrank_eq0. +have [a defvg]: exists a, v *m rZ 1%R = a *: v. + by apply/sub_rVP; rewrite -defMi mxmodule_trans ?socle_module ?defMi. +have: a ^+ h - 1 == 0. + apply: contraR nz_v => nz_pZa; rewrite -(eqmx_eq0 (eqmx_scale _ nz_pZa)). + by rewrite scalerBl scale1r -lin_rZ // subr_eq0 char_Zp ?mulmx1. +rewrite subr_eq0; move/eqP; case/(prim_rootP prim_eps) => k def_a. +by rewrite defMi (sumsmx_sup k) // /V_ -def_a; exact/eigenspaceP. +Qed. + +(* This is B & G, Proposition 2.4(b). *) +Proposition rank_step_eigenspace_cycle i : 'n_ (i + h) = 'n_ i. +Proof. by rewrite /n_ -Vi_mod modnDr Vi_mod. Qed. + +Let sumE := (\sum_(it : 'I_h * 'I_h) 'E_(it.1, it.2))%MS. + +(* This is B & G, Proposition 2.4(c). *) +Proposition mxdirect_sum_proj_eigenspace_cycle : + (sumE :=: 1%:M)%MS /\ mxdirect sumE. +Proof. +have [def1V] := mxdirect_sum_eigenspace_cycle; move/mxdirect_sumsP=> dxV. +pose p (i : 'I_h) := proj_mx 'V_i (\sum_(j | j != i) 'V_j)%MS. +have def1p: 1%:M = \sum_i p i. + rewrite -[\sum_i _]mul1mx; move/eqmxP: def1V; rewrite submx1. + case/sub_sumsmxP=> u ->; rewrite mulmx_sumr; apply: eq_bigr => i _. + rewrite (bigD1 i) //= mulmxDl proj_mx_id ?submxMl ?dxV //. + rewrite proj_mx_0 ?dxV ?addr0 ?summx_sub // => j ne_ji. + by rewrite (sumsmx_sup j) ?submxMl. +split; first do [apply/eqmxP; rewrite submx1]. + apply/(@memmx_subP F _ _ q)=> A _; apply/memmx_sumsP. + pose B i t := p i *m A *m p t. + exists (fun it => B it.1 it.2) => [|[i t] /=]. + rewrite -(pair_bigA _ B) /= -[A]mul1mx def1p mulmx_suml. + by apply: eq_bigr => i _; rewrite -mulmx_sumr -def1p mulmx1. + apply/E2iP; split=> [|j ne_ji]; first by rewrite mulmxA proj_mx_sub. + rewrite 2!mulmxA -mulmxA proj_mx_0 ?dxV ?mul0mx //. + rewrite (sumsmx_sup (inh j)) ?Vi_mod //. + by rewrite (modn_small (valP i)) in ne_ji. +apply/mxdirect_sumsP=> [[i t] _] /=. +apply/eqP; rewrite -submx0; apply/(@memmx_subP F _ _ q)=> A. +rewrite sub_capmx submx0 mxvec_eq0 -submx0. +case/andP=> /E2iP[ViA Vi'A] /memmx_sumsP[B /= defA sBE]. +rewrite -[A]mul1mx -(eqmxMr A def1V) sumsmxMr (bigD1 i) //=. +rewrite big1 ?addsmx0 => [|j ne_ij]; last by rewrite Vi'A ?modn_small. +rewrite -[_ *m A]mulmx1 def1p mulmx_sumr (bigD1 t) //=. +rewrite big1 ?addr0 => [|u ne_ut]; last first. + by rewrite proj_mx_0 ?dxV ?(sumsmx_sup t) // eq_sym. +rewrite {A ViA Vi'A}defA mulmx_sumr mulmx_suml summx_sub // => [[j u]]. +case/E2iP: (sBE (j, u)); rewrite eqE /=; case: eqP => [-> sBu _ ne_ut|]. + by rewrite proj_mx_0 ?dxV ?(sumsmx_sup u). +by move/eqP=> ne_ji _ ->; rewrite ?mul0mx // eq_sym !modn_small. +Qed. + +(* This is B & G, Proposition 2.4(d). *) +Proposition rank_proj_eigenspace_cycle i t : \rank 'E_(i, t) = ('n_i * 'n_t)%N. +Proof. +have [def1V] := mxdirect_sum_eigenspace_cycle; move/mxdirect_sumsP=> dxV. +pose p (i : 'I_h) := proj_mx 'V_i (\sum_(j | j != i) 'V_j)%MS. +have def1p: 1%:M = \sum_i p i. + rewrite -[\sum_i _]mul1mx; move/eqmxP: def1V; rewrite submx1. + case/sub_sumsmxP=> u ->; rewrite mulmx_sumr; apply: eq_bigr => j _. + rewrite (bigD1 j) //= mulmxDl proj_mx_id ?submxMl ?dxV //. + rewrite proj_mx_0 ?dxV ?addr0 ?summx_sub // => k ne_kj. + by rewrite (sumsmx_sup k) ?submxMl. +move: i t => i0 t0; pose i := inh i0; pose t := inh t0. +transitivity (\rank 'E_(i, t)); first by rewrite /E2_ !Vi_mod modn_mod. +transitivity ('n_i * 'n_t)%N; last by rewrite /n_ !Vi_mod. +move: {i0 t0}i t => i t; pose Bi := row_base 'V_i; pose Bt := row_base 'V_t. +pose B := lin_mx (mulmx (p i *m pinvmx Bi) \o mulmxr Bt). +pose B' := lin_mx (mulmx Bi \o mulmxr (pinvmx Bt)). +have Bk : B *m B' = 1%:M. + have frVpK m (C : 'M[F]_(m, q)) : row_free C -> C *m pinvmx C = 1%:M. + by move/row_free_inj; apply; rewrite mul1mx mulmxKpV. + apply/row_matrixP=> k; rewrite row_mul mul_rV_lin /= rowE mx_rV_lin /= -row1. + rewrite (mulmxA _ _ Bt) -(mulmxA _ Bt) [Bt *m _]frVpK ?row_base_free //. + rewrite mulmx1 2!mulmxA proj_mx_id ?dxV ?eq_row_base //. + by rewrite frVpK ?row_base_free // mul1mx vec_mxK. +have <-: \rank B = ('n_i * 'n_t)%N by apply/eqP; apply/row_freeP; exists B'. +apply/eqP; rewrite eqn_leq !mxrankS //. + apply/row_subP=> k; rewrite rowE mul_rV_lin /=. + apply/E2iP; split=> [|j ne_ji]. + rewrite 3!mulmxA mulmx_sub ?eq_row_base //. + rewrite 2!(mulmxA 'V_j) proj_mx_0 ?dxV ?mul0mx //. + rewrite (sumsmx_sup (inh j)) ?Vi_mod //. + by rewrite (modn_small (valP i)) in ne_ji. +apply/(@memmx_subP F _ _ q) => A /E2iP[ViA Vi'A]. +apply/submxP; exists (mxvec (Bi *m A *m pinvmx Bt)); rewrite mul_vec_lin /=. +rewrite mulmxKpV; last by rewrite eq_row_base (eqmxMr _ (eq_row_base _)). +rewrite mulmxA -[p i]mul1mx mulmxKpV ?eq_row_base ?proj_mx_sub // mul1mx. +rewrite -{1}[A]mul1mx def1p mulmx_suml (bigD1 i) //= big1 ?addr0 // => j neji. +rewrite -[p j]mul1mx -(mulmxKpV (proj_mx_sub _ _ _)) -mulmxA Vi'A ?mulmx0 //. +by rewrite !modn_small. +Qed. + +(* This is B & G, Proposition 2.4(e). *) +Proposition proj_eigenspace_cycle_sub_quasi_cent i j : + ('E_(i, i + j) <= 'E_j)%MS. +Proof. +apply/(@memmx_subP F _ _ q)=> A /E2iP[ViA Vi'A]. +apply/EiP; apply: canLR (mulKmx inj_g) _; rewrite -{1}[A]mul1mx -{2}[g]mul1mx. +have: (1%:M <= sumV)%MS by have [->] := mxdirect_sum_eigenspace_cycle. +case/sub_sumsmxP=> p ->; rewrite -!mulmxA !mulmx_suml. +apply: eq_bigr=> k _; have [-> | ne_ki] := eqVneq (k : nat) (i %% h)%N. + rewrite Vi_mod -mulmxA (mulmxA _ A) (eigenspaceP ViA). + rewrite (mulmxA _ g) (eigenspaceP (submxMl _ _)). + by rewrite -!(scalemxAl, scalemxAr) scalerA mulmxA exprD. +rewrite 2!mulmxA (eigenspaceP (submxMl _ _)) -!(scalemxAr, scalemxAl). +by rewrite -(mulmxA _ 'V_k A) Vi'A ?linear0 ?mul0mx ?scaler0 // modn_small. +Qed. + +Let diagE m := + (\sum_(it : 'I_h * 'I_h | it.1 + m == it.2 %[mod h]) 'E_(it.1, it.2))%MS. + +(* This is B & G, Proposition 2.4(f). *) +Proposition diag_sum_proj_eigenspace_cycle m : + (diagE m :=: 'E_m)%MS /\ mxdirect (diagE m). +Proof. +have sub_diagE n: (diagE n <= 'E_n)%MS. + apply/sumsmx_subP=> [[i t] /= def_t]. + apply: submx_trans (proj_eigenspace_cycle_sub_quasi_cent i n). + by rewrite /E2_ -(Vi_mod (i + n)) (eqP def_t) Vi_mod. +pose sum_diagE := (\sum_(n < h) diagE n)%MS. +pose p (it : 'I_h * 'I_h) := inh (h - it.1 + it.2). +have def_diag: sum_diagE = sumE. + rewrite /sumE (partition_big p xpredT) //. + apply: eq_bigr => n _; apply: eq_bigl => [[i t]] /=. + rewrite /p -val_eqE /= -(eqn_modDl (h - i)). + by rewrite addnA subnK 1?ltnW // modnDl modn_small. +have [Efull dxE] := mxdirect_sum_proj_eigenspace_cycle. +have /mxdirect_sumsE[/= dx_diag rank_diag]: mxdirect sum_diagE. + apply/mxdirectP; rewrite /= -/sum_diagE def_diag (mxdirectP dxE) /=. + rewrite (partition_big p xpredT) //. + apply: eq_bigr => n _; apply: eq_bigl => [[i t]] /=. + symmetry; rewrite /p -val_eqE /= -(eqn_modDl (h - i)). + by rewrite addnA subnK 1?ltnW // modnDl modn_small. +have dx_sumE1: mxdirect (\sum_(i < h) 'E_i). + by apply: mxdirect_sum_eigenspace => i j _ _; exact: inj_eps. +have diag_mod n: diagE (n %% h) = diagE n. + by apply: eq_bigl=> it; rewrite modnDmr. +split; last first. + apply/mxdirectP; rewrite /= -/(diagE m) -diag_mod. + rewrite (mxdirectP (dx_diag (inh m) _)) //=. + by apply: eq_bigl=> it; rewrite modnDmr. +apply/eqmxP; rewrite sub_diagE /=. +rewrite -(capmx_idPl (_ : _ <= sumE))%MS ?Efull ?submx1 //. +rewrite -def_diag /sum_diagE (bigD1 (inh m)) //= addsmxC. +rewrite diag_mod -matrix_modr ?sub_diagE //. +rewrite ((_ :&: _ =P 0)%MS _) ?adds0mx // -submx0. +rewrite -{2}(mxdirect_sumsP dx_sumE1 (inh m)) ?capmxS //. + by rewrite /E_ eps_mod_h. +by apply/sumsmx_subP=> i ne_i_m; rewrite (sumsmx_sup i) ?sub_diagE. +Qed. + +(* This is B & G, Proposition 2.4(g). *) +Proposition rank_quasi_cent_cycle m : + \rank 'E_m = (\sum_(i < h) 'n_i * 'n_(i + m))%N. +Proof. +have [<- dx_diag] := diag_sum_proj_eigenspace_cycle m. +rewrite (mxdirectP dx_diag) /= (reindex (fun i : 'I_h => (i, inh (i + m)))) /=. + apply: eq_big => [i | i _]; first by rewrite modn_mod eqxx. + by rewrite rank_proj_eigenspace_cycle /n_ Vi_mod. +exists (@fst _ _) => // [] [i t] /=. +by rewrite !inE /= (modn_small (valP t)) => def_t; exact/eqP/andP. +Qed. + +(* This is B & G, Proposition 2.4(h). *) +Proposition diff_rank_quasi_cent_cycle m : + (2 * \rank 'E_0 = 2 * \rank 'E_m + \sum_(i < h) `|'n_i - 'n_(i + m)| ^ 2)%N. +Proof. +rewrite !rank_quasi_cent_cycle !{1}mul2n -addnn. +rewrite {1}(reindex (fun i : 'I_h => inh (i + m))) /=; last first. + exists (fun i : 'I_h => inh (i + (h - m %% h))%N) => i _. + apply: val_inj; rewrite /= modnDml -addnA addnCA -modnDml addnCA. + by rewrite subnKC 1?ltnW ?ltn_mod // modnDr modn_small. + apply: val_inj; rewrite /= modnDml -modnDmr -addnA. + by rewrite subnK 1?ltnW ?ltn_mod // modnDr modn_small. +rewrite -mul2n big_distrr -!big_split /=; apply: eq_bigr => i _. +by rewrite !addn0 (addnC (2 * _)%N) sqrn_dist addnC /n_ Vi_mod. +Qed. + +Hypothesis rankEm : forall m, m != 0 %[mod h] -> \rank 'E_0 = (\rank 'E_m).+1. + +(* This is B & G, Proposition 2.4(j). *) +Proposition rank_eigenspaces_quasi_homocyclic : + exists2 n, `|q - h * n| = 1%N & + exists i : 'I_h, [/\ `|'n_i - n| = 1%N, (q < h * n) = ('n_i < n) + & forall j, j != i -> 'n_j = n]. +Proof. +have [defV dxV] := mxdirect_sum_eigenspace_cycle. +have sum_n: (\sum_(i < h) 'n_i)%N = q by rewrite -(mxdirectP dxV) defV mxrank1. +suffices [n [i]]: exists n : nat, exists2 i : 'I_h, + `|'n_i - n| == 1%N & forall i', i' != i -> 'n_i' = n. +- move/eqP=> n_i n_i'; rewrite -{1 5}(prednK h_gt0). + rewrite -sum_n (bigD1 i) //= (eq_bigr _ n_i') sum_nat_const cardC1 card_ord. + by exists n; last exists i; rewrite ?distnDr ?ltn_add2r. +case: (leqP h 1) sum_n {defV dxV} => [|h_gt1 _]. + rewrite leq_eqVlt ltnNge h_gt0 orbF; move/eqP->; rewrite big_ord1 => n_0. + by exists q', 0 => [|i']; rewrite ?(ord1 i') // n_0 distSn. +pose dn1 i := `|'n_i - 'n_(i + 1)|. +have sum_dn1: (\sum_(0 <= i < h) dn1 i ^ 2 == 2)%N. + rewrite big_mkord -(eqn_add2l (2 * \rank 'E_1)) -diff_rank_quasi_cent_cycle. + by rewrite -mulnSr -rankEm ?modn_small. +pose diff_n := [seq i <- index_iota 0 h | dn1 i != 0%N]. +have diff_n_1: all (fun i => dn1 i == 1%N) diff_n. + apply: contraLR sum_dn1; case/allPn=> i; rewrite mem_filter. + case def_i: (dn1 i) => [|[|ni]] //=; case/splitPr=> e e' _. + by rewrite big_cat big_cons /= addnCA def_i -add2n sqrnD. +have: sorted ltn diff_n. + by rewrite (sorted_filter ltn_trans) // /index_iota subn0 iota_ltn_sorted. +have: all (ltn^~ h) diff_n. + by apply/allP=> i; rewrite mem_filter mem_index_iota; case/andP. +have: size diff_n = 2%N. + move: diff_n_1; rewrite size_filter -(eqnP sum_dn1) /diff_n. + elim: (index_iota 0 h) => [|i e IHe]; rewrite (big_nil, big_cons) //=. + by case def_i: (dn1 i) => [|[]] //=; rewrite def_i //; move/IHe->. +case def_jk: diff_n diff_n_1 => [|j [|k []]] //=; case/and3P=> dn1j dn1k _ _. +case/and3P=> lt_jh lt_kh _ /andP[lt_jk _]. +have def_n i: + i <= h -> 'n_i = if i <= j then 'n_0 else if i <= k then 'n_j.+1 else 'n_k.+1. +- elim: i => // i IHi lt_ik; have:= IHi (ltnW lt_ik); rewrite !(leq_eqVlt i). + have:= erefl (i \in diff_n); rewrite {2}def_jk !inE mem_filter mem_index_iota. + case: (i =P j) => [-> _ _ | _]; first by rewrite ltnn lt_jk. + case: (i =P k) => [-> _ _ | _]; first by rewrite ltnNge ltnW // ltnn. + by rewrite distn_eq0 lt_ik addn1; case: eqP => [->|]. +have n_j1: 'n_j.+1 = 'n_k by rewrite (def_n k (ltnW lt_kh)) leqnn leqNgt lt_jk. +have n_k1: 'n_k.+1 = 'n_0. + rewrite -(rank_step_eigenspace_cycle 0) (def_n h (leqnn h)). + by rewrite leqNgt lt_jh leqNgt lt_kh; split. +case: (leqP k j.+1) => [ | lt_j1_k]. + rewrite leq_eqVlt ltnNge lt_jk orbF; move/eqP=> def_k. + exists 'n_(k + 1); exists (Ordinal lt_kh) => [|i' ne_i'k]; first exact: dn1k. + rewrite addn1 {1}(def_n _ (ltnW (valP i'))) n_k1. + by rewrite -ltnS -def_k ltn_neqAle ne_i'k /=; case: leqP; split. +case: (leqP h.-1 (k - j)) => [le_h1_kj | lt_kj_h1]. + have k_h1: k = h.-1. + apply/eqP; rewrite eqn_leq -ltnS (prednK h_gt0) lt_kh. + exact: leq_trans (leq_subr j k). + have j0: j = 0%N. + apply/eqP; rewrite -leqn0 -(leq_add2l k) -{2}(subnK (ltnW lt_jk)). + by rewrite addn0 leq_add2r {1}k_h1. + exists 'n_(j + 1); exists (Ordinal lt_jh) => [|i' ne_i'j]; first exact: dn1j. + rewrite addn1 {1}(def_n _ (ltnW (valP i'))) j0 leqNgt lt0n -j0. + by rewrite ne_i'j -ltnS k_h1 (prednK h_gt0) (valP i'); split. +suffices: \sum_(i < h) `|'n_i - 'n_(i + 2)| ^ 2 > 2. + rewrite -(ltn_add2l (2 * \rank 'E_2)) -diff_rank_quasi_cent_cycle. + rewrite -mulnSr -rankEm ?ltnn ?modn_small //. + by rewrite -(prednK h_gt0) ltnS (leq_trans _ lt_kj_h1) // ltnS subn_gt0. +have lt_k1h: k.-1 < h by rewrite ltnW // (ltn_predK lt_jk). +rewrite (bigD1 (Ordinal lt_jh)) // (bigD1 (Ordinal lt_k1h)) /=; last first. + by rewrite -val_eqE neq_ltn /= orbC -subn1 ltn_subRL lt_j1_k. +rewrite (bigD1 (Ordinal lt_kh)) /=; last first. + by rewrite -!val_eqE !neq_ltn /= lt_jk (ltn_predK lt_jk) leqnn !orbT. +rewrite !addnA ltn_addr // !addn2 (ltn_predK lt_jk) n_k1. +rewrite (def_n j (ltnW lt_jh)) leqnn (def_n _ (ltn_trans lt_j1_k lt_kh)). +rewrite lt_j1_k -if_neg -leqNgt leqnSn n_j1. +rewrite (def_n _ (ltnW lt_k1h)) leq_pred -if_neg -ltnNge. +rewrite -subn1 ltn_subRL lt_j1_k n_j1. +suffices ->: 'n_k.+2 = 'n_k.+1. + by rewrite distnC -n_k1 -(addn1 k) -/(dn1 k) (eqP dn1k). +case: (leqP k.+2 h) => [le_k2h | ]. + by rewrite (def_n _ le_k2h) (leqNgt _ k) leqnSn n_k1 if_same. +rewrite ltnS leq_eqVlt ltnNge lt_kh orbF; move/eqP=> def_h. +rewrite -{1}def_h -add1n rank_step_eigenspace_cycle (def_n _ h_gt0). +rewrite -(subSn (ltnW lt_jk)) def_h leq_subLR in lt_kj_h1. +by rewrite -(leq_add2r k) lt_kj_h1 n_k1. +Qed. + +(* This is B & G, Proposition 2.4(k). *) +Proposition rank_eigenspaces_free_quasi_homocyclic : + q > 1 -> 'n_0 = 0%N -> h = q.+1 /\ (forall j, j != 0 %[mod h] -> 'n_j = 1%N). +Proof. +move=> q_gt1 n_0; rewrite mod0n. +have [n d_q_hn [i [n_i lt_q_hn n_i']]] := rank_eigenspaces_quasi_homocyclic. +move/eqP: d_q_hn; rewrite distn_eq1 {}lt_q_hn. +case: (eqVneq (Ordinal h_gt0) i) n_i n_i' => [<- | ne0i _ n_i']; last first. + by rewrite -(n_i' _ ne0i) n_0 /= muln0 -(subnKC q_gt1). +rewrite n_0 dist0n => -> n_i'; rewrite muln1 => /eqP->; split=> // i'. +by move/(n_i' (inh i')); rewrite /n_ Vi_mod. +Qed. + +End QuasiRegularCyclic. + +(* This is B & G, Theorem 2.5, used for Theorems 3.4 and 15.7. *) +Theorem repr_extraspecial_prime_sdprod_cycle p n gT (G P H : {group gT}) : + p.-group P -> extraspecial P -> P ><| H = G -> cyclic H -> + let h := #|H| in #|P| = (p ^ n.*2.+1)%N -> coprime p h -> + {in H^#, forall x, 'C_P[x] = 'Z(P)} -> + (h %| p ^ n + 1) || (h %| p ^ n - 1) + /\ ((h != p ^ n + 1)%N -> forall F q (rG : mx_representation F G q), + [char F]^'.-group G -> mx_faithful rG -> rfix_mx rG H != 0). +Proof. +move=> pP esP sdPH_G cycH h oPpn co_p_h primeHP. +set dvd_h_pn := _ || _; set neq_h_pn := h != _. +suffices IH F q (rG : mx_representation F G q): + [char F]^'.-group G -> mx_faithful rG -> + dvd_h_pn && (neq_h_pn ==> (rfix_mx rG H != 0)). +- split=> [|ne_h F q rG F'G ffulG]; last first. + by case/andP: (IH F q rG F'G ffulG) => _; rewrite ne_h. + pose r := pdiv #|G|.+1. + have r_pr: prime r by rewrite pdiv_prime // ltnS cardG_gt0. + have F'G: [char 'F_r]^'.-group G. + rewrite /pgroup (eq_pnat _ (eq_negn (charf_eq (char_Fp r_pr)))). + rewrite p'natE // -prime_coprime // (coprime_dvdl (pdiv_dvd _)) //. + by rewrite /coprime -addn1 gcdnC gcdnDl gcdn1. + by case/andP: (IH _ _ _ F'G (regular_mx_faithful _ _)). +move=> F'G ffulG. +without loss closF: F rG F'G ffulG / group_closure_field F gT. + move=> IH; apply: (@group_closure_field_exists gT F) => [[Fs f clFs]]. + rewrite -(map_mx_eq0 f) map_rfix_mx {}IH ?map_mx_faithful //. + by rewrite (eq_p'group _ (fmorph_char f)). +have p_pr := extraspecial_prime pP esP; have p_gt1 := prime_gt1 p_pr. +have oZp := card_center_extraspecial pP esP; have[_ prZ] := esP. +have{sdPH_G} [nsPG sHG defG nPH tiPH] := sdprod_context sdPH_G. +have sPG := normal_sub nsPG. +have coPH: coprime #|P| #|H| by rewrite oPpn coprime_pexpl. +have nsZG: 'Z(P) <| G := char_normal_trans (center_char _) nsPG. +have defCP: 'C_G(P) = 'Z(P). + apply/eqP; rewrite eqEsubset andbC setSI //=. + rewrite -(coprime_mulG_setI_norm defG) ?norms_cent ?normal_norm //=. + rewrite mul_subG // -(setD1K (group1 H)). + apply/subsetP=> x; case/setIP; case/setU1P=> [-> // | H'x]. + rewrite -sub_cent1; move/setIidPl; rewrite primeHP // => defP. + by have:= min_card_extraspecial pP esP; rewrite -defP oZp (leq_exp2l 3 1). +have F'P: [char F]^'.-group P by exact: pgroupS sPG F'G. +have F'H: [char F]^'.-group H by exact: pgroupS sHG F'G. +wlog{ffulG F'G} [irrG regZ]: q rG / mx_irreducible rG /\ rfix_mx rG 'Z(P) = 0. + move=> IH; wlog [I W /= simW defV _]: / mxsemisimple rG 1%:M. + exact: (mx_reducible_semisimple (mxmodule1 _) (mx_Maschke rG F'G)). + have [z Zz ntz]: exists2 z, z \in 'Z(P) & z != 1%g. + by apply/trivgPn; rewrite -cardG_gt1 oZp prime_gt1. + have Gz := subsetP sPG z (subsetP (center_sub P) z Zz). + case: (pickP (fun i => z \notin rstab rG (W i))) => [i ffZ | z1]; last first. + case/negP: ntz; rewrite -in_set1 (subsetP ffulG) // inE Gz /=. + apply/eqP; move/eqmxP: defV; case/andP=> _; case/sub_sumsmxP=> w ->. + rewrite mulmx_suml; apply: eq_bigr => i _. + by move/negbFE: (z1 i) => /rstab_act-> //; rewrite submxMl. + have [modW _ _] := simW i; pose rW := submod_repr modW. + rewrite -(eqmx_rstab _ (val_submod1 (W i))) -(rstab_submod modW) in ffZ. + have irrW: mx_irreducible rW by exact/submod_mx_irr. + have regZ: rfix_mx rW 'Z(P)%g = 0. + apply/eqP; apply: contraR ffZ; case/mx_irrP: irrW => _ minW /minW. + by rewrite normal_rfix_mx_module // -sub1mx inE Gz /= => /implyP/rfix_mxP->. + have ffulP: P :&: rker rW = 1%g. + apply: (TI_center_nil (pgroup_nil pP)). + by rewrite /normal subsetIl normsI ?normG ?(subset_trans _ (rker_norm _)). + rewrite /= setIC setIA (setIidPl (center_sub _)); apply: prime_TIg=> //. + by apply: contra ffZ => /subsetP->. + have cPker: rker rW \subset 'C_G(P). + rewrite subsetI rstab_sub (sameP commG1P trivgP) /= -ffulP subsetI. + rewrite commg_subl commg_subr (subset_trans sPG) ?rker_norm //. + by rewrite (subset_trans (rstab_sub _ _)) ?normal_norm. + have [->] := andP (IH _ _ (conj irrW regZ)); case: (neq_h_pn) => //. + apply: contra; rewrite (eqmx_eq0 (rfix_submod modW sHG)) => /eqP->. + by rewrite capmx0 linear0. +pose rP := subg_repr rG sPG; pose rH := subg_repr rG sHG. +wlog [M simM _]: / exists2 M, mxsimple rP M & (M <= 1%:M)%MS. + by apply: (mxsimple_exists (mxmodule1 _)); last case irrG. +have{M simM irrG regZ F'P} [irrP def_q]: mx_irreducible rP /\ q = (p ^ n)%N. + have [modM nzM _]:= simM. + have [] := faithful_repr_extraspecial _ _ oPpn _ _ simM => // [|<- isoM]. + apply/eqP; apply: (TI_center_nil (pgroup_nil pP)). + rewrite /= -(eqmx_rstab _ (val_submod1 M)) -(rstab_submod modM). + exact: rker_normal. + rewrite setIC prime_TIg //=; apply: contra nzM => cMZ. + rewrite -submx0 -regZ; apply/rfix_mxP=> z; move/(subsetP cMZ)=> cMz. + by rewrite (rstab_act cMz). + suffices irrP: mx_irreducible rP. + by split=> //; apply/eqP; rewrite eq_sym; case/mx_irrP: irrP => _; exact. + apply: (@mx_irr_prime_index F _ G P _ M nsPG) => // [|x Gx]. + by rewrite -defG quotientMidl quotient_cyclic. + rewrite (bool_irrelevance (normal_sub nsPG) sPG). + apply: isoM; first exact: (@Clifford_simple _ _ _ _ nsPG). + have cZx: x \in 'C_G('Z(P)). + rewrite (setIidPl _) // -defG mulG_subG centsC subsetIr. + rewrite -(setD1K (group1 H)) subUset sub1G /=. + by apply/subsetP=> y H'y; rewrite -sub_cent1 -(primeHP y H'y) subsetIr. + by have [f] := Clifford_iso nsZG rG M cZx; exists f. +pose E_P := enveloping_algebra_mx rP; have{irrP} absP := closF P _ _ irrP. +have [q_gt0 EPfull]: q > 0 /\ (1%:M <= E_P)%MS by apply/andP; rewrite sub1mx. +pose Z := 'Z(P); have [sZP nZP] := andP (center_normal P : Z <| P). +have nHZ: H \subset 'N(Z) := subset_trans sHG (normal_norm nsZG). +pose clPqH := [set Zx ^: (H / Z) | Zx in P / Z]%g. +pose b (ZxH : {set coset_of Z}) := repr (repr ZxH). +have Pb ZxH: ZxH \in clPqH -> b ZxH \in P. + case/imsetP=> Zx P_Zx ->{ZxH}. + rewrite -(quotientGK (center_normal P)) /= -/Z inE repr_coset_norm /=. + rewrite inE coset_reprK; apply: subsetP (mem_repr _ (class_refl _ _)). + rewrite -class_support_set1l class_support_sub_norm ?sub1set //. + by rewrite quotient_norms. +have{primeHP coPH} card_clPqH ZxH: ZxH \in clPqH^# -> #|ZxH| = #|H|. + case/setD1P=> ntZxH P_ZxH. + case/imsetP: P_ZxH ntZxH => Zx P_Zx ->{ZxH}; rewrite classG_eq1 => ntZx. + rewrite -index_cent1 ['C__[_]](trivgP _). + rewrite indexg1 card_quotient // -indexgI setICA setIA tiPH. + by rewrite (setIidPl (sub1G _)) indexg1. + apply/subsetP=> Zy => /setIP[/morphimP[y Ny]]; rewrite -(setD1K (group1 H)). + case/setU1P=> [-> | Hy] ->{Zy} cZxy; first by rewrite morph1 set11. + have: Zx \in 'C_(P / Z)(<[y]> / Z). + by rewrite inE P_Zx quotient_cycle // cent_cycle cent1C. + case/idPn; rewrite -coprime_quotient_cent ?cycle_subG ?(pgroup_sol pP) //. + by rewrite /= cent_cycle primeHP // trivg_quotient inE. + by apply: coprimegS coPH; rewrite cycle_subG; case/setD1P: Hy. +pose B x := \matrix_(i < #|H|) mxvec (rP (x ^ enum_val i)%g). +have{E_P EPfull absP} sumB: (\sum_(ZxH in clPqH) <<B (b ZxH)>> :=: 1%:M)%MS. + apply/eqmxP; rewrite submx1 (submx_trans EPfull) //. + apply/row_subP=> ix; set x := enum_val ix; pose ZxH := coset Z x ^: (H / Z)%g. + have Px: x \in P by [rewrite enum_valP]; have nZx := subsetP nZP _ Px. + have P_ZxH: ZxH \in clPqH by apply: mem_imset; rewrite mem_quotient. + have Pbx := Pb _ P_ZxH; have nZbx := subsetP nZP _ Pbx. + rewrite rowK (sumsmx_sup ZxH) {P_ZxH}// genmxE -/x. + have: coset Z x \in coset Z (b ZxH) ^: (H / Z)%g. + by rewrite class_sym coset_reprK (mem_repr _ (class_refl _ _)). + case/imsetP=> _ /morphimP[y Ny Hy ->]. + rewrite -morphJ //; case/kercoset_rcoset; rewrite ?groupJ // => z Zz ->. + have [Pz cPz] := setIP Zz; rewrite repr_mxM ?memJ_norm ?(subsetP nPH) //. + have [a ->]: exists a, rP z = a%:M. + apply/is_scalar_mxP; apply: (mx_abs_irr_cent_scalar absP). + by apply/centgmxP=> t Pt; rewrite -!repr_mxM ?(centP cPz). + rewrite mul_scalar_mx linearZ scalemx_sub //. + by rewrite (eq_row_sub (gring_index H y)) // rowK gring_indexK. +have{card_clPqH} Bfree_if ZxH: + ZxH \in clPqH^# -> \rank <<B (b ZxH)>> <= #|ZxH| ?= iff row_free (B (b ZxH)). +- by move=> P_ZxH; rewrite genmxE card_clPqH // /leqif rank_leq_row. +have B1_if: \rank <<B (b 1%g)>> <= 1 ?= iff (<<B (b 1%g)>> == mxvec 1%:M)%MS. + have r1: \rank (mxvec 1%:M : 'rV[F]_(q ^ 2)) = 1%N. + by rewrite rank_rV mxvec_eq0 -mxrank_eq0 mxrank1 -lt0n q_gt0. + rewrite -{1}r1; apply: mxrank_leqif_eq; rewrite genmxE. + have ->: b 1%g = 1%g by rewrite /b repr_set1 repr_coset1. + by apply/row_subP=> i; rewrite rowK conj1g repr_mx1. +have rankEP: \rank (1%:M : 'A[F]_q) = (\sum_(ZxH in clPqH) #|ZxH|)%N. + rewrite acts_sum_card_orbit ?astabsJ ?quotient_norms // card_quotient //. + rewrite mxrank1 -divgS // -mulnn oPpn oZp expnS -muln2 expnM -def_q. + by rewrite mulKn // ltnW. +have cl1: 1%g \in clPqH by apply/imsetP; exists 1%g; rewrite ?group1 ?class1G. +have{B1_if Bfree_if}:= leqif_add B1_if (leqif_sum Bfree_if). +case/(leqif_trans (mxrank_sum_leqif _)) => _ /=. +rewrite -{1}(big_setD1 _ cl1) sumB {}rankEP (big_setD1 1%g) // cards1 eqxx. +case/esym/and3P=> dxB /eqmxP defB1 /forall_inP/= Bfree. +have [yg defH] := cyclicP cycH; pose g := rG yg. +have Hxg: yg \in H by [rewrite defH cycle_id]; have Gyg := subsetP sHG _ Hxg. +pose gE : 'A_q := lin_mx (mulmx (invmx g) \o mulmxr g). +pose yr := regular_repr F H yg. +have mulBg x: x \in P -> B x *m gE = yr *m B x. + move/(subsetP sPG)=> Gx. + apply/row_matrixP=> i; have Hi := enum_valP i; have Gi := subsetP sHG _ Hi. + rewrite 2!row_mul !rowK mul_vec_lin /= -rowE rowK gring_indexK ?groupM //. + by rewrite conjgM -repr_mxV // -!repr_mxM // ?(groupJ, groupM, groupV). +wlog sH: / irrType F H by exact: socle_exists. +have{cycH} linH: irr_degree (_ : sH) = 1%N. + exact: irr_degree_abelian (cyclic_abelian cycH). +have baseH := linear_irr_comp F'H (closF H) (linH _). +have{linH} linH (W : sH): \rank W = 1%N by rewrite baseH; exact: linH. +have [w] := cycle_repr_structure sH defH F'H (closF H). +rewrite -/h => prim_w [Wi [bijWi _ _ Wi_yg]]. +have{Wi_yg baseH} Wi_yr i: Wi i *m yr = w ^+ i *: (Wi i : 'M_h). + have /submxP[u ->]: (Wi i <= val_submod (irr_repr (Wi i) 1%g))%MS. + by rewrite repr_mx1 val_submod1 -baseH. + rewrite repr_mx1 -mulmxA -2!linearZ; congr (u *m _). + by rewrite -mul_mx_scalar -Wi_yg /= val_submodJ. +pose E_ m := eigenspace gE (w ^+ m). +have dxE: mxdirect (\sum_(m < h) E_ m)%MS. + apply: mxdirect_sum_eigenspace => m1 m2 _ _ eq_m12; apply/eqP. + by move/eqP: eq_m12; rewrite (eq_prim_root_expr prim_w) !modn_small. +pose B2 ZxH i : 'A_q := <<Wi i *m B (b ZxH)>>%MS. +pose B1 i : 'A_q := (\sum_(ZxH in clPqH^#) B2 ZxH i)%MS. +pose SB := (<<B (b 1%g)>> + \sum_i B1 i)%MS. +have{yr Wi_yr Pb mulBg} sB1E i: (B1 i <= E_ i)%MS. + apply/sumsmx_subP=> ZxH /setIdP[_]; rewrite genmxE => P_ZxH. + by apply/eigenspaceP; rewrite -mulmxA mulBg ?Pb // mulmxA Wi_yr scalemxAl. +have{bijWi sumB cl1 F'H} defSB: (SB :=: 1%:M)%MS. + apply/eqmxP; rewrite submx1 -sumB (big_setD1 _ cl1) addsmxS //=. + rewrite exchange_big sumsmxS // => ZxH _; rewrite genmxE /= -sumsmxMr_gen. + rewrite -((reindex Wi) xpredT val) /=; last by exact: onW_bij. + by rewrite -/(Socle _) (reducible_Socle1 sH (mx_Maschke _ F'H)) mul1mx. +rewrite mxdirect_addsE /= in dxB; case/and3P: dxB => _ dxB dxB1. +have{linH Bfree dxB} rankB1 i: \rank (B1 i) = #|clPqH^#|. + rewrite -sum1_card (mxdirectP _) /=. + by apply: eq_bigr => ZxH P_ZxH; rewrite genmxE mxrankMfree ?Bfree. + apply/mxdirect_sumsP=> ZxH P_ZxH. + apply/eqP; rewrite -submx0 -{2}(mxdirect_sumsP dxB _ P_ZxH) capmxS //. + by rewrite !genmxE submxMl. + by rewrite sumsmxS // => ZyH _; rewrite !genmxE submxMl. +have rankEi (i : 'I_h) : i != 0%N :> nat -> \rank (E_ i) = #|clPqH^#|. + move=> i_gt0; apply/eqP; rewrite -(rankB1 i) (mxrank_leqif_sup _) ?sB1E //. + rewrite -[E_ i]cap1mx -(cap_eqmx defSB (eqmx_refl _)) /SB. + rewrite (bigD1 i) //= (addsmxC (B1 i)) addsmxA addsmxC -matrix_modl //. + rewrite -(addsmx0 (q ^ 2) (B1 i)) addsmxS //. + rewrite capmxC -{2}(mxdirect_sumsP dxE i) // capmxS // addsmx_sub // . + rewrite (sumsmx_sup (Ordinal (cardG_gt0 H))) ?sumsmxS 1?eq_sym //. + rewrite defB1; apply/eigenspaceP; rewrite mul_vec_lin scale1r /=. + by rewrite mul1mx mulVmx ?repr_mx_unit. +have{b B defB1 rP rH sH Wi rankB1 dxB1 defSB sB1E B1 B2 dxE SB} rankE0 i: + (i : 'I_h) == 0%N :> nat -> \rank (E_ i) = #|clPqH^#|.+1. +- move=> i_eq0; rewrite -[E_ i]cap1mx -(cap_eqmx defSB (eqmx_refl _)) /SB. + rewrite (bigD1 i) // addsmxA -matrix_modl; last first. + rewrite addsmx_sub // sB1E andbT defB1; apply/eigenspaceP. + by rewrite mul_vec_lin (eqP i_eq0) scale1r /= mul1mx mulVmx ?repr_mx_unit. + rewrite (((_ :&: _)%MS =P 0) _). + rewrite addsmx0 mxrank_disjoint_sum /=. + by rewrite defB1 rank_rV rankB1 mxvec_eq0 -mxrank_eq0 mxrank1 -lt0n q_gt0. + apply/eqP; rewrite -submx0 -(eqP dxB1) capmxS // sumsmxS // => ZxH _. + by rewrite !genmxE ?submxMl. + by rewrite -submx0 capmxC /= -{2}(mxdirect_sumsP dxE i) // capmxS ?sumsmxS. +have{clPqH rankE0 rankEi} (m): + m != 0 %[mod h] -> \rank (E_ 0%N) = (\rank (E_ m)).+1. +- move=> nz_m; rewrite (rankE0 (Ordinal (cardG_gt0 H))) //. + rewrite /E_ -(prim_expr_mod prim_w); rewrite mod0n in nz_m. + have lt_m: m %% h < h by rewrite ltn_mod ?cardG_gt0. + by rewrite (rankEi (Ordinal lt_m)). +have: q > 1. + rewrite def_q (ltn_exp2l 0) // lt0n. + apply: contraL (min_card_extraspecial pP esP). + by rewrite oPpn; move/eqP->; rewrite leq_exp2l. +rewrite {}/E_ {}/gE {}/dvd_h_pn {}/neq_h_pn -{n oPpn}def_q subn1 addn1 /=. +case: q q_gt0 => // q' _ in rG g * => q_gt1 rankE. +have gh1: g ^+ h = 1 by rewrite -repr_mxX // /h defH expg_order repr_mx1. +apply/andP; split. + have [n' def_q _]:= rank_eigenspaces_quasi_homocyclic gh1 prim_w rankE. + move/eqP: def_q; rewrite distn_eq1 eqSS. + by case: ifP => _ /eqP->; rewrite dvdn_mulr ?orbT. +apply/implyP; apply: contra => regH. +have [|-> //]:= rank_eigenspaces_free_quasi_homocyclic gh1 prim_w rankE q_gt1. +apply/eqP; rewrite mxrank_eq0 -submx0 -(eqP regH). +apply/rV_subP=> v /eigenspaceP; rewrite scale1r => cvg. +apply/rfix_mxP=> y Hy; apply: rstab_act (submx_refl v); apply: subsetP y Hy. +by rewrite defH cycle_subG !inE Gyg /= cvg. +Qed. + +(* This is the main part of B & G, Theorem 2.6; it implies 2.6(a) and most of *) +(* 2.6(b). *) +Theorem der1_odd_GL2_charf F gT (G : {group gT}) + (rG : mx_representation F G 2) : + odd #|G| -> mx_faithful rG -> [char F].-group G^`(1)%g. +Proof. +move=> oddG ffulG. +without loss closF: F rG ffulG / group_closure_field F gT. + move=> IH; apply: (@group_closure_field_exists gT F) => [[Fc f closFc]]. + rewrite -(eq_pgroup _ (fmorph_char f)). + by rewrite -(map_mx_faithful f) in ffulG; exact: IH ffulG closFc. +elim: {G}_.+1 {-2}G (ltnSn #|G|) => // m IHm G le_g_m in rG oddG ffulG *. +apply/pgroupP=> p p_pr pG'; rewrite !inE p_pr /=; apply: wlog_neg => p_nz. +have [P sylP] := Sylow_exists p G. +have nPG: G \subset 'N(P). + apply/idPn=> ltNG; pose N := 'N_G(P); have sNG: N \subset G := subsetIl _ _. + have{IHm ltNG} p'N': [char F].-group N^`(1)%g. + apply: IHm (subg_mx_faithful sNG ffulG); last exact: oddSg oddG. + rewrite -ltnS (leq_trans _ le_g_m) // ltnS proper_card //. + by rewrite /proper sNG subsetI subxx. + have{p'N'} tiPN': P :&: N^`(1)%g = 1%g. + rewrite coprime_TIg ?(pnat_coprime (pHall_pgroup sylP)) //= -/N. + apply: sub_in_pnat p'N' => q _; apply: contraL; move/eqnP->. + by rewrite !inE p_pr. + have sPN: P \subset N by rewrite subsetI normG (pHall_sub sylP). + have{tiPN'} cPN: N \subset 'C(P). + rewrite (sameP commG1P trivgP) -tiPN' subsetI commgS //. + by rewrite commg_subr subsetIr. + have /sdprodP[_ /= defG nKP _] := Burnside_normal_complement sylP cPN. + set K := 'O_p^'(G) in defG nKP; have nKG: G \subset 'N(K) by exact: gFnorm. + suffices p'G': p^'.-group G^`(1)%g by case/eqnP: (pgroupP p'G' p p_pr pG'). + apply: pgroupS (pcore_pgroup p^' G); rewrite -quotient_cents2 //= -/K. + by rewrite -defG quotientMidl /= -/K quotient_cents ?(subset_trans sPN). +pose Q := G^`(1)%g :&: P; have sQG: Q \subset G by rewrite subIset ?der_subS. +have nQG: G \subset 'N(Q) by rewrite normsI // normal_norm ?der_normalS. +have pQ: (p %| #|Q|)%N. + have sylQ: p.-Sylow(G^`(1)%g) Q. + by apply: Sylow_setI_normal (der_normalS _ _) _. + apply: contraLR pG'; rewrite -!p'natE // (card_Hall sylQ) -!partn_eq1 //. + by rewrite part_pnat_id ?part_pnat. +have{IHm} abelQ: abelian Q. + apply/commG1P/eqP/idPn => ntQ'. + have{IHm} p'Q': [char F].-group Q^`(1)%g. + apply: IHm (subg_mx_faithful sQG ffulG); last exact: oddSg oddG. + rewrite -ltnS (leq_trans _ le_g_m) // ltnS proper_card //. + rewrite /proper sQG subsetI //= andbC subEproper. + case: eqP => [-> /= | _]; last by rewrite /proper (pHall_sub sylP) andbF. + have: nilpotent P by rewrite (pgroup_nil (pHall_pgroup sylP)). + move/forallP/(_ P); apply: contraL; rewrite subsetI subxx => -> /=. + apply: contra ntQ'; rewrite /Q => /eqP->. + by rewrite (setIidPr _) ?sub1G // commG1. + case/eqP: ntQ'; have{p'Q'}: P :&: Q^`(1)%g = 1%g. + rewrite coprime_TIg ?(pnat_coprime (pHall_pgroup sylP)) //= -/Q. + by rewrite (pi_p'nat p'Q') // !inE p_pr. + by rewrite (setIidPr _) // comm_subG ?subsetIr. +pose rQ := subg_repr rG sQG. +wlog [U simU sU1]: / exists2 U, mxsimple rQ U & (U <= 1%:M)%MS. + by apply: mxsimple_exists; rewrite ?mxmodule1 ?oner_eq0. +have Uscal: \rank U = 1%N by exact: (mxsimple_abelian_linear (closF _)) simU. +have{simU} [Umod _ _] := simU. +have{sU1} [|V Vmod sumUV dxUV] := mx_Maschke _ _ Umod sU1. + have: p.-group Q by apply: pgroupS (pHall_pgroup sylP); rewrite subsetIr. + by apply: sub_in_pnat=> q _; move/eqnP->; rewrite !inE p_pr. +have [u defU]: exists u : 'rV_2, (u :=: U)%MS. + by move: (row_base U) (eq_row_base U); rewrite Uscal => u; exists u. +have{dxUV Uscal} [v defV]: exists v : 'rV_2, (v :=: V)%MS. + move/mxdirectP: dxUV; rewrite /= Uscal sumUV mxrank1 => [[Vscal]]. + by move: (row_base V) (eq_row_base V); rewrite -Vscal => v; exists v. +pose B : 'M_(1 + 1) := col_mx u v; have{sumUV} uB: B \in unitmx. + rewrite -row_full_unit /row_full eqn_leq rank_leq_row {1}addn1. + by rewrite -addsmxE -(mxrank1 F 2) -sumUV mxrankS // addsmxS ?defU ?defV. +pose Qfix (w : 'rV_2) := {in Q, forall y, w *m rG y <= w}%MS. +have{U defU Umod} u_fix: Qfix u. + by move=> y Qy; rewrite /= (eqmxMr _ defU) defU (mxmoduleP Umod). +have{V defV Vmod} v_fix: Qfix v. + by move=> y Qy; rewrite /= (eqmxMr _ defV) defV (mxmoduleP Vmod). +case/Cauchy: pQ => // x Qx oxp; have Gx := subsetP sQG x Qx. +case/submxP: (u_fix x Qx) => a def_ux. +case/submxP: (v_fix x Qx) => b def_vx. +have def_x: rG x = B^-1 *m block_mx a 0 0 b *m B. + rewrite -mulmxA -[2]/(1 + 1)%N mul_block_col !mul0mx addr0 add0r. + by rewrite -def_ux -def_vx -mul_col_mx mulKmx. +have ap1: a ^+ p = 1. + suff: B^-1 *m block_mx (a ^+ p) 0 0 (b ^+ p) *m B = 1. + move/(canRL (mulmxK uB))/(canRL (mulKVmx uB)); rewrite mul1mx. + by rewrite mulmxV // scalar_mx_block; case/eq_block_mx. + transitivity (rG x ^+ p); last first. + by rewrite -(repr_mxX (subg_repr rG sQG)) // -oxp expg_order repr_mx1. + elim: (p) => [|k IHk]; first by rewrite -scalar_mx_block mulmx1 mulVmx. + rewrite !exprS -IHk def_x -!mulmxE !mulmxA mulmxK // -2!(mulmxA B^-1). + by rewrite -[2]/(1 + 1)%N mulmx_block !mulmx0 !mul0mx !addr0 mulmxA add0r. +have ab1: a * b = 1. + have: Q \subset <<[set y in G | \det (rG y) == 1]>>. + rewrite subIset // genS //; apply/subsetP=> yz; case/imset2P=> y z Gy Gz ->. + rewrite inE !repr_mxM ?groupM ?groupV //= !detM (mulrCA _ (\det (rG y))). + rewrite -!det_mulmx -!repr_mxM ?groupM ?groupV //. + by rewrite mulKg mulVg repr_mx1 det1. + rewrite gen_set_id; last first. + apply/group_setP; split=> [|y z /setIdP[Gy /eqP y1] /setIdP[Gz /eqP z1]]. + by rewrite inE group1 /= repr_mx1 det1. + by rewrite inE groupM ?repr_mxM //= detM y1 z1 mulr1. + case/subsetP/(_ x Qx)/setIdP=> _. + rewrite def_x !detM mulrAC -!detM -mulrA mulKr // -!mulmxE. + rewrite -[2]/(1 + 1)%N det_lblock // [a]mx11_scalar [b]mx11_scalar. + by rewrite !det_scalar1 -scalar_mxM => /eqP->. +have{ab1 ap1 def_x} ne_ab: a != b. + apply/eqP=> defa; have defb: b = 1. + rewrite -ap1 (divn_eq p 2) modn2. + have ->: odd p by rewrite -oxp (oddSg _ oddG) // cycle_subG. + by rewrite addn1 exprS mulnC exprM exprS {1 3}defa ab1 expr1n mulr1. + suff x1: x \in [1] by rewrite -oxp (set1P x1) order1 in p_pr. + rewrite (subsetP ffulG) // inE Gx def_x defa defb -scalar_mx_block mulmx1. + by rewrite mul1mx mulVmx ?eqxx. +have{a b ne_ab def_ux def_vx} nx_uv (w : 'rV_2): + (w *m rG x <= w -> w <= u \/ w <= v)%MS. +- case/submxP=> c; have:= mulmxKV uB w. + rewrite -[_ *m invmx B]hsubmxK [lsubmx _]mx11_scalar [rsubmx _]mx11_scalar. + move: (_ 0) (_ 0) => dv du; rewrite mul_row_col !mul_scalar_mx => <-{w}. + rewrite mulmxDl -!scalemxAl def_ux def_vx mulmxDr -!scalemxAr. + rewrite !scalemxAl -!mul_row_col; move/(can_inj (mulmxK uB)). + case/eq_row_mx => eqac eqbc; apply/orP. + have [-> | nz_dv] := eqVneq dv 0; first by rewrite scale0r addr0 scalemx_sub. + have [-> | nz_du] := eqVneq du 0. + by rewrite orbC scale0r add0r scalemx_sub. + case/eqP: ne_ab; rewrite -[b]scale1r -(mulVf nz_dv) -[a]scale1r. + by rewrite -(mulVf nz_du) -!scalerA eqac eqbc !scalerA !mulVf. +have{x Gx Qx oxp nx_uv} redG y (A := rG y): + y \in G -> (u *m A <= u /\ v *m A <= v)%MS. +- move=> Gy; have uA: row_free A by rewrite row_free_unit repr_mx_unit. + have Ainj (w t : 'rV_2): (w *m A <= w -> t *m A <= w -> t *m A <= t)%MS. + case/sub_rVP=> [c ryww] /sub_rVP[d rytw]. + rewrite -(submxMfree _ _ uA) rytw -scalemxAl ryww scalerA mulrC. + by rewrite -scalerA scalemx_sub. + have{Qx nx_uv} nAx w: Qfix w -> (w *m A <= u \/ w *m A <= v)%MS. + move=> nwQ; apply: nx_uv; rewrite -mulmxA -repr_mxM // conjgCV. + rewrite repr_mxM ?groupJ ?groupV // mulmxA submxMr // nwQ // -mem_conjg. + by rewrite (normsP nQG). + have [uAu | uAv] := nAx _ u_fix; have [vAu | vAv] := nAx _ v_fix; eauto. + have [k ->]: exists k, A = A ^+ k.*2. + exists #[y].+1./2; rewrite -mul2n -divn2 mulnC divnK. + by rewrite -repr_mxX // expgS expg_order mulg1. + by rewrite dvdn2 negbK; apply: oddSg oddG; rewrite cycle_subG. + elim: k => [|k [IHu IHv]]; first by rewrite !mulmx1. + case/sub_rVP: uAv => c uAc; case/sub_rVP: vAu => d vAd. + rewrite doubleS !exprS !mulmxA; do 2!rewrite uAc vAd -!scalemxAl. + by rewrite !scalemx_sub. +suffices trivG': G^`(1)%g = 1%g. + by rewrite /= trivG' cards1 gtnNdvd ?prime_gt1 in pG'. +apply/trivgP; apply: subset_trans ffulG; rewrite gen_subG. +apply/subsetP=> _ /imset2P[y z Gy Gz ->]; rewrite inE groupR //=. +rewrite -(inj_eq (can_inj (mulKmx (repr_mx_unit rG (groupM Gz Gy))))). +rewrite mul1mx mulmx1 -repr_mxM ?(groupR, groupM) // -commgC !repr_mxM //. +rewrite -(inj_eq (can_inj (mulKmx uB))) !mulmxA !mul_col_mx. +case/redG: Gy => /sub_rVP[a uya] /sub_rVP[b vyb]. +case/redG: Gz => /sub_rVP[c uzc] /sub_rVP[d vzd]. +by do 2!rewrite uya vyb uzc vzd -?scalemxAl; rewrite !scalerA mulrC (mulrC d). +Qed. + +(* This is B & G, Theorem 2.6(a) *) +Theorem charf'_GL2_abelian F gT (G : {group gT}) + (rG : mx_representation F G 2) : + odd #|G| -> mx_faithful rG -> [char F]^'.-group G -> abelian G. +Proof. +move=> oddG ffG char'G; apply/commG1P/eqP. +rewrite trivg_card1 (pnat_1 _ (pgroupS _ char'G)) ?comm_subG //=. +exact: der1_odd_GL2_charf ffG. +Qed. + +(* This is B & G, Theorem 2.6(b) *) +Theorem charf_GL2_der_subS_abelian_Sylow p F gT (G : {group gT}) + (rG : mx_representation F G 2) : + odd #|G| -> mx_faithful rG -> p \in [char F] -> + exists P : {group gT}, [/\ p.-Sylow(G) P, abelian P & G^`(1)%g \subset P]. +Proof. +move=> oddG ffG charFp. +have{oddG} pG': p.-group G^`(1)%g. + rewrite /pgroup -(eq_pnat _ (charf_eq charFp)). + exact: der1_odd_GL2_charf ffG. +have{pG'} [P SylP sG'P]:= Sylow_superset (der_sub _ _) pG'. +exists P; split=> {sG'P}//; case/and3P: SylP => sPG pP _. +apply/commG1P/trivgP; apply: subset_trans ffG; rewrite gen_subG. +apply/subsetP=> _ /imset2P[y z Py Pz ->]; rewrite inE (subsetP sPG) ?groupR //=. +pose rP := subg_repr rG sPG; pose U := rfix_mx rP P. +rewrite -(inj_eq (can_inj (mulKmx (repr_mx_unit rP (groupM Pz Py))))). +rewrite mul1mx mulmx1 -repr_mxM ?(groupR, groupM) // -commgC !repr_mxM //. +have: U != 0 by apply: (rfix_pgroup_char charFp). +rewrite -mxrank_eq0 -lt0n 2!leq_eqVlt ltnNge rank_leq_row orbF orbC eq_sym. +case/orP=> [Ufull | Uscal]. + suffices{y z Py Pz} rP1 y: y \in P -> rP y = 1%:M by rewrite !rP1 ?mulmx1. + move=> Py; apply/row_matrixP=> i. + by rewrite rowE -row1 (rfix_mxP P _) ?submx_full. +have [u defU]: exists u : 'rV_2, (u :=: U)%MS. + by move: (row_base U) (eq_row_base U); rewrite -(eqP Uscal) => u; exists u. +have fix_u: {in P, forall x, u *m rP x = u}. + by move/eqmxP: defU; case/andP; move/rfix_mxP. +have [v defUc]: exists u : 'rV_2, (u :=: U^C)%MS. + have UCscal: \rank U^C = 1%N by rewrite mxrank_compl -(eqP Uscal). + by move: (row_base _)%MS (eq_row_base U^C)%MS; rewrite UCscal => v; exists v. +pose B := col_mx u v; have uB: B \in unitmx. + rewrite -row_full_unit -sub1mx -(eqmxMfull _ (addsmx_compl_full U)). + by rewrite mulmx1 -addsmxE addsmxS ?defU ?defUc. +have Umod: mxmodule rP U by exact: rfix_mx_module. +pose W := rfix_mx (factmod_repr Umod) P. +have ntW: W != 0. + apply: (rfix_pgroup_char charFp) => //. + rewrite eqmxMfull ?row_full_unit ?unitmx_inv ?row_ebase_unit //. + by rewrite rank_copid_mx -(eqP Uscal). +have{ntW} Wfull: row_full W. + by rewrite -col_leq_rank {1}mxrank_coker -(eqP Uscal) lt0n mxrank_eq0. +have svW: (in_factmod U v <= W)%MS by rewrite submx_full. +have fix_v: {in P, forall x, v *m rG x - v <= u}%MS. + move=> x Px /=; rewrite -[v *m _](add_sub_fact_mod U) (in_factmodJ Umod) //. + move/rfix_mxP: svW => -> //; rewrite in_factmodK ?defUc // addrK. + by rewrite defU val_submodP. +have fixB: {in P, forall x, exists2 a, u *m rG x = u & v *m rG x = v + a *: u}. + move=> x Px; case/submxP: (fix_v x Px) => a def_vx. + exists (a 0 0); first exact: fix_u. + by rewrite addrC -mul_scalar_mx -mx11_scalar -def_vx subrK. +rewrite -(inj_eq (can_inj (mulKmx uB))) // !mulmxA !mul_col_mx. +case/fixB: Py => a uy vy; case/fixB: Pz => b uz vz. +by rewrite uy uz vy vz !mulmxDl -!scalemxAl uy uz vy vz addrAC. +Qed. + +(* This is B & G, Lemma 2.7. *) +Lemma regular_abelem2_on_abelem2 p q gT (P Q : {group gT}) : + p.-abelem P -> q.-abelem Q -> 'r_p(P) = 2 ->'r_q(Q) = 2 -> + Q \subset 'N(P) -> 'C_Q(P) = 1%g -> + (q %| p.-1)%N + /\ (exists2 a, a \in Q^# & exists r, + [/\ {in P, forall x, x ^ a = x ^+ r}%g, + r ^ q = 1 %[mod p] & r != 1 %[mod p]]). +Proof. +move=> abelP abelQ; rewrite !p_rank_abelem // => logP logQ nPQ regQ. +have ntP: P :!=: 1%g by case: eqP logP => // ->; rewrite cards1 logn1. +have [p_pr _ _]:= pgroup_pdiv (abelem_pgroup abelP) ntP. +have ntQ: Q :!=: 1%g by case: eqP logQ => // ->; rewrite cards1 logn1. +have [q_pr _ _]:= pgroup_pdiv (abelem_pgroup abelQ) ntQ. +pose rQ := abelem_repr abelP ntP nPQ. +have [|P1 simP1 _] := dec_mxsimple_exists (mxmodule1 rQ). + by rewrite oner_eq0. +have [modP1 nzP1 _] := simP1. +have ffulQ: mx_faithful rQ by exact: abelem_mx_faithful. +have linP1: \rank P1 = 1%N. + apply/eqP; have:= abelem_cyclic abelQ; rewrite logQ; apply: contraFT. + rewrite neq_ltn ltnNge lt0n mxrank_eq0 nzP1 => P1full. + have irrQ: mx_irreducible rQ. + apply: mx_iso_simple simP1; apply: eqmx_iso; apply/eqmxP. + by rewrite submx1 sub1mx -col_leq_rank {1}(dim_abelemE abelP ntP) logP. + exact: mx_faithful_irr_abelian_cyclic irrQ (abelem_abelian abelQ). +have ne_qp: q != p. + move/implyP: (logn_quotient_cent_abelem nPQ abelP). + by rewrite logP regQ indexg1 /=; case: eqP => // <-; rewrite logQ. +have redQ: mx_completely_reducible rQ 1%:M. + apply: mx_Maschke; apply: pi_pnat (abelem_pgroup abelQ) _. + by rewrite inE /= (charf_eq (char_Fp p_pr)). +have [P2 modP2 sumP12 dxP12] := redQ _ modP1 (submx1 _). +have{dxP12} linP2: \rank P2 = 1%N. + apply: (@addnI 1%N); rewrite -{1}linP1 -(mxdirectP dxP12) /= sumP12. + by rewrite mxrank1 (dim_abelemE abelP ntP) logP. +have{sumP12} [u def1]: exists u, 1%:M = u.1 *m P1 + u.2 *m P2. + by apply/sub_addsmxP; rewrite sumP12. +pose lam (Pi : 'M(P)) b := (nz_row Pi *m rQ b *m pinvmx (nz_row Pi)) 0 0. +have rQ_lam Pi b: + mxmodule rQ Pi -> \rank Pi = 1%N -> b \in Q -> Pi *m rQ b = lam Pi b *: Pi. +- rewrite /lam => modPi linPi Qb; set v := nz_row Pi; set a := _ 0. + have nz_v: v != 0 by rewrite nz_row_eq0 -mxrank_eq0 linPi. + have sPi_v: (Pi <= v)%MS. + by rewrite -mxrank_leqif_sup ?nz_row_sub // rank_rV nz_v linPi. + have [v' defPi] := submxP sPi_v; rewrite {2}defPi scalemxAr -mul_scalar_mx. + rewrite -mx11_scalar !(mulmxA v') -defPi mulmxKpV ?(submx_trans _ sPi_v) //. + exact: (mxmoduleP modPi). +have lam_q Pi b: + mxmodule rQ Pi -> \rank Pi = 1%N -> b \in Q -> lam Pi b ^+ q = 1. +- move=> modPi linPi Qb; apply/eqP; rewrite eq_sym -subr_eq0. + have: \rank Pi != 0%N by rewrite linPi. + apply: contraR; move/eqmx_scale=> <-. + rewrite mxrank_eq0 scalerBl subr_eq0 -mul_mx_scalar -(repr_mx1 rQ). + have <-: (b ^+ q = 1)%g by case/and3P: abelQ => _ _; move/exponentP->. + apply/eqP; rewrite repr_mxX //. + elim: (q) => [|k IHk]; first by rewrite scale1r mulmx1. + by rewrite !exprS mulmxA rQ_lam // -scalemxAl IHk scalerA. +pose f b := (lam P1 b, lam P2 b). +have inj_f: {in Q &, injective f}. + move=> b c Qb Qc /= [eq_bc1 eq_bc2]; apply: (mx_faithful_inj ffulQ) => //. + rewrite -[rQ b]mul1mx -[rQ c]mul1mx {}def1 !mulmxDl -!mulmxA. + by rewrite !{1}rQ_lam ?eq_bc1 ?eq_bc2. +pose rs := [set x : 'F_p | x ^+ q == 1]. +have s_fQ_rs: f @: Q \subset setX rs rs. + apply/subsetP=> _ /imsetP[b Qb ->]. + by rewrite !{1}inE /= !{1}lam_q ?eqxx. +have le_rs_q: #|rs| <= q ?= iff (#|rs| == q). + split; rewrite // cardE max_unity_roots ?enum_uniq ?prime_gt0 //. + by apply/allP=> x; rewrite mem_enum inE unity_rootE. +have:= subset_leqif_card s_fQ_rs. +rewrite card_in_imset // (card_pgroup (abelem_pgroup abelQ)) logQ. +case/(leqif_trans (leqif_mul le_rs_q le_rs_q))=> _; move/esym. +rewrite cardsX eqxx andbb muln_eq0 orbb eqn0Ngt prime_gt0 //= => /andP[rs_q]. +rewrite subEproper /proper {}s_fQ_rs andbF orbF => /eqP rs2_Q. +have: ~~ (rs \subset [set 1 : 'F_p]). + apply: contraL (prime_gt1 q_pr); move/subset_leq_card. + by rewrite cards1 (eqnP rs_q) leqNgt. +case/subsetPn => r rs_r; rewrite inE => ne_r_1. +have rq1: r ^+ q = 1 by apply/eqP; rewrite inE in rs_r. +split. + have Ur: r \in GRing.unit. + by rewrite -(unitrX_pos _ (prime_gt0 q_pr)) rq1 unitr1. + pose u_r : {unit 'F_p} := Sub r Ur; have:= order_dvdG (in_setT u_r). + rewrite card_units_Zp ?pdiv_gt0 // {2}/pdiv primes_prime //=. + rewrite (@totient_pfactor p 1) // muln1; apply: dvdn_trans. + have: (u_r ^+ q == 1)%g. + by rewrite -val_eqE unit_Zp_expg -Zp_nat natrX natr_Zp rq1. + case/primeP: q_pr => _ q_min; rewrite -order_dvdn; move/q_min. + by rewrite order_eq1 -val_eqE (negPf ne_r_1) /=; move/eqnP->. +have /imsetP[a Qa [def_a1 def_a2]]: (r, r) \in f @: Q. + by rewrite -rs2_Q inE andbb. +have rQa: rQ a = r%:M. + rewrite -[rQ a]mul1mx def1 mulmxDl -!mulmxA !rQ_lam //. + by rewrite -def_a1 -def_a2 !linearZ -scalerDr -def1 /= scalemx1. +exists a. + rewrite !inE Qa andbT; apply: contra ne_r_1 => a1. + by rewrite (eqP a1) repr_mx1 in rQa; rewrite (fmorph_inj _ rQa). +exists r; rewrite -!val_Fp_nat // natrX natr_Zp rq1. +split=> // x Px; apply: (@abelem_rV_inj _ _ _ abelP ntP); rewrite ?groupX //. + by rewrite memJ_norm ?(subsetP nPQ). +by rewrite abelem_rV_X // -mul_mx_scalar natr_Zp -rQa -abelem_rV_J. +Qed. + +End BGsection2. diff --git a/mathcomp/odd_order/BGsection3.v b/mathcomp/odd_order/BGsection3.v new file mode 100644 index 0000000..25879a6 --- /dev/null +++ b/mathcomp/odd_order/BGsection3.v @@ -0,0 +1,1831 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div. +Require Import fintype tuple bigop prime binomial finset ssralg fingroup finalg. +Require Import morphism perm automorphism quotient action commutator gproduct. +Require Import zmodp cyclic gfunctor center pgroup gseries nilpotent sylow. +Require Import finmodule abelian frobenius maximal extremal hall. +Require Import matrix mxalgebra mxrepresentation mxabelem wielandt_fixpoint. +Require Import BGsection1 BGsection2. + +(******************************************************************************) +(* This file covers the material in B & G, Section 3. *) +(* Note that in spite of the use of Gorenstein 2.7.6, the material in all *) +(* of Section 3, and in all likelyhood the whole of B & G, does NOT depend on *) +(* the general proof of existence of Frobenius kernels, because results on *) +(* Frobenius groups are only used when the semidirect product decomposition *) +(* is already known, and (see file frobenius.v) in this case the kernel is *) +(* equal to the normal complement of the Frobenius complement. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope ring_scope. +Import GroupScope GRing.Theory. + +Section BGsection3. + +Implicit Type F : fieldType. +Implicit Type gT : finGroupType. +Implicit Type p : nat. + +(* B & G, Lemma 3.1 is covered by frobenius.Frobenius_semiregularP. *) + +(* This is B & G, Lemma 3.2. *) +Section FrobeniusQuotient. + +Variables (gT : finGroupType) (G K R : {group gT}). +Implicit Type N : {group gT}. + +(* This is a special case of B & G, Lemma 3.2 (b). *) +Lemma Frobenius_proper_quotient N : + [Frobenius G = K ><| R] -> solvable K -> N <| G -> N \proper K -> + [Frobenius G / N = (K / N) ><| (R / N)]. +Proof. +move=> frobG solK nsNG /andP[sNK ltNK]. +have [defG _ ntR _ _] := Frobenius_context frobG. +have [nsKG sRG defKR nKR tiKR] := sdprod_context defG; have [sKG _]:= andP nsKG. +have nsNK := normalS sNK sKG nsNG. +apply/Frobenius_semiregularP=> [|||Nx]. +- rewrite sdprodE ?quotient_norms //. + by rewrite -quotientMl ?defKR ?normal_norm. + by rewrite -quotientGI // tiKR quotient1. +- by rewrite -subG1 quotient_sub1 ?normal_norm. +- rewrite -subG1 quotient_sub1; last by rewrite (subset_trans sRG) ?normal_norm. + apply: contra ntR => sRN. + by rewrite -subG1 -tiKR subsetI (subset_trans sRN) /=. +rewrite !inE andbC => /andP[/morphimP[x nNx Rx ->{Nx}] notNx]. +apply/trivgP; rewrite /= -cent_cycle -quotient_cycle //. +rewrite -coprime_quotient_cent ?cycle_subG //; last first. + by apply: coprimegS (Frobenius_coprime frobG); rewrite cycle_subG. +rewrite cent_cycle (Frobenius_reg_ker frobG) ?quotient1 // !inE Rx andbT. +by apply: contraNneq notNx => ->; rewrite morph1. +Qed. + +(* This is B & G, Lemma 3.2 (a). *) +Lemma Frobenius_normal_proper_ker N : + [Frobenius G = K ><| R] -> solvable K -> N <| G -> ~~ (K \subset N) -> + N \proper K. +Proof. +move=> frobG solK nsNG ltNK; have [sNG nNG] := andP nsNG; pose H := N :&: K. +have [defG _ ntR _ _] := Frobenius_context frobG. +have [nsKG _ /mulG_sub[sKG _] nKR tiKR] := sdprod_context defG. +have nsHG: H <| G := normalI nsNG nsKG; have [_ nHG] := andP nsHG. +have ltHK: H \proper K by rewrite /proper subsetIr subsetI subxx andbT. +suffices /eqP tiNR: N :&: R == 1. + rewrite /proper ltNK andbT -(setIidPl sNG). + rewrite -(cover_partition (Frobenius_partition frobG)) big_distrr /=. + apply/bigcupsP=> _ /setU1P[->| /imsetP[x Kx ->]]; first exact: subsetIr. + rewrite conjD1g setIDA subDset -(normsP (subset_trans sKG nNG) x) //. + by rewrite -conjIg tiNR conjs1g subsetUl. +suffices: (N :&: R) / H \subset [1]. + by rewrite -subG1 quotient_sub1 ?normsGI // -subsetIidr setIACA tiKR setIg1. +have frobGq := Frobenius_proper_quotient frobG solK nsHG ltHK. +have [_ ntKq _ _ _] := Frobenius_context frobGq. +rewrite -(cent_semiregular (Frobenius_reg_compl frobGq) _ ntKq) //. +rewrite subsetI quotientS ?subsetIr // quotient_cents2r //. +by rewrite commg_subI ?setIS // subsetIidl (subset_trans sKG). +Qed. + +(* This is B & G, Lemma 3.2 (b). *) +Lemma Frobenius_quotient N : + [Frobenius G = K ><| R] -> solvable K -> N <| G -> ~~ (K \subset N) -> + [Frobenius G / N = (K / N) ><| (R / N)]. +Proof. +move=> frobG solK nsNG ltKN; apply: Frobenius_proper_quotient => //. +exact: (Frobenius_normal_proper_ker frobG). +Qed. + +End FrobeniusQuotient. + +(* This is B & G, Lemma 3.3. *) +Lemma Frobenius_rfix_compl F gT (G K R : {group gT}) n + (rG : mx_representation F G n) : + [Frobenius G = K ><| R] -> [char F]^'.-group K -> + ~~ (K \subset rker rG) -> rfix_mx rG R != 0. +Proof. +rewrite /pgroup charf'_nat => frobG nzK. +have [defG _ _ ltKG ltRG]:= Frobenius_context frobG. +have{ltKG ltRG} [sKG sRG]: K \subset G /\ R \subset G by rewrite !proper_sub. +apply: contraNneq => fixR0; rewrite rfix_mx_rstabC // -(eqmx_scale _ nzK). +pose gsum H := gring_op rG (gset_mx F G H). +have fixsum (H : {group gT}): H \subset G -> (gsum H <= rfix_mx rG H)%MS. + move/subsetP=> sHG; apply/rfix_mxP=> x Hx; have Gx := sHG x Hx. + rewrite -gring_opG // -gring_opM ?envelop_mx_id //; congr (gring_op _ _). + rewrite {2}/gset_mx (reindex_acts 'R _ Hx) ?astabsR //= mulmx_suml. + by apply:eq_bigr=> y; move/sHG=> Gy; rewrite repr_mxM. +have: gsum G + rG 1 *+ #|K| = gsum K + \sum_(x in K) gsum (R :^ x). + rewrite -gring_opG // -sumr_const -!linear_sum -!linearD; congr gring_op. + rewrite {1}/gset_mx (set_partition_big _ (Frobenius_partition frobG)) /=. + rewrite big_setU1 -?addrA /=; last first. + by apply: contraL (group1 K) => /imsetP[x _ ->]; rewrite conjD1g !inE eqxx. + congr (_ + _); rewrite big_imset /= => [|x y Kx Ky /= eqRxy]; last first. + have [/eqP/sdprodP[_ _ _ tiKR] _ _ _ /eqP snRG] := and5P frobG. + apply/eqP; rewrite eq_mulgV1 -in_set1 -set1gE -tiKR -snRG setIA. + by rewrite (setIidPl sKG) !inE conjsgM eqRxy actK groupM /= ?groupV. + rewrite -big_split; apply: eq_bigr => x Kx /=. + by rewrite addrC conjD1g -big_setD1 ?group1. +have ->: gsum G = 0. + apply/eqP; rewrite -submx0 -fixR0; apply: submx_trans (rfix_mxS rG sRG). + exact: fixsum. +rewrite repr_mx1 -scaler_nat add0r => ->. +rewrite big1 ?addr0 ?fixsum // => x Kx; have Gx := subsetP sKG x Kx. +apply/eqP; rewrite -submx0 (submx_trans (fixsum _ _)) ?conj_subG //. +by rewrite -(mul0mx _ (rG x)) -fixR0 rfix_mx_conjsg. +Qed. + +(* This is Aschbacher (40.6)(3), or G. (3.14)(iii). *) +Lemma regular_pq_group_cyclic gT p q (H R : {group gT}) : + [/\ prime p, prime q & p != q] -> #|R| = (p * q)%N -> + H :!=: 1 -> R \subset 'N(H) -> semiregular H R -> + cyclic R. +Proof. +case=> pr_p pr_q q'p oR ntH nHR regHR. +without loss{q'p} ltpq: p q pr_p pr_q oR / p < q. + by move=> IH; case: ltngtP q'p => // /IH-> //; rewrite mulnC. +have [p_gt0 q_gt0]: 0 < p /\ 0 < q by rewrite !prime_gt0. +have [[P sylP] [Q sylQ]] := (Sylow_exists p R, Sylow_exists q R). +have [sPR sQR] := (pHall_sub sylP, pHall_sub sylQ). +have [oP oQ]: #|P| = p /\ #|Q| = q. + rewrite (card_Hall sylQ) (card_Hall sylP) oR !p_part !lognM ?logn_prime //. + by rewrite !eqxx eq_sym gtn_eqF. +have [ntP ntQ]: P :!=: 1 /\ Q :!=: 1 by rewrite -!cardG_gt1 oP oQ !prime_gt1. +have nQR: R \subset 'N(Q). + rewrite -subsetIidl -indexg_eq1 -(card_Syl_mod R pr_q) (card_Syl sylQ) /=. + rewrite modn_small // -divgS ?subsetIl ?ltn_divLR // mulnC oR ltn_pmul2r //. + by rewrite (leq_trans ltpq) // -oQ subset_leq_card // subsetI sQR normG. +have coQP: coprime #|Q| #|P|. + by rewrite oP oQ prime_coprime ?dvdn_prime2 ?gtn_eqF. +have defR: Q ><| P = R. + rewrite sdprodE ?coprime_TIg ?(subset_trans sPR) //. + by apply/eqP; rewrite eqEcard mul_subG //= oR coprime_cardMg // oP oQ mulnC. +have [cycP cycQ]: cyclic P /\ cyclic Q by rewrite !prime_cyclic ?oP ?oQ. +suffices cQP: P \subset 'C(Q) by rewrite (@cyclic_dprod _ Q P) ?dprodEsd. +without loss /is_abelemP[r pr_r abelH]: H ntH nHR regHR / is_abelem H. + move=> IH; have [r _ rH] := rank_witness H. + have solR: solvable R. + apply/metacyclic_sol/metacyclicP; exists Q. + by rewrite /(Q <| R) sQR -(isog_cyclic (sdprod_isog defR)). + have coHR: coprime #|H| #|R| := regular_norm_coprime nHR regHR. + have [H1 sylH1 nH1R] := sol_coprime_Sylow_exists r solR nHR coHR. + have ntH1: H1 :!=: 1 by rewrite -rank_gt0 (rank_Sylow sylH1) -rH rank_gt0. + have [H2 minH2 sH21] := minnormal_exists ntH1 nH1R. + have [sH1H rH1 _] := and3P sylH1; have sH2H := subset_trans sH21 sH1H. + have [nH2R ntH2 abelH2] := minnormal_solvable minH2 sH21 (pgroup_sol rH1). + by apply: IH abelH2 => //; apply: semiregularS regHR. +have: rfix_mx (abelem_repr abelH ntH nHR) P == 0. + rewrite -mxrank_eq0 rfix_abelem // mxrank_eq0 rowg_mx_eq0 /=. + by rewrite (cent_semiregular regHR) ?morphim1. +apply: contraLR => not_cQP; have{not_cQP} frobR: [Frobenius R = Q ><| P]. + by apply/prime_FrobeniusP; rewrite ?prime_TIg ?oP ?oQ // centsC. +apply: (Frobenius_rfix_compl frobR). + rewrite (eq_p'group _ (charf_eq (char_Fp pr_r))). + rewrite (coprime_p'group _ (abelem_pgroup abelH)) //. + by rewrite coprime_sym (coprimegS sQR) ?regular_norm_coprime. +rewrite rker_abelem subsetI sQR centsC. +by rewrite -subsetIidl (cent_semiregular regHR) ?subG1. +Qed. + +(* This is B & G, Theorem 3.4. *) +Theorem odd_prime_sdprod_rfix0 F gT (G K R : {group gT}) n + (rG : mx_representation F G n) : + K ><| R = G -> solvable G -> odd #|G| -> coprime #|K| #|R| -> prime #|R| -> + [char F]^'.-group G -> rfix_mx rG R = 0 -> + [~: R, K] \subset rker rG. +Proof. +move: {2}_.+1 (ltnSn #|G|) => m; elim: m => // m IHm in gT G K R n rG *. +rewrite ltnS; set p := #|R| => leGm defG solG oddG coKR p_pr F'G regR. +have [nsKG sRG defKR nKR tiKR] := sdprod_context defG. +have [sKG nKG] := andP nsKG; have solK := solvableS sKG solG. +have [-> | ntK] := eqsVneq K 1; first by rewrite commG1 sub1G. +have ker_ltK (H : {group gT}): + H \proper K -> R \subset 'N(H) -> [~: R, H] \subset rker rG. +- move=> ltKH nHR; have sHK := proper_sub ltKH; set G1 := H <*> R. + have sG1G: G1 \subset G by rewrite join_subG (subset_trans sHK). + have coHR := coprimeSg sHK coKR. + have defG1: H ><| R = G1 by rewrite sdprodEY // coprime_TIg. + apply: subset_trans (subsetIr G1 _); rewrite -(rker_subg _ sG1G). + apply: IHm; rewrite ?(solvableS sG1G) ?(oddSg sG1G) ?(pgroupS sG1G) //. + apply: leq_trans leGm; rewrite /= norm_joinEr // -defKR !coprime_cardMg //. + by rewrite ltn_pmul2r ?proper_card. +without loss [q q_pr qK]: / exists2 q, prime q & q.-group K. + move=> IH; set q := pdiv #|K|. + have q_pr: prime q by rewrite pdiv_prime ?cardG_gt1. + have exHall := coprime_Hall_exists _ nKR coKR solK. + have [Q sylQ nQR] := exHall q; have [Q' hallQ' nQ'R] := exHall q^'. + have [sQK qQ _] := and3P sylQ; have [sQ'K q'Q' _] := and3P hallQ'. + without loss{IH} ltQK: / Q \proper K. + by rewrite properEneq; case: eqP IH => [<- -> | _ _ ->] //; exists q. + have ltQ'K: Q' \proper K. + rewrite properEneq; case: eqP (pgroupP q'Q' q q_pr) => //= ->. + by rewrite !inE pdiv_dvd eqxx; apply. + have nkerG := subset_trans _ (rker_norm rG). + rewrite -quotient_cents2 ?nkerG //. + have <-: Q * Q' = K. + apply/eqP; rewrite eqEcard mulG_subG sQK sQ'K. + rewrite coprime_cardMg ?(pnat_coprime qQ) //=. + by rewrite (card_Hall sylQ) (card_Hall hallQ') partnC. + rewrite quotientMl ?nkerG ?(subset_trans sQK) // centM subsetI. + by rewrite !quotient_cents2r ?ker_ltK. +without loss{m IHm leGm} [ffulG cycZ]: / rker rG = 1 /\ cyclic 'Z(G). + move=> IH; wlog [I M /= simM sumM _]: / mxsemisimple rG 1%:M. + exact: (mx_reducible_semisimple (mxmodule1 _) (mx_Maschke _ F'G)). + pose not_cRK_M i := ~~ ([~: R, K] \subset rstab rG (M i)). + case: (pickP not_cRK_M) => [i | cRK_M]; last first. + rewrite rfix_mx_rstabC ?comm_subG // -sumM. + apply/sumsmx_subP=> i _; move/negbFE: (cRK_M i). + by rewrite rfix_mx_rstabC ?comm_subG. + have [modM ntM _] := simM i; pose rM := kquo_repr (submod_repr modM). + do [rewrite {+}/not_cRK_M -(rker_submod modM) /=; set N := rker _] in rM *. + have [N1 _ | ntN] := eqVneq N 1. + apply: IH; split. + by apply/trivgP; rewrite -N1 /N rker_submod rstabS ?submx1. + have: mx_irreducible (submod_repr modM) by exact/submod_mx_irr. + by apply: mx_faithful_irr_center_cyclic; exact/trivgP. + have tiRN: R :&: N = 1. + by apply: prime_TIg; rewrite //= rker_submod rfix_mx_rstabC // regR submx0. + have nsNG: N <| G := rker_normal _; have [sNG nNG] := andP nsNG. + have nNR := subset_trans sRG nNG. + have sNK: N \subset K. + have [pi hallK]: exists pi, pi.-Hall(G) K. + by apply: HallP; rewrite -(coprime_sdprod_Hall_l defG). + rewrite (sub_normal_Hall hallK) //=. + apply: pnat_dvd (pHall_pgroup hallK). + rewrite -(dvdn_pmul2r (prime_gt0 p_pr)) -!TI_cardMg // 1?setIC // defKR. + by rewrite -norm_joinEr // cardSg // join_subG sNG. + have defGq: (K / N) ><| (R / N) = G / N. + rewrite sdprodE ?quotient_norms -?quotientMr ?defKR //. + by rewrite -quotientGI // tiKR quotient1. + case/negP; rewrite -quotient_cents2 ?(subset_trans _ nNG) //= -/N. + rewrite (sameP commG1P trivgP). + apply: subset_trans (kquo_mx_faithful (submod_repr modM)). + rewrite IHm ?quotient_sol ?coprime_morph ?morphim_odd ?quotient_pgroup //. + - apply: leq_trans leGm; exact: ltn_quotient. + - by rewrite card_quotient // -indexgI tiRN indexg1. + apply/eqP; rewrite -submx0 rfix_quo // rfix_submod //. + by rewrite regR capmx0 linear0 sub0mx. +without loss perfectK: / [~: K, R] = K. + move=> IH; have: [~: K, R] \subset K by rewrite commg_subl. + rewrite subEproper; case/predU1P=> //; move/ker_ltK. + by rewrite commGC commg_normr coprime_commGid // commGC => ->. +have primeR: {in R^#, forall x, 'C_K[x] = 'C_K(R)}. + move=> x; case/setD1P=> nt_x Rx; rewrite -cent_cycle ((<[x]> =P R) _) //. + rewrite eqEsubset cycle_subG Rx; apply: contraR nt_x; move/prime_TIg. + by rewrite -cycle_eq1 (setIidPr _) ?cycle_subG // => ->. +case cKK: (abelian K). + rewrite commGC perfectK; move/eqP: regR; apply: contraLR. + apply: Frobenius_rfix_compl => //; last exact: pgroupS F'G. + rewrite -{2 4}perfectK coprime_abel_cent_TI // in primeR. + by apply/Frobenius_semiregularP; rewrite // -cardG_gt1 prime_gt1. +have [spK defZK]: special K /\ 'C_K(R) = 'Z(K). + apply: (abelian_charsimple_special qK) => //. + apply/bigcupsP=> H; case/andP=> chHK cHH. + have:= char_sub chHK; rewrite subEproper. + case/predU1P=> [eqHK | ltHK]; first by rewrite eqHK cKK in cHH. + have nHR: R \subset 'N(H) := char_norm_trans chHK nKR. + by rewrite (sameP commG1P trivgP) /= commGC -ffulG ker_ltK. +have{spK} esK: extraspecial K. + have abelZK := center_special_abelem qK spK; have [qZK _] := andP abelZK. + have /(pgroup_pdiv qZK)[_ _ []]: 'Z(K) != 1. + by case: spK => _ <-; rewrite (sameP eqP commG1P) -abelianE cKK. + case=> [|e] oK; first by split; rewrite ?oK. + suffices: cyclic 'Z(K) by rewrite (abelem_cyclic abelZK) oK pfactorK. + rewrite (cyclicS _ cycZ) // subsetI subIset ?sKG //=. + by rewrite -defKR centM subsetI -{2}defZK !subsetIr. +have [e e_gt0 oKqe] := card_extraspecial qK esK. +have cycR: cyclic R := prime_cyclic p_pr. +have co_q_p: coprime q p by rewrite oKqe coprime_pexpl in coKR. +move/eqP: regR; case/idPn. +rewrite defZK in primeR. +case: (repr_extraspecial_prime_sdprod_cycle _ _ defG _ oKqe) => // _. +apply=> //; last exact/trivgP. +apply: contraL (oddSg sRG oddG); move/eqP->; have:= oddSg sKG oddG. +by rewrite oKqe addn1 /= !odd_exp /= orbC => ->. +Qed. + +(* Internal action version of B & G, Theorem 3.4. *) +Theorem odd_prime_sdprod_abelem_cent1 k gT (G K R V : {group gT}) : + solvable G -> odd #|G| -> K ><| R = G -> coprime #|K| #|R| -> prime #|R| -> + k.-abelem V -> G \subset 'N(V) -> k^'.-group G -> 'C_V(R) = 1 -> + [~: R, K] \subset 'C_K(V). +Proof. +move=> solG oddG defG coKR prR abelV nVG k'G regR. +have [_ sRG _ nKR _] := sdprod_context defG; rewrite subsetI commg_subr nKR. +case: (eqsVneq V 1) => [-> | ntV]; first exact: cents1. +pose rV := abelem_repr abelV ntV nVG. +apply: subset_trans (_ : rker rV \subset _); last first. + by rewrite rker_abelem subsetIr. +apply: odd_prime_sdprod_rfix0 => //. + have k_pr: prime k by case/pgroup_pdiv: (abelem_pgroup abelV). + by rewrite (eq_pgroup G (eq_negn (charf_eq (char_Fp k_pr)))). +by apply/eqP; rewrite -submx0 rfix_abelem //= regR morphim1 rowg_mx1. +Qed. + +(* This is B & G, Theorem 3.5. *) +Theorem Frobenius_prime_rfix1 F gT (G K R : {group gT}) n + (rG : mx_representation F G n) : + K ><| R = G -> solvable G -> prime #|R| -> 'C_K(R) = 1 -> + [char F]^'.-group G -> \rank (rfix_mx rG R) = 1%N -> + K^`(1) \subset rker rG. +Proof. +move=> defG solG p_pr regR F'G fixRlin. +wlog closF: F rG F'G fixRlin / group_closure_field F gT. + move=> IH; apply: (@group_closure_field_exists gT F) => [[Fc f closFc]]. + rewrite -(rker_map f) IH //; last by rewrite -map_rfix_mx mxrank_map. + by rewrite (eq_p'group _ (fmorph_char f)). +move: {2}_.+1 (ltnSn #|K|) => m. +elim: m => // m IHm in gT G K R rG solG p_pr regR F'G closF fixRlin defG *. +rewrite ltnS => leKm. +have [nsKG sRG defKR nKR tiKR] := sdprod_context defG. +have [sKG nKG] := andP nsKG; have solK := solvableS sKG solG. +have cycR := prime_cyclic p_pr. +case: (eqsVneq K 1) => [-> | ntK]; first by rewrite derg1 commG1 sub1G. +have defR x: x \in R^# -> <[x]> = R. + case/setD1P; rewrite -cycle_subG -cycle_eq1 => ntX sXR. + apply/eqP; rewrite eqEsubset sXR; apply: contraR ntX => /(prime_TIg p_pr). + by rewrite /= (setIidPr sXR) => ->. +have ntR: R :!=: 1 by rewrite -cardG_gt1 prime_gt1. +have frobG: [Frobenius G = K ><| R]. + by apply/Frobenius_semiregularP=> // x Rx; rewrite -cent_cycle defR. +case: (eqVneq (rker rG) 1) => [ffulG | ntC]; last first. + set C := rker rG in ntC *; have nsCG: C <| G := rker_normal rG. + have [sCG nCG] := andP nsCG. + have nCK := subset_trans sKG nCG; have nCR := subset_trans sRG nCG. + case sKC: (K \subset C); first exact: subset_trans (der_sub _ _) _. + have sCK: C \subset K. + by rewrite proper_sub // (Frobenius_normal_proper_ker frobG) ?sKC. + have frobGq: [Frobenius G / C = (K / C) ><| (R / C)]. + by apply: Frobenius_quotient; rewrite ?sKC. + have [defGq _ ntRq _ _] := Frobenius_context frobGq. + rewrite -quotient_sub1 ?comm_subG ?quotient_der //= -/C. + apply: subset_trans (kquo_mx_faithful rG). + apply: IHm defGq _; rewrite 1?(quotient_sol, quotient_pgroup, rfix_quo) //. + - rewrite card_quotient // -indexgI /= -/C setIC. + by rewrite -(setIidPl sCK) -setIA tiKR (setIidPr (sub1G _)) indexg1. + - have: cyclic (R / C) by [rewrite quotient_cyclic]; case/cyclicP=> Cx defRq. + rewrite /= defRq cent_cycle (Frobenius_reg_ker frobGq) //= !inE defRq. + by rewrite cycle_id -cycle_eq1 -defRq ntRq. + - move=> Hq; rewrite -(group_inj (cosetpreK Hq)). + by apply: quotient_splitting_field; rewrite ?subsetIl. + by apply: leq_trans leKm; exact: ltn_quotient. +have ltK_abelian (N : {group gT}): R \subset 'N(N) -> N \proper K -> abelian N. + move=> nNR ltNK; have [sNK _] := andP ltNK; apply/commG1P/trivgP. + rewrite -(setIidPr (sub1G (N <*> R))) /= -ffulG; set G1 := N <*> R. + have sG1: G1 \subset G by rewrite join_subG (subset_trans sNK). + have defG1: N ><| R = G1. + by rewrite sdprodEY //; apply/trivgP; rewrite -tiKR setSI. + rewrite -(rker_subg _ sG1). + apply: IHm defG1 _; rewrite ?(solvableS sG1) ?(pgroupS sG1) //. + by apply/trivgP; rewrite -regR setSI. + by apply: leq_trans leKm; exact: proper_card. +have cK'K': abelian K^`(1). + apply: ltK_abelian; first exact: char_norm_trans (der_char _ _) nKR. + exact: (sol_der1_proper solK). +pose fixG := rfix_mx rG; pose NRmod N (U : 'M_n) := N <*> R \subset rstabs rG U. +have dx_modK_rfix (N : {group gT}) U V: + N \subset K -> R \subset 'N(N) -> NRmod N U -> NRmod N V -> + mxdirect (U + V) -> (U <= fixG N)%MS || (V <= fixG N)%MS. +- move=> sNK nNR nUNR nVNR dxUV. + have [-> | ntN] := eqsVneq N 1; first by rewrite -rfix_mx_rstabC sub1G. + have sNRG: N <*> R \subset G by rewrite join_subG (subset_trans sNK). + pose rNR := subg_repr rG sNRG. + have nfixU W: NRmod N W -> ~~ (W <= fixG N)%MS -> (fixG R <= W)%MS. + move=> nWN not_cWN; rewrite (sameP capmx_idPr eqmxP). + rewrite -(geq_leqif (mxrank_leqif_eq (capmxSr _ _))) fixRlin lt0n. + rewrite mxrank_eq0 -(in_submodK (capmxSl _ _)) val_submod_eq0. + have modW: mxmodule rNR W by rewrite /mxmodule rstabs_subg subsetI subxx. + rewrite -(eqmx_eq0 (rfix_submod modW _)) ?joing_subr //. + apply: Frobenius_rfix_compl (pgroupS (subset_trans sNK sKG) F'G) _. + apply/Frobenius_semiregularP=> // [|x Rx]. + by rewrite sdprodEY //; apply/trivgP; rewrite -tiKR setSI. + by apply/trivgP; rewrite -regR /= -cent_cycle defR ?setSI. + by rewrite rker_submod rfix_mx_rstabC ?joing_subl. + have: fixG R != 0 by rewrite -mxrank_eq0 fixRlin. + apply: contraR; case/norP=> not_fixU not_fixW. + by rewrite -submx0 -(mxdirect_addsP dxUV) sub_capmx !nfixU. +have redG := mx_Maschke rG F'G. +wlog [U simU nfixU]: / exists2 U, mxsimple rG U & ~~ (U <= fixG K)%MS. + move=> IH; wlog [I U /= simU sumU _]: / mxsemisimple rG 1%:M. + exact: (mx_reducible_semisimple (mxmodule1 _) redG). + case: (pickP (fun i => ~~ (U i <= fixG K))%MS) => [i nfixU | fixK]. + by apply: IH; exists (U i). + apply: (subset_trans (der_sub _ _)); rewrite rfix_mx_rstabC // -sumU. + by apply/sumsmx_subP=> i _; apply/idPn; rewrite fixK. +have [modU ntU minU] := simU; pose rU := submod_repr modU. +have irrU: mx_irreducible rU by exact/submod_mx_irr. +have [W modW sumUW dxUW] := redG U modU (submx1 U). +have cWK: (W <= fixG K)%MS. + have:= dx_modK_rfix _ _ _ (subxx _) nKR _ _ dxUW. + by rewrite /NRmod /= norm_joinEr // defKR (negPf nfixU); exact. +have nsK'G: K^`(1) <| G by exact: char_normal_trans (der_char _ _) nsKG. +have [sK'G nK'G] := andP nsK'G. +suffices nregK'U: (rfix_mx rU K^`(1))%MS != 0. + rewrite rfix_mx_rstabC ?normal_sub // -sumUW addsmx_sub andbC. + rewrite (submx_trans cWK) ?rfix_mxS ?der_sub //= (sameP capmx_idPl eqmxP). + rewrite minU ?capmxSl ?capmx_module ?normal_rfix_mx_module //. + apply: contra nregK'U => cUK'; rewrite (eqmx_eq0 (rfix_submod _ _)) //. + by rewrite (eqP cUK') linear0. +pose rK := subg_repr rU (normal_sub nsKG); set p := #|R| in p_pr. +wlog sK: / socleType rK by exact: socle_exists. +have [i _ def_sK]: exists2 i, i \in setT & [set: sK] = orbit 'Cl G i. + by apply/imsetP; exact: Clifford_atrans. +have card_sK: #|[set: sK]| = #|G : 'C[i | 'Cl]|. + by rewrite def_sK card_orbit_in ?indexgI. +have ciK: K \subset 'C[i | 'Cl]. + apply: subset_trans (astabS _ (subsetT _)). + by apply: subset_trans (Clifford_astab _); exact: joing_subl. +pose M := socle_base i; have simM: mxsimple rK M := socle_simple i. +have [sKp | sK1 {ciK card_sK}]: #|[set: sK]| = p \/ #|[set: sK]| = 1%N. +- apply/pred2P; rewrite orbC card_sK; case/primeP: p_pr => _; apply. + by rewrite (_ : p = #|G : K|) ?indexgS // -divgS // -(sdprod_card defG) mulKn. +- have{def_sK} def_sK: [set: sK] = orbit 'Cl R i. + apply/eqP; rewrite eq_sym -subTset def_sK. + apply/subsetP=> i_yz; case/imsetP=> yz; rewrite -{1}defKR. + case/imset2P=> y z; move/(subsetP ciK); rewrite !inE sub1set inE. + case/andP=> Gy; move/eqP=> ciy Rz -> ->{yz i_yz}. + by rewrite actMin ?(subsetP sRG z Rz) // ciy mem_orbit. + have inj_i: {in R &, injective ('Cl%act i)}. + apply/dinjectiveP; apply/card_uniqP; rewrite size_map -cardE -/p. + by rewrite -sKp def_sK /orbit Imset.imsetE cardsE. + pose sM := (\sum_(y in R) M *m rU y)%MS. + have dxM: mxdirect sM. + apply/mxdirect_sumsP=> y Ry; have Gy := subsetP sRG y Ry. + pose j := 'Cl%act i y. + apply/eqP; rewrite -submx0 -{2}(mxdirect_sumsP (Socle_direct sK) j) //. + rewrite capmxS ?val_Clifford_act // ?submxMr ?component_mx_id //. + apply/sumsmx_subP => z; case/andP=> Rz ne_z_y; have Gz := subsetP sRG z Rz. + rewrite (sumsmx_sup ('Cl%act i z)) ?(inj_in_eq inj_i) //. + by rewrite val_Clifford_act // ?submxMr // ?component_mx_id. + pose inCR := \sum_(x in R) rU x. + have im_inCR: (inCR <= rfix_mx rU R)%MS. + apply/rfix_mxP=> x Rx; have Gx := subsetP sRG x Rx. + rewrite {2}[inCR](reindex_astabs 'R x) ?astabsR //= mulmx_suml. + by apply: eq_bigr => y; move/(subsetP sRG)=> Gy; rewrite repr_mxM. + pose inM := proj_mx M (\sum_(x in R | x != 1) M *m rU x)%MS. + have dxM1 := mxdirect_sumsP dxM _ (group1 R). + rewrite repr_mx1 mulmx1 in dxM1. + have inCR_K: M *m inCR *m inM = M. + rewrite mulmx_sumr (bigD1 1) //= repr_mx1 mulmx1 mulmxDl proj_mx_id //. + by rewrite proj_mx_0 ?addr0 // summx_sub_sums. + have [modM ntM _] := simM. + have linM: \rank M = 1%N. + apply/eqP; rewrite eqn_leq lt0n mxrank_eq0 ntM andbT. + rewrite -inCR_K; apply: leq_trans (mxrankM_maxl _ _) _. + apply: leq_trans (mxrankS (mulmx_sub _ im_inCR)) _. + rewrite rfix_submod //; apply: leq_trans (mxrankM_maxl _ _) _. + by rewrite -fixRlin mxrankS ?capmxSr. + apply: contra (ntM); move/eqP; rewrite -submx0 => <-. + by rewrite -(rfix_mx_rstabC rK) ?der_sub // -(rker_submod modM) rker_linear. +have{sK i M simM sK1 def_sK} irrK: mx_irreducible rK. + have cycGq: cyclic (G / K) by rewrite -defKR quotientMidl quotient_cyclic. + apply: (mx_irr_prime_index closF irrU cycGq simM) => x Gx /=. + apply: (component_mx_iso simM); first exact: Clifford_simple. + have jP: component_mx rK (M *m rU x) \in socle_enum sK. + by apply: component_socle; exact: Clifford_simple. + pose j := PackSocle jP; apply: submx_trans (_ : j <= _)%MS. + by rewrite PackSocleK component_mx_id //; exact: Clifford_simple. + have def_i: [set i] == [set: sK] by rewrite eqEcard subsetT cards1 sK1. + by rewrite ((j =P i) _) // -in_set1 (eqP def_i) inE. +pose G' := K^`(1) <*> R. +have sG'G: G' \subset G by rewrite join_subG sK'G. +pose rG' := subg_repr rU sG'G. +wlog irrG': / mx_irreducible rG'. + move=> IH; wlog [M simM sM1]: / exists2 M, mxsimple rG' M & (M <= 1%:M)%MS. + by apply: mxsimple_exists; rewrite ?mxmodule1; case: irrK. + have [modM ntM _] := simM. + have [M' modM' sumM dxM] := mx_Maschke rG' (pgroupS sG'G F'G) modM sM1. + wlog{IH} ntM': / M' != 0. + case: eqP sumM => [-> M1 _ | _ _ -> //]; apply: IH. + by apply: mx_iso_simple simM; apply: eqmx_iso; rewrite addsmx0_id in M1. + suffices: (K^`(1) \subset rstab rG' M) || (K^`(1) \subset rstab rG' M'). + rewrite !rfix_mx_rstabC ?joing_subl //; rewrite -!submx0 in ntM ntM' *. + by case/orP; move/submx_trans=> sM; apply: (contra (sM _ _)). + rewrite !rstab_subg !rstab_submod !subsetI joing_subl !rfix_mx_rstabC //. + rewrite /mxmodule !rstabs_subg !rstabs_submod !subsetI !subxx in modM modM'. + do 2!rewrite orbC -genmxE. + rewrite dx_modK_rfix // /NRmod ?(eqmx_rstabs _ (genmxE _)) ?der_sub //. + exact: subset_trans sRG nK'G. + apply/mxdirect_addsP; apply/eqP; rewrite -genmx_cap (eqmx_eq0 (genmxE _)). + rewrite -(in_submodK (submx_trans (capmxSl _ _) (val_submodP _))). + rewrite val_submod_eq0 in_submodE -submx0 (submx_trans (capmxMr _ _ _)) //. + by rewrite -!in_submodE !val_submodK (mxdirect_addsP dxM). +have nsK'K: K^`(1) <| K by apply: der_normal. +pose rK'K := subg_repr rK (normal_sub nsK'K). +have irrK'K: mx_absolutely_irreducible rK'K. + wlog sK'K: / socleType rK'K by apply: socle_exists. + have sK'_dv_K: #|[set: sK'K]| %| #|K|. + exact: atrans_dvd_in (Clifford_atrans _ _). + have nsK'G': K^`(1) <| G' := normalS (joing_subl _ _) sG'G nsK'G. + pose rK'G' := subg_repr rG' (normal_sub nsK'G'). + wlog sK'G': / socleType rK'G' by exact: socle_exists. + have coKp: coprime #|K| p := Frobenius_coprime frobG. + have nK'R := subset_trans sRG nK'G. + have sK'_dv_p: #|[set: sK'G']| %| p. + suffices: #|G' : 'C([set: sK'G'] | 'Cl)| %| #|G' : K^`(1)|. + rewrite -(divgS (joing_subl _ _)) /= {2}norm_joinEr //. + rewrite coprime_cardMg ?(coprimeSg (normal_sub nsK'K)) //. + rewrite mulKn ?cardG_gt0 // -indexgI; apply: dvdn_trans. + exact: atrans_dvd_index_in (Clifford_atrans _ _). + rewrite indexgS //; apply: subset_trans (Clifford_astab sK'G'). + exact: joing_subl. + have eq_sK': #|[set: sK'K]| = #|[set: sK'G']|. + rewrite !cardsT !cardE -!(size_map (fun i => socle_val i)). + apply: perm_eq_size. + rewrite uniq_perm_eq 1?(map_inj_uniq val_inj) 1?enum_uniq // => V. + apply/mapP/mapP=> [] [i _ ->{V}]. + exists (PackSocle (component_socle sK'G' (socle_simple i))). + by rewrite mem_enum. + by rewrite PackSocleK. + exists (PackSocle (component_socle sK'K (socle_simple i))). + by rewrite mem_enum. + by rewrite PackSocleK. + have [i def_i]: exists i, [set: sK'G'] = [set i]. + apply/cards1P; rewrite -dvdn1 -{7}(eqnP coKp) dvdn_gcd. + by rewrite -{1}eq_sK' sK'_dv_K sK'_dv_p. + pose M := socle_base i; have simM : mxsimple rK'G' M := socle_simple i. + have cycGq: cyclic (G' / K^`(1)). + by rewrite /G' joingC quotientYidr ?quotient_cyclic. + apply closF; apply: (mx_irr_prime_index closF irrG' cycGq simM) => x K'x /=. + apply: (component_mx_iso simM); first exact: Clifford_simple. + have jP: component_mx rK'G' (M *m rG' x) \in socle_enum sK'G'. + by apply: component_socle; exact: Clifford_simple. + pose j := PackSocle jP; apply: submx_trans (_ : j <= _)%MS. + by rewrite PackSocleK component_mx_id //; exact: Clifford_simple. + by rewrite ((j =P i) _) // -in_set1 -def_i inE. +have linU: \rank U = 1%N by apply/eqP; rewrite abelian_abs_irr in irrK'K. +case: irrU => _ nz1 _; apply: contra nz1; move/eqP=> fix0. +by rewrite -submx0 -fix0 -(rfix_mx_rstabC rK) ?der_sub // rker_linear. +Qed. + +(* Internal action version of B & G, Theorem 3.5. *) +Theorem Frobenius_prime_cent_prime k gT (G K R V : {group gT}) : + solvable G -> K ><| R = G -> prime #|R| -> 'C_K(R) = 1 -> + k.-abelem V -> G \subset 'N(V) -> k^'.-group G -> #|'C_V(R)| = k -> + K^`(1) \subset 'C_K(V). +Proof. +move=> solG defG prR regRK abelV nVG k'G primeRV. +have [_ sRG _ nKR _] := sdprod_context defG; rewrite subsetI der_sub. +have [-> | ntV] := eqsVneq V 1; first exact: cents1. +pose rV := abelem_repr abelV ntV nVG. +apply: subset_trans (_ : rker rV \subset _); last first. + by rewrite rker_abelem subsetIr. +have k_pr: prime k by case/pgroup_pdiv: (abelem_pgroup abelV). +apply: (Frobenius_prime_rfix1 defG) => //. + by rewrite (eq_pgroup G (eq_negn (charf_eq (char_Fp k_pr)))). +apply/eqP; rewrite rfix_abelem // -(eqn_exp2l _ _ (prime_gt1 k_pr)). +rewrite -{1}(card_Fp k_pr) -card_rowg rowg_mxK. +by rewrite card_injm ?abelem_rV_injm ?subsetIl ?primeRV. +Qed. + +Section Theorem_3_6. +(* Limit the scope of the FiniteModule notations *) +Import FiniteModule. + +(* This is B & G, Theorem 3.6. *) +Theorem odd_sdprod_Zgroup_cent_prime_plength1 p gT (G H R R0 : {group gT}) : + solvable G -> odd #|G| -> H ><| R = G -> coprime #|H| #|R| -> + R0 \subset R -> prime #|R0| -> Zgroup 'C_H(R0) -> + p.-length_1 [~: H, R]. +Proof. +move: {2}_.+1 (ltnSn #|G|) => n; elim: n => // n IHn in gT G H R R0 *. +rewrite ltnS; move oR0: #|R0| => r leGn solG oddG defG coHR sR0R r_pr ZgrCHR0. +have rR0: r.-group R0 by rewrite /pgroup oR0 pnat_id. +have [nsHG sRG mulHR nHR tiHR]:= sdprod_context defG. +have [sHG nHG] := andP nsHG; have solH := solvableS sHG solG. +have IHsub (H1 R1 : {group gT}): + H1 \subset H -> H1 * R1 \subset 'N(H1) -> R0 \subset R1 -> R1 \subset R -> + (#|H1| < #|H|) || (#|R1| < #|R|) -> p.-length_1 [~: H1, R1]. +- move=> sH1 nH1 sR01 sR1 ltG1; set G1 := H1 <*> R1. + have coHR1: coprime #|H1| #|R1| by rewrite (coprimeSg sH1) // (coprimegS sR1). + have defG1: H1 ><| R1 = G1. + by rewrite sdprodEY ?coprime_TIg ?(subset_trans (mulG_subr H1 R1)). + have sG1: G1 \subset G by rewrite join_subG -mulG_subG -mulHR mulgSS. + have{ltG1} ltG1n: #|G1| < n. + rewrite (leq_trans _ leGn) // -(sdprod_card defG1) -(sdprod_card defG). + have leqifS := leqif_geq (subset_leq_card _). + rewrite ltn_neqAle !(leqif_mul (leqifS _ _ _ sH1) (leqifS _ _ _ sR1)). + by rewrite muln_eq0 !negb_or negb_and -!ltnNge ltG1 -!lt0n !cardG_gt0. + apply: IHn defG1 _ sR01 _ _; rewrite ?oR0 ?(solvableS sG1) ?(oddSg sG1) //. + exact: ZgroupS (setSI _ sH1) ZgrCHR0. +without loss defHR: / [~: H, R] = H; last rewrite defHR. + have sHR_H: [~: H, R] \subset H by rewrite commg_subl. + have:= sHR_H; rewrite subEproper; case/predU1P=> [-> -> //|ltHR_H _]. + rewrite -coprime_commGid // IHsub 1?proper_card //. + by apply: subset_trans (commg_norm H R); rewrite norm_joinEr ?mulSg. +have{n leGn IHn tiHR} IHquo (X : {group gT}): + X :!=: 1 -> X \subset H -> G \subset 'N(X) -> p.-length_1 (H / X). +- move=> ntX sXH nXG; have nXH := subset_trans sHG nXG. + have nXR := subset_trans sRG nXG; have nXR0 := subset_trans sR0R nXR. + rewrite -defHR quotientE morphimR // -!quotientE. + have ltGbn: #|G / X| < n. + exact: leq_trans (ltn_quotient ntX (subset_trans sXH sHG)) _. + have defGb: (H / X) ><| (R / X) = G / X by exact: quotient_coprime_sdprod. + have pr_R0b: prime #|R0 / X|. + have tiXR0: X :&: R0 = 1 by apply/trivgP; rewrite -tiHR setISS. + by rewrite card_quotient // -indexgI setIC tiXR0 indexg1 oR0. + have solGb: solvable (G / X) by exact: quotient_sol. + have coHRb: coprime #|H / X| #|R / X| by exact: coprime_morph. + apply: IHn defGb coHRb _ pr_R0b _; rewrite ?quotientS ?quotient_odd //. + by rewrite -coprime_quotient_cent ?(coprimegS sR0R) // morphim_Zgroup. +without loss Op'H: / 'O_p^'(H) = 1. + have [_ -> // | ntO _] := eqVneq 'O_p^'(H) 1. + suffices: p.-length_1 (H / 'O_p^'(H)). + by rewrite p'quo_plength1 ?pcore_normal ?pcore_pgroup. + apply: IHquo => //; first by rewrite normal_sub ?pcore_normal. + by rewrite normal_norm // (char_normal_trans (pcore_char _ _)). +move defV: 'F(H)%G => V. +have charV: V \char H by rewrite -defV Fitting_char. +have [sVH nVH]: V \subset H /\ H \subset 'N(V) := andP (char_normal charV). +have nsVG: V <| G := char_normal_trans charV nsHG. +have [_ nVG] := andP nsVG; have nVR: R \subset 'N(V) := subset_trans sRG nVG. +without loss ntV: / V :!=: 1. + by rewrite -defV trivg_Fitting //; case: eqP => [|_] ->; rewrite ?plength1_1. +have scVHV: 'C_H(V) \subset V by rewrite -defV cent_sub_Fitting. +have{defV Op'H} defV: 'O_p(H) = V by rewrite -(Fitting_eq_pcore Op'H) -defV. +have pV: p.-group V by rewrite -defV pcore_pgroup. +have [p_pr p_dv_V _] := pgroup_pdiv pV ntV. +have p'r: r != p. + rewrite eq_sym -dvdn_prime2 // -prime_coprime // (coprime_dvdl p_dv_V) //. + by rewrite -oR0 (coprimegS sR0R) // (coprimeSg sVH). +without loss{charV} abelV: / p.-abelem V; last have [_ cVV eV] := and3P abelV. + move/implyP; rewrite implybE -trivg_Phi //; case/orP=> // ntPhi. + have charPhi: 'Phi(V) \char H := char_trans (Phi_char _) charV. + have nsPhiH := char_normal charPhi; have [sPhiH nPhiH] := andP nsPhiH. + have{charPhi} nPhiG: G \subset 'N('Phi(V)):= char_norm_trans charPhi nHG. + rewrite -(pquo_plength1 nsPhiH) 1?IHquo ?(pgroupS (Phi_sub _)) //. + have [/= W defW sPhiW nsWH] := inv_quotientN nsPhiH (pcore_normal p^' _). + have p'Wb: p^'.-group (W / 'Phi(V)) by rewrite -defW pcore_pgroup. + have{p'Wb} tiWb := coprime_TIg (pnat_coprime (quotient_pgroup _ _) p'Wb). + suffices pW: p.-group W by rewrite -(tiWb W pW) setIid. + apply/pgroupP=> q q_pr; case/Cauchy=> // x Wx ox; apply: wlog_neg => q'p. + suffices Vx: x \in V by rewrite (pgroupP pV) // -ox order_dvdG. + have [sWH nWH] := andP nsWH; rewrite (subsetP scVHV) // inE (subsetP sWH) //=. + have coVx: coprime #|V| #[x] by rewrite ox (pnat_coprime pV) // pnatE. + rewrite -cycle_subG (coprime_cent_Phi pV coVx) //. + have: V :&: W \subset 'Phi(V); last apply: subset_trans. + rewrite -quotient_sub1; last by rewrite subIset ?(subset_trans sWH) ?orbT. + by rewrite quotientIG ?tiWb. + rewrite commg_subI //; first by rewrite subsetI subxx (subset_trans sVH). + by rewrite cycle_subG inE Wx (subsetP nVH) // (subsetP sWH). +have{scVHV} scVH: 'C_H(V) = V by apply/eqP; rewrite eqEsubset scVHV subsetI sVH. +without loss{IHquo} indecomposableV: / forall U W, + U \x W = V -> G \subset 'N(U) :&: 'N(W) -> U = 1 \/ U = V. +- pose decV UW := let: (U, W) := UW in + [&& U \x W == V, G \subset 'N(U) :&: 'N(W), U != 1 & W != 1]. + case: (pickP decV) => [[A B /=] | indecV]; last first. + apply=> U W defUW nUW_G; have:= indecV (U, W); rewrite /= -defUW nUW_G eqxx. + by rewrite -negb_or; case/pred2P=> ->; [left | right; rewrite dprodg1]. + rewrite subsetI -!andbA => /and5P[/eqP/dprodP[[U W -> ->{A B}]]]. + move=> defUW _ tiUW nUG nWG ntU ntW _. + have [sUH sWH]: U \subset H /\ W \subset H. + by apply/andP; rewrite -mulG_subG defUW. + have [nsUH nsWH]: U <| H /\ W <| H. + by rewrite /normal !(subset_trans sHG) ?andbT. + by rewrite -(quo2_plength1 _ nsUH nsWH) ?tiUW ?IHquo. +have nsFb: 'F(H / V) <| G / V. + exact: char_normal_trans (Fitting_char _) (morphim_normal _ _). +have{nsVG nsFb} [/= U defU sVU nsUG] := inv_quotientN nsVG nsFb. +have{nsUG} [sUG nUG] := andP nsUG. +have [solU nVU] := (solvableS sUG solG, subset_trans sUG nVG). +have sUH: U \subset H by rewrite -(quotientSGK nVU sVH) -defU Fitting_sub. +have [K hallK nKR]: exists2 K : {group gT}, p^'.-Hall(U) K & R \subset 'N(K). + by apply: coprime_Hall_exists; rewrite ?(coprimeSg sUH) ?(subset_trans sRG). +have [sKU p'K _] := and3P hallK; have{sUG} sKG := subset_trans sKU sUG. +have coVK: coprime #|V| #|K| := pnat_coprime pV p'K. +have [sKH nVK] := (subset_trans sKU sUH, subset_trans sKU nVU). +have{defV} p'Ub: p^'.-group (U / V). + rewrite -defU -['F(H / V)](nilpotent_pcoreC p (Fitting_nil _)) /=. + by rewrite p_core_Fitting -defV trivg_pcore_quotient dprod1g pcore_pgroup. +have{p'Ub} sylV: p.-Sylow(U) V by rewrite /pHall sVU pV -card_quotient. +have{sKU} mulVK: V * K = U. + apply/eqP; rewrite eqEcard mul_subG //= coprime_cardMg //. + by rewrite (card_Hall sylV) (card_Hall hallK) partnC. +have [sKN sNH]: K \subset 'N_H(K) /\ 'N_H(K) \subset H. + by rewrite subsetIl subsetI sKH normG. +have [solN nVN] := (solvableS sNH solH, subset_trans sNH nVH). +have{solU hallK sUH nUG} defH: V * 'N_H(K) = H. + have nsUH: U <| H by apply/andP; rewrite (subset_trans sHG). + by rewrite -(mulSGid sKN) mulgA mulVK (Hall_Frattini_arg solU nsUH hallK). +have [P sylP nPR]: exists2 P : {group _}, p.-Sylow('N_H(K)) P & R \subset 'N(P). + apply: coprime_Hall_exists (coprimeSg sNH coHR) solN. + by rewrite normsI ?norms_norm. +have [sPN pP _]: [/\ P \subset 'N_H(K), p.-group P & _] := and3P sylP. +have [sPH nKP]: P \subset H /\ P \subset 'N(K) by apply/andP; rewrite -subsetI. +have nVP := subset_trans sPH nVH. +have coKP: coprime #|K| #|P| by rewrite coprime_sym (pnat_coprime pP). +have{sylP} sylVP: p.-Sylow(H) (V <*> P). + rewrite pHallE /= norm_joinEr ?mul_subG //= -defH -!LagrangeMl. + rewrite partnM // part_pnat_id // -!card_quotient //. + by apply/eqP; congr (_ * _)%N; apply: card_Hall; exact: quotient_pHall. +have [trKP | {sylV sVU nVU}ntKP] := eqVneq [~: K, P] 1. + suffices sylVH: p.-Sylow(H) V. + rewrite p_elt_gen_length1 // (_ : p_elt_gen p H = V). + rewrite /pHall pcore_sub pcore_pgroup /= pnatNK. + by apply: pnat_dvd pV; exact: dvdn_indexg. + rewrite -(genGid V) -(setIidPr sVH); congr <<_>>; apply/setP=> x. + rewrite !inE; apply: andb_id2l => Hx. + by rewrite (mem_normal_Hall sylVH) /normal ?sVH. + suffices sPV: P \subset V by rewrite -(joing_idPl sPV). + suffices sPU: P \subset U by rewrite (sub_normal_Hall sylV) //; exact/andP. + have cUPb: P / V \subset 'C_(H / V)(U / V). + rewrite subsetI morphimS // -mulVK quotientMidl quotient_cents2r //. + by rewrite commGC trKP sub1G. + rewrite -(quotientSGK nVP sVU) (subset_trans cUPb) //. + by rewrite -defU cent_sub_Fitting ?quotient_sol. +have{sylVP} dxV: [~: V, K] \x 'C_V(K) = V by exact: coprime_abelian_cent_dprod. +have tiVsub_VcK: 'C_V(K) = 1 \/ 'C_V(K) = V. + apply: (indecomposableV _ [~: V, K]); first by rewrite dprodC. + rewrite -mulHR -defH -mulgA mul_subG // subsetI. + by rewrite commg_norml cents_norm // centsC subIset // -abelianE cVV. + have nK_NR: 'N_H(K) * R \subset 'N(K) by rewrite mul_subG ?subsetIr. + have nV_NR: 'N_H(K) * R \subset 'N(V) by rewrite mul_subG. + by rewrite normsR // normsI ?norms_cent. +have{tiVsub_VcK dxV} [defVK tiVcK]: [~: V, K] = V /\ 'C_V(K) = 1. + have [tiVcK | eqC] := tiVsub_VcK; first by rewrite -{2}dxV // tiVcK dprodg1. + rewrite (card1_trivg (pnat_1 (pgroupS _ pV) p'K)) ?comm1G ?eqxx // in ntKP. + by rewrite -scVH subsetI sKH centsC -eqC subsetIr. +have eqVncK: 'N_V(K) = 'C_V(K) := coprime_norm_cent nVK (pnat_coprime pV p'K). +have{eqVncK} tiVN: V :&: 'N_H(K) = 1 by rewrite setIA (setIidPl sVH) eqVncK. +have{sPN} tiVP: V :&: P = 1 by apply/trivgP; rewrite -tiVN setIS. +have{U defU mulVK} defK: 'F('N_H(K)) = K. + have [injV imV] := isomP (quotient_isom nVN tiVN). + rewrite -(im_invm injV) -injm_Fitting ?injm_invm //= {2}imV /=. + rewrite -quotientMidl defH defU -mulVK quotientMidl morphim_invmE. + by rewrite morphpre_restrm quotientK // -group_modr // setIC tiVN mul1g. +have scKH: 'C_H(K) \subset K. + rewrite -{2}defK; apply: subset_trans (cent_sub_Fitting _) => //. + by rewrite defK subsetI subsetIr setIS // cent_sub. +have{nVN} ntKR0: [~: K, R0] != 1. + rewrite (sameP eqP commG1P); apply: contra ntKP => cR0K. + have ZgrK: Zgroup K by apply: ZgroupS ZgrCHR0; rewrite subsetI sKH. + have{ZgrK} cycK: cyclic K by rewrite nil_Zgroup_cyclic // -defK Fitting_nil. + have{cycK} sNR_K: [~: 'N_H(K), R] \subset K. + apply: subset_trans scKH; rewrite subsetI; apply/andP; split. + by rewrite (subset_trans (commSg R sNH)) // commGC commg_subr. + suffices: 'N(K)^`(1) \subset 'C(K). + by apply: subset_trans; rewrite commgSS ?subsetIr. + rewrite der1_min ?cent_norm //= -ker_conj_aut (isog_abelian (first_isog _)). + exact: abelianS (Aut_conj_aut K 'N(K)) (Aut_cyclic_abelian cycK). + suffices sPV: P \subset V by rewrite -(setIidPr sPV) tiVP commG1. + have pPV: p.-group (P / V) := quotient_pgroup V pP. + rewrite -quotient_sub1 // subG1 (card1_trivg (pnat_1 pPV _)) //. + apply: pgroupS (quotient_pgroup V p'K). + apply: subset_trans (quotientS V sNR_K). + by rewrite quotientR // -quotientMidl defH -quotientR ?defHR ?quotientS. +have nKR0: R0 \subset 'N(K) := subset_trans sR0R nKR. +have mulKR0: K * R0 = K <*> R0 by rewrite norm_joinEr. +have sKR0_G : K <*> R0 \subset G by rewrite -mulKR0 -mulHR mulgSS. +have nV_KR0: K <*> R0 \subset 'N(V) := subset_trans sKR0_G nVG. +have solKR0: solvable (K <*> R0) by exact: solvableS solG. +have coKR0: coprime #|K| #|R0| by rewrite (coprimeSg sKH) ?(coprimegS sR0R). +have r'K: r^'.-group K. + by rewrite /pgroup p'natE -?prime_coprime // coprime_sym -oR0. +have tiKcV: 'C_K(V) = 1. + by apply/trivgP; rewrite -tiVN -{2}scVH -setIIr setICA setIC setSI. +have tiKR0cV: 'C_(K <*> R0)(V) = 1. + set C := 'C_(K <*> R0)(V); apply/eqP; apply: contraR ntKR0 => ntC. + have nC_KR0: K <*> R0 \subset 'N(C) by rewrite normsI ?normG ?norms_cent. + rewrite -subG1 -(coprime_TIg coKR0) commg_subI ?subsetI ?subxx //=. + suff defC: C == R0 by rewrite -(eqP defC) (subset_trans (joing_subl K R0)). + have sC_R0: C \subset R0. + rewrite -[C](coprime_mulG_setI_norm mulKR0) ?norms_cent //= tiKcV mul1g. + exact: subsetIl. + rewrite eqEsubset sC_R0; apply: contraR ntC => not_sR0C. + by rewrite -(setIidPr sC_R0) prime_TIg ?oR0. +have{nKR0 mulKR0 sKR0_G solKR0 nV_KR0} oCVR0: #|'C_V(R0)| = p. + case: (eqVneq 'C_V(R0) 1) => [tiVcR0 | ntCVR0]. + case/negP: ntKR0; rewrite -subG1/= commGC -tiKcV. + have defKR0: K ><| R0 = K <*> R0 by rewrite sdprodE ?coprime_TIg. + have odd_KR0: odd #|K <*> R0| := oddSg sKR0_G oddG. + apply: odd_prime_sdprod_abelem_cent1 abelV nV_KR0 _ _; rewrite // ?oR0 //=. + by rewrite -mulKR0 pgroupM p'K /pgroup oR0 pnatE. + have [x defC]: exists x, 'C_V(R0) = <[x]>. + have ZgrC: Zgroup 'C_V(R0) by apply: ZgroupS ZgrCHR0; exact: setSI. + apply/cyclicP; apply: (forall_inP ZgrC); apply/SylowP; exists p => //. + by rewrite /pHall subxx indexgg (pgroupS (subsetIl V _)). + rewrite defC; apply: nt_prime_order => //; last by rewrite -cycle_eq1 -defC. + by rewrite (exponentP eV) // -cycle_subG -defC subsetIl. +have tiPcR0: 'C_P(R0) = 1. + rewrite -(setIidPl (joing_subl P V)) setIIl TI_Ohm1 //=. + set C := 'C_(P <*> V)(R0); suffices <-: 'C_V(R0) = 'Ohm_1(C). + by rewrite setIC -setIIl tiVP (setIidPl (sub1G _)). + have pPV: p.-group (P <*> V) by rewrite norm_joinEl // pgroupM pP. + have pC: p.-group C := pgroupS (subsetIl _ _) pPV. + have abelCVR0: p.-abelem 'C_V(R0) by rewrite prime_abelem ?oCVR0. + have sCV_C: 'C_V(R0) \subset C by rewrite setSI ?joing_subr. + apply/eqP; rewrite eqEcard -(Ohm1_id abelCVR0) OhmS //=. + have [-> | ntC] := eqVneq C 1; first by rewrite subset_leq_card ?OhmS ?sub1G. + rewrite (Ohm1_id abelCVR0) oCVR0 (Ohm1_cyclic_pgroup_prime _ pC) //=. + have ZgrC: Zgroup C by rewrite (ZgroupS _ ZgrCHR0) ?setSI // join_subG sPH. + apply: (forall_inP ZgrC); apply/SylowP; exists p => //. + by apply/pHallP; rewrite part_pnat_id. +have defP: [~: P, R0] = P. + have solvP := pgroup_sol pP; have nPR0 := subset_trans sR0R nPR. + have coPR0: coprime #|P| #|R0| by rewrite (coprimeSg sPH) ?(coprimegS sR0R). + by rewrite -{2}(coprime_cent_prod nPR0) // tiPcR0 mulg1. +have{IHsub nVH} IHsub: forall X : {group gT}, + P <*> R0 \subset 'N(X) -> X \subset K -> + (#|V <*> X <*> P| < #|H|) || (#|R0| < #|R|) -> [~: X, P] = 1. +- move=> X; rewrite join_subG; case/andP=> nXP nXR0 sXK. + set H0 := V <*> X <*> P => ltG0G; have sXH := subset_trans sXK sKH. + have sXH0: X \subset H0 by rewrite /H0 joingC joingA joing_subr. + have sH0H: H0 \subset H by rewrite !join_subG sVH sXH. + have nH0R0: R0 \subset 'N(H0). + by rewrite 2?normsY ?nXR0 ?(subset_trans sR0R) // (subset_trans sRG). + have Op'H0: 'O_p^'(H0) = 1. + have [sOp' nOp'] := andP (pcore_normal _ _ : 'O_p^'(H0) <| H0). + have p'Op': p^'.-group 'O_p^'(H0) by exact: pcore_pgroup. + apply: card1_trivg (pnat_1 (pgroupS _ pV) p'Op'). + rewrite -scVH subsetI (subset_trans sOp') //= centsC; apply/setIidPl. + rewrite -coprime_norm_cent ?(pnat_coprime pV p'Op') //. + by rewrite (setIidPl (subset_trans _ nOp')) // /H0 -joingA joing_subl. + exact: subset_trans (subset_trans sH0H nVH). + have Op'HR0: 'O_p^'([~: H0, R0]) = 1. + apply/trivgP; rewrite -Op'H0 pcore_max ?pcore_pgroup //. + apply: char_normal_trans (pcore_char _ _) _. + by rewrite /(_ <| _) commg_norml andbT commg_subl. + have{ltG0G IHsub} p1_HR0: p.-length_1 [~: H0, R0]. + by apply: IHsub ltG0G => //=; rewrite mul_subG ?normG. + have{p1_HR0} sPOpHR0: P \subset 'O_p([~: H0, R0]). + rewrite sub_Hall_pcore //; last by rewrite -defP commSg ?joing_subr. + rewrite /pHall pcore_sub pcore_pgroup /= -(pseries_pop2 _ Op'HR0). + rewrite -card_quotient ?normal_norm ?pseries_normal // -/(pgroup _ _). + by rewrite -{1}((_ :=P: _) p1_HR0) (quotient_pseries [::_;_]) pcore_pgroup. + apply/trivgP; have <-: K :&: 'O_p([~: H0, R0]) = 1. + by rewrite setIC coprime_TIg // (pnat_coprime (pcore_pgroup p _)). + rewrite commg_subI // subsetI ?sPOpHR0 ?sXK //=. + by rewrite (char_norm_trans (pcore_char _ _)) // normsRl. +have{defH sR0R} [defH defR0]: V * K * P = H /\ R0 :=: R. + suffices: (V * K * P == H) && (R0 :==: R) by do 2!case: eqP => // ->. + apply: contraR ntKP; rewrite -subG1 !eqEcard sR0R ?mul_subG //= negb_and. + rewrite -!ltnNge -!norm_joinEr // 1?normsY //; move/IHsub=> -> //. + by rewrite join_subG nKP (subset_trans sR0R). +move: IHsub defP oR0 rR0 ZgrCHR0 coKR0 ntKR0 tiKR0cV oCVR0 tiPcR0. +rewrite {R0}defR0 ltnn => IHsub defP oR rR ZgrCHR coKR ntKR tiKRcV oCVR tiPcR. +have mulVK: V * K = V <*> K by rewrite norm_joinEr. +have oVK: #|V <*> K| = (#|V| * #|K|)%N by rewrite -mulVK coprime_cardMg. +have tiVK_P: V <*> K :&: P = 1. + have sylV: p.-Sylow(V <*> K) V. + by rewrite /pHall pV -divgS joing_subl //= oVK mulKn. + apply/trivgP; rewrite -tiVP subsetI subsetIr. + rewrite (sub_normal_Hall sylV) ?subsetIl ?(pgroupS (subsetIr _ P)) //=. + by rewrite /normal joing_subl join_subG normG. +have{mulVK oVK} oH: (#|H| = #|V| * #|K| * #|P|)%N. + by rewrite -defH mulVK -oVK (TI_cardMg tiVK_P). +have{oH tiVK_P IHsub} IHsub: forall X : {group gT}, + P <*> R \subset 'N(X) -> X \subset K -> X :=: K \/ X \subset 'C(P). +- move=> X nX_PR sXK; have p'X: p^'.-group X := pgroupS sXK p'K. + have nXP: P \subset 'N(X) := subset_trans (joing_subl P R) nX_PR. + apply/predU1P; rewrite eqEcard sXK; case: leqP => //= ltXK. + apply/commG1P; rewrite {}IHsub // orbF (norm_joinEr (normsY _ _)) //=. + rewrite TI_cardMg /=; last first. + by apply/trivgP; rewrite -tiVK_P setSI ?genS ?setUS. + rewrite oH ltn_pmul2r ?cardG_gt0 // norm_joinEr ?(subset_trans sXK) //. + by rewrite coprime_cardMg ?ltn_pmul2l ?(pnat_coprime pV). +have defKP: [~: K, P] = K. + have sKP_K: [~: K, P] \subset K by rewrite commg_subl. + have{sKP_K} [|//|cP_KP] := IHsub _ _ sKP_K. + by rewrite join_subG /= commg_normr normsR. + by case/eqP: ntKP; rewrite -coprime_commGid ?(commG1P cP_KP) ?(solvableS sKH). +have nKPR: P <*> R \subset 'N(K) by rewrite join_subG nKP. +have coPR: coprime #|P| #|R| by rewrite (coprimeSg sPH). +have{scKH} tiPRcK: 'C_(P <*> R)(K) = 1. + have tiPK: P :&: K = 1 by rewrite setIC coprime_TIg. + have tiPcK: 'C_P(K) = 1. + by apply/trivgP; rewrite /= -{1}(setIidPl sPH) -setIA -tiPK setIS. + have tiRcK: 'C_R(K) = 1. + by rewrite prime_TIg ?oR // centsC (sameP commG1P eqP). + have mulPR: P * R = P <*> R by rewrite norm_joinEr. + by rewrite -(coprime_mulG_setI_norm mulPR) ?tiPcK ?mul1g ?norms_cent. +have [K1 | ntK]:= eqsVneq K 1; first by rewrite K1 comm1G eqxx in ntKR. +have [K1 | [q q_pr q_dv_K]] := trivgVpdiv K; first by case/eqP: ntK. +have q_gt1 := prime_gt1 q_pr. +have p'q: q != p by exact: (pgroupP p'K). +have{r'K} q'r: r != q by rewrite eq_sym; exact: (pgroupP r'K). +have{defK} qK: q.-group K. + have{defK} nilK: nilpotent K by rewrite -defK Fitting_nil. + have{nilK} [_ defK _ _] := dprodP (nilpotent_pcoreC q nilK). + have{IHsub} IHpi: forall pi, 'O_pi(K) = K \/ 'O_pi(K) \subset 'C(P). + move=> pi; apply: IHsub (pcore_sub _ _). + by apply: char_norm_trans (pcore_char _ _) _; rewrite join_subG nKP. + case: (IHpi q) => [<-| cPKq]; first exact: pcore_pgroup. + case/eqP: ntKP; apply/commG1P; rewrite -{}defK mul_subG //. + case: (IHpi q^') => // defK; case/idPn: q_dv_K. + rewrite -p'natE // -defK; exact: pcore_pgroup. +pose K' := K^`(1); have charK': K' \char K := der_char 1 K. +have nsK'K: K' <| K := der_normal 1 K; have [sK'K nK'K] := andP nsK'K. +have nK'PR: P <*> R \subset 'N(K') := char_norm_trans charK' nKPR. +have iK'K: 'C_(P <*> R / K')(K / K') = 1 -> #|K / K'| > q ^ 2. + have qKb: q.-group (K / K') by exact: morphim_pgroup qK. + rewrite ltnNge => trCK'; apply: contra ntKP => Kq_le_q2. + suffices sPR_K': [~: P, R] \subset K'. + rewrite -defP -(setIidPl sPR_K') coprime_TIg ?commG1 //. + by rewrite (pnat_coprime (pgroupS _ pP) (pgroupS sK'K p'K)) ?commg_subl. + rewrite -quotient_cents2 ?(char_norm_trans charK') //. + suffices cPRbPrb: abelian (P <*> R / K'). + by rewrite (sub_abelian_cent2 cPRbPrb) ?quotientS ?joing_subl ?joing_subr. + have nKbPR: P <*> R / K' \subset 'N(K / K') by exact: quotient_norms. + case cycK: (cyclic (K / K')). + rewrite (isog_abelian (quotient1_isog _)) -trCK' -ker_conj_aut. + rewrite (isog_abelian (first_isog_loc _ _)) //. + by rewrite (abelianS (Aut_conj_aut _ _)) ?Aut_cyclic_abelian. + have{cycK} [oKb abelKb]: #|K / K'| = (q ^ 2)%N /\ q.-abelem (K / K'). + have sKb1: 'Ohm_1(K / K') \subset K / K' by exact: Ohm_sub. + have cKbKb: abelian (K / K') by rewrite sub_der1_abelian. + have: #|'Ohm_1(K / K')| >= q ^ 2. + rewrite (card_pgroup (pgroupS sKb1 qKb)) leq_exp2l // ltnNge. + by rewrite -p_rank_abelian -?rank_pgroup // -abelian_rank1_cyclic ?cycK. + rewrite (geq_leqif (leqif_trans (subset_leqif_card sKb1) (leqif_eq _))) //. + by case/andP=> sKbKb1; move/eqP->; rewrite (abelemS sKbKb1) ?Ohm1_abelem. + have ntKb: K / K' != 1 by rewrite -cardG_gt1 oKb (ltn_exp2l 0). + pose rPR := abelem_repr abelKb ntKb nKbPR. + have: mx_faithful rPR by rewrite abelem_mx_faithful. + move: rPR; rewrite (dim_abelemE abelKb ntKb) oKb pfactorK // => rPR ffPR. + apply: charf'_GL2_abelian ffPR _. + by rewrite quotient_odd ?(oddSg _ oddG) // join_subG (subset_trans sPH). + rewrite (eq_pgroup _ (eq_negn (charf_eq (char_Fp q_pr)))). + rewrite quotient_pgroup //= norm_joinEr // pgroupM. + by rewrite /pgroup (pi_pnat rR) // (pi_pnat pP) // !inE eq_sym. +case cKK: (abelian K); last first. + have [|[dPhiK dK'] dCKP] := abelian_charsimple_special qK coKP defKP. + apply/bigcupsP=> L; case/andP=> charL; have sLK := char_sub charL. + by case/IHsub: sLK cKK => // [|-> -> //]; exact: (char_norm_trans charL). + have eK: exponent K %| q. + have oddK: odd #|K| := oddSg sKG oddG. + have [Q [charQ _ _ eQ qCKQ]] := critical_odd qK oddK ntK; rewrite -eQ. + have sQK: Q \subset K := char_sub charQ. + have [<- // | cQP] := IHsub Q (char_norm_trans charQ nKPR) sQK. + case/negP: ntKP; rewrite (sameP eqP commG1P) centsC. + rewrite -ker_conj_aut -sub_morphim_pre // subG1 trivg_card1. + rewrite (pnat_1 (morphim_pgroup _ pP) (pi_pnat (pgroupS _ qCKQ) _)) //. + apply/subsetP=> a; case/morphimP=> x nKx Px ->{a}. + rewrite /= astab_ract inE /= Aut_aut; apply/astabP=> y Qy. + rewrite [_ y _]norm_conj_autE ?(subsetP sQK) //. + by rewrite /conjg (centsP cQP y) ?mulKg. + have tiPRcKb: 'C_(P <*> R / K')(K / K') = 1. + rewrite -quotient_astabQ -quotientIG /=; last first. + by rewrite sub_astabQ normG trivg_quotient sub1G. + apply/trivgP; rewrite -quotient1 quotientS // -tiPRcK subsetI subsetIl /=. + rewrite (coprime_cent_Phi qK) ?(coprimegS (subsetIl _ _)) //=. + by rewrite norm_joinEr // coprime_cardMg // coprime_mulr coKP. + rewrite dPhiK -dK' -/K' (subset_trans (commgS _ (subsetIr _ _))) //. + by rewrite astabQ -quotient_cents2 ?subsetIl // cosetpreK centsC /=. + have [nK'P nK'R] := (char_norm_trans charK' nKP, char_norm_trans charK' nKR). + have solK: solvable K := pgroup_sol qK. + have dCKRb: 'C_K(R) / K' = 'C_(K / K')(R / K'). + by rewrite coprime_quotient_cent. + have abelKb: q.-abelem (K / K') by rewrite [K']dK' -dPhiK Phi_quotient_abelem. + have [qKb cKbKb _] := and3P abelKb. + have [tiKcRb | ntCKRb]:= eqVneq 'C_(K / K')(R / K') 1. + have coK'P: coprime #|K'| #|P| by rewrite (coprimeSg sK'K). + suffices sPK': P \subset K'. + by case/negP: ntKP; rewrite -(setIidPr sPK') coprime_TIg ?commG1. + rewrite -quotient_sub1 // -defP commGC quotientR //= -/K'. + have <-: 'C_(P / K')(K / K') = 1. + by apply/trivgP; rewrite -tiPRcKb setSI ?morphimS ?joing_subl. + have q'P: q^'.-group P by rewrite /pgroup (pi_pnat pP) // !inE eq_sym. + move: tiKcRb; have: q^'.-group (P <*> R / K'). + rewrite quotient_pgroup //= norm_joinEr //. + by rewrite pgroupM q'P /pgroup oR pnatE. + have sPRG: P <*> R \subset G by rewrite join_subG sRG (subset_trans sPH). + have coPRb: coprime #|P / K'| #|R / K'| by rewrite coprime_morph. + apply: odd_prime_sdprod_abelem_cent1 abelKb _; rewrite ?quotient_norms //. + - by rewrite quotient_sol // (solvableS sPRG). + - by rewrite quotient_odd // (oddSg sPRG). + - by rewrite /= quotientY // sdprodEY ?quotient_norms ?coprime_TIg. + rewrite -(card_isog (quotient_isog nK'R _)) ?oR //. + by rewrite coprime_TIg // (coprimeSg sK'K). + have{ntCKRb} not_sCKR_K': ~~ ('C_K(R) \subset K'). + by rewrite -quotient_sub1 ?subIset ?nK'K // dCKRb subG1. + have oCKR: #|'C_K(R)| = q. + have [x defCKR]: exists x, 'C_K(R) = <[x]>. + have ZgrCKR: Zgroup 'C_K(R) := ZgroupS (setSI _ sKH) ZgrCHR. + have qCKR: q.-group 'C_K(R) by rewrite (pgroupS (subsetIl K _)). + by apply/cyclicP; exact: nil_Zgroup_cyclic (pgroup_nil qCKR). + have Kx: x \in K by rewrite -cycle_subG -defCKR subsetIl. + rewrite defCKR cycle_subG in not_sCKR_K' *. + exact: nt_prime_order (exponentP eK x Kx) (group1_contra not_sCKR_K'). + have tiCKR_K': 'C_K(R) :&: K' = 1 by rewrite prime_TIg ?oCKR. + have sKR_K: [~: K, R] \subset K by rewrite commg_subl nKR. + have ziKRcR: 'C_K(R) :&: [~: K, R] \subset K'. + rewrite -quotient_sub1 ?subIset ?nK'K // setIC. + rewrite (subset_trans (quotientI _ _ _)) // dCKRb setIA. + rewrite (setIidPl (quotientS _ sKR_K)) // ?quotientR //= -/K'. + by rewrite coprime_abel_cent_TI ?quotient_norms ?coprime_morph. + have not_sK_KR: ~~ (K \subset [~: K, R]). + by apply: contra not_sCKR_K' => sK_KR; rewrite -{1}(setIidPl sK_KR) setIAC. + have tiKRcR: 'C_[~: K, R](R) = 1. + rewrite -(setIidPr sKR_K) setIAC -(setIidPl ziKRcR) setIAC tiCKR_K'. + by rewrite (setIidPl (sub1G _)). + have cKR_KR: abelian [~: K, R]. + have: 'C_[~: K, R](V) \subset [1]. + rewrite -tiVN -{2}scVH -setIIr setICA setIC setIS //. + exact: subset_trans sKR_K sKN. + rewrite /abelian (sameP commG1P trivgP) /= -derg1; apply: subset_trans. + have nKR_R: R \subset 'N([~: K, R]) by rewrite commg_normr. + have sKRR_G: [~: K, R] <*> R \subset G by rewrite join_subG comm_subG. + move: oCVR; have: p^'.-group ([~: K, R] <*> R). + by rewrite norm_joinEr // pgroupM (pgroupS sKR_K p'K) /pgroup oR pnatE. + have solKR_R := solvableS sKRR_G solG. + apply: Frobenius_prime_cent_prime; rewrite ?oR ?(subset_trans _ nVG) //. + by rewrite sdprodEY // coprime_TIg // (coprimeSg sKR_K). + case nKR_P: (P \subset 'N([~: K, R])). + have{nKR_P} nKR_PR: P <*> R \subset 'N([~: K, R]). + by rewrite join_subG nKR_P commg_normr. + have{nKR_PR} [dKR | cP_KR] := IHsub _ nKR_PR sKR_K. + by rewrite dKR subxx in not_sK_KR. + have{cP_KR} cKRb: R / K' \subset 'C(K / K'). + by rewrite quotient_cents2r //= dK' -dCKP commGC subsetI sKR_K. + case/negP: ntKR; rewrite (sameP eqP commG1P) centsC. + by rewrite (coprime_cent_Phi qK) // dPhiK -dK' commGC -quotient_cents2. + have{nKR_P} [x Px not_nKRx] := subsetPn (negbT nKR_P). + have iKR: #|K : [~: K, R]| = q. + rewrite -divgS // -{1}(coprime_cent_prod nKR) // TI_cardMg ?mulKn //. + by rewrite setIA (setIidPl sKR_K). + have sKRx_K: [~: K, R] :^ x \subset K by rewrite -{2}(normsP nKP x Px) conjSg. + have nKR_K: K \subset 'N([~: K, R]) by exact: commg_norml. + have mulKR_Krx: [~: K, R] * [~: K, R] :^ x = K. + have maxKR: maximal [~: K, R] K by rewrite p_index_maximal ?iKR. + apply: mulg_normal_maximal; rewrite ?(p_maximal_normal qK) //. + by rewrite inE in not_nKRx. + have ziKR_KRx: [~: K, R] :&: [~: K, R] :^ x \subset K'. + rewrite /K' dK' subsetI subIset ?sKR_K // -{3}mulKR_Krx centM centJ. + by rewrite setISS ?conjSg. + suffices: q ^ 2 >= #|K / K'| by rewrite leqNgt iK'K. + rewrite -divg_normal // leq_divLR ?cardSg //. + rewrite -(@leq_pmul2l (#|[~: K, R]| ^ 2)) ?expn_gt0 ?cardG_gt0 // mulnA. + rewrite -expnMn -iKR Lagrange // -mulnn -{2}(cardJg _ x) mul_cardG. + by rewrite mulKR_Krx mulnAC leq_pmul2l ?muln_gt0 ?cardG_gt0 ?subset_leq_card. +have tiKcP: 'C_K(P) = 1 by rewrite -defKP coprime_abel_cent_TI. +have{IHsub} abelK: q.-abelem K. + have [|cPK1] := IHsub _ (char_norm_trans (Ohm_char 1 K) nKPR) (Ohm_sub 1 K). + by move/abelem_Ohm1P->. + rewrite -(setIid K) TI_Ohm1 ?eqxx // in ntK. + by apply/eqP; rewrite -subG1 -tiKcP setIS. +have{K' iK'K charK' nsK'K sK'K nK'K nK'PR} oKq2: q ^ 2 < #|K|. + have K'1: K' :=: 1 by exact/commG1P. + rewrite -indexg1 -K'1 -card_quotient ?normal_norm // iK'K // K'1. + by rewrite -injm_subcent ?coset1_injm ?norms1 //= tiPRcK morphim1. +pose S := [set Vi : {group gT} | 'C_V('C_K(Vi)) == Vi & maximal 'C_K(Vi) K]. +have defSV Vi: Vi \in S -> 'C_V('C_K(Vi)) = Vi by rewrite inE; case: eqP. +have maxSK Vi: Vi \in S -> maximal 'C_K(Vi) K by case/setIdP. +have sSV Vi: Vi \in S -> Vi \subset V by move/defSV <-; rewrite subsetIl. +have ntSV Vi: Vi \in S -> Vi :!=: 1. + move=> Si; apply: contraTneq (maxgroupp (maxSK _ Si)) => ->. + by rewrite /= cent1T setIT proper_irrefl. +have nSK Vi: Vi \in S -> K \subset 'N(Vi). + by move/defSV <-; rewrite normsI ?norms_cent // sub_abelian_norm ?subsetIl. +have defV: <<\bigcup_(Vi in S) Vi>> = V. + apply/eqP; rewrite eqEsubset gen_subG. + apply/andP; split; first by apply/bigcupsP; apply: sSV. + rewrite -(coprime_abelian_gen_cent cKK nVK) ?(pnat_coprime pV) // gen_subG. + apply/bigcupsP=> Kj /= /and3P[cycKbj sKjK nKjK]. + have [xb defKbj] := cyclicP cycKbj. + have Kxb: xb \in K / Kj by rewrite defKbj cycle_id. + set Vj := 'C_V(Kj); have [-> | ntVj] := eqsVneq Vj 1; first exact: sub1G. + have nt_xb: xb != 1. + apply: contra ntVj; rewrite -cycle_eq1 -defKbj -!subG1 -tiVcK. + by rewrite quotient_sub1 // => sKKj; rewrite setIS ?centS. + have maxKj: maximal Kj K. + rewrite p_index_maximal // -card_quotient // defKbj -orderE. + by rewrite (abelem_order_p (quotient_abelem Kj abelK) Kxb nt_xb). + suffices defKj: 'C_K(Vj) = Kj. + by rewrite sub_gen // (bigcup_max 'C_V(Kj))%G // inE defKj eqxx. + have{maxKj} [_ maxKj] := maxgroupP maxKj. + rewrite ['C_K(Vj)]maxKj //; last by rewrite subsetI sKjK centsC subsetIr. + rewrite properEneq subsetIl andbT (sameP eqP setIidPl) centsC. + by apply: contra ntVj; rewrite -subG1 -tiVcK subsetI subsetIl. +pose dxp := [fun D : {set {group gT}} => \big[dprod/1]_(Vi in D) Vi]. +have{defV} defV: \big[dprod/1]_(Vi in S) Vi = V. + have [D maxD]: {D | maxset [pred E | group_set (dxp E) & E \subset S] D}. + by apply: ex_maxset; exists set0; rewrite /= sub0set big_set0 groupP. + have [gW sDS] := andP (maxsetp maxD); have{maxD} [_ maxD] := maxsetP maxD. + have{gW} [W /= defW]: {W : {group gT} | dxp D = W} by exists (Group gW). + have [eqDS | ltDS] := eqVproper sDS. + by rewrite eqDS in defW; rewrite defW -(bigdprodWY defW). + have{ltDS} [_ [Vi Si notDi]] := properP ltDS. + have sWV: W \subset V. + rewrite -(bigdprodWY defW) gen_subG. + by apply/bigcupsP=> Vj Dj; rewrite sSV ?(subsetP sDS). + suffices{maxD sWV defV} tiWcKi: 'C_W('C_K(Vi)) = 1. + have:= notDi; rewrite -(maxD (Vi |: D)) ?setU11 ?subsetUr //= subUset sDS. + rewrite sub1set Si big_setU1 //= defW dprodEY ?groupP //. + by rewrite (sub_abelian_cent2 cVV) // sSV. + by rewrite -(defSV Vi Si) setIAC (setIidPr sWV). + apply/trivgP/subsetP=> w /setIP[Ww cKi_w]. + have [v [Vv def_w v_uniq]] := mem_bigdprod defW Ww. + rewrite def_w big1 ?inE // => Vj Dj; have Sj := subsetP sDS Vj Dj. + have cKi_vj: v Vj \in 'C('C_K(Vi)). + apply/centP=> x Ki_x; apply/commgP/conjg_fixP. + apply: (v_uniq (fun Vk => v Vk ^ x)) => // [Vk Dk|]. + have [[Kx _] Sk]:= (setIP Ki_x, subsetP sDS Vk Dk). + by rewrite memJ_norm ?Vv // (subsetP (nSK Vk Sk)). + rewrite -(mulKg x w) -(centP cKi_w) // -conjgE def_w. + by apply: (big_morph (conjg^~ x)) => [y z|]; rewrite ?conj1g ?conjMg. + suffices mulKji: 'C_K(Vj) * 'C_K(Vi) = K. + by apply/set1gP; rewrite -tiVcK -mulKji centM setIA defSV // inE Vv. + have maxKj := maxSK Vj Sj; have [_ maxKi] := maxgroupP (maxSK Vi Si). + rewrite (mulg_normal_maximal _ maxKj) -?sub_abelian_normal ?subsetIl //. + have [eqVji|] := eqVneq Vj Vi; first by rewrite -eqVji Dj in notDi. + apply: contra => /= sKiKj; rewrite -val_eqE /= -(defSV Vj Sj). + by rewrite (maxKi _ (maxgroupp maxKj) sKiKj) defSV. +have nVPR: P <*> R \subset 'N(V) by rewrite join_subG nVP. +have actsPR: [acts P <*> R, on S | 'JG]. + apply/subsetP=> x PRx; rewrite !inE; apply/subsetP=> Vi. + rewrite !inE /= => Si; rewrite -(normsP nKPR x PRx) !centJ -!conjIg centJ . + by rewrite -(normsP nVPR x PRx) -conjIg (inj_eq (@conjsg_inj _ _)) maximalJ. +have transPR: [transitive P <*> R, on S | 'JG]. + pose ndxp D (U A B : {group gT}) := dxp (S :&: D) = U -> A * B \subset 'N(U). + have nV_VK D U: ndxp D U V K. + move/bigdprodWY <-; rewrite norms_gen ?norms_bigcup //. + apply/bigcapsP=> Vi /setIP[Si _]. + by rewrite mulG_subG nSK // sub_abelian_norm // sSV. + have nV_PR D U: [acts P <*> R, on S :&: D | 'JG] -> ndxp D U P R. + move=> actsU /bigdprodWY<-; rewrite -norm_joinEr ?norms_gen //. + apply/subsetP=> x PRx; rewrite inE sub_conjg; apply/bigcupsP=> Vi Di. + by rewrite -sub_conjg (bigcup_max (Vi :^ x)%G) //= (acts_act actsU). + have [S0 | [V1 S1]] := set_0Vmem S. + by case/eqP: ntV; rewrite -defV S0 big_set0. + apply/imsetP; exists V1 => //; set D := orbit _ _ _. + rewrite (big_setID D) /= setDE in defV. + have [[U W defU defW] _ _ tiUW] := dprodP defV. + rewrite defU defW in defV tiUW. + have [|U1|eqUV]:= indecomposableV _ _ defV. + - rewrite -mulHR -defH -mulgA mul_subG //. + by rewrite subsetI (nV_VK _ _ defU) (nV_VK _ _ defW). + rewrite subsetI (nV_PR _ _ _ defU) ?actsI ?acts_orbit ?subsetT //=. + by rewrite (nV_PR _ _ _ defW) // actsI ?astabsC ?acts_orbit ?subsetT /=. + - case/negP: (ntSV V1 S1); rewrite -subG1 -U1 -(bigdprodWY defU) sub_gen //. + by rewrite (bigcup_max V1) // inE S1 orbit_refl. + apply/eqP; rewrite eqEsubset (acts_sub_orbit _ actsPR) S1 andbT. + apply/subsetP=> Vi Si; apply: contraR (ntSV Vi Si) => D'i; rewrite -subG1. + rewrite -tiUW eqUV subsetI sSV // -(bigdprodWY defW). + by rewrite (bigD1 Vi) ?joing_subl // inE Si inE. +have [cSR | not_cSR] := boolP (R \subset 'C(S | 'JG)). + have{cSR} sRnSV: R \subset \bigcap_(Vi in S) 'N(Vi). + apply/bigcapsP=> Vi Si. + by rewrite -astab1JG (subset_trans cSR) ?astabS ?sub1set. + have sPRnSV: P <*> R \subset 'N(\bigcap_(Vi in S) 'N(Vi)). + apply/subsetP=> x PRx; rewrite inE; apply/bigcapsP=> Vi Si. + by rewrite sub_conjg -normJ bigcap_inf ?(acts_act actsPR) ?groupV. + have [V1 S1] := imsetP transPR. + have: P <*> R \subset 'N(V1). + rewrite join_subG (subset_trans sRnSV) /= ?bigcap_inf // andbT -defP. + apply: (subset_trans (commgS P sRnSV)). + have:= subset_trans (joing_subl P R) sPRnSV; rewrite -commg_subr /=. + move/subset_trans; apply; exact: bigcap_inf. + rewrite -afixJG; move/orbit1P => -> allV1. + have defV1: V1 = V by apply: group_inj; rewrite /= -defV allV1 big_set1. + case/idPn: oKq2; rewrite -(Lagrange (subsetIl K 'C(V1))). + rewrite (p_maximal_index qK (maxSK V1 S1)) defV1 /= tiKcV cards1 mul1n. + by rewrite (ltn_exp2l 2 1). +have actsR: [acts R, on S | 'JG] := subset_trans (joing_subr P R) actsPR. +have ntSRcR Vi: + Vi \in S -> ~~ (R \subset 'N(Vi)) -> + #|Vi| = p /\ 'C_V(R) \subset <<class_support Vi R>>. +- move=> Si not_nViR; have [sVi nV] := (subsetP (sSV Vi Si), subsetP nVR). + pose f v := fmval (\sum_(x in R) fmod cVV v ^@ x). + have fM: {in Vi &, {morph f: u v / u * v}}. + move=> u v /sVi Vu /sVi Vv; rewrite -fmvalA -big_split. + by congr (fmval _); apply: eq_bigr => x Rx; rewrite /= -actAr fmodM. + have injf: 'injm (Morphism fM). + apply/subsetP=> v /morphpreP[Vi_v]; have Vv := sVi v Vi_v. + rewrite (bigD1 Vi) //= in defV; have [[_ W _ dW] _ _ _] := dprodP defV. + have [u [w [_ _ uw Uuw]]] := mem_dprod defV (group1 V). + case: (Uuw 1 1) => // [||u1 w1]; rewrite ?dW ?mulg1 // !inE eq_sym /f /=. + move/eqP; rewrite (big_setD1 1) // actr1 ?fmodK // fmvalA //= fmval_sum. + do [case/Uuw; rewrite ?dW ?fmodK -?u1 ?group_prod //] => [x R'x | ->] //. + rewrite (nt_gen_prime _ R'x) ?cycle_subG ?oR // inE in not_nViR nVR actsR. + rewrite fmvalJ ?fmodK // -(bigdprodWY dW) ?mem_gen //; apply/bigcupP. + exists (Vi :^ x)%G; rewrite ?memJ_conjg // (astabs_act _ actsR) Si. + by apply: contraNneq not_nViR => /congr_group->. + have im_f: Morphism fM @* Vi \subset 'C_V(R). + apply/subsetP=> _ /morphimP[v _ Vi_v ->]; rewrite inE fmodP. + apply/centP=> x Rx; red; rewrite conjgC -fmvalJ ?nV //; congr (x * fmval _). + rewrite {2}(reindex_acts 'R _ Rx) ?astabsR //= actr_sum. + by apply: eq_bigr => y Ry; rewrite actrM ?nV. + have defCVR: Morphism fM @* Vi = 'C_V(R). + apply/eqP; rewrite eqEcard im_f (prime_nt_dvdP _ _ (cardSg im_f)) ?oCVR //=. + by rewrite -trivg_card1 morphim_injm_eq1 ?ntSV. + rewrite -oCVR -defCVR; split; first by rewrite card_injm. + apply/subsetP=> _ /morphimP[v _ Vi_v ->] /=; rewrite /f fmval_sum. + have Vv := sVi v Vi_v; apply: group_prod => x Rx. + by rewrite fmvalJ ?fmodK ?nV // mem_gen // mem_imset2. +have{not_cSR} [V1 S1 not_nV1R]: exists2 V1, V1 \in S & ~~ (R \subset 'N(V1)). + by move: not_cSR; rewrite astabC; case/subsetPn=> v; rewrite afixJG; exists v. +set D := orbit 'JG%act R V1. +have oD: #|D| = r by rewrite card_orbit astab1JG prime_TIg ?indexg1 ?oR. +have oSV Vi: Vi \in S -> #|Vi| = p. + move=> Si; have [z _ ->]:= atransP2 transPR S1 Si. + by rewrite cardJg; case/ntSRcR: not_nV1R. +have cSnS' Vi: Vi \in S -> 'N(Vi)^`(1) \subset 'C(Vi). + move=> Si; rewrite der1_min ?cent_norm //= -ker_conj_aut. + rewrite (isog_abelian (first_isog _)) (abelianS (Aut_conj_aut _ _)) //. + by rewrite Aut_cyclic_abelian // prime_cyclic // oSV. +have nVjR Vj: Vj \in S :\: D -> 'C_K(Vj) = [~: K, R]. + case/setDP=> Sj notDj; set Kj := 'C_K(Vj). + have [nVjR|] := boolP (R \subset 'N(Vj)). + have{nVjR} sKRVj: [~: K, R] \subset Kj. + rewrite subsetI {1}commGC commg_subr nKR. + by rewrite (subset_trans _ (cSnS' Vj Sj)) // commgSS ?nSK. + have iKj: #|K : Kj| = q by rewrite (p_maximal_index qK (maxSK Vj Sj)). + have dxKR: [~: K, R] \x 'C_K(R) = K by rewrite coprime_abelian_cent_dprod. + have{dxKR} [_ defKR _ tiKRcR] := dprodP dxKR. + have Z_CK: Zgroup 'C_K(R) by apply: ZgroupS ZgrCHR; exact: setSI. + have abelCKR: q.-abelem 'C_K(R) := abelemS (subsetIl _ _) abelK. + have [qCKR _] := andP abelCKR. + apply/eqP; rewrite eq_sym eqEcard sKRVj -(leq_pmul2r (ltnW q_gt1)). + rewrite -{1}iKj Lagrange ?subsetIl // -{1}defKR (TI_cardMg tiKRcR). + rewrite leq_pmul2l ?cardG_gt0 //= (card_pgroup qCKR). + rewrite (leq_exp2l _ 1) // -abelem_cyclic // (forall_inP Z_CK) //. + by rewrite (@p_Sylow _ q) // /pHall subxx indexgg qCKR. + case/ntSRcR=> // _ sCVj; case/ntSRcR: not_nV1R => // _ sCV1. + suffices trCVR: 'C_V(R) = 1 by rewrite -oCVR trCVR cards1 in p_pr. + apply/trivgP; rewrite (big_setID D) in defV. + have{defV} [[W U /= defW defU] _ _ <-] := dprodP defV. + rewrite defW defU subsetI (subset_trans sCV1) /=; last first. + rewrite class_supportEr -(bigdprodWY defW) genS //. + apply/bigcupsP=> x Rx; rewrite (bigcup_max (V1 :^ x)%G) // inE. + by rewrite (actsP actsR) //= S1 mem_imset. + rewrite (subset_trans sCVj) // class_supportEr -(bigdprodWY defU) genS //. + apply/bigcupsP=> x Rx; rewrite (bigcup_max (Vj :^ x)%G) // inE. + by rewrite (actsP actsR) // Sj andbT (orbit_transr _ (mem_orbit 'JG Vj Rx)). +have sDS: D \subset S. + by rewrite acts_sub_orbit //; apply: subset_trans actsPR; exact: joing_subr. +have [eqDS | ltDS] := eqVproper sDS. + have [fix0 | [Vj cVjP]] := set_0Vmem 'Fix_(S | 'JG)(P). + case/negP: p'r; rewrite eq_sym -dvdn_prime2 // -oD eqDS /dvdn. + rewrite (pgroup_fix_mod pP (subset_trans (joing_subl P R) actsPR)). + by rewrite fix0 cards0 mod0n. + have{cVjP} [Sj nVjP] := setIP cVjP; rewrite afixJG in nVjP. + case/negP: (ntSV Vj Sj); rewrite -subG1 -tiVcK subsetI sSV // centsC -defKP. + by rewrite (subset_trans _ (cSnS' Vj Sj)) // commgSS ?nSK. +have [_ [Vj Sj notDj]] := properP ltDS. +have defS: S = Vj |: D. + apply/eqP; rewrite eqEsubset andbC subUset sub1set Sj sDS. + apply/subsetP=> Vi Si; rewrite !inE orbC /= -val_eqE /= -(defSV Vi Si). + have [//|notDi] := boolP (Vi \in _); rewrite -(defSV Vj Sj) /=. + by rewrite !nVjR // inE ?notDi ?notDj. +suffices: odd #|S| by rewrite defS cardsU1 (negPf notDj) /= oD -oR (oddSg sRG). +rewrite (dvdn_odd (atrans_dvd transPR)) // (oddSg _ oddG) //. +by rewrite join_subG (subset_trans sPH). +Qed. + +End Theorem_3_6. + +(* This is B & G, Theorem 3.7. *) +Theorem prime_Frobenius_sol_kernel_nil gT (G K R : {group gT}) : + K ><| R = G -> solvable G -> prime #|R| -> 'C_K(R) = 1 -> nilpotent K. +Proof. +move=> defG solG R_pr regR. +elim: {K}_.+1 {-2}K (ltnSn #|K|) => // m IHm K leKm in G defG solG regR *. +have [nsKG sRG defKR nKR tiKR] := sdprod_context defG. +have [sKG nKG] := andP nsKG. +wlog ntK: / K :!=: 1 by case: eqP => [-> _ | _ ->] //; exact: nilpotent1. +have [L maxL _]: {L : {group gT} | maxnormal L K G & [1] \subset L}. + by apply: maxgroup_exists; rewrite proper1G ntK norms1. +have [ltLK nLG]:= andP (maxgroupp maxL); have [sLK not_sKL]:= andP ltLK. +have{m leKm IHm}nilL: nilpotent L. + pose G1 := L <*> R; have nLR := subset_trans sRG nLG. + have sG1G: G1 \subset G by rewrite join_subG (subset_trans sLK). + have defG1: L ><| R = G1. + by rewrite sdprodEY //; apply/eqP; rewrite -subG1 -tiKR setSI. + apply: (IHm _ _ _ defG1); rewrite ?(solvableS sG1G) ?(oddSg sG1G) //. + exact: leq_trans (proper_card ltLK) _. + by apply/eqP; rewrite -subG1 -regR setSI. +have sLG := subset_trans sLK sKG; have nsLG: L <| G by apply/andP. +have sLF: L \subset 'F(G) by apply: Fitting_max. +have frobG: [Frobenius G = K ><| R] by apply/prime_FrobeniusP. +have solK := solvableS sKG solG. +have frobGq := Frobenius_quotient frobG solK nsLG not_sKL. +suffices sKF: K \subset 'F(K) by apply: nilpotentS sKF (Fitting_nil K). +apply: subset_trans (chief_stab_sub_Fitting solG nsKG). +rewrite subsetI subxx; apply/bigcapsP=> [[X Y] /= /andP[chiefXY sXF]]. +set V := X / Y; have [maxY nsXG] := andP chiefXY. +have [ltYX nYG] := andP (maxgroupp maxY); have [sYX _]:= andP ltYX. +have [sXG nXG] := andP nsXG; have sXK := subset_trans sXF (Fitting_sub K). +have minV := chief_factor_minnormal chiefXY. +have cVL: L \subset 'C(V | 'Q). + apply: subset_trans (subset_trans sLF (Fitting_stab_chief solG _)) _ => //. + exact: (bigcap_inf (X, Y)). +have nVG: {acts G, on group V | 'Q}. + by split; rewrite ?quotientS ?subsetT // actsQ // normal_norm. +pose V1 := sdpair1 <[nVG]> @* V. +have [p p_pr abelV]: exists2 p, prime p & p.-abelem V. + apply/is_abelemP; apply: charsimple_solvable (quotient_sol _ _). + exact: minnormal_charsimple minV. + exact: nilpotent_sol (nilpotentS sXF (Fitting_nil _)). +have abelV1: p.-abelem V1 by rewrite morphim_abelem. +have injV1 := injm_sdpair1 <[nVG]>. +have ntV1: V1 :!=: 1. + by rewrite -cardG_gt1 card_injm // cardG_gt1; case/andP: (mingroupp minV). +have nV1_G1 := im_sdpair_norm <[nVG]>. +pose rV := morphim_repr (abelem_repr abelV1 ntV1 nV1_G1) (subxx G). +have def_kerV: rker rV = 'C_G(V | 'Q). + rewrite rker_morphim rker_abelem morphpreIdom morphpreIim -astabEsd //. + by rewrite astab_actby setIid. +have kerL: L \subset rker rV by rewrite def_kerV subsetI sLG. +pose rVq := quo_repr kerL nLG. +suffices: K / L \subset rker rVq. + rewrite rker_quo def_kerV quotientSGK //= 1?subsetI 1?(subset_trans sKG) //. + by rewrite sLG. +have irrVq: mx_irreducible rVq. + apply/quo_mx_irr; apply/morphim_mx_irr; apply/abelem_mx_irrP. + apply/mingroupP; rewrite ntV1; split=> // U1; case/andP=> ntU1 nU1G sU1V. + rewrite -(morphpreK sU1V); congr (_ @* _). + case/mingroupP: minV => _; apply; last by rewrite sub_morphpre_injm. + rewrite -subG1 sub_morphpre_injm ?sub1G // morphim1 subG1 ntU1 /=. + set U := _ @*^-1 U1; rewrite -(cosetpreK U) quotient_norms //. + have: [acts G, on U | <[nVG]>] by rewrite actsEsd ?subsetIl // morphpreK. + rewrite astabs_actby subsetI subxx (setIidPr _) ?subsetIl //=. + by rewrite -{1}(cosetpreK U) astabsQ ?normal_cosetpre //= -/U subsetI nYG. +have [q q_pr abelKq]: exists2 q, prime q & q.-abelem (K / L). + apply/is_abelemP; apply: charsimple_solvable (quotient_sol _ solK). + exact: maxnormal_charsimple maxL. +case (eqVneq q p) => [def_q | neq_qp]. + have sKGq: K / L \subset G / L by apply: quotientS. + rewrite rfix_mx_rstabC //; have [_ _]:= irrVq; apply; rewrite ?submx1 //. + by rewrite normal_rfix_mx_module ?quotient_normal. + rewrite -(rfix_subg _ sKGq) (@rfix_pgroup_char _ p) ?char_Fp -?def_q //. + exact: (abelem_pgroup abelKq). +suffices: rfix_mx rVq (R / L) == 0. + apply: contraLR; apply: (Frobenius_rfix_compl frobGq). + apply: pi_pnat (abelem_pgroup abelKq) _. + by rewrite inE /= (charf_eq (char_Fp p_pr)). +rewrite -mxrank_eq0 (rfix_quo _ _ sRG) (rfix_morphim _ _ sRG). +rewrite (rfix_abelem _ _ _ (morphimS _ sRG)) mxrank_eq0 rowg_mx_eq0 -subG1. +rewrite (sub_abelem_rV_im _ _ _ (subsetIl _ _)) -(morphpreSK _ (subsetIl _ _)). +rewrite morphpreIim -gacentEsd gacent_actby gacentQ (setIidPr sRG) /=. +rewrite -coprime_quotient_cent ?(solvableS sXG) ?(subset_trans sRG) //. + by rewrite {1}['C_X(R)](trivgP _) ?quotient1 ?sub1G // -regR setSI. +by apply: coprimeSg sXK _; apply: Frobenius_coprime frobG. +Qed. + +Corollary Frobenius_sol_kernel_nil gT (G K H : {group gT}) : + [Frobenius G = K ><| H] -> solvable G -> nilpotent K. +Proof. +move=> frobG solG; have [defG ntK ntH _ _] := Frobenius_context frobG. +have{defG} /sdprodP[_ defG nKH tiKH] := defG. +have[H1 | [p p_pr]] := trivgVpdiv H; first by case/eqP: ntH. +case/Cauchy=> // x Hx ox; rewrite -ox in p_pr. +have nKx: <[x]> \subset 'N(K) by rewrite cycle_subG (subsetP nKH). +have tiKx: K :&: <[x]> = 1 by apply/trivgP; rewrite -tiKH setIS ?cycle_subG. +apply: (prime_Frobenius_sol_kernel_nil (sdprodEY nKx tiKx)) => //. + by rewrite (solvableS _ solG) // join_subG -mulG_subG -defG mulgS ?cycle_subG. +by rewrite cent_cycle (Frobenius_reg_ker frobG) // !inE -order_gt1 prime_gt1. +Qed. + +(* This is B & G, Theorem 3.8. *) +Theorem odd_sdprod_primact_commg_sub_Fitting gT (G K R : {group gT}) : + K ><| R = G -> odd #|G| -> solvable G -> + (*1*) coprime #|K| #|R| -> + (*2*) semiprime K R -> + (*3*) 'C_('F(K))(R) = 1 -> + [~: K, R] \subset 'F(K). +Proof. +elim: {G}_.+1 {-2}G (ltnSn #|G|) K R => // n IHn G. +rewrite ltnS => leGn K R defG oddG solG coKR primR regR_F. +have [nsKG sRG defKR nKR tiKR] := sdprod_context defG. +have [sKG nKG] := andP nsKG. +have chF: 'F(K) \char K := Fitting_char K; have nFR := char_norm_trans chF nKR. +have nsFK := char_normal chF; have [sFK nFK] := andP nsFK. +pose KqF := K / 'F(K); have solK := solvableS sKG solG. +without loss [p p_pr pKqF]: / exists2 p, prime p & p.-group KqF. + move=> IHp; apply: wlog_neg => IH_KR; rewrite -quotient_cents2 //= -/KqF. + set Rq := R / 'F(K); have nKRq: Rq \subset 'N(KqF) by exact: quotient_norms. + rewrite centsC. + apply: subset_trans (coprime_cent_Fitting nKRq _ _); last first. + - exact: quotient_sol. + - exact: coprime_morph. + rewrite subsetI subxx centsC -['F(KqF)]Sylow_gen gen_subG. + apply/bigcupsP=> Pq /SylowP[p p_pr /= sylPq]; rewrite -/KqF in sylPq. + have chPq: Pq \char KqF. + apply: char_trans (Fitting_char _); rewrite /= -/KqF. + by rewrite (nilpotent_Hall_pcore (Fitting_nil _) sylPq) ?pcore_char. + have [P defPq sFP sPK] := inv_quotientS nsFK (char_sub chPq). + have nsFP: 'F(K) <| P by rewrite /normal sFP (subset_trans sPK). + have{chPq} chP: P \char K. + by apply: char_from_quotient nsFP (Fitting_char _) _; rewrite -defPq. + have defFP: 'F(P) = 'F(K). + apply/eqP; rewrite eqEsubset !Fitting_max ?Fitting_nil //. + by rewrite char_normal ?(char_trans (Fitting_char _)). + have coPR := coprimeSg sPK coKR. + have nPR: R \subset 'N(P) := char_norm_trans chP nKR. + pose G1 := P <*> R. + have sG1G: G1 \subset G by rewrite /G1 -defKR norm_joinEr ?mulSg. + have defG1: P ><| R = G1 by rewrite sdprodEY ?coprime_TIg. + rewrite defPq quotient_cents2r //= -defFP. + have:= sPK; rewrite subEproper; case/predU1P=> [defP | ltPK]. + rewrite IHp // in IH_KR; exists p => //. + by rewrite /KqF -{2}defP -defPq (pHall_pgroup sylPq). + move/IHn: defG1 => ->; rewrite ?(oddSg sG1G) ?(solvableS sG1G) ?defFP //. + apply: leq_trans leGn; rewrite /= norm_joinEr //. + by rewrite -defKR !coprime_cardMg // ltn_pmul2r ?proper_card. + by move=> x Rx; rewrite -(setIidPl sPK) -!setIA primR. +without loss r_pr: / prime #|R|; last set r := #|R| in r_pr. + have [-> _ | [r r_pr]] := trivgVpdiv R; first by rewrite commG1 sub1G. + case/Cauchy=> // x; rewrite -cycle_subG subEproper orderE; set X := <[x]>. + case/predU1P=> [-> -> -> // | ltXR rX _]; have sXR := proper_sub ltXR. + have defCX: 'C_K(X) = 'C_K(R). + rewrite cent_cycle primR // !inE -order_gt1 orderE rX prime_gt1 //=. + by rewrite -cycle_subG. + have primX: semiprime K X. + by move=> y; case/setD1P=> nty Xy; rewrite primR // !inE nty (subsetP sXR). + have nKX := subset_trans sXR nKR; have coKX := coprimegS sXR coKR. + pose H := K <*> X; have defH: K ><| X = H by rewrite sdprodEY ?coprime_TIg. + have sHG: H \subset G by rewrite /H -defKR norm_joinEr ?mulgSS. + have ltHn: #|H| < n. + rewrite (leq_trans _ leGn) /H ?norm_joinEr // -defKR !coprime_cardMg //. + by rewrite ltn_pmul2l ?proper_card. + have oddH := oddSg sHG oddG; have solH := solvableS sHG solG. + have regX_F: 'C_('F(K))(X) = 1. + by rewrite -regR_F -(setIidPl sFK) -!setIA defCX. + have:= IHn _ ltHn _ _ defH oddH solH coKX primX regX_F. + rewrite -!quotient_cents2 ?(subset_trans sXR) //; move/setIidPl <-. + rewrite -coprime_quotient_cent ?(subset_trans sXR) // defCX. + by rewrite coprime_quotient_cent ?subsetIr. +apply: subset_trans (chief_stab_sub_Fitting solG nsKG) => //. +rewrite subsetI commg_subl nKR; apply/bigcapsP => [[U V]] /=. +case/andP=> chiefUV sUF; set W := U / V. +have minW := chief_factor_minnormal chiefUV. +have [ntW nWG] := andP (mingroupp minW). +have /andP[/maxgroupp/andP[/andP[sVU _] nVG] nsUG] := chiefUV. +have sUK := subset_trans sUF sFK; have sVK := subset_trans sVU sUK. +have nVK := subset_trans sKG nVG; have nVR := subset_trans sRG nVG. +have [q q_pr abelW]: exists2 q, prime q & q.-abelem W. + apply/is_abelemP; apply: charsimple_solvable (minnormal_charsimple minW) _. + by rewrite quotient_sol // (solvableS sUK). +have regR_W: 'C_(W)(R / V) = 1. + rewrite -coprime_quotient_cent ?(coprimeSg sUK) ?(solvableS sUK) //. + by rewrite -(setIidPl sUF) -setIA regR_F (setIidPr _) ?quotient1 ?sub1G. +rewrite sub_astabQ comm_subG ?quotientR //=. +have defGv: (K / V) * (R / V) = G / V by rewrite -defKR quotientMl. +have oRv: #|R / V| = r. + rewrite card_quotient // -indexgI -(setIidPr sVK) setICA setIA tiKR. + by rewrite (setIidPl (sub1G _)) indexg1. +have defCW: 'C_(G / V)(W) = 'C_(K / V)(W). + apply/eqP; rewrite eqEsubset andbC setSI ?quotientS //=. + rewrite subsetI subsetIr /= andbT. + rewrite -(coprime_mulG_setI_norm defGv) ?coprime_morph ?norms_cent //=. + suffices ->: 'C_(R / V)(W) = 1 by rewrite mulg1 subsetIl. + apply/trivgP; apply/subsetP=> x; case/setIP=> Rx cWx. + apply: contraR ntW => ntx; rewrite -subG1 -regR_W subsetI subxx centsC /= -/W. + by apply: contraR ntx; move/prime_TIg <-; rewrite ?oRv // inE Rx. +have [P sylP nPR] := coprime_Hall_exists p nKR coKR solK. +have [sPK pP _] := and3P sylP. +have nVP := subset_trans sPK nVK; have nFP := subset_trans sPK nFK. +have sylPv: p.-Sylow(K / V) (P / V) by rewrite quotient_pHall. +have defKv: (P / V) * 'C_(G / V)(W) = (K / V). + rewrite defCW; apply/eqP; rewrite eqEsubset mulG_subG subsetIl quotientS //=. + have sK_PF: K \subset P * 'F(K). + rewrite (normC nFP) -quotientSK // subEproper eq_sym eqEcard quotientS //=. + by rewrite (card_Hall (quotient_pHall nFP sylP)) part_pnat_id ?leqnn. + rewrite (subset_trans (quotientS _ sK_PF)) // quotientMl // mulgS //. + rewrite subsetI -quotient_astabQ !quotientS //. + by rewrite (subset_trans (Fitting_stab_chief solG nsKG)) ?(bigcap_inf (U, V)). +have nW_ := subset_trans (quotientS _ _) nWG; have nWK := nW_ _ sKG. +rewrite -quotient_cents2 ?norms_cent ?(nW_ _ sRG) //. +have [eq_qp | p'q] := eqVneq q p. + apply: subset_trans (sub1G _); rewrite -trivg_quotient quotientS // centsC. + apply/setIidPl; case/mingroupP: minW => _; apply; last exact: subsetIl. + rewrite andbC normsI ?norms_cent // ?quotient_norms //=. + have nsWK: W <| K / V by rewrite /normal quotientS. + have sWP: W \subset P / V. + by rewrite (normal_sub_max_pgroup (Hall_max sylPv)) -?eq_qp ?abelem_pgroup. + rewrite -defKv centM setIA setIAC /=. + rewrite ['C_W(_)](setIidPl _); last by rewrite centsC subsetIr. + have nilPv: nilpotent (P / V) := pgroup_nil (pHall_pgroup sylPv). + rewrite -/W -(setIidPl sWP) -setIA meet_center_nil //. + exact: normalS (quotientS V sPK) nsWK. +rewrite -defKv -quotientMidr -mulgA mulSGid ?subsetIr // quotientMidr. +have sPG := subset_trans sPK sKG. +rewrite quotient_cents2 ?norms_cent ?nW_ //= commGC. +pose Hv := (P / V) <*> (R / V). +have sHGv: Hv \subset G / V by rewrite join_subG !quotientS. +have solHv: solvable Hv := solvableS sHGv (quotient_sol V solG). +have sPHv: P / V \subset Hv by exact: joing_subl. +have nPRv: R / V \subset 'N(P / V) := quotient_norms _ nPR. +have coPRv: coprime #|P / V| #|R / V| := coprime_morph _ (coprimeSg sPK coKR). +apply: subset_trans (subsetIr (P / V) _). +have oHv: #|Hv| = (#|P / V| * #|R / V|)%N. + by rewrite /Hv norm_joinEr // coprime_cardMg // oRv. +move/(odd_prime_sdprod_abelem_cent1 solHv): (abelW); apply=> //. +- exact: oddSg sHGv (quotient_odd _ _). +- by rewrite sdprodEY ?quotient_norms // coprime_TIg. +- by rewrite oRv. +- by rewrite (subset_trans _ nWG) ?join_subG ?quotientS. +rewrite /= norm_joinEr // pgroupM /pgroup. +rewrite (pi_pnat (quotient_pgroup _ pP)) ?inE 1?eq_sym //=. +apply: coprime_p'group (abelem_pgroup abelW) ntW. +by rewrite coprime_sym coprime_morph // (coprimeSg sUK). +Qed. + +(* This is B & G, Proposition 3.9 (for external action), with the incorrectly *) +(* omitted nontriviality assumption reinstated. *) +Proposition ext_odd_regular_pgroup_cyclic (aT rT : finGroupType) p + (D R : {group aT}) (K H : {group rT}) (to : groupAction D K) : + p.-group R -> odd #|R| -> H :!=: 1 -> + {acts R, on group H | to} -> {in R^#, forall x, 'C_(H | to)[x] = 1} -> + cyclic R. +Proof. +move: R H => R0 H0 pR0 oddR0 ntH0 actsR0 regR0. +pose gT := sdprod_groupType <[actsR0]>. +pose H : {group gT} := (sdpair1 <[actsR0]> @* H0)%G. +pose R : {group gT} := (sdpair2 <[actsR0]> @* R0)%G. +pose G : {group gT} := [set: gT]%G. +have{pR0} pR: p.-group R by rewrite morphim_pgroup. +have{oddR0} oddR: odd #|R| by rewrite morphim_odd. +have [R1 | ntR] := eqsVneq R 1. + by rewrite -(im_invm (injm_sdpair2 <[actsR0]>)) {2}R1 morphim1 cyclic1. +have{ntH0} ntH: H :!=: 1. + apply: contraNneq ntH0 => H1. + by rewrite -(im_invm (injm_sdpair1 <[actsR0]>)) {2}H1 morphim1. +have{regR0 ntR} frobG: [Frobenius G = H ><| R]. + apply/Frobenius_semiregularP => // [|x]; first exact: sdprod_sdpair. + case/setD1P=> nt_x; case/morphimP=> x2 _ Rx2 def_x. + apply/trivgP; rewrite -(morphpreSK _ (subsetIl _ _)) morphpreI. + rewrite /= -cent_cycle def_x -morphim_cycle // -gacentEsd. + rewrite injmK ?injm_sdpair1 // (trivgP (injm_sdpair1 _)). + rewrite -(regR0 x2) ?inE ?Rx2 ?andbT; last first. + by apply: contra nt_x; rewrite def_x; move/eqP->; rewrite morph1. + have [sRD sHK]: R0 \subset D /\ H0 \subset K by case actsR0; move/acts_dom. + have sx2R: <[x2]> \subset R0 by rewrite cycle_subG. + rewrite gacent_actby setIA setIid (setIidPr sx2R). + rewrite !gacentE ?cycle_subG ?sub1set ?(subsetP sRD) //. + by rewrite !setIS ?afixS ?sub_gen. +suffices: cyclic R by rewrite (injm_cyclic (injm_sdpair2 _)). +move: gT H R G => {aT rT to D K H0 R0 actsR0} gT H R G in ntH pR oddR frobG *. +have [defG _ _ _ _] := Frobenius_context frobG; case/sdprodP: defG => _ _ nHR _. +have coHR := Frobenius_coprime frobG. +rewrite (odd_pgroup_rank1_cyclic pR oddR) leqNgt. +apply: contra ntH => /p_rank_geP[E /pnElemP[sER abelE dimE2]]. +have ncycE: ~~ cyclic E by rewrite (abelem_cyclic abelE) dimE2. +have nHE := subset_trans sER nHR; have coHE := coprimegS sER coHR. +rewrite -subG1 -(coprime_abelian_gen_cent1 _ _ nHE) ?(abelem_abelian abelE) //. +rewrite -bigprodGE big1 // => x /setD1P[nt_x Ex]; apply: val_inj => /=. +by apply: (Frobenius_reg_ker frobG); rewrite !inE nt_x (subsetP sER). +Qed. + +(* Internal action version of B & G, Proposition 3.9 (possibly, the only one *) +(* we should keep). *) +Proposition odd_regular_pgroup_cyclic gT p (H R : {group gT}) : + p.-group R -> odd #|R| -> H :!=: 1 -> R \subset 'N(H) -> semiregular H R -> + cyclic R. +Proof. +move=> pR oddR ntH nHR regR. +have actsR: {acts R, on group H | 'J} by split; rewrite ?subsetT ?astabsJ. +apply: ext_odd_regular_pgroup_cyclic pR oddR ntH actsR _ => // x Rx. +by rewrite gacentJ cent_set1 regR. +Qed. + +(* Another proof of Proposition 3.9, which avoids Frobenius groups entirely. *) +Proposition simple_odd_regular_pgroup_cyclic gT p (H R : {group gT}) : + p.-group R -> odd #|R| -> H :!=: 1 -> R \subset 'N(H) -> semiregular H R -> + cyclic R. +Proof. +move=> pR oddR ntH nHR regR; rewrite (odd_pgroup_rank1_cyclic pR oddR) leqNgt. +apply: contra ntH => /p_rank_geP[E /pnElemP[sER abelE dimE2]]. +have ncycE: ~~ cyclic E by rewrite (abelem_cyclic abelE) dimE2. +have nHE := subset_trans sER nHR. +have coHE := coprimegS sER (regular_norm_coprime nHR regR). +rewrite -subG1 -(coprime_abelian_gen_cent1 _ _ nHE) ?(abelem_abelian abelE) //. +rewrite -bigprodGE big1 // => x; case/setD1P=> nt_x Ex; apply: val_inj => /=. +by rewrite regR // !inE nt_x (subsetP sER). +Qed. + +(* This is Aschbacher (40.6)(4). *) +Lemma odd_regular_metacyclic gT (H R : {group gT}) : + odd #|R| -> H :!=: 1 -> R \subset 'N(H) -> semiregular H R -> + metacyclic R. +Proof. +move=> oddR ntH nHR regHR. +apply/Zgroup_metacyclic/forall_inP=> P /SylowP[p pr_p /and3P[sPR pP _]]. +have [oddP nHP] := (oddSg sPR oddR, subset_trans sPR nHR). +exact: odd_regular_pgroup_cyclic pP oddP ntH nHP (semiregularS _ sPR regHR). +Qed. + +(* This is Huppert, Kapitel V, Satz 18.8 b (used in Peterfalvi, Section 13). *) +Lemma prime_odd_regular_normal gT (H R P : {group gT}) : + prime #|P| -> odd #|R| -> P \subset R -> + H :!=: 1 -> R \subset 'N(H) -> semiregular H R -> + P <| R. +Proof. +set p := #|P| => pr_p oddR sPR ntH nHR regHR. +have pP: p.-group P := pnat_id pr_p. +have cycQ (q : nat) (Q : {group gT}) : q.-group Q -> Q \subset R -> cyclic Q. + move=> qQ sQR; have [oddQ nHQ] := (oddSg sQR oddR, subset_trans sQR nHR). + exact: odd_regular_pgroup_cyclic qQ oddQ ntH nHQ (semiregularS _ sQR regHR). +have cycRq (q : nat): cyclic 'O_q(R) by rewrite (cycQ q) ?pcore_pgroup ?gFsub. +suffices cFP: P \subset 'C('F(R)). + have nilF: nilpotent 'F(R) := Fitting_nil R. + have hallRp: p.-Hall('F(R)) 'O_p('F(R)) := nilpotent_pcore_Hall p nilF. + apply: char_normal_trans (pcore_normal p R); rewrite sub_cyclic_char //=. + rewrite -p_core_Fitting (sub_normal_Hall hallRp) ?gFnormal //. + have solR: solvable R. + by apply: metacyclic_sol; apply: odd_regular_metacyclic regHR. + by apply: subset_trans (cent_sub_Fitting solR); rewrite subsetI sPR. +rewrite centsC -(bigdprodWY (erefl 'F(R))) gen_subG big_tnth. +apply/bigcupsP=> i _; move: {i}(tuple.tnth _ i) => q. +have [<- | q'p] := eqVneq p q. + have [Q sylQ sPQ] := Sylow_superset sPR pP; have [sQR pQ _] := and3P sylQ. + rewrite (sub_abelian_cent2 (cyclic_abelian (cycQ _ _ pQ sQR))) //. + by rewrite pcore_sub_Hall. +have [-> | ntRq] := eqVneq 'O_q(R) 1%g; first exact: sub1G. +have /andP[sRqR qRq]: q.-subgroup(R) 'O_q(R) by apply: pcore_psubgroup. +have [pr_q _ _] := pgroup_pdiv qRq ntRq. +have coRqP: coprime #|'O_q(R)| p by rewrite (pnat_coprime qRq) // pnatE. +have nRqP: P \subset 'N('O_q(R)) by rewrite (subset_trans sPR) ?gFnorm. +rewrite centsC (coprime_odd_faithful_Ohm1 qRq) ?(oddSg sRqR) //. +apply: sub_abelian_cent2 (joing_subl _ _) (joing_subr _ _) => /=. +set PQ := P <*> _; have oPQ: #|PQ| = (p * q)%N. + rewrite /PQ norm_joinEl ?(char_norm_trans (Ohm_char 1 _)) //. + rewrite coprime_cardMg 1?coprime_sym ?(coprimeSg (Ohm_sub 1 _)) // -/p. + by congr (p * _)%N; apply: Ohm1_cyclic_pgroup_prime => /=. +have sPQ_R: PQ \subset R by rewrite join_subG sPR (subset_trans (Ohm_sub 1 _)). +have nH_PQ := subset_trans sPQ_R nHR. +apply: cyclic_abelian; apply: regular_pq_group_cyclic oPQ ntH nH_PQ _ => //. +exact: semiregularS regHR. +Qed. + +Section Wielandt_Frobenius. + +Variables (gT : finGroupType) (G K R : {group gT}). +Implicit Type A : {group gT}. + +(* This is Peterfalvi (9.1). *) +Lemma Frobenius_Wielandt_fixpoint (M : {group gT}) : + [Frobenius G = K ><| R] -> + G \subset 'N(M) -> coprime #|M| #|G| -> solvable M -> + [/\ (#|'C_M(G)| ^ #|R| * #|M| = #|'C_M(R)| ^ #|R| * #|'C_M(K)|)%N, + 'C_M(R) = 1 -> K \subset 'C(M) + & 'C_M(K) = 1 -> (#|M| = #|'C_M(R)| ^ #|R|)%N]. +Proof. +move=> frobG nMG coMG solM; have [defG _ ntR _ _] := Frobenius_context frobG. +have [_ _ _ _ /eqP snRG] := and5P frobG. +have [nsKG sRG _ _ tiKR] := sdprod_context defG; have [sKG _] := andP nsKG. +pose m0 (_ : {group gT}) := 0%N. +pose Dm := [set 1%G; G]; pose Dn := K |: orbit 'JG K R. +pose m := [fun A => 0%N with 1%G |-> #|K|, G |-> 1%N]. +pose n A : nat := A \in Dn. +have out_m: {in [predC Dm], m =1 m0}. + by move=> A; rewrite !inE /=; case/norP; do 2!move/negbTE->. +have out_n: {in [predC Dn], n =1 m0}. + by rewrite /n => A /=; move/negbTE=> /= ->. +have ntG: G != 1%G by case: eqP sRG => // -> <-; rewrite subG1. +have neqKR: K \notin orbit 'JG K R. + apply/imsetP=> [[x _ defK]]; have:= Frobenius_dvd_ker1 frobG. + by rewrite defK cardJg gtnNdvd // ?prednK // -subn1 subn_gt0 cardG_gt1. +have Gmn A: m A + n A > 0 -> A \subset G. + rewrite /=; case: eqP => [-> | ] _; first by rewrite sub1G. + rewrite /n 2!inE; do 2!case: eqP => [-> // | ] _. + case R_A: (A \in _) => // _; case/imsetP: R_A => x Kx ->{A}. + by rewrite conj_subG ?(subsetP sKG). +have partG: {in G, forall a, + \sum_(A | a \in gval A) m A = \sum_(A | a \in gval A) n A}%N. +- move=> a Ga; have [-> | nt_a] := eqVneq a 1. + rewrite (bigD1 1%G) ?inE ?eqxx //= (bigD1 G) ?inE ?group1 //=. + rewrite (negbTE ntG) !eqxx big1 ?addn1 => [|A]; last first. + by rewrite group1 -negb_or -in_set2; apply: out_m. + rewrite (bigID (mem Dn)) /= addnC big1 => [|A]; last first. + by rewrite group1; apply: out_n. + transitivity #|Dn|. + rewrite cardsU1 neqKR card_orbit astab1JG. + by rewrite -{3}(setIidPl sKG) -setIA -normD1 snRG tiKR indexg1. + by rewrite -sum1_card /n; apply: eq_big => [A | A ->]; rewrite ?group1. + rewrite (bigD1 G) //= (negbTE ntG) eqxx big1 => [|A]; last first. + case/andP=> Aa neAG; apply: out_m; rewrite !inE; case: eqP => // A1. + by rewrite A1 inE (negbTE nt_a) in Aa. + have [partG tiG _] := and3P (Frobenius_partition frobG). + do [rewrite -(eqP partG); set pG := _ |: _] in Ga tiG. + rewrite (bigD1 <<pblock pG a>>%G) /=; last by rewrite mem_gen // mem_pblock. + rewrite big1 => [|B]; last first. + case/andP=> Ba neqBA; rewrite -/(false : nat); congr (nat_of_bool _). + apply: contraTF neqBA; rewrite negbK -val_eqE /=. + case/setU1P=> [BK | /imsetP[x Kx defB]]. + by rewrite (def_pblock tiG _ Ba) BK ?setU11 ?genGid. + have Rxa: a \in R^# :^ x by rewrite conjD1g !inE nt_a -(congr_group defB). + rewrite (def_pblock tiG _ Rxa) ?setU1r ?mem_imset // conjD1g. + by rewrite genD1 ?group1 // genGid defB. + rewrite /n !inE -val_eqE /= -/(true : nat); congr ((_ : bool) + _)%N. + case/setU1P: (pblock_mem Ga) => [-> |]; first by rewrite genGid eqxx. + case/imsetP=> x Kx ->; symmetry; apply/orP; right. + apply/imsetP; exists x => //. + by apply: val_inj; rewrite conjD1g /= genD1 ?group1 // genGid. +move/eqP: (solvable_Wielandt_fixpoint Gmn nMG coMG solM partG). +rewrite (bigD1 1%G) // (bigD1 G) //= eqxx (setIidPl (cents1 _)) cards1 muln1. +rewrite (negbTE ntG) eqxx mul1n -(sdprod_card defG) (mulnC #|K|) expnM. +rewrite mulnA -expnMn big1 ?muln1 => [|A]; last first. + by rewrite -negb_or -in_set2; move/out_m; rewrite /m => /= ->. +rewrite mulnC eq_sym (bigID (mem Dn)) /= mulnC. +rewrite big1 ?mul1n => [|A]; last by move/out_n->. +rewrite big_setU1 //= /n setU11 mul1n. +rewrite (eq_bigr (fun _ => #|'C_M(R)| ^ #|R|)%N) => [|A R_A]; last first. + rewrite inE R_A orbT mul1n; case/imsetP: R_A => x Kx ->. + suffices nMx: x \in 'N(M) by rewrite -{1}(normP nMx) centJ -conjIg !cardJg. + exact: subsetP (subsetP sKG x Kx). +rewrite mulnC prod_nat_const card_orbit astab1JG. +have ->: 'N_K(R) = 1 by rewrite -(setIidPl sKG) -setIA -normD1 snRG tiKR. +rewrite indexg1 -expnMn eq_sym eqn_exp2r ?cardG_gt0 //; move/eqP=> eq_fix. +split=> // [regR | regK]. + rewrite centsC (sameP setIidPl eqP) eqEcard subsetIl /=. + move: eq_fix; rewrite regR cards1 exp1n mul1n => <-. + suffices ->: 'C_M(G) = 1 by rewrite cards1 exp1n mul1n. + by apply/trivgP; rewrite -regR setIS ?centS //; case/sdprod_context: defG. +move: eq_fix; rewrite regK cards1 muln1 => <-. +suffices ->: 'C_M(G) = 1 by rewrite cards1 exp1n mul1n. +by apply/trivgP; rewrite -regK setIS ?centS. +Qed. + +End Wielandt_Frobenius. + +(* This is B & G, Theorem 3.10. *) +Theorem Frobenius_primact gT (G K R M : {group gT}) : + [Frobenius G = K ><| R] -> solvable G -> + G \subset 'N(M) -> solvable M -> M :!=: 1 -> + (*1*) coprime #|M| #|G| -> + (*2*) semiprime M R -> + (*3*) 'C_M(K) = 1 -> + [/\ prime #|R|, + #|M| = (#|'C_M(R)| ^ #|R|)%N + & cyclic 'C_M(R) -> K^`(1) \subset 'C_K(M)]. +Proof. +move: {2}_.+1 (ltnSn #|M|) => n; elim: n => // n IHn in gT G K R M *. +rewrite ltnS => leMn frobG solG nMG solM ntM coMG primRM tcKM. +case: (Frobenius_Wielandt_fixpoint frobG nMG) => // _ _ /(_ tcKM) oM. +have [defG ntK ntR ltKG _]:= Frobenius_context frobG. +have Rpr: prime #|R|. + have [R1 | [r r_pr]] := trivgVpdiv R; first by case/eqP: ntR. + case/Cauchy=> // x Rx ox; pose R0 := <[x]>; pose G0 := K <*> R0. + have [_ defKR nKR tiKR] := sdprodP defG. + have sR0R: R0 \subset R by rewrite cycle_subG. + have sG0G: G0 \subset G by rewrite /G0 -genM_join gen_subG -defKR mulgS. + have nKR0 := subset_trans sR0R nKR; have nMG0 := subset_trans sG0G nMG. + have ntx: <[x]> != 1 by rewrite cycle_eq1 -order_gt1 ox prime_gt1. + have [tcRM | ntcRM] := eqVneq 'C_M(R) 1. + by rewrite -cardG_gt1 oM tcRM cards1 exp1n in ntM. + have frobG0: [Frobenius G0 = K ><| R0]. + apply/Frobenius_semiregularP=> // [|y /setD1P[nty x_y]]. + by apply: sdprodEY nKR0 (trivgP _); rewrite -tiKR setIS. + by apply: (Frobenius_reg_ker frobG); rewrite !inE nty (subsetP sR0R). + case: (Frobenius_Wielandt_fixpoint frobG0 nMG0 (coprimegS _ coMG)) => // _ _. + move/(_ tcKM)/eqP; rewrite oM cent_cycle. + rewrite primRM; last by rewrite !inE Rx andbT -cycle_eq1. + by rewrite eqn_exp2l ?cardG_gt1 // -orderE ox => /eqP->. +split=> // cyc_cMR. +have nM_MG: M <*> G \subset 'N(M) by rewrite join_subG normG. +have [V minV sVM] := minnormal_exists ntM nM_MG. +have [] := minnormal_solvable minV sVM solM. +rewrite join_subG; case/andP=> nVM nVG ntV; case/is_abelemP=> [q q_pr abelV]. +have coVG := coprimeSg sVM coMG; have solV := solvableS sVM solM. +have cVK': K^`(1) \subset 'C_K(V). + case: (eqVneq 'C_V(R) 1) => [tcVR | ntcRV]. + case: (Frobenius_Wielandt_fixpoint frobG nVG) => // _. + by move/(_ tcVR)=> cVK _; rewrite (setIidPl cVK) der_sub. + have ocVR: #|'C_V(R)| = q. + have [u def_u]: exists u, 'C_V(R) = <[u]>. + by apply/cyclicP; apply: cyclicS (setSI _ sVM) cyc_cMR. + rewrite def_u -orderE (abelem_order_p abelV) -?cycle_eq1 -?def_u //. + by rewrite -cycle_subG -def_u subsetIl. + apply: (Frobenius_prime_cent_prime _ defG _ _ abelV) => //. + by case/prime_FrobeniusP: frobG. + by rewrite (coprime_p'group _ (abelem_pgroup abelV) ntV) // coprime_sym. +have cMK': K^`(1) / V \subset 'C_(K / V)(M / V). + have [-> | ntMV] := eqVneq (M / V) 1. + by rewrite subsetI cents1 quotientS ?der_sub. + have coKR := Frobenius_coprime frobG. + case/prime_FrobeniusP: frobG => //. + case/sdprod_context=> nsKG sRG defKR nKR tiKR regR; have [sKG _] := andP nsKG. + have nVK := subset_trans sKG nVG; have nVR := subset_trans sRG nVG. + have RVpr: prime #|R / V|. + rewrite card_quotient // -indexgI setIC coprime_TIg ?(coprimegS sRG) //. + by rewrite indexg1. + have frobGV: [Frobenius G / V = (K / V) ><| (R / V)]. + apply/prime_FrobeniusP; rewrite // -?cardG_gt1 ?card_quotient //. + rewrite -indexgI setIC coprime_TIg ?(coprimegS sKG) //. + by rewrite indexg1 cardG_gt1. + rewrite -coprime_norm_quotient_cent ?(coprimegS sRG) //= regR quotient1. + rewrite -defKR quotientMl // sdprodE ?quotient_norms //. + by rewrite coprime_TIg ?coprime_morph. + have ltMVn: #|M / V| < n by apply: leq_trans leMn; rewrite ltn_quotient. + rewrite quotient_der //; move/IHn: frobGV. + case/(_ _ ltMVn); rewrite ?quotient_sol ?quotient_norms ?coprime_morph //. + - move=> Vx; case/setD1P=> ntVx; case/morphimP=> x nVx Rx defVx. + rewrite defVx /= -cent_cycle -quotient_cycle //; congr 'C__(_ / V). + apply/eqP; rewrite eqEsubset cycle_subG Rx /=. + apply: contraR ntVx; move/(prime_TIg Rpr); move/trivgP. + rewrite defVx /= (setIidPr _) cycle_subG //; move/set1P->. + by rewrite morph1. + - rewrite -coprime_norm_quotient_cent ?(coprimegS sKG) ?(subset_trans sKG) //. + by rewrite tcKM quotient1. + move=> _ _ -> //; rewrite -coprime_quotient_cent ?quotient_cyclic //. + by rewrite (coprimegS sRG). +rewrite !subsetI in cVK' cMK' *. +case/andP: cVK' => sK'K cVK'; case/andP: cMK' => _ cMVK'; rewrite sK'K. +have sK'G: K^`(1) \subset G by rewrite (subset_trans sK'K) ?proper_sub. +have coMK': coprime #|M| #|K^`(1)| := coprimegS sK'G coMG. +rewrite (stable_factor_cent cVK') // /stable_factor /normal sVM nVM !andbT. +by rewrite commGC -quotient_cents2 // (subset_trans sK'G). +Qed. + +End BGsection3. diff --git a/mathcomp/odd_order/BGsection4.v b/mathcomp/odd_order/BGsection4.v new file mode 100644 index 0000000..c33bd2f --- /dev/null +++ b/mathcomp/odd_order/BGsection4.v @@ -0,0 +1,1413 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div. +Require Import fintype finfun bigop ssralg finset prime binomial. +Require Import fingroup morphism automorphism perm quotient action gproduct. +Require Import gfunctor commutator zmodp cyclic center pgroup gseries nilpotent. +Require Import sylow abelian maximal extremal hall. +Require Import matrix mxalgebra mxrepresentation mxabelem. +Require Import BGsection1 BGsection2. + +(******************************************************************************) +(* This file covers B & G, Section 4, i.e., the proof of structure theorems *) +(* for solvable groups with a small (of rank at most 2) Fitting subgroup. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Section Section4. + +Implicit Type gT : finGroupType. +Implicit Type p : nat. + +(* B & G, Lemma 4.1 (also, Gorenstein, 1.3.4, and Aschbacher, ex. 2.4) is *) +(* covered by Lemma center_cyclic_abelian, in center.v. *) + +(* B & G, Lemma 4.2 is covered by Lemmas commXg, commgX, commXXg (for 4.2(a)) *) +(* and expMg_Rmul (for 4.2(b)) in commutators.v. *) + +(* This is B & G, Proposition 4.3. *) +Proposition exponent_odd_nil23 gT (R : {group gT}) p : + p.-group R -> odd #|R| -> nil_class R <= 2 + (p > 3) -> + [/\ (*a*) exponent 'Ohm_1(R) %| p + & (*b*) R^`(1) \subset 'Ohm_1(R) -> + {in R &, {morph expgn^~ p : x y / x * y}}]. +Proof. +move=> pR oddR classR. +pose f n := 'C(n, 3); pose g n := 'C(n, 3).*2 + 'C(n, 2). +have fS n: f n.+1 = 'C(n, 2) + f n by rewrite /f binS addnC. +have gS n: g n.+1 = 'C(n, 2).*2 + 'C(n, 1) + g n. + by rewrite /g !binS doubleD -!addnA; do 3!nat_congr. +have [-> | ntR] := eqsVneq R 1. + rewrite Ohm1 exponent1 der_sub dvd1n; split=> // _ _ _ /set1P-> /set1P->. + by rewrite !(mulg1, expg1n). +have{ntR} [p_pr p_dv_R _] := pgroup_pdiv pR ntR. +have pdivbin2: p %| 'C(p, 2). + by rewrite prime_dvd_bin //= odd_prime_gt2 // (dvdn_odd p_dv_R). +have p_dv_fp: p > 3 -> p %| f p by move=> pgt3; apply: prime_dvd_bin. +have p_dv_gp: p > 3 -> p %| g p. + by move=> pgt3; rewrite dvdn_add // -muln2 dvdn_mulr // p_dv_fp. +have exp_dv_p x m (S : {group gT}): + exponent S %| p -> p %| m -> x \in S -> x ^+ m = 1. +- move=> expSp p_dv_m Sx; apply/eqP; rewrite -order_dvdn. + by apply: dvdn_trans (dvdn_trans expSp p_dv_m); apply: dvdn_exponent. +have p3_L21: p <= 3 -> {in R & &, forall u v w, [~ u, v, w] = 1}. + move=> lep3 u v w Ru Rv Rw; rewrite (ltnNge 3) lep3 nil_class2 in classR. + by apply/eqP/commgP; red; rewrite (centerC Rw) // (subsetP classR) ?mem_commg. +have{fS gS} expMR_fg: {in R &, forall u v n (r := [~ v, u]), + (u * v) ^+ n = u ^+ n * v ^+ n * r ^+ 'C(n, 2) + * [~ r, u] ^+ f n * [~ r, v] ^+ g n}. +- move=> u v Ru Rv n r; have Rr: r \in R by exact: groupR. + have cRr: {in R &, forall x y, commute x [~ r, y]}. + move=> x y Rx Ry /=; red; rewrite (centerC Rx) //. + have: [~ r, y] \in 'L_3(R) by rewrite !mem_commg. + by apply: subsetP; rewrite -nil_class3 (leq_trans classR) // !ltnS leq_b1. + elim: n => [|n IHn]; first by rewrite !mulg1. + rewrite 3!expgSr {}IHn -!mulgA (mulgA (_ ^+ f n)); congr (_ * _). + rewrite -commuteM; try by apply: commuteX; red; rewrite cRr ?groupM. + rewrite -mulgA; do 2!rewrite (mulgA _ u) (commgC _ u) -2!mulgA. + congr (_ * (_ * _)); rewrite (mulgA _ v). + have ->: [~ v ^+ n, u] = r ^+ n * [~ r, v] ^+ 'C(n, 2). + elim: n => [|n IHn]; first by rewrite comm1g mulg1. + rewrite !expgS commMgR -/r {}IHn commgX; last exact: cRr. + rewrite binS bin1 addnC expgD -2!mulgA; congr (_ * _); rewrite 2!mulgA. + by rewrite commuteX2 // /commute cRr. + rewrite commXg 1?commuteX2 -?[_ * v]commuteX; try exact: cRr. + rewrite mulgA {1}[mulg]lock mulgA -mulgA -(mulgA v) -!expgD -fS -lock. + rewrite -{2}(bin1 n) addnC -binS -2!mulgA (mulgA _ v) (commgC _ v). + rewrite -commuteX; last by red; rewrite cRr ?(Rr, groupR, groupX, groupM). + rewrite -3!mulgA; congr (_ * (_ * _)); rewrite 2!mulgA. + rewrite commXg 1?commuteX2; try by red; rewrite cRr 1?groupR. + by rewrite -!mulgA -!expgD addnCA binS addnAC addnn addnC -gS. +have expR1p: exponent 'Ohm_1(R) %| p. + elim: _.+1 {-2 4}R (ltnSn #|R|) (subxx R) => // n IHn Q leQn sQR. + rewrite (OhmE 1 (pgroupS sQR pR)) expn1 -sub_LdivT. + rewrite gen_set_id ?subsetIr //. + apply/group_setP; rewrite !inE group1 expg1n /=. + split=> // x y /LdivP[Qx xp1] /LdivP[Qy yp1]; rewrite !inE groupM //=. + have sxQ: <[x]> \subset Q by rewrite cycle_subG. + have [{sxQ}defQ|[S maxS /= sxS]] := maximal_exists sxQ. + rewrite expgMn; first by rewrite xp1 yp1 mulg1. + by apply: (centsP (cycle_abelian x)); rewrite ?defQ. + have:= maxgroupp maxS; rewrite properEcard => /andP[sSQ ltSQ]. + have pQ := pgroupS sQR pR; have pS := pgroupS sSQ pQ. + have{ltSQ leQn} ltSn: #|S| < n by exact: leq_trans ltSQ _. + have expS1p := IHn _ ltSn (subset_trans sSQ sQR). + have defS1 := Ohm1Eexponent p_pr expS1p; move/exp_dv_p: expS1p => expS1p. + have nS1Q: [~: Q, 'Ohm_1(S)] \subset 'Ohm_1(S). + rewrite commg_subr (char_norm_trans (Ohm_char 1 S)) ?normal_norm //. + exact: p_maximal_normal pQ maxS. + have S1x : x \in 'Ohm_1(S) by rewrite defS1 !inE -cycle_subG sxS xp1 /=. + have S1yx : [~ y, x] \in 'Ohm_1(S) by rewrite (subsetP nS1Q) ?mem_commg. + have S1yxx : [~ y, x, x] \in 'Ohm_1(S) by rewrite groupR. + have S1yxy : [~ y, x, y] \in 'Ohm_1(S). + by rewrite -invg_comm groupV (subsetP nS1Q) 1?mem_commg. + rewrite expMR_fg ?(subsetP sQR) //= xp1 yp1 expS1p ?mul1g //. + case: (leqP p 3) => [p_le3 | p_gt3]; last by rewrite ?expS1p ?mul1g; auto. + by rewrite !p3_L21 // ?(subsetP sQR) // !expg1n mulg1. +split=> // sR'R1 x y Rx Ry; rewrite -[x ^+ p * _]mulg1 expMR_fg // -2!mulgA //. +have expR'p := exp_dv_p _ _ _ (dvdn_trans (exponentS sR'R1) expR1p). +congr (_ * _); rewrite expR'p ?mem_commg // mul1g. +case: (leqP p 3) => [p_le3 | p_gt3]. + by rewrite !p3_L21 // ?(subsetP sQR) // !expg1n mulg1. +by rewrite !expR'p 1?mem_commg ?groupR ?mulg1; auto. +Qed. + +(* Part (a) of B & G, Proposition 4.4 is covered in file maximal.v by lemmas *) +(* max_SCN and SCN_max. *) + +(* This is B & G, Proposition 4.4(b), or Gorenstein 7.6.5. *) +Proposition SCN_Sylow_cent_dprod gT (R G A : {group gT}) p : + p.-Sylow(G) R -> A \in 'SCN(R) -> 'O_p^'('C_G(A)) \x A = 'C_G(A). +Proof. +move=> sylR scnA; have [[sRG _] [nAR CRA_A]] := (andP sylR, SCN_P scnA). +set C := 'C_G(A); have /maxgroupP[/andP[nAG abelA] maxA] := SCN_max scnA. +have CiP_eq : C :&: R = A by rewrite -CRA_A setIC setIA (setIidPl sRG). +have sylA: p.-Sylow(C) A. + rewrite -CiP_eq; apply: (Sylow_setI_normal (subcent_normal _ _)). + by apply: pHall_subl sylR; rewrite ?subsetIl // subsetI sRG normal_norm. +rewrite dprodEsd; last first. + by rewrite centsC (subset_trans (pcore_sub _ _)) ?subsetIr. +by apply: Burnside_normal_complement; rewrite // subIset ?subsetIr. +Qed. + +(* This is B & G, Lemma 4.5(b), or Gorenstein, 5.4.4 and 5.5.5. *) +Lemma Ohm1_extremal_odd gT (R : {group gT}) p x : + p.-group R -> odd #|R| -> ~~ cyclic R -> x \in R -> #|R : <[x]>| = p -> + ('Ohm_1(R))%G \in 'E_p^2(R). +Proof. +move=> pR oddR ncycR Rx ixR; rewrite -cycle_subG in Rx. +have ntR: R :!=: 1 by apply: contra ncycR; move/eqP->; exact: cyclic1. +have [p_pr _ [e oR]]:= pgroup_pdiv pR ntR. +case p2: (p == 2); first by rewrite oR odd_exp (eqP p2) in oddR. +have [cRR | not_cRR] := orP (orbN (abelian R)). + rewrite 2!inE Ohm_sub Ohm1_abelem // -p_rank_abelian //= eqn_leq. + rewrite -rank_pgroup // ltnNge -abelian_rank1_cyclic // ncycR andbT. + have maxX: maximal <[x]> R by rewrite (p_index_maximal Rx) ?ixR. + have nsXR: <[x]> <| R := p_maximal_normal pR maxX. + have [_ [y Ry notXy]] := properP (maxgroupp maxX). + have defR: <[x]> * <[y]> = R. + by apply: mulg_normal_maximal; rewrite ?cycle_subG. + rewrite -grank_abelian // -(genGid R) -defR genM_join joing_idl joing_idr. + by rewrite (leq_trans (grank_min _)) // cards2 ltnS leq_b1. +have{x Rx ixR} [e_gt1 isoR]: 2 < e.+1 /\ R \isog 'Mod_(p ^ e.+1). + have:= maximal_cycle_extremal pR not_cRR (cycle_cyclic x) Rx ixR. + rewrite p2 orbF /extremal_class oR pfactorKpdiv // pdiv_pfactor //. + by do 4!case: andP => //. +have [[x y] genR modR] := generators_modular_group p_pr e_gt1 isoR. +have [_ _ _ _] := modular_group_structure p_pr e_gt1 genR isoR modR. +rewrite xpair_eqE p2; case/(_ 1%N) => // _ oR1. +by rewrite 2!inE Ohm_sub oR1 pfactorK ?abelem_Ohm1 ?(card_p2group_abelian p_pr). +Qed. + +Section OddNonCyclic. + +Variables (gT : finGroupType) (p : nat) (R : {group gT}). +Hypotheses (pR : p.-group R) (oddR : odd #|R|) (ncycR : ~~ cyclic R). + +(* This is B & G, Lemma 4.5(a), or Gorenstein 5.4.10. *) +Lemma ex_odd_normal_p2Elem : {S : {group gT} | S <| R & S \in 'E_p^2(R)}. +Proof. +have [M minM]: {M | [min M | M <| R & ~~ cyclic M]}. + by apply: ex_mingroup; exists R; rewrite normal_refl. +have{minM} [[nsMR ncycM] [_ minM]] := (andP (mingroupp minM), mingroupP minM). +have [sMR _] := andP nsMR; have pM := pgroupS sMR pR. +exists ('Ohm_1(M))%G; first exact: char_normal_trans (Ohm_char 1 M) nsMR. +apply: (subsetP (pnElemS _ _ sMR)). +have [M1 | ntM] := eqsVneq M 1; first by rewrite M1 cyclic1 in ncycM. +have{ntM} [p_pr _ [e oM]] := pgroup_pdiv pM ntM. +have le_e_M: e <= logn p #|M| by rewrite ltnW // oM pfactorK. +have{le_e_M} [N [sNM nsNR] oN] := normal_pgroup pR nsMR le_e_M. +have ltNM: ~~ (#|N| >= #|M|) by rewrite -ltnNge oM oN ltn_exp2l ?prime_gt1. +have cycN : cyclic N by apply: contraR ltNM => ncycN; rewrite minM //= nsNR. +case/cyclicP: cycN => x defN; have Mx : x \in M by rewrite -cycle_subG -defN. +apply: Ohm1_extremal_odd Mx _; rewrite ?(oddSg sMR) //. +by rewrite -divgS /= -defN // oM oN expnS mulnK // expn_gt0 prime_gt0. +Qed. + +(* This is B & G, Lemma 4.5(c). *) +Lemma Ohm1_odd_ucn2 (Z := 'Ohm_1('Z_2(R))) : ~~ cyclic Z /\ exponent Z %| p. +Proof. +have [S nsSR Ep2S] := ex_odd_normal_p2Elem; have p_pr := pnElem_prime Ep2S. +have [sSR abelS dimS] := pnElemP Ep2S; have [pS cSS expSp]:= and3P abelS. +pose SR := [~: S, R]; pose SRR := [~: SR, R]. +have nilR := pgroup_nil pR. +have ntS: S :!=: 1 by rewrite -rank_gt0 (rank_abelem abelS) dimS. +have sSR_S: SR \subset S by rewrite commg_subl normal_norm. +have sSRR_SR: SRR \subset SR by rewrite commSg. +have sSR_R := subset_trans sSR_S sSR. +have{ntS} prSR: SR \proper S. + by rewrite (nil_comm_properl nilR) // subsetI subxx -commg_subl. +have SRR1: SRR = 1. + have [SR1 | ntSR] := eqVneq SR 1; first by rewrite /SRR SR1 comm1G. + have prSRR: SRR \proper SR. + rewrite /proper sSRR_SR; apply: contra ntSR => sSR_SRR. + by rewrite (forall_inP nilR) // subsetI sSR_R. + have pSR := pgroupS sSR_R pR; have pSRR := pgroupS sSRR_SR pSR. + have [_ _ [e oSR]] := pgroup_pdiv pSR ntSR; have [f oSRR] := p_natP pSRR. + have e0: e = 0. + have:= proper_card prSR; rewrite oSR (card_pnElem Ep2S). + by rewrite ltn_exp2l ?prime_gt1 // !ltnS leqn0; move/eqP. + apply/eqP; have:= proper_card prSRR; rewrite trivg_card1 oSR oSRR e0. + by rewrite ltn_exp2l ?prime_gt1 // ltnS; case f. +have sSR_ZR: [~: S, R] \subset 'Z(R). + by rewrite subsetI sSR_R /=; apply/commG1P. +have sS_Z2R: S \subset 'Z_2(R). + rewrite ucnSnR; apply/subsetP=> s Ss; rewrite inE (subsetP sSR) //= ucn1. + by rewrite (subset_trans _ sSR_ZR) ?commSg ?sub1set. +have sZ2R_R := ucn_sub 2 R; have pZ2R := pgroupS sZ2R_R pR. +have pZ: p.-group Z. + apply: pgroupS pR; apply: subset_trans (Ohm_sub _ _) (ucn_sub 2 R). +have sSZ: S \subset Z. + by rewrite /Z (OhmE 1 pZ2R) sub_gen // subsetI sS_Z2R sub_LdivT. +have ncycX: ~~ cyclic S by rewrite (abelem_cyclic abelS) dimS. +split; first by apply: contra ncycX; exact: cyclicS. +have nclZ2R : nil_class 'Z_2(R) <= 2 + _ := leq_trans (nil_class_ucn _ _) _. +by have [] := exponent_odd_nil23 pZ2R (oddSg sZ2R_R oddR) (nclZ2R _ _). +Qed. + +End OddNonCyclic. + +(* Some "obvious" consequences of the above, which are used casually and *) +(* pervasively throughout B & G. *) +Definition odd_pgroup_rank1_cyclic := odd_pgroup_rank1_cyclic. (* in extremal *) + +Lemma odd_rank1_Zgroup gT (G : {group gT}) : + odd #|G| -> Zgroup G = ('r(G) <= 1). +Proof. +move=> oddG; apply/forallP/idP=> [ZgG | rG_1 P]. + have [p p_pr ->]:= rank_witness G; have [P sylP]:= Sylow_exists p G. + have [sPG pP _] := and3P sylP; have oddP := oddSg sPG oddG. + rewrite -(p_rank_Sylow sylP) -(odd_pgroup_rank1_cyclic pP) //. + by apply: (implyP (ZgG P)); apply: (p_Sylow sylP). +apply/implyP=> /SylowP[p p_pr /and3P[sPG pP _]]. +rewrite (odd_pgroup_rank1_cyclic pP (oddSg sPG oddG)). +by apply: leq_trans (leq_trans (p_rank_le_rank p G) rG_1); apply: p_rankS. +Qed. + +(* This is B & G, Proposition 4.6 (a stronger version of Lemma 4.5(a)). *) +Proposition odd_normal_p2Elem_exists gT p (R S : {group gT}) : + p.-group R -> odd #|R| -> S <| R -> ~~ cyclic S -> + exists E : {group gT}, [/\ E \subset S, E <| R & E \in 'E_p^2(R)]. +Proof. +move=> pR oddR nsSR ncycS; have sSR := normal_sub nsSR. +have{sSR ncycS} []:= Ohm1_odd_ucn2 (pgroupS sSR pR) (oddSg sSR oddR) ncycS. +set Z := 'Ohm_1(_) => ncycZ expZp. +have chZS: Z \char S := char_trans (Ohm_char 1 _) (ucn_char 2 S). +have{nsSR} nsZR: Z <| R := char_normal_trans chZS nsSR. +have [sZR _] := andP nsZR; have pZ: p.-group Z := pgroupS sZR pR. +have geZ2: 2 <= logn p #|Z|. + rewrite (odd_pgroup_rank1_cyclic pZ (oddSg sZR oddR)) -ltnNge /= -/Z in ncycZ. + by case/p_rank_geP: ncycZ => E; case/pnElemP=> sEZ _ <-; rewrite lognSg. +have [E [sEZ nsER oE]] := normal_pgroup pR nsZR geZ2. +have [sER _] := andP nsER; have{pR} pE := pgroupS sER pR. +have{geZ2} p_pr: prime p by move: geZ2; rewrite lognE; case: (prime p). +have{oE p_pr} dimE2: logn p #|E| = 2 by rewrite oE pfactorK. +exists E; split; rewrite ?(subset_trans _ (char_sub chZS)) {chZS nsZR}//. +rewrite !inE /abelem sER pE (p2group_abelian pE) dimE2 //= andbT. +by apply: (dvdn_trans _ expZp); apply: exponentS. +Qed. + +(* This is B & G, Lemma 4.7, and (except for the trivial converse) Gorenstein *) +(* 5.4.15 and Aschbacher 23.17. *) +Lemma rank2_SCN3_empty gT p (R : {group gT}) : + p.-group R -> odd #|R| -> ('r(R) <= 2) = ('SCN_3(R) == set0). +Proof. +move=> pR oddR; apply/idP/idP=> [leR2 | SCN_3_empty]. + apply/set0Pn=> [[A /setIdP[/SCN_P[/andP[sAR _] _]]]]. + by rewrite ltnNge (leq_trans (rankS sAR)). +rewrite (rank_pgroup pR) leqNgt; apply/negP=> gtR2. +have ncycR: ~~ cyclic R by rewrite (odd_pgroup_rank1_cyclic pR) // -ltnNge ltnW. +have{ncycR} [Z nsZR] := ex_odd_normal_p2Elem pR oddR ncycR. +case/pnElemP=> sZR abelZ dimZ2; have [pZ cZZ _] := and3P abelZ. +have{SCN_3_empty} defZ: 'Ohm_1('C_R(Z)) = Z. + apply: (Ohm1_cent_max_normal_abelem _ pR). + by have:= oddSg sZR oddR; rewrite (card_pgroup pZ) dimZ2 odd_exp. + apply/maxgroupP; split=> [|H /andP[nsHR abelH] sZH]; first exact/andP. + have [pH cHH _] := and3P abelH; apply/eqP; rewrite eq_sym eqEproper sZH /=. + pose normal_abelian := [pred K : {group gT} | K <| R & abelian K]. + have [|K maxK sHK] := @maxgroup_exists _ normal_abelian H; first exact/andP. + apply: contraL SCN_3_empty => ltZR; apply/set0Pn; exists K. + rewrite inE (max_SCN pR) {maxK}//= -dimZ2 (leq_trans _ (rankS sHK)) //. + by rewrite (rank_abelem abelH) properG_ltn_log. +have{gtR2} [A] := p_rank_geP gtR2; pose H := 'C_A(Z); pose K := H <*> Z. +case/pnElemP=> sAR abelA dimA3; have [pA cAA _] := and3P abelA. +have{nsZR} nZA := subset_trans sAR (normal_norm nsZR). +have sHA: H \subset A := subsetIl A _; have abelH := abelemS sHA abelA. +have geH2: logn p #|H| >= 2. + rewrite -ltnS -dimA3 -(Lagrange sHA) lognM // -addn1 leq_add2l /= -/H. + by rewrite logn_quotient_cent_abelem ?dimZ2. +have{abelH} abelK : p.-abelem K. + by rewrite (cprod_abelem _ (cprodEY _)) 1?centsC ?subsetIr ?abelH. +suffices{sZR cZZ defZ}: 'r(Z) < 'r(K). + by rewrite ltnNge -defZ rank_Ohm1 rankS // join_subG setSI // subsetI sZR. +rewrite !(@rank_abelem _ p) // properG_ltn_log ?abelem_pgroup //= -/K properE. +rewrite joing_subr join_subG subxx andbT subEproper; apply: contraL geH2. +case/predU1P=> [defH | ltHZ]; last by rewrite -ltnNge -dimZ2 properG_ltn_log. +rewrite -defH [H](setIidPl _) ?dimA3 // in dimZ2. +by rewrite centsC -defH subIset // -abelianE cAA. +Qed. + +(* This is B & G, Proposition 4.8(a). *) +Proposition rank2_exponent_p_p3group gT (R : {group gT}) p : + p.-group R -> rank R <= 2 -> exponent R %| p -> logn p #|R| <= 3. +Proof. +move=> pR rankR expR. +have [A max_na_A]: {A | [max A | A <| R & abelian A]}. + by apply: ex_maxgroup; exists 1%G; rewrite normal1 abelian1. +have {max_na_A} SCN_A := max_SCN pR max_na_A. +have cAA := SCN_abelian SCN_A; case/SCN_P: SCN_A => nAR cRAA. +have sAR := normal_sub nAR; have pA := pgroupS sAR pR. +have abelA : p.-abelem A. + by rewrite /abelem pA cAA /= (dvdn_trans (exponentS sAR) expR). +have cardA : logn p #|A| <= 2. + by rewrite -rank_abelem // (leq_trans (rankS sAR) rankR). +have cardRA : logn p #|R : A| <= 1. + by rewrite -cRAA logn_quotient_cent_abelem // (normal_norm nAR). +rewrite -(Lagrange sAR) lognM ?cardG_gt0 //. +by apply: (leq_trans (leq_add cardA cardRA)). +Qed. + +(* This is B & G, Proposition 4.8(b). *) +Proposition exponent_Ohm1_rank2 gT p (R : {group gT}) : + p.-group R -> 'r(R) <= 2 -> p > 3 -> exponent 'Ohm_1(R) %| p. +Proof. +move=> pR rR p_gt3; wlog p_pr: / prime p. + have [-> | ntR] := eqsVneq R 1; first by rewrite Ohm1 exponent1 dvd1n. + by apply; have [->] := pgroup_pdiv pR ntR. +wlog minR: R pR rR / forall S, gval S \proper R -> exponent 'Ohm_1(S) %| p. + elim: {R}_.+1 {-2}R (ltnSn #|R|) => // m IHm R leRm in pR rR * => IH. + apply: (IH) => // S; rewrite properEcard; case/andP=> sSR ltSR. + exact: IHm (leq_trans ltSR _) (pgroupS sSR pR) (leq_trans (rankS sSR) rR) IH. +wlog not_clR_le3: / ~~ (nil_class R <= 3). + case: leqP => [clR_le3 _ | _ -> //]. + have [||-> //] := exponent_odd_nil23 pR; last by rewrite p_gt3. + by apply: odd_pgroup_odd pR; case/even_prime: p_pr p_gt3 => ->. +rewrite -sub_LdivT (OhmE 1 pR) gen_set_id ?subsetIr //. +apply/group_setP; rewrite !inE group1 expg1n. +split=> //= x y; case/LdivP=> Rx xp1; case/LdivP=> Ry yp1. +rewrite !inE groupM //=; apply: contraR not_clR_le3 => nt_xyp. +pose XY := <[x]> <*> <[y]>. +have [XYx XYy]: x \in XY /\ y \in XY by rewrite -!cycle_subG; exact/joing_subP. +have{nt_xyp} defR: XY = R. + have sXY_R : XY \subset R by rewrite join_subG !cycle_subG Rx Ry. + have pXY := pgroupS sXY_R pR; have [// | ltXY_R] := eqVproper sXY_R. + rewrite (exponentP (minR _ ltXY_R)) ?eqxx // in nt_xyp. + by rewrite (OhmE 1 pXY) groupM ?mem_gen ?inE ?XYx ?XYy /= ?xp1 ?yp1. +have sXR: <[x]> \subset R by rewrite cycle_subG. +have [<- | ltXR] := eqVproper sXR. + by rewrite 2?leqW // nil_class1 cycle_abelian. +have [S maxS sXS]: {S : {group gT} | maximal S R & <[x]> \subset S}. + exact: maxgroup_exists. +have nsSR: S <| R := p_maximal_normal pR maxS; have [sSR _] := andP nsSR. +have{nsSR} nsS1R: 'Ohm_1(S) <| R := char_normal_trans (Ohm_char 1 S) nsSR. +have [sS1R nS1R] := andP nsS1R; have pS1 := pgroupS sS1R pR. +have expS1p: exponent 'Ohm_1(S) %| p := minR S (maxgroupp maxS). +have{expS1p} dimS1: logn p #|'Ohm_1(S)| <= 3. + exact: rank2_exponent_p_p3group pS1 (leq_trans (rankS sS1R) rR) expS1p. +have sXS1: <[x]> \subset 'Ohm_1(S). + rewrite cycle_subG /= (OhmE 1 (pgroupS sSR pR)) mem_gen //. + by rewrite !inE -cycle_subG sXS xp1 /=. +have dimS1b: logn p #|R / 'Ohm_1(S)| <= 1. + rewrite -quotientYidl // -defR joingA (joing_idPl sXS1). + rewrite quotientYidl ?cycle_subG ?(subsetP nS1R) //. + rewrite (leq_trans (logn_quotient _ _ _)) // -(pfactorK 1 p_pr). + by rewrite dvdn_leq_log ?prime_gt0 // order_dvdn yp1. +rewrite (leq_trans (nil_class_pgroup pR)) // geq_max /= -subn1 leq_subLR. +by rewrite -(Lagrange sS1R) lognM // -card_quotient // addnC leq_add. +Qed. + +(* This is B & G, Lemma 4.9. *) +Lemma quotient_p2_Ohm1 gT p (R : {group gT}) : + p.-group R -> p > 3 -> logn p #|'Ohm_1(R)| <= 2 -> + forall T : {group gT}, T <| R -> logn p #|'Ohm_1(R / T)| <= 2. +Proof. +move=> pR p_gt3 dimR1; move: {2}_.+1 (ltnSn #|R|) => n. +elim: n => // n IHn in gT R pR dimR1 *; rewrite ltnS => leRn. +apply/forall_inP/idPn; rewrite negb_forall_in. +case/existsP/ex_mingroup=> T /mingroupP[/andP[nsTR dimRb1] minT]. +have [sTR nTR] := andP nsTR; have pT: p.-group T := pgroupS sTR pR. +pose card_iso_Ohm := card_isog (gFisog [igFun of Ohm 1] _). +have ntT: T :!=: 1; last have p_pr: prime p by have [] := pgroup_pdiv pT ntT. + apply: contraNneq dimRb1 => ->. + by rewrite -(card_iso_Ohm _ _ _ _ (quotient1_isog R)). +have{minT} dimT: logn p #|T| = 1%N. + have [Z EpZ]: exists Z, Z \in 'E_p^1(T :&: 'Z(R)). + apply/p_rank_geP; rewrite -rank_pgroup ?(pgroupS (subsetIl T _)) //. + by rewrite rank_gt0 (meet_center_nil (pgroup_nil pR)). + have [sZ_ZT _ dimZ] := pnElemP EpZ; have [sZT sZZ] := subsetIP sZ_ZT. + have{sZZ} nsZR: Z <| R := sub_center_normal sZZ. + rewrite -(minT Z) // nsZR; apply: contra dimRb1 => dimRz1. + rewrite -(card_iso_Ohm _ _ _ _ (third_isog sZT nsZR nsTR)) /=. + rewrite IHn ?quotient_pgroup ?quotient_normal ?(leq_trans _ leRn) //. + by rewrite ltn_quotient ?(subset_trans sZT) // (nt_pnElem EpZ). +have pRb: p.-group (R / T) by apply: quotient_pgroup. +have{IHn} minR (Ub : {group coset_of T}): + Ub \subset R / T -> ~~ (logn p #|'Ohm_1(Ub)| <= 2) -> R / T = Ub. +- case/inv_quotientS=> // U -> sTU sUR dimUb; congr (_ / T). + apply/eqP; rewrite eq_sym eqEcard sUR leqNgt; apply: contra dimUb => ltUR. + rewrite IHn ?(pgroupS sUR) ?(normalS _ sUR) ?(leq_trans ltUR) //. + by rewrite (leq_trans _ dimR1) ?lognSg ?OhmS. +have [dimRb eRb]: logn p #|R / T| = 3 /\ exponent (R / T) %| p. + have [Rb_gt2 | Rb_le2] := ltnP 2 'r_p(R / T). + have [Ub Ep3Ub] := p_rank_geP Rb_gt2. + have [sUbR abelUb dimUb] := pnElemP Ep3Ub; have [_ _ eUb] := and3P abelUb. + by rewrite (minR Ub) // (Ohm1_id abelUb) dimUb. + rewrite -rank_pgroup // in Rb_le2. + have eRb: exponent (R / T) %| p. + by rewrite (minR _ (Ohm_sub 1 _)) ?exponent_Ohm1_rank2 ?Ohm_id. + split=> //; apply/eqP; rewrite eqn_leq rank2_exponent_p_p3group // ltnNge. + by apply: contra (leq_trans _) dimRb1; rewrite lognSg ?Ohm_sub. +have ntRb: (R / T) != 1. + by rewrite -cardG_gt1 (card_pgroup pRb) dimRb (ltn_exp2l 0) ?prime_gt1. +have{dimRb} dimR: logn p #|R| = 4. + by rewrite -(Lagrange sTR) lognM ?cardG_gt0 // dimT -card_quotient ?dimRb. +have nsR1R: 'Ohm_1(R) <| R := Ohm_normal 1 R; have [sR1R nR1R] := andP nsR1R. +have pR1: p.-group 'Ohm_1(R) := pgroupS sR1R pR. +have p_odd: odd p by case/even_prime: p_pr p_gt3 => ->. +have{p_odd} oddR: odd #|R| := odd_pgroup_odd p_odd pR. +have{dimR1} dimR1: logn p #|'Ohm_1(R)| = 2. + apply/eqP; rewrite eqn_leq dimR1 -p_rank_abelem; last first. + by rewrite abelem_Ohm1 // (p2group_abelian pR1). + rewrite ltnNge p_rank_Ohm1 -odd_pgroup_rank1_cyclic //. + apply: contra dimRb1 => cycR; have cycRb := quotient_cyclic T cycR. + by rewrite (Ohm1_cyclic_pgroup_prime cycRb pRb ntRb) (pfactorK 1). +have pRs: p.-group (R / 'Ohm_1(R)) by rewrite quotient_pgroup. +have dimRs: logn p #|R / 'Ohm_1(R)| = 2. + by rewrite -divg_normal // logn_div ?cardSg // dimR1 dimR. +have sR'R1: R^`(1) \subset 'Ohm_1(R). + by rewrite der1_min // (p2group_abelian pRs) ?dimRs. +have [|_ phiM] := exponent_odd_nil23 pR oddR. + by rewrite (leq_trans (nil_class_pgroup pR)) // dimR p_gt3. +pose phi := Morphism (phiM sR'R1). +suffices: logn p #|R / 'Ohm_1(R)| <= logn p #|T| by rewrite dimT dimRs. +have ->: 'Ohm_1(R) = 'ker phi. + rewrite -['ker phi]genGid (OhmE 1 pR); congr <<_>>. + by apply/setP=> x; rewrite !inE. +rewrite (card_isog (first_isog phi)) lognSg //=. +apply/subsetP=> _ /morphimP[x _ Rx ->] /=. +apply: coset_idr; first by rewrite groupX ?(subsetP nTR). +by rewrite morphX ?(subsetP nTR) // (exponentP eRb) // mem_quotient. +Qed. + +(* This is B & G, Lemma 4.10. *) +Lemma Ohm1_metacyclic_p2Elem gT p (R : {group gT}) : + metacyclic R -> p.-group R -> odd #|R| -> ~~ cyclic R -> + 'Ohm_1(R)%G \in 'E_p^2(R). +Proof. +case/metacyclicP=> S [cycS nsSR cycRb] pR oddR ncycR. +have [[sSR nSR] [s defS]] := (andP nsSR, cyclicP cycS). +have [T defTb sST sTR] := inv_quotientS nsSR (Ohm_sub 1 (R / S)). +have [pT oddT] := (pgroupS sTR pR, oddSg sTR oddR). +have Ts: s \in T by rewrite -cycle_subG -defS. +have iTs: #|T : <[s]>| = p. + rewrite -defS -card_quotient ?(subset_trans sTR) // -defTb. + rewrite (Ohm1_cyclic_pgroup_prime cycRb (quotient_pgroup _ pR)) // -subG1. + by rewrite quotient_sub1 ?(contra (fun sRS => cyclicS sRS cycS)). +have defR1: 'Ohm_1(R) = 'Ohm_1(T). + apply/eqP; rewrite eqEsubset (OhmS _ sTR) andbT -Ohm_id OhmS //. + rewrite -(quotientSGK _ sST); last by rewrite (subset_trans _ nSR) ?Ohm_sub. + by rewrite -defTb morphim_Ohm. +rewrite (subsetP (pnElemS _ _ sTR)) // (group_inj defR1). +apply: Ohm1_extremal_odd iTs => //; apply: contra ncycR. +by rewrite !(@odd_pgroup_rank1_cyclic _ p) // -p_rank_Ohm1 -defR1 p_rank_Ohm1. +Qed. + +(* This is B & G, Proposition 4.11 (due to Huppert). *) +Proposition p2_Ohm1_metacyclic gT p (R : {group gT}) : + p.-group R -> p > 3 -> logn p #|'Ohm_1(R)| <= 2 -> metacyclic R. +Proof. +move=> pR p_gt3 dimR1; move: {2}_.+1 (ltnSn #|R|) => n. +elim: n => // n IHn in gT R pR dimR1 *; rewrite ltnS => leRn. +have pR1: p.-group 'Ohm_1(R) := pgroupS (Ohm_sub 1 R) pR. +have [cRR | not_cRR] := boolP (abelian R). + have [b defR typeR] := abelian_structure cRR; move: dimR1 defR. + rewrite -(rank_abelian_pgroup pR cRR) -(size_abelian_type cRR) -{}typeR. + case: b => [|a [|b []]] //= _; first by move <-; rewrite big_nil metacyclic1. + by rewrite big_seq1 => <-; rewrite cyclic_metacyclic ?cycle_cyclic. + rewrite big_cons big_seq1; case/dprodP=> _ <- cAB _. + apply/existsP; exists <[a]>%G; rewrite cycle_cyclic /=. + rewrite /normal mulG_subl mulG_subG normG cents_norm //= quotientMidl. + by rewrite quotient_cycle ?cycle_cyclic // -cycle_subG cents_norm. +pose R' := R^`(1); pose e := 'Mho^1(R') != 1. +have nsR'R: R' <| R := der_normal 1 R; have [sR'R nR'R] := andP nsR'R. +have [T EpT]: exists T, T \in 'E_p^1('Mho^e(R') :&: 'Z(R)). + apply/p_rank_geP; rewrite -rank_pgroup; last first. + by rewrite (pgroupS _ pR) //= setIC subIset ?center_sub. + rewrite rank_gt0 (meet_center_nil (pgroup_nil pR)) //. + exact: char_normal_trans (Mho_char e _) nsR'R. + by case ntR'1: e; rewrite //= Mho0 (sameP eqP derG1P). +have [p_gt1 p_pr] := (ltnW (ltnW p_gt3), pnElem_prime EpT). +have p_odd: odd p by case/even_prime: p_pr p_gt3 => ->. +have{p_odd} oddR: odd #|R| := odd_pgroup_odd p_odd pR. +have [sTR'eZ abelT oT] := pnElemPcard EpT; rewrite expn1 in oT. +have{sTR'eZ abelT} [[sTR'e sTZ] [pT cTT eT]] := (subsetIP sTR'eZ, and3P abelT). +have sTR': T \subset R' := subset_trans sTR'e (Mho_sub e _). +have nsTR := sub_center_normal sTZ; have [sTR cRT] := subsetIP sTZ. +have cTR: R \subset 'C(T) by rewrite centsC. +have{n IHn leRn EpT} metaRb: metacyclic (R / T). + have pRb: p.-group (R / T) := quotient_pgroup T pR. + have dimRb: logn p #|'Ohm_1(R / T)| <= 2 by apply: quotient_p2_Ohm1. + by rewrite IHn ?(leq_trans (ltn_quotient _ _)) ?(nt_pnElem EpT). +have{metaRb} [Xb [cycXb nsXbR cycRs]] := metacyclicP metaRb. +have{cycRs} [yb defRb]: exists yb, R / T = Xb <*> <[yb]>. + have [ys defRs] := cyclicP cycRs; have [yb nXyb def_ys] := cosetP ys. + exists yb; rewrite -quotientYK ?cycle_subG ?quotient_cycle // -def_ys -defRs. + by rewrite quotientGK. +have{sTZ} ntXb: Xb :!=: 1. + apply: contraNneq not_cRR => Xb1. + by rewrite (cyclic_factor_abelian sTZ) // defRb Xb1 joing1G cycle_cyclic. +have [TX defTXb sTTX nsTXR] := inv_quotientN nsTR nsXbR. +have{cycXb} [[sTXR nTXR] [xb defXb]] := (andP nsTXR, cyclicP cycXb). +have [[x nTx def_xb] [y nTy def_yb]] := (cosetP xb, cosetP yb). +have{defTXb} defTX: T <*> <[x]> = TX. + rewrite -quotientYK ?cycle_subG ?quotient_cycle // -def_xb -defXb defTXb. + by rewrite quotientGK // (normalS _ sTXR). +have{yb defRb def_yb} defR: TX <*> <[y]> = R. + rewrite -defTX -joingA -quotientYK ?join_subG ?quotientY ?cycle_subG ?nTx //. + by rewrite !quotient_cycle // -def_xb -def_yb -defXb -defRb quotientGK. +have sXYR: <[x]> <*> <[y]> \subset R by rewrite -defR -defTX -joingA joing_subr. +have [Rx Ry]: x \in R /\ y \in R by rewrite -!cycle_subG; exact/joing_subP. +have cTXY := subset_trans sXYR cTR; have [cTX cTY] := joing_subP cTXY. +have [R'1_1 {e sTR'e} | ntR'1] := eqVneq 'Mho^1(R') 1; last first. + have sR'TX: R' \subset TX. + rewrite der1_min // -defR quotientYidl ?cycle_subG ?(subsetP nTXR) //. + by rewrite quotient_abelian // cycle_abelian. + have sTX : T \subset <[x]>. + rewrite (subset_trans (subset_trans sTR'e (MhoS e sR'TX))) // /e ntR'1. + have{defTX} defTX: T \* <[x]> = TX by rewrite cprodEY // centsC. + rewrite -(Mho_cprod 1 defTX) ['Mho^1(_)](trivgP _) ?cprod1g ?Mho_sub //=. + rewrite (MhoE 1 pT) gen_subG; apply/subsetP=> tp; case/imsetP=> t Tt ->{tp}. + by rewrite inE (exponentP eT). + apply/metacyclicP; exists TX; split=> //. + by rewrite -defTX (joing_idPr sTX) cycle_cyclic. + rewrite -defR quotientYidl ?cycle_subG ?(subsetP nTXR) //. + by rewrite quotient_cyclic ?cycle_cyclic. +have{R'1_1} eR': exponent R' %| p. + have <-: 'Ohm_1(R') = R' by apply/eqP; rewrite trivg_Mho ?R'1_1. + rewrite -sub_LdivT (OhmEabelian (pgroupS sR'R pR)) ?subsetIr //. + by rewrite (abelianS (OhmS 1 sR'R)) // (p2group_abelian pR1). +pose r := [~ x, y]; have Rr: r \in R by exact: groupR. +have{defXb ntXb nsXbR} [i def_rb]: exists i, coset T r = (xb ^+ p) ^+ i. + have p_xb: p.-elt xb by rewrite def_xb morph_p_elt ?(mem_p_elt pR). + have pRbb: p.-group (R / T / 'Mho^1(Xb)) by rewrite !quotient_pgroup. + have [_ nXb1R] := andP (char_normal_trans (Mho_char 1 Xb) nsXbR). + apply/cycleP; rewrite -(Mho_p_cycle 1 p_xb) -defXb. + apply: coset_idr; first by rewrite (subsetP nXb1R) ?mem_quotient. + apply/eqP; rewrite !morphR ?(subsetP nXb1R) ?mem_quotient //=; apply/commgP. + red; rewrite -(@centerC _ (R / T / _)) ?mem_quotient // -cycle_subG. + rewrite -quotient_cycle ?(subsetP nXb1R) ?mem_quotient // -def_xb -defXb. + suffices oXbb: #|Xb / 'Mho^1(Xb)| = p. + apply: prime_meetG; first by rewrite oXbb. + rewrite (meet_center_nil (pgroup_nil pRbb)) ?quotient_normal //. + by rewrite -cardG_gt1 oXbb. + rewrite -divg_normal ?Mho_normal //= defXb. + rewrite -(mul_card_Ohm_Mho_abelian 1) ?cycle_abelian ?mulnK ?cardG_gt0 //. + by rewrite (Ohm1_cyclic_pgroup_prime _ p_xb) ?cycle_cyclic //= -defXb. +have{xb def_xb def_rb} [t Tt def_r]: exists2 t, t \in T & r = t * x ^+ (p * i). + apply/rcosetP; rewrite -val_coset ?groupX ?morphX //= -def_xb. + by rewrite expgM -def_rb val_coset ?groupR // rcoset_refl. +have{eR' def_r cTT} defR': R' = <[r]>. + have R'r : r \in R' by exact: mem_commg. + have cxt: t \in 'C[x] by apply/cent1P; exact: (centsP cRT). + have crx: x \in 'C[r] by rewrite cent1C def_r groupM ?groupX ?cent1id. + have def_xy: x ^ y = t * x ^+ (p * i).+1. + by rewrite conjg_mulR -/r def_r expgS !mulgA (cent1P cxt). + have crR : R \subset 'C[r]. + rewrite -defR -defTX !join_subG sub_cent1 (subsetP cTR) //= !cycle_subG. + rewrite crx cent1C (sameP cent1P commgP); apply/conjg_fixP. + rewrite def_r conjMg conjXg conjgE (centsP cRT t) // mulKg conjg_mulR -/r. + by rewrite (expgMn _ (cent1P crx)) (expgM r) (exponentP eR') ?expg1n ?mulg1. + apply/eqP; rewrite eqEsubset cycle_subG R'r andbT. + have nrR : R \subset 'N(<[r]>) by rewrite cents_norm ?cent_cycle. + rewrite der1_min // -defR -defTX -joingA. + rewrite norm_joinEr ?(subset_trans sXYR) ?normal_norm //. + rewrite quotientMl ?(subset_trans sTR) // abelianM quotient_abelian //=. + rewrite quotient_cents //= -der1_joing_cycles ?der_abelian //. + by rewrite -sub_cent1 (subset_trans sXYR). +have [S maxS sR'S] : {S | [max S | S \subset R & cyclic S] & R' \subset S}. + by apply: maxgroup_exists; rewrite sR'R /= -/R' defR' cycle_cyclic. +case/maxgroupP: maxS; case/andP=> sSR cycS maxS. +have nsSR: S <| R := sub_der1_normal sR'S sSR; have [_ nSR] := andP nsSR. +apply/existsP; exists S; rewrite cycS nsSR //=. +suffices uniqRs1 Us: Us \in 'E_p^1(R / S) -> 'Ohm_1(R) / S = Us. + have pRs: p.-group (R / S) := quotient_pgroup S pR. + rewrite abelian_rank1_cyclic ?sub_der1_abelian ?(rank_pgroup pRs) // leqNgt. + apply: contraFN (ltnn 1) => rRs; have [Us EpUs] := p_rank_geP (ltnW rRs). + have [Vs EpVs] := p_rank_geP rRs; have [sVsR abelVs <-] := pnElemP EpVs. + have [_ _ <-] := pnElemP EpUs; apply: lognSg; apply/subsetP=> vs Vvs. + apply: wlog_neg => notUvs; rewrite -cycle_subG -(uniqRs1 _ EpUs). + rewrite (uniqRs1 <[vs]>%G) ?p1ElemE // !inE cycle_subG (subsetP sVsR) //=. + by rewrite -orderE (abelem_order_p abelVs Vvs (group1_contra notUvs)). +case/pnElemPcard; rewrite expn1 => sUsR _ oUs. +have [U defUs sSU sUR] := inv_quotientS nsSR sUsR. +have [cycU | {maxS} ncycU] := boolP (cyclic U). + by rewrite -[p]oUs defUs (maxS U) ?sUR // trivg_quotient cards1 in p_gt1. +have EpU1: 'Ohm_1(U)%G \in 'E_p^2(U). + have [u defS] := cyclicP cycS; rewrite defS cycle_subG in sSU. + rewrite (Ohm1_extremal_odd (pgroupS sUR pR) (oddSg sUR oddR) _ sSU) //. + by rewrite -defS -card_quotient -?defUs // (subset_trans sUR). +have defU1: 'Ohm_1(U) = 'Ohm_1(R). + apply/eqP; rewrite eqEcard OhmS // (card_pnElem EpU1). + by rewrite (card_pgroup pR1) leq_exp2l. +apply/eqP; rewrite eqEcard oUs defUs -{1}defU1 quotientS ?Ohm_sub //. +rewrite dvdn_leq ?cardG_gt0 //; case/pgroup_pdiv: (quotient_pgroup S pR1) => //. +rewrite -subG1 quotient_sub1 ?(subset_trans (Ohm_sub 1 R) nSR) //. +apply: contraL (cycS) => sR1S; rewrite abelian_rank1_cyclic ?cyclic_abelian //. +rewrite -ltnNge (rank_pgroup (pgroupS sSR pR)); apply/p_rank_geP. +by exists 'Ohm_1(U)%G; rewrite -(setIidPr sSU) pnElemI inE EpU1 inE /= defU1. +Qed. + +(* This is B & G, Theorem 4.12 (also due to Huppert), for internal action. *) +Theorem coprime_metacyclic_cent_sdprod gT p (R A : {group gT}) : + p.-group R -> odd #|R| -> metacyclic R -> p^'.-group A -> A \subset 'N(R) -> + let T := [~: R, A] in let C := 'C_R(A) in + [/\ (*a*) abelian T, + (*b*) T ><| C = R + & (*c*) ~~ abelian R -> T != 1 -> + [/\ C != 1, cyclic T, cyclic C & R^`(1) \subset T]]. +Proof. +move=> pR oddR metaR p'A nRA T C. +suffices{C T} cTT: abelian [~: R, A]. + have sTR: T \subset R by rewrite commg_subl. + have nTR: R \subset 'N(T) by rewrite commg_norml. + have coRA: coprime #|R| #|A| := pnat_coprime pR p'A. + have solR: solvable R := pgroup_sol pR. + have defR: T * C = R by rewrite coprime_cent_prod. + have sCR: C \subset R := subsetIl _ _; have nTC := subset_trans sCR nTR. + have tiTC: T :&: C = 1. + have defTA: [~: T, A] = T by rewrite coprime_commGid. + have coTA: coprime #|T| #|A| := coprimeSg sTR coRA. + by rewrite setIA (setIidPl sTR) -defTA coprime_abel_cent_TI ?commg_normr. + split=> // [|not_cRR ntT]; first by rewrite sdprodE. + have ntC: C != 1 by apply: contraNneq not_cRR => C1; rewrite -defR C1 mulg1. + suffices [cycT cycC]: cyclic T /\ cyclic C. + split=> //; rewrite der1_min //= -/T -defR quotientMidl. + by rewrite cyclic_abelian ?quotient_cyclic. + have [pT pC]: p.-group T /\ p.-group C by rewrite !(pgroupS _ pR). + apply/andP; rewrite (odd_pgroup_rank1_cyclic pC (oddSg sCR oddR)). + rewrite abelian_rank1_cyclic // -rank_pgroup //. + rewrite -(geq_leqif (leqif_add (leqif_geq _) (leqif_geq _))) ?rank_gt0 //. + have le_rTC_dimTC1: 'r(T) + 'r(C) <= logn p #|'Ohm_1(T) * 'Ohm_1(C)|. + rewrite (rank_pgroup pC) -p_rank_Ohm1 (rank_abelian_pgroup pT cTT). + rewrite TI_cardMg; last by apply/trivgP; rewrite -tiTC setISS ?Ohm_sub. + by rewrite lognM ?cardG_gt0 // leq_add2l p_rank_le_logn. + apply: leq_trans le_rTC_dimTC1 _; rewrite add1n. + have ncycR: ~~ cyclic R by apply: contra not_cRR; apply: cyclic_abelian. + have: 'Ohm_1(R)%G \in 'E_p^2(R) by apply: Ohm1_metacyclic_p2Elem. + have nT1C1: 'Ohm_1(C) \subset 'N('Ohm_1(T)). + by rewrite (subset_trans (Ohm_sub 1 _)) ?(char_norm_trans (Ohm_char 1 _)). + by case/pnElemP=> _ _ <-; rewrite -norm_joinEr ?lognSg // join_subG !OhmS. +without loss defR: R pR oddR metaR nRA / [~: R, A] = R. + set T := [~: R, A] => IH; have sTR: T \subset R by rewrite commg_subl. + have defTA: [~: T, A] = T. + by rewrite coprime_commGid ?(pgroup_sol pR) ?(pnat_coprime pR). + rewrite -defTA IH ?(pgroupS sTR) ?(oddSg sTR) ?(metacyclicS sTR) //. + exact: commg_normr. +rewrite defR; apply: wlog_neg => not_cRR. +have ncycR: ~~ cyclic R := contra (@cyclic_abelian _ R) not_cRR. +pose cycR_nA S := [&& cyclic S, S \subset R & A \subset 'N(S)]. +have [S maxS sR'S] : {S | [max S | cycR_nA S] & R^`(1) \subset S}. + apply: maxgroup_exists; rewrite {}/cycR_nA der_sub /=. + rewrite (char_norm_trans (der_char 1 _)) // andbT. + have [K [cycK nsKR cycKR]] := metacyclicP metaR. + by rewrite (cyclicS _ cycK) // der1_min ?normal_norm // cyclic_abelian. +case/maxgroupP: maxS; case/and3P=> cycS sSR nSA maxS. +have ntS: S :!=: 1 by rewrite (subG1_contra sR'S) // (sameP eqP derG1P). +have nSR: R \subset 'N(S) := sub_der1_norm sR'S sSR. +have nsSR: S <| R by exact/andP. +have sSZ: S \subset 'Z(R). + have sR_NS': R \subset 'N(S)^`(1) by rewrite -{1}defR commgSS. + rewrite subsetI sSR centsC (subset_trans sR_NS') // der1_min ?cent_norm //=. + rewrite -ker_conj_aut (isog_abelian (first_isog _)). + by rewrite (abelianS (Aut_conj_aut _ _)) ?Aut_cyclic_abelian. +have cRbRb: abelian (R / S) by rewrite sub_der1_abelian. +have pRb: p.-group (R / S) := quotient_pgroup S pR. +pose R1 := 'Ohm_1(R); pose Rb1 := 'Ohm_1(R / S). +have [Xb]: exists2 Xb, R1 / S \x gval Xb = Rb1 & A / S \subset 'N(Xb). + have MaschkeRb1 := Maschke_abelem (Ohm1_abelem pRb cRbRb). + pose normOhm1 := (char_norm_trans (Ohm_char 1 _), quotient_norms S). + by apply: MaschkeRb1; rewrite ?quotient_pgroup ?morphim_Ohm ?normOhm1. +case/dprodP=> _ defRb1 _ tiR1bX nXbA. +have sXbR: Xb \subset R / S. + by apply: subset_trans (Ohm_sub 1 _); rewrite -defRb1 mulG_subr. +have{sXbR} [X defXb sSX sXR] := inv_quotientS nsSR sXbR. +have{nXbA nsSR} nXA: A \subset 'N(X). + rewrite (subset_trans (mulG_subr S A)) // -quotientK //. + by rewrite -(quotientGK (normalS sSX sXR nsSR)) -defXb morphpre_norms. +have{tiR1bX} cycX: cyclic X. + have sX1_XR1: 'Ohm_1(X) \subset X :&: R1 by rewrite subsetI Ohm_sub OhmS. + have cyc_sR := odd_pgroup_rank1_cyclic (pgroupS _ pR) (oddSg _ oddR). + have:= cycS; rewrite !{}cyc_sR //; apply: leq_trans. + rewrite -p_rank_Ohm1 p_rankS // (subset_trans sX1_XR1) //. + rewrite -quotient_sub1 ?subIset ?(subset_trans sXR) //. + by rewrite quotientGI // setIC -defXb tiR1bX. +rewrite (cyclic_factor_abelian sSZ) // abelian_rank1_cyclic //. +rewrite (rank_abelian_pgroup pRb cRbRb) -defRb1 defXb. +rewrite (maxS X) ?trivg_quotient ?mulg1 //; last exact/and3P. +have EpR1: 'Ohm_1(R)%G \in 'E_p^2(R) by exact: Ohm1_metacyclic_p2Elem. +have [sR1R _ dimR1] := pnElemP EpR1; have pR1 := pgroupS sR1R pR. +rewrite -(card_isog (second_isog _)) ?(subset_trans sR1R) // -ltnS -dimR1. +by rewrite (ltn_log_quotient pR1) ?subsetIr //= meet_Ohm1 // (setIidPl sSR). +Qed. + +(* This covers B & G, Lemmas 4.13 and 4.14. *) +Lemma pi_Aut_rank2_pgroup gT p q (R : {group gT}) : + p.-group R -> odd #|R| -> 'r(R) <= 2 -> q \in \pi(Aut R) -> q != p -> + [/\ q %| (p ^ 2).-1, q < p & q %| p.+1./2 \/ q %| p.-1./2]. +Proof. +move=> pR oddR rR q_Ar p'q; rewrite /= in q_Ar. +have [R1 | ntR] := eqsVneq R 1; first by rewrite R1 Aut1 cards1 in q_Ar. +have{ntR} [p_pr p_dv_R _] := pgroup_pdiv pR ntR. +have{oddR p_dv_R} [p_odd p_gt1] := (dvdn_odd p_dv_R oddR, prime_gt1 p_pr). +have{q_Ar} [q_pr q_dv_Ar]: prime q /\ q %| #|Aut R|. + by move: q_Ar; rewrite mem_primes; case/and3P. +suffices q_dv_p2: q %| (p ^ 2).-1. + have q_dv_p1: q %| p.+1./2 \/ q %| p.-1./2. + apply/orP; have:= q_dv_p2; rewrite -subn1 (subn_sqr p 1). + rewrite -[p]odd_double_half p_odd /= !doubleK addKn addn1 -doubleS -!mul2n. + rewrite mulnC !Euclid_dvdM // dvdn_prime2 // -orbA; case: eqP => // -> _. + by rewrite -Euclid_dvdM // /dvdn modn2 mulnC odd_mul andbN. + have p_gt2: p > 2 by rewrite ltn_neqAle; case: eqP p_odd => // <-. + have p1_ltp: p.+1./2 < p. + by rewrite -divn2 ltn_divLR // muln2 -addnn -addn2 leq_add2l. + split=> //; apply: leq_ltn_trans p1_ltp. + move/orP: q_dv_p1; rewrite -(subnKC p_gt2) leqNgt. + by apply: contraL => lt_p1q; rewrite negb_or !gtnNdvd // ltnW. +wlog{q_dv_Ar} [a oa nRa]: gT R pR rR / {a | #[a] = q & a \in 'N(R) :\: 'C(R)}. + have [a Ar_a oa] := Cauchy q_pr q_dv_Ar. + rewrite -(injm_rank (injm_sdpair1 [Aut R])) // in rR. + move=> IH; apply: IH rR _; rewrite ?morphim_pgroup ?morphim_odd //. + exists (sdpair2 [Aut R] a); rewrite ?(order_injm (injm_sdpair2 _)) //. + rewrite inE (subsetP (im_sdpair_norm _)) ?mem_morphim //= andbT. + apply: contraL (prime_gt1 q_pr) => cRa; rewrite -oa order_gt1 negbK. + apply/eqP; apply: (eq_Aut Ar_a (group1 _)) => x Rx. + by rewrite perm1 [a x](@astab_act _ _ _ [Aut R] R) ?astabEsd ?mem_morphpre. +move: {2}_.+1 (ltnSn #|R|) => n. +elim: n => // n IHn in gT a R pR rR nRa oa *; rewrite ltnS => leRn. +case recR: [exists (S : {group gT} | S \proper R), a \in 'N(S) :\: 'C(S)]. + have [S ltSR nSa] := exists_inP recR; rewrite properEcard in ltSR. + have{ltSR} [sSR ltSR] := andP ltSR; have rS := leq_trans (rankS sSR) rR. + by apply: IHn nSa oa _; rewrite ?(pgroupS sSR) ?(leq_trans ltSR). +do [rewrite inE -!cycle_subG orderE; set A := <[a]>] in nRa oa. +have{nRa oa} [[not_cRA nRA] oA] := (andP nRa, oa). +have coRA : coprime #|R| #|A| by rewrite oA (pnat_coprime pR) ?pnatE. +have{recR} IH: forall S, gval S \proper R -> A \subset 'N(S) -> A \subset 'C(S). + move=> S ltSR; rewrite !cycle_subG => nSa; apply: contraFT recR => not_cSa. + by apply/exists_inP; exists S; rewrite // inE not_cSa nSa. +have defR1: 'Ohm_1(R) = R. +apply: contraNeq not_cRA; rewrite eqEproper Ohm_sub negbK => ltR1R. + rewrite (coprime_odd_faithful_Ohm1 pR) ?IH ?(odd_pgroup_odd p_odd) //. + by rewrite (char_norm_trans (Ohm_char 1 R)). +have defRA: [~: R, A] = R. + apply: contraNeq not_cRA; rewrite eqEproper commg_subl nRA negbK => ltRAR. + rewrite centsC; apply/setIidPl. + rewrite -{2}(coprime_cent_prod nRA) ?(pgroup_sol pR) //. + by rewrite mulSGid // subsetI commg_subl nRA centsC IH ?commg_normr. +have [cRR | not_cRR] := boolP (abelian R). + rewrite -subn1 (subn_sqr p 1) Euclid_dvdM //. + have abelR: p.-abelem R by rewrite -defR1 Ohm1_abelem. + have ntR: R :!=: 1 by apply: contraNneq not_cRA => ->; apply: cents1. + pose rAR := reprGLm (abelem_repr abelR ntR nRA). + have:= cardSg (subsetT (rAR @* A)); rewrite card_GL ?card_Fp //. + rewrite card_injm ?ker_reprGLm ?rker_abelem ?prime_TIg ?oA // unlock. + rewrite Gauss_dvdr; last by rewrite coprime_expr ?prime_coprime ?dvdn_prime2. + move: rR; rewrite -ltnS -[_ < _](mem_iota 0) !inE eqn0Ngt rank_gt0 ntR. + rewrite (dim_abelemE abelR ntR) (rank_abelem abelR). + do [case/pred2P=> ->; rewrite /= muln1] => [-> // | ]. + by rewrite (subn_sqr p 1) mulnA !Euclid_dvdM ?orbb. +have [[defPhi defR'] _]: special R /\ 'C_R(A) = 'Z(R). + apply: (abelian_charsimple_special pR) => //. + apply/bigcupsP=> S; case/andP=> charS cSS. + rewrite centsC IH ?(char_norm_trans charS) // properEneq char_sub // andbT. + by apply: contraNneq not_cRR => <-. +have ntZ: 'Z(R) != 1 by rewrite -defR' (sameP eqP derG1P). +have ltRbR: #|R / 'Z(R)| < #|R| by rewrite ltn_quotient ?center_sub. +have pRb: p.-group (R / 'Z(R)) by apply: quotient_pgroup. +have nAZ: A \subset 'N('Z(R)) by rewrite (char_norm_trans (center_char R)). +have defAb: A / 'Z(R) = <[coset _ a]> by rewrite quotient_cycle -?cycle_subG. +have oab: #[coset 'Z(R) a] = q. + rewrite orderE -defAb -(card_isog (quotient_isog _ _)) //. + by rewrite coprime_TIg ?(coprimeSg (center_sub R)). +have rRb: 'r(R / 'Z(R)) <= 2. + rewrite (rank_pgroup pRb) (leq_trans (p_rank_le_logn _ _)) // -ltnS. + apply: leq_trans (rank2_exponent_p_p3group pR rR _). + by rewrite -(ltn_exp2l _ _ p_gt1) -!card_pgroup. + by rewrite -defR1 (exponent_Ohm1_class2 p_odd) // nil_class2 defR'. +apply: IHn oab (leq_trans ltRbR leRn) => //. +rewrite inE -!cycle_subG -defAb quotient_norms ?andbT //. +apply: contra not_cRA => cRAb; rewrite (coprime_cent_Phi pR coRA) // defPhi. +by rewrite commGC -quotient_cents2 ?gFnorm. +Qed. + +(* B & G, Lemma 4.15 is covered by maximal/critical_extraspecial. *) + +(* This is B & G, Theorem 4.16 (due to Blackburn). *) +Theorem rank2_coprime_comm_cprod gT p (R A : {group gT}) : + p.-group R -> odd #|R| -> R :!=: 1 -> 'r(R) <= 2 -> + [~: R, A] = R -> p^'.-group A -> odd #|A| -> + [/\ p > 3 + & [\/ abelian R + | exists2 S : {group gT}, + [/\ ~~ abelian S, logn p #|S| = 3 & exponent S %| p] + & exists C : {group gT}, + [/\ S \* C = R, cyclic C & 'Ohm_1(C) = S^`(1)]]]. +Proof. +move=> pR oddR ntR rR defRA p'A oddA. +have [p_pr _ _] := pgroup_pdiv pR ntR; have p_gt1 := prime_gt1 p_pr. +have nilR: nilpotent R := pgroup_nil pR. +have nRA: A \subset 'N(R) by rewrite -commg_subl defRA. +have p_gt3: p > 3; last split => //. + have [Ab1 | [q q_pr q_dv_Ab]] := trivgVpdiv (A / 'C_A(R)). + case/eqP: ntR; rewrite -defRA commGC; apply/commG1P. + by rewrite -subsetIidl -quotient_sub1 ?Ab1 ?normsI ?norms_cent ?normG. + have odd_q := dvdn_odd q_dv_Ab (quotient_odd _ oddA). + have p'q := pgroupP (quotient_pgroup _ p'A) q q_pr q_dv_Ab. + have q_gt1: q > 1 := prime_gt1 q_pr. + have q_gt2: q > 2 by rewrite ltn_neqAle; case: eqP odd_q => // <-. + apply: leq_ltn_trans q_gt2 _. + rewrite /= -ker_conj_aut (card_isog (first_isog_loc _ _)) // in q_dv_Ab. + have q_dv_A := dvdn_trans q_dv_Ab (cardSg (Aut_conj_aut _ _)). + by case/(pi_Aut_rank2_pgroup pR): (pgroupP (pgroup_pi _) q q_pr q_dv_A). +pose S := 'Ohm_1(R); pose S' := S^`(1); pose C := 'C_R(S). +have pS: p.-group S := pgroupS (Ohm_sub 1 _) pR. +have nsSR: S <| R := Ohm_normal 1 R. +have nsS'R: S' <| R := char_normal_trans (der_char 1 _) nsSR. +have [sSR nSR] := andP nsSR; have [_ nS'R] := andP nsS'R. +have [Sle2 | Sgt2] := leqP (logn p #|S|) 2. + have metaR: metacyclic R := p2_Ohm1_metacyclic pR p_gt3 Sle2. + have [cRR _ _] := coprime_metacyclic_cent_sdprod pR oddR metaR p'A nRA. + by left; rewrite -defRA. +have{p_gt3} eS: exponent S %| p by apply: exponent_Ohm1_rank2. +have{rR} rS: 'r(S) <= 2 by rewrite rank_Ohm1. +have{Sgt2} dimS: logn p #|S| = 3. + by apply/eqP; rewrite eqn_leq rank2_exponent_p_p3group. +have{rS} not_cSS: ~~ abelian S. + by apply: contraL rS => cSS; rewrite -ltnNge -dimS -rank_abelem ?abelem_Ohm1. +have esS: extraspecial S by apply: (p3group_extraspecial pS); rewrite ?dimS. +have defS': S' = 'Z(S) by case: esS; case. +have oS': #|S'| = p by rewrite defS' (card_center_extraspecial pS esS). +have dimS': logn p #|S'| = 1%N by rewrite oS' (pfactorK 1). +have nsCR: C <| R := normalGI nSR (cent_normal _); have [sCR nCR] := andP nsCR. +have [pC oddC]: p.-group C * odd #|C| := (pgroupS sCR pR, oddSg sCR oddR). +have defC1: 'Ohm_1(C) = S'. + apply/eqP; rewrite eqEsubset defS' subsetI OhmS ?(OhmE 1 pC) //= -/C. + by rewrite gen_subG setIAC subsetIr sub_gen ?setSI // subsetI sSR sub_LdivT. +have{pC oddC} cycC: cyclic C. + rewrite (odd_pgroup_rank1_cyclic pC) //. + by rewrite -p_rank_Ohm1 defC1 -dimS' p_rank_le_logn. +pose T := [~: S, R]; have nsTR: T <| R by rewrite /normal commg_normr comm_subG. +have [sTR nTR] := andP nsTR; have pT: p.-group T := pgroupS sTR pR. +have [sTS' | not_sTS' {esS}] := boolP (T \subset S'). + right; exists [group of S] => //; exists [group of C]. + by rewrite (critical_extraspecial pR sSR esS sTS'). +have ltTS: T \proper S by rewrite (nil_comm_properl nilR) ?Ohm1_eq1 ?subsetIidl. +have sTS: T \subset S := proper_sub ltTS. +have [sS'T ltS'T]: S' \subset T /\ S' \proper T by rewrite /proper commgS. +have{ltS'T ltTS} dimT: logn p #|T| = 2. + by apply/eqP; rewrite eqn_leq -ltnS -dimS -dimS' !properG_ltn_log. +have{eS} eT: exponent T %| p := dvdn_trans (exponentS sTS) eS. +have cTT: abelian T by rewrite (p2group_abelian pT) ?dimT. +have abelT: p.-abelem T by apply/and3P. +pose B := 'C_R(T); have sTB: T \subset B by rewrite subsetI sTR. +have nsBR: B <| R := normalGI nTR (cent_normal _); have [sBR nBR] := andP nsBR. +have not_sSB: ~~ (S \subset B). + by rewrite defS' !subsetI sTS sSR centsC in not_sTS' *. +have maxB: maximal B R. + rewrite p_index_maximal // (_ : #|R : B| = p) //; apply/prime_nt_dvdP=> //. + by apply: contra not_sSB; rewrite indexg_eq1; apply: subset_trans. + rewrite -(part_pnat_id (pnat_dvd (dvdn_indexg _ _) pR)) p_part. + by rewrite (@dvdn_exp2l _ _ 1) // logn_quotient_cent_abelem ?dimT //. +have{maxB nsBR} defR: B * S = R := mulg_normal_maximal nsBR maxB sSR not_sSB. +have cBbBb: abelian (B / C). + rewrite sub_der1_abelian // subsetI comm_subG ?subsetIl //=; apply/commG1P. + suff cB_SB: [~: S, B, B] = 1 by rewrite three_subgroup // [[~: _, S]]commGC. + by apply/commG1P; rewrite centsC subIset // centS ?orbT // commgS. +have{cBbBb} abelBb: p.-abelem (B / C). + apply/abelemP=> //; split=> // Cg; case/morphimP=> x Nx Bx /= ->. + have [Rx cTx] := setIP Bx; rewrite -morphX //= coset_id // inE groupX //=. + apply/centP=> y Sy; symmetry; have Tyx : [~ y, x] \in T by apply: mem_commg. + by apply/commgP; rewrite commgX ?(exponentP eT) //; apply: (centP cTx). +have nsCB: C <| B by rewrite (normalS _ _ nsCR) ?setIS ?subsetIl // centS. +have p'Ab: p^'.-group (A / C) by apply: quotient_pgroup. +have sTbB: T / C \subset B / C by rewrite quotientS. +have nSA: A \subset 'N(S) := char_norm_trans (Ohm_char 1 _) nRA. +have nTA: A \subset 'N(T) := normsR nSA nRA. +have nTbA: A / C \subset 'N(T / C) := quotient_norms _ nTA. +have nBbA: A / C \subset 'N(B / C). + by rewrite quotient_norms ?normsI ?norms_cent. +have{p'Ab sTbB nBbA abelBb nTbA} + [Xb defBb nXbA] := Maschke_abelem abelBb p'Ab sTbB nBbA nTbA. +have{defBb} [_] := dprodP defBb; rewrite /= -/T -/B => defBb _ tiTbX. +have sXbB: Xb \subset B / C by rewrite -defBb mulG_subr. +have{sXbB} [X] := inv_quotientS nsCB sXbB; rewrite /= -/C -/B => defXb sCX sXB. +have sXR: X \subset R := subset_trans sXB sBR; have pX := pgroupS sXR pR. +have nsCX: C <| X := normalS sCX sXR nsCR. +have{tiTbX} ziTX: T :&: X \subset C. + rewrite -quotient_sub1 ?subIset ?(subset_trans sTR) ?normal_norm //= -/C. + by rewrite quotientIG -?defXb ?tiTbX. +have{nXbA} nXA: A \subset 'N(X). + have nCA: A \subset 'N(C) by rewrite normsI ?norms_cent. + by rewrite -(quotientSGK nCA) ?normsG // quotient_normG -?defXb. +have{abelT} defB1: 'Ohm_1(B) = T. + apply/eqP; rewrite eq_sym eqEcard -{1}[T](Ohm1_id abelT) OhmS //. + have pB1: p.-group 'Ohm_1(B) := pgroupS (subset_trans (Ohm_sub 1 _) sBR) pR. + rewrite (card_pgroup pT) (card_pgroup pB1) leq_exp2l //= -/T -/B. + rewrite dimT -ltnS -dimS properG_ltn_log // properEneq OhmS ?subsetIl //= -/S. + by case: eqP not_sSB => // <-; rewrite Ohm_sub. +have{ziTX defB1} cycX: cyclic X; last have [x defX]:= cyclicP cycX. + rewrite (odd_pgroup_rank1_cyclic pX (oddSg sXR oddR)) -p_rank_Ohm1. + have:= cycC; rewrite abelian_rank1_cyclic ?cyclic_abelian //= -/C. + apply: leq_trans (leq_trans (p_rank_le_rank p _) (rankS _)). + by apply: subset_trans ziTX; rewrite subsetI Ohm_sub -defB1 OhmS. +have{Xb defXb defBb nsCX} mulSX: S * X = R. + have nCT: T \subset 'N(C) := subset_trans sTR nCR. + rewrite -defR -(normC (subset_trans sSR nBR)) -[B](quotientGK nsCB) -defBb. + rewrite cosetpreM quotientK // defXb quotientGK // -(normC nCT). + by rewrite -mulgA (mulSGid sCX) mulgA (mulGSid sTS). +have{mulSX} not_sXS_S': ~~ ([~: X, S] \subset S'). + apply: contra not_sTS' => sXS_S'; rewrite /T -mulSX. + by rewrite commGC commMG ?(subset_trans sXR) // mul_subG. +have [oSb oTb] : #|S / T| = p /\ #|T / S'| = p. + rewrite (card_pgroup (quotient_pgroup _ pS)) -divg_normal ?(normalS _ sSR) //. + rewrite (card_pgroup (quotient_pgroup _ pT)) -divg_normal ?(normalS _ sTR) //. + by rewrite !logn_div ?cardSg // dimS dimT dimS'. +have [Ty defSb]: exists Ty, S / T = <[Ty]>. + by apply/cyclicP; rewrite prime_cyclic ?oSb. +have SbTy: Ty \in S / T by rewrite defSb cycle_id. +have{SbTy} [y nTy Sy defTy] := morphimP SbTy. +have [S'z defTb]: exists S'z, T / S' = <[S'z]>. + apply/cyclicP; rewrite prime_cyclic ?oTb //. +have TbS'z: S'z \in T / S' by rewrite defTb cycle_id. +have{TbS'z} [z nS'z Tz defS'z] := morphimP TbS'z. +have [Ta AbTa not_cSbTa]: exists2 Ta, Ta \in A / T & Ta \notin 'C(S / T). + apply: subsetPn; rewrite quotient_cents2 ?commg_norml //= -/T commGC. + apply: contra not_sSB => sSA_T; rewrite (subset_trans sSR) // -defRA -defR. + rewrite -(normC (subset_trans sSR nBR)) commMG /= -/S -/B; last first. + by rewrite cents_norm ?subIset ?centS ?orbT. + by rewrite mul_subG ?commg_subl ?normsI ?norms_cent // (subset_trans sSA_T). +have [a nTa Aa defTa] := morphimP AbTa. +have nS'a: a \in 'N(S') := subsetP (char_norm_trans (der_char 1 _) nSA) a Aa. +have [i xa]: exists i, x ^ a = x ^+ i. + by apply/cycleP; rewrite -cycle_subG cycleJ /= -defX (normsP nXA). +have [j Tya]: exists j, Ty ^ Ta = Ty ^+ j. + apply/cycleP; rewrite -cycle_subG cycleJ /= -defSb. + by rewrite (normsP (quotient_norms _ nSA)). +suffices {oSb oddA not_cSbTa} j2_1: j ^ 2 == 1 %[mod p]. + have Tya2: Ty ^ coset T (a ^+ 2) = Ty ^+ (j ^ 2). + by rewrite morphX // conjgM -defTa Tya conjXg Tya expgM. + have coA2: coprime #|A| 2 by rewrite coprime_sym prime_coprime // dvdn2 oddA. + case/negP: not_cSbTa; rewrite defTa -(expgK coA2 Aa) morphX groupX //=. + rewrite defSb cent_cycle inE conjg_set1 Tya2 sub1set inE. + by rewrite (eq_expg_mod_order _ _ 1) orderE -defSb oSb. +have {Tya Ta defTa AbTa} [u Tu yj]: exists2 u, u \in T & y ^+ j = u * y ^ a. + apply: rcosetP; apply/rcoset_kercosetP; rewrite ?groupX ?groupJ //. + by rewrite morphX ?morphJ -?defTy // -defTa. +have{Ty defTy defSb} defS: T * <[y]> = S. + rewrite -quotientK ?cycle_subG ?quotient_cycle // -defTy -defSb /= -/T. + by rewrite quotientGK // /normal sTS /= commg_norml. +have{nTA} [k S'zk]: exists k, S'z ^ coset S' a = S'z ^+ k. + apply/cycleP; rewrite -cycle_subG cycleJ /= -defTb. + by rewrite (normsP (quotient_norms _ nTA)) ?mem_quotient. +have S'yz: [~ y, z] \in S' by rewrite mem_commg // (subsetP sTS). +have [v Zv zk]: exists2 v, v \in 'Z(S) & z ^+ k = v * z ^ a. + apply: rcosetP; rewrite -defS'. + by apply/rcoset_kercosetP; rewrite ?groupX ?groupJ ?morphX ?morphJ -?defS'z. +have defT: S' * <[z]> = T. + rewrite -quotientK ?cycle_subG ?quotient_cycle // -defS'z -defTb /= -/S'. + by rewrite quotientGK // (normalS _ sTR) // proper_sub. +have nt_yz: [~ y, z] != 1. + apply: contra not_cSS; rewrite (sameP commgP cent1P) => cyz. + rewrite -defS abelianM cTT cycle_abelian /= -/T -defT centM /= -/S' defS'. + by rewrite cent_cycle subsetI centsC subIset ?centS ?cycle_subG ?orbT. +have sS'X1: S' \subset 'Ohm_1(X) by rewrite -defC1 OhmS. +have i_neq0: i != 0 %[mod p]. + have: 'Ohm_1(X) != 1 by rewrite (subG1_contra sS'X1) //= -cardG_gt1 oS'. + rewrite defX in pX *; rewrite (Ohm_p_cycle 1 pX) subn1 trivg_card1 -orderE. + rewrite -(orderJ _ a) conjXg xa order_eq1 -expgM -order_dvdn mod0n. + apply: contra; case/dvdnP=> m ->; rewrite -mulnA -expnS dvdn_mull //. + by rewrite {1}[#[x]](card_pgroup pX) dvdn_exp2l ?leqSpred. +have Txy: [~ x, y] \in T by rewrite [T]commGC mem_commg // -cycle_subG -defX. +have [Rx Ry]: x \in R /\ y \in R by rewrite -cycle_subG -defX (subsetP sSR). +have [nS'x nS'y] := (subsetP nS'R x Rx, subsetP nS'R y Ry). +have{not_sXS_S'} not_S'xy: [~ x, y] \notin S'. + apply: contra not_sXS_S' => S'xy. + rewrite -quotient_cents2 ?(subset_trans _ nS'R) //= -/S'. + rewrite -defS quotientMl ?(subset_trans _ nS'R) // centM /= -/S' -/T. + rewrite subsetI quotient_cents; last by rewrite (subset_trans sXB) ?subsetIr. + rewrite defX !quotient_cycle // cent_cycle cycle_subG /= -/S'. + by rewrite (sameP cent1P commgP) -morphR /= ?coset_id. +have jk_eq_i: j * k = i %[mod p]. + have Zyz: [~ y, z] \in 'Z(S) by rewrite -defS'. + have Sz: z \in S := subsetP sTS z Tz. + have yz_p: [~ y, z] ^+ p == 1 by rewrite -order_dvdn -oS' order_dvdG. + have <-: #[[~ y, z]] = p by apply: nt_prime_order => //; apply: eqP. + apply: eqP; rewrite -eq_expg_mod_order -commXXg; try exact: centerC Zyz. + have cyv: [~ y ^+ j, v] = 1 by apply/eqP/commgP/(centerC (groupX j Sy) Zv). + have cuz: [~ u, z ^ a] = 1. + by apply/eqP/commgP/(centsP cTT); rewrite ?memJ_norm. + rewrite zk commgMJ cyv yj commMgJ cuz !conj1g mulg1 mul1g -conjRg. + suffices [m ->]: exists m, [~ y, z] = x ^+ m by rewrite conjXg xa expgAC. + by apply/cycleP; rewrite -defX (subsetP (Ohm_sub 1 X)) ?(subsetP sS'X1). +have ij_eq_k: i * j = k %[mod p]. + have <-: #[coset S' [~ x, y]] = p. + apply: nt_prime_order => //. + by apply: eqP; rewrite -order_dvdn -oTb order_dvdG 1?mem_quotient. + by apply: contraNneq not_S'xy; apply: coset_idr; rewrite groupR. + have sTbZ: T / S' \subset 'Z(R / S'). + rewrite prime_meetG ?oTb // (meet_center_nil (quotient_nil _ nilR)) //. + by rewrite quotient_normal //; apply/andP. + by rewrite -cardG_gt1 oTb. + have ZRxyb: [~ coset S' x, coset S' y] \in 'Z(R / S'). + by rewrite -morphR // (subsetP sTbZ) ?mem_quotient. + apply: eqP; rewrite -eq_expg_mod_order {1}morphR //. + rewrite -commXXg; try by apply: centerC ZRxyb; apply: mem_quotient. + have [Ru nRa] := (subsetP sTR u Tu, subsetP nRA a Aa). + rewrite -2?morphX // yj morphM ?(subsetP nS'R) ?memJ_norm //. + have cxu_b: [~ coset S' (x ^+ i), coset S' u] = 1. + apply: eqP; apply/commgP. + by apply: centerC (subsetP sTbZ _ _); rewrite mem_quotient ?groupX. + rewrite commgMJ cxu_b conj1g mulg1 -xa !morphJ // -conjRg -morphR //=. + have: coset S' [~ x, y] \in <[S'z]> by rewrite -defTb mem_quotient. + by case/cycleP=> m ->; rewrite conjXg S'zk expgAC. +have j2_gt0: j ^ 2 > 0. + rewrite expn_gt0 orbF lt0n; apply: contraNneq i_neq0 => j0. + by rewrite -jk_eq_i j0. +have{i_neq0} co_p_i: coprime p i by rewrite mod0n prime_coprime in i_neq0 *. +rewrite eqn_mod_dvd // -(Gauss_dvdr _ co_p_i) mulnBr -eqn_mod_dvd ?leq_mul //. +by rewrite muln1 mulnCA -modnMmr ij_eq_k modnMmr jk_eq_i. +Qed. + +(* This is B & G, Theorem 4.17. *) +Theorem der1_Aut_rank2_pgroup gT p (R : {group gT}) (A : {group {perm gT}}) : + p.-group R -> odd #|R| -> 'r(R) <= 2 -> + A \subset Aut R -> solvable A -> odd #|A| -> + p.-group A^`(1). +Proof. +move=> pR oddR rR AutA solA oddA. +without loss ntR: / R :!=: 1. + case: eqP AutA => [-> | ntR _ -> //]; rewrite Aut1. + by move/trivgP=> ->; rewrite derg1 commG1 pgroup1. +have [p_pr _ _] := pgroup_pdiv pR ntR; have p_gt1 := prime_gt1 p_pr. +have{ntR oddR} [H [charH _] _ eH pCH] := critical_odd pR oddR ntR. +have sHR := char_sub charH; have pH := pgroupS sHR pR. +have{rR} rH: 'r(H) <= 2 := leq_trans (rankS (char_sub charH)) rR. +have dimH: logn p #|H| <= 3 by rewrite rank2_exponent_p_p3group ?eH. +have{eH} ntH: H :!=: 1 by rewrite trivg_exponent eH gtnNdvd. +have charP := Phi_char H; have [sPH nPH] := andP (Phi_normal H : 'Phi(H) <| H). +have nHA: {acts A, on group H | [Aut R]} := gacts_char _ AutA charH. +pose B := 'C(H | <[nHA]>); pose V := H / 'Phi(H); pose C := 'C(V | <[nHA]> / _). +have{pCH} pB: p.-group B. + by rewrite (pgroupS _ pCH) //= astab_actby setIid subsetIr. +have s_p'C_B X: gval X \subset C -> p^'.-group X -> X \subset B. + move=> sXC p'X; have [sDX _] := subsetIP sXC; have [sXA _] := subsetIP sDX. + rewrite -gacentC //; apply/setIidPl; rewrite -[H :&: _]genGid //. + apply: Phi_nongen; apply/eqP; rewrite eqEsubset join_subG sPH subsetIl. + rewrite -quotientYK 1?subIset ?nPH //= -sub_quotient_pre //= -/V gacentIim. + have pP := pgroupS sPH pH; have coPX := pnat_coprime pP p'X. + rewrite -(setIid X) -(gacent_ract _ sXA). + rewrite ext_coprime_quotient_cent ?(pgroup_sol pP) ?acts_char //. + have domXb: X \subset qact_dom (<[nHA]> \ sXA) 'Phi(H). + by rewrite qact_domE ?acts_char. + rewrite gacentE // subsetIidl -/V; apply/subsetP=> v Vv; apply/afixP=> a Xa. + have [cVa dom_a] := (subsetP sXC a Xa, subsetP domXb a Xa). + have [x Nx Hx def_v] := morphimP Vv; rewrite {1}def_v qactE //=. + by rewrite -qactE ?(astab_dom cVa) ?(astab_act cVa) -?def_v. +have{B pB s_p'C_B} pC : p.-group C. + apply/pgroupP=> q q_pr; case/Cauchy=> // a Ca oa; apply: wlog_neg => p'q. + apply: (pgroupP pB) => //; rewrite -oa cardSg // s_p'C_B ?cycle_subG //. + by rewrite /pgroup -orderE oa pnatE. +have nVA: A \subset qact_dom <[nHA]> 'Phi(H) by rewrite qact_domE // acts_char. +have nCA: A \subset 'N(C). + by rewrite (subset_trans _ (astab_norm _ _)) // astabs_range. +suffices{pC nCA}: p.-group (A / C)^`(1). + by rewrite -quotient_der ?pquotient_pgroup // (subset_trans (der_sub 1 A)). +pose toAV := ((<[nHA]> / 'Phi(H)) \ nVA)%gact. +have defC: C = 'C(V | toAV). + by symmetry; rewrite astab_ract; apply/setIidPr; rewrite subIset ?subsetIl. +have abelV: p.-abelem V := Phi_quotient_abelem pH. +have ntV: V != 1 by rewrite -subG1 quotient_sub1 // proper_subn ?Phi_proper. +have: 'r(V) \in iota 1 2. + rewrite mem_iota rank_gt0 ntV (rank_abelem abelV). + have [abelH | not_abelH] := boolP (p.-abelem H). + by rewrite ltnS (leq_trans _ rH) // (rank_abelem abelH) logn_quotient. + by rewrite (leq_trans _ dimH) // ltn_log_quotient // (trivg_Phi pH). +rewrite !inE; case/pred2P=> dimV. + have isoAb: A / C \isog actperm toAV @* A. + by rewrite defC astab_range -ker_actperm first_isog. + rewrite (derG1P _) ?pgroup1 // (isog_abelian isoAb). + apply: abelianS (im_actperm_Aut _) (Aut_cyclic_abelian _). + by rewrite (abelem_cyclic abelV) -rank_abelem ?dimV. +pose Vb := sdpair1 toAV @* V; pose Ab := sdpair2 toAV @* A. +have [injV injA] := (injm_sdpair1 toAV, injm_sdpair2 toAV). +have abelVb: p.-abelem Vb := morphim_abelem _ abelV. +have ntVb: Vb != 1 by rewrite morphim_injm_eq1. +have nVbA: Ab \subset 'N(Vb) := im_sdpair_norm toAV. +pose rV := morphim_repr (abelem_repr abelVb ntVb nVbA) (subxx A). +have{defC} <-: rker rV = C; last move: rV. + rewrite rker_morphim rker_abelem morphpreI morphimK //=. + by rewrite (trivgP injA) mul1g -astabEsd // defC astab_ract 2!setIA !setIid. +have ->: 'dim Vb = 2 by rewrite (dim_abelemE abelVb) // card_injm -?rank_abelem. +move=> rV; rewrite -(eq_pgroup _ (GRing.charf_eq (char_Fp p_pr))). +by apply: der1_odd_GL2_charf (kquo_mx_faithful rV); rewrite !morphim_odd. +Qed. + +(* This is B & G, Theorem 4.18(a). *) +Theorem rank2_max_pdiv gT p q (G : {group gT}) : + solvable G -> odd #|G| -> 'r_p(G) <= 2 -> q \in \pi(G / 'O_p^'(G)) -> q <= p. +Proof. +rewrite mem_primes => solG oddG rG /and3P[pr_q _ /= q_dv_G]. +without loss Gp'1: gT G solG oddG rG q_dv_G / 'O_p^'(G) = 1. + move/(_ _ (G / 'O_p^'(G))%G); rewrite quotient_odd ?quotient_sol //. + rewrite trivg_pcore_quotient -(card_isog (quotient1_isog _)). + by rewrite p_rank_p'quotient ?pcore_pgroup ?gFnorm //; apply. +set R := 'O_p(G); have pR: p.-group R := pcore_pgroup p G. +have [sRG nRG] := andP (pcore_normal p G : R <| G). +have oddR: odd #|R| := oddSg sRG oddG. +have rR: 'r(R) <= 2 by rewrite (rank_pgroup pR) (leq_trans _ rG) ?p_rankS. +rewrite leq_eqVlt -implyNb; apply/implyP=> p'q. +have [|//] := pi_Aut_rank2_pgroup pR oddR rR _ p'q; rewrite eq_sym in p'q. +apply: (piSg (Aut_conj_aut _ G)); apply: contraLR q_dv_G. +rewrite -p'groupEpi -p'natE // Gp'1 -(card_isog (quotient1_isog _)) /pgroup. +rewrite -(card_isog (first_isog_loc _ _)) // -!pgroupE ker_conj_aut /= -/R. +set C := 'C_G(R); rewrite pquotient_pgroup ?normsI ?norms_cent ?normG //= -/C. +suffices sCR: C \subset R by rewrite (pgroupS sCR (pi_pnat pR _)). +by rewrite /C /R -(Fitting_eq_pcore _) ?cent_sub_Fitting. +Qed. + +(* This is B & G, Theorem 4.18(c,e) *) +Theorem rank2_der1_complement gT p (G : {group gT}) : + solvable G -> odd #|G| -> 'r_p(G) <= 2 -> + [/\ (*c*) p^'.-Hall(G^`(1)) 'O_p^'(G^`(1)), + (*e1*) abelian (G / 'O_{p^',p}(G)) + & (*e2*) p^'.-group (G / 'O_{p^',p}(G))]. +Proof. +move=> solG oddG rG; rewrite /pHall pcore_sub pcore_pgroup /= pnatNK. +rewrite -(pcore_setI_normal _ (der_normal 1 G)) // setIC indexgI /=. +without loss Gp'1: gT G solG oddG rG / 'O_p^'(G) = 1. + have nsGp': 'O_p^'(G) <| G := pcore_normal p^' G; have [_ nGp'] := andP nsGp'. + move/(_ _ (G / 'O_p^'(G))%G); rewrite quotient_sol // quotient_odd //=. + have Gp'1 := trivg_pcore_quotient p^' G. + rewrite p_rank_p'quotient ?pcore_pgroup // Gp'1 indexg1; case=> //=. + rewrite -quotient_der // card_quotient ?(subset_trans (der_sub 1 G)) // => ->. + rewrite (pseries_pop2 _ Gp'1) /= -pseries1 -quotient_pseries /= /pgroup. + pose isos := (isog_abelian (third_isog _ _ _), card_isog (third_isog _ _ _)). + by rewrite !{}isos ?pseries_normal ?pseries_sub_catl. +rewrite pseries_pop2 // Gp'1 indexg1 -pgroupE /=. +set R := 'O_p(G); pose C := 'C_G(R). +have [sRG nRG] := andP (pcore_normal p G : R <| G). +have sCR: C \subset R by rewrite /C /R -(Fitting_eq_pcore _) ?cent_sub_Fitting. +have pR: p.-group R := pcore_pgroup p G; have pC: p.-group C := pgroupS sCR pR. +have nCG: G \subset 'N(C) by rewrite normsI ?normG ?norms_cent. +have nsG'G: G^`(1) <| G := der_normal 1 G; have [sG'G nG'G] := andP nsG'G. +suffices sG'R: G^`(1) \subset R. + have cGbGb: abelian (G / R) := sub_der1_abelian sG'R. + rewrite -{2}(nilpotent_pcoreC p (abelian_nil cGbGb)) trivg_pcore_quotient. + by rewrite dprod1g pcore_pgroup (pgroupS sG'R pR). +rewrite pcore_max // -(pquotient_pgroup pC (subset_trans sG'G nCG)) /= -/C. +pose A := conj_aut 'O_p(G) @* G; have AutA: A \subset Aut R := Aut_conj_aut _ G. +have isoGbA: G / C \isog A by rewrite /C -ker_conj_aut first_isog_loc. +have{isoGbA} [f injf defA] := isogP isoGbA; rewrite /= -/A in defA. +rewrite quotient_der // /pgroup -(card_injm injf) ?der_sub ?morphim_der //. +have [? ?]: odd #|A| /\ solvable A by rewrite -defA !morphim_odd ?morphim_sol. +have rR: 'r(R) <= 2 by rewrite (rank_pgroup pR) (leq_trans (p_rankS p sRG)). +by rewrite defA -pgroupE (der1_Aut_rank2_pgroup pR) ?(oddSg sRG). +Qed. + +(* This is B & G, Theorem 4.18(b) *) +Theorem rank2_min_p_complement gT (G : {group gT}) (p := pdiv #|G|) : + solvable G -> odd #|G| -> 'r_p(G) <= 2 -> p^'.-Hall(G) 'O_p^'(G). +Proof. +move=> solG oddG rG; rewrite /pHall pcore_pgroup pcore_sub pnatNK /=. +rewrite -card_quotient ?gFnorm //; apply/pgroupP=> q q_pr q_dv_Gb. +rewrite inE /= eqn_leq (rank2_max_pdiv _ _ rG) ?mem_primes ?q_pr ?cardG_gt0 //. +by rewrite pdiv_min_dvd ?prime_gt1 ?(dvdn_trans q_dv_Gb) ?dvdn_quotient. +Qed. + +(* This is B & G, Theorem 4.18(d) *) +Theorem rank2_sub_p'core_der1 gT (G A : {group gT}) p : + solvable G -> odd #|G| -> 'r_p(G) <= 2 -> p^'.-subgroup(G^`(1)) A -> + A \subset 'O_p^'(G^`(1)). +Proof. +move=> solG oddG rG /andP[sAG' p'A]; rewrite sub_Hall_pcore //. +by have [-> _ _] := rank2_der1_complement solG oddG rG. +Qed. + +(* This is B & G, Corollary 4.19 *) +Corollary rank2_der1_cent_chief gT p (G Gs U V : {group gT}) : + odd #|G| -> solvable G -> Gs <| G -> 'r_p(Gs) <= 2 -> + chief_factor G V U -> p.-group (U / V) -> U \subset Gs -> + G^`(1) \subset 'C(U / V | 'Q). +Proof. +move=> oddG solG nsGsG rGs chiefUf pUf sUGs. +wlog Gs_p'_1: gT G Gs U V oddG solG nsGsG rGs chiefUf pUf sUGs / 'O_p^'(Gs) = 1. + pose K := 'O_p^'(Gs)%G; move/(_ _ (G / K) (Gs / K) (U / K) (V / K))%G. + rewrite trivg_pcore_quotient quotient_odd ?quotient_sol ?quotientS //. + have p'K: p^'.-group K := pcore_pgroup p^' Gs. + have tiUfK := coprime_TIg (pnat_coprime pUf (quotient_pgroup V p'K)). + have nsKG: K <| G := char_normal_trans (pcore_char p^' Gs) nsGsG. + have [[sG'G sGsG] nKG] := (der_sub 1 G, normal_sub nsGsG, normal_norm nsKG). + have{sGsG} [nKG' nKGs] := (subset_trans sG'G nKG, subset_trans sGsG nKG). + case/andP: chiefUf; case/maxgroupP; case/andP=> ltVU nVG maxV nsUG. + have [sUG nUG] := andP nsUG; have [sVU not_sUV] := andP ltVU. + have [nUG' nVG'] := (subset_trans sG'G nUG, subset_trans sG'G nVG). + have [sVG nVU] := (subset_trans sVU sUG, subset_trans sUG nVG). + have [nKU nKV] := (subset_trans sUG nKG, subset_trans sVG nKG). + have nsVU: V <| U by apply/andP. + rewrite p_rank_p'quotient // /chief_factor -quotient_der ?quotient_normal //. + rewrite andbT !sub_astabQR ?quotient_norms // -quotientR // => IH. + rewrite -quotient_sub1 ?comm_subG // -tiUfK subsetI quotientS ?commg_subr //. + rewrite quotientSK ?(comm_subG nVG') // (normC nKV) -quotientSK ?comm_subG //. + apply: IH => //=; last first. + rewrite -(setIid U) -(setIidPr sVU) -. + by rewrite -(morphim_quotm _ nsVU) morphim_pgroup. + apply/maxgroupP; rewrite /proper quotientS ?quotient_norms //= andbT. + rewrite quotientSK // -(normC nKV) -quotientSK // -subsetIidl tiUfK. + split=> [|Wb]; first by rewrite quotient_sub1. + do 2![case/andP]=> sWbU not_sUWb nWbG sVWb; apply/eqP; rewrite eqEsubset sVWb. + have nsWbG: Wb <| G / K by rewrite /normal (subset_trans sWbU) ?quotientS. + have [W defWb sKW] := inv_quotientN nsKG nsWbG; case/andP=> sWG nWG. + rewrite -(setIidPl sWbU) defWb -quotientGI // quotientS //. + rewrite (maxV (W :&: U))%G ?normsI //; last first. + by rewrite subsetI sVU andbT -(quotientSGK nKV sKW) -defWb. + by rewrite andbT /proper subsetIr subsetIidr -(quotientSGK nKU sKW) -defWb. +pose R := 'O_p(Gs); have pR: p.-group R := pcore_pgroup p Gs. +have nsRG: R <| G := char_normal_trans (pcore_char p Gs) nsGsG. +have [[sGsG nGsG] [sRG nRG]] := (andP nsGsG, andP nsRG). +have nsRGs: R <| Gs := pcore_normal p Gs; have [sRGs nRGs] := andP nsRGs. +have sylR: p.-Sylow(Gs) R. + have [solGs oddGs] := (solvableS sGsG solG, oddSg sGsG oddG). + have [_ _ p'Gsb] := rank2_der1_complement solGs oddGs rGs. + by rewrite /pHall pcore_sub pR -card_quotient //= -(pseries_pop2 p Gs_p'_1). +case/andP: (chiefUf); case/maxgroupP; case/andP=> ltVU nVG maxV nsUG. +have [sUG nUG] := andP nsUG; have [sVU not_sUV] := andP ltVU. +have [sVG nVU] := (subset_trans sVU sUG, subset_trans sUG nVG). +have nsVU: V <| U by apply/andP. +have nVGs := subset_trans sGsG nVG; have nVR := subset_trans sRGs nVGs. +have{sylR} sUfR: U / V \subset R / V. + have sylRb: p.-Sylow(Gs / V) (R / V) by rewrite quotient_pHall. + by rewrite (sub_normal_Hall sylRb) ?quotientS ?quotient_normal. +have pGb: p.-group((G / 'C_G(R))^`(1)). + pose A := conj_aut 'O_p(Gs) @* G. + have AA: A \subset Aut R := Aut_conj_aut _ G. + have isoGbA: G / 'C_G(R) \isog A by rewrite -ker_conj_aut first_isog_loc. + have{isoGbA} [f injf defA] := isogP isoGbA; rewrite /= -/A in defA. + rewrite /pgroup -(card_injm injf) ?der_sub ?morphim_der //. + have [? ?]: odd #|A| /\ solvable A by rewrite -defA !morphim_odd ?morphim_sol. + have rR: 'r(R) <= 2 by rewrite (rank_pgroup pR) (leq_trans (p_rankS p sRGs)). + by rewrite defA -pgroupE (der1_Aut_rank2_pgroup pR) ?(oddSg sRG). +set C := 'C_G(U / V | 'Q). +have nUfG: [acts G, on U / V | 'Q] by rewrite actsQ. +have nCG: G \subset 'N(C) by rewrite -(setIidPl nUfG) normsGI ?astab_norm. +have{pGb sUfR} pGa': p.-group (G / C)^`(1). + have nCRG : G \subset 'N('C_G(R)) by rewrite normsI ?normG ?norms_cent. + have sCR_C: 'C_G(R) \subset C. + rewrite subsetI subsetIl sub_astabQ ?subIset ?nVG ?(centsS sUfR) //=. + by rewrite quotient_cents ?subsetIr. + have [f /= <-]:= homgP (homg_quotientS nCRG nCG sCR_C). + by rewrite -morphim_der //= morphim_pgroup. +have irrG: acts_irreducibly (G / C) (U / V) ('Q %% _). + by rewrite acts_irr_mod_astab // acts_irrQ // chief_factor_minnormal. +have Ga_p_1: 'O_p(G / C) = 1. + rewrite (pcore_faithful_irr_act pUf _ irrG) ?modact_faithful //. + by rewrite gacentC ?quotientS ?subsetT ?subsetIr //= setICA subsetIl. +have sG'G := der_sub 1 G; have nCG' := subset_trans sG'G nCG. +rewrite -subsetIidl -{2}(setIidPl sG'G) -setIA subsetIidl -/C. +by rewrite -quotient_sub1 /= ?quotient_der //= -Ga_p_1 pcore_max ?der_normal. +Qed. + +(* This is B & G, Theorem 4.20(a) *) +Theorem rank2_der1_sub_Fitting gT (G : {group gT}) : + odd #|G| -> solvable G -> 'r('F(G)) <= 2 -> G^`(1) \subset 'F(G). +Proof. +move=> oddG solG Fle2; have nsFG := Fitting_normal G. +apply: subset_trans (chief_stab_sub_Fitting solG _) => //. +rewrite subsetI der_sub; apply/bigcapsP=> [[U V] /= /andP[chiefUV sUF]]. +have [p p_pr /andP[pUV _]] := is_abelemP (sol_chief_abelem solG chiefUV). +apply: rank2_der1_cent_chief nsFG _ _ pUV sUF => //. +exact: leq_trans (p_rank_le_rank p _) Fle2. +Qed. + +(* This is B & G, Theorem 4.20(b) *) +Theorem rank2_char_Sylow_normal gT (G S T : {group gT}) : + odd #|G| -> solvable G -> 'r('F(G)) <= 2 -> + Sylow G S -> T \char S -> T \subset S^`(1) -> T <| G. +Proof. +set F := 'F(G) => oddG solG Fle2 /SylowP[p p_pr sylS] charT sTS'. +have [sSG pS _] := and3P sylS. +have nsFG: F <| G := Fitting_normal G; have [sFG nFG] := andP nsFG. +have nFS := subset_trans sSG nFG; have nilF: nilpotent F := Fitting_nil _. +have cGGq: abelian (G / F). + by rewrite sub_der1_abelian ?rank2_der1_sub_Fitting. +have nsFS_G: F <*> S <| G. + rewrite -(quotientGK nsFG) norm_joinEr // -(quotientK nFS) cosetpre_normal. + by rewrite -sub_abelian_normal ?quotientS. +have sylSF: p.-Sylow(F <*> S) S. + by rewrite (pHall_subl _ _ sylS) ?joing_subr // join_subG sFG. +have defG: G :=: F * 'N_G(S). + rewrite -{1}(Frattini_arg nsFS_G sylSF) /= norm_joinEr // -mulgA. + by congr (_ * _); rewrite mulSGid // subsetI sSG normG. +rewrite /normal (subset_trans (char_sub charT)) //= defG mulG_subG /= -/F. +rewrite setIC andbC subIset /=; last by rewrite (char_norm_trans charT). +case/dprodP: (nilpotent_pcoreC p nilF); rewrite /= -/F => _ defF cFpFp' _. +have defFp: 'O_p(F) = F :&: S. + rewrite -{2}defF -group_modl ?coprime_TIg ?mulg1 //. + by rewrite coprime_sym (pnat_coprime pS (pcore_pgroup _ _)). + by rewrite p_core_Fitting pcore_sub_Hall. +rewrite -defF mulG_subG /= -/F defFp setIC subIset ?(char_norm charT) //=. +rewrite cents_norm // (subset_trans cFpFp') // defFp centS // subsetI. +rewrite (char_sub charT) (subset_trans (subset_trans sTS' (dergS 1 sSG))) //. +exact: rank2_der1_sub_Fitting. +Qed. + +(* This is B & G, Theorem 4.20(c), for the last factor of the series. *) +Theorem rank2_min_p'core_Hall gT (G : {group gT}) (p := pdiv #|G|) : + odd #|G| -> solvable G -> 'r('F(G)) <= 2 -> p^'.-Hall(G) 'O_p^'(G). +Proof. +set F := 'F(G) => oddG solG Fle2. +have nsFG: F <| G := Fitting_normal G; have [sFG nFG] := andP nsFG. +have [H] := inv_quotientN nsFG (pcore_normal p^' _). +rewrite /= -/F => defH sFH nsHG; have [sHG nHG] := andP nsHG. +have [P sylP] := Sylow_exists p H; have [sPH pP _] := and3P sylP. +have sPF: P \subset F. + rewrite -quotient_sub1 ?(subset_trans (subset_trans sPH sHG)) //. + rewrite -(setIidPl (quotientS _ sPH)) -defH coprime_TIg //. + by rewrite coprime_morphl // (pnat_coprime pP (pcore_pgroup _ _)). +have nilGq: nilpotent (G / F). + by rewrite abelian_nil ?sub_der1_abelian ?rank2_der1_sub_Fitting. +have pGq: p.-group (G / H). + rewrite /pgroup -(card_isog (third_isog sFH nsFG nsHG)) /= -/F -/(pgroup _ _). + rewrite -(dprodW (nilpotent_pcoreC p nilGq)) defH quotientMidr. + by rewrite quotient_pgroup ?pcore_pgroup. +rewrite pHallE pcore_sub -(Lagrange sHG) partnM // -card_quotient //=. +have hallHp': p^'.-Hall(H) 'O_p^'(H). + case p'H: (p^'.-group H). + by rewrite pHallE /= pcore_pgroup_id ?subxx //= part_pnat_id. + have def_p: p = pdiv #|H|. + have [p_pr pH]: prime p /\ p %| #|H|. + apply/andP; apply: contraFT p'H => p'H; apply/pgroupP=> q q_pr qH. + by apply: contraNneq p'H => <-; rewrite q_pr qH. + apply/eqP; rewrite eqn_leq ?pdiv_min_dvd ?prime_gt1 //. + rewrite pdiv_prime // cardG_gt1. + by case: eqP p'H => // ->; rewrite pgroup1. + exact: dvdn_trans (pdiv_dvd _) (cardSg (normal_sub nsHG)). + rewrite def_p rank2_min_p_complement ?(oddSg sHG) ?(solvableS sHG) -?def_p //. + rewrite -(p_rank_Sylow sylP) (leq_trans (p_rank_le_rank _ _)) //. + exact: leq_trans (rankS sPF) Fle2. +rewrite -(card_Hall hallHp') part_p'nat ?pnatNK ?muln1 // subset_leqif_card. + by rewrite pcore_max ?pcore_pgroup ?(char_normal_trans (pcore_char _ _)). +rewrite pcore_max ?pcore_pgroup // (normalS _ _ (pcore_normal _ _)) //. +rewrite -quotient_sub1 ?(subset_trans (pcore_sub _ _)) //. +rewrite -(setIidPr (quotientS _ (pcore_sub _ _))) coprime_TIg //. +by rewrite coprime_morphr // (pnat_coprime pGq (pcore_pgroup _ _)). +Qed. + +(* This is B & G, Theorem 4.20(c), for intermediate factors. *) +Theorem rank2_ge_pcore_Hall gT m (G : {group gT}) (pi := [pred p | p >= m]) : + odd #|G| -> solvable G -> 'r('F(G)) <= 2 -> pi.-Hall(G) 'O_pi(G). +Proof. +elim: {G}_.+1 {-2}G (ltnSn #|G|) => // n IHn G. +rewrite ltnS => leGn oddG solG Fle2; pose p := pdiv #|G|. +have [defGpi | not_pi_G] := eqVneq 'O_pi(G) G. + by rewrite /pHall pcore_sub pcore_pgroup defGpi indexgg. +have pi'_p: (p \in pi^'). + apply: contra not_pi_G => pi_p; rewrite eqEsubset pcore_sub pcore_max //. + apply/pgroupP=> q q_pr qG; apply: leq_trans pi_p _. + by rewrite pdiv_min_dvd ?prime_gt1. +pose Gp' := 'O_p^'(G); have sGp'G: Gp' \subset G := pcore_sub _ _. +have hallGp'pi: pi.-Hall(Gp') 'O_pi(Gp'). + apply: IHn; rewrite ?(oddSg sGp'G) ?(solvableS sGp'G) //; last first. + by apply: leq_trans (rankS _) Fle2; rewrite /= Fitting_pcore pcore_sub. + apply: leq_trans (proper_card _) leGn. + rewrite properEneq pcore_sub andbT; apply/eqP=> defG. + suff: p \in p^' by case/eqnP. + have p'G: p^'.-group G by rewrite -defG pcore_pgroup. + rewrite (pgroupP p'G) ?pdiv_dvd ?pdiv_prime // cardG_gt1. + by apply: contra not_pi_G; move/eqP->; rewrite (trivgP (pcore_sub _ _)). +have defGp'pi: 'O_pi(Gp') = 'O_pi(G). + rewrite -pcoreI; apply: eq_pcore => q; apply: andb_idr. + by apply: contraL => /=; move/eqnP->. +have hallGp': p^'.-Hall(G) Gp' by rewrite rank2_min_p'core_Hall. +rewrite pHallE pcore_sub /= -defGp'pi (card_Hall hallGp'pi) (card_Hall hallGp'). +by rewrite partn_part // => q; apply: contraL => /=; move/eqnP->. +Qed. + +(* This is B & G, Theorem 4.20(c), for the first factor of the series. *) +Theorem rank2_max_pcore_Sylow gT (G : {group gT}) (p := max_pdiv #|G|) : + odd #|G| -> solvable G -> 'r('F(G)) <= 2 -> p.-Sylow(G) 'O_p(G). +Proof. +move=> oddG solG Fle2; pose pi := [pred q | p <= q]. +rewrite pHallE pcore_sub eqn_leq -{1}(part_pnat_id (pcore_pgroup _ _)). +rewrite dvdn_leq ?partn_dvd ?cardSg ?pcore_sub // /=. +rewrite (@eq_in_partn _ pi) => [|q piGq]; last first. + by rewrite !inE eqn_leq; apply: andb_idl => le_q_p; exact: max_pdiv_max. +rewrite -(card_Hall (rank2_ge_pcore_Hall p oddG solG Fle2)) -/pi. +rewrite subset_leq_card // pcore_max ?pcore_normal //. +apply: sub_in_pnat (pcore_pgroup _ _) => q; move/(piSg (pcore_sub _ _)) => piGq. +by rewrite !inE eqn_leq max_pdiv_max. +Qed. + +End Section4. diff --git a/mathcomp/odd_order/BGsection5.v b/mathcomp/odd_order/BGsection5.v new file mode 100644 index 0000000..ab5a14a --- /dev/null +++ b/mathcomp/odd_order/BGsection5.v @@ -0,0 +1,536 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div. +Require Import fintype finset prime fingroup morphism perm automorphism action. +Require Import quotient cyclic gfunctor pgroup gproduct center commutator. +Require Import gseries nilpotent sylow abelian maximal hall. +Require Import BGsection1 BGsection4. + +(******************************************************************************) +(* This file covers Section 5 of B & G, except for some technical results *) +(* that are not actually used in the proof of the Odd Order Theorem, namely *) +(* part (c) of Theorem 5.5, parts (b), (d) and (e) of Theorem 5.5, and all of *) +(* Theorem 5.7. We also make the following change: in B & G, narrow p-groups *) +(* of rank at least 3 are defined by the structure of the centralisers of *) +(* their prime subgroups, then characterized by their rank 2 elementary *) +(* abelian subgroups in Theorem 5.3. We exchange the two, because the latter *) +(* condition is easier to check, and is the only one used later in the proof. *) +(* *) +(* p.-narrow G == G has a maximal elementary abelian p-subgroup of *) +(* p-rank at most 2. *) +(* := ('r_p(G) > 2) ==> ('E_p^2(G) :&: 'E*_p(G) != set0) *) +(* *) +(* narrow_structure p G <-> G has a subgroup S of order p whose centraliser *) +(* is the direct product of S and a cyclic group C, *) +(* i.e., S \x C = 'C_G(S). This is the condition used *) +(* in the definition of "narrow" in B & G, p. 2. *) +(* Theorem 5.3 states that for odd p this definition *) +(* is equivalent to ours, and this property is not *) +(* used outside of Section 5. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Reserved Notation "p .-narrow" (at level 2, format "p .-narrow"). + +Section Definitions. + +Variables (gT : finGroupType) (p : nat) (A : {set gT}). + +Definition narrow := ('r_p(A) > 2) ==> ('E_p^2(A) :&: 'E*_p(A) != set0). + +Inductive narrow_structure : Prop := + NarrowStructure (S C : {group gT}) of + S \subset A & C \subset A & #|S| = p & cyclic C & S \x C = 'C_A(S). + +End Definitions. + +Notation "p .-narrow" := (narrow p) : group_scope. + +Section IsoDef. + +Variables (gT rT : finGroupType) (p : nat). +Implicit Types G H : {group gT}. +Implicit Type R : {group rT}. + +Lemma injm_narrow G H (f : {morphism G >-> rT}) : + 'injm f -> H \subset G -> p.-narrow (f @* H) = p.-narrow H. +Proof. +move=> injf sHG; rewrite /narrow injm_p_rank //; congr (_ ==> _). +apply/set0Pn/set0Pn=> [] [E /setIP[Ep2E maxE]]. + exists (invm injf @* E)%G; rewrite -[H](group_inj (morphim_invm injf _)) //. + have sEfG: E \subset f @* G. + by rewrite (subset_trans _ (morphimS _ sHG)) //; case/pnElemP: Ep2E. + by rewrite inE injm_pnElem ?injm_pmaxElem ?injm_invm ?morphimS // Ep2E. +have sEG: E \subset G by rewrite (subset_trans _ sHG) //; case/pnElemP: Ep2E. +by exists (f @* E)%G; rewrite inE injm_pnElem ?injm_pmaxElem // Ep2E. +Qed. + +Lemma isog_narrow G R : G \isog R -> p.-narrow G = p.-narrow R. +Proof. by case/isogP=> f injf <-; rewrite injm_narrow. Qed. + +(* No isomorphism theorems for narrow_structure, which is not used outside of *) +(* this file. *) + +End IsoDef. + +Section Five. + +Implicit Type gT : finGroupType. +Implicit Type p : nat. + +Section OneGroup. + +Variables (gT : finGroupType) (p : nat) (R : {group gT}). +Implicit Types B E S : {group gT}. + +Lemma narrowJ x : p.-narrow (R :^ x) = p.-narrow R. +Proof. by apply: isog_narrow (isog_symr (conj_isog R x)). Qed. + +Hypotheses (pR : p.-group R) (oddR : odd #|R|). + +Section Rank3. + +Hypothesis rR : 2 < 'r_p(R). + +(* This lemma uses only the rR hypothesis. *) +Lemma narrow_pmaxElem : p.-narrow R -> exists E, E \in 'E_p^2(R) :&: 'E*_p(R). +Proof. by move=> nnP; apply: set0Pn; apply: implyP rR. Qed. + +Let ntR : R :!=: 1. Proof. by case: eqP rR => // ->; rewrite p_rank1. Qed. +Let p_pr : prime p. Proof. by case: (pgroup_pdiv pR ntR). Qed. +Let p_gt1 : p > 1. Proof. exact: prime_gt1. Qed. + +(* This is B & G, Lemma 5.1(a). *) +Lemma rank3_SCN3 : exists B, B \in 'SCN_3(R). +Proof. +by apply/set0Pn; rewrite -(rank2_SCN3_empty pR oddR) leqNgt (rank_pgroup pR) rR. +Qed. + +(* This is B & G, Lemma 5.1(b). *) +Lemma normal_p2Elem_SCN3 E : + E \in 'E_p^2(R) -> E <| R -> exists2 B, B \in 'SCN_3(R) & E \subset B. +Proof. +move=> Ep2E /andP[sER nER]; have [_ abelE dimE] := pnElemP Ep2E. +have [B Ep3B nBR]: exists2 B, B \in 'E_p^3(R) & R \subset 'N(B). + have [C] := rank3_SCN3; case/setIdP=> SCN_C rC. + have [nsCR cCC] := andP (maxgroupp (SCN_max SCN_C)). + have [sCR _] := andP nsCR; have pC: p.-group C := pgroupS sCR pR. + have{pC cCC} abelC1: p.-abelem 'Ohm_1(C) := Ohm1_abelem pC cCC. + have dimC1: 3 <= logn p #|'Ohm_1(C)| by rewrite -rank_abelem // rank_Ohm1. + have nsC1R: 'Ohm_1(C) <| R := char_normal_trans (Ohm_char 1 _) nsCR. + have [B [sBC1 nsBR oB]] := normal_pgroup pR nsC1R dimC1. + have [sBR nBR] := andP nsBR; exists B => //; apply/pnElemP. + by rewrite oB pfactorK // (abelemS sBC1). +have [sBR abelB dimB] := pnElemP Ep3B; have [pB cBB _] := and3P abelB. +have [oB oE] := (card_pnElem Ep3B, card_pnElem Ep2E). +pose Bs := (E <*> 'C_B(E))%G; have sCB: 'C_B(E) \subset B := subsetIl B _. +have sBsR: Bs \subset R by rewrite join_subG sER subIset ?sBR. +suffices Bs_gt2: 2 < logn p #|Bs|. + have nBsR: Bs <| R by rewrite /normal sBsR // normsY ?normsI ?norms_cent. + have abelBs: p.-abelem Bs. + by rewrite (cprod_abelem p (cprodEY _)) ?subsetIr // abelE (abelemS sCB). + have [C maxC sBsC] : {H | [max H | H <| R & abelian H ] & Bs \subset H}. + by apply: maxgroup_exists; rewrite nBsR (abelem_abelian abelBs). + exists C; last by rewrite (subset_trans _ sBsC) ?joing_subl. + by rewrite inE (max_SCN pR) ?(leq_trans Bs_gt2) // -rank_abelem ?rankS. +apply: contraFT (ltnn 2); rewrite -leqNgt => Bs_le2. +have{Bs_le2} sCE: 'C_B(E) \subset E. + rewrite (sameP joing_idPl eqP) eq_sym eqEcard joing_subl /=. + by rewrite (card_pgroup (pgroupS sBsR pR)) oE leq_exp2l. +have dimCBE: 2 <= logn p #|'C_B(E)|. + rewrite -ltnS -dimB -addn1 -leq_subLR -logn_div ?divgS ?cardSg //. + by rewrite logn_quotient_cent_abelem ?dimE ?(subset_trans sBR nER). +have defE: 'C_B(E) = E. + apply/eqP; rewrite eqEcard sCE oE /=. + by rewrite (card_pgroup (pgroupS sCB pB)) leq_exp2l. +by rewrite -dimB -dimE -defE lognSg // subsetIidl sub_abelian_cent // -defE. +Qed. + +Let Z := 'Ohm_1('Z(R)). +Let W := 'Ohm_1('Z_2(R)). +Let T := 'C_R(W). + +Let ntZ : Z != 1. +Proof. by rewrite Ohm1_eq1 (center_nil_eq1 (pgroup_nil pR)). Qed. + +Let sZR : Z \subset R. +Proof. exact: subset_trans (Ohm_sub 1 _) (center_sub R). Qed. + +Let abelZ : p.-abelem (Z). +Proof. by rewrite (Ohm1_abelem (pgroupS _ pR)) ?center_sub ?center_abelian. Qed. + +Let pZ : p.-group Z. +Proof. exact: abelem_pgroup abelZ. Qed. + +Let defCRZ : 'C_R(Z) = R. +Proof. +apply/eqP; rewrite eqEsubset subsetIl subsetIidl centsC. +by rewrite (subset_trans (Ohm_sub 1 _)) ?subsetIr. +Qed. + +Let sWR : W \subset R. +Proof. exact: subset_trans (Ohm_sub 1 _) (ucn_sub 2 R). Qed. + +Let nWR : R \subset 'N(W). +Proof. exact: char_norm_trans (Ohm_char 1 _) (char_norm (ucn_char 2 R)). Qed. + +(* This is B & G, Lemma 5.2. *) +Lemma Ohm1_ucn_p2maxElem E : + E \in 'E_p^2(R) :&: 'E*_p(R) -> + [/\ (*a*) ~~ (E \subset T), + (*b*) #|Z| = p /\ [group of W] \in 'E_p^2(R) + & (*c*) T \char R /\ #|R : T| = p ]. +Proof. +case/setIP=> Ep2E maxE; have defCRE1 := Ohm1_cent_max maxE pR. +have [[sER abelE dimE] oE] := (pnElemP Ep2E, card_pnElem Ep2E). +have [[sZR_R nZR_R] [pE _ eE]] := (andP (center_normal R), and3P abelE). +have{nZR_R} nZR: R \subset 'N(Z) := char_norm_trans (Ohm_char 1 _) nZR_R. +have{sZR_R} [pZR pW] := (pgroupS sZR_R pR, pgroupS sWR pR). +have sZE: Z \subset E by rewrite -defCRE1 OhmS ?setIS // centS. +have rCRE : 'r_p('C_R(E)) = 2 by rewrite -p_rank_Ohm1 defCRE1 p_rank_abelem. +have oZ: #|Z| = p. + apply/prime_nt_dvdP; rewrite -?trivg_card1 // (card_pgroup pZ) /= -/Z. + rewrite (@dvdn_exp2l _ _ 1) // -ltnS -dimE properG_ltn_log //= -/Z. + by case/eqVproper: sZE rR => // defZ; rewrite -defCRZ defZ rCRE ltnn. +have ncycR: ~~ cyclic R by rewrite (odd_pgroup_rank1_cyclic pR) // -(subnKC rR). +have [ncycW eW] := Ohm1_odd_ucn2 pR oddR ncycR; rewrite -/W in ncycW eW. +have sWRZ: [~: W, R] \subset Z. + rewrite [Z](OhmE 1 pZR) sub_gen //= -ucn1 subsetI. + rewrite (subset_trans _ (ucn_comm 1 _)) ?commSg ?Ohm_sub //. + by move: nWR eW; rewrite -commg_subl -sub_LdivT; apply: subset_trans. +have sZW: Z \subset W by rewrite OhmS /= -?ucn1 ?ucn_subS //. +have ltZW: Z \proper W. + by rewrite properEneq; case: eqP ncycW => // <-; rewrite prime_cyclic ?oZ. +have sWRE := subset_trans sWRZ sZE. +have nEW: W \subset 'N(E) by rewrite -commg_subr (subset_trans _ sWRE) ?commgSS. +have defZ: 'C_W(E) = Z. + have sCE: 'C_W(E) \subset E. + rewrite -{2}defCRE1 (OhmE 1 (pgroupS (subsetIl R _) pR)) sub_gen //. + by rewrite subsetI setSI // subIset // sub_LdivT eW. + have [defC | ltCE] := eqVproper sCE. + have sEW: E \subset W by rewrite -defC subsetIl. + have nsER: E <| R. + by rewrite /normal sER -commg_subl (subset_trans (commSg R sEW)). + have [B scn3B sEB] := normal_p2Elem_SCN3 Ep2E nsER. + have [scnB dimB] := setIdP scn3B; have [_ scBR] := SCN_P scnB. + rewrite ltnNge -rank_Ohm1 -dimE -rank_abelem ?rankS // in dimB. + by rewrite -scBR -defCRE1 OhmS // setIS ?centS. + apply/eqP; rewrite eq_sym eqEcard oZ (card_pgroup (pgroupS sCE pE)) /= -/W. + rewrite subsetI sZW (centsS sER); last by rewrite centsC -subsetIidl defCRZ. + by rewrite (leq_exp2l _ 1) // -ltnS -dimE properG_ltn_log. +have dimW: logn p #|W| = 2. + apply/eqP; rewrite -(Lagrange sZW) lognM ?cardG_gt0 // oZ (pfactorK 1) //=. + rewrite -/Z eqSS eqn_leq -{1}defZ logn_quotient_cent_abelem ?dimE // -/W. + by rewrite -divgS // logn_div ?cardSg // subn_gt0 properG_ltn_log. +have abelW: p.-abelem W. + by rewrite (abelem_Ohm1 (pgroupS _ pR)) ?(p2group_abelian pW) ?dimW ?ucn_sub. +have charT: T \char R. + by rewrite subcent_char ?char_refl //= (char_trans (Ohm_char 1 _)) ?ucn_char. +rewrite 2!inE sWR abelW dimW; do 2?split => //. + by apply: contra (proper_subn ltZW); rewrite -defZ !subsetI subxx sER centsC. +apply/prime_nt_dvdP=> //. + rewrite indexg_eq1 subsetIidl centsC; apply: contraFN (ltnn 1) => cRW. + by rewrite -dimW -(setIidPl (centsS sER cRW)) defZ oZ (pfactorK 1). +rewrite -(part_pnat_id (pnat_dvd (dvdn_indexg _ _) pR)) p_part. +by rewrite (@dvdn_exp2l p _ 1) ?logn_quotient_cent_abelem ?dimW. +Qed. + +(* This is B & G, Theorem 5.3(d); we omit parts (a)-(c) as they are mostly *) +(* redundant with Lemma 5.2, given our definition of "narrow". *) +Theorem narrow_cent_dprod S : + p.-narrow R -> #|S| = p -> S \subset R -> 'r_p('C_R(S)) <= 2 -> + [/\ cyclic 'C_T(S), S :&: R^`(1) = 1, S :&: T = 1 & S \x 'C_T(S) = 'C_R(S)]. +Proof. +move=> nnR oS sSR rS; have pS : p.-group S := pgroupS sSR pR. +have [E maxEp2E] := narrow_pmaxElem nnR; have [Ep2E maxE] := setIP maxEp2E. +have [not_sET [oZ Ep2W] [charT maxT]] := Ohm1_ucn_p2maxElem maxEp2E. +have cZS : S \subset 'C(Z) by rewrite (subset_trans sSR) // -defCRZ subsetIr. +have nZS : S \subset 'N(Z) by rewrite cents_norm. +have cSS : abelian S by rewrite cyclic_abelian ?prime_cyclic // oS. +pose SZ := (S <*> [group of Z])%G; have sSSZ: S \subset SZ := joing_subl _ _. +have sSZ_R: SZ \subset R by rewrite join_subG sSR sZR. +have abelSZ : p.-abelem SZ. + by rewrite /= joingC (cprod_abelem p (cprodEY cZS)) abelZ prime_abelem. +have tiSZ: S :&: Z = 1. + rewrite prime_TIg ?oS //= -/Z; apply: contraL rR => sZS. + by rewrite -leqNgt (leq_trans _ rS) ?p_rankS // -{1}defCRZ setIS ?centS. +have{tiSZ} oSZ: #|SZ| = (p ^ 2)%N by rewrite /= norm_joinEl ?TI_cardMg ?oS ?oZ. +have Ep2SZ: SZ \in 'E_p^2(R) by rewrite pnElemE // !inE sSZ_R abelSZ oSZ eqxx. +have{oSZ Ep2SZ abelSZ sSZ_R} maxSZ: SZ \in 'E_p^2(R) :&: 'E*_p(R). + rewrite inE Ep2SZ; apply/pmaxElemP; rewrite inE sSZ_R abelSZ. + split=> // H /setIdP[sHR abelH] sSZH. + have [[_ _ dimSZ] [cHH pH _]] := (pnElemP Ep2SZ, and3P abelH). + have sSH: S \subset H := subset_trans sSSZ sSZH. + have{sSH} sH_CRS: H \subset 'C_R(S) by rewrite subsetI sHR (centsS sSH). + have{sH_CRS}: 'r_p(H) <= 2 by rewrite (leq_trans _ rS) ?p_rankS. + apply: contraTeq; rewrite eq_sym eqEproper sSZH negbK => lSZH. + by rewrite -ltnNge p_rank_abelem // -dimSZ properG_ltn_log. +have sZT: Z \subset T. + by rewrite subsetI sZR (centsS sWR) // centsC -defCRZ subsetIr. +have{SZ sSSZ maxSZ} not_sST: ~~ (S \subset T). + have: ~~ (SZ \subset T) by case/Ohm1_ucn_p2maxElem: maxSZ. + by rewrite join_subG sZT andbT. +have tiST: S :&: T :=: 1 by rewrite prime_TIg ?oS. +have defST: S * T = R. + apply/eqP; rewrite eqEcard TI_cardMg ?mul_subG ?subsetIl //=. + by rewrite mulnC oS -maxT Lagrange ?subsetIl. +have cRRb: abelian (R / T) by rewrite -defST quotientMidr quotient_abelian. +have sR'T: R^`(1) \subset T by rewrite der1_min ?char_norm. +have TI_SR': S :&: R^`(1) :=: 1. + by rewrite prime_TIg ?oS // (contra _ not_sST) //; move/subset_trans->. +have defCRS : S \x 'C_T(S) = 'C_R(S). + rewrite (dprodE _ _) ?subsetIr //= -/T; last by rewrite setIA tiST setI1g. + rewrite -{1}(center_idP cSS) subcent_TImulg ?defST //. + by rewrite subsetI normG (subset_trans sSR) ?char_norm. +have sCTSR: 'C_T(S) \subset R by rewrite subIset ?subsetIl. +split; rewrite ?(odd_pgroup_rank1_cyclic (pgroupS _ pR) (oddSg _ oddR)) //= -/T. +rewrite -ltnS (leq_trans _ rS) //= -(p_rank_dprod p defCRS) -add1n leq_add2r. +by rewrite -rank_pgroup // rank_gt0 -cardG_gt1 oS. +Qed. + +(* This is B & G, Corollary 5.4. Given our definition of narrow, this is used *) +(* directly in the proof of the main part of Theorem 5.3. *) +Corollary narrow_centP : + reflect (exists S, [/\ gval S \subset R, #|S| = p & 'r_p('C_R(S)) <= 2]) + (p.-narrow R). +Proof. +rewrite /narrow rR; apply: (iffP (set0Pn _)) => [[E maxEp2E]|[S [sSR oS rCRS]]]. + have [Ep2E maxE] := setIP maxEp2E. + have{maxEp2E} [_ [oZ _] _] := Ohm1_ucn_p2maxElem maxEp2E. + have [sER abelE dimE] := pnElemP Ep2E; have oE := card_pnElem Ep2E. + have sZE: Z \subset E by rewrite -(Ohm1_cent_max maxE pR) OhmS ?setIS ?centS. + have [S defE] := abelem_split_dprod abelE sZE; exists S. + have{defE} [[_ defZS _ _] oZS] := (dprodP defE, dprod_card defE). + split; first by rewrite (subset_trans _ sER) // -defZS mulG_subr. + by apply/eqP; rewrite -(eqn_pmul2l (ltnW p_gt1)) -{1}oZ oZS oE. + rewrite -dimE -p_rank_abelem // -(Ohm1_cent_max maxE pR) p_rank_Ohm1. + by rewrite -defZS /= centM setIA defCRZ. +have abelS := prime_abelem p_pr oS. +have cSZ: Z \subset 'C(S) by rewrite (centsS sSR) // centsC -defCRZ subsetIr. +have sSZR: S <*> Z \subset R by rewrite join_subG sSR. +have defSZ: S \x Z = S <*> Z. + rewrite dprodEY ?prime_TIg ?oS //= -/Z; apply: contraL rR => sSZ. + by rewrite -leqNgt (leq_trans _ rCRS) ?p_rankS // -{1}defCRZ setIS ?centS. +have abelSZ: p.-abelem (S <*> Z) by rewrite (dprod_abelem p defSZ) abelS. +have [pSZ cSZSZ _] := and3P abelSZ. +have dimSZ: logn p #|S <*> Z| = 2. + apply/eqP; rewrite -p_rank_abelem // eqn_leq (leq_trans (p_rankS _ _) rCRS). + rewrite -(p_rank_dprod p defSZ) p_rank_abelem // oS (pfactorK 1) // ltnS. + by rewrite -rank_pgroup // rank_gt0. + by rewrite subsetI sSZR sub_abelian_cent ?joing_subl. +exists [group of S <*> Z]; rewrite 3!inE sSZR abelSZ dimSZ /=. +apply/pmaxElemP; rewrite inE sSZR; split=> // E; case/pElemP=> sER abelE sSZE. +apply: contraTeq rCRS; rewrite eq_sym -ltnNge -dimSZ => neqSZE. +have [[pE cEE _] sSE] := (and3P abelE, subset_trans (joing_subl S Z) sSZE). +rewrite (leq_trans (properG_ltn_log pE _)) ?properEneq ?neqSZE //. +by rewrite -p_rank_abelem ?p_rankS // subsetI sER sub_abelian_cent. +Qed. + +(* This is the main statement of B & G, Theorem 5.3, stating the equivalence *) +(* of the structural and rank characterizations of the "narrow" property. Due *) +(* to our definition of "narrow", the equivalence is the converse of that in *) +(* B & G (we define narrow in terms of maximal elementary abelian subgroups). *) +Lemma narrow_structureP : reflect (narrow_structure p R) (p.-narrow R). +Proof. +apply: (iffP idP) => [nnR | [S C sSR sCR oS cycC defSC]]. + have [S [sSR oS rCRS]] := narrow_centP nnR. + have [cycC _ _ defCRS] := narrow_cent_dprod nnR oS sSR rCRS. + by exists S [group of 'C_T(S)]; rewrite //= -setIA subsetIl. +apply/narrow_centP; exists S; split=> //. +have cycS: cyclic S by rewrite prime_cyclic ?oS. +rewrite -(p_rank_dprod p defSC) -!(rank_pgroup (pgroupS _ pR)) // -addn1. +rewrite leq_add -?abelian_rank1_cyclic ?cyclic_abelian //. +Qed. + +End Rank3. + +(* This is B & G, Theoren 5.5 (a) and (b). Part (c), which is not used in the *) +(* proof of the Odd Order Theorem, is omitted. *) +Theorem Aut_narrow (A : {group {perm gT}}) : + p.-narrow R -> solvable A -> A \subset Aut R -> odd #|A| -> + [/\ (*a*) p^'.-group (A / 'O_p(A)), abelian (A / 'O_p(A)) + & (*b*) 2 < 'r(R) -> forall x, x \in A -> p^'.-elt x -> #[x] %| p.-1]. +Proof. +move=> nnR solA AutA oddA; have nilR := pgroup_nil pR. +have [rR | rR] := leqP 'r(R) 2. + have pA' := der1_Aut_rank2_pgroup pR oddR rR AutA solA oddA. + have sA'Ap: A^`(1) \subset 'O_p(A) by rewrite pcore_max ?der_normal. + have cAbAb: abelian (A / 'O_p(A)) by rewrite sub_der1_abelian. + split; rewrite // -(nilpotent_pcoreC p (abelian_nil cAbAb)). + by rewrite trivg_pcore_quotient dprod1g pcore_pgroup. +have ntR: R :!=: 1 by rewrite -rank_gt0 2?ltnW. +rewrite (rank_pgroup pR) in rR. +have [H [charH sHRZ] _ eH pCH] := critical_odd pR oddR ntR. +have{ntR} [[p_pr _ _] sHR] := (pgroup_pdiv pR ntR, char_sub charH). +have ntH: H :!=: 1 by rewrite trivg_exponent eH -prime_coprime ?coprimen1. +have{nnR} [S C sSR sCR oS cycC defSC] := narrow_structureP rR nnR. +have [_ mulSC cSC tiSC] := dprodP defSC. +have abelS: p.-abelem S := prime_abelem p_pr oS; have [pS cSS _] := and3P abelS. +have cycS: cyclic S by rewrite prime_cyclic ?oS. +have tiHS: H :&: S = 1. + have rCRS: 'r_p('C_R(S)) <= 2. + rewrite -(p_rank_dprod p defSC) -addn1 -!rank_pgroup ?(pgroupS _ pR) //. + by rewrite leq_add -?abelian_rank1_cyclic ?cyclic_abelian. + rewrite setIC prime_TIg ?oS //; apply: contraL (rCRS) => sSH; rewrite -ltnNge. + have cZHS: S \subset 'C('Z(H)) by rewrite centsC (centsS sSH) ?subsetIr. + pose U := S <*> 'Z(H). + have sUH: U \subset H by rewrite join_subG sSH subsetIl. + have cUU: abelian U by rewrite abelianY cSS center_abelian centsC. + have abelU: p.-abelem U by rewrite abelemE // cUU -eH exponentS. + have sUR: U \subset R := subset_trans sUH sHR. + have rU: 'r_p(U) <= 'r_p('C_R(S)). + by rewrite p_rankS //= subsetI sUR (centsS (joing_subl S 'Z(H))). + have nsUR: U <| R. + rewrite /normal sUR -commg_subl (subset_trans (commSg _ sUH)) //= -/U. + by rewrite (subset_trans sHRZ) // joing_subr. + have{rU}:= leq_trans rU rCRS; rewrite leq_eqVlt => /predU1P[] rU. + have Ep2U: [group of U] \in 'E_p^2(R). + by rewrite !inE /= sUR abelU -(p_rank_abelem abelU) rU. + have [F scn3F sUF] := normal_p2Elem_SCN3 rR Ep2U nsUR. + have [scnF rF] := setIdP scn3F; have [_ scF] := SCN_P scnF. + rewrite (leq_trans rF) // -scF -rank_pgroup ?(pgroupS (subsetIl _ _)) //. + by rewrite rankS ?setIS ?centS // (subset_trans _ sUF) ?joing_subl. + have defU: S :=: U. + apply/eqP; rewrite eqEcard oS joing_subl (card_pgroup (pgroupS sUR pR)). + by rewrite -p_rank_abelem // (leq_exp2l _ 1) // prime_gt1. + have ntS: S :!=: 1 by rewrite -cardG_gt1 oS prime_gt1. + have sSZ: S \subset 'Z(R) by rewrite prime_meetG ?oS ?meet_center_nil // defU. + by rewrite (setIidPl _) // centsC (subset_trans sSZ) ?subsetIr. +have{tiHS eH} oCHS: #|'C_H(S)| = p. + have ntCHS: 'C_H(S) != 1. + have: H :&: 'Z(R) != 1 by rewrite meet_center_nil ?char_normal. + by apply: subG1_contra; rewrite setIS // (centsS sSR) ?subsetIr. + have cycCHS: cyclic 'C_H(S). + have tiS_CHS: S :&: 'C_H(S) = 1 by rewrite setICA setIA tiHS setI1g. + rewrite (isog_cyclic (quotient_isog _ tiS_CHS)) ?subIset ?cent_sub ?orbT //. + rewrite (cyclicS _ (quotient_cyclic S cycC)) //= -(quotientMidl S C). + by rewrite mulSC quotientS // setSI // char_sub. + have abelCHS: p.-abelem 'C_H(S). + by rewrite abelemE ?cyclic_abelian // -eH exponentS ?subsetIl. + rewrite -(Ohm1_id abelCHS). + by rewrite (Ohm1_cyclic_pgroup_prime _ (abelem_pgroup abelCHS)). +pose B := A^`(1) <*> [set a ^+ p.-1 | a in A]. +have sBA: B \subset A. + rewrite join_subG (der_sub 1 A) /=. + by apply/subsetP=> _ /imsetP[a Aa ->]; rewrite groupX. +have AutB: B \subset Aut R := subset_trans sBA AutA. +suffices pB (X : {group {perm gT}}): X \subset B -> p^'.-group X -> X :=: 1. + have cAbAb: abelian (A / 'O_p(A)). + rewrite sub_der1_abelian // pcore_max ?der_normal //. + apply/pgroupP=> q q_pr; apply: contraLR => p'q; rewrite -p'natE //. + have [X sylX] := Sylow_exists q A^`(1); have [sXA' qX _] := and3P sylX. + rewrite -partn_eq1 ?cardG_gt0 // -(card_Hall sylX). + by rewrite (pB X) ?cards1 ?(pi_pgroup qX) ?(subset_trans sXA') ?joing_subl. + rewrite cAbAb -(nilpotent_pcoreC p (abelian_nil cAbAb)) trivg_pcore_quotient. + rewrite dprod1g pcore_pgroup; split=> //_ a Aa p'a. + rewrite order_dvdn -cycle_eq1 [<[_]>]pB ?(pgroupS (cycleX _ _) p'a) //. + by rewrite genS // sub1set inE orbC (mem_imset (expgn^~ _)). +move=> sXB p'X; have AutX := subset_trans sXB AutB. +pose toX := ([Aut R] \ AutX)%gact; pose CX := 'C_(H | toX)(X). +suffices sHCX: H \subset CX. + rewrite -(setIid X) coprime_TIg ?(pnat_coprime (pgroupS _ pCH)) //. + by rewrite subsetIidl gacent_ract setIid gacentC in sHCX. +elim: _.+1 {1 2 4 6}H (charH) (subxx H) (ltnSn #|H|) => // n IHn L charL sLH. +rewrite ltnS => leLn; have sLR := char_sub charL; pose K := [~: L, R]. +wlog ntL: / L :!=: 1 by case: eqP => [-> | _ -> //]; rewrite sub1G. +have charK: K \char R by rewrite charR ?char_refl. +have ltKL: K \proper L. + have nLR: R \subset 'N_R(L) by rewrite subsetIidl char_norm. + exact: nil_comm_properl nilR sLR ntL nLR. +have [sKL sKR] := (proper_sub ltKL, char_sub charK). +have [sKH pK] := (subset_trans sKL sLH, pgroupS sKR pR : p.-group K). +have nsKH: K <| H := normalS sKH sHR (char_normal charK). +have sKCX: K \subset CX by rewrite IHn ?(leq_trans (proper_card ltKL)) ?leLn. +have pL := pgroupS sLR pR; have nKL: L \subset 'N(K) := commg_norml _ _. +have{pS cSS} oLb: #|L / K| = p. + have [v defS] := cyclicP cycS; rewrite defS cycle_subG in sSR. + have ntLb: L / K != 1 by rewrite -subG1 quotient_sub1 ?proper_subn. + have [_ p_dv_Lb _] := pgroup_pdiv (quotient_pgroup _ pL) ntLb. + apply/eqP; rewrite eqn_leq {p_dv_Lb}(dvdn_leq _ p_dv_Lb) // andbT. + rewrite -divg_normal ?(normalS sKL sLH nsKH) // leq_divLR ?cardSg //= -/K. + rewrite -(card_lcoset K v) -(LagrangeI L 'C(S)) -indexgI /= -oCHS /K commGC. + rewrite {2}defS cent_cycle index_cent1 leq_mul ?subset_leq_card ?setSI //. + by apply/subsetP=> vx; case/imsetP=> x Lx ->; rewrite mem_lcoset mem_commg. +have cycLb: cyclic (L / K) by rewrite prime_cyclic ?oLb. +rewrite -(quotientSGK _ sKCX) // quotientGI // subsetI quotientS //= -/K. +have actsXK: [acts X, on K | toX] by rewrite acts_ract subxx acts_char. +rewrite ext_coprime_quotient_cent ?(pnat_coprime pK p'X) ?(pgroup_sol pK) //. +have actsAL : {acts A, on group L | [Aut R]} by exact: gacts_char. +have sAD: A \subset qact_dom <[actsAL]> [~: L, R]. + by rewrite qact_domE // acts_actby subxx (setIidPr sKL) acts_char. +suffices cLbX: X \subset 'C(L / K | <[actsAL]> / _). + rewrite gacentE ?qact_domE // subsetI quotientS //=. + apply/subsetP=> Ku LbKu; rewrite inE; apply/subsetP=> x Xx; rewrite inE. + have [Dx cLx] := setIdP (subsetP cLbX x Xx); have [Ax _] := setIdP Dx. + rewrite inE in cLx; have:= subsetP cLx Ku LbKu; rewrite inE /=. + have [u Nu Lu ->] := morphimP LbKu. + by rewrite !{1}qactE // ?actbyE // qact_domE ?(subsetP actsXK). +rewrite (subset_trans sXB) // astab_range -ker_actperm gen_subG. +rewrite -sub_morphim_pre; last by rewrite -gen_subG ?(subset_trans sBA). +rewrite morphimU subUset morphim_der // (sameP trivgP derG1P). +rewrite (abelianS _ (Aut_cyclic_abelian cycLb)); last first. + exact: subset_trans (morphim_sub _ _) (im_actperm_Aut _). +apply/subsetP=> _ /morphimP[_ _ /imsetP[x Ax ->] ->]. +have Dx := subsetP sAD x Ax; rewrite inE morphX //= -order_dvdn. +apply: dvdn_trans (order_dvdG (actperm_Aut _ Dx)) _. +by rewrite card_Aut_cyclic // oLb (@totient_pfactor p 1) ?muln1. +Qed. + +End OneGroup. + +(* This is B & G, Theorem 5.6, parts (a) and (c). We do not prove parts (b), *) +(* (d) and (e), as they are not used in the proof of the Odd Order Theorem. *) +Theorem narrow_der1_complement_max_pdiv gT p (G S : {group gT}) : + odd #|G| -> solvable G -> p.-Sylow(G) S -> p.-narrow S -> + (2 < 'r(S)) ==> p.-length_1 G -> + [/\ (*a*) p^'.-Hall(G^`(1)) 'O_p^'(G^`(1)) + & (*c*) forall q, q \in \pi(G / 'O_p^'(G)) -> q <= p]. +Proof. +move=> oddG solG sylS nnS; case: (leqP 'r(S) 2) => /= rS pl1G. + have rG: 'r_p(G) <= 2 by rewrite -(rank_Sylow sylS). + split=> [|q]; first by have [-> _ _] := rank2_der1_complement solG oddG rG. + exact: rank2_max_pdiv solG oddG rG. +rewrite /pHall pcore_sub pcore_pgroup pnatNK /=. +rewrite -(pcore_setI_normal p^' (der_normal 1 G)) // setIC indexgI /=. +wlog Gp'1: gT G S oddG nnS solG sylS rS pl1G / 'O_p^'(G) = 1. + set K := 'O_p^'(G); have [_ nKG] := andP (pcore_normal _ G : K <| G). + move/(_ _ (G / K) (S / K))%G; rewrite quotient_sol ?quotient_odd //. + have [[sSG pS _] p'K] := (and3P sylS, pcore_pgroup _ G : p^'.-group K). + have [nKS nKG'] := (subset_trans sSG nKG, subset_trans (der_sub 1 G) nKG). + have tiKS: K :&: S = 1 := coprime_TIg (p'nat_coprime p'K pS). + have isoS := isog_symr (quotient_isog nKS tiKS). + rewrite (isog_narrow p isoS) {isoS}(isog_rank isoS) quotient_pHall //. + rewrite plength1_quo // trivg_pcore_quotient indexg1 /= -quotient_der //. + by rewrite card_quotient //= -/K -(card_isog (quotient1_isog _)); exact. +rewrite Gp'1 indexg1 -(card_isog (quotient1_isog _)) -pgroupE. +have [sSG pS _] := and3P sylS; have oddS: odd #|S| := oddSg sSG oddG. +have ntS: S :!=: 1 by rewrite -rank_gt0 (leq_trans _ rS). +have [p_pr _ _] := pgroup_pdiv pS ntS; have p_gt1 := prime_gt1 p_pr. +have{pl1G} defS: 'O_p(G) = S. + by rewrite (eq_Hall_pcore _ sylS) -?plength1_pcore_Sylow. +have nSG: G \subset 'N(S) by rewrite -defS gFnorm. +pose fA := restrm nSG (conj_aut S); pose A := fA @* G. +have AutA: A \subset Aut S by rewrite [A]im_restrm Aut_conj_aut. +have [solA oddA]: solvable A /\ odd #|A| by rewrite morphim_sol ?morphim_odd. +have [/= _ cAbAb p'A_dv_p1] := Aut_narrow pS oddS nnS solA AutA oddA. +have{defS} pKfA: p.-group ('ker fA). + rewrite (pgroupS _ pS) //= ker_restrm ker_conj_aut. + by rewrite -defS -Fitting_eq_pcore ?cent_sub_Fitting. +split=> [|q]. + rewrite -(pmorphim_pgroup pKfA) ?der_sub // morphim_der //. + by rewrite (pgroupS (der1_min (char_norm _) cAbAb)) ?pcore_pgroup ?pcore_char. +rewrite mem_primes; case/and3P=> q_pr _; case/Cauchy=> // x Gx ox. +rewrite leq_eqVlt -implyNb; apply/implyP=> p'q; rewrite -(ltn_predK p_gt1) ltnS. +have ofAx: #[fA x] = q. + apply/prime_nt_dvdP=> //; last by rewrite -ox morph_order. + rewrite order_eq1; apply: contraNneq p'q => fAx1. + by apply: (pgroupP pKfA); rewrite // -ox order_dvdG //; exact/kerP. +have p'fAx: p^'.-elt (fA x) by rewrite /p_elt ofAx pnatE. +by rewrite -ofAx dvdn_leq ?p'A_dv_p1 ?mem_morphim // -(subnKC p_gt1). +Qed. + +End Five. diff --git a/mathcomp/odd_order/BGsection6.v b/mathcomp/odd_order/BGsection6.v new file mode 100644 index 0000000..234313c --- /dev/null +++ b/mathcomp/odd_order/BGsection6.v @@ -0,0 +1,315 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div fintype finset. +Require Import prime fingroup morphism automorphism quotient gproduct gfunctor. +Require Import cyclic center commutator pgroup nilpotent sylow abelian hall. +Require Import maximal. +Require Import BGsection1 BGappendixAB. + +(******************************************************************************) +(* This file covers most of B & G section 6. *) +(* Theorem 6.4 is not proved, since it is not needed for the revised proof of *) +(* the odd order theorem. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Section Six. + +Implicit Type gT : finGroupType. +Implicit Type p : nat. + +Section OneType. + +Variable gT : finGroupType. +Implicit Types G H K S U : {group gT}. + +(* This is B & G, Theorem A.4(b) and 6.1, or Gorenstein 6.5.2, the main Hall- *) +(* Higman style p-stability result used in the proof of the Odd Order Theorem *) +Theorem odd_p_abelian_constrained p G : + odd #|G| -> solvable G -> p.-abelian_constrained G. +Proof. +move/odd_p_stable=> stabG /solvable_p_constrained constrG. +exact: p_stable_abelian_constrained. +Qed. + +(* Auxiliary results from AppendixAB, necessary to exploit the results below. *) +Definition center_Puig_char := BGappendixAB.center_Puig_char. +Definition trivg_center_Puig_pgroup := BGappendixAB.trivg_center_Puig_pgroup. + +(* The two parts of B & G, Theorem 6.2 are established in BGappendixAB. *) +Theorem Puig_factorisation p G S : + odd #|G| -> solvable G -> p.-Sylow(G) S -> 'O_p^'(G) * 'N_G('Z('L(S))) = G. +Proof. exact: BGappendixAB.Puig_factorization. Qed. + +(* This is the main statement of B & G, Theorem 6.2. It is not used in the *) +(* actual proof. *) +Theorem Puig_center_p'core_normal p G S : + odd #|G| -> solvable G -> p.-Sylow(G) S -> 'O_p^'(G) * 'Z('L(S)) <| G. +Proof. +move=> oddG solG sylS; rewrite -{2}(Puig_factorisation _ _ sylS) //. +have sZL_G := subset_trans (char_sub (center_Puig_char S)) (pHall_sub sylS). +rewrite -!quotientK ?(subset_trans _ (gFnorm _ _)) ?subsetIl //. +by rewrite cosetpre_normal quotient_normal // normalSG. +Qed. + +(* This is the second part (special case) of B & G, Theorem 6.2. *) +Theorem Puig_center_normal p G S : + odd #|G| -> solvable G -> p.-Sylow(G) S -> 'O_p^'(G) = 1 -> 'Z('L(S)) <| G. +Proof. exact: BGappendixAB.Puig_center_normal. Qed. + +(* This is B & G, Lemma 6.3(a). *) +Lemma coprime_der1_sdprod K H G : + K ><| H = G -> coprime #|K| #|H| -> solvable K -> K \subset G^`(1) -> + [~: K, H] = K /\ 'C_K(H) \subset K^`(1). +Proof. +case/sdprodP=> _ defG nKH tiKH coKH solK sKG'. +set K' := K^`(1); have [sK'K nK'K] := andP (der_normal 1 K : K' <| K). +have nK'H: H \subset 'N(K') := char_norm_trans (der_char 1 K) nKH. +set R := [~: K, H]; have sRK: R \subset K by rewrite commg_subl. +have [nRK nRH] := joing_subP (commg_norm K H : K <*> H \subset 'N(R)). +have sKbK'H': K / R \subset (K / R)^`(1) * (H / R)^`(1). + have defGb: (K / R) \* (H / R) = G / R. + by rewrite -defG quotientMl ?cprodE // centsC quotient_cents2r. + have [_ -> _ /=] := cprodP (der_cprod 1 defGb). + by rewrite -quotient_der ?quotientS // -defG mul_subG. +have tiKbHb': K / R :&: (H / R)^`(1) = 1. + by rewrite coprime_TIg // (coprimegS (der_sub 1 _)) ?coprime_morph. +have{sKbK'H' tiKbHb'} derKb: (K / R)^`(1) = K / R. + by rewrite -{2}(setIidPr sKbK'H') -group_modl ?der_sub // setIC tiKbHb' mulg1. +have{derKb} Kb1: K / R = 1. + rewrite (contraNeq (sol_der1_proper _ (subxx (K / R)))) ?quotient_sol //. + by rewrite derKb properxx. +have{Kb1 sRK} defKH: [~: K, H] = K. + by apply/eqP; rewrite eqEsubset sRK -quotient_sub1 ?Kb1 //=. +split=> //; rewrite -quotient_sub1 ?subIset ?nK'K //= -/K'. +have cKaKa: abelian (K / K') := der_abelian 0 K. +rewrite coprime_quotient_cent ?quotient_norms ?coprime_morph //= -/K' -defKH. +by rewrite quotientR // coprime_abel_cent_TI ?quotient_norms ?coprime_morph. +Qed. + +(* This is B & G, Lemma 6.3(b). It is apparently not used later. *) +Lemma prime_nil_der1_factor G : + nilpotent G^`(1) -> prime #|G / G^`(1)| -> + Hall G G^`(1) /\ (forall H, G^`(1) ><| H = G -> G^`(1) = [~: G, H]). +Proof. +move=> nilG' /=; set G' := G^`(1); set p := #|G / G'| => p_pr. +have nsG'G: G' <| G := der_normal 1 G; have [sG'G nG'G] := andP nsG'G. +have nsG'p'G: 'O_p^'(G') <| G := char_normal_trans (pcore_char _ _) nsG'G. +have nG'p'G := normal_norm nsG'p'G; have solG' := nilpotent_sol nilG'. +have{nilG'} pGb: p.-group (G / 'O_p^'(G')). + rewrite /pgroup card_quotient -?(Lagrange_index sG'G (pcore_sub _ _)) //=. + rewrite pnat_mul // -card_quotient // pnat_id //= -pnatNK. + by case/and3P: (nilpotent_pcore_Hall p^' nilG'). +have{pGb} cycGb: cyclic (G / 'O_p^'(G')). + apply: (cyclic_nilpotent_quo_der1_cyclic (pgroup_nil pGb)). + rewrite -quotient_der // (isog_cyclic (third_isog _ _ _)) ?pcore_sub //. + by apply: prime_cyclic. +have defG': G' = 'O_p^'(G'). + by apply/eqP; rewrite eqEsubset pcore_sub der1_min ?cyclic_abelian. +have hallG': Hall G G'. + rewrite /Hall sG'G -?card_quotient // defG' //= -/p. + by rewrite (p'nat_coprime (pcore_pgroup _ _)) ?pnat_id. +split=> // H defG; have [_ mulG'H nG'H tiG'H] := sdprodP defG. +rewrite -mulG'H commMG ?commg_normr // -derg1 (derG1P _) ?mulg1 //. + by case/coprime_der1_sdprod: (defG); rewrite ?(coprime_sdprod_Hall_l defG). +rewrite (isog_abelian (quotient_isog nG'H tiG'H)) /= -/G'. +by rewrite -quotientMidl mulG'H der_abelian. +Qed. + +Section PprodSubCoprime. + +Variables K U H G : {group gT}. +Hypotheses (defG : K * U = G) (nsKG : K <| G). +Hypotheses (sHU : H \subset U) (coKH : coprime #|K| #|H|). +Let nKG : G \subset 'N(K). Proof. by case/andP: nsKG. Qed. +Let sKG : K \subset G. Proof. by case/mulG_sub: defG. Qed. +Let sUG : U \subset G. Proof. by case/mulG_sub: defG. Qed. +Let nKU : U \subset 'N(K). Proof. exact: subset_trans sUG nKG. Qed. +Let nKH : H \subset 'N(K). Proof. exact: subset_trans sHU nKU. Qed. + +(* This is B & G, Lemma 6.5(a); note that we do not assume solvability. *) +Lemma pprod_focal_coprime : H :&: G^`(1) = H :&: U^`(1). +Proof. +set G' := G^`(1); set U' := U^`(1). +have [sU'U nU'U] := andP (der_normal 1 U : U' <| U). +have{nU'U} nU_U': U :&: _ \subset 'N(U') by move=> A; rewrite subIset ?nU'U. +suffices sHG'U': H :&: G' \subset U'. + by rewrite -(setIidPl sHG'U') -setIA (setIidPr (dergS 1 sUG)). +rewrite -(setIidPr sHU) -setIA -quotient_sub1 // setICA setIC. +rewrite quotientGI ?subsetI ?sU'U ?dergS ?coprime_TIg //= -/G' -/U'. +have sUG'_UKb: (U :&: G') / U' \subset (U :&: K) / U'. + rewrite quotientSK // -normC ?group_modr ?setIS //. + by rewrite -quotientSK ?comm_subG ?quotient_der // -defG quotientMidl. +rewrite (coprimeSg sUG'_UKb) // -(card_isog (second_isog _)) //=. +rewrite setIA (setIidPl sU'U) coprime_morphl ?coprime_morphr //. +exact: coprimeSg (subsetIr U K) coKH. +Qed. + +Hypothesis solG : solvable G. + +(* This is B & G, Lemma 6.5(c). *) +Lemma pprod_trans_coprime g : + g \in G -> H :^ g \subset U -> + exists2 c, c \in 'C_K(H) & exists2 u, u \in U & g = c * u. +Proof. +rewrite -{1}defG => /mulsgP[k u Kk Uu defg] sHgU. +have [sK_KH sH_KH] := joing_sub (erefl (K <*> H)). +have hallH: \pi(H).-Hall(K <*> H :&: U) H. + rewrite (pHall_subl _ (subsetIl _ _)) ?subsetI ?sH_KH //. + rewrite /pHall sH_KH pgroup_pi /= joingC norm_joinEl // indexMg -indexgI. + by rewrite -coprime_pi' ?cardG_gt0 //= coprime_sym coprime_TIg ?indexg1. +have{sHgU} hallHk: \pi(H).-Hall(K <*> H :&: U) (H :^ k). + rewrite pHallE cardJg (card_Hall hallH) eqxx andbT subsetI andbC. + rewrite -(conjSg _ _ u) (conjGid Uu) -conjsgM -defg sHgU. + by rewrite sub_conjg conjGid // groupV (subsetP sK_KH). +have{hallH hallHk} [w KUw defHk]: exists2 w, w \in K :&: U & H :^ k = H :^ w. + have sKHU_G: K <*> H :&: U \subset G by rewrite setIC subIset ?sUG. + have [hw KHUhw ->] := Hall_trans (solvableS sKHU_G solG) hallHk hallH. + have: hw \in H * (K :&: U) by rewrite group_modl // -norm_joinEl // joingC. + by case/mulsgP=> h w Hh KUw ->; exists w; rewrite // conjsgM (conjGid Hh). +have{KUw} [Kw Uw] := setIP KUw. +exists (k * w^-1); last by exists (w * u); rewrite ?groupM // -mulgA mulKg. +by rewrite -coprime_norm_cent // !inE groupM ?groupV //= conjsgM defHk conjsgK. +Qed. + +(* This is B & G, Lemma 6.5(b). *) +Lemma pprod_norm_coprime_prod : 'C_K(H) * 'N_U(H) = 'N_G(H). +Proof. +apply/eqP; rewrite eqEsubset mul_subG ?setISS ?cent_sub //=. +apply/subsetP=> g /setIP[Gg /normP nHg]. +have [|c Cc [u Uu defg]] := pprod_trans_coprime Gg; first by rewrite nHg. +rewrite defg mem_mulg // !inE Uu -{2}nHg defg conjsgM conjSg (normP _) //=. +by case/setIP: Cc => _; exact: (subsetP (cent_sub H)). +Qed. + +End PprodSubCoprime. + +Section Plength1Prod. + +Variables (p : nat) (G S : {group gT}). +Hypotheses (sylS : p.-Sylow(G) S) (pl1G : p.-length_1 G). +Let K := 'O_p^'(G). +Let sSG : S \subset G. Proof. by case/andP: sylS. Qed. +Let nsKG : K <| G. Proof. exact: pcore_normal. Qed. +Let sKG : K \subset G. Proof. by case/andP: nsKG. Qed. +Let nKG : G \subset 'N(K). Proof. by case/andP: nsKG. Qed. +Let nKS : S \subset 'N(K). Proof. exact: subset_trans sSG nKG. Qed. +Let coKS : coprime #|K| #|S|. +Proof. exact: p'nat_coprime (pcore_pgroup _ G) (pHall_pgroup sylS). Qed. +Let sSN : S \subset 'N_G(S). Proof. by rewrite subsetI sSG normG. Qed. + +Let sylGbp : p.-Sylow(G / K) 'O_p(G / K). +Proof. by rewrite -plength1_pcore_quo_Sylow. Qed. + +(* This is B & G, Lemma 6.6(a1); note that we do not assume solvability. *) +Lemma plength1_Sylow_prod : K * S = 'O_{p^',p}(G). +Proof. +by rewrite -quotientK 1?(eq_Hall_pcore sylGbp) ?quotient_pHall //= /K -pseries1. +Qed. + +Let sylS_Gp'p : p.-Sylow('O_{p^',p}(G)) S. +Proof. +have [_ sSGp'p] := mulG_sub plength1_Sylow_prod. +exact: pHall_subl sSGp'p (pseries_sub _ _) sylS. +Qed. + +(* This is B & G, Lemma 6.6(a2); note that we do not assume solvability. *) +Lemma plength1_Frattini : K * 'N_G(S) = G. +Proof. +rewrite -{2}(Frattini_arg _ sylS_Gp'p) ?pseries_normal //= -plength1_Sylow_prod. +by rewrite -mulgA [S * _]mulSGid // subsetI sSG normG. +Qed. +Local Notation defG := plength1_Frattini. + +(* This is B & G, Lemma 6.6(b); note that we do not assume solvability. *) +Lemma plength1_Sylow_sub_der1 : S \subset G^`(1) -> S \subset ('N_G(S))^`(1). +Proof. +by move/setIidPl=> sSG'; apply/setIidPl; rewrite -(pprod_focal_coprime defG). +Qed. + +Hypothesis solG : solvable G. + +(* This is B & G, Lemma 6.6(c). *) +Lemma plength1_Sylow_trans (Y : {set gT}) g : + Y \subset S -> g \in G -> Y :^ g \subset S -> + exists2 c, c \in 'C_G(Y) & exists2 u, u \in 'N_G(S) & g = c * u. +Proof. +rewrite -gen_subG -(gen_subG (Y :^ g)) genJ => sYS Gg sYgS. +have coKY := coprimegS sYS coKS. +have [sYN sYgN] := (subset_trans sYS sSN, subset_trans sYgS sSN). +have [c Cc defg] := pprod_trans_coprime defG nsKG sYN coKY solG Gg sYgN. +by exists c => //; apply: subsetP Cc; rewrite cent_gen setSI. +Qed. + +(* This is B & G, Lemma 6.6(d). *) +Lemma plength1_Sylow_Jsub (Q : {group gT}) : + Q \subset G -> p.-group Q -> + exists2 x, x \in 'C_G(Q :&: S) & Q :^ x \subset S. +Proof. +move=> sQG pQ; have sQ_Gp'p: Q \subset 'O_{p^',p}(G). + rewrite -sub_quotient_pre /= pcore_mod1 ?(subset_trans sQG) //. + by rewrite (sub_Hall_pcore sylGbp) ?quotientS ?quotient_pgroup. +have [xy /= KSxy sQxyS] := Sylow_Jsub sylS_Gp'p sQ_Gp'p pQ. +rewrite -plength1_Sylow_prod in KSxy; have [x y Kx Sy def_xy] := mulsgP KSxy. +have{sQxyS} sQxS: Q :^ x \subset S. + by rewrite -(conjSg _ _ y) (conjGid Sy) -conjsgM -def_xy. +exists x; rewrite // inE (subsetP sKG) //; apply/centP=> z; case/setIP=> Qz Sz. +apply/commgP; rewrite -in_set1 -set1gE -(coprime_TIg coKS) inE. +rewrite groupMl ?groupV ?memJ_norm ?(subsetP nKS) ?Kx //=. +by rewrite commgEr groupMr // (subsetP sQxS) ?memJ_conjg ?groupV. +Qed. + +End Plength1Prod. + +End OneType. + +(* This is B & G, Theorem 6.7 *) +Theorem plength1_norm_pmaxElem gT p (G E L : {group gT}) : + E \in 'E*_p(G) -> odd p -> solvable G -> p.-length_1 G -> + L \subset G -> E \subset 'N(L) -> p^'.-group L -> + L \subset 'O_p^'(G). +Proof. +move=> maxE p_odd solG pl1G sLG nEL p'L. +case p_pr: (prime p); last first. + by rewrite pcore_pgroup_id // p'groupEpi mem_primes p_pr. +wlog Gp'1: gT G E L maxE solG pl1G sLG nEL p'L / 'O_p^'(G) = 1. + set K := 'O_p^'(G); have [sKG nKG] := andP (pcore_normal _ G : K <| G). + move/(_ _ (G / K) (E / K) (L / K))%G; rewrite morphim_sol ?plength1_quo //. + rewrite morphimS ?morphim_norms ?quotient_pgroup // trivg_pcore_quotient. + rewrite (quotient_sub1 (subset_trans sLG nKG)) => -> //. + have [EpE _] := pmaxElemP maxE; have{EpE} [sEG abelE] := pElemP EpE. + apply/pmaxElemP; rewrite inE quotient_abelem ?quotientS //. + split=> // Fb; case/pElemP=> sFbG abelFb; have [pFb _ _] := and3P abelFb. + have [S sylS sES] := Sylow_superset sEG (abelem_pgroup abelE). + have [sSG pS _] := and3P sylS; have nKS := subset_trans sSG nKG. + have: (E / K)%G \in 'E*_p(S / K). + have: E \in 'E*_p(S) by rewrite (subsetP (pmaxElemS p sSG)) // inE maxE inE. + have coKS: coprime #|K| #|S| := p'nat_coprime (pcore_pgroup _ _) pS. + have [injK imK] := isomP (quotient_isom nKS (coprime_TIg coKS)). + by rewrite -(injm_pmaxElem injK) ?imK ?inE //= morphim_restrm (setIidPr _). + case/pmaxElemP=> _; apply; rewrite inE abelFb andbT. + rewrite (sub_normal_Hall (quotient_pHall _ sylS)) //= -quotientMidl /= -/K. + by rewrite plength1_Sylow_prod // quotient_pseries2 pcore_normal. +have [EpE _] := pmaxElemP maxE; have{EpE} [sEG abelE] := pElemP EpE. +have [S sylS sES] := Sylow_superset sEG (abelem_pgroup abelE). +have [sSG pS _] := and3P sylS; have oddS: odd #|S| := odd_pgroup_odd p_odd pS. +have defS: S :=: 'O_p(G) by apply eq_Hall_pcore; rewrite -?plength1_pcore_Sylow. +have coSL: coprime #|S| #|L| := pnat_coprime pS p'L. +have tiSL: S :&: L = 1 := coprime_TIg coSL. +have{solG} scSG: 'C_G(S) \subset S. + by rewrite defS -Fitting_eq_pcore ?cent_sub_Fitting. +rewrite Gp'1 -tiSL subsetIidr (subset_trans _ scSG) // subsetI sLG /=. +have nSL: L \subset 'N(S) by rewrite (subset_trans sLG) // defS gFnorm. +have cLE: L \subset 'C(E). + by rewrite (sameP commG1P trivgP) -tiSL setIC commg_subI ?(introT subsetIP). +have maxES: E \in 'E*_p(S) by rewrite (subsetP (pmaxElemS p sSG)) ?(maxE, inE). +have EpE: E \in 'E_p(S) by apply/setIdP. +by rewrite (coprime_odd_faithful_cent_abelem EpE) ?(pmaxElem_LdivP p_pr maxES). +Qed. + +End Six. + diff --git a/mathcomp/odd_order/BGsection7.v b/mathcomp/odd_order/BGsection7.v new file mode 100644 index 0000000..9982283 --- /dev/null +++ b/mathcomp/odd_order/BGsection7.v @@ -0,0 +1,979 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype bigop. +Require Import finset prime fingroup morphism automorphism action quotient. +Require Import gfunctor cyclic pgroup center commutator gseries nilpotent. +Require Import sylow abelian maximal hall. +Require Import BGsection1 BGsection6. + +(******************************************************************************) +(* This file covers B & G, section 7, i.e., the proof of the Thompson *) +(* Transitivity Theorem, as well as some generalisations used later in the *) +(* proof. *) +(* This is the first section of the proof that applies to a (hypothetical) *) +(* minimally simple odd group, so we also introduce at this point some *) +(* infrastructure to carry over this assumption into the rest of the proof. *) +(* minSimpleOddGroupType == a finGroupType that ranges exactly over the *) +(* elements of a minimal counter-example to the *) +(* Odd Order Theorem. *) +(* G == the group of all the elements in a *) +(* minSimpleOddGroupType (this is a local notation *) +(* that must be reestablished for each such Type). *) +(* 'M == the set of all (proper) maximal subgroups of G *) +(* 'M(H) == the set of all elements of 'M that contain H *) +(* 'U == the set of all H such that 'M(H) contains a *) +(* single (unique) maximal subgroup of G. *) +(* 'SCN_n[p] == the set of all SCN subgroups of rank at least n *) +(* of all the Sylow p-subgroups of G. *) +(* |/|_H(A, pi) == the set of all pi-subgroups of H that are *) +(* normalised by A. *) +(* |/|*(A, pi) == the set of pi-subgroups of G, normalised by A, *) +(* and maximal subject to this condition. *) +(* normed_constrained A == A is a nontrivial proper subgroup of G, such *) +(* that for any proper subgroup X containing A, *) +(* all Y in |/|_X(A, pi') lie in the pi'-core of X *) +(* (here pi is the set of prime divisors of #|A|). *) +(* This is Hypothesis 7.1 in B & G. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Reserved Notation "''M'" (at level 8, format "''M'"). +Reserved Notation "''M' ( H )" (at level 8, format "''M' ( H )"). +Reserved Notation "''U'" (at level 8). +Reserved Notation "''SCN_' n [ p ]" + (at level 8, n at level 2, format "''SCN_' n [ p ]"). +Reserved Notation "|/|_ X ( A ; pi )" + (at level 8, X at level 2, format "|/|_ X ( A ; pi )"). +Reserved Notation "|/|* ( A ; pi )" + (at level 8, format "|/|* ( A ; pi )"). + +(* The generic setup for the whole Odd Order Theorem proof. *) +Section InitialReduction. + +Implicit Type gT : finGroupType. + +Record minSimpleOddGroupMixin gT : Prop := MinSimpleOddGroupMixin { + _ : odd #|[set: gT]|; + _ : simple [set: gT]; + _ : ~~ solvable [set: gT]; + _ : forall M : {group gT}, M \proper [set: gT] -> solvable M +}. + +Structure minSimpleOddGroupType := MinSimpleOddGroupType { + minSimpleOddGroupType_base :> finGroupType; + _ : minSimpleOddGroupMixin minSimpleOddGroupType_base +}. + +Hypothesis IH_FT : minSimpleOddGroupType -> False. + +Lemma minSimpleOdd_ind gT (G : {group gT}) : odd #|G| -> solvable G. +Proof. +move: {2}_.+1 (ltnSn #|G|) => n. +elim: n => // n IHn in gT G *; rewrite ltnS => leGn oddG. +have oG: #|[subg G]| = #|G| by rewrite (card_isog (isog_subg G)). +apply/idPn=> nsolG; case: IH_FT; exists [finGroupType of subg_of G]. +do [split; rewrite ?oG //=] => [||M]. +- rewrite -(isog_simple (isog_subg _)); apply/simpleP; split=> [|H nsHG]. + by apply: contra nsolG; move/eqP->; rewrite abelian_sol ?abelian1. + have [sHG _]:= andP nsHG; apply/pred2P; apply: contraR nsolG; case/norP=> ntH. + rewrite eqEcard sHG -ltnNge (series_sol nsHG) => ltHG. + by rewrite !IHn ?(oddSg sHG) ?quotient_odd ?(leq_trans _ leGn) ?ltn_quotient. +- by apply: contra nsolG => solG; rewrite -(im_sgval G) morphim_sol. +rewrite properEcard oG; case/andP=> sMG ltMG. +by apply: IHn (leq_trans ltMG leGn) (oddSg sMG _); rewrite oG. +Qed. + +Lemma minSimpleOdd_prime gT (G : {group gT}) : + odd #|G| -> simple G -> prime #|G|. +Proof. by move/minSimpleOdd_ind; apply: simple_sol_prime. Qed. + +End InitialReduction. + +Notation TheMinSimpleOddGroup gT := + [set: FinGroup.arg_sort (FinGroup.base (minSimpleOddGroupType_base gT))] + (only parsing). + +(* Elementary properties of the minimal counter example. *) +Section MinSimpleOdd. + +Variable gT : minSimpleOddGroupType. +Notation G := (TheMinSimpleOddGroup gT). +Implicit Types H K D M P U V X : {group gT}. +Local Notation sT := {set gT}. +Implicit Type p : nat. + +Lemma mFT_odd H : odd #|H|. +Proof. by apply: (oddSg (subsetT H)); case: gT => ? []. Qed. + +Lemma mFT_simple : simple G. +Proof. by case: gT => ? []. Qed. + +Lemma mFT_nonSolvable : ~~ solvable G. +Proof. by case: gT => ? []. Qed. + +Lemma mFT_sol M : M \proper G -> solvable M. +Proof. by case: gT M => ? []. Qed. + +Lemma mFT_nonAbelian : ~~ abelian G. +Proof. apply: contra mFT_nonSolvable; exact: abelian_sol. Qed. + +Lemma mFT_neq1 : G != 1. +Proof. by apply: contraNneq mFT_nonAbelian => ->; exact: abelian1. Qed. + +Lemma mFT_gt1 : [1] \proper G. Proof. by rewrite proper1G mFT_neq1. Qed. + +Lemma mFT_quo_odd M H : odd #|M / H|. +Proof. by rewrite quotient_odd ?mFT_odd. Qed. + +Lemma mFT_sol_proper M : (M \proper G) = solvable M. +Proof. +apply/idP/idP; first exact: mFT_sol. +by rewrite properT; apply: contraL; move/eqP->; exact: mFT_nonSolvable. +Qed. + +Lemma mFT_pgroup_proper p P : p.-group P -> P \proper G. +Proof. by move/pgroup_sol; rewrite mFT_sol_proper. Qed. + +Lemma mFT_norm_proper H : H :!=: 1 -> H \proper G -> 'N(H) \proper G. +Proof. +move=> ntH; rewrite !properT; apply: contra; move/eqP=> nHG; apply/eqP. +move/eqP: ntH; case/simpleP: mFT_simple => _; case/(_ H) => //=. +by rewrite -nHG normalG. +Qed. + +Lemma cent_mFT_trivial : 'C(G) = 1. +Proof. +apply/eqP; apply: contraR mFT_nonAbelian => ntC. +rewrite /abelian subTset /= eqEproper subsetT /=; apply/negP=> prC. +have:= mFT_norm_proper ntC prC. +by rewrite /proper subsetT norms_cent ?normG. +Qed. + +Lemma mFT_cent_proper H : H :!=: 1 -> 'C(H) \proper G. +Proof. +case: (eqsVneq H G) => [-> | ]. + by rewrite cent_mFT_trivial properT eq_sym. +rewrite -properT => prH ntH; apply: sub_proper_trans (cent_sub H) _. +exact: mFT_norm_proper. +Qed. + +Lemma mFT_cent1_proper x : x != 1 -> 'C[x] \proper G. +Proof. by rewrite -cycle_eq1 -cent_cycle; exact: mFT_cent_proper. Qed. + +Lemma mFT_quo_sol M H : H :!=: 1 -> solvable (M / H). +Proof. +move=> ntH; case: (eqsVneq H G) => [-> |]. + rewrite [_ / _](trivgP _) ?abelian_sol ?abelian1 //. + by rewrite quotient_sub1 ?normsG ?subsetT. +rewrite -properT => prH; rewrite -quotientInorm morphim_sol //. +by apply: solvableS (subsetIr _ _) (mFT_sol _); rewrite mFT_norm_proper. +Qed. + +(* Maximal groups of the minimal FT counterexample, as defined at the start *) +(* of B & G, section 7. *) +Definition minSimple_max_groups := [set M : {group gT} | maximal M G]. +Local Notation "'M" := minSimple_max_groups : group_scope. + +Definition minSimple_max_groups_of (H : sT) := [set M in 'M | H \subset M]. +Local Notation "''M' ( H )" := (minSimple_max_groups_of H) : group_scope. + +Definition minSimple_uniq_max_groups := [set U : {group gT} | #|'M(U)| == 1%N]. +Local Notation "'U" := minSimple_uniq_max_groups : group_scope. + +Definition minSimple_SCN_at n p := \bigcup_(P in 'Syl_p(G)) 'SCN_n(P). + +Lemma mmax_exists H : H \proper G -> {M | M \in 'M(H)}. +Proof. +case/(@maxgroup_exists _ (fun M => M \proper G)) => M maxM sHM. +by exists M; rewrite !inE sHM andbT. +Qed. + +Lemma any_mmax : {M : {group gT} | M \in 'M}. +Proof. by have [M] := mmax_exists mFT_gt1; case/setIdP; exists M. Qed. + +Lemma mmax_proper M : M \in 'M -> M \proper G. +Proof. by rewrite inE; apply: maxgroupp. Qed. + +Lemma mmax_sol M : M \in 'M -> solvable M. +Proof. by move/mmax_proper/mFT_sol. Qed. + +Lemma mmax_max M H : M \in 'M -> H \proper G -> M \subset H -> H :=: M. +Proof. by rewrite inE; case/maxgroupP=> _; apply. Qed. + +Lemma eq_mmax : {in 'M &, forall M H, M \subset H -> M :=: H}. +Proof. by move=> M H Mmax; move/mmax_proper=> prH; move/mmax_max->. Qed. + +Lemma sub_mmax_proper M H : M \in 'M -> H \subset M -> H \proper G. +Proof. by move=> maxM sHM; apply: sub_proper_trans (mmax_proper maxM). Qed. + +Lemma mmax_norm X M : + M \in 'M -> X :!=: 1 -> X \proper G -> M \subset 'N(X) -> 'N(X) = M. +Proof. by move=> maxM ntX prX; exact: mmax_max (mFT_norm_proper _ _). Qed. + +Lemma mmax_normal_subset A M : + M \in 'M -> A <| M -> ~~ (A \subset [1]) -> 'N(A) = M. +Proof. +rewrite -gen_subG subG1 => maxM /andP[sAM nAM] ntGA. +rewrite (mmax_max maxM) // (sub_proper_trans (norm_gen _)) ?mFT_norm_proper //. +by rewrite (sub_mmax_proper maxM) // gen_subG. +Qed. + +Lemma mmax_normal M H : M \in 'M -> H <| M -> H :!=: 1 -> 'N(H) = M. +Proof. by rewrite -subG1; apply: mmax_normal_subset. Qed. + +Lemma mmax_sigma_Sylow p P M : + M \in 'M -> p.-Sylow(M) P -> 'N(P) \subset M -> p.-Sylow(G) P. +Proof. +by move=> maxM sylP sNM; rewrite -Sylow_subnorm setTI (pHall_subl _ sNM) ?normG. +Qed. + +Lemma mmax_neq1 M : M \in 'M -> M :!=: 1. +Proof. +move=> maxM; apply: contra mFT_nonAbelian; move/eqP=> M1. +case: (eqVneq G 1) => [-> | ]; first exact: abelian1. +case/trivgPn=> x; rewrite -cycle_subG -cycle_eq1 subEproper /=. +case/predU1P=> [<- | ]; first by rewrite cycle_abelian. +by move/(mmax_max maxM)=> ->; rewrite M1 ?sub1G ?eqxx. +Qed. + +Lemma norm_mmax M : M \in 'M -> 'N(M) = M. +Proof. +move=> maxM; apply: mmax_max (normG M) => //. +exact: (mFT_norm_proper (mmax_neq1 maxM) (mmax_proper maxM)). +Qed. + +Lemma mmaxJ M x : (M :^ x \in 'M)%G = (M \in 'M). +Proof. by rewrite !inE /= -{1}[G](@conjGid _ _ x) ?maximalJ ?inE. Qed. + +Lemma mmax_ofS H K : H \subset K -> 'M(K) \subset 'M(H). +Proof. +move=> sHK; apply/subsetP=> M; rewrite !inE => /andP[->]. +exact: subset_trans. +Qed. + +Lemma mmax_ofJ K x M : ((M :^ x)%G \in 'M(K :^ x)) = (M \in 'M(K)). +Proof. by rewrite inE mmaxJ conjSg !inE. Qed. + +Lemma uniq_mmaxP U : reflect (exists M, 'M(U) = [set M]) (U \in 'U). +Proof. by rewrite inE; apply: cards1P. Qed. +Implicit Arguments uniq_mmaxP [U]. + +Lemma mem_uniq_mmax U M : 'M(U) = [set M] -> M \in 'M /\ U \subset M. +Proof. by move/setP/(_ M); rewrite set11 => /setIdP. Qed. + +Lemma eq_uniq_mmax U M H : + 'M(U) = [set M] -> H \in 'M -> U \subset H -> H :=: M. +Proof. +by move=> uU_M maxH sUH; apply/congr_group/set1P; rewrite -uU_M inE maxH. +Qed. + +Lemma def_uniq_mmax U M : + U \in 'U -> M \in 'M -> U \subset M -> 'M(U) = [set M]. +Proof. +case/uniq_mmaxP=> D uU_D maxM sUM. +by rewrite (group_inj (eq_uniq_mmax uU_D maxM sUM)). +Qed. + +Lemma uniq_mmax_subset1 U M : + M \in 'M -> U \subset M -> (U \in 'U) = ('M(U) \subset [set M]). +Proof. +move=> maxM sUM; apply/idP/idP=> uU; first by rewrite -(def_uniq_mmax uU). +by apply/uniq_mmaxP; exists M; apply/eqP; rewrite eqEsubset uU sub1set inE maxM. +Qed. + +Lemma sub_uniq_mmax U M H : + 'M(U) = [set M] -> U \subset H -> H \proper G -> H \subset M. +Proof. +move=> uU_M sUH; case/mmax_exists=> D; case/setIdP=> maxD sHD. +by rewrite -(eq_uniq_mmax uU_M maxD) ?(subset_trans sUH). +Qed. + +Lemma mmax_sup_id M : M \in 'M -> 'M(M) = [set M]. +Proof. +move=> maxM; apply/eqP; rewrite eqEsubset sub1set inE maxM subxx !andbT. +apply/subsetP=> H; case/setIdP=> maxH; rewrite inE -val_eqE /=. +by move/eq_mmax=> ->. +Qed. + +Lemma mmax_uniq_id : {subset 'M <= 'U}. +Proof. by move=> M maxM; apply/uniq_mmaxP; exists M; exact: mmax_sup_id. Qed. + +Lemma def_uniq_mmaxJ M K x : 'M(K) = [set M] -> 'M(K :^ x) = [set M :^ x]%G. +Proof. +move=> uK_M; apply/setP=> L; rewrite -(actKV 'JG x L) mmax_ofJ uK_M. +by rewrite !inE (inj_eq (act_inj 'JG x)). +Qed. + +Lemma uniq_mmaxJ K x :((K :^ x)%G \in 'U) = (K \in 'U). +Proof. +apply/uniq_mmaxP/uniq_mmaxP=> [] [M uK_M]. + exists (M :^ x^-1)%G; rewrite -(conjsgK x K); exact: def_uniq_mmaxJ. +by exists (M :^ x)%G; exact: def_uniq_mmaxJ. +Qed. + +Lemma uniq_mmax_norm_sub (M U : {group gT}) : + 'M(U) = [set M] -> 'N(U) \subset M. +Proof. +move=> uU_M; have [maxM _] := mem_uniq_mmax uU_M. +apply/subsetP=> x nUx; rewrite -(norm_mmax maxM) inE. +have:= set11 M; rewrite -uU_M -(mmax_ofJ _ x) (normP nUx) uU_M. +by move/set1P/congr_group->. +Qed. + +Lemma uniq_mmax_neq1 (U : {group gT}) : U \in 'U -> U :!=: 1. +Proof. +case/uniq_mmaxP=> M uU_M; have [maxM _] := mem_uniq_mmax uU_M. +apply: contraL (uniq_mmax_norm_sub uU_M); move/eqP->. +by rewrite norm1 subTset -properT mmax_proper. +Qed. + +Lemma def_uniq_mmaxS M U V : + U \subset V -> V \proper G -> 'M(U) = [set M] -> 'M(V) = [set M]. +Proof. +move=> sUV prV uU_M; apply/eqP; rewrite eqEsubset sub1set -uU_M. +rewrite mmax_ofS //= inE (sub_uniq_mmax uU_M) //. +by case/mem_uniq_mmax: uU_M => ->. +Qed. + +Lemma uniq_mmaxS U V : U \subset V -> V \proper G -> U \in 'U -> V \in 'U. +Proof. +move=> sUV prV /uniq_mmaxP[M uU_M]; apply/uniq_mmaxP; exists M. +exact: def_uniq_mmaxS uU_M. +Qed. + +End MinSimpleOdd. + +Implicit Arguments uniq_mmaxP [gT U]. +Prenex Implicits uniq_mmaxP. + +Notation "''M'" := (minSimple_max_groups _) : group_scope. +Notation "''M' ( H )" := (minSimple_max_groups_of H) : group_scope. +Notation "''U'" := (minSimple_uniq_max_groups _) : group_scope. +Notation "''SCN_' n [ p ]" := (minSimple_SCN_at _ n p) : group_scope. + +Section Hypothesis7_1. + +Variable gT : finGroupType. +Implicit Types X Y A P Q : {group gT}. +Local Notation G := [set: gT]. + +Definition normed_pgroups (X A : {set gT}) pi := + [set Y : {group gT} | pi.-subgroup(X) Y & A \subset 'N(Y)]. +Local Notation "|/|_ X ( A ; pi )" := (normed_pgroups X A pi) : group_scope. + +Definition max_normed_pgroups (A : {set gT}) pi := + [set Y : {group gT} | [max Y | pi.-group Y & A \subset 'N(Y)]]. +Local Notation "|/|* ( A ; pi )" := (max_normed_pgroups A pi) : group_scope. + +(* This is the statement for B & G, Hypothesis 7.1. *) +Inductive normed_constrained (A : {set gT}) := + NormedConstrained (pi := \pi(A)) of A != 1 & A \proper G + & forall X Y : {group gT}, + A \subset X -> X \proper G -> Y \in |/|_X(A; pi^') -> Y \subset 'O_pi^'(X). + +Variable pi : nat_pred. + +Lemma max_normed_exists A X : + pi.-group X -> A \subset 'N(X) -> {Y | Y \in |/|*(A; pi) & X \subset Y}. +Proof. +move=> piX nXA; pose piAn Y := pi.-group(Y) && (A \subset 'N(Y)). +have [|Y] := @maxgroup_exists _ piAn X; first by rewrite /piAn piX. +by exists Y; rewrite // inE. +Qed. + +Lemma mem_max_normed A X : X \in |/|*(A; pi) -> pi.-group X /\ A \subset 'N(X). +Proof. by rewrite inE; move/maxgroupp; move/andP. Qed. + +Lemma norm_acts_max_norm P : [acts 'N(P), on |/|*(P; pi) | 'JG]. +Proof. +apply/subsetP=> z Nz; rewrite !inE; apply/subsetP=> Q; rewrite !inE. +case/maxgroupP=> qQ maxQ; apply/maxgroupP; rewrite pgroupJ norm_conj_norm //. +split=> // Y; rewrite sub_conjg /= => qY; move/maxQ=> <-; rewrite ?conjsgKV //. +by rewrite pgroupJ norm_conj_norm ?groupV. +Qed. + +Lemma trivg_max_norm P : 1%G \in |/|*(P; pi) -> |/|*(P; pi) = [set 1%G]. +Proof. +move=> max1; apply/eqP; rewrite eqEsubset sub1set max1 andbT. +apply/subsetP=> Q; rewrite !inE -val_eqE /= in max1 *. +by case/maxgroupP: max1 => _ max1; move/maxgroupp; move/max1->; rewrite ?sub1G. +Qed. + +Lemma max_normed_uniq A P Q : + |/|*(A; pi) = [set Q] -> A \subset P -> P \subset 'N(Q) -> + |/|*(P; pi) = [set Q]. +Proof. +move=> defAmax sAP nQP; have: Q \in |/|*(A; pi) by rewrite defAmax set11. +rewrite inE; case/maxgroupP; case/andP=> piQ _ maxQ. +apply/setP=> X; rewrite !inE -val_eqE /=; apply/maxgroupP/eqP=> [[]|->{X}]. + case/andP=> piX nXP maxX; have nXA := subset_trans sAP nXP. + have [Y] := max_normed_exists piX nXA. + by rewrite defAmax; move/set1P->; move/maxX=> -> //; rewrite piQ. +rewrite piQ; split=> // X; case/andP=> piX nXP sQX. +by rewrite (maxQ X) // piX (subset_trans sAP). +Qed. + +End Hypothesis7_1. + +Notation "|/|_ X ( A ; pi )" := (normed_pgroups X A pi) : group_scope. +Notation "|/|* ( A ; pi )" := (max_normed_pgroups A pi) : group_scope. + +Section Seven. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Local Notation grT := {group gT}. +Implicit Types H P Q R K M A B : grT. +Implicit Type p q : nat. + +Section NormedConstrained. + +Variables (q : nat) (A : grT). +Let pi := Eval simpl in \pi(A). +Let K := 'O_pi^'('C(A)). +Let nsKC : K <| 'C(A) := pcore_normal _ _. + +Lemma cent_core_acts_max_norm : [acts K, on |/|*(A; q) | 'JG]. +Proof. +by rewrite (subset_trans _ (norm_acts_max_norm _ _)) ?cents_norm ?pcore_sub. +Qed. +Let actsKmax := actsP cent_core_acts_max_norm. + +Hypotheses (cstrA : normed_constrained A) (pi'q : q \notin pi). + +Let hyp71 H R : + A \subset H -> H \proper G -> R \in |/|_H(A; pi^') -> R \subset 'O_pi^'(H). +Proof. by case: cstrA H R. Qed. + +(* This is the observation between B & G, Hypothesis 7.1 and Lemma 7.1. *) +Remark normed_constrained_Hall : pi^'.-Hall('C(A)) K. +Proof. +have [_ ntA prA _] := cstrA; rewrite -[setT]/G in prA. +rewrite /pHall pcore_pgroup pcore_sub pnatNK /=. +rewrite -card_quotient ?gFnorm //= -/K. +apply/pgroupP=> p p_pr; case/Cauchy=> // Kx; case/morphimP=> x Nx Cx ->{Kx}. +rewrite /order -quotient_cycle //= -/K => def_p; apply/idPn=> pi'p. +have [P sylP] := Sylow_exists p <[x]>; have [sPx pP _]:= and3P sylP. +suffices: P \subset K. + have nKP: P \subset 'N(K) by rewrite (subset_trans sPx) ?cycle_subG. + rewrite -quotient_sub1 //= -/K (sameP trivgP eqP) trivg_card1. + rewrite (card_Hall (morphim_pHall _ nKP sylP)) def_p part_pnat_id ?pnat_id //. + by case: eqP p_pr => // ->. +suffices sP_pAC: P \subset 'O_pi^'(A <*> 'C(A)). + rewrite (subset_trans sP_pAC) ?pcore_max ?pcore_pgroup //. + rewrite /normal (char_norm_trans (pcore_char _ _)) ?normsG ?joing_subr //. + rewrite andbT -quotient_sub1; last first. + rewrite (subset_trans (pcore_sub _ _)) // join_subG normG cents_norm //. + by rewrite centsC. + rewrite /= -(setIidPr (pcore_sub _ _)) quotientGI ?joing_subr //=. + rewrite {1}cent_joinEr // quotientMidr coprime_TIg // coprime_morph //. + by rewrite coprime_pi' ?cardG_gt0 //= -/pi [pnat _ _]pcore_pgroup. +apply: hyp71; first exact: joing_subl. + apply: sub_proper_trans (mFT_norm_proper ntA prA). + by rewrite join_subG normG cent_sub. +have sPC: P \subset 'C(A) by rewrite (subset_trans sPx) ?cycle_subG. +rewrite inE /psubgroup cents_norm 1?centsC // andbT. +rewrite (subset_trans sPC) ?joing_subr //=. +by apply: sub_in_pnat pP => p' _; move/eqnP->. +Qed. +Let hallK := normed_constrained_Hall. + +(* This is B & G, Lemma 7.1. *) +Lemma normed_constrained_meet_trans Q1 Q2 H : + A \subset H -> H \proper G -> Q1 \in |/|*(A; q) -> Q2 \in |/|*(A; q) -> + Q1 :&: H != 1 -> Q2 :&: H != 1 -> + exists2 k, k \in K & Q2 :=: Q1 :^ k. +Proof. +move: {2}_.+1 (ltnSn (#|G| - #|Q1 :&: Q2|)) => m. +elim: m => // m IHm in H Q1 Q2 * => geQ12m sAH prHG maxQ1 maxQ2 ntHQ1 ntHQ2. +have:= maxQ1; rewrite inE => /maxgroupP[/andP[qQ1 nQ1A] maxQ1P]. +have:= maxQ2; rewrite inE => /maxgroupP[/andP[qQ2 nQ2A] maxQ2P]. +have prQ12: Q1 :&: Q2 \proper G. + rewrite properT; apply: contraNneq (mFT_nonSolvable gT) => <-. + by apply: pgroup_sol (pgroupS _ qQ1); rewrite subsetIl. +wlog defH: H prHG sAH ntHQ1 ntHQ2 / Q1 :&: Q2 != 1 -> H :=: 'N(Q1 :&: Q2). + case: (eqVneq (Q1 :&: Q2) 1) => [-> | ntQ12] IH. + by apply: (IH H) => //; case/eqP. + apply: (IH 'N(Q1 :&: Q2)%G); rewrite ?normsI ?mFT_norm_proper //; + apply: contra ntQ12; rewrite -!subG1; apply: subset_trans; + by rewrite subsetI normG (subsetIl, subsetIr). +pose L := 'O_pi^'(H); have sLH: L \subset H := pcore_sub _ _. +have [nLA coLA solL]: [/\ A \subset 'N(L), coprime #|L| #|A| & solvable L]. +- rewrite (char_norm_trans (pcore_char _ _)) ?normsG //. + rewrite coprime_sym coprime_pi' ?cardG_gt0 ?[pnat _ _]pcore_pgroup //. + by rewrite (solvableS sLH) ?mFT_sol. +have Qsyl Q: Q \in |/|*(A; q) -> Q :&: H != 1 -> + exists R : {group _}, [/\ q.-Sylow(L) R, A \subset 'N(R) & Q :&: H \subset R]. +- case/mem_max_normed=> qQ nQA ntQH. + have qQH: q.-group (Q :&: H) by rewrite (pgroupS _ qQ) ?subsetIl. + have nQHA: A \subset 'N(Q :&: H) by rewrite normsI // normsG. + apply: coprime_Hall_subset => //; apply: (hyp71) => //. + rewrite inE nQHA /psubgroup subsetIr andbT. + by apply: sub_in_pnat qQH => p _; move/eqnP->. +have [R1 [sylR1 nR1A sQR1]] := Qsyl _ maxQ1 ntHQ1. +have [R2 [sylR2 nR2A sQR2]] := Qsyl _ maxQ2 ntHQ2. +have [h Ch defR2] := coprime_Hall_trans nLA coLA solL sylR2 nR2A sylR1 nR1A. +have{Ch} [Hh Kh]: h \in H /\ h \in K. + case/setIP: Ch => Lh Ch; rewrite (subsetP sLH) //. + rewrite (mem_normal_Hall hallK (pcore_normal _ _)) //. + by rewrite (mem_p_elt _ Lh) ?pcore_pgroup. +have [Q3 maxQ3 sR2Q3] := max_normed_exists (pHall_pgroup sylR2) nR2A. +have maxQ1h: (Q1 :^ h)%G \in |/|*(A; q) by rewrite actsKmax. +case: (eqsVneq Q1 Q2) => [| neQ12]; first by exists 1; rewrite ?group1 ?conjsg1. +have ntHQ3: Q3 :&: H != 1. + apply: contra ntHQ2; rewrite -!subG1; apply: subset_trans. + by rewrite subsetI subsetIr (subset_trans sQR2). +have ntHQ1h: (Q1 :^ h) :&: H != 1. + by move: ntHQ1; rewrite !trivg_card1 -(cardJg _ h) conjIg (conjGid Hh). +suff [prI1 prI2]: Q1 :&: Q2 \proper Q1 :&: R1 /\ Q1 :&: Q2 \proper Q2 :&: R2. + have: #|G| - #|(Q1 :^ h) :&: Q3| < m. + rewrite ltnS in geQ12m; apply: leq_trans geQ12m. + rewrite ltn_sub2l ?(proper_card prQ12) // -(cardJg _ h) proper_card //. + by rewrite (proper_sub_trans _ (setIS _ sR2Q3)) // defR2 -conjIg properJ. + have: #|G| - #|Q3 :&: Q2| < m. + rewrite ltnS in geQ12m; apply: leq_trans geQ12m. + rewrite ltn_sub2l ?proper_card // (proper_sub_trans prI2) //. + by rewrite setIC setISS. + case/(IHm H) => // k2 Kk2 defQ2; case/(IHm H) => // k3 Kk3 defQ3. + by exists (h * k3 * k2); rewrite ?groupM ?conjsgM // -defQ3. +case: (eqVneq (Q1 :&: Q2) 1) => [-> | ntQ12]. + rewrite !proper1G; split; [apply: contra ntHQ1 | apply: contra ntHQ2]; + by rewrite -!subG1; apply: subset_trans; rewrite subsetI subsetIl. +rewrite -(setIidPr (subset_trans (pHall_sub sylR1) sLH)) setIA. +rewrite -(setIidPr (subset_trans (pHall_sub sylR2) sLH)) setIA. +rewrite (setIidPl sQR1) (setIidPl sQR2) {}defH //. +have nilQ1 := pgroup_nil qQ1; have nilQ2 := pgroup_nil qQ2. +rewrite !nilpotent_proper_norm /proper ?subsetIl ?subsetIr ?subsetI ?subxx //=. + by rewrite andbT; apply: contra neQ12 => sQ21; rewrite (maxQ2P Q1) ?qQ1. +by apply: contra neQ12 => sQ12; rewrite (maxQ1P Q2) ?qQ2. +Qed. + +(* This is B & G, Theorem 7.2. *) +Theorem normed_constrained_rank3_trans : + 'r('Z(A)) >= 3 -> [transitive K, on |/|*(A; q) | 'JG]. +Proof. +case/rank_geP=> B /nElemP[p]; rewrite !inE subsetI -2!andbA. +case/and4P=> sBA cAB abelB mB3; have [_ cBB _] := and3P abelB. +have q'B: forall Q, q.-group Q -> coprime #|Q| #|B|. + move=> Q qQ; rewrite coprime_sym (coprimeSg sBA) ?coprime_pi' //. + exact: pi_pnat qQ _. +have [Q1 maxQ1 _] := max_normed_exists (pgroup1 _ q) (norms1 A). +apply/imsetP; exists Q1 => //; apply/setP=> Q2. +apply/idP/imsetP=> [maxQ2|[k Kk] ->]; last by rewrite actsKmax. +have [qQ1 nQ1A]:= mem_max_normed maxQ1; have [qQ2 nQ2A]:= mem_max_normed maxQ2. +case: (eqVneq Q1 1%G) => [trQ1 | ntQ1]. + exists 1; rewrite ?group1 // act1; apply/eqP. + by rewrite trivg_max_norm -trQ1 // inE in maxQ2. +case: (eqVneq Q2 1%G) => [trQ2 | ntQ2]. + by case/negP: ntQ1; rewrite trivg_max_norm -trQ2 // inE in maxQ1 *. +have: [exists (C : grT | 'C_Q1(C) != 1), cyclic (B / C) && (C <| B)]. + apply: contraR ntQ1 => trQ1; have: B \subset 'N(Q1) := subset_trans sBA nQ1A. + rewrite -val_eqE -subG1 /=; move/coprime_abelian_gen_cent <-; rewrite ?q'B //. + rewrite gen_subG; apply/bigcupsP=> C cocyC; rewrite subG1. + by apply: contraR trQ1 => ntCC; apply/existsP; exists C; rewrite ntCC. +case/existsP=> C /and3P[ntCQ1 cycBC nsCB]; have [sCB nCB]:= andP nsCB. +have{mB3} ncycC: ~~ cyclic C. + rewrite (abelem_cyclic (quotient_abelem _ abelB)) ?card_quotient // in cycBC. + rewrite -divgS // logn_div ?cardSg // leq_subLR addn1 (eqP mB3) in cycBC. + by rewrite (abelem_cyclic (abelemS sCB abelB)) -ltnNge. +have: [exists (z | 'C_Q2[z] != 1), z \in C^#]. + apply: contraR ntQ2 => trQ2; have:= subset_trans sCB (subset_trans sBA nQ2A). + rewrite -[_ == _]subG1 /=. + move/coprime_abelian_gen_cent1 <-; rewrite ?(abelianS sCB) //; last first. + by rewrite (coprimegS sCB) ?q'B. + rewrite gen_subG; apply/bigcupsP=> z Cz. + by apply: contraR trQ2 => ntCz; apply/existsP; exists z; rewrite -subG1 ntCz. +case/existsP=> z; rewrite !inE => /and3P[ntzQ2 ntz Cz]. +have prCz: 'C[z] \proper G by rewrite -cent_cycle mFT_cent_proper ?cycle_eq1. +have sACz: A \subset 'C[z] by rewrite sub_cent1 (subsetP cAB) ?(subsetP sCB). +have [|//|k Kk defQ2]:= normed_constrained_meet_trans sACz prCz maxQ1 maxQ2. + apply: contra ntCQ1; rewrite -!subG1; apply: subset_trans. + by rewrite setIS //= -cent_cycle centS ?cycle_subG. +exists k => //; exact: val_inj. +Qed. + +(* This is B & G, Theorem 7.3. *) +Theorem normed_constrained_rank2_trans : + q %| #|'C(A)| -> 'r('Z(A)) >= 2 -> [transitive K, on |/|*(A; q) | 'JG]. +Proof. +move=> qC; case/rank_geP=> B; case/nElemP=> p; do 2![case/setIdP]. +rewrite subsetI; case/andP=> sBA cAB abelB mB2; have [_ cBB _] := and3P abelB. +have{abelB mB2} ncycB: ~~ cyclic B by rewrite (abelem_cyclic abelB) (eqP mB2). +have [R0 sylR0] := Sylow_exists q 'C(A); have [cAR0 qR0 _] := and3P sylR0. +have nR0A: A \subset 'N(R0) by rewrite cents_norm // centsC. +have{nR0A} [R maxR sR0R] := max_normed_exists qR0 nR0A. +apply/imsetP; exists R => //; apply/setP=> Q. +apply/idP/imsetP=> [maxQ|[k Kk] ->]; last by rewrite actsKmax. +have [qR nRA]:= mem_max_normed maxR; have [qQ nQA]:= mem_max_normed maxQ. +have [R1 | ntR] := eqVneq R 1%G. + rewrite trivg_max_norm -R1 // in maxQ. + by exists 1; rewrite ?group1 ?act1 ?(set1P maxQ). +have ntQ: Q != 1%G. + by apply: contra ntR => Q1; rewrite trivg_max_norm -(eqP Q1) // inE in maxR *. +have ntRC: 'C_R(A) != 1. + have sR0CR: R0 \subset 'C_R(A) by rewrite subsetI sR0R. + suffices: R0 :!=: 1 by rewrite -!proper1G; move/proper_sub_trans->. + move: ntR; rewrite -!cardG_gt1 -(part_pnat_id qR) (card_Hall sylR0). + by rewrite !p_part_gt1 !mem_primes !cardG_gt0 qC; case/and3P=> ->. +have: [exists (z | 'C_Q[z] != 1), z \in B^#]. + apply: contraR ntQ => trQ; have:= subset_trans sBA nQA. + rewrite -[_ == _]subG1; move/coprime_abelian_gen_cent1 <- => //; last first. + by rewrite coprime_sym (coprimeSg sBA) ?coprime_pi' /pgroup ?(pi_pnat qQ). + rewrite gen_subG; apply/bigcupsP=> z Cz; rewrite subG1. + by apply: contraR trQ => ntCz; apply/existsP; exists z; rewrite ntCz. +case/existsP=> z; rewrite 2!inE => /and3P[ntzQ ntz Bz]. +have prCz: 'C[z] \proper G by rewrite -cent_cycle mFT_cent_proper ?cycle_eq1. +have sACz: A \subset 'C[z] by rewrite sub_cent1 (subsetP cAB). +have [|//|k Kk defQ2]:= normed_constrained_meet_trans sACz prCz maxR maxQ. + apply: contra ntRC; rewrite -!subG1; apply: subset_trans. + by rewrite setIS //= -cent_cycle centS // cycle_subG (subsetP sBA). +exists k => //; exact: val_inj. +Qed. + +(* This is B & G, Theorem 7.4. *) +Theorem normed_trans_superset P : + A <|<| P -> pi.-group P -> [transitive K, on |/|*(A; q) | 'JG] -> + [/\ 'C_K(P) = 'O_pi^'('C(P)), + [transitive 'O_pi^'('C(P)), on |/|*(P; q) | 'JG], + |/|*(P; q) \subset |/|*(A; q) + & {in |/|*(P; q), forall Q, P :&: 'N(P)^`(1) \subset 'N(Q)^`(1) + /\ 'N(P) = 'C_K(P) * 'N_('N(P))(Q)}]. +Proof. +move=> snAP piP trnK; set KP := 'O_pi^'('C(P)). +have defK: forall B, A \subset B -> 'C_K(B) = 'O_pi^'('C(B)). + move=> B sAB; apply/eqP; rewrite eqEsubset {1}setIC pcoreS ?centS //. + rewrite subsetI pcore_sub (sub_Hall_pcore hallK) ?pcore_pgroup //. + by rewrite (subset_trans (pcore_sub _ _)) ?centS. +suffices: [transitive KP, on |/|*(P; q) | 'JG] /\ |/|*(P; q) \subset |/|*(A; q). + have nsKPN: KP <| 'N(P) := char_normal_trans (pcore_char _ _) (cent_normal _). + case=> trKP smnPA; rewrite (defK _ (subnormal_sub snAP)); split=> // Q maxQ. + have defNP: KP * 'N_('N(P))(Q) = 'N(P). + rewrite -(astab1JG Q) -normC; last by rewrite subIset 1?normal_norm. + apply/(subgroup_transitiveP maxQ); rewrite ?normal_sub //=. + by rewrite (atrans_supgroup _ trKP) ?norm_acts_max_norm ?normal_sub. + split=> //; move/pprod_focal_coprime: defNP => -> //. + - by rewrite subIset // orbC commgSS ?subsetIr. + - by rewrite subsetI normG; case/mem_max_normed: maxQ. + by rewrite (p'nat_coprime (pcore_pgroup _ _)). +elim: {P}_.+1 {-2}P (ltnSn #|P|) => // m IHm P lePm in KP piP snAP *. +wlog{snAP} [B maxnB snAB]: / {B : grT | maxnormal B P P & A <|<| B}. + case/subnormalEr: snAP => [|[D [snAD nDP prDP]]]; first by rewrite /KP => <-. + have [B maxnB sDB]: {B : grT | maxnormal B P P & D \subset B}. + by apply: maxgroup_exists; rewrite prDP normal_norm. + apply; exists B => //; apply: subnormal_trans snAD (normal_subnormal _). + by apply: normalS sDB _ nDP; case/andP: (maxgroupp maxnB); case/andP. +have [prBP nBP] := andP (maxgroupp maxnB); have sBP := proper_sub prBP. +have{lePm}: #|B| < m by exact: leq_trans (proper_card prBP) _. +case/IHm=> {IHm}// [|trnB smnBA]; first by rewrite (pgroupS sBP). +have{maxnB} abelPB: is_abelem (P / B). + apply: charsimple_solvable (maxnormal_charsimple _ maxnB) _ => //. + have [_ ntA _ _] := cstrA; have sAB := subnormal_sub snAB. + by apply: mFT_quo_sol; apply: contraL sAB; move/eqP->; rewrite subG1. +have{abelPB} [p p_pr pPB]: exists2 p, prime p & p.-group (P / B). + by case/is_abelemP: abelPB => p p_pr; case/andP; exists p. +have{prBP} pi_p: p \in pi. + case/pgroup_pdiv: pPB => [|_ pPB _]. + by rewrite -subG1 quotient_sub1 // proper_subn. + by apply: pgroupP p_pr pPB; exact: quotient_pgroup. +pose S := |/|*(B; q); have p'S: #|S| %% p != 0. + have pi'S: pi^'.-nat #|S| := pnat_dvd (atrans_dvd trnB) (pcore_pgroup _ _). + by rewrite -prime_coprime // (pnat_coprime _ pi'S) ?pnatE. +have{p'S} [Q S_Q nQP]: exists2 Q, Q \in S & P \subset 'N(Q). + have sTSB: setT \subset G / B by rewrite -im_quotient quotientS ?subsetT. + have modBE: {in P & S, forall x Q, ('JG %% B) Q (coset B x) = 'JG Q x}%act. + move=> x Q Px; rewrite inE; move/maxgroupp; case/andP=> _ nQB. + by rewrite /= modactE ?(subsetP nBP) ?afixJG ?setTI ?inE. + have actsPB: [acts P / B, on S | 'JG %% B \ sTSB]. + apply/subsetP=> _ /morphimP[x Nx Px ->]. + rewrite !inE; apply/subsetP=> Q S_Q; rewrite inE /= modBE //. + by rewrite (actsP (norm_acts_max_norm q B)). + move: p'S; rewrite (pgroup_fix_mod pPB actsPB); set nQ := #|_|. + case: (posnP nQ) => [->|]; first by rewrite mod0n. + rewrite lt0n; case/existsP=> Q /setIP[Q_S fixQ]; exists Q => //. + apply/normsP=> x Px; apply: congr_group; have Nx := subsetP nBP x Px. + by have:= afixP fixQ (coset B x); rewrite /= modBE ?mem_morphim //= => ->. +have [qQ _]:= mem_max_normed S_Q. +have{qQ nQP} [Q0 maxQ0 sQQ0] := max_normed_exists qQ nQP. +have [_ nQ0P]:= mem_max_normed maxQ0. +have actsKmnP: [acts 'O_pi^'('C(P)), on |/|*(P; q) | 'JG]. + by rewrite (subset_trans _ (norm_acts_max_norm q P)) // cents_norm ?pcore_sub. +case nt_mnP: (1%G \in |/|*(P; q)) => [|{Q S_Q sQQ0}]. + rewrite atrans_acts_card actsKmnP trivg_max_norm // imset_set1 in maxQ0 *. + have <-: Q = 1%G by apply/trivGP; rewrite -(congr_group (set1P maxQ0)). + by rewrite cards1 sub1set (subsetP smnBA). +have sAB := subnormal_sub snAB; have sAP := subset_trans sAB sBP. +have smnP_S: |/|*(P; q) \subset S. + apply/subsetP=> Q1 maxQ1; have [qQ1 nQ1P] := mem_max_normed maxQ1. + have ntQ1: Q1 != 1%G by case: eqP nt_mnP maxQ1 => // -> ->. + have prNQ1: 'N(Q1) \proper G := mFT_norm_proper ntQ1 (mFT_pgroup_proper qQ1). + have nQ1A: A \subset 'N(Q1) := subset_trans sAP nQ1P. + have [Q2 maxQ2 sQ12] := max_normed_exists qQ1 (subset_trans sBP nQ1P). + have [qQ2 nQ2B] := mem_max_normed maxQ2; apply: etrans maxQ2; congr in_mem. + apply: val_inj; suffices: q.-Sylow(Q2) Q1 by move/pHall_id=> /= ->. + have qNQ2: q.-group 'N_Q2(Q1) by rewrite (pgroupS _ qQ2) ?subsetIl. + pose KN := 'O_pi^'('N(Q1)); have sNQ2_KN: 'N_Q2(Q1) \subset KN. + rewrite hyp71 // inE normsI ?norms_norm ?(subset_trans sAB nQ2B) //=. + by rewrite /psubgroup subsetIr andbT; exact: pi_pnat qNQ2 _. + rewrite -Sylow_subnorm (pHall_subl _ sNQ2_KN) ?subsetI ?sQ12 ?normG //= -/KN. + suff: exists Q3 : grT, [/\ q.-Sylow(KN) Q3, P \subset 'N(Q3) & Q1 \subset Q3]. + move: maxQ1; rewrite inE; case/maxgroupP=> _ maxQ1 [Q3 [sylQ3 nQ3P sQ13]]. + by rewrite -(maxQ1 Q3) // (pHall_pgroup sylQ3). + apply: coprime_Hall_subset; rewrite //= -/KN. + - by rewrite (char_norm_trans (pcore_char _ _)) ?norms_norm. + - by rewrite coprime_sym (pnat_coprime piP (pcore_pgroup _ _)). + - by rewrite (solvableS (pcore_sub _ _)) ?mFT_sol. + by rewrite pcore_max ?normalG // /pgroup (pi_pnat qQ1). +split; last exact: subset_trans smnP_S smnBA. +apply/imsetP; exists Q0 => //; apply/setP=> Q2. +apply/idP/imsetP=> [maxQ2 | [k Pk ->]]; last by rewrite (actsP actsKmnP). +have [S_Q0 S_Q2]: Q0 \in S /\ Q2 \in S by rewrite !(subsetP smnP_S). +pose KB := 'O_pi^'('C(B)); pose KBP := KB <*> P. +have pi'KB: pi^'.-group KB by exact: pcore_pgroup. +have nKB_P: P \subset 'N(KB). + by rewrite (char_norm_trans (pcore_char _ _)) ?norms_cent. +have [k KBk defQ2]:= atransP2 trnB S_Q0 S_Q2. +have [qQ2 nQ2P] := mem_max_normed maxQ2. +have hallP: pi.-Hall('N_KBP(Q2)) P. + have sPN: P \subset 'N_KBP(Q2) by rewrite subsetI joing_subr. + rewrite pHallE eqn_leq -{1}(part_pnat_id piP) dvdn_leq ?partn_dvd ?cardSg //. + have ->: #|P| = #|KBP|`_pi. + rewrite /KBP joingC norm_joinEl // coprime_cardMg ?(pnat_coprime piP) //. + by rewrite partnM // part_pnat_id // part_p'nat // muln1. + by rewrite sPN dvdn_leq ?partn_dvd ?cardSg ?cardG_gt0 ?subsetIl. +have hallPk: pi.-Hall('N_KBP(Q2)) (P :^ k). + rewrite pHallE -(card_Hall hallP) cardJg eqxx andbT subsetI /=. + by rewrite defQ2 normJ conjSg conj_subG ?joing_subr // mem_gen // inE KBk. +have [gz]: exists2 gz, gz \in 'N_KBP(Q2) & P :=: (P :^ k) :^ gz. + apply: Hall_trans (solvableS (subsetIr _ _) _) hallP hallPk. + have ntQ2: Q2 != 1%G by case: eqP nt_mnP maxQ2 => // -> ->. + exact: mFT_sol (mFT_norm_proper ntQ2 (mFT_pgroup_proper qQ2)). +rewrite [KBP]norm_joinEr //= setIC -group_modr //= setIC -/KB. +case/imset2P=> g z; case/setIP=> KBg nQ2g Pz ->{gz} defP. +exists (k * g); last first. + by apply: val_inj; rewrite /= conjsgM -(normP nQ2g) defQ2. +rewrite /KP -defK // (subsetP (subsetIl _ 'C(B))) //= setIAC defK // -/KB. +rewrite -coprime_norm_cent 1?coprime_sym ?(pnat_coprime piP) //= -/KB. +rewrite inE groupM //; apply/normP. +by rewrite -{2}(conjsgK z P) (conjGid Pz) {2}defP /= !conjsgM conjsgK. +Qed. + +End NormedConstrained. + +(* This is B & G, Proposition 7.5(a). As this is only used in Proposition *) +(* 10.10, under the assumption A \in E*_p(G), we avoid the in_pmaxElemE *) +(* detour A = [set x in 'C_G(A) | x ^+ p == 1], and just use A \in E*_p(G). *) +Proposition plength_1_normed_constrained p A : + A :!=: 1 -> A \in 'E*_p(G) -> (forall M, M \proper G -> p.-length_1 M) -> + normed_constrained A. +Proof. +move=> ntA EpA pl1subG. +case/pmaxElemP: (EpA); case/pElemP=> sAG; case/and3P=> pA cAA _ _. +have prA: A \proper G := sub_proper_trans cAA (mFT_cent_proper ntA). +split=> // X Y sAX prX; case/setIdP; case/andP=> sYX p'Y nYA. +have pl1X := pl1subG _ prX; have solX := mFT_sol prX. +have [p_pr _ [r oApr]] := pgroup_pdiv pA ntA. +have oddp: odd p by move: (mFT_odd A); rewrite oApr odd_exp. +have def_pi: \pi(A)^' =i p^'. + by move=> q; rewrite inE /= oApr pi_of_exp // pi_of_prime. +have{p'Y} p'Y : p^'.-group Y by rewrite -(eq_pgroup _ def_pi). +rewrite (eq_pcore _ def_pi) (@plength1_norm_pmaxElem _ p X A) //. +by rewrite (subsetP (pmaxElemS p (subsetT _))) // setIC 2!inE sAX. +Qed. + +(* This is B & G, Proposition 7.5(b). *) +Proposition SCN_normed_constrained p P A : + p.-Sylow(G) P -> A \in 'SCN_2(P) -> normed_constrained A. +Proof. +move=> sylP; rewrite 2!inE -andbA => /and3P[nsAP /eqP defCA lt1mA]. +have [sAP nAP]:= andP nsAP. +have pP := pHall_pgroup sylP; have pA := pgroupS sAP pP. +have abA: abelian A by rewrite /abelian -{1}defCA subsetIr. +have prP: P \proper G := mFT_pgroup_proper pP. +have ntA: A :!=: 1 by rewrite -rank_gt0 ltnW. +pose pi := \pi(A); simpl in pi. +have [p_pr pdvA [r oApr]] := pgroup_pdiv pA ntA. +have{r oApr} def_pi: pi =i (p : nat_pred). + by move=> p'; rewrite !inE oApr primes_exp // primes_prime ?inE. +have def_pi' := eq_negn def_pi; have defK := eq_pcore _ def_pi'. +pose Z := 'Ohm_1('Z(P)); have sZ_ZP: Z \subset 'Z(P) by exact: Ohm_sub. +have sZP_A: 'Z(P) \subset A by rewrite -defCA setIS ?centS. +have sZA := subset_trans sZ_ZP sZP_A. +have nsA1: 'Ohm_1(A) <| P by exact: (char_normal_trans (Ohm_char _ _)). +pose inZor1 B := B \subset Z \/ #|Z| = p /\ Z \subset B. +have [B [E2_B nsBP sBZ]]: exists B, [/\ B \in 'E_p^2(A), B <| P & inZor1 B]. + have pZP: p.-group 'Z(P) by exact: pgroupS (center_sub _) pP. + have pZ: p.-group Z by exact: pgroupS sZ_ZP pZP. + have abelZ: p.-abelem Z by rewrite Ohm1_abelem ?center_abelian. + have nsZP: Z <| P := sub_center_normal sZ_ZP; have [sZP nZP] := andP nsZP. + case: (eqVneq Z 1). + rewrite -(setIidPr sZ_ZP); move/TI_Ohm1; rewrite setIid. + by move/(trivg_center_pgroup pP)=> P1; rewrite -subG1 -P1 sAP in ntA. + case/(pgroup_pdiv pZ)=> _ _ [[|k] /=]; rewrite -/Z => oZ; last first. + have: 2 <= 'r_p(Z) by rewrite p_rank_abelem // oZ pfactorK. + case/p_rank_geP=> B; rewrite /= -/Z => Ep2Z_B; exists B. + rewrite (subsetP (pnElemS _ _ sZA)) //. + case/setIdP: Ep2Z_B; case/setIdP=> sBZ _ _; split=> //; last by left. + by rewrite sub_center_normal ?(subset_trans sBZ). + pose BZ := ('Ohm_1(A) / Z) :&: 'Z(P / Z). + have ntBz: BZ != 1. + rewrite meet_center_nil ?quotient_nil ?(pgroup_nil pP) ?quotient_normal //. + rewrite -subG1 quotient_sub1 ?(subset_trans (normal_sub nsA1) nZP) //= -/Z. + apply: contraL lt1mA => sA1Z; rewrite -(pfactorK 1 p_pr) -oZ -rank_Ohm1. + by rewrite -(rank_abelem abelZ) -leqNgt rankS. + have lt1A1: 1 < logn p #|'Ohm_1(A)| by rewrite -p_rank_abelian -?rank_pgroup. + have [B [sBA1 nsBP oB]] := normal_pgroup pP nsA1 lt1A1. + exists B; split=> //; last do [right; split=> //]. + rewrite 2!inE (subset_trans sBA1) ?Ohm_sub // oB pfactorK //. + by rewrite (abelemS sBA1) ?Ohm1_abelem. + apply/idPn=> s'BZ; have: B :&: Z = 1 by rewrite setIC prime_TIg ?oZ. + move/TI_Ohm1; apply/eqP; rewrite meet_center_nil ?(pgroup_nil pP) //. + by rewrite -cardG_gt1 oB (ltn_exp2l 0 _ (prime_gt1 p_pr)). +split; rewrite ?(sub_proper_trans sAP) // => X Y sAX prX. +rewrite inE defK -andbA (eq_pgroup _ def_pi'); case/and3P=> sYX p'Y nYA. +move: E2_B; rewrite 2!inE -andbA; case/and3P=> sBA abelB dimB2. +have [pB cBB _] := and3P abelB. +have ntB: B :!=: 1 by case: (eqsVneq B 1) dimB2 => // ->; rewrite cards1 logn1. +have cBA b: b \in B -> A \subset 'C[b]. + by move=> Bb; rewrite -cent_set1 centsC sub1set (subsetP abA) ?(subsetP sBA). +have solCB (b : gT): b != 1 -> solvable 'C[b]. + by move=> ntb; rewrite mFT_sol ?mFT_cent1_proper. +wlog{sAX prX} [b B'b defX]: X Y p'Y nYA sYX / exists2 b, b \in B^# & 'C[b] = X. + move=> IH; have nYB := subset_trans sBA nYA. + rewrite -(coprime_abelian_gen_cent1 cBB _ nYB); last first. + - by rewrite coprime_sym (pnat_coprime pB). + - apply: contraL dimB2 => /cyclicP[x defB]. + have Bx: x \in B by rewrite defB cycle_id. + rewrite defB -orderE (abelem_order_p abelB Bx) ?(pfactorK 1) //. + by rewrite -cycle_eq1 -defB. + rewrite gen_subG; apply/bigcupsP=> b B'b. + have [ntb Bb]:= setD1P B'b; have sYbY: 'C_Y[b] \subset Y := subsetIl _ _. + have{IH} sYbKb: 'C_Y[b] \subset 'O_p^'('C[b]). + rewrite IH ?(pgroupS sYbY) ?subsetIr //; last by exists b. + by rewrite normsI // ?normsG ?cBA. + have{sYbKb} sYbKXb: 'C_Y[b] \subset 'O_p^'('C_X(<[b]>)). + apply: subset_trans (pcoreS _ (subsetIr _ _)). + by rewrite /= cent_gen cent_set1 subsetI setSI. + rewrite (subset_trans sYbKXb) // p'core_cent_pgroup ?mFT_sol //. + rewrite /psubgroup ?(pgroupS _ pB) cycle_subG //. + by rewrite (subsetP sAX) ?(subsetP sBA). +wlog Zb: b X Y defX B'b p'Y nYA sYX / b \in Z. + move=> IH; case Zb: (b \in Z); first exact: IH Zb. + case/setD1P: B'b => ntb Bb; have solX := solCB b ntb; rewrite defX in solX. + case: sBZ => [sBZ | [oZ sZB]]; first by rewrite (subsetP sBZ) in Zb. + have defB: Z * <[b]> = B. + apply/eqP; rewrite eqEcard mulG_subG sZB cycle_subG Bb. + have obp := abelem_order_p abelB Bb ntb. + rewrite (card_pgroup pB) /= (eqP dimB2) TI_cardMg -/#[_] ?oZ ?obp //. + rewrite -obp in p_pr; case: (prime_subgroupVti [group of Z] p_pr) => //. + by rewrite cycle_subG Zb. + pose P1 := P :&: X; have sP1P: P1 \subset P := subsetIl _ _. + have pP1 := pgroupS sP1P pP. + have [P2 sylP2 sP12] := Sylow_superset (subsetIr _ _) pP1. + have defP1: P1 = 'C_P(B). + rewrite -defB centM /= -/Z setIA /cycle cent_gen cent_set1 defX. + by rewrite [P :&: _](setIidPl _) // centsC (subset_trans sZ_ZP) ?subsetIr. + have dimPP1: logn p #|P : P1| <= 1. + by rewrite defP1 logn_quotient_cent_abelem ?normal_norm ?(eqP dimB2). + have{dimPP1} nsP12: P1 <| P2. + have pP2 := pHall_pgroup sylP2. + have: logn p #|P2 : P1| <= 1. + apply: leq_trans dimPP1; rewrite dvdn_leq_log //. + rewrite -(dvdn_pmul2l (cardG_gt0 [group of P1])) !Lagrange ?subsetIl //. + rewrite -(part_pnat_id pP2) (card_Hall sylP). + by rewrite partn_dvd ?cardSg ?subsetT. + rewrite -(pfactorK 1 p_pr) -pfactor_dvdn ?prime_gt0 // -p_part. + rewrite part_pnat_id ?(pnat_dvd (dvdn_indexg _ _)) //=. + case: (primeP p_pr) => _ dv_p; move/dv_p=> {dv_p}. + case/pred2P=> oP21; first by rewrite -(index1g sP12 oP21) normal_refl. + by rewrite (p_maximal_normal pP2) ?p_index_maximal ?oP21. + have nsZP1_2: 'Z(P1) <| P2 by rewrite (char_normal_trans (center_char _)). + have sZKp: Z \subset 'O_{p^', p}(X). + suff: 'Z(P1) \subset 'O_{p^', p}(X). + apply: subset_trans; rewrite subsetI {1}defP1 (subset_trans sZB). + by rewrite (subset_trans sZ_ZP) ?subIset // orbC centS. + by rewrite subsetI normal_sub. + apply: odd_p_abelian_constrained sylP2 (center_abelian _) nsZP1_2 => //. + exact: mFT_odd. + have coYZ: coprime #|Y| #|Z|. + by rewrite oZ coprime_sym (pnat_coprime _ p'Y) ?pnatE ?inE. + have nYZ := subset_trans sZA nYA. + have <-: [~: Y, Z] * 'C_Y(Z) = Y. + exact: coprime_cent_prod (solvableS sYX solX). + set K := 'O_p^'(X); have [nKY nKZ]: Y \subset 'N(K) /\ Z \subset 'N(K). + rewrite !(char_norm_trans (pcore_char _ _)) ?(subset_trans sZA) ?normsG //. + by rewrite -defX cBA. + rewrite mul_subG //. + have coYZK: coprime #|Y / K| #|'O_p(X / K)|. + by rewrite coprime_sym coprime_morphr ?(pnat_coprime (pcore_pgroup _ _)). + rewrite -quotient_sub1 ?comm_subG // -(coprime_TIg coYZK) subsetI. + rewrite /= -quotient_pseries2 !quotientS ?commg_subl //. + by rewrite (subset_trans (commgSS sYX sZKp)) ?commg_subr //= gFnorm. + have: 'O_p^'('C_X(Z)) \subset K. + rewrite p'core_cent_pgroup // /psubgroup /pgroup oZ pnat_id //. + by rewrite -defX (subset_trans sZA) ?cBA. + apply: subset_trans; apply: subset_trans (pcoreS _ (subsetIr _ _)). + have: cyclic Z by rewrite prime_cyclic ?oZ. + case/cyclicP=> z defZ; have Zz: z \in Z by rewrite defZ cycle_id. + rewrite subsetI setSI //= (IH z) ?subsetIr ?(pgroupS (subsetIl _ _)) //. + - by rewrite defZ /= cent_gen cent_set1. + - rewrite !inE -cycle_eq1 -defZ trivg_card_le1 oZ -ltnNge prime_gt1 //=. + by rewrite (subsetP sZB). + by rewrite normsI // norms_cent // cents_norm // centsC (subset_trans sZA). +set K := 'O_p^'(X); have nsKX: K <| X by exact: pcore_normal. +case/setD1P: B'b => ntb Bb. +have [sAX solX]: A \subset X /\ solvable X by rewrite -defX cBA ?solCB. +have sPX: P \subset X. + by rewrite -defX -cent_set1 centsC sub1set; case/setIP: (subsetP sZ_ZP b Zb). +have [nKA nKY nKP]: [/\ A \subset 'N(K), Y \subset 'N(K) & P \subset 'N(K)]. + by rewrite !(subset_trans _ (normal_norm nsKX)). +have sylPX: p.-Sylow(X) P by exact: pHall_subl (subsetT _) sylP. +have sAKb: A \subset 'O_{p^', p}(X). + exact: (odd_p_abelian_constrained (mFT_odd _)) abA nsAP. +have coYZK: coprime #|Y / K| #|'O_p(X / K)|. + by rewrite coprime_sym coprime_morphr ?(pnat_coprime (pcore_pgroup _ _)). +have cYAq: A / K \subset 'C_('O_p(X / K))(Y / K). + rewrite subsetI -quotient_pseries2 quotientS //= (sameP commG1P trivgP). + rewrite /= -quotientR // -(coprime_TIg coYZK) subsetI /= -quotient_pseries2. + rewrite !quotientS ?commg_subr // (subset_trans (commgSS sAKb sYX)) //. + by rewrite commg_subl /= gFnorm. +have cYKq: Y / K \subset 'C('O_p(X / K)). + apply: coprime_nil_faithful_cent_stab => /=. + - by rewrite (char_norm_trans (pcore_char _ _)) ?normsG ?quotientS. + - by rewrite coprime_morphr ?(pnat_coprime (pcore_pgroup _ _)). + - exact: pgroup_nil (pcore_pgroup _ _). + apply: subset_trans (cYAq); rewrite -defCA -['C_P(A) / K](morphim_restrm nKP). + rewrite injm_cent ?ker_restrm ?ker_coset ?morphim_restrm -?quotientE //. + rewrite setIid (setIidPr sAP) setISS ?centS //. + by rewrite pcore_sub_Hall ?morphim_pHall. + by rewrite coprime_TIg ?(pnat_coprime _ (pcore_pgroup _ _)). +rewrite -quotient_sub1 //= -/K -(coprime_TIg coYZK) subsetI subxx /=. +rewrite -Fitting_eq_pcore ?trivg_pcore_quotient // in cYKq *. +apply: subset_trans (cent_sub_Fitting (quotient_sol _ solX)). +by rewrite subsetI quotientS. +Qed. + +(* This is B & G, Theorem 7.6 (the Thompson Transitivity Theorem). *) +Theorem Thompson_transitivity p q A : + A \in 'SCN_3[p] -> q \in p^' -> + [transitive 'O_p^'('C(A)), on |/|*(A; q) | 'JG]. +Proof. +case/bigcupP=> P; rewrite 2!inE => sylP /andP[SCN_A mA3]. +have [defZ def_pi']: 'Z(A) = A /\ \pi(A)^' =i p^'. + rewrite inE -andbA in SCN_A; case/and3P: SCN_A => sAP _ /eqP defCA. + case: (eqsVneq A 1) mA3 => /= [-> | ntA _]. + rewrite /rank big1_seq // => p1 _; rewrite /p_rank big1 // => E. + by rewrite inE; case/andP; move/trivgP->; rewrite cards1 logn1. + have [p_pr _ [k ->]] := pgroup_pdiv (pgroupS sAP (pHall_pgroup sylP)) ntA. + split=> [|p1]; last by rewrite !inE primes_exp // primes_prime ?inE. + by apply/eqP; rewrite eqEsubset subsetIl subsetI subxx -{1}defCA subsetIr. +rewrite -(eq_pcore _ def_pi') -def_pi' => pi'q. +apply: normed_constrained_rank3_trans; rewrite ?defZ //. +by apply: SCN_normed_constrained sylP _; rewrite inE SCN_A ltnW. +Qed. + +End Seven. + diff --git a/mathcomp/odd_order/BGsection8.v b/mathcomp/odd_order/BGsection8.v new file mode 100644 index 0000000..8e306fa --- /dev/null +++ b/mathcomp/odd_order/BGsection8.v @@ -0,0 +1,404 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype path. +Require Import finset prime fingroup automorphism action gproduct gfunctor. +Require Import center commutator pgroup gseries nilpotent sylow abelian maximal. +Require Import BGsection1 BGsection5 BGsection6 BGsection7. + +(******************************************************************************) +(* This file covers B & G, section 8, i.e., the proof of two special cases *) +(* of the Uniqueness Theorem, for maximal groups with Fitting subgroups of *) +(* rank at least 3. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Section Eight. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types H M A X P : {group gT}. +Implicit Types p q r : nat. + +Local Notation "K ` p" := 'O_(nat_pred_of_nat p)(K) + (at level 8, p at level 2, format "K ` p") : group_scope. +Local Notation "K ` p" := 'O_(nat_pred_of_nat p)(K)%G : Group_scope. + +(* This is B & G, Theorem 8.1(a). *) +Theorem non_pcore_Fitting_Uniqueness p M A0 : + M \in 'M -> ~~ p.-group ('F(M)) -> A0 \in 'E*_p('F(M)) -> 'r_p(A0) >= 3 -> + 'C_('F(M))(A0)%G \in 'U. +Proof. +set F := 'F(M) => maxM p'F /pmaxElemP[/=/setIdP[sA0F abelA0] maxA0]. +have [pA0 cA0A0 _] := and3P abelA0; rewrite (p_rank_abelem abelA0) => dimA0_3. +rewrite (uniq_mmax_subset1 maxM) //= -/F; last by rewrite subIset ?Fitting_sub. +set A := 'C_F(A0); pose pi := \pi(A). +have [sZA sAF]: 'Z(F) \subset A /\ A \subset F by rewrite subsetIl setIS ?centS. +have nilF: nilpotent F := Fitting_nil _. +have nilZ := nilpotentS (center_sub _) nilF. +have piZ: \pi('Z(F)) = \pi(F) by rewrite pi_center_nilpotent. +have def_pi: pi = \pi(F). + by apply/eq_piP=> q; apply/idP/idP; last rewrite -piZ; exact: piSg. +have def_nZq: forall q, q \in pi -> 'N('Z(F)`q) = M. + move=> q; rewrite def_pi -piZ -p_part_gt1. + rewrite -(card_Hall (nilpotent_pcore_Hall _ nilZ)) cardG_gt1 /= -/F => ntZ. + apply: mmax_normal => //=; apply: char_normal_trans (Fitting_normal _). + exact: char_trans (pcore_char _ _) (center_char _). +have sCqM: forall q, q \in pi -> 'C(A`q) \subset M. + move=> q; move/def_nZq <-; rewrite cents_norm // centS //. + rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ _)) ?pcore_pgroup //. + by apply: nilpotentS (Fitting_nil M); exact: subsetIl. + exact: subset_trans (pcore_sub _ _) _. +have sA0A: A0 \subset A by rewrite subsetI sA0F. +have pi_p: p \in pi. + by apply: (piSg sA0A); rewrite -[p \in _]logn_gt0 (leq_trans _ dimA0_3). +have sCAM: 'C(A) \subset M. + by rewrite (subset_trans (centS (pcore_sub p _))) ?sCqM. +have prM: M \proper G := mmax_proper maxM; have solM := mFT_sol prM. +have piCA: pi.-group('C(A)). + apply/pgroupP=> q q_pr; case/Cauchy=> // x cAx oxq; apply/idPn=> pi'q. + have Mx := subsetP sCAM x cAx; pose C := 'C_F(<[x]>). + have sAC: A \subset C by rewrite subsetI sAF centsC cycle_subG. + have sCFC_C: 'C_F(C) \subset C. + by rewrite (subset_trans _ sAC) ?setIS // centS ?(subset_trans _ sAC). + have cFx: x \in 'C_M(F). + rewrite inE Mx -cycle_subG coprime_nil_faithful_cent_stab //=. + by rewrite cycle_subG (subsetP (gFnorm _ _)). + by rewrite -orderE coprime_pi' ?cardG_gt0 // -def_pi oxq pnatE. + case/negP: pi'q; rewrite def_pi mem_primes q_pr cardG_gt0 -oxq cardSg //. + by rewrite cycle_subG (subsetP (cent_sub_Fitting _)). +have{p'F} pi_alt q: exists2 r, r \in pi & r != q. + have [<-{q} | ] := eqVneq p q; last by exists p. + rewrite def_pi; apply/allPn; apply: contra p'F => /allP/=pF. + by apply/pgroupP=> q q_pr qF; rewrite !inE pF // mem_primes q_pr cardG_gt0. +have sNZqXq' q X: + A \subset X -> X \proper G -> 'O_q^'('N_X('Z(F)`q)) \subset 'O_q^'(X). +- move=> sAX prX; have sZqX: 'Z(F)`q \subset X. + exact: subset_trans (pcore_sub _ _) (subset_trans sZA sAX). + have cZqNXZ: 'O_q^'('N_X('Z(F)`q)) \subset 'C('Z(F)`q). + have coNq'Zq: coprime #|'O_q^'('N_X('Z(F)`q))| #|'Z(F)`q|. + by rewrite coprime_sym coprime_pcoreC. + rewrite (sameP commG1P trivgP) -(coprime_TIg coNq'Zq) subsetI commg_subl /=. + rewrite commg_subr /= andbC (subset_trans (pcore_sub _ _)) ?subsetIr //=. + by rewrite (char_norm_trans (pcore_char _ _)) ?normsG // subsetI sZqX normG. + have: 'O_q^'('C_X(('Z(F))`q)) \subset 'O_q^'(X). + by rewrite p'core_cent_pgroup ?mFT_sol // /psubgroup sZqX pcore_pgroup. + apply: subset_trans; apply: subset_trans (pcoreS _ (subcent_sub _ _)). + by rewrite !subsetI subxx cZqNXZ (subset_trans (pcore_sub _ _)) ?subsetIl. +have sArXq' q r X: + q \in pi -> q != r -> A \subset X -> X \proper G -> A`r \subset 'O_q^'(X). +- move=> pi_q r'q sAX prX; apply: subset_trans (sNZqXq' q X sAX prX). + apply: subset_trans (pcoreS _ (subsetIr _ _)). + rewrite -setIA (setIidPr (pcore_sub _ _)) subsetI. + rewrite (subset_trans (pcore_sub _ _)) //= def_nZq //. + apply: subset_trans (pcore_Fitting _ _); rewrite -/F. + rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ nilF)) //; last first. + exact: subset_trans (pcore_sub _ _) sAF. + by apply: (pi_pnat (pcore_pgroup _ _)); rewrite !inE eq_sym. +have cstrA: normed_constrained A. + split=> [||X Y sAX prX]. + - by apply/eqP=> A1; rewrite /pi /= A1 cards1 in pi_p. + - exact: sub_proper_trans (subset_trans sAF (Fitting_sub _)) prM. + rewrite !inE -/pi -andbA; case/and3P=> sYX pi'Y nYA. + rewrite -bigcap_p'core subsetI sYX; apply/bigcapsP=> [[q /= _] pi_q]. + have [r pi_r q'r] := pi_alt q. + have{sArXq'} sArXq': A`r \subset 'O_q^'(X) by apply: sArXq'; rewrite 1?eq_sym. + have cA_CYr: 'C_Y(A`r) \subset 'C(A). + have coYF: coprime #|Y| #|F|. + by rewrite coprime_sym coprime_pi' ?cardG_gt0 // -def_pi. + rewrite (sameP commG1P trivgP) -(coprime_TIg coYF) commg_subI //. + by rewrite setIS // (subset_trans (sCqM r pi_r)) // gFnorm. + by rewrite subsetI subsetIl. + have{cA_CYr} CYr1: 'C_Y(A`r) = 1. + rewrite -(setIid Y) setIAC coprime_TIg // (coprimeSg cA_CYr) //. + by rewrite (pnat_coprime piCA). + have{CYr1} ->: Y :=: [~: Y, A`r]. + rewrite -(mulg1 [~: Y, _]) -CYr1 coprime_cent_prod //. + - by rewrite (subset_trans (pcore_sub _ _)). + - rewrite coprime_sym (coprimeSg (pcore_sub _ _)) //= -/A. + by rewrite coprime_pi' ?cardG_gt0. + by rewrite mFT_sol // (sub_proper_trans sYX). + rewrite (subset_trans (commgS _ sArXq')) // commg_subr. + by rewrite (char_norm_trans (pcore_char _ _)) ?normsG. +have{cstrA} nbyApi'1 q: q \in pi^' -> |/|*(A; q) = [set 1%G]. + move=> pi'q; have trA: [transitive 'O_pi^'('C(A)), on |/|*(A; q) | 'JG]. + apply: normed_constrained_rank3_trans; rewrite //= -/A. + rewrite -rank_abelem // in dimA0_3; apply: leq_trans dimA0_3 (rankS _). + by rewrite /= -/A subsetI sA0A centsC subsetIr. + have [Q maxQ defAmax]: exists2 Q, Q \in |/|*(A; q) & |/|*(A; q) = [set Q]. + case/imsetP: trA => Q maxQ defAmax; exists Q; rewrite // {maxQ}defAmax. + suffices ->: 'O_pi^'('C(A)) = 1 by rewrite /orbit imset_set1 act1. + rewrite -(setIidPr (pcore_sub _ _)) coprime_TIg //. + exact: pnat_coprime piCA (pcore_pgroup _ _). + have{maxQ} qQ: q.-group Q by move: maxQ; rewrite inE => /maxgroupp/andP[]. + have [<- // |] := eqVneq Q 1%G; rewrite -val_eqE /= => ntQ. + have{defAmax trA} defFmax: |/|*(F; q) = [set Q]. + apply/eqP; rewrite eqEcard cards1 -defAmax. + have snAF: A <|<| F by rewrite nilpotent_subnormal ?Fitting_nil. + have piF: pi.-group F by rewrite def_pi /pgroup pnat_pi ?cardG_gt0. + case/(normed_trans_superset _ _ snAF): trA => //= _ /imsetP[R maxR _] -> _. + by rewrite (cardsD1 R) maxR. + have nQM: M \subset 'N(Q). + apply/normsP=> x Mx; apply: congr_group; apply/set1P. + rewrite -defFmax (acts_act (norm_acts_max_norm _ _)) ?defFmax ?set11 //. + by apply: subsetP Mx; exact: gFnorm. + have{nQM} nsQM: Q <| M. + rewrite inE in maxM; case/maxgroupP: maxM => _ maxM. + rewrite -(maxM 'N(Q)%G) ?normalG ?mFT_norm_proper //. + exact: mFT_pgroup_proper qQ. + have sQF: Q \subset F by rewrite Fitting_max ?(pgroup_nil qQ). + rewrite -(setIidPr sQF) coprime_TIg ?eqxx // in ntQ. + by rewrite coprime_pi' ?cardG_gt0 // -def_pi (pi_pnat qQ). +apply/subsetP=> H /setIdP[maxH sAH]; rewrite inE -val_eqE /=. +have prH: H \proper G := mmax_proper maxH; have solH := mFT_sol prH. +pose D := 'F(H); have nilD: nilpotent D := Fitting_nil H. +have card_pcore_nil := card_Hall (nilpotent_pcore_Hall _ _). +have piD: \pi(D) = pi. + set sigma := \pi(_); have pi_sig: {subset sigma <= pi}. + move=> q; rewrite -p_part_gt1 -card_pcore_nil // cardG_gt1 /= -/D. + apply: contraR => /nbyApi'1 defAmax. + have nDqA: A \subset 'N(D`q). + rewrite (char_norm_trans (pcore_char _ _)) //. + by rewrite (subset_trans sAH) ?gFnorm. + have [Q]:= max_normed_exists (pcore_pgroup _ _) nDqA. + by rewrite defAmax -subG1; move/set1P->. + apply/eq_piP=> q; apply/idP/idP=> [|pi_q]; first exact: pi_sig. + apply: contraLR (pi_q) => sig'q; have nilA := nilpotentS sAF nilF. + rewrite -p_part_eq1 -card_pcore_nil // -trivg_card1 -subG1 /= -/A. + have <-: 'O_sigma^'(H) = 1. + apply/eqP; rewrite -trivg_Fitting ?(solvableS (pcore_sub _ _)) //. + rewrite Fitting_pcore -(setIidPr (pcore_sub _ _)) coprime_TIg //. + by rewrite coprime_pi' ?cardG_gt0 //; exact: pcore_pgroup. + rewrite -bigcap_p'core subsetI (subset_trans (pcore_sub _ _)) //=. + apply/bigcapsP=> [[r /= _] sig_r]; apply: sArXq' => //; first exact: pi_sig. + by apply: contra sig'q; move/eqP <-. +have cAD q r: q != r -> D`q \subset 'C(A`r). + move=> r'q; have [-> |] := eqVneq D`q 1; first by rewrite sub1G. + rewrite -cardG_gt1 card_pcore_nil // p_part_gt1 piD => pi_q. + have sArHq': A`r \subset 'O_q^'(H) by rewrite sArXq'. + have coHqHq': coprime #|D`q| #|'O_q^'(H)| by rewrite coprime_pcoreC. + rewrite (sameP commG1P trivgP) -(coprime_TIg coHqHq') commg_subI //. + rewrite subsetI subxx /= p_core_Fitting (subset_trans (pcore_sub _ _)) //. + exact: gFnorm. + rewrite subsetI sArHq' (subset_trans (subset_trans (pcore_sub _ _) sAH)) //. + by rewrite /= p_core_Fitting gFnorm. +have sDM: D \subset M. + rewrite [D]FittingEgen gen_subG; apply/bigcupsP=> [[q /= _] _]. + rewrite -p_core_Fitting -/D; have [r pi_r r'q] := pi_alt q. + by apply: subset_trans (sCqM r pi_r); apply: cAD; rewrite eq_sym. +have cApHp': A`p \subset 'C('O_p^'(H)). + have coApHp': coprime #|'O_p^'(H)| #|A`p|. + by rewrite coprime_sym coprime_pcoreC. + have solHp': solvable 'O_p^'(H) by rewrite (solvableS (pcore_sub _ _)). + have nHp'Ap: A`p \subset 'N('O_p^'(H)). + by rewrite (subset_trans (subset_trans (pcore_sub _ _) sAH)) ?gFnorm. + apply: subset_trans (coprime_cent_Fitting nHp'Ap coApHp' solHp'). + rewrite subsetI subxx centsC /= FittingEgen gen_subG. + apply/bigcupsP=> [[q /= _] _]; have [-> | /cAD] := eqVneq q p. + by rewrite -(setIidPl (pcore_sub p _)) TI_pcoreC sub1G. + apply: subset_trans; rewrite p_core_Fitting -pcoreI. + by apply: sub_pcore => r /andP[]. +have sHp'M: 'O_p^'(H) \subset M. + by apply: subset_trans (sCqM p pi_p); rewrite centsC. +have ntDp: D`p != 1 by rewrite -cardG_gt1 card_pcore_nil // p_part_gt1 piD. +have sHp'_NMDp': 'O_p^'(H) \subset 'O_p^'('N_M(D`p)). + apply: subset_trans (pcoreS _ (subsetIr _ _)). + rewrite -setIA (setIidPr (pcore_sub _ _)) /= (mmax_normal maxH) //. + by rewrite subsetI sHp'M subxx. + by rewrite /= p_core_Fitting pcore_normal. +have{sHp'_NMDp'} sHp'Mp': 'O_p^'(H) \subset 'O_p^'(M). + have pM_D: p.-subgroup(M) D`p. + by rewrite /psubgroup pcore_pgroup (subset_trans (pcore_sub _ _)). + apply: subset_trans (p'core_cent_pgroup pM_D (mFT_sol prM)). + apply: subset_trans (pcoreS _ (subcent_sub _ _)). + rewrite !subsetI sHp'_NMDp' sHp'M andbT /= (sameP commG1P trivgP). + have coHp'Dp: coprime #|'O_p^'(H)| #|D`p|. + by rewrite coprime_sym coprime_pcoreC. + rewrite -(coprime_TIg coHp'Dp) subsetI commg_subl commg_subr /=. + by rewrite p_core_Fitting !(subset_trans (pcore_sub _ _)) ?gFnorm. +have sMp'H: 'O_p^'(M) \subset H. + rewrite -(mmax_normal maxH (pcore_normal p H)) /= -p_core_Fitting //. + rewrite -/D (subset_trans _ (cent_sub _)) // centsC. + have solMp' := solvableS (pcore_sub p^' _) (mFT_sol prM). + have coMp'Dp: coprime #|'O_p^'(M)| #|D`p|. + by rewrite coprime_sym coprime_pcoreC. + have nMp'Dp: D`p \subset 'N('O_p^'(M)). + by rewrite (subset_trans (subset_trans (pcore_sub _ _) sDM)) ?gFnorm. + apply: subset_trans (coprime_cent_Fitting nMp'Dp coMp'Dp solMp'). + rewrite subsetI subxx centsC /= FittingEgen gen_subG. + apply/bigcupsP=> [[q /= _] _]; have [<- | /cAD] := eqVneq p q. + by rewrite -(setIidPl (pcore_sub p _)) TI_pcoreC sub1G. + rewrite centsC; apply: subset_trans. + rewrite -p_core_Fitting Fitting_pcore pcore_max ?pcore_pgroup //=. + rewrite /normal subsetI -pcoreI pcore_sub subIset ?gFnorm //=. + rewrite pcoreI (subset_trans (pcore_sub _ _)) //= -/F centsC. + case/dprodP: (nilpotent_pcoreC p nilF) => _ _ /= cFpp' _. + rewrite centsC (subset_trans cFpp' (centS _)) //. + have hallFp := nilpotent_pcore_Hall p nilF. + by rewrite (sub_Hall_pcore hallFp). +have{sHp'Mp' sMp'H} eqHp'Mp': 'O_p^'(H) = 'O_p^'(M). + apply/eqP; rewrite eqEsubset sHp'Mp'. + apply: subset_trans (sNZqXq' p H sAH prH). + apply: subset_trans (pcoreS _ (subsetIr _ _)). + rewrite -setIA (setIidPr (pcore_sub _ _)) subsetI sMp'H /=. + rewrite (mmax_normal maxM (char_normal_trans (pcore_char _ _) _)) //. + by rewrite (char_normal_trans (center_char _)) ?Fitting_normal. + by rewrite -cardG_gt1 card_pcore_nil // p_part_gt1 piZ -def_pi. +have ntHp': 'O_p^'(H) != 1. + have [q pi_q p'q] := pi_alt p; have: D`q \subset 'O_p^'(H). + by rewrite p_core_Fitting sub_pcore // => r; move/eqnP->. + rewrite -proper1G; apply: proper_sub_trans; rewrite proper1G. + by rewrite -cardG_gt1 card_pcore_nil // p_part_gt1 piD. +rewrite -(mmax_normal maxH (pcore_normal p^' H)) //= eqHp'Mp'. +by rewrite (mmax_normal maxM (pcore_normal _ _)) //= -eqHp'Mp'. +Qed. + +(* This is B & G, Theorem 8.1(b). *) +Theorem SCN_Fitting_Uniqueness p M P A : + M \in 'M -> p.-group ('F(M)) -> p.-Sylow(M) P -> + 'r_p('F(M)) >= 3 -> A \in 'SCN_3(P) -> + [/\ p.-Sylow(G) P, A \subset 'F(M) & A \in 'U]. +Proof. +set F := 'F(M) => maxM pF sylP dimFp3 scn3_A. +have [scnA dimA3] := setIdP scn3_A; have [nsAP defCA] := SCN_P scnA. +have cAA := SCN_abelian scnA; have sAP := normal_sub nsAP. +have [sPM pP _] := and3P sylP; have sAM := subset_trans sAP sPM. +have{dimA3} ntA: A :!=: 1 by case: eqP dimA3 => // ->; rewrite rank1. +have prM := mmax_proper maxM; have solM := mFT_sol prM. +have{pF} Mp'1: 'O_p^'(M) = 1. + apply/eqP; rewrite -trivg_Fitting ?(solvableS (pcore_sub _ _)) //. + rewrite Fitting_pcore -(setIidPr (pcore_sub _ _)) coprime_TIg //. + exact: pnat_coprime (pcore_pgroup _ _). +have defF: F = M`p := Fitting_eq_pcore Mp'1. +have sFP: F \subset P by rewrite defF (pcore_sub_Hall sylP). +have sAF: A \subset F. + rewrite defF -(pseries_pop2 _ Mp'1). + exact: (odd_p_abelian_constrained (mFT_odd _) solM sylP cAA nsAP). +have sZA: 'Z(F) \subset A. + by rewrite -defCA setISS ?centS // defF pcore_sub_Hall. +have sCAM: 'C(A) \subset M. + have nsZM: 'Z(F) <| M := char_normal_trans (center_char _) (Fitting_normal _). + rewrite -(mmax_normal maxM nsZM); last first. + rewrite /= -(setIidPr (center_sub _)) meet_center_nil ?Fitting_nil //. + by rewrite -proper1G (proper_sub_trans _ sAF) ?proper1G. + by rewrite (subset_trans _ (cent_sub _)) ?centS. +have nsZL_M: 'Z('L(P)) <| M. + by rewrite (Puig_center_normal (mFT_odd _) solM sylP). +have sNPM: 'N(P) \subset M. + rewrite -(mmax_normal maxM nsZL_M). + by rewrite (char_norm_trans (center_Puig_char P)). + apply/eqP; move/(trivg_center_Puig_pgroup (pHall_pgroup sylP))=> P1. + by rewrite -subG1 -P1 sAP in ntA. +have sylPG: p.-Sylow(G) P := mmax_sigma_Sylow maxM sylP sNPM. +split; rewrite // (uniq_mmax_subset1 maxM sAM). +have{scn3_A} scn3_A: A \in 'SCN_3[p] by apply/bigcupP; exists P; rewrite // inE. +pose K := 'O_p^'('C(A)); have sKF: K \subset F. + have sKM: K \subset M := subset_trans (pcore_sub _ _) sCAM. + apply: subset_trans (cent_sub_Fitting solM). + rewrite subsetI sKM coprime_nil_faithful_cent_stab ?Fitting_nil //. + - by rewrite (subset_trans (subset_trans (pcore_sub _ _) sCAM)) ?gFnorm. + - by rewrite /= -/F defF coprime_pcoreC. + have sACK: A \subset 'C_F(K) by rewrite subsetI sAF centsC pcore_sub. + by rewrite /= -/F -/K (subset_trans _ sACK) //= -defCA setISS ?centS. +have{sKF} K1: K = 1 by rewrite -(setIidPr sKF) defF TI_pcoreC. +have p'nbyA_1 q: q != p -> |/|*(A; q) = [set 1%G]. + move=> p'q. + have: [transitive K, on |/|*(A; q) | 'JG] by apply: Thompson_transitivity. + case/imsetP=> Q maxQ; rewrite K1 /orbit imset_set1 act1 => defAmax. + have nQNA: 'N(A) \subset 'N(Q). + apply/normsP=> x Nx; apply: congr_group; apply/set1P; rewrite -defAmax. + by rewrite (acts_act (norm_acts_max_norm _ _)). + have{nQNA} nQF: F \subset 'N(Q). + exact: subset_trans (subset_trans (normal_norm nsAP) nQNA). + have defFmax: |/|*(F; q) = [set Q] := max_normed_uniq defAmax sAF nQF. + have nQM: M \subset 'N(Q). + apply/normsP=> x Mx; apply: congr_group; apply/set1P; rewrite -defFmax. + rewrite (acts_act (norm_acts_max_norm _ _)) ?defFmax ?set11 //. + by rewrite (subsetP (gFnorm _ _)). + have [<- // | ntQ] := eqVneq Q 1%G. + rewrite inE in maxQ; have [qQ _] := andP (maxgroupp maxQ). + have{nQM} defNQ: 'N(Q) = M. + by rewrite (mmax_norm maxM) // (mFT_pgroup_proper qQ). + case/negP: ntQ; rewrite -[_ == _]subG1 -Mp'1 -defNQ pcore_max ?normalG //. + exact: pi_pnat qQ _. +have{p'nbyA_1} p'nbyA_1 X: + X \proper G -> p^'.-group X -> A \subset 'N(X) -> X :=: 1. +- move=> prX p'X nXA; have solX := mFT_sol prX. + apply/eqP; rewrite -trivg_Fitting // -subG1 /= FittingEgen gen_subG. + apply/bigcupsP=> [[q /= _] _]; have [-> | p'q] := eqVneq q p. + rewrite -(setIidPl (pcore_sub _ _)) coprime_TIg //. + by rewrite (pnat_coprime (pcore_pgroup _ _)). + have [|R] := max_normed_exists (pcore_pgroup q X) (char_norm_trans _ nXA). + exact: pcore_char. + by rewrite p'nbyA_1 // => /set1P->. +apply/subsetPn=> [[H0 MA_H0 neH0M]]. +have:= erefl [arg max_(H > H0 | (H \in 'M(A)) && (H != M)) #|H :&: M|`_p]. +case: arg_maxP => [|H {H0 MA_H0 neH0M}]; first by rewrite MA_H0 -in_set1. +rewrite /= inE -andbA; case/and3P=> maxH sAH neHM maxHM _. +have prH: H \proper G by rewrite inE in maxH; exact: maxgroupp maxH. +have sAHM: A \subset H :&: M by rewrite subsetI sAH. +have [R sylR_HM sAR]:= Sylow_superset sAHM (pgroupS sAP pP). +have [/subsetIP[sRH sRM] pR _] := and3P sylR_HM. +have{sylR_HM} sylR_H: p.-Sylow(H) R. + have [Q sylQ] := Sylow_superset sRM pR; have [sQM pQ _] := and3P sylQ. + case/eqVproper=> [defR | /(nilpotent_proper_norm (pgroup_nil pQ)) sRN]. + apply: (pHall_subl sRH (subsetT _)); rewrite pHallE subsetT /=. + by rewrite -(card_Hall sylPG) (card_Hall sylP) defR (card_Hall sylQ). + case/maximal_exists: (subsetT 'N(R)) => [nRG | [D maxD sND]]. + case/negP: (proper_irrefl (mem G)); rewrite -{1}nRG. + rewrite mFT_norm_proper ?(mFT_pgroup_proper pR) //. + by rewrite -proper1G (proper_sub_trans _ sAR) ?proper1G. + move/implyP: (maxHM D); rewrite 2!inE {}maxD leqNgt. + case: eqP sND => [->{D} sNM _ | _ sND]. + rewrite -Sylow_subnorm (pHall_subl _ _ sylR_HM) ?setIS //. + by rewrite subsetI sRH normG. + rewrite (subset_trans (subset_trans sAR (normG R)) sND); case/negP. + rewrite -(card_Hall sylR_HM) (leq_trans (proper_card sRN)) //. + rewrite -(part_pnat_id (pgroupS (subsetIl _ _) pQ)) dvdn_leq //. + by rewrite partn_dvd ?cardG_gt0 // cardSg //= setIC setISS. +have Hp'1: 'O_p^'(H) = 1. + apply: p'nbyA_1 (pcore_pgroup _ _) (subset_trans sAH (gFnorm _ _)). + exact: sub_proper_trans (pcore_sub _ _) prH. +have nsZLR_H: 'Z('L(R)) <| H. + exact: Puig_center_normal (mFT_odd _) (mFT_sol prH) sylR_H _. +have ntZLR: 'Z('L(R)) != 1. + apply/eqP=> /(trivg_center_Puig_pgroup pR) R1. + by rewrite -subG1 -R1 sAR in ntA. +have defH: 'N('Z('L(R))) = H := mmax_normal maxH nsZLR_H ntZLR. +have{sylR_H} sylR: p.-Sylow(G) R. + rewrite -Sylow_subnorm setTI (pHall_subl _ _ sylR_H) ?normG //=. + by rewrite -defH (char_norm_trans (center_Puig_char R)). +have nsZLR_M: 'Z('L(R)) <| M. + have sylR_M := pHall_subl sRM (subsetT _) sylR. + exact: Puig_center_normal (mFT_odd _) solM sylR_M _. +case/eqP: neHM; apply: group_inj. +by rewrite -defH (mmax_normal maxM nsZLR_M). +Qed. + +(* This summarizes the two branches of B & G, Theorem 8.1. *) +Theorem Fitting_Uniqueness M : M \in 'M -> 'r('F(M)) >= 3 -> 'F(M)%G \in 'U. +Proof. +move=> maxM; have [p _ -> dimF3] := rank_witness 'F(M). +have prF: 'F(M) \proper G := sub_mmax_proper maxM (Fitting_sub M). +have [pF | npF] := boolP (p.-group 'F(M)). + have [P sylP] := Sylow_exists p M; have [sPM pP _] := and3P sylP. + have dimP3: 'r_p(P) >= 3. + rewrite (p_rank_Sylow sylP) (leq_trans dimF3) //. + by rewrite p_rankS ?Fitting_sub. + have [A] := rank3_SCN3 pP (mFT_odd _) dimP3. + by case/(SCN_Fitting_Uniqueness maxM pF)=> // _ sAF; exact: uniq_mmaxS. +case/p_rank_geP: dimF3 => A /setIdP[EpA dimA3]. +have [A0 maxA0 sAA0] := @maxgroup_exists _ [pred X in 'E_p('F(M))] _ EpA. +have [_ abelA] := pElemP EpA; have pmaxA0: A0 \in 'E*_p('F(M)) by rewrite inE. +case/pElemP: (maxgroupp maxA0) => sA0F; case/and3P=> _ cA0A0 _. +have dimA0_3: 'r_p(A0) >= 3. + by rewrite -(eqP dimA3) -(p_rank_abelem abelA) p_rankS. +have:= non_pcore_Fitting_Uniqueness maxM npF pmaxA0 dimA0_3. +exact: uniq_mmaxS (subsetIl _ _) prF. +Qed. + +End Eight. + diff --git a/mathcomp/odd_order/BGsection9.v b/mathcomp/odd_order/BGsection9.v new file mode 100644 index 0000000..dba9344 --- /dev/null +++ b/mathcomp/odd_order/BGsection9.v @@ -0,0 +1,470 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype path. +Require Import finset prime fingroup action automorphism quotient cyclic. +Require Import gproduct gfunctor pgroup center commutator gseries nilpotent. +Require Import sylow abelian maximal hall. +Require Import BGsection1 BGsection4 BGsection5 BGsection6. +Require Import BGsection7 BGsection8. + +(******************************************************************************) +(* This file covers B & G, section 9, i.e., the proof the Uniqueness *) +(* Theorem, along with the several variants and auxiliary results. Note that *) +(* this is the only file to import BGsection8. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope. + +Section Nine. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types H K L M A B P Q R : {group gT}. +Implicit Types p q r : nat. + +(* This is B & G, Theorem 9.1(b). *) +Theorem noncyclic_normed_sub_Uniqueness p M B : + M \in 'M -> B \in 'E_p(M) -> ~~ cyclic B -> + \bigcup_(K in |/|_G(B; p^')) K \subset M -> + B \in 'U. +Proof. +move=> maxM /pElemP[sBM abelB] ncycB snbBp'_M; have [pB cBB _] := and3P abelB. +have prM := mmax_proper maxM; have solM := mFT_sol prM. +apply/uniq_mmaxP; exists M; symmetry; apply/eqP. +rewrite eqEsubset sub1set inE maxM sBM; apply/subsetPn=> [[H0 MB_H0 neH0M]]. +have:= erefl [arg max_(H > H0 | (H \in 'M(B)) && (H :!=: M)) #|H :&: M|`_p]. +have [|H] := arg_maxP; first by rewrite MB_H0; rewrite inE in neH0M. +rewrite inE -andbA => /and3P[maxH sBH neHM] maxHM _ {H0 MB_H0 neH0M}. +have sB_HM: B \subset H :&: M by rewrite subsetI sBH. +have{sB_HM} [R sylR sBR] := Sylow_superset sB_HM pB. +have [/subsetIP[sRH sRM] pR _] := and3P sylR. +have [P sylP sRP] := Sylow_superset sRM pR; have [sPM pP _] := and3P sylP. +have sHp'M: 'O_p^'(H) \subset M. + apply: subset_trans snbBp'_M; rewrite (bigcup_max 'O_p^'(H)%G) // inE -andbA. + by rewrite subsetT pcore_pgroup (subset_trans sBH) ?gFnorm. +have{snbBp'_M} defMp': <<\bigcup_(K in |/|_G(P; p^')) K>> = 'O_p^'(M). + have nMp'M: M \subset 'N('O_p^'(M)) by exact: gFnorm. + have nMp'P := subset_trans sPM nMp'M. + apply/eqP; rewrite eqEsubset gen_subG sub_gen ?andbT; last first. + by rewrite (bigcup_max 'O_p^'(M)%G) // inE -andbA subsetT pcore_pgroup. + apply/bigcupsP=> K; rewrite inE -andbA => /and3P[_ p'K nKP]. + have sKM: K \subset M. + apply: subset_trans snbBp'_M; rewrite (bigcup_max K) // inE -andbA subsetT. + by rewrite p'K (subset_trans (subset_trans sBR sRP)). + rewrite -quotient_sub1 ?(subset_trans sKM) //=; set Mp' := 'O__(M). + have tiKp: 'O_p(M / Mp') :&: (K / _) = 1. + exact: coprime_TIg (pnat_coprime (pcore_pgroup _ _) (quotient_pgroup _ _)). + suffices sKMp: K / _ \subset 'O_p(M / Mp') by rewrite -(setIidPr sKMp) tiKp. + rewrite -Fitting_eq_pcore ?trivg_pcore_quotient //. + apply: subset_trans (cent_sub_Fitting (quotient_sol _ solM)). + rewrite subsetI quotientS //= (Fitting_eq_pcore (trivg_pcore_quotient _ _)). + rewrite (sameP commG1P trivgP) /= -/Mp' -tiKp subsetI commg_subl commg_subr. + rewrite (subset_trans (quotientS _ sKM)) ?gFnorm //=. + apply: subset_trans (pcore_sub_Hall (quotient_pHall nMp'P sylP)) _. + by rewrite quotient_norms. +have ntR: R :!=: 1 by case: eqP sBR ncycB => // -> /trivgP->; rewrite cyclic1. +have{defMp'} sNPM: 'N(P) \subset M. + have [Mp'1 | ntMp'] := eqVneq 'O_p^'(M) 1. + have nsZLP: 'Z('L(P)) <| M. + by apply: Puig_center_normal Mp'1 => //; apply: mFT_odd. + rewrite -(mmax_normal maxM nsZLP). + exact: char_norm_trans (center_Puig_char P) _. + apply: contraNneq ntR => /(trivg_center_Puig_pgroup pP) P1. + by rewrite -subG1 -P1. + rewrite -(mmax_normal maxM (pcore_normal _ _) ntMp') /= -defMp' norms_gen //. + apply/subsetP=> x nPx; rewrite inE sub_conjg; apply/bigcupsP=> K. + rewrite inE -andbA -sub_conjg => /and3P[_ p'K nKP]. + rewrite (bigcup_max (K :^ x)%G) // inE -andbA subsetT pgroupJ p'K /=. + by rewrite -(normP nPx) normJ conjSg. +have sylPG := mmax_sigma_Sylow maxM sylP sNPM. +have{sNPM} [sNRM sylRH]: 'N(R) \subset M /\ p.-Sylow(H) R. + have [defR | ltRP] := eqVproper sRP. + by split; rewrite defR // (pHall_subl _ (subsetT _)) // -defR. + have [| D /setIdP[maxD sND]]:= @mmax_exists _ 'N(R). + by rewrite mFT_norm_proper // (mFT_pgroup_proper pR). + have/implyP := maxHM D; rewrite inE {}maxD /= leqNgt. + rewrite (subset_trans (subset_trans sBR (normG R))) //= implybNN. + have ltRN := nilpotent_proper_norm (pgroup_nil pP) ltRP. + rewrite -(card_Hall sylR) (leq_trans (proper_card ltRN)) /=; last first. + rewrite setIC -(part_pnat_id (pgroupS (subsetIr _ _) pP)) dvdn_leq //. + by rewrite partn_dvd ?cardG_gt0 // cardSg // setISS. + move/eqP=> defD; rewrite defD in sND; split; rewrite // -Sylow_subnorm. + by rewrite (pHall_subl _ _ sylR) ?setIS // subsetI sRH normG. +have sFH_RHp': 'F(H) \subset R * 'O_p^'(H). + case/dprodP: (nilpotent_pcoreC p (Fitting_nil H)) => _ /= <- _ _. + by rewrite p_core_Fitting mulgSS ?(pcore_sub_Hall sylRH) ?pcore_Fitting. +have sFH_M: 'F(H) \subset M by rewrite (subset_trans sFH_RHp') ?mul_subG. +case/(H :=P: M): neHM; have [le3r | ge2r] := ltnP 2 'r('F(H)). + have [D uF_D] := uniq_mmaxP (Fitting_Uniqueness maxH le3r). + by rewrite (eq_uniq_mmax uF_D maxM) // (eq_uniq_mmax uF_D maxH) ?Fitting_sub. +have nHp'R: R \subset 'N('O_p^'(H)) by rewrite (subset_trans sRH) ?gFnorm. +have nsRHp'H: R <*> 'O_p^'(H) <| H. + rewrite sub_der1_normal //= ?join_subG ?sRH ?pcore_sub //. + rewrite norm_joinEl // (subset_trans _ sFH_RHp') //. + by rewrite rank2_der1_sub_Fitting ?mFT_odd // mFT_sol ?mmax_proper. +have sylR_RHp': p.-Sylow(R <*> 'O_p^'(H)) R. + by apply: (pHall_subl _ _ sylRH); rewrite ?joing_subl // normal_sub. +rewrite (mmax_max maxH) // -(Frattini_arg nsRHp'H sylR_RHp') /=. +by rewrite mulG_subG join_subG sRM sHp'M /= setIC subIset ?sNRM. +Qed. + +(* This is B & G, Theorem 9.1(a). *) +Theorem noncyclic_cent1_sub_Uniqueness p M B : + M \in 'M -> B \in 'E_p(M) -> ~~ cyclic B -> + \bigcup_(b in B^#) 'C[b] \subset M -> + B \in 'U. +Proof. +move=> maxM EpB ncycB sCB_M. +apply: (noncyclic_normed_sub_Uniqueness maxM EpB) => //. +apply/bigcupsP=> K; rewrite inE -andbA => /and3P[_ p'K nKB]. +case/pElemP: EpB => _ /and3P[pB cBB _]. +rewrite -(coprime_abelian_gen_cent1 cBB ncycB nKB); last first. + by rewrite coprime_sym (pnat_coprime pB). +rewrite gen_subG (subset_trans _ sCB_M) //. +by apply/bigcupsP=> b Bb; rewrite (bigcup_max b) // subsetIr. +Qed. + +(* This is B & G, Corollary 9.2. *) +Corollary cent_uniq_Uniqueness K L : + L \in 'U -> K \subset 'C(L) -> 'r(K) >= 2 -> K \in 'U. +Proof. +move=> uL; have ntL := uniq_mmax_neq1 uL. +case/uniq_mmaxP: uL => H uL_H cLK; have [maxH sLH] := mem_uniq_mmax uL_H. +case/rank_geP=> B /nElemP[p /pnElemP[sBK abelB /eqP dimB2]]. +have scBH: \bigcup_(b in B^#) 'C[b] \subset H. + apply/bigcupsP=> b /setIdP[]; rewrite inE -cycle_eq1 => ntb Bb. + apply: (sub_uniq_mmax uL_H); last by rewrite /= -cent_cycle mFT_cent_proper. + by rewrite sub_cent1 (subsetP cLK) ?(subsetP sBK). +have EpB: B \in 'E_p(H). + apply/pElemP; split=> //; rewrite -(setD1K (group1 B)) subUset sub1G /=. + apply/subsetP=> b Bb; apply: (subsetP scBH). + by apply/bigcupP; exists b => //; apply/cent1P. +have prK: K \proper G by rewrite (sub_proper_trans cLK) ?mFT_cent_proper. +apply: uniq_mmaxS prK (noncyclic_cent1_sub_Uniqueness _ EpB _ _) => //. +by rewrite (abelem_cyclic abelB) (eqP dimB2). +Qed. + +(* This is B & G, Corollary 9.3. *) +Corollary any_cent_rank3_Uniquness p A B : + abelian A -> p.-group A -> 'r(A) >= 3 -> A \in 'U -> + p.-group B -> ~~ cyclic B -> 'r_p('C(B)) >= 3 -> + B \in 'U. +Proof. +move=> cAA pA rA3 uA pB ncycB /p_rank_geP[C /= Ep3C]. +have [cBC abelC dimC3] := pnElemP Ep3C; have [pC cCC _] := and3P abelC. +have [P /= sylP sCP] := Sylow_superset (subsetT _) pC. +wlog sAP: A pA cAA rA3 uA / A \subset P. + move=> IHA; have [x _] := Sylow_Jsub sylP (subsetT _) pA. + by apply: IHA; rewrite ?pgroupJ ?abelianJ ?rankJ ?uniq_mmaxJ. +have ncycC: ~~ cyclic C by rewrite (abelem_cyclic abelC) dimC3. +have ncycP: ~~ cyclic P := contra (cyclicS sCP) ncycC. +have [D] := ex_odd_normal_p2Elem (pHall_pgroup sylP) (mFT_odd _) ncycP. +case/andP=> sDP nDP /pnElemP[_ abelD dimD2]. +have CADge2: 'r('C_A(D)) >= 2. + move: rA3; rewrite (rank_pgroup pA) => /p_rank_geP[E]. + case/pnElemP=> sEA abelE dimE3; apply: leq_trans (rankS (setSI _ sEA)). + rewrite (rank_abelem (abelemS (subsetIl _ _) abelE)) -(leq_add2r 1) addn1. + rewrite -dimE3 -leq_subLR -logn_div ?cardSg ?divgS ?subsetIl //. + rewrite logn_quotient_cent_abelem ?dimD2 //. + exact: subset_trans (subset_trans sAP nDP). +have CCDge2: 'r('C_C(D)) >= 2. + rewrite (rank_abelem (abelemS (subsetIl _ _) abelC)) -(leq_add2r 1) addn1. + rewrite -dimC3 -leq_subLR -logn_div ?cardSg ?divgS ?subsetIl //. + by rewrite logn_quotient_cent_abelem ?dimD2 //; apply: subset_trans nDP. +rewrite centsC in cBC; apply: cent_uniq_Uniqueness cBC _; last first. + by rewrite ltnNge (rank_pgroup pB) -odd_pgroup_rank1_cyclic ?mFT_odd. +have cCDC: C \subset 'C('C_C(D)) + by rewrite (sub_abelian_cent (abelem_abelian abelC)) ?subsetIl. +apply: cent_uniq_Uniqueness cCDC _; last by rewrite (rank_abelem abelC) dimC3. +apply: cent_uniq_Uniqueness (subsetIr _ _) CCDge2. +have cDCA: D \subset 'C('C_A(D)) by rewrite centsC subsetIr. +apply: cent_uniq_Uniqueness cDCA _; last by rewrite (rank_abelem abelD) dimD2. +by apply: cent_uniq_Uniqueness uA _ CADge2; rewrite subIset // -abelianE cAA. +Qed. + +(* This is B & G, Lemma 9.4. *) +Lemma any_rank3_Fitting_Uniqueness p M P : + M \in 'M -> 'r_p('F(M)) >= 3 -> p.-group P -> 'r(P) >= 3 -> P \in 'U. +Proof. +move=> maxM FMge3 pP; rewrite (rank_pgroup pP) => /p_rank_geP[B]. +case/pnElemP=> sBP abelB dimB3; have [pB cBB _] := and3P abelB. +have CBge3: 'r_p('C(B)) >= 3 by rewrite -dimB3 -(p_rank_abelem abelB) p_rankS. +have ncycB: ~~ cyclic B by rewrite (abelem_cyclic abelB) dimB3. +apply: {P pP}uniq_mmaxS sBP (mFT_pgroup_proper pP) _. +case/orP: (orbN (p.-group 'F(M))) => [pFM | pFM']. + have [P sylP sFP] := Sylow_superset (Fitting_sub _) pFM. + have pP := pHall_pgroup sylP. + have [|A SCN_A]:= rank3_SCN3 pP (mFT_odd _). + by rewrite (leq_trans FMge3) ?p_rankS. + have [_ _ uA] := SCN_Fitting_Uniqueness maxM pFM sylP FMge3 SCN_A. + case/setIdP: SCN_A => SCN_A dimA3; case: (setIdP SCN_A); case/andP=> sAP _ _. + have cAA := SCN_abelian SCN_A; have pA := pgroupS sAP pP. + exact: (any_cent_rank3_Uniquness cAA pA). +have [A0 EpA0 A0ge3] := p_rank_pmaxElem_exists FMge3. +have uA := non_pcore_Fitting_Uniqueness maxM pFM' EpA0 A0ge3. +case/pmaxElemP: EpA0; case/setIdP=> _ abelA0 _. +have [pA0 cA0A0 _] := and3P abelA0; rewrite -rank_pgroup // in A0ge3. +rewrite (any_cent_rank3_Uniquness _ pA0) // (cent_uniq_Uniqueness uA) 1?ltnW //. +by rewrite centsC subsetIr. +Qed. + +(* This is B & G, Lemma 9.5. *) +Lemma SCN_3_Uniqueness p A : A \in 'SCN_3[p] -> A \in 'U. +Proof. +move=> SCN3_A; apply/idPn=> uA'. +have [P] := bigcupP SCN3_A; rewrite inE => sylP /setIdP[SCN_A Age3]. +have [nsAP _] := setIdP SCN_A; have [sAP nAP] := andP nsAP. +have cAA := SCN_abelian SCN_A. +have pP := pHall_pgroup sylP; have pA := pgroupS sAP pP. +have ntA: A :!=: 1 by rewrite -rank_gt0 -(subnKC Age3). +have [p_pr _ [e oA]] := pgroup_pdiv pA ntA. +have{e oA} def_piA: \pi(A) =i (p : nat_pred). + by rewrite /= oA pi_of_exp //; exact: pi_of_prime. +have FmCAp_le2 M: M \in 'M('C(A)) -> 'r_p('F(M)) <= 2. + case/setIdP=> maxM cCAM; rewrite leqNgt; apply: contra uA' => Fge3. + exact: (any_rank3_Fitting_Uniqueness maxM Fge3). +have sNP_mCA M: M \in 'M('C(A)) -> 'N(P) \subset M. + move=> mCA_M; have Fple2 := FmCAp_le2 M mCA_M. + case/setIdP: mCA_M => maxM sCAM; set F := 'F(M) in Fple2. + have sNR_M R: A \subset R -> R \subset P :&: M -> 'N(R) \subset M. + move=> sAR /subsetIP[sRP sRM]. + pose q := if 'r(F) <= 2 then max_pdiv #|M| else s2val (rank_witness 'F(M)). + have nMqR: R \subset 'N('O_q(M)) := subset_trans sRM (gFnorm _ _). + have{nMqR} [Q maxQ sMqQ] := max_normed_exists (pcore_pgroup _ _) nMqR. + have [p'q sNQ_M]: q != p /\ 'N(Q) \subset M. + case/mem_max_normed: maxQ sMqQ; rewrite {}/q. + case: leqP => [Fle2 | ]; last first. + case: rank_witness => q /= q_pr -> Fge3 qQ _ sMqQ; split=> //. + by case: eqP Fge3 => // ->; rewrite ltnNge Fple2. + have Mqge3: 'r('O_q(M)) >= 3. + rewrite (rank_pgroup (pcore_pgroup _ _)) /= -p_core_Fitting. + by rewrite (p_rank_Sylow (nilpotent_pcore_Hall _ (Fitting_nil _))). + have uMq: 'O_q(M)%G \in 'U. + exact: (any_rank3_Fitting_Uniqueness _ Fge3 (pcore_pgroup _ _)). + have uMqM := def_uniq_mmax uMq maxM (pcore_sub _ _). + apply: sub_uniq_mmax (subset_trans sMqQ (normG _)) _ => //. + apply: mFT_norm_proper (mFT_pgroup_proper qQ). + by rewrite -rank_gt0 2?ltnW ?(leq_trans Mqge3) ?rankS. + set q := max_pdiv _ => qQ _ sMqQ. + have sylMq: q.-Sylow(M) 'O_q(M). + by rewrite [pHall _ _ _]rank2_max_pcore_Sylow ?mFT_odd ?mmax_sol. + have defNMq: 'N('O_q(M)) = M. + rewrite (mmax_normal maxM (pcore_normal _ _)) // -rank_gt0. + rewrite (rank_pgroup (pcore_pgroup _ _)) (p_rank_Sylow sylMq). + by rewrite p_rank_gt0 pi_max_pdiv cardG_gt1 mmax_neq1. + have sylMqG: q.-Sylow(G) 'O_q(M). + by rewrite (mmax_sigma_Sylow maxM) ?defNMq. + rewrite (sub_pHall sylMqG qQ) ?subsetT // defNMq; split=> //. + have: 'r_p(G) > 2. + by rewrite (leq_trans Age3) // (rank_pgroup pA) p_rankS ?subsetT. + apply: contraTneq => <-; rewrite -(p_rank_Sylow sylMqG). + rewrite -leqNgt -(rank_pgroup (pcore_pgroup _ _)) /=. + by rewrite -p_core_Fitting (leq_trans _ Fle2) // rankS ?pcore_sub. + have trCRq': [transitive 'O_p^'('C(R)), on |/|*(R; q) | 'JG]. + have cstrA: normed_constrained A. + by apply: SCN_normed_constrained sylP _; rewrite inE SCN_A ltnW. + have pR: p.-group R := pgroupS sRP pP. + have snAR: A <|<| R by rewrite (nilpotent_subnormal (pgroup_nil pR)). + have A'q: q \notin \pi(A) by rewrite def_piA. + rewrite -(eq_pgroup _ def_piA) in pR. + have [|?] := normed_trans_superset cstrA A'q snAR pR. + by rewrite (eq_pcore _ (eq_negn def_piA)) Thompson_transitivity. + by rewrite (eq_pcore _ (eq_negn def_piA)). + apply/subsetP=> x nRx; have maxQx: (Q :^ x)%G \in |/|*(R; q). + by rewrite (actsP (norm_acts_max_norm _ _)). + have [y cRy [defQx]] := atransP2 trCRq' maxQ maxQx. + rewrite -(mulgKV y x) groupMr. + by rewrite (subsetP sNQ_M) // inE conjsgM defQx conjsgK. + apply: subsetP cRy; apply: (subset_trans (pcore_sub _ _)). + exact: subset_trans (centS _) sCAM. + have sNA_M: 'N(A) \subset M. + by rewrite sNR_M // subsetI sAP (subset_trans cAA). + by rewrite sNR_M // subsetI subxx (subset_trans nAP). +pose P0 := [~: P, 'N(P)]. +have ntP0: P0 != 1. + apply/eqP=> /commG1P; rewrite centsC -(setIidPr (subsetT 'N(P))) /=. + case/(Burnside_normal_complement sylP)/sdprodP=> _ /= defG nGp'P _. + have prGp': 'O_p^'(G) \proper G. + rewrite properT; apply: contra ntA; move/eqP=> defG'. + rewrite -(setIidPl (subsetT A)) /= -defG'. + by rewrite coprime_TIg // (pnat_coprime pA (pcore_pgroup _ _)). + have ntGp': 'O_p^'(G) != 1. + apply: contraTneq (mFT_pgroup_proper pP); rewrite -{2}defG => ->. + by rewrite mul1g proper_irrefl. + by have:= mFT_norm_proper ntGp' prGp'; rewrite properE gFnorm andbF. +have sP0P: P0 \subset P by rewrite commg_subl. +have pP0: p.-group P0 := pgroupS sP0P pP. +have uNP0_mCA M: M \in 'M('C(A)) -> 'M('N(P0)) = [set M]. + move=> mCA_M; have [maxM sCAM] := setIdP mCA_M. + have sAM := subset_trans cAA sCAM. + pose F := 'F(M); pose D := 'O_p^'(F). + have cDP0: P0 \subset 'C(D). + have sA1A := Ohm_sub 1 A. + have nDA1: 'Ohm_1(A) \subset 'N(D). + apply: subset_trans sA1A (subset_trans sAM (char_norm _)). + exact: char_trans (pcore_char _ _) (Fitting_char _). + have abelA1: p.-abelem 'Ohm_1(A) by rewrite Ohm1_abelem. + have dimA1ge3: logn p #|'Ohm_1(A)| >= 3. + by rewrite -(rank_abelem abelA1) rank_Ohm1. + have coDA1: coprime #|D| #|'Ohm_1(A)|. + rewrite coprime_sym (coprimeSg sA1A) //. + exact: pnat_coprime pA (pcore_pgroup _ _). + rewrite centsC -[D](coprime_abelian_gen_cent (abelianS sA1A cAA) nDA1) //=. + rewrite gen_subG /= -/D; apply/bigcupsP=> B /and3P[cycqB sBA1 nBA1]. + have abelB := abelemS sBA1 abelA1; have sBA := subset_trans sBA1 sA1A. + have{cycqB} ncycB: ~~ cyclic B. + move: cycqB; rewrite (abelem_cyclic (quotient_abelem _ abelA1)). + rewrite card_quotient // -divgS // logn_div ?cardSg // leq_subLR addn1. + by move/(leq_trans dimA1ge3); rewrite ltnS ltnNge -(abelem_cyclic abelB). + have [x Bx sCxM']: exists2 x, x \in B^# & ~~ ('C[x] \subset M). + suff: ~~ (\bigcup_(x in B^#) 'C[x] \subset M). + case/subsetPn=> y /bigcupP[x Bx cxy] My'. + by exists x; last by apply/subsetPn; exists y. + have EpB: B \in 'E_p(M) by rewrite inE (subset_trans sBA sAM). + apply: contra uA' => sCB_M. + apply: uniq_mmaxS sBA (mFT_pgroup_proper pA) _. + exact: noncyclic_cent1_sub_Uniqueness maxM EpB ncycB sCB_M. + case/setD1P: Bx; rewrite -cycle_eq1 => ntx Bx. + have{ntx} [L /setIdP[maxL /=]] := mmax_exists (mFT_cent_proper ntx). + rewrite cent_cycle => sCxL. + have{sCxM'} neLM : L != M by case: eqP sCxL sCxM' => // -> ->. + have sNP_LM: 'N(P) \subset L :&: M. + rewrite subsetI !sNP_mCA // inE maxL (subset_trans _ sCxL) // -cent_set1. + by rewrite centS // sub1set (subsetP sBA). + have sP0_LM': P0 \subset (L :&: M)^`(1). + exact: subset_trans (commSg _ (normG _)) (dergS 1 sNP_LM). + have DLle2: 'r(D :&: L) <= 2. + apply: contraR neLM; rewrite -ltnNge -in_set1 => /rank_geP[E /nElemP[q]]. + rewrite /= -/D => /pnElemP[/subsetIP[sED sEL] abelE dimE3]. + have sEF: E \subset F := subset_trans sED (pcore_sub _ _). + have Fge3: 'r_q(F) >= 3 by rewrite -dimE3 -p_rank_abelem // p_rankS. + have qE := abelem_pgroup abelE. + have uE: E \in 'U. + apply: any_rank3_Fitting_Uniqueness Fge3 _ _ => //. + by rewrite (rank_pgroup qE) p_rank_abelem ?dimE3. + rewrite -(def_uniq_mmax uE maxM (subset_trans sEF (Fitting_sub _))). + by rewrite inE maxL. + have cDL_P0: P0 \subset 'C(D :&: L). + have nsDM: D <| M:= char_normal_trans (pcore_char _ _) (Fitting_normal M). + have{nsDM} [sDM nDM] := andP nsDM. + have sDL: D :&: L \subset L :&: M by rewrite setIC setIS. + have nsDL: D :&: L <| L :&: M by rewrite /normal sDL setIC normsIG. + have [s ch_s last_s_DL] := chief_series_exists nsDL. + have solLM := solvableS (subsetIl L M) (mmax_sol maxL). + have solDL := solvableS sDL solLM. + apply: (stable_series_cent (congr_group last_s_DL)) => //; first 1 last. + rewrite coprime_sym (coprimegS (subsetIl _ _)) //. + exact: pnat_coprime (pcore_pgroup _ _). + have{last_s_DL}: last 1%G s \subset D :&: L by rewrite last_s_DL. + rewrite /= -/P0; elim/last_ind: s ch_s => //= s U IHs. + rewrite !rcons_path last_rcons /=; set V := last _ s. + case/andP=> ch_s chUV sUDL; have [maxU _ nU_LM] := and3P chUV. + have{maxU} /andP[/andP[sVU _] nV_LM] := maxgroupp maxU. + have nVU := subset_trans sUDL (subset_trans sDL nV_LM). + rewrite IHs ?(subset_trans sVU) // /stable_factor /normal sVU nVU !andbT. + have nVP0 := subset_trans (subset_trans sP0_LM' (der_sub _ _)) nV_LM. + rewrite commGC -sub_astabQR // (subset_trans sP0_LM') //. + have /is_abelemP[q _ /andP[qUV _]]: is_abelem (U / V). + exact: sol_chief_abelem solLM chUV. + apply: rank2_der1_cent_chief qUV sUDL; rewrite ?mFT_odd //. + exact: leq_trans (p_rank_le_rank _ _) DLle2. + rewrite centsC (subset_trans cDL_P0) ?centS ?setIS //. + by rewrite (subset_trans _ sCxL) // -cent_set1 centS ?sub1set. + case: (ltnP 2 'r(F)) => [| Fle2]. + have [q q_pr -> /= Fq3] := rank_witness [group of F]. + have Mq3: 'r('O_q(M)) >= 3. + rewrite (rank_pgroup (pcore_pgroup _ _)) /= -p_core_Fitting. + by rewrite (p_rank_Sylow (nilpotent_pcore_Hall _ (Fitting_nil _))). + have uMq: 'O_q(M)%G \in 'U. + exact: any_rank3_Fitting_Uniqueness Fq3 (pcore_pgroup _ _) Mq3. + apply: def_uniq_mmaxS (def_uniq_mmax uMq maxM (pcore_sub q _)); last first. + exact: mFT_norm_proper ntP0 (mFT_pgroup_proper pP0). + rewrite cents_norm // centsC (subset_trans cDP0) ?centS //=. + rewrite -p_core_Fitting sub_pcore // => q1; move/eqnP=> ->{q1}. + by apply/eqnP=> def_q; rewrite ltnNge def_q FmCAp_le2 in Fq3. + rewrite (mmax_normal maxM) ?mmax_sup_id //. + have sNP_M := sNP_mCA M mCA_M; have sPM := subset_trans (normG P) sNP_M. + rewrite /normal comm_subG //= -/P0. + have nFP: P \subset 'N(F) by rewrite (subset_trans _ (gFnorm _ _)). + have <-: F <*> P * 'N_M(P) = M. + apply: Frattini_arg (pHall_subl (joing_subr _ _) (subsetT _) sylP). + rewrite -(quotientGK (Fitting_normal M)) /= norm_joinEr //= -/F. + rewrite -quotientK // cosetpre_normal -sub_abelian_normal ?quotientS //. + by rewrite sub_der1_abelian ?rank2_der1_sub_Fitting ?mFT_odd ?mmax_sol. + case/dprodP: (nilpotent_pcoreC p (Fitting_nil M)) => _ /= defF cDFp _. + rewrite norm_joinEr //= -{}defF -(centC cDFp) -/D p_core_Fitting /= -/F. + rewrite -!mulgA mul_subG //; first by rewrite cents_norm // centsC. + rewrite mulgA [_ * P]mulSGid ?pcore_sub_Hall 1?(pHall_subl _ (subsetT _)) //. + by rewrite mulSGid ?subsetI ?sPM ?normG // subIset // orbC normsRr. +have [M mCA_M] := mmax_exists (mFT_cent_proper ntA). +have [maxM sCAM] := setIdP mCA_M; have sAM := subset_trans cAA sCAM. +have abelA1: p.-abelem 'Ohm_1(A) by rewrite Ohm1_abelem. +have sA1A := Ohm_sub 1 A. +have EpA1: 'Ohm_1(A)%G \in 'E_p(M) by rewrite inE (subset_trans sA1A). +have ncycA1: ~~ cyclic 'Ohm_1(A). + rewrite (abelem_cyclic abelA1) -(rank_abelem abelA1) rank_Ohm1. + by rewrite -(subnKC Age3). +have [x A1x sCxM']: exists2 x, x \in 'Ohm_1(A)^# & ~~ ('C[x] \subset M). + suff: ~~ (\bigcup_(x in 'Ohm_1(A)^#) 'C[x] \subset M). + case/subsetPn=> y /bigcupP[x A1 cxy] My'. + by exists x; last by apply/subsetPn; exists y. + apply: contra uA' => sCA1_M. + apply: uniq_mmaxS sA1A (mFT_pgroup_proper pA) _. + exact: noncyclic_cent1_sub_Uniqueness maxM EpA1 ncycA1 sCA1_M. +case/setD1P: A1x; rewrite -cycle_eq1 => ntx A1x. +have: 'C[x] \proper G by rewrite -cent_cycle mFT_cent_proper. +case/mmax_exists=> L /setIdP[maxL sCxL]. +have mCA_L: L \in 'M('C(A)). + rewrite inE maxL (subset_trans _ sCxL) //= -cent_set1 centS // sub1set. + by rewrite (subsetP sA1A). +case/negP: sCxM'; have/uNP0_mCA := mCA_L. +by rewrite (uNP0_mCA M) // => /set1_inj->. +Qed. + +(* This is B & G, Theorem 9.6, first assertion; note that B & G omit the *) +(* (necessary!) condition K \proper G. *) +Theorem rank3_Uniqueness K : K \proper G -> 'r(K) >= 3 -> K \in 'U. +Proof. +move=> prK /rank_geP[B /nElemP[p /pnElemP[sBK abelB dimB3]]]. +have [pB cBB _] := and3P abelB. +suffices: B \in 'U by apply: uniq_mmaxS. +have [P sylP sBP] := Sylow_superset (subsetT _) pB. +have pP := pHall_pgroup sylP. +have [|A SCN3_A] := rank3_SCN3 pP (mFT_odd _). + by rewrite -dimB3 -(rank_abelem abelB) (rank_pgroup pB) p_rankS. +have [SCN_A Age3] := setIdP SCN3_A. +have: A \in 'SCN_3[p] by apply/bigcupP; exists P; rewrite // inE. +move/SCN_3_Uniqueness=> uA; have cAA := SCN_abelian SCN_A. +case/setIdP: SCN_A; case/andP=> sAP _ _; have pA := pgroupS sAP pP. +apply: any_cent_rank3_Uniquness uA pB _ _ => //. + by rewrite (abelem_cyclic abelB) dimB3. +by rewrite -dimB3 -p_rank_abelem ?p_rankS. +Qed. + +(* This is B & G, Theorem 9.6, second assertion *) +Theorem cent_rank3_Uniqueness K : 'r(K) >= 2 -> 'r('C(K)) >= 3 -> K \in 'U. +Proof. +move=> Kge2 CKge3; have cCK_K: K \subset 'C('C(K)) by rewrite centsC. +apply: cent_uniq_Uniqueness cCK_K _ => //. +apply: rank3_Uniqueness (mFT_cent_proper _) CKge3. +by rewrite -rank_gt0 ltnW. +Qed. + +(* This is B & G, Theorem 9.6, final observation *) +Theorem nonmaxElem2_Uniqueness p A : A \in 'E_p^2(G) :\: 'E*_p(G) -> A \in 'U. +Proof. +case/setDP=> EpA nmaxA; have [_ abelA dimA2]:= pnElemP EpA. +case/setIdP: EpA => EpA _; have [pA _] := andP abelA. +apply: cent_rank3_Uniqueness; first by rewrite -dimA2 -(rank_abelem abelA). +have [E maxE sAE] := pmaxElem_exists EpA. +have [/pElemP[_ abelE _]] := pmaxElemP maxE; have [pE cEE _] := and3P abelE. +have: 'r(E) <= 'r('C(A)) by rewrite rankS // (subset_trans cEE) ?centS. +apply: leq_trans; rewrite (rank_abelem abelE) -dimA2 properG_ltn_log //. +by rewrite properEneq; case: eqP maxE nmaxA => // => /group_inj-> ->. +Qed. + +End Nine. + diff --git a/mathcomp/odd_order/PFsection1.v b/mathcomp/odd_order/PFsection1.v new file mode 100644 index 0000000..dc5ce95 --- /dev/null +++ b/mathcomp/odd_order/PFsection1.v @@ -0,0 +1,809 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg finset fingroup morphism. +Require Import perm automorphism quotient action zmodp center commutator. +Require Import poly cyclic pgroup nilpotent matrix mxalgebra mxrepresentation. +Require Import vector falgebra fieldext ssrnum algC rat algnum galois. +Require Import classfun character inertia integral_char vcharacter. +Require ssrint. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 1: Preliminary results. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Local Notation algCF := [fieldType of algC]. + +Section Main. + +Variable gT : finGroupType. + +(* This is Peterfalvi (1.1). *) +Lemma odd_eq_conj_irr1 (G : {group gT}) t : + odd #|G| -> (('chi[G]_t)^*%CF == 'chi_t) = ('chi_t == 1). +Proof. +move=> OG; apply/eqP/eqP=> [Ht | ->]; last exact: cfConjC_cfun1. +pose a := (@Zp1 1). +have Aito: + is_action <[a]> (fun (t : Iirr G) v => if v == a then conjC_Iirr t else t). + split=> [[[|[]]] //= _ t1 t2 Hj |j [[|[]]] // HH1 [[|[]]] // HH2 ] //=. + by apply: (inv_inj (@conjC_IirrK _ _)). + by rewrite conjC_IirrK. +pose ito := Action Aito. +have Acto: + is_action <[a]> (fun (c : {set gT}) v => if v == a then c^-1%g else c). + split=> [[[|[]]] //= _ t1 t2 Hj |j [[|[]]] // HH1 [[|[]]] // HH2 ] //=. + by rewrite -[t1]invgK Hj invgK. + by rewrite invgK. +pose cto := Action Acto. +have F1: [acts <[a]>, on (classes G) | cto]. + apply/subsetP=> j Hj. + rewrite !inE Hj; apply/subsetP=> u. + case/imsetP=> g GiG ->. + by rewrite inE /=; case: (_ == _) => //; + rewrite -?classVg mem_classes // ?groupV. +have F2 u x y: x \in G -> y \in cto (x ^: G) a -> 'chi_u x = 'chi_(ito u a) y. + rewrite mem_invg -{2}[y]invgK => Gx {y}/imsetP[y Gy ->]. + by rewrite -conjVg cfunJ {y Gy}//= conjC_IirrE cfunE -irr_inv invgK. +have F3: forall c, c \in classes G -> c^-1%g = c -> c = 1%g. + move=> c; case/imsetP => g GiG ->; rewrite -classVg => Hg. + move: (class_refl G g^-1); rewrite Hg; case/imsetP=> x XiG Hx. + have F4: (x ^+ 2)%g \in 'C_G[g]. + apply/subcent1P; split; rewrite ?groupM //. + apply: (mulgI (x * x * g)^-1)%g. + rewrite mulVg !invMg Hx conjgE !mulgA mulgK. + rewrite -[(_ * g * x)%g]mulgA -[(_ * (g * _))%g]mulgA -conjgE. + by rewrite -Hx mulgK mulVg. + have F5 : x \in 'C_G[g]. + suff->: (x = (x ^+ 2) ^+ (#|G| %/2).+1)%g by apply: groupX. + rewrite -expgM -[(_%/_).+1]addn1 mulnDr muln1 -{3}addn1 addnA. + move: (modn2 #|G|); rewrite {1}OG /= => HH; rewrite -{3}HH. + rewrite [(2 * _)%N]mulnC -divn_eq expgD expg1. + by move: (order_dvdG XiG); rewrite order_dvdn; move/eqP->; rewrite mul1g. + move: Hx; rewrite conjgE; case/subcent1P: F5=> _ ->. + rewrite mulgA mulVg mul1g => HH. + have F6: (g ^+ 2 == 1)%g by rewrite expgS -{1}HH expg1 mulVg. + suff: #[g] == 1%N by rewrite order_eq1; move/eqP->; apply: class1G. + move: F6 (order_gt0 g) (order_dvdG GiG); rewrite -order_dvdn. + move/(dvdn_leq (isT : (0 < 2)%N)); case: #[_]=> // [[|[]]] //. + by rewrite dvdn2 OG. +apply/eqP; case: (boolP (t == 0))=> // Hd. + by move/eqP: Hd->; rewrite irr0. +have:= card_afix_irr_classes (cycle_id a) F1 F2. +have->: #|'Fix_(classes G | cto)[a]| = 1%N. + apply: (@eq_card1 _ 1%g)=> c; apply/idP/idP; rewrite !inE. + case/andP=> GiG HH; apply/eqP; apply: F3=> //; apply/eqP. + by move/subsetP: HH; move/(_ a); rewrite !inE eqxx; apply. + move/eqP->; rewrite classes1. + apply/subsetP=> b; rewrite !inE; move/eqP=> -> /=. + by rewrite invg1. +rewrite (cardD1 (0 : Iirr _)). +have->: 0 \in 'Fix_ito[a]. + apply/afixP=> b; rewrite !inE; move/eqP->; rewrite /=. + apply: irr_inj; apply/cfunP=> g. + by rewrite conjC_IirrE cfConjCE irr0 cfun1E conjC_nat. +rewrite (cardD1 t) //. +suff->: t \in [predD1 'Fix_ito[a] & 0] by []. +rewrite inE /= Hd. +apply/afixP=> b; rewrite !inE; move/eqP->; rewrite /=. +apply: irr_inj; apply/cfunP=> g. +by rewrite conjC_IirrE Ht. +Qed. + +Variables G H : {group gT}. + +(* This is Peterfalvi (1.2). *) +Lemma not_in_ker_char0 t g : g \in G -> + H <| G -> ~~ (H \subset cfker 'chi[G]_t) -> 'C_H[g] = 1%g -> 'chi_t g = 0. +Proof. +move=> GiG HnG nHsC CH1. +have: (#|'C_G[g]| <= #|'C_(G/H)[coset H g]|)%N. + suff->: #|'C_G[g]| = #|'C_G[g] / H|%G. + by apply: (subset_leq_card (quotient_subcent1 H G g)). + apply: card_isog. + apply: isog_trans (second_isog _); last first. + apply: subset_trans (normal_norm HnG). + by apply: subcent1_sub. + suff->: H :&: 'C_G[g] = 1%g by exact: quotient1_isog. + rewrite -CH1. + apply/setP=> h; rewrite inE. + apply/andP/subcent1P; case=> H1 H2; split=> //. + by move/subcent1P: H2; case. + apply/subcent1P; split=> //. + by apply: (subsetP (normal_sub HnG)). +have F1: coset H g \in (G / H)%g by exact: mem_quotient. +rewrite -leC_nat. +have:= second_orthogonality_relation g GiG. +rewrite mulrb class_refl => <-. +have:= second_orthogonality_relation (coset H g) F1. +rewrite mulrb class_refl => <-; rewrite -!(eq_bigr _ (fun _ _ => normCK _)). +rewrite sum_norm_irr_quo // (bigID (fun i => H \subset cfker 'chi_i)) //=. +set S := \sum_(i | ~~ _) _; set S' := \sum_(i | _) _ => HH. +have /eqP F2: S == 0. + rewrite eqr_le -(ler_add2l S') addr0 HH /=. + by apply: sumr_ge0 => j _; rewrite mulr_ge0 ?normr_ge0. +apply/eqP; have: `|'chi_t g| ^+ 2 == 0. + apply/eqP; apply: (psumr_eq0P _ F2) nHsC => j _. + by rewrite mulr_ge0 ?normr_ge0. +by rewrite mulf_eq0 orbb normr_eq0. +Qed. + +(* This is Peterfalvi (1.3)(a). *) +Lemma equiv_restrict_compl A m (Phi : m.-tuple 'CF(H)) (mu : 'CF(G)) d : + H \subset G -> A <| H -> basis_of 'CF(H, A) Phi -> + ({in A, mu =1 \sum_i d i *: 'chi_i} <-> + (forall j : 'I_m, + \sum_i '[Phi`_j, 'chi_i] * (d i)^* = '['Ind[G] Phi`_j, mu])). +Proof. +move=> sHG nsAH BPhi; have [sAH nAH] := andP nsAH. +have APhi (i : 'I_m) : Phi`_i \in 'CF(H, A). + by apply: (basis_mem BPhi _); apply: mem_nth; rewrite size_tuple. +pose D := 'Res[H] mu - \sum_i d i *: 'chi_i. +transitivity (D \in 'CF(H, H :\: A)). + split=> [A'D | /cfun_onP A'D x Ax]. + apply/cfun_onP=> x; rewrite inE negb_and negbK. + case/orP=> [Ax | /cfun0-> //]; rewrite !cfunE -A'D //. + by rewrite cfResE ?subrr ?(subsetP sAH). + have:= A'D x; rewrite !cfunE !inE Ax => /(_ isT)/(canRL (subrK _)). + by rewrite add0r cfResE // ?(subsetP sAH). +have F0 (j : 'I_m) : + (\sum_i '[Phi`_j, 'chi_i] * (d i)^* == '['Ind Phi`_j, mu]) + = ('[Phi`_j, D] == 0). + rewrite raddfB raddf_sum /= Frobenius_reciprocity subr_eq0 eq_sym. + by congr (_ == _); apply: eq_bigr=> i _; rewrite cfdotZr mulrC. +split=> [HH j | HH]. + by apply/eqP; rewrite F0; apply/eqP; apply: cfdot_complement. +have{F0} F1 (j : 'I_m) : '[Phi`_j, D]_H = 0. + by have/eqP := HH j; rewrite F0 => /eqP. +have: (D \in 'CF(H))%VS by rewrite memvf. +rewrite -(cfun_complement nsAH) => /memv_addP[f Cf [g Cg defD]]. +have: '[f, f + g] = 0. + rewrite -defD (coord_basis BPhi Cf) cfdot_suml. + by rewrite big1 // => i _; rewrite cfdotZl F1 mulr0. +rewrite raddfD /= {1}(cfdot_complement Cf Cg) addr0 => /eqP. +by rewrite cfnorm_eq0 defD => /eqP->; rewrite add0r. +Qed. + +(* This is Peterfalvi (1.3)(b). *) +Lemma equiv_restrict_compl_ortho A m (Phi : m.-tuple 'CF(H)) mu_ : + H \subset G -> A <| H -> basis_of 'CF(H, A) Phi -> + (forall i j, '[mu_ i, mu_ j] = (i == j)%:R) -> + (forall j : 'I_m, 'Ind[G] Phi`_j = \sum_i '[Phi`_j, 'chi_i] *: mu_ i) -> + [/\ forall i, {in A, mu_ i =1 'chi_i} + & forall mu, (forall i, '[mu, mu_ i] = 0) -> {in A, forall x, mu x = 0}]. +Proof. +move=> HsG nsAH /equiv_restrict_compl Phi_A Mo IP; split=> [/= i | mu Cmu x Ax]. + have->: 'chi[H]_i = \sum_j (j == i)%:R *: 'chi_j. + rewrite (bigD1 i) //= eqxx scale1r big1 ?addr0 // => j /negPf->. + by rewrite scale0r. + apply/Phi_A=> // j; rewrite IP cfdot_suml. + by apply: eq_bigr=> k _; rewrite cfdotZl rmorph_nat Mo. +transitivity ((\sum_j 0 *: 'chi[H]_j) x); last first. + by rewrite sum_cfunE big1 // => j _; rewrite cfunE mul0r. +move: x Ax; apply/Phi_A=> // j. +rewrite -mulr_suml rmorph0 mulr0 IP cfdot_suml big1 // => k _. +by rewrite cfdotZl [d in _ * d]cfdotC Cmu rmorph0 mulr0. +Qed. + +Let vchar_isometry_base3 f f' : + f \in 'Z[irr G, G^#] -> '[f]_G = 2%:R -> + f' \in 'Z[irr G, G^#] -> '[f']_G = 2%:R -> + '[f, f'] = 1 -> + exists es : _ * bool, let: (i, j, k, epsilon) := es in + [/\ f = (-1) ^+ epsilon *: ('chi_j - 'chi_i), + f' = (-1) ^+ epsilon *: ('chi_j - 'chi_k) + & uniq [:: i; j; k]]. +Proof. +move=> Hf H2f Hf1 H2f1. +have [j [i neq_ij ->]] := vchar_norm2 Hf H2f. +have [j' [k neq_kj' ->]] := vchar_norm2 Hf1 H2f1. +rewrite cfdotBl !cfdotBr !cfdot_irr opprB addrAC !addrA. +do 2!move/(canRL (subrK _)); rewrite -(natrD _ 1) -!natrD => /eqP. +rewrite eqr_nat; have [eq_jj' | neq_jj'] := altP (j =P j'). + rewrite (eq_sym j) -eq_jj' {1}eq_jj' (negbTE neq_ij) (negbTE neq_kj'). + rewrite eqSS (can_eq oddb) => /eqP neq_ik; exists (i, j, k, false). + by rewrite !scaler_sign /= !inE neq_ik orbF neq_ij eq_sym eq_jj' neq_kj'. +case: (i =P k) => // eq_ik; exists (j, i, j', true). +rewrite !scaler_sign !opprB /= !inE eq_sym negb_or neq_ij neq_jj'. +by rewrite eq_ik neq_kj'. +Qed. + +Let vchar_isometry_base4 (eps : bool) i j k n m : + let f1 := 'chi_j - 'chi_i in + let f2 := 'chi_k - 'chi_i in + let f3 := 'chi_n - 'chi_m in + j != k -> '[f3, f1]_G = (-1) ^+ eps -> '[f3, f2] = (-1) ^+ eps -> + if eps then n == i else m == i. +Proof. +move=> /= Hjk; wlog ->: eps n m / eps = false. + case: eps; last exact; move/(_ false m n)=> IH nm_ji nm_ki. + by apply: IH; rewrite // -opprB cfdotNl (nm_ji, nm_ki) opprK. +rewrite !cfdotBl !cfdotBr !cfdot_irr !opprB addrAC addrA. +do 2!move/(canRL (subrK _)); rewrite -(natrD _ 1) -!natrD. +move/(can_inj natCK); case: (m == i) => //. +case: eqP => // ->; case: (j == i) => // _. +rewrite subr0 add0r => /(canRL (subrK _)); rewrite -(natrD _ 1). +by move/(can_inj natCK); rewrite (negbTE Hjk). +Qed. + +(* This is Peterfalvi (1.4). *) +Lemma vchar_isometry_base m L (Chi : m.-tuple 'CF(H)) + (tau : {linear 'CF(H) -> 'CF(G)}) : + (1 < m)%N -> {subset Chi <= irr H} -> free Chi -> + (forall chi, chi \in Chi -> chi 1%g = Chi`_0 1%g) -> + (forall i : 'I_m, (Chi`_i - Chi`_0) \in 'CF(H, L)) -> + {in 'Z[Chi, L], isometry tau, to 'Z[irr G, G^#]} -> + exists2 mu : m.-tuple (Iirr G), + uniq mu + & exists epsilon : bool, forall i : 'I_m, + tau (Chi`_i - Chi`_0) = (-1) ^+ epsilon *: ('chi_(mu`_i) - 'chi_(mu`_0)). +Proof. +case: m Chi => [|[|m]] // Chi _ irrChi Chifree Chi1 ChiCF [iso_tau Ztau]. +rewrite -(tnth_nth 0 _ 0); set chi := tnth Chi. +have chiE i: chi i = Chi`_i by rewrite -tnth_nth. +have inChi i: chi i \in Chi by exact: mem_tnth. +have{irrChi} irrChi i: chi i \in irr H by exact: irrChi. +have eq_chi i j: (chi i == chi j) = (i == j). + by rewrite /chi !(tnth_nth 0) nth_uniq ?size_tuple ?free_uniq. +have dot_chi i j: '[chi i, chi j] = (i == j)%:R. + rewrite -eq_chi; have [/irrP[{i}i ->] /irrP[{j}j ->]] := (irrChi i,irrChi j). + by rewrite cfdot_irr inj_eq //; exact: irr_inj. +pose F i j := chi i - chi j. +have DF i j : F i j = F i 0 - F j 0 by rewrite /F opprB addrA subrK. +have ZF i j: F i j \in 'Z[Chi, L]. + by rewrite zchar_split rpredB ?mem_zchar // DF memvB // /F !chiE. +have htau2 i j: i != j -> '[tau (F i j)] = 2%:R. + rewrite iso_tau // cfnormB -cfdotC !dot_chi !eqxx eq_sym => /negbTE->. + by rewrite -!natrD subr0. +have htau1 i j: j != 0 -> j != i -> i != 0 -> '[tau (F i 0), tau (F j 0)] = 1. + rewrite iso_tau // cfdotBl !cfdotBr opprB !dot_chi !(eq_sym j). + by do 3!move/negbTE->; rewrite !subr0 add0r. +have [m0 | nz_m] := boolP (m == 0%N). + rewrite -2!eqSS eq_sym in m0; move: (htau2 1 0 isT). + case/(vchar_norm2 (Ztau _ (ZF 1 0))) => [k1 [k0 neq_k01 eq_mu]]. + pose mu := @Tuple _ _ [:: k0; k1] m0. + exists mu; first by rewrite /= andbT inE. + exists false => i; rewrite scale1r chiE. + have: (i : nat) \in iota 0 2 by rewrite mem_iota (eqP m0) (valP i). + rewrite !inE; case/pred2P=> ->; first by rewrite !subrr linear0. + by rewrite -eq_mu /F !chiE. +have m_gt2: (2 < m.+2)%N by rewrite !ltnS lt0n. +pose i2 := Ordinal m_gt2. +case: (@vchar_isometry_base3 (tau (F 1 0)) (tau (F i2 0))); auto. +case=> [[[k1 k0] k2] e] []; set d := (-1) ^+ e => eq10 eq20. +rewrite /= !inE => /and3P[/norP[nek10 nek12]]; rewrite eq_sym => nek20 _. +have muP i: + {k | (i == 0) ==> (k == k0) & tau (F i 0) == d *: ('chi_k0 - 'chi_k)}. +- apply: sig2W; have [-> | nei0] := altP (i =P 0). + by exists k0; rewrite ?eqxx // /F !subrr !linear0. + have /(vchar_norm2 (Ztau _ (ZF i 0)))[k [k' nekk' eqFkk']] := htau2 i 0 nei0. + have [-> | neq_i1] := eqVneq i 1; first by exists k1; rewrite // -eq10. + have [-> | neq_i2] := eqVneq i i2; first by exists k2; rewrite // -eq20. + have:= @vchar_isometry_base4 (~~ e) k0 k1 k2 k k' nek12. + have ZdK u v w: '[u, v - w]_G = (-1) ^+ (~~ e) * '[u, d *: (w - v)]. + rewrite cfdotZr rmorph_sign mulrA -signr_addb addNb addbb mulN1r. + by rewrite -cfdotNr opprB. + rewrite -eqFkk' ZdK -eq10 {}ZdK -eq20 !htau1 //; try by rewrite eq_sym. + move/(_ (mulr1 _) (mulr1 _)); rewrite /d eqFkk'. + by case e => /eqP <-; [exists k | exists k']; rewrite ?scaler_sign ?opprB. +pose mu := [tuple of [seq s2val (muP i) | i <- ord_tuple m.+2]]; exists mu. + rewrite map_inj_uniq ?enum_uniq // => i j. + case: (muP i) (muP j) => /= ki _ /eqP eq_i0 [/= kj _ /eqP eq_j0] eq_kij. + apply/eqP; rewrite -eq_chi -subr_eq0 -cfnorm_eq0 -iso_tau ?ZF //. + rewrite -[chi i](subrK (chi 0)) -addrA linearD eq_i0 eq_kij -eq_j0. + by rewrite -linearD -opprB subrr !raddf0. +exists (~~ e) => i; rewrite -addbT signr_addb -/d -scalerA scaleN1r opprB. +rewrite -!tnth_nth -/(F i 0) tnth_map tnth_ord_tuple. +suffices /= ->: mu`_0 = k0 by case: (muP i) => /= k _ /eqP. +rewrite -(tnth_nth 0 _ 0) tnth_map tnth_ord_tuple. +by case: (muP 0) => /= k /(k =P k0). +Qed. + +(* This is Peterfalvi (1.5)(a). *) +Lemma cfResInd_sum_cfclass t : H <| G -> + 'Res[H] ('Ind[G] 'chi_t) + = #|'I_G['chi_t] : H|%:R *: \sum_(xi <- ('chi_t ^: G)%CF) xi. +Proof. +set T := 'I_G['chi_t] => nsHG; have [sHG nHG] := andP nsHG. +apply/cfun_inP=> h Hh; rewrite cfResE ?cfIndE // cfunE sum_cfunE. +apply: (canLR (mulKf (neq0CG H))). +rewrite mulrA -natrM Lagrange ?sub_Inertia //= -/T reindex_cfclass //=. +rewrite mulr_sumr [s in _ = s]big_mkcond /= (reindex_inj invg_inj). +rewrite (partition_big (conjg_Iirr t) xpredT) //=; apply: eq_bigr => i _. +have [[y Gy chi_i] | not_i_t] := cfclassP _ _ _; last first. + apply: big1 => z; rewrite groupV => /andP[Gz /eqP def_i]. + by case: not_i_t; exists z; rewrite // -def_i conjg_IirrE. +rewrite -(card_rcoset _ y) mulr_natl -sumr_const; apply: eq_big => z. + rewrite -(inj_eq irr_inj) conjg_IirrE chi_i mem_rcoset inE groupMr ?groupV //. + apply: andb_id2l => Gz; rewrite eq_sym (cfConjg_eqE _ nsHG) //. + by rewrite mem_rcoset inE groupM ?groupV. +rewrite groupV => /andP[Gz /eqP <-]. +by rewrite conjg_IirrE cfConjgE ?(subsetP nHG). +Qed. + +(* This is Peterfalvi (1.5)(b), main formula. *) +Lemma cfnorm_Ind_irr t : + H <| G -> '['Ind[G] 'chi[H]_t] = #|'I_G['chi_t] : H|%:R. +Proof. +set r := _%:R => HnG; have HsG := normal_sub HnG. +rewrite -Frobenius_reciprocity cfResInd_sum_cfclass //= cfdotZr rmorph_nat -/r. +rewrite reindex_cfclass // cfdot_sumr (bigD1 t) ?cfclass_refl //= cfnorm_irr. +rewrite big1 ?addr0 ?mulr1 // => j /andP[_ /negbTE]. +by rewrite eq_sym cfdot_irr => ->. +Qed. + +(* This is Peterfalvi (1.5)(b), irreducibility remark. *) +Lemma inertia_Ind_irr t : + H <| G -> 'I_G['chi[H]_t] \subset H -> 'Ind[G] 'chi_t \in irr G. +Proof. +rewrite -indexg_eq1 => nsHG /eqP r1. +by rewrite irrEchar cfInd_char ?irr_char //= cfnorm_Ind_irr ?r1. +Qed. + +(* This is Peterfalvi (1.5)(c). *) +Lemma cfclass_Ind_cases t1 t2 : H <| G -> + if 'chi_t2 \in ('chi[H]_t1 ^: G)%CF + then 'Ind[G] 'chi_t1 = 'Ind[G] 'chi_t2 + else '['Ind[G] 'chi_t1, 'Ind[G] 'chi_t2] = 0. +Proof. +move=> nsHG; have [/cfclass_Ind-> // | not_ch1Gt2] := ifPn. +rewrite -Frobenius_reciprocity cfResInd_sum_cfclass // cfdotZr rmorph_nat. +rewrite cfdot_sumr reindex_cfclass // big1 ?mulr0 // => j; rewrite cfdot_irr. +case: eqP => // <- /idPn[]; apply: contra not_ch1Gt2 => /cfclassP[y Gy ->]. +by apply/cfclassP; exists y^-1%g; rewrite ?groupV ?cfConjgK. +Qed. + +(* Useful consequences of (1.5)(c) *) +Lemma not_cfclass_Ind_ortho i j : + H <| G -> ('chi_i \notin 'chi_j ^: G)%CF -> + '['Ind[G, H] 'chi_i, 'Ind[G, H] 'chi_j] = 0. +Proof. by move/(cfclass_Ind_cases i j); rewrite cfclass_sym; case: ifP. Qed. + +Lemma cfclass_Ind_irrP i j : + H <| G -> + reflect ('Ind[G, H] 'chi_i = 'Ind[G, H] 'chi_j) ('chi_i \in 'chi_j ^: G)%CF. +Proof. +move=> nsHG; have [sHG _] := andP nsHG. +case: ifP (cfclass_Ind_cases j i nsHG) => [|_ Oji]; first by left. +right=> eq_chijG; have /negP[]: 'Ind[G] 'chi_i != 0 by exact: Ind_irr_neq0. +by rewrite -cfnorm_eq0 {1}eq_chijG Oji. +Qed. + +Lemma card_imset_Ind_irr (calX : {set Iirr H}) : + H <| G -> {in calX, forall i, 'Ind 'chi_i \in irr G} -> + {in calX & G, forall i y, conjg_Iirr i y \in calX} -> + #|calX| = (#|G : H| * #|[set cfIirr ('Ind[G] 'chi_i) | i in calX]|)%N. +Proof. +move=> nsHG irrIndX sXGX; have [sHG _] := andP nsHG; set f := fun i => cfIirr _. +rewrite -sum1_card (partition_big_imset f) /= mulnC -sum_nat_const. +apply: eq_bigr => _ /imsetP[i Xi ->]; transitivity (size (cfclass 'chi_i G)). + rewrite -sum1_size reindex_cfclass //; apply: eq_bigl => j. + case Xj: (j \in calX). + rewrite -(inj_eq irr_inj) !(cfIirrPE irrIndX) //. + exact/eqP/cfclass_Ind_irrP. + apply/esym/(contraFF _ Xj)=> /cfclassP[y Gy Dj]. + by rewrite -conjg_IirrE in Dj; rewrite (irr_inj Dj) sXGX. +rewrite -(Lagrange_index (Inertia_sub G 'chi_i)) ?sub_Inertia //. +rewrite -size_cfclass ((#|_ : _| =P 1)%N _) ?muln1 // -eqC_nat. +by rewrite -cfnorm_Ind_irr // -(cfIirrPE irrIndX) ?cfnorm_irr. +Qed. + +(* This is Peterfalvi (1.5)(d). *) +Lemma scaled_cfResInd_sum_cfclass t : H <| G -> + let chiG := 'Ind[G] 'chi_t in + (chiG 1%g / '[chiG]) *: 'Res[H] chiG + = #|G : H|%:R *: (\sum_(xi <- ('chi_t ^: G)%CF) xi 1%g *: xi). +Proof. +move=> nsHG chiG; have [sHG _] := andP nsHG. +rewrite cfResInd_sum_cfclass // cfnorm_Ind_irr // scalerA cfInd1 //. +rewrite divfK ?pnatr_eq0 -?lt0n // -scalerA linear_sum !reindex_cfclass //=. +congr (_ *: _); apply: eq_bigr => _ /cfclassP[y _ ->]. +by rewrite cfConjg1. +Qed. + +(* This is Peterfalvi (1.5)(e). *) +Lemma odd_induced_orthogonal t : + H <| G -> odd #|G| -> t != 0 -> + '['Ind[G, H] 'chi_t, ('Ind[G] 'chi_t)^*] = 0. +Proof. +move=> nsHG oddG nz_t; have [sHG _] := andP nsHG. +have:= cfclass_Ind_cases t (conjC_Iirr t) nsHG. +rewrite conjC_IirrE conj_cfInd; case: cfclassP => // [[g Gg id_cht]]. +have oddH: odd #|H| := pgroup.oddSg sHG oddG. +case/eqP: nz_t; apply: irr_inj; rewrite irr0. +apply/eqP; rewrite -odd_eq_conj_irr1 // id_cht; apply/eqP. +have F1: ('chi_t ^ (g ^+ 2))%CF = 'chi_t. + rewrite (cfConjgM _ nsHG) // -id_cht -conj_cfConjg -id_cht. + exact: cfConjCK. +suffices /eqP->: g == ((g ^+ 2) ^+ #|G|./2.+1)%g. + elim: _./2.+1 => [|n IHn]; first exact: cfConjgJ1. + by rewrite expgS (cfConjgM _ nsHG) ?groupX // F1. +rewrite eq_mulVg1 expgS -expgM mul2n -mulgA mulKg -expgS -order_dvdn. +by rewrite -add1n -[1%N](congr1 nat_of_bool oddG) odd_double_half order_dvdG. +Qed. + +(* This is Peterfalvi (1.6)(a). *) +Lemma sub_cfker_Ind_irr A i : + H \subset G -> G \subset 'N(A) -> + (A \subset cfker ('Ind[G, H] 'chi_i)) = (A \subset cfker 'chi_i). +Proof. by move=> sHG nAG; rewrite cfker_Ind_irr ?sub_gcore. Qed. + +(* Some consequences and related results. *) +Lemma sub_cfker_Ind (A : {set gT}) chi : + A \subset H -> H \subset G -> G \subset 'N(A) -> chi \is a character -> + (A \subset cfker ('Ind[G, H] chi)) = (A \subset cfker chi). +Proof. +move=> sAH sHG nAG Nchi; have [-> | nz_chi] := eqVneq chi 0. + by rewrite raddf0 !cfker_cfun0 !(subset_trans sAH). +by rewrite cfker_Ind ?sub_gcore. +Qed. + +Lemma cfInd_irr_eq1 i : + H <| G -> ('Ind[G, H] 'chi_i == 'Ind[G, H] 1) = (i == 0). +Proof. +case/andP=> sHG nHG; apply/eqP/idP=> [chi1 | /eqP->]; last by rewrite irr0. +rewrite -subGcfker -(sub_cfker_Ind_irr _ sHG nHG) chi1 -irr0. +by rewrite sub_cfker_Ind_irr ?cfker_irr0. +Qed. + +Lemma sub_cfker_constt_Res_irr (A : {set gT}) i j : + j \in irr_constt ('Res[H, G] 'chi_i) -> + A \subset H -> H \subset G -> G \subset 'N(A) -> + (A \subset cfker 'chi_j) = (A \subset cfker 'chi_i). +Proof. +move=> iHj sAH sHG nAG; apply/idP/idP=> kerA. + have jGi: i \in irr_constt ('Ind 'chi_j) by rewrite constt_Ind_Res. + rewrite (subset_trans _ (cfker_constt _ jGi)) ?cfInd_char ?irr_char //=. + by rewrite sub_cfker_Ind_irr. +rewrite (subset_trans _ (cfker_constt _ iHj)) ?cfRes_char ?irr_char //=. +by rewrite cfker_Res ?irr_char // subsetI sAH. +Qed. + +Lemma sub_cfker_constt_Ind_irr (A : {set gT}) i j : + i \in irr_constt ('Ind[G, H] 'chi_j) -> + A \subset H -> H \subset G -> G \subset 'N(A) -> + (A \subset cfker 'chi_j) = (A \subset cfker 'chi_i). +Proof. by rewrite constt_Ind_Res; apply: sub_cfker_constt_Res_irr. Qed. + +(* This is a stronger version of Peterfalvi (1.6)(b). *) +Lemma cfIndMod (K : {group gT}) (phi : 'CF(H / K)) : + K \subset H -> H \subset G -> K <| G -> + 'Ind[G] (phi %% K)%CF = ('Ind[G / K] phi %% K)%CF. +Proof. by move=> sKH sHG /andP[_ nKG]; rewrite cfIndMorph ?ker_coset. Qed. + +Lemma cfIndQuo (K : {group gT}) (phi : 'CF(H)) : + K \subset cfker phi -> H \subset G -> K <| G -> + 'Ind[G / K] (phi / K)%CF = ('Ind[G] phi / K)%CF. +Proof. +move=> kerK sHG nsKG; have sKH := subset_trans kerK (cfker_sub phi). +have nsKH := normalS sKH sHG nsKG. +by apply: canRL (cfModK nsKG) _; rewrite -cfIndMod // cfQuoK. +Qed. + +Section IndSumInertia. + +Variable s : Iirr H. + +Let theta := 'chi_s. +Let T := 'I_G[theta]. +Let calA := irr_constt ('Ind[T] theta). +Let calB := irr_constt ('Ind[G] theta). +Let AtoB (t : Iirr T) := Ind_Iirr G t. +Let e_ t := '['Ind theta, 'chi[T]_t]. + +Hypothesis nsHG: H <| G. +(* begin hide *) +Let sHG : H \subset G. Proof. exact: normal_sub. Qed. +Let nHG : G \subset 'N(H). Proof. exact: normal_norm. Qed. +Let nsHT : H <| T. Proof. exact: normal_Inertia. Qed. +Let sHT : H \subset T. Proof. exact: normal_sub. Qed. +Let nHT : T \subset 'N(H). Proof. exact: normal_norm. Qed. +Let sTG : T \subset G. Proof. exact: subsetIl. Qed. +(* end hide *) + +(* This is Peterfalvi (1.7)(a). *) +Lemma cfInd_sum_Inertia : + [/\ {in calA, forall t, 'Ind 'chi_t \in irr G}, + {in calA, forall t, 'chi_(AtoB t) = 'Ind 'chi_t}, + {in calA &, injective AtoB}, + AtoB @: calA =i calB + & 'Ind[G] theta = \sum_(t in calA) e_ t *: 'Ind 'chi_t]. +Proof. +have [AtoBirr AtoBinj defB _ _] := constt_Inertia_bijection s nsHG. +split=> // [i Ai|]; first exact/cfIirrE/AtoBirr. +rewrite -(cfIndInd _ sTG sHT) {1}['Ind theta]cfun_sum_constt linear_sum. +by apply: eq_bigr => i _; rewrite linearZ. +Qed. + +Hypothesis abTbar : abelian (T / H). + +(* This is Peterfalvi (1.7)(b). *) +Lemma cfInd_central_Inertia : + exists2 e, [/\ e \in Cnat, e != 0 & {in calA, forall t, e_ t = e}] + & [/\ 'Ind[G] theta = e *: \sum_(j in calB) 'chi_j, + #|calB|%:R = #|T : H|%:R / e ^+ 2 + & {in calB, forall i, 'chi_i 1%g = #|G : T|%:R * e * theta 1%g}]. +Proof. +have [t1 At1] := constt_cfInd_irr s sHT; pose psi1 := 'chi_t1. +pose e := '['Ind theta, psi1]. +have NthT: 'Ind[T] theta \is a character by rewrite cfInd_char ?irr_char. +have Ne: e \in Cnat by rewrite Cnat_cfdot_char_irr. +have Dpsi1H: 'Res[H] psi1 = e *: theta. + have psi1Hs: s \in irr_constt ('Res psi1) by rewrite -constt_Ind_Res. + rewrite (Clifford_Res_sum_cfclass nsHT psi1Hs) cfclass_invariant ?subsetIr //. + by rewrite big_seq1 cfdot_Res_l cfdotC conj_Cnat. +have linL j: 'chi[T / H]_j \is a linear_char by apply/char_abelianP. +have linLH j: ('chi_j %% H)%CF \is a linear_char := cfMod_lin_char (linL j). +pose LtoT (j : Iirr (T / H)) := mul_mod_Iirr t1 j. +have LtoTE j: 'chi_(LtoT j) = ('chi_j %% H)%CF * psi1. + by rewrite !(mod_IirrE, cfIirrE) // mul_lin_irr ?mem_irr ?cfMod_lin_char. +have psiHG: 'Ind ('Res[H] psi1) = \sum_j 'chi_(LtoT j). + transitivity ((cfReg (T / H) %% H)%CF * psi1); last first. + rewrite cfReg_sum linear_sum /= mulr_suml; apply: eq_bigr => i _. + by rewrite LtoTE // lin_char1 ?scale1r. + apply/cfun_inP=> x Tx; rewrite cfunE cfModE // cfRegE mulrnAl mulrb. + rewrite (sameP eqP (kerP _ (subsetP nHT x Tx))) ker_coset. + case: ifPn => [Hx | H'x]; last by rewrite (cfun_on0 (cfInd_normal _ _)). + rewrite card_quotient // -!(cfResE _ sHT) // cfRes_Ind_invariant ?cfunE //. + by rewrite -subsetIidl (subset_trans _ (sub_inertia_Res _ _)) ?sub_Inertia. +have imLtoT: {subset calA <= codom LtoT}. + move=> t At; apply/codomP/exists_eqP. + have{At}: t \in irr_constt ('Ind ('Res[H] 'chi_t1)). + by rewrite Dpsi1H linearZ irr_consttE cfdotZl mulf_neq0. + apply: contraR; rewrite negb_exists => /forallP imL't. + by rewrite psiHG cfdot_suml big1 // => j _; rewrite cfdot_irr mulrb ifN_eqC. +have De_ t: t \in calA -> e_ t = e. + case/imLtoT/codomP=> j ->; rewrite /e_ LtoTE /e -!cfdot_Res_r rmorphM /=. + by rewrite cfRes_sub_ker ?cfker_mod // mulr_algl lin_char1 ?scale1r. +have{imLtoT} A_1 t: t \in calA -> 'chi_t 1%g = e * theta 1%g. + case/imLtoT/codomP=> j ->; rewrite LtoTE //= cfunE. + by rewrite (lin_char1 (linLH j)) mul1r -(cfRes1 H) Dpsi1H cfunE. +exists e => //; have [_ defAtoB injAtoB imAtoB ->] := cfInd_sum_Inertia. +rewrite -(eq_bigl _ _ imAtoB) -(eq_card imAtoB) big_imset //= scaler_sumr. +split=> [||i]; first by apply: eq_bigr => t2 At2; rewrite De_ ?defAtoB. + apply: (mulIf (irr1_neq0 s)); rewrite mulrAC -cfInd1 // mulr_natl mulrC invfM. + rewrite ['Ind _]cfun_sum_constt sum_cfunE mulr_sumr card_in_imset //. + rewrite -sumr_const; apply: eq_bigr => t At. + by rewrite -mulrA -/(e_ t) De_ // cfunE A_1 ?mulKf. +by rewrite -imAtoB => /imsetP[t At ->]; rewrite defAtoB ?cfInd1 ?A_1 ?mulrA. +Qed. + +(* This is Peterfalvi (1.7)(c). *) +Lemma cfInd_Hall_central_Inertia : + Hall T H -> + [/\ 'Ind[G] theta = \sum_(i in calB) 'chi_i, #|calB| = #|T : H| + & {in calB, forall i, 'chi_i 1%g = #|G : T|%:R * theta 1%g}]. +Proof. +case/andP=> _ hallH; have [e [_ _ De]] := cfInd_central_Inertia. +suffices ->: e = 1. + by case=> -> /eqP; rewrite scale1r expr1n divr1 mulr1 eqC_nat => /eqP. +suffices{De} [t Dtheta]: exists i, 'Res[H, T] 'chi_i = theta. + have e_t_1: e_ t = 1 by rewrite /e_ -cfdot_Res_r Dtheta cfnorm_irr. + by rewrite -(De t) // irr_consttE -/(e_ t) e_t_1 oner_eq0. +have ITtheta: T \subset 'I[theta] := subsetIr _ _. +have solT: solvable (T / H) := abelian_sol abTbar. +have [|t []] := extend_solvable_coprime_irr nsHT solT ITtheta; last by exists t. +rewrite coprime_sym coprime_mull !(coprime_dvdl _ hallH) ?cfDet_order_dvdG //. +by rewrite -dvdC_nat !CdivE truncCK ?Cnat_irr1 // dvd_irr1_cardG. +Qed. + +End IndSumInertia. + +(* This is Peterfalvi (1.8). *) +Lemma irr1_bound_quo (B C D : {group gT}) i : + B <| C -> B \subset cfker 'chi[G]_i -> + B \subset D -> D \subset C -> C \subset G -> (D / B \subset 'Z(C / B))%g -> + 'chi_i 1%g <= #|G : C|%:R * sqrtC #|C : D|%:R. +Proof. +move=> BnC BsK BsD DsC CsG QsZ. +case: (boolP ('Res[C] 'chi_i == 0))=> [HH|]. + have: ('Res[C] 'chi_i) 1%g = 0 by rewrite (eqP HH) cfunE. + by rewrite cfResE // => HH1; case/eqP: (irr1_neq0 i). +have IC := cfRes_char C (irr_char i). +case/neq0_has_constt=> i1 Hi1. +have CIr: i \in irr_constt ('Ind[G] 'chi_i1). + by rewrite inE /= -Frobenius_reciprocity /= cfdotC conjC_eq0. +have BsKi : B \subset cfker 'chi_i1. + suff BsKri: B \subset cfker ('Res[C] 'chi_i). + by apply: (subset_trans BsKri); exact: (cfker_constt _ Hi1). + apply/subsetP=> g GiG. + have F: g \in C by rewrite (subsetP (subset_trans BsD _)). + rewrite cfkerEchar // inE F !cfResE //. + by move: (subsetP BsK _ GiG); rewrite cfkerEirr inE. +pose i2 := quo_Iirr B i1. +have ZsC: 'Z(C / B)%g \subset 'Z('chi_i2)%CF. + by rewrite -(cap_cfcenter_irr (C / B)); apply: bigcap_inf. +have CBsH: C :&: B \subset D. + apply/subsetP=> g; rewrite inE; case/andP=> _ HH. + by apply: (subsetP (BsD)). +have I1B: 'chi_i1 1%g ^+ 2 <= #|C : D|%:R. + case: (irr1_bound i2)=> HH _; move: HH. + have ->: 'chi_i2 1%g = 'chi_i1 1%g. + by rewrite quo_IirrE // -(coset_id (group1 B)) cfQuoE. + move/ler_trans; apply. + rewrite ler_nat // -(index_quotient_eq CBsH) ?normal_norm //. + rewrite -(@leq_pmul2l #|'Z('chi_i2)%CF|) ?cardG_gt0 ?cfcenter_sub //. + rewrite Lagrange ?quotientS ?cfcenter_sub //. + rewrite -(@leq_pmul2l #|(D / B)%g|) ?cardG_gt0 //. + rewrite mulnA mulnAC Lagrange ?quotientS //. + rewrite mulnC leq_pmul2l ?cardG_gt0 // subset_leq_card //. + exact: subset_trans QsZ ZsC. +have IC': 'Ind[G] 'chi_i1 \is a character := cfInd_char G (irr_char i1). +move: (char1_ge_constt IC' CIr); rewrite cfInd1 //= => /ler_trans-> //. +have chi1_1_ge0: 0 <= 'chi_i1 1%g by rewrite ltrW ?irr1_gt0. +rewrite ler_pmul2l ?gt0CiG //. +by rewrite -(@ler_pexpn2r _ 2) -?topredE /= ?sqrtC_ge0 ?ler0n ?sqrtCK. +Qed. + +(* This is Peterfalvi (1.9)(a). *) +Lemma extend_coprime_Qn_aut a b (Qa Qb : fieldExtType rat) w_a w_b + (QaC : {rmorphism Qa -> algC}) (QbC : {rmorphism Qb -> algC}) + (mu : {rmorphism algC -> algC}) : + coprime a b -> + a.-primitive_root w_a /\ <<1; w_a>>%VS = {:Qa}%VS -> + b.-primitive_root w_b /\ <<1; w_b>>%VS = {:Qb}%VS -> + {nu : {rmorphism algC -> algC} | forall x, nu (QaC x) = mu (QaC x) + & forall y, nu (QbC y) = QbC y}. +Proof. +move=> coab [pr_w_a genQa] [pr_w_b genQb]. +have [k co_k_a Dmu]: {k | coprime k a & mu (QaC w_a) = QaC (w_a ^+ k)}. + have prCw: a.-primitive_root (QaC w_a) by rewrite fmorph_primitive_root. + by have [k coka ->] := aut_prim_rootP mu prCw; rewrite -rmorphX; exists k. +pose k1 := chinese a b k 1; have /Qn_aut_exists[nu Dnu]: coprime k1 (a * b). + rewrite coprime_mulr -!(coprime_modl k1) chinese_modl ?chinese_modr //. + by rewrite !coprime_modl co_k_a coprime1n. +exists nu => [x | y]. + have /Fadjoin_polyP[p Qp ->]: x \in <<1; w_a>>%VS by rewrite genQa memvf. + rewrite -!horner_map -!map_poly_comp !map_Qnum_poly // Dmu Dnu -rmorphX /=. + by rewrite -(prim_expr_mod pr_w_a) chinese_modl // prim_expr_mod. + by rewrite exprM (prim_expr_order pr_w_a) expr1n rmorph1. +have /Fadjoin_polyP[p Qp ->]: y \in <<1; w_b>>%VS by rewrite genQb memvf. +rewrite -!horner_map -!map_poly_comp !map_Qnum_poly // Dnu -rmorphX /=. + by rewrite -(prim_expr_mod pr_w_b) chinese_modr // prim_expr_mod. +by rewrite mulnC exprM (prim_expr_order pr_w_b) expr1n rmorph1. +Qed. + +(* This intermediate result in the proof of Peterfalvi (1.9)(b) is used in *) +(* he proof of (3.9)(c). *) +Lemma dvd_restrict_cfAut a (v : {rmorphism algC -> algC}) : + exists2 u : {rmorphism algC -> algC}, + forall gT0 G0 chi x, + chi \in 'Z[irr (@gval gT0 G0)] -> #[x] %| a -> u (chi x) = v (chi x) + & forall chi x, chi \in 'Z[irr G] -> coprime #[x] a -> u (chi x) = chi x. +Proof. +have [-> | a_gt0] := posnP a. + exists v => // chi x Zchi; rewrite /coprime gcdn0 order_eq1 => /eqP->. + by rewrite aut_Cint ?Cint_vchar1. +pose b := (#|G|`_(\pi(a)^'))%N. +have co_a_b: coprime a b := pnat_coprime (pnat_pi a_gt0) (part_pnat _ _). +have [Qa _ [QaC _ [w_a genQa memQa]]] := group_num_field_exists [group of Zp a]. +have [Qb _ [QbC _ [w_b genQb memQb]]] := group_num_field_exists [group of Zp b]. +rewrite !card_Zp ?part_gt0 // in Qa QaC w_a genQa memQa Qb QbC w_b genQb memQb. +have [nu nuQa nuQb] := extend_coprime_Qn_aut QaC QbC v co_a_b genQa genQb. +exists nu => [gt0 G0 chi x Zchi x_dv_a | chi x Zchi co_x_a]. + without loss{Zchi} Nchi: chi / chi \is a character. + move=> IH; case/vcharP: Zchi => [chi1 Nchi1 [chi2 Nchi2 ->]]. + by rewrite !cfunE !rmorphB !IH. + by have [xa <-] := memQa _ _ _ Nchi x x_dv_a; rewrite nuQa. +without loss{Zchi} Nchi: chi / chi \is a character. + move=> IH; case/vcharP: Zchi => [chi1 Nchi1 [chi2 Nchi2 ->]]. + by rewrite !cfunE rmorphB !IH. +have [Gx | /cfun0->] := boolP (x \in G); last by rewrite rmorph0. +have{Gx} x_dv_b: (#[x] %| b)%N. + rewrite coprime_sym coprime_pi' // in co_x_a. + by rewrite -(part_pnat_id co_x_a) partn_dvd ?order_dvdG. +by have [xb <-] := memQb _ _ _ Nchi x x_dv_b; rewrite nuQb. +Qed. + +(* This is Peterfalvi (1.9)(b). *) +(* We have strengthened the statement of this lemma so that it can be used *) +(* rather than reproved for Peterfalvi (3.9). In particular we corrected a *) +(* quantifier inversion in the original statement: the automorphism is *) +(* constructed uniformly for all (virtual) characters. We have also removed *) +(* the spurrious condition that a be a \pi(a) part of #|G| -- the proof works *) +(* for all a, and indeed the first part holds uniformaly for all groups! *) +Lemma make_pi_cfAut a k : + coprime k a -> + exists2 u : {rmorphism algC -> algC}, + forall (gT0 : finGroupType) (G0 : {group gT0}) chi x, + chi \in 'Z[irr G0] -> #[x] %| a -> cfAut u chi x = chi (x ^+ k)%g + & forall chi x, chi \in 'Z[irr G] -> coprime #[x] a -> cfAut u chi x = chi x. +Proof. +move=> co_k_a; have [v Dv] := Qn_aut_exists co_k_a. +have [u Du_a Du_a'] := dvd_restrict_cfAut a v. +exists u => [gt0 G0 | ] chi x Zchi a_x; last by rewrite cfunE Du_a'. +rewrite cfunE {u Du_a'}Du_a //. +without loss{Zchi} Nchi: chi / chi \is a character. + move=> IH; case/vcharP: Zchi => [chi1 Nchi1 [chi2 Nchi2 ->]]. + by rewrite !cfunE rmorphB !IH. +have [sXG0 | G0'x] := boolP (<[x]> \subset G0); last first. + have /(<[x]> =P _) gen_xk: generator <[x]> (x ^+ k). + by rewrite generator_coprime coprime_sym (coprime_dvdr a_x). + by rewrite !cfun0 ?rmorph0 -?cycle_subG -?gen_xk. +rewrite -!(cfResE chi sXG0) ?cycle_id ?mem_cycle //. +rewrite ['Res _]cfun_sum_cfdot !sum_cfunE rmorph_sum; apply: eq_bigr => i _. +have chiX := lin_charX (char_abelianP _ (cycle_abelian x) i) _ (cycle_id x). +rewrite !cfunE rmorphM aut_Cnat ?Cnat_cfdot_char_irr ?cfRes_char //. +by congr (_ * _); rewrite Dv -chiX // -expg_mod_order (eqnP a_x) chiX. +Qed. + +Section ANT. +Import ssrint. + +(* This section covers Peterfalvi (1.10). *) +(* We have simplified the statement somewhat by substituting the global ring *) +(* of algebraic integers for the specific ring Z[eta]. Formally this amounts *) +(* to strengthening (b) and weakening (a) accordingly, but since actually the *) +(* Z[eta] is equal to the ring of integers of Q[eta] (cf. Theorem 6.4 in J.S. *) +(* Milne's course notes on Algebraic Number Theory), the simplified statement *) +(* is actually equivalent to the textbook one. *) +Variable (p : nat) (eps : algC). +Hypothesis (pr_eps : p.-primitive_root eps). +Local Notation e := (1 - eps). + +(* This is Peterfalvi (1.10) (a). *) +Lemma vchar_ker_mod_prim : {in G & G & 'Z[irr G], forall x y (chi : 'CF(G)), + #[x] = p -> y \in 'C[x] -> chi (x * y)%g == chi y %[mod e]}%A. +Proof. +move=> x y chi Gx Gy Zchi ox cxy; pose X := <<[set x; y]>>%G. +have [Xx Xy]: x \in X /\ y \in X by apply/andP; rewrite -!sub1set -join_subG. +have sXG: X \subset G by rewrite join_subG !sub1set Gx. +suffices{chi Zchi} IHiX i: ('chi[X]_i (x * y)%g == 'chi_i y %[mod e])%A. + rewrite -!(cfResE _ sXG) ?groupM //. + have irr_free := (free_uniq (basis_free (irr_basis X))). + have [c Zc ->] := (zchar_expansion irr_free (cfRes_vchar X Zchi)). + rewrite !sum_cfunE /eqAmod -sumrB big_seq rpred_sum // => _ /irrP[i ->]. + by rewrite !cfunE [(_ %| _)%A]eqAmodMl // rpred_Cint. +have lin_chi: 'chi_i \is a linear_char. + apply/char_abelianP; rewrite -[gval X]joing_idl -joing_idr abelianY. + by rewrite !cycle_abelian cycle_subG /= cent_cycle. +rewrite lin_charM // -{2}['chi_i y]mul1r eqAmodMr ?Aint_irr //. +have [|k ->] := (prim_rootP pr_eps) ('chi_i x). + by rewrite -lin_charX // -ox expg_order lin_char1. +rewrite -[_ ^+ k](subrK 1) subrX1 -[_ - 1]opprB mulNr -mulrN mulrC. +rewrite eqAmod_addl_mul // rpredN rpred_sum // => n _. +by rewrite rpredX ?(Aint_prim_root pr_eps). +Qed. + +(* This is Peterfalvi (1.10)(b); the primality condition is only needed here. *) +Lemma int_eqAmod_prime_prim n : + prime p -> n \in Cint -> (n == 0 %[mod e])%A -> (p %| n)%C. +Proof. +move=> p_pr Zn; rewrite /eqAmod unfold_in subr0. +have p_gt0 := prime_gt0 p_pr. +case: ifPn => [_ /eqP->// | nz_e e_dv_n]. +suffices: (n ^+ p.-1 == 0 %[mod p])%A. + rewrite eqAmod0_rat ?rpredX ?rpred_nat 1?rpred_Cint // !dvdC_int ?rpredX //. + by rewrite floorCX // abszX Euclid_dvdX // => /andP[]. +rewrite /eqAmod subr0 unfold_in pnatr_eq0 eqn0Ngt p_gt0 /=. +pose F := \prod_(1 <= i < p) ('X - (eps ^+ i)%:P). +have defF: F = \sum_(i < p) 'X^i. + apply: (mulfI (monic_neq0 (monicXsubC 1))); rewrite -subrX1. + by rewrite -(factor_Xn_sub_1 pr_eps) big_ltn. +have{defF} <-: F.[1] = p :> Algebraics.divisor. + rewrite -[p]card_ord -[rhs in _ = rhs]sumr_const defF horner_sum. + by apply: eq_bigr => i _; rewrite hornerXn expr1n. +rewrite -[p.-1]card_ord {F}horner_prod big_add1 big_mkord -prodfV. +rewrite -prodr_const -big_split rpred_prod //= => k _; rewrite !hornerE. +rewrite -[n](divfK nz_e) -[_ * _ / _]mulrA rpredM {e_dv_n}//. +have p'k: ~~ (p %| k.+1)%N by rewrite gtnNdvd // -{2}(prednK p_gt0) ltnS. +have [r {1}->]: exists r, eps = eps ^+ k.+1 ^+ r. + have [q _ /dvdnP[r Dr]] := Bezoutl p (ltn0Sn k); exists r; apply/esym/eqP. + rewrite -exprM (eq_prim_root_expr pr_eps _ 1) mulnC -Dr addnC gcdnC. + by rewrite -prime_coprime // in p'k; rewrite (eqnP p'k) modnMDl. +rewrite -[1 - _]opprB subrX1 -mulNr opprB mulrC. +rewrite mulKf; last by rewrite subr_eq0 eq_sym -(prim_order_dvd pr_eps). +by apply: rpred_sum => // i _; rewrite !rpredX ?(Aint_prim_root pr_eps). +Qed. + +End ANT. + +End Main. + + diff --git a/mathcomp/odd_order/PFsection10.v b/mathcomp/odd_order/PFsection10.v new file mode 100644 index 0000000..d380e47 --- /dev/null +++ b/mathcomp/odd_order/PFsection10.v @@ -0,0 +1,1215 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg poly finset center. +Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +Require Import gfunctor gproduct cyclic commutator gseries nilpotent pgroup. +Require Import sylow hall abelian maximal frobenius. +Require Import matrix mxalgebra mxrepresentation mxabelem vector. +Require Import BGsection1 BGsection3 BGsection7 BGsection15 BGsection16. +Require Import ssrnum algC classfun character integral_char inertia vcharacter. +Require Import PFsection1 PFsection2 PFsection3 PFsection4. +Require Import PFsection5 PFsection6 PFsection7 PFsection8 PFsection9. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 10: Maximal subgroups of Types III, *) +(* IV and V. For defW : W1 \x W2 = W and MtypeP : of_typeP M U defW, and *) +(* setting ptiW := FT_primeTI_hyp MtypeP, mu2_ i j := primeTIirr ptiW i j and *) +(* delta_ j := primeTIsign j, we define here, for M of type III-V: *) +(* FTtype345_TIirr_degree MtypeP == the common degree of the components of *) +(* (locally) d the images of characters of irr W that don't have *) +(* W2 in their kernel by the cyclicTI isometry to M. *) +(* Thus mu2_ i j 1%g = d%:R for all j != 0. *) +(* FTtype345_TIsign MtypeP == the common sign of the images of characters *) +(* (locally) delta of characters of irr W that don't have W2 in *) +(* their kernel by the cyclicTI isometry to M. *) +(* Thus delta_ j = delta for all j != 0. *) +(* FTtype345_ratio MtypeP == the ratio (d - delta) / #|W1|. Even though it *) +(* (locally) n is always a positive integer we take n : algC. *) +(* FTtype345_bridge MtypeP s i j == a virtual character that can be used to *) +(* (locally) alpha_ i j bridge coherence between the mu2_ i j and other *) +(* irreducibles of M; here s should be the index of *) +(* an irreducible character of M induced from M^(1). *) +(* := mu2_ i j - delta *: mu2_ i 0 -n *: 'chi_s. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. + +Section Ten. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types (p q : nat) (x y z : gT). +Implicit Types H K L N P Q R S T U W : {group gT}. + +Local Notation "#1" := (inord 1) (at level 0). + +Section OneMaximal. + +(* These assumptions correspond to Peterfalvi, Hypothesis (10.1). *) +(* We also declare the group U_M, even though it is not used in this section, *) +(* because it is a parameter to the theorems and definitions of PFsection8 *) +(* and PFsection9. *) +Variables M U_M W W1 W2 : {group gT}. +Hypotheses (maxM : M \in 'M) (defW : W1 \x W2 = W). +Hypotheses (MtypeP : of_typeP M U_M defW) (notMtype2: FTtype M != 2). + +Local Notation "` 'M'" := (gval M) (at level 0, only parsing) : group_scope. +Local Notation "` 'W1'" := (gval W1) (at level 0, only parsing) : group_scope. +Local Notation "` 'W2'" := (gval W2) (at level 0, only parsing) : group_scope. +Local Notation "` 'W'" := (gval W) (at level 0, only parsing) : group_scope. +Local Notation V := (cyclicTIset defW). +Local Notation M' := M^`(1)%G. +Local Notation "` 'M''" := `M^`(1) (at level 0) : group_scope. +Local Notation M'' := M^`(2)%G. +Local Notation "` 'M'''" := `M^`(2) (at level 0) : group_scope. + +Let defM : M' ><| W1 = M. Proof. by have [[]] := MtypeP. Qed. +Let nsM''M' : M'' <| M'. Proof. exact: (der_normal 1 M'). Qed. +Let nsM'M : M' <| M. Proof. exact: (der_normal 1 M). Qed. +Let sM'M : M' \subset M. Proof. exact: der_sub. Qed. +Let nsM''M : M'' <| M. Proof. exact: der_normal 2 M. Qed. + +Let notMtype1 : FTtype M != 1%N. Proof. exact: FTtypeP_neq1 MtypeP. Qed. +Let typeMgt2 : FTtype M > 2. +Proof. by move: (FTtype M) (FTtype_range M) notMtype1 notMtype2=> [|[|[]]]. Qed. + +Let defA1 : 'A1(M) = M'^#. Proof. by rewrite /= -FTcore_eq_der1. Qed. +Let defA : 'A(M) = M'^#. Proof. by rewrite FTsupp_eq1 ?defA1. Qed. +Let defA0 : 'A0(M) = M'^# :|: class_support V M. +Proof. by rewrite -defA (FTtypeP_supp0_def _ MtypeP). Qed. +Let defMs : M`_\s :=: M'. Proof. exact: FTcore_type_gt2. Qed. + +Let pddM := FT_prDade_hyp maxM MtypeP. +Let ptiWM : primeTI_hypothesis M M' defW := FT_primeTI_hyp MtypeP. +Let ctiWG : cyclicTI_hypothesis G defW := pddM. +Let ctiWM : cyclicTI_hypothesis M defW := prime_cycTIhyp ptiWM. + +Let ntW1 : W1 :!=: 1. Proof. by have [[]] := MtypeP. Qed. +Let ntW2 : W2 :!=: 1. Proof. by have [_ _ _ []] := MtypeP. Qed. +Let cycW1 : cyclic W1. Proof. by have [[]] := MtypeP. Qed. +Let cycW2 : cyclic W2. Proof. by have [_ _ _ []] := MtypeP. Qed. + +Let w1 := #|W1|. +Let w2 := #|W2|. +Let nirrW1 : #|Iirr W1| = w1. Proof. by rewrite card_Iirr_cyclic. Qed. +Let nirrW2 : #|Iirr W2| = w2. Proof. by rewrite card_Iirr_cyclic. Qed. +Let NirrW1 : Nirr W1 = w1. Proof. by rewrite -nirrW1 card_ord. Qed. +Let NirrW2 : Nirr W2 = w2. Proof. by rewrite -nirrW2 card_ord. Qed. + +Let w1gt2 : w1 > 2. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. +Let w2gt2 : w2 > 2. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. + +Let coM'w1 : coprime #|M'| w1. +Proof. by rewrite (coprime_sdprod_Hall_r defM); have [[]] := MtypeP. Qed. + +(* This is used both in (10.2) and (10.8). *) +Let frobMbar : [Frobenius M / M'' = (M' / M'') ><| (W1 / M'')]. +Proof. +have [[_ hallW1 _ _] _ _ [_ _ _ sW2M'' regM'W1 ] _] := MtypeP. +apply: Frobenius_coprime_quotient => //. +split=> [|w /regM'W1-> //]; apply: (sol_der1_proper (mmax_sol maxM)) => //. +by apply: subG1_contra ntW2; apply: subset_trans sW2M'' (der_sub 1 M'). +Qed. + +Local Open Scope ring_scope. + +Let sigma := (cyclicTIiso ctiWG). +Let w_ i j := (cyclicTIirr defW i j). +Local Notation eta_ i j := (sigma (w_ i j)). + +Local Notation Imu2 := (primeTI_Iirr ptiWM). +Let mu2_ i j := primeTIirr ptiWM i j. +Let mu_ := primeTIred ptiWM. +Local Notation chi_ j := (primeTIres ptiWM j). + +Local Notation Idelta := (primeTI_Isign ptiWM). +Local Notation delta_ j := (primeTIsign ptiWM j). + +Local Notation tau := (FT_Dade0 maxM). +Local Notation "chi ^\tau" := (tau chi). + +Let calS0 := seqIndD M' M M`_\s 1. +Let rmR := FTtypeP_coh_base maxM MtypeP. +Let scohS0 : subcoherent calS0 tau rmR. +Proof. exact: FTtypeP_subcoherent MtypeP. Qed. + +Let calS := seqIndD M' M M' 1. +Let sSS0 : cfConjC_subset calS calS0. +Proof. by apply: seqInd_conjC_subset1; rewrite /= ?defMs. Qed. + +Let mem_calS s : ('Ind 'chi[M']_s \in calS) = (s != 0). +Proof. +rewrite mem_seqInd ?normal1 ?FTcore_normal //=. +by rewrite !inE sub1G subGcfker andbT. +Qed. + +Let calSmu j : j != 0 -> mu_ j \in calS. +Proof. +move=> nz_j; rewrite -[mu_ j]cfInd_prTIres mem_calS -irr_eq1. +by rewrite -(prTIres0 ptiWM) (inj_eq irr_inj) (inj_eq (prTIres_inj _)). +Qed. + +Let tauM' : {subset 'Z[calS, M'^#] <= 'CF(M, 'A0(M))}. +Proof. by rewrite defA0 => phi /zchar_on/(cfun_onS (subsetUl _ _))->. Qed. + +(* This is Peterfalvi (10.2). *) +(* Note that this result is also valid for type II groups. *) +Lemma FTtypeP_ref_irr : + {zeta | [/\ zeta \in irr M, zeta \in calS & zeta 1%g = w1%:R]}. +Proof. +have [_ /has_nonprincipal_irr[s nz_s] _ _ _] := Frobenius_context frobMbar. +exists ('Ind 'chi_s %% M'')%CF; split. +- exact/cfMod_irr/(irr_induced_Frobenius_ker (FrobeniusWker frobMbar)). +- by rewrite -cfIndMod ?normal_sub // -mod_IirrE // mem_calS mod_Iirr_eq0. +rewrite -cfIndMod ?cfInd1 ?normal_sub // -(index_sdprod defM) cfMod1. +by rewrite lin_char1 ?mulr1 //; apply/char_abelianP/sub_der1_abelian. +Qed. + +(* This is Peterfalvi (10.3), first assertion. *) +Lemma FTtype345_core_prime : prime w2. +Proof. +have [S pairMS [xdefW [U StypeP]]] := FTtypeP_pair_witness maxM MtypeP. +have [[_ _ maxS] _] := pairMS; rewrite {1}(negPf notMtype2) /= => Stype2 _ _. +by have [[]] := compl_of_typeII maxS StypeP Stype2. +Qed. +Let w2_pr := FTtype345_core_prime. + +Definition FTtype345_TIirr_degree := truncC (mu2_ 0 #1 1%g). +Definition FTtype345_TIsign := delta_ #1. +Local Notation d := FTtype345_TIirr_degree. +Local Notation delta := FTtype345_TIsign. +Definition FTtype345_ratio := (d%:R - delta) / w1%:R. +Local Notation n := FTtype345_ratio. + +(* This is the remainder of Peterfalvi (10.3). *) +Lemma FTtype345_constants : + [/\ forall i j, j != 0 -> mu2_ i j 1%g = d%:R, + forall j, j != 0 -> delta_ j = delta, + (d > 1)%N + & n \in Cnat]. +Proof. +have nz_j1 : #1 != 0 :> Iirr W2 by rewrite Iirr1_neq0. +have invj j: j != 0 -> mu2_ 0 j 1%g = d%:R /\ delta_ j = delta. + move=> nz_j; have [k co_k_j1 Dj] := cfExp_prime_transitive w2_pr nz_j1 nz_j. + rewrite -(cforder_dprodr defW) -dprod_IirrEr in co_k_j1. + have{co_k_j1} [[u Dj1u] _] := cycTIiso_aut_exists ctiWM co_k_j1. + rewrite dprod_IirrEr -rmorphX -Dj /= -!dprod_IirrEr -!/(w_ _ _) in Dj1u. + rewrite truncCK ?Cnat_irr1 //. + have: delta_ j *: mu2_ 0 j == cfAut u (delta_ #1 *: mu2_ 0 #1). + by rewrite -!(cycTIiso_prTIirr pddM) -/ctiWM -Dj1u. + rewrite raddfZsign /= -prTIirr_aut eq_scaled_irr signr_eq0 /= /mu2_. + by case/andP=> /eqP-> /eqP->; rewrite prTIirr_aut cfunE aut_Cnat ?Cnat_irr1. +have d_gt1: (d > 1)%N. + rewrite ltn_neqAle andbC -eqC_nat -ltC_nat truncCK ?Cnat_irr1 //. + rewrite irr1_gt0 /= eq_sym; apply: contraNneq nz_j1 => mu2_lin. + have: mu2_ 0 #1 \is a linear_char by rewrite qualifE irr_char /= mu2_lin. + by rewrite lin_irr_der1 => /(prTIirr0P ptiWM)[i /irr_inj/prTIirr_inj[_ ->]]. +split=> // [i j /invj[<- _] | _ /invj[//] | ]; first by rewrite prTIirr_1. +have: (d%:R == delta %[mod w1])%C by rewrite truncCK ?Cnat_irr1 ?prTIirr1_mod. +rewrite /eqCmod unfold_in -/n (negPf (neq0CG W1)) CnatEint => ->. +rewrite divr_ge0 ?ler0n // [delta]signrE opprB addrA -natrD subr_ge0 ler1n. +by rewrite -(subnKC d_gt1). +Qed. + +Let o_mu2_irr zeta i j : + zeta \in calS -> zeta \in irr M -> '[mu2_ i j, zeta] = 0. +Proof. +case/seqIndP=> s _ -> irr_sM; rewrite -cfdot_Res_l cfRes_prTIirr cfdot_irr. +rewrite (negPf (contraNneq _ (prTIred_not_irr ptiWM j))) // => Ds. +by rewrite -cfInd_prTIres Ds. +Qed. + +Let ZmuBzeta zeta j : + zeta \in calS -> zeta 1%g = w1%:R -> j != 0 -> + mu_ j - d%:R *: zeta \in 'Z[calS, M'^#]. +Proof. +move=> Szeta zeta1w1 nz_j; have [mu1 _ _ _] := FTtype345_constants. +rewrite -[d%:R](mulKf (neq0CiG M M')) mulrC -(mu1 0 j nz_j). +rewrite -(cfResE _ sM'M) // cfRes_prTIirr -cfInd1 // cfInd_prTIres. +by rewrite (seqInd_sub_lin_vchar _ Szeta) ?calSmu // -(index_sdprod defM). +Qed. + +Let mu0Bzeta_on zeta : + zeta \in calS -> zeta 1%g = w1%:R -> mu_ 0 - zeta \in 'CF(M, 'A(M)). +Proof. +move/seqInd_on=> M'zeta zeta1w1; rewrite [mu_ 0]prTIred0 defA cfun_onD1. +rewrite !cfunE zeta1w1 cfuniE // group1 mulr1 subrr rpredB ?M'zeta //=. +by rewrite rpredZ ?cfuni_on. +Qed. + +(* We need to prove (10.5) - (10.7) for an arbitrary choice of zeta, to allow *) +(* part of the proof of (10.5) to be reused in that of (11.8). *) +Variable zeta : 'CF(M). +Hypotheses (irr_zeta : zeta \in irr M) (Szeta : zeta \in calS). +Hypothesis (zeta1w1 : zeta 1%g = w1%:R). + +Let o_mu2_zeta i j : '[mu2_ i j, zeta] = 0. Proof. exact: o_mu2_irr. Qed. + +Let o_mu_zeta j : '[mu_ j, zeta] = 0. +Proof. by rewrite cfdot_suml big1 // => i _; apply: o_mu2_zeta. Qed. + +Definition FTtype345_bridge i j := mu2_ i j - delta *: mu2_ i 0 - n *: zeta. +Local Notation alpha_ := FTtype345_bridge. + +(* This is the first part of Peterfalvi (10.5), which does not depend on the *) +(* coherence assumption that will ultimately be refuted by (10.8). *) +Lemma supp_FTtype345_bridge i j : j != 0 -> alpha_ i j \in 'CF(M, 'A0(M)). +Proof. +move=> nz_j; have [Dd Ddelta _ _] := FTtype345_constants. +have Dmu2 := prTIirr_id pddM. +have W1a0 x: x \in W1 -> alpha_ i j x = 0. + move=> W1x; rewrite !cfunE; have [-> | ntx] := eqVneq x 1%g. + by rewrite Dd // prTIirr0_1 mulr1 zeta1w1 divfK ?neq0CG ?subrr. + have notM'x: x \notin M'. + apply: contra ntx => M'x; have: x \in M' :&: W1 by apply/setIP. + by rewrite coprime_TIg ?inE. + have /sdprod_context[_ sW1W _ _ tiW21] := dprodWsdC defW. + have abW2: abelian W2 := cyclic_abelian cycW2. + have Wx: x \in W :\: W2. + rewrite inE (contra _ ntx) ?(subsetP sW1W) // => W2x. + by rewrite -in_set1 -set1gE -tiW21 inE W2x. + rewrite !Dmu2 {Wx}// Ddelta // prTIsign0 scale1r !dprod_IirrE cfunE. + rewrite -!(cfResE _ sW1W) ?cfDprodKl_abelian // subrr. + have [s _ ->] := seqIndP Szeta. + by rewrite (cfun_on0 (cfInd_normal _ _)) ?mulr0 ?subrr. +apply/cfun_onP=> x; rewrite !inE defA notMtype1 /= => /norP[notM'x]. +set pi := \pi(M'); have [Mx /= pi_x | /cfun0->//] := boolP (x \in M). +have hallM': pi.-Hall(M) M' by rewrite Hall_pi -?(coprime_sdprod_Hall_l defM). +have hallW1: pi^'.-Hall(M) W1 by rewrite -(compl_pHall _ hallM') sdprod_compl. +have{pi_x} pi'x: pi^'.-elt x. + apply: contraR notM'x => not_pi'x; rewrite !inE (mem_normal_Hall hallM') //. + rewrite not_pi'x andbT negbK in pi_x. + by rewrite (contraNneq _ not_pi'x) // => ->; apply: p_elt1. +have [|y My] := Hall_subJ (mmax_sol maxM) hallW1 _ pi'x; rewrite cycle_subG //. +by case/imsetP=> z Wz ->; rewrite cfunJ ?W1a0. +Qed. +Local Notation alpha_on := supp_FTtype345_bridge. + +Lemma vchar_FTtype345_bridge i j : alpha_ i j \in 'Z[irr M]. +Proof. +have [_ _ _ Nn] := FTtype345_constants. +by rewrite !rpredB ?rpredZsign ?rpredZ_Cnat ?irr_vchar ?mem_zchar. +Qed. +Local Notation Zalpha := vchar_FTtype345_bridge. +Local Hint Resolve Zalpha. + +Lemma vchar_Dade_FTtype345_bridge i j : + j != 0 -> (alpha_ i j)^\tau \in 'Z[irr G]. +Proof. by move=> nz_j; rewrite Dade_vchar // zchar_split Zalpha alpha_on. Qed. +Local Notation Zalpha_tau := vchar_Dade_FTtype345_bridge. + +(* This covers the last paragraph in the proof of (10.5); it's isolated here *) +(* because it is reused in the proof of (10.10) and (11.8). *) + +Lemma norm_FTtype345_bridge i j : + j != 0 -> '[(alpha_ i j)^\tau] = 2%:R + n ^+ 2. +Proof. +move=> nz_j; rewrite Dade_isometry ?alpha_on // cfnormBd ?cfnormZ; last first. + by rewrite cfdotZr cfdotBl cfdotZl !o_mu2_zeta !(mulr0, subr0). +have [_ _ _ /Cnat_ge0 n_ge0] := FTtype345_constants. +rewrite ger0_norm // cfnormBd ?cfnorm_sign ?cfnorm_irr ?irrWnorm ?mulr1 //. +by rewrite cfdotZr (cfdot_prTIirr pddM) (negPf nz_j) andbF ?mulr0. +Qed. +Local Notation norm_alpha := norm_FTtype345_bridge. + +Implicit Type tau : {additive 'CF(M) -> 'CF(G)}. + +(* This exported version is adapted to its use in (11.8). *) +Lemma FTtype345_bridge_coherence calS1 tau1 i j X Y : + coherent_with calS1 M^# tau tau1 -> (alpha_ i j)^\tau = X + Y -> + cfConjC_subset calS1 calS0 -> {subset calS1 <= irr M} -> + j != 0 -> Y \in 'Z[map tau1 calS1] -> '[Y, X] = 0 -> '[Y] = n ^+ 2 -> + X = delta *: (eta_ i j - eta_ i 0). +Proof. +move=> cohS1 Dalpha sS10 irrS1 nz_j S1_Y oYX nY_n2. +have [[_ Ddelta _ Nn] [[Itau1 Ztau1] _]] := (FTtype345_constants, cohS1). +have [|z Zz defY] := zchar_expansion _ S1_Y. + rewrite map_inj_in_uniq; first by case: sS10. + by apply: sub_in2 (Zisometry_inj Itau1); apply: mem_zchar. +have nX_2: '[X] = 2%:R. + apply: (addrI '[Y]); rewrite -cfnormDd // addrC -Dalpha norm_alpha //. + by rewrite addrC nY_n2. +have Z_X: X \in 'Z[irr G]. + rewrite -[X](addrK Y) -Dalpha rpredB ?Zalpha_tau // defY big_map big_seq. + by apply: rpred_sum => psi S1psi; rewrite rpredZ_Cint // Ztau1 ?mem_zchar. +apply: eq_signed_sub_cTIiso => // y Vy; rewrite -[X](addrK Y) -Dalpha -/delta. +rewrite !cfunE !cycTIiso_restrict //; set rhs := delta * _. +rewrite Dade_id ?defA0 //; last by rewrite setUC inE mem_class_support. +have notM'y: y \notin M'. + by have:= subsetP (prDade_supp_disjoint pddM) y Vy; rewrite inE. +have Wy: y \in W :\: W2 by move: Vy; rewrite !inE => /andP[/norP[_ ->]]. +rewrite !cfunE 2?{1}prTIirr_id // prTIsign0 scale1r Ddelta // cfunE -mulrBr. +rewrite -/rhs (cfun_on0 (seqInd_on _ Szeta)) // mulr0 subr0. +rewrite (ortho_cycTIiso_vanish ctiWG) ?subr0 // -/sigma. +apply/orthoPl=> _ /mapP[_ /(cycTIirrP defW)[i1 [j1 ->]] ->]. +rewrite defY cfdot_suml big_map big1_seq //= => psi S1psi. +by rewrite cfdotZl (coherent_ortho_cycTIiso MtypeP sS10) ?irrS1 ?mulr0. +Qed. + +(* This is a specialization of the above, used in (10.5) and (10.10). *) +Let def_tau_alpha calS1 tau1 i j : + coherent_with calS1 M^# tau tau1 -> cfConjC_subset calS1 calS0 -> + j != 0 -> zeta \in calS1 -> '[(alpha_ i j)^\tau, tau1 zeta] = - n -> + (alpha_ i j)^\tau = delta *: (eta_ i j - eta_ i 0) - n *: tau1 zeta. +Proof. +move=> cohS1 [_ sS10 ccS1] nz_j S1zeta alpha_zeta_n. +have [[_ _ _ Nn] [[Itau1 _] _]] := (FTtype345_constants, cohS1). +set Y := - (n *: _); apply: canRL (addrK _) _; set X := _ + _. +have Dalpha: (alpha_ i j)^\tau = X + Y by rewrite addrK. +have nY_n2: '[Y] = n ^+ 2. + by rewrite cfnormN cfnormZ norm_Cnat // Itau1 ?mem_zchar // irrWnorm ?mulr1. +pose S2 := zeta :: zeta^*%CF; pose S2tau1 := map tau1 S2. +have S2_Y: Y \in 'Z[S2tau1] by rewrite rpredN rpredZ_Cnat ?mem_zchar ?mem_head. +have sS21: {subset S2 <= calS1} by apply/allP; rewrite /= ccS1 ?S1zeta. +have cohS2 : coherent_with S2 M^# tau tau1 := subset_coherent_with sS21 cohS1. +have irrS2: {subset S2 <= irr M} by apply/allP; rewrite /= cfAut_irr irr_zeta. +rewrite (FTtype345_bridge_coherence cohS2 Dalpha) //; last first. + rewrite -[X]opprK cfdotNr opprD cfdotDr nY_n2 cfdotNl cfdotNr opprK cfdotZl. + by rewrite cfdotC alpha_zeta_n rmorphN conj_Cnat // mulrN addNr oppr0. +split=> [|_ /sS21/sS10//|]; last first. + by apply/allP; rewrite /= !inE cfConjCK !eqxx orbT. +by rewrite /= inE eq_sym; have [[_ /hasPn-> //]] := scohS0; apply: sS10. +Qed. + +Section NonCoherence. + +(* We will ultimately contradict these assumptions. *) +(* Note that we do not need to export any lemma save the final contradiction. *) +Variable tau1 : {additive 'CF(M) -> 'CF(G)}. +Hypothesis cohS : coherent_with calS M^# tau tau1. + +Local Notation "mu ^\tau1" := (tau1 mu%CF) + (at level 2, format "mu ^\tau1") : ring_scope. + +Let Dtau1 : {in 'Z[calS, M'^#], tau1 =1 tau}. +Proof. by case: cohS => _; apply: sub_in1; apply: zchar_onS; apply: setSD. Qed. + +Let o_zeta_s: '[zeta, zeta^*] = 0. +Proof. by rewrite (seqInd_conjC_ortho _ _ _ Szeta) ?mFT_odd /= ?defMs. Qed. + +Import ssrint rat. + +(* This is the second part of Peterfalvi (10.5). *) +Let tau_alpha i j : j != 0 -> + (alpha_ i j)^\tau = delta *: (eta_ i j - eta_ i 0) - n *: zeta^\tau1. +Proof. +move=> nz_j; set al_ij := alpha_ i j; have [[Itau1 Ztau1] _] := cohS. +have [mu1 Ddelta d_gt1 Nn] := FTtype345_constants. +pose a := '[al_ij^\tau, zeta^\tau1] + n. +have al_ij_zeta_s: '[al_ij^\tau, zeta^*^\tau1] = a. + apply: canRL (addNKr _) _; rewrite addrC -opprB -cfdotBr -raddfB. + have M'dz: zeta - zeta^*%CF \in 'Z[calS, M'^#] by apply: seqInd_sub_aut_zchar. + rewrite Dtau1 // Dade_isometry ?alpha_on ?tauM' //. + rewrite cfdotBr opprB cfdotBl cfdot_conjCr rmorphB linearZ /=. + rewrite -!prTIirr_aut !cfdotBl !cfdotZl !o_mu2_zeta o_zeta_s !mulr0. + by rewrite opprB !(subr0, rmorph0) add0r irrWnorm ?mulr1. +have Zal_ij: al_ij^\tau \in 'Z[irr G] by apply: Zalpha_tau. +have Za: a \in Cint. + by rewrite rpredD ?(Cint_Cnat Nn) ?Cint_cfdot_vchar ?Ztau1 ?(mem_zchar Szeta). +have{al_ij_zeta_s} ub_da2: (d ^ 2)%:R * a ^+ 2 <= (2%:R + n ^+ 2) * w1%:R. + have [k nz_k j'k]: exists2 k, k != 0 & k != j. + have:= w2gt2; rewrite -nirrW2 (cardD1 0) (cardD1 j) !inE nz_j !ltnS lt0n. + by case/pred0Pn=> k /and3P[]; exists k. + have muk_1: mu_ k 1%g = (d * w1)%:R. + by rewrite (prTIred_1 pddM) mu1 // mulrC -natrM. + rewrite natrX -exprMn; have <-: '[al_ij^\tau, (mu_ k)^\tau1] = d%:R * a. + rewrite mulrDr mulr_natl -raddfMn /=; apply: canRL (addNKr _) _. + rewrite addrC -cfdotBr -raddfMn -raddfB -scaler_nat. + rewrite Dtau1 ?Dade_isometry ?alpha_on ?tauM' ?ZmuBzeta // cfdotBr cfdotZr. + rewrite rmorph_nat !cfdotBl !cfdotZl !o_mu2_zeta irrWnorm //. + rewrite !(cfdot_prTIirr_red pddM) cfdotC o_mu_zeta conjC0 !mulr0 mulr1. + by rewrite 2 1?eq_sym // mulr0 -mulrN opprB !subr0 add0r. + have ZSmuk: mu_ k \in 'Z[calS] by rewrite mem_zchar ?calSmu. + have <-: '[al_ij^\tau] * '[(mu_ k)^\tau1] = (2%:R + n ^+ 2) * w1%:R. + by rewrite Itau1 // cfdot_prTIred eqxx mul1n norm_alpha. + by rewrite -Cint_normK ?cfCauchySchwarz // Cint_cfdot_vchar // Ztau1. +suffices a0 : a = 0. + by apply: (def_tau_alpha _ sSS0); rewrite // -sub0r -a0 addrK. +apply: contraTeq (d_gt1) => /(sqr_Cint_ge1 Za) a2_ge1. +suffices: n == 0. + rewrite mulf_eq0 invr_eq0 orbC -implyNb neq0CG /= subr_eq0 => /eqP Dd. + by rewrite -ltC_nat -(normr_nat _ d) Dd normr_sign ltrr. +suffices: n ^+ 2 < n + 1. + have d_dv_M: (d%:R %| #|M|)%C by rewrite -(mu1 0 j) // ?dvd_irr1_cardG. + have{d_dv_M} d_odd: odd d by apply: dvdn_odd (mFT_odd M); rewrite -dvdC_nat. + have: (2 %| n * w1%:R)%C. + rewrite divfK ?neq0CG // -signrN signrE addrA -(natrD _ d 1). + by rewrite rpredB // dvdC_nat dvdn2 ?odd_double // odd_add d_odd. + rewrite -(truncCK Nn) -mulrSr -natrM -natrX ltC_nat (dvdC_nat 2) pnatr_eq0. + rewrite dvdn2 odd_mul mFT_odd; case: (truncC n) => [|[|n1]] // _ /idPn[]. + by rewrite -leqNgt (ltn_exp2l 1). +apply: ltr_le_trans (_ : n * - delta + 1 <= _); last first. + have ->: n + 1 = n * `|- delta| + 1 by rewrite normrN normr_sign mulr1. + rewrite ler_add2r ler_wpmul2l ?Cnat_ge0 ?real_ler_norm //. + by rewrite rpredN ?rpred_sign. +rewrite -(ltr_pmul2r (ltC_nat 0 2)) mulrDl mul1r -[rhs in rhs + _]mulrA. +apply: ler_lt_trans (_ : n ^+ 2 * (w1%:R - 1) < _). + rewrite -(subnKC w1gt2) -(@natrB _ _ 1) // ler_wpmul2l ?leC_nat //. + by rewrite Cnat_ge0 ?rpredX. +rewrite -(ltr_pmul2l (gt0CG W1)) -/w1 2!mulrBr mulr1 mulrCA -exprMn. +rewrite mulrDr ltr_subl_addl addrCA -mulrDr mulrCA mulrA -ltr_subl_addl. +rewrite -mulrBr mulNr opprK divfK ?neq0CG // mulr_natr addrA subrK -subr_sqr. +rewrite sqrr_sign mulrC [_ + 2%:R]addrC (ltr_le_trans _ ub_da2) //. +apply: ltr_le_trans (ler_wpmul2l (ler0n _ _) a2_ge1). +by rewrite mulr1 ltr_subl_addl -mulrS -natrX ltC_nat. +Qed. + +(* This is the first part of Peterfalvi (10.6)(a). *) +Let tau1mu j : j != 0 -> (mu_ j)^\tau1 = delta *: \sum_i eta_ i j. +Proof. +move=> nz_j; have [[[Itau1 _] _] Smu_j] := (cohS, calSmu nz_j). +have eta_mu i: '[delta *: (eta_ i j - eta_ i 0), (mu_ j)^\tau1] = 1. + have Szeta_s: zeta^*%CF \in calS by rewrite cfAut_seqInd. + have o_zeta_s_w k: '[eta_ i k, d%:R *: zeta^*^\tau1] = 0. + have o_S_eta_ := coherent_ortho_cycTIiso MtypeP sSS0 cohS. + by rewrite cfdotZr cfdotC o_S_eta_ ?conjC0 ?mulr0 // cfAut_irr. + pose psi := mu_ j - d%:R *: zeta^*%CF; rewrite (canRL (subrK _) (erefl psi)). + rewrite (raddfD tau1) raddfZnat cfdotDr addrC cfdotZl cfdotBl !{}o_zeta_s_w. + rewrite subr0 mulr0 add0r -(canLR (subrK _) (tau_alpha i nz_j)). + have Zpsi: psi \in 'Z[calS, M'^#]. + by rewrite ZmuBzeta // cfunE zeta1w1 rmorph_nat. + rewrite cfdotDl cfdotZl Itau1 ?(zcharW Zpsi) ?mem_zchar // -cfdotZl Dtau1 //. + rewrite Dade_isometry ?alpha_on ?tauM' {Zpsi}// -cfdotDl cfdotBr cfdotZr. + rewrite subrK !cfdotBl !cfdotZl !cfdot_prTIirr_red eq_sym (negPf nz_j). + by rewrite !o_mu2_irr ?cfAut_irr // !(mulr0, subr0) eqxx. +have [_ Ddel _ _] := FTtype345_constants. +have [[d1 k] Dtau1mu] := FTtypeP_coherent_TIred sSS0 cohS irr_zeta Szeta Smu_j. +case=> [[Dd1 Dk] | [_ Dk _]]; first by rewrite Dtau1mu Dd1 Dk [_ ^+ _]Ddel. +have /esym/eqP/idPn[] := eta_mu 0; rewrite Dtau1mu Dk /= cfdotZl cfdotZr. +rewrite cfdot_sumr big1 ?mulr0 ?oner_eq0 // => i _; rewrite -/sigma -/(w_ i _). +rewrite cfdotBl !(cfdot_cycTIiso pddM) !(eq_sym 0) conjC_Iirr_eq0 -!irr_eq1. +rewrite (eq_sym j) -(inj_eq irr_inj) conjC_IirrE. +by rewrite odd_eq_conj_irr1 ?mFT_odd ?subrr. +Qed. + +(* This is the second part of Peterfalvi (10.6)(a). *) +Let tau1mu0 : (mu_ 0 - zeta)^\tau = \sum_i eta_ i 0 - zeta^\tau1. +Proof. +have [j nz_j] := has_nonprincipal_irr ntW2. +have sum_al: \sum_i alpha_ i j = mu_ j - d%:R *: zeta - delta *: (mu_ 0 - zeta). + rewrite scalerBr opprD addrACA scaler_sumr !sumrB sumr_const; congr (_ + _). + by rewrite -opprD -scalerBl nirrW1 -scaler_nat scalerA mulrC divfK ?neq0CG. +have ->: mu_ 0 - zeta = delta *: (mu_ j - d%:R *: zeta - \sum_i alpha_ i j). + by rewrite sum_al opprD addNKr opprK signrZK. +rewrite linearZ linearB; apply: canLR (signrZK _) _; rewrite -/delta /=. +rewrite linear_sum -Dtau1 ?ZmuBzeta //= raddfB raddfZnat addrAC scalerBr. +rewrite (eq_bigr _ (fun i _ => tau_alpha i nz_j)) sumrB sumr_const nirrW1 opprD. +rewrite -scaler_sumr sumrB scalerBr -tau1mu // opprD !opprK -!addrA addNKr. +congr (_ + _); rewrite -scaler_nat scalerA mulrC divfK ?neq0CG //. +by rewrite addrC -!scaleNr -scalerDl addKr. +Qed. + +(* This is Peterfalvi (10.6)(b). *) +Let zeta_tau1_coprime g : + g \notin 'A~(M) -> coprime #[g] w1 -> `|zeta^\tau1 g| >= 1. +Proof. +move=> notAg co_g_w1; have Amu0zeta := mu0Bzeta_on Szeta zeta1w1. +have mu0_zeta_g: (mu_ 0 - zeta)^\tau g = 0. + have [ | ] := boolP (g \in 'A0~(M)); rewrite -FT_Dade0_supportE; last first. + by apply: cfun_on0; apply: Dade_cfunS. + case/bigcupP=> x A0x xRg; rewrite (DadeE _ A0x) // (cfun_on0 Amu0zeta) //. + apply: contra notAg => Ax; apply/bigcupP; exists x => //. + by rewrite -def_FTsignalizer0. +have{mu0_zeta_g} zeta_g: zeta^\tau1 g = \sum_i eta_ i 0 g. + by apply/esym/eqP; rewrite -subr_eq0 -{2}mu0_zeta_g tau1mu0 !cfunE sum_cfunE. +have Zwg i: eta_ i 0 g \in Cint. + have Lchi: 'chi_i \is a linear_char by apply: irr_cyclic_lin. + rewrite Cint_cycTIiso_coprime // dprod_IirrE irr0 cfDprod_cfun1r. + rewrite (coprime_dvdr _ co_g_w1) // dvdn_cforder. + rewrite -rmorphX cfDprodl_eq1 -dvdn_cforder; apply/dvdn_cforderP=> x W1x. + by rewrite -lin_charX // -expg_mod_order (eqnP (order_dvdG W1x)) lin_char1. +have odd_zeta_g: (zeta^\tau1 g == 1 %[mod 2])%C. + rewrite zeta_g (bigD1 0) //= [w_ 0 0]cycTIirr00 cycTIiso1 cfun1E inE. + pose eW1 := [pred i : Iirr W1 | conjC_Iirr i < i]%N. + rewrite (bigID eW1) (reindex_inj (can_inj (@conjC_IirrK _ _))) /=. + set s1 := \sum_(i | _) _; set s2 := \sum_(i | _) _; suffices ->: s1 = s2. + by rewrite -mulr2n addrC -(mulr_natr _ 2) eqCmod_addl_mul ?rpred_sum. + apply/eq_big=> [i | i _]. + rewrite (canF_eq (@conjC_IirrK _ _)) conjC_Iirr0 conjC_IirrK -leqNgt. + rewrite ltn_neqAle val_eqE -irr_eq1 (eq_sym i) -(inj_eq irr_inj) andbA. + by rewrite aut_IirrE odd_eq_conj_irr1 ?mFT_odd ?andbb. + rewrite -{1}conjC_Iirr0 [w_ _ _]cycTIirr_aut -cfAut_cycTIiso. + by rewrite cfunE conj_Cint ?Zwg. +rewrite norm_Cint_ge1 //; first by rewrite zeta_g rpred_sum. +apply: contraTneq odd_zeta_g => ->. +by rewrite eqCmod_sym /eqCmod subr0 /= (dvdC_nat 2 1). +Qed. + +(* This is Peterfalvi (10.7). *) +Let Frob_der1_type2 S : + S \in 'M -> FTtype S == 2 -> [Frobenius S^`(1) with kernel S`_\F]. +Proof. +move: S => L maxL /eqP Ltype2. +have [S pairMS [xdefW [U StypeP]]] := FTtypeP_pair_witness maxM MtypeP. +have [[_ _ maxS] _] := pairMS; rewrite {1}(negPf notMtype2) /= => Stype2 _. +move/(_ L maxL)/implyP; rewrite Ltype2 /= => /setUP[] /imsetP[x0 _ defL]. + by case/eqP/idPn: Ltype2; rewrite defL FTtypeJ. +pose H := (S`_\F)%G; pose HU := (S^`(1))%G. +suffices{L Ltype2 maxL x0 defL}: [Frobenius HU = H ><| U]. + by rewrite defL derJ FcoreJ FrobeniusJker; apply: FrobeniusWker. +have sHHU: H \subset HU by have [_ [_ _ _ /sdprodW/mulG_sub[]]] := StypeP. +pose calT := seqIndD HU S H 1; pose tauS := FT_Dade0 maxS. +have DcalTs: calT = seqIndD HU S S`_\s 1. + by congr (seqIndD _ _ _ _); apply: val_inj; rewrite /= FTcore_type2. +have notFrobM: ~~ [Frobenius M with kernel M`_\F]. + by apply/existsP=> [[U1 /Frobenius_of_typeF/(typePF_exclusion MtypeP)]]. +have{notFrobM} notSsupportM: ~~ [exists x, FTsupports M (S :^ x)]. + apply: contra notFrobM => /'exists_existsP[x [y /and3P[Ay not_sCyM sCySx]]]. + have [_ [_ /(_ y)uMS] /(_ y)] := FTsupport_facts maxM. + rewrite inE (subsetP (FTsupp_sub0 _)) //= in uMS *. + rewrite -(eq_uniq_mmax (uMS not_sCyM) _ sCySx) ?mmaxJ // FTtypeJ. + by case=> // _ _ _ [_ ->]. +have{notSsupportM} tiA1M_AS: [disjoint 'A1~(M) & 'A~(S)]. + have notMG_S: gval S \notin M :^: G. + by apply: contraL Stype2 => /imsetP[x _ ->]; rewrite FTtypeJ. + by apply: negbNE; have [_ <- _] := FT_Dade_support_disjoint maxM maxS notMG_S. +pose pddS := FT_prDade_hypF maxS StypeP; pose nu := primeTIred pddS. +have{tiA1M_AS} oST phi psi: + phi \in 'Z[calS, M^#] -> psi \in 'Z[calT, S^#] -> '[phi^\tau, tauS psi] = 0. +- rewrite zcharD1_seqInd // -[seqInd _ _]/calS => Sphi. + rewrite zcharD1E => /andP[Tpsi psi1_0]. + rewrite -FT_Dade1E ?defA1 ?(zchar_on Sphi) //. + apply: cfdot_complement (Dade_cfunS _ _) _; rewrite FT_Dade1_supportE setTD. + rewrite -[tauS _]FT_DadeE ?(cfun_onS _ (Dade_cfunS _ _)) ?FT_Dade_supportE //. + by rewrite -disjoints_subset disjoint_sym. + have /subsetD1P[_ /setU1K <-] := FTsupp_sub S; rewrite cfun_onD1 {}psi1_0. + rewrite -Tpsi andbC -zchar_split {psi Tpsi}(zchar_trans_on _ Tpsi) //. + move=> psi Tpsi; rewrite zchar_split mem_zchar //=. + have [s /setDP[_ kerH's] ->] := seqIndP Tpsi. + by rewrite inE in kerH's; rewrite (prDade_Ind_irr_on pddS). +have notStype5: FTtype S != 5 by rewrite (eqP Stype2). +have [|[_ _ _ _ -> //]] := typeP_reducible_core_cases maxS StypeP notStype5. +case=> t []; set lambda := 'chi_t => T0C'lam lam_1 _. +have{T0C'lam} Tlam: lambda \in calT. + by apply: seqIndS T0C'lam; rewrite Iirr_kerDS ?sub1G. +have{lam_1} [r [nz_r Tnu_r nu_r_1]]: + exists r, [/\ r != 0, nu r \in calT & nu r 1%g = lambda 1%g]. +- have [_] := typeP_reducible_core_Ind maxS StypeP notStype5. + set H0 := Ptype_Fcore_kernel _; set nuT := filter _ _; rewrite -/nu. + case/hasP=> nu_r nuTr _ /(_ _ nuTr)/imageP[r nz_r Dr] /(_ _ nuTr)[nu_r1 _ _]. + have{nuTr} Tnu_r := mem_subseq (filter_subseq _ _) nuTr. + by exists r; rewrite -Dr nu_r1 (seqIndS _ Tnu_r) // Iirr_kerDS ?sub1G. +pose T2 := [:: lambda; lambda^*; nu r; (nu r)^*]%CF. +have [rmRS scohT]: exists rmRS, subcoherent calT tauS rmRS. + move: (FTtypeP_coh_base _ _) (FTtypeP_subcoherent maxS StypeP) => RS scohT. + by rewrite DcalTs; exists RS. +have [lam_irr nu_red]: lambda \in irr S /\ nu r \notin irr S. + by rewrite mem_irr prTIred_not_irr. +have [lam'nu lams'nu]: lambda != nu r /\ lambda^*%CF != nu r. + by rewrite -conjC_IirrE !(contraNneq _ nu_red) // => <-; apply: mem_irr. +have [[_ nRT ccT] _ _ _ _] := scohT. +have{ccT} sT2T: {subset T2 <= calT} by apply/allP; rewrite /= ?Tlam ?Tnu_r ?ccT. +have{nRT} uccT2: cfConjC_subset T2 calT. + split; last 1 [by [] | by apply/allP; rewrite /= !inE !cfConjCK !eqxx !orbT]. + rewrite /uniq /T2 !inE !negb_or -!(inv_eq (@cfConjCK _ S)) !cfConjCK. + by rewrite lam'nu lams'nu !(hasPn nRT). +have scohT2 := subset_subcoherent scohT uccT2. +have [tau2 cohT2]: coherent T2 S^# tauS. + apply: (uniform_degree_coherence scohT2); rewrite /= !cfunE nu_r_1 eqxx. + by rewrite conj_Cnat ?Cnat_irr1 ?eqxx. +have [s nz_s] := has_nonprincipal_irr ntW2; have Smu_s := calSmu nz_s. +pose alpha := mu_ s - d%:R *: zeta; pose beta := nu r - lambda. +have Salpha: alpha \in 'Z[calS, M^#] by rewrite zcharD1_seqInd ?ZmuBzeta. +have [T2lam T2nu_r]: lambda \in T2 /\ nu r \in T2 by rewrite !inE !eqxx !orbT. +have Tbeta: beta \in 'Z[T2, S^#]. + by rewrite zcharD1E rpredB ?mem_zchar //= !cfunE nu_r_1 subrr. +have /eqP/idPn[] := oST _ _ Salpha (zchar_subset sT2T Tbeta). +have [[_ <- //] [_ <- //]] := (cohS, cohT2). +rewrite !raddfB raddfZnat /= subr_eq0 !cfdotBl !cfdotZl. +have [|[dr r'] -> _] := FTtypeP_coherent_TIred _ cohT2 lam_irr T2lam T2nu_r. + by rewrite -DcalTs. +set sigS := cyclicTIiso _ => /=. +have etaC i j: sigS (cyclicTIirr xdefW j i) = eta_ i j by apply: cycTIisoC. +rewrite !cfdotZr addrC cfdot_sumr big1 => [|j _]; last first. + by rewrite etaC (coherent_ortho_cycTIiso _ sSS0) ?mem_irr. +rewrite !mulr0 oppr0 add0r rmorph_sign. +have ->: '[zeta^\tau1, tau2 lambda] = 0. + pose X1 := (zeta :: zeta^*)%CF; pose X2 := (lambda :: lambda^*)%CF. + pose Y1 := map tau1 X1; pose Y2 := map tau2 X2; have [_ _ ccS] := sSS0. + have [sX1S sX2T]: {subset X1 <= calS} /\ {subset X2 <= T2}. + by split; apply/allP; rewrite /= ?inE ?eqxx ?orbT // Szeta ccS. + have [/(sub_iso_to (zchar_subset sX1S) sub_refl)[Itau1 Ztau1] Dtau1L] := cohS. + have [/(sub_iso_to (zchar_subset sX2T) sub_refl)[Itau2 Ztau2] Dtau2] := cohT2. + have Z_Y12: {subset Y1 <= 'Z[irr G]} /\ {subset Y2 <= 'Z[irr G]}. + by rewrite /Y1 /Y2; split=> ? /mapP[xi /mem_zchar] => [/Ztau1|/Ztau2] ? ->. + have o1Y12: orthonormal Y1 && orthonormal Y2. + rewrite !map_orthonormal //. + by apply: seqInd_conjC_ortho2 Tlam; rewrite ?gFnormal ?mFT_odd. + by apply: seqInd_conjC_ortho2 Szeta; rewrite ?gFnormal ?mFT_odd ?mem_irr. + apply: orthonormal_vchar_diff_ortho Z_Y12 o1Y12 _; rewrite -2!raddfB. + have SzetaBs: zeta - zeta^*%CF \in 'Z[calS, M^#]. + by rewrite zcharD1_seqInd // seqInd_sub_aut_zchar. + have T2lamBs: lambda - lambda^*%CF \in 'Z[T2, S^#]. + rewrite sub_aut_zchar ?zchar_onG ?mem_zchar ?inE ?eqxx ?orbT //. + by move=> xi /sT2T/seqInd_vcharW. + by rewrite Dtau1L // Dtau2 // !Dade1 oST ?(zchar_subset sT2T) ?eqxx. +have [[ds s'] /= -> _] := FTtypeP_coherent_TIred sSS0 cohS irr_zeta Szeta Smu_s. +rewrite mulr0 subr0 !cfdotZl mulrA -signr_addb !cfdot_suml. +rewrite (bigD1 r') //= cfdot_sumr (bigD1 s') //=. +rewrite etaC cfdot_cycTIiso !eqxx big1 => [|j ne_s'_j]; last first. + by rewrite etaC cfdot_cycTIiso andbC eq_sym (negPf ne_s'_j). +rewrite big1 => [|i ne_i_r']; last first. + rewrite cfdot_sumr big1 // => j _. + by rewrite etaC cfdot_cycTIiso (negPf ne_i_r'). +rewrite !addr0 mulr1 big1 ?mulr0 ?signr_eq0 // => i _. +by rewrite -etaC cfdotC (coherent_ortho_cycTIiso _ _ cohT2) ?conjC0 -?DcalTs. +Qed. + +(* This is the bulk of the proof of Peterfalvi (10.8); however the result *) +(* will be restated below to avoid the quantification on zeta and tau1. *) +Lemma FTtype345_noncoherence_main : False. +Proof. +have [S pairMS [xdefW [U StypeP]]] := FTtypeP_pair_witness maxM MtypeP. +have [[_ _ maxS] _] := pairMS; rewrite {1}(negPf notMtype2) /= => Stype2 _ _. +pose H := (S`_\F)%G; pose HU := (S^`(1))%G. +have [[_ hallW2 _ defS] [_ _ nUW2 defHU] _ [_ _ sW1H _ _] _] := StypeP. +have ntU: U :!=: 1%g by have [[]] := compl_of_typeII maxS StypeP Stype2. +pose G01 := [set g : gT | coprime #[g] w1]. +pose G0 := ~: 'A~(M) :&: G01; pose G1 := ~: 'A~(M) :\: G01. +pose chi := zeta^\tau1; pose ddAM := FT_Dade_hyp maxM; pose rho := invDade ddAM. +have Suzuki: + #|G|%:R^-1 * (\sum_(g in ~: 'A~(M)) `|chi g| ^+ 2 - #|~: 'A~(M)|%:R) + + '[rho chi] - #|'A(M)|%:R / #|M|%:R <= 0. +- pose A_ (_ : 'I_1) := ddAM; pose Atau i := Dade_support (A_ i). + have tiA i j : i != j -> [disjoint Atau i & Atau j] by rewrite !ord1. + have Nchi1: '[chi] = 1 by have [[->]] := cohS; rewrite ?mem_zchar ?irrWnorm. + have:= Dade_cover_inequality tiA Nchi1; rewrite /= !big_ord1 -/rho -addrA. + by congr (_ * _ + _ <= 0); rewrite FT_Dade_supportE setTD. +have{Suzuki} ub_rho: '[rho chi] <= #|'A(M)|%:R / #|M|%:R + #|G1|%:R / #|G|%:R. + rewrite addrC -subr_le0 opprD addrCA (ler_trans _ Suzuki) // -addrA. + rewrite ler_add2r -(cardsID G01 (~: _)) (big_setID G01) -/G0 -/G1 /=. + rewrite mulrC mulrBr ler_subr_addl -mulrBr natrD addrK. + rewrite ler_wpmul2l ?invr_ge0 ?ler0n // -sumr_const ler_paddr //. + by apply: sumr_ge0 => g; rewrite exprn_ge0 ?normr_ge0. + apply: ler_sum => g; rewrite !inE => /andP[notAg] /(zeta_tau1_coprime notAg). + by rewrite expr_ge1 ?normr_ge0. +have lb_M'bar: (w1 * 2 <= #|M' / M''|%g.-1)%N. + suffices ->: w1 = #|W1 / M''|%g. + rewrite muln2 -ltnS prednK ?cardG_gt0 //. + by rewrite (ltn_odd_Frobenius_ker frobMbar) ?quotient_odd ?mFT_odd. + have [_ sW1M _ _ tiM'W1] := sdprod_context defM. + apply/card_isog/quotient_isog; first exact: subset_trans (der_norm 2 M). + by apply/trivgP; rewrite -tiM'W1 setSI ?normal_sub. +have lb_rho: 1 - w1%:R / #|M'|%:R <= '[rho chi]. + have cohS_A: coherent_with calS M^# (Dade ddAM) tau1. + have [Itau1 _] := cohS; split=> // phi; rewrite zcharD1_seqInd // => Sphi. + by rewrite Dtau1 // FT_DadeE // defA (zchar_on Sphi). + rewrite {ub_rho}/rho [w1](index_sdprod defM); rewrite defA in (ddAM) cohS_A *. + have [||_ [_ _ [] //]] := Dade_Ind1_sub_lin cohS_A _ irr_zeta Szeta. + - by apply: seqInd_nontrivial Szeta; rewrite ?mem_irr ?mFT_odd. + - by rewrite -(index_sdprod defM). + rewrite -(index_sdprod defM) ler_pdivl_mulr ?ltr0n // -natrM. + rewrite -leC_nat in lb_M'bar; apply: ler_trans lb_M'bar _. + rewrite ler_subr_addl -mulrS prednK ?cardG_gt0 // leC_nat. + by rewrite dvdn_leq ?dvdn_quotient. +have{lb_rho ub_rho}: 1 - #|G1|%:R/ #|G|%:R - w1%:R^-1 < w1%:R / #|M'|%:R :> algC. + rewrite -addrA -opprD ltr_subl_addr -ltr_subl_addl. + apply: ler_lt_trans (ler_trans lb_rho ub_rho) _; rewrite addrC ltr_add2l. + rewrite ltr_pdivr_mulr ?gt0CG // mulrC -(sdprod_card defM) natrM. + by rewrite mulfK ?neq0CG // defA ltC_nat (cardsD1 1%g M') group1. +have frobHU: [Frobenius HU with kernel H] by apply: Frob_der1_type2. +have tiH: normedTI H^# G S. + by have [_ _] := FTtypeII_ker_TI maxS Stype2; rewrite FTsupp1_type2. +have sG1_HVG: G1 \subset class_support H^# G :|: class_support V G. + apply/subsetP=> x; rewrite !inE coprime_has_primes ?cardG_gt0 // negbK. + case/andP=> /hasP[p W1p]; rewrite /= mem_primes => /and3P[p_pr _ p_dv_x] _. + have [a x_a a_p] := Cauchy p_pr p_dv_x. + have nta: a != 1%g by rewrite -order_gt1 a_p prime_gt1. + have ntx: x != 1%g by apply: contraTneq x_a => ->; rewrite /= cycle1 inE. + have cxa: a \in 'C[x] by rewrite -cent_cycle (subsetP (cycle_abelian x)). + have hallH: \pi(H).-Hall(G) H by apply: Hall_pi; have [] := FTcore_facts maxS. + have{a_p} p_a: p.-elt a by rewrite /p_elt a_p pnat_id. + have piHp: p \in \pi(H) by rewrite (piSg _ W1p). + have [y _ Hay] := Hall_pJsub hallH piHp (subsetT _) p_a. + do [rewrite -cycleJ cycle_subG; set ay := (a ^ y)%g] in Hay. + rewrite -[x](conjgK y); set xy := (x ^ y)%g. + have caxy: xy \in 'C[ay] by rewrite cent1J memJ_conjg cent1C. + have [ntxy ntay]: xy != 1%g /\ ay != 1%g by rewrite !conjg_eq1. + have Sxy: xy \in S. + have H1ay: ay \in H^# by apply/setD1P. + by rewrite (subsetP (cent1_normedTI tiH H1ay)) ?setTI. + have [HUxy | notHUxy] := boolP (xy \in HU). + rewrite memJ_class_support ?inE ?ntxy //=. + have [_ _ _ regHUH] := Frobenius_kerP frobHU. + by rewrite (subsetP (regHUH ay _)) // inE ?HUxy // inE ntay. + suffices /imset2P[xyz z Vxzy _ ->]: xy \in class_support V S. + by rewrite -conjgM orbC memJ_class_support. + rewrite /V setUC -(FTsupp0_typeP maxS StypeP) !inE Sxy. + rewrite andb_orr andNb (contra (subsetP _ _) notHUxy) /=; last first. + by apply/bigcupsP=> z _; rewrite (eqP Stype2) setDE -setIA subsetIl. + have /Hall_pi hallHU: Hall S HU by rewrite (sdprod_Hall defS). + rewrite (eqP Stype2) -(mem_normal_Hall hallHU) ?gFnormal // notHUxy. + have /mulG_sub[sHHU _] := sdprodW defHU. + rewrite (contra (fun p'xy => pi'_p'group p'xy (piSg sHHU piHp))) //. + by rewrite pgroupE p'natE // cycleJ cardJg p_dv_x. +have ub_G1: #|G1|%:R / #|G|%:R <= #|H|%:R / #|S|%:R + #|V|%:R / #|W|%:R :> algC. + rewrite ler_pdivr_mulr ?ltr0n ?cardG_gt0 // mulrC mulrDr !mulrA. + rewrite ![_ * _ / _]mulrAC -!natf_indexg ?subsetT //= -!natrM -natrD ler_nat. + apply: leq_trans (subset_leq_card sG1_HVG) _. + rewrite cardsU (leq_trans (leq_subr _ _)) //. + have unifJG B C: C \in B :^: G -> #|C| = #|B|. + by case/imsetP=> z _ ->; rewrite cardJg. + have oTI := card_uniform_partition (unifJG _) (partition_class_support _ _). + have{tiH} [ntH tiH /eqP defNH] := and3P tiH. + have [_ _ /and3P[ntV tiV /eqP defNV]] := ctiWG. + rewrite !oTI // !card_conjugates defNH defNV /= leq_add2r ?leq_mul //. + by rewrite subset_leq_card ?subsetDl. +rewrite ler_gtF // addrAC ler_subr_addl -ler_subr_addr (ler_trans ub_G1) //. +rewrite -(sdprod_card defS) -(sdprod_card defHU) addrC. +rewrite -mulnA !natrM invfM mulVKf ?natrG_neq0 // -/w1 -/w2. +have sW12_W: W1 :|: W2 \subset W by rewrite -(dprodWY defW) sub_gen. +rewrite cardsD (setIidPr sW12_W) natrB ?subset_leq_card // mulrBl. +rewrite divff ?natrG_neq0 // -!addrA ler_add2l. +rewrite cardsU -(dprod_card defW) -/w1 -/w2; have [_ _ _ ->] := dprodP defW. +rewrite cards1 natrB ?addn_gt0 ?cardG_gt0 // addnC natrD -addrA mulrDl mulrBl. +rewrite {1}mulnC !natrM !invfM !mulVKf ?natrG_neq0 // opprD -addrA ler_add2l. +rewrite mul1r -{1}[_^-1]mul1r addrC ler_oppr [- _]opprB -!mulrBl. +rewrite -addrA -opprD ler_pdivl_mulr; last by rewrite natrG_gt0. +apply: ler_trans (_ : 1 - (3%:R^-1 + 7%:R^-1) <= _); last first. + rewrite ler_add2l ler_opp2. + rewrite ler_add // lef_pinv ?qualifE ?gt0CG ?ltr0n ?ler_nat //. + have notStype5: FTtype S != 5 by rewrite (eqP Stype2). + have frobUW2 := Ptype_compl_Frobenius maxS StypeP notStype5. + apply: leq_ltn_trans (ltn_odd_Frobenius_ker frobUW2 (mFT_odd _)). + by rewrite (leq_double 3). +apply: ler_trans (_ : 2%:R^-1 <= _); last by rewrite -!CratrE; compute. +rewrite mulrAC ler_pdivr_mulr 1?gt0CG // ler_pdivl_mull ?ltr0n //. +rewrite -!natrM ler_nat mulnA -(Lagrange (normal_sub nsM''M')) mulnC leq_mul //. + by rewrite subset_leq_card //; have [_ _ _ []] := MtypeP. +by rewrite -card_quotient ?normal_norm // mulnC -(prednK (cardG_gt0 _)) leqW. +Qed. + +End NonCoherence. + +(* This is Peterfalvi (10.9). *) +Lemma FTtype345_Dade_bridge0 : + (w1 < w2)%N -> + {chi | [/\ (mu_ 0 - zeta)^\tau = \sum_i eta_ i 0 - chi, + chi \in 'Z[irr G], '[chi] = 1 + & forall i j, '[chi, eta_ i j] = 0]}. +Proof. +move=> w1_lt_w2; set psi := mu_ 0 - zeta; pose Wsig := map sigma (irr W). +have [X wsigX [chi [DpsiG _ o_chiW]]] := orthogonal_split Wsig psi^\tau. +exists (- chi); rewrite opprK rpredN cfnormN. +have o_chi_w i j: '[chi, eta_ i j] = 0. + by rewrite (orthoPl o_chiW) ?map_f ?mem_irr. +have [Isigma Zsigma] := cycTI_Zisometry ctiWG. +have o1Wsig: orthonormal Wsig by rewrite map_orthonormal ?irr_orthonormal. +have [a_ Da defX] := orthonormal_span o1Wsig wsigX. +have{Da} Da i j: a_ (eta_ i j) = '[psi^\tau, eta_ i j]. + by rewrite DpsiG cfdotDl o_chi_w addr0 Da. +have sumX: X = \sum_i \sum_j a_ (eta_ i j) *: eta_ i j. + rewrite pair_bigA defX big_map (big_nth 0) size_tuple big_mkord /=. + rewrite (reindex (dprod_Iirr defW)) /=. + by apply: eq_bigr => [[i j] /= _]; rewrite -tnth_nth. + by exists (inv_dprod_Iirr defW) => ij; rewrite (inv_dprod_IirrK, dprod_IirrK). +have Zpsi: psi \in 'Z[irr M]. + by rewrite rpredB ?irr_vchar ?(mem_zchar irr_zeta) ?char_vchar ?prTIred_char. +have{Zpsi} M'psi: psi \in 'Z[irr M, M'^#]. + by rewrite -defA zchar_split Zpsi mu0Bzeta_on. +have A0psi: psi \in 'CF(M, 'A0(M)). + by apply: cfun_onS (zchar_on M'psi); rewrite defA0 subsetUl. +have a_00: a_ (eta_ 0 0) = 1. + rewrite Da [w_ 0 0](cycTIirr00 defW) [sigma 1]cycTIiso1. + rewrite Dade_reciprocity // => [|x _ y _]; last by rewrite !cfun1E !inE. + rewrite rmorph1 /= -(prTIirr00 ptiWM) -/(mu2_ 0 0) cfdotC. + by rewrite cfdotBr o_mu2_zeta subr0 cfdot_prTIirr_red rmorph1. +have n2psiG: '[psi^\tau] = w1.+1%:R. + rewrite Dade_isometry // cfnormBd ?o_mu_zeta //. + by rewrite cfnorm_prTIred irrWnorm // -/w1 mulrSr. +have psiG_V0 x: x \in V -> psi^\tau x = 0. + move=> Vx; rewrite Dade_id ?defA0; last first. + by rewrite inE orbC mem_class_support. + rewrite (cfun_on0 (zchar_on M'psi)) // -defA. + suffices /setDP[]: x \in 'A0(M) :\: 'A(M) by []. + by rewrite (FTsupp0_typeP maxM MtypeP) // mem_class_support. +have ZpsiG: psi^\tau \in 'Z[irr G]. + by rewrite Dade_vchar // zchar_split (zcharW M'psi). +have n2psiGsum: '[psi^\tau] = \sum_i \sum_j `|a_ (eta_ i j)| ^+ 2 + '[chi]. + rewrite DpsiG addrC cfnormDd; last first. + by rewrite (span_orthogonal o_chiW) ?memv_span1. + rewrite addrC defX cfnorm_sum_orthonormal // big_map pair_bigA; congr (_ + _). + rewrite big_tuple /= (reindex (dprod_Iirr defW)) //. + by exists (inv_dprod_Iirr defW) => ij; rewrite (inv_dprod_IirrK, dprod_IirrK). +have NCpsiG: (cyclicTI_NC ctiWG psi^\tau < 2 * minn w1 w2)%N. + apply: (@leq_ltn_trans w1.+1); last first. + by rewrite /minn w1_lt_w2 mul2n -addnn (leq_add2r w1 2) cardG_gt1. + pose z_a := [pred ij | a_ (eta_ ij.1 ij.2) == 0]. + have ->: cyclicTI_NC ctiWG psi^\tau = #|[predC z_a]|. + by apply: eq_card => ij; rewrite !inE -Da. + rewrite -leC_nat -n2psiG n2psiGsum ler_paddr ?cfnorm_ge0 // pair_bigA. + rewrite (bigID z_a) big1 /= => [|ij /eqP->]; last by rewrite normCK mul0r. + rewrite add0r -sumr_const ler_sum // => [[i j] nz_ij]. + by rewrite expr_ge1 ?norm_Cint_ge1 // Da Cint_cfdot_vchar ?Zsigma ?irr_vchar. +have nz_psiG00: '[psi^\tau, eta_ 0 0] != 0 by rewrite -Da a_00 oner_eq0. +have [a_i|a_j] := small_cycTI_NC psiG_V0 NCpsiG nz_psiG00. + have psiGi: psi^\tau = \sum_i eta_ i 0 + chi. + rewrite DpsiG sumX; congr (_ + _); apply: eq_bigr => i _. + rewrite big_ord_recl /= Da a_i -Da a_00 mul1r scale1r. + by rewrite big1 ?addr0 // => j1 _; rewrite Da a_i mul0r scale0r. + split=> // [||i j]; last by rewrite cfdotNl o_chi_w oppr0. + rewrite -(canLR (addKr _) psiGi) rpredD // rpredN rpred_sum // => j _. + by rewrite Zsigma ?irr_vchar. + apply: (addrI w1%:R); rewrite -mulrSr -n2psiG n2psiGsum; congr (_ + _). + rewrite -nirrW1 // -sumr_const; apply: eq_bigr => i _. + rewrite big_ord_recl /= Da a_i -Da a_00 mul1r normr1. + by rewrite expr1n big1 ?addr0 // => j1 _; rewrite Da a_i normCK !mul0r. +suffices /idPn[]: '[psi^\tau] >= w2%:R. + rewrite odd_geq /= ?uphalf_half mFT_odd //= in w1_lt_w2. + by rewrite n2psiG leC_nat -ltnNge odd_geq ?mFT_odd. +rewrite n2psiGsum exchange_big /= ler_paddr ?cfnorm_ge0 //. +rewrite -nirrW2 -sumr_const; apply: ler_sum => i _. +rewrite big_ord_recl /= Da a_j -Da a_00 mul1r normr1. +by rewrite expr1n big1 ?addr0 // => j1 _; rewrite Da a_j normCK !mul0r. +Qed. + +Local Notation H := M'. +Local Notation "` 'H'" := `M' (at level 0) : group_scope. +Local Notation H' := M''. +Local Notation "` 'H''" := `M'' (at level 0) : group_scope. + +(* This is the bulk of the proof of Peterfalvi, Theorem (10.10); as with *) +(* (10.8), it will be restated below in order to remove dependencies on zeta, *) +(* U_M and W1. *) +Lemma FTtype5_exclusion_main : FTtype M != 5. +Proof. +apply/negP=> Mtype5. +suffices [tau1]: coherent calS M^# tau by case/FTtype345_noncoherence_main. +have [[_ U_M_1] MtypeV] := compl_of_typeV maxM MtypeP Mtype5. +have [_ [_ _ _ defH] _ [_ _ _ sW2H' _] _] := MtypeP. +have{U_M_1 defH} defMF: M`_\F = H by rewrite /= -defH U_M_1 sdprodg1. +have nilH: nilpotent `H by rewrite -defMF Fcore_nil. +have scohS := subset_subcoherent scohS0 sSS0. +have [|//|[[_]]] := (non_coherent_chief nsM'M (nilpotent_sol nilH) scohS) 1%G. + split; rewrite ?mFT_odd ?normal1 ?sub1G ?quotient_nil //. + by rewrite joingG1 (FrobeniusWker frobMbar). +rewrite /= joingG1 -(index_sdprod defM) /= -/w1 -[H^`(1)%g]/`H' => ubHbar [p[]]. +rewrite -(isog_abelian (quotient1_isog H)) -(isog_pgroup p (quotient1_isog H)). +rewrite subn1 => pH not_cHH /negP not_w1_dv_p'1. +have ntH: H :!=: 1%g by apply: contraNneq not_cHH => ->; apply: abelian1. +have [sH'H nH'H] := andP nsM''M'; have sW2H := subset_trans sW2H' sH'H. +have def_w2: w2 = p by apply/eqP; have:= pgroupS sW2H pH; rewrite pgroupE pnatE. +have [p_pr _ [e oH]] := pgroup_pdiv pH ntH. +rewrite -/w1 /= defMF oH pi_of_exp {e oH}// /pi_of primes_prime // in MtypeV. +have [tiHG | [_ /predU1P[->[]|]]// | [_ /predU1P[->|//] [oH w1p1 _]]] := MtypeV. + suffices [tau1 [Itau1 Dtau1]]: coherent (seqIndD H M H 1) M^# 'Ind[G]. + exists tau1; split=> // phi Sphi; rewrite {}Dtau1 //. + rewrite zcharD1_seqInd // -subG1 -setD_eq0 -defA in Sphi tiHG ntH. + by have Aphi := zchar_on Sphi; rewrite -FT_DadeE // Dade_Ind. + apply: (@Sibley_coherence _ [set:_] M H W1); first by rewrite mFT_odd. + right; exists W2 => //; exists 'A0(M), W, defW. + by rewrite -defA -{2}(group_inj defMs). +rewrite pcore_pgroup_id // in oH. +have esH: extraspecial H. + by apply: (p3group_extraspecial pH); rewrite // oH pfactorK. +have oH': #|H'| = p. + by rewrite -(card_center_extraspecial pH esH); have [[_ <-]] := esH. +have defW2: W2 :=: H' by apply/eqP; rewrite eqEcard sW2H' oH' -def_w2 /=. +have iH'H: #|H : H'|%g = (p ^ 2)%N by rewrite -divgS // oH oH' mulKn ?prime_gt0. +have w1_gt0: (0 < w1)%N by apply: cardG_gt0. +(* This is step (10.10.1). *) +have{ubHbar} [def_p_w1 w1_lt_w2]: (p = 2 * w1 - 1 /\ w1 < w2)%N. + have /dvdnP[k def_p]: 2 * w1 %| p.+1. + by rewrite Gauss_dvd ?coprime2n ?mFT_odd ?dvdn2 //= -{1}def_w2 mFT_odd. + suffices k1: k = 1%N. + rewrite k1 mul1n in def_p; rewrite -ltn_double -mul2n -def_p -addn1 addnK. + by rewrite -addnS -addnn def_w2 leq_add2l prime_gt1. + have [k0 | k_gt0] := posnP k; first by rewrite k0 in def_p. + apply/eqP; rewrite eqn_leq k_gt0 andbT -ltnS -ltn_double -mul2n. + rewrite -[(2 * k)%N]prednK ?muln_gt0 // ltnS -ltn_sqr 3?leqW //=. + rewrite -subn1 sqrn_sub ?muln_gt0 // expnMn muln1 mulnA ltnS leq_subLR. + rewrite addn1 addnS ltnS -mulnSr leq_pmul2l // -(leq_subLR _ 1). + rewrite (leq_trans (leq_pmulr _ w1_gt0)) // -(leq_pmul2r w1_gt0). + rewrite -mulnA mulnBl mul1n -2!leq_double -!mul2n mulnA mulnBr -!expnMn. + rewrite -(expnMn 2 _ 2) mulnCA -def_p -addn1 leq_subLR sqrnD muln1. + by rewrite (addnC p) mulnDr addnA leq_add2r addn1 addnS -iH'H. +(* This is step (10.10.2). *) +pose S1 := seqIndD H M H H'. +have sS1S: {subset S1 <= calS} by apply: seqIndS; rewrite Iirr_kerDS ?sub1G. +have irrS1: {subset S1 <= irr M}. + move=> _ /seqIndP[s /setDP[kerH' ker'H] ->]; rewrite !inE in kerH' ker'H. + rewrite -(quo_IirrK _ kerH') // mod_IirrE // cfIndMod // cfMod_irr //. + rewrite (irr_induced_Frobenius_ker (FrobeniusWker frobMbar)) //. + by rewrite quo_Iirr_eq0 // -subGcfker. +have S1w1: {in S1, forall xi : 'CF(M), xi 1%g = w1%:R}. + move=> _ /seqIndP[s /setDP[kerH' _] ->]; rewrite !inE in kerH'. + by rewrite cfInd1 // -(index_sdprod defM) lin_char1 ?mulr1 // lin_irr_der1. +have sS10: cfConjC_subset S1 calS0. + by apply: seqInd_conjC_subset1; rewrite /= defMs. +pose S2 := [seq mu_ j | j in predC1 0]. +have szS2: size S2 = p.-1. + by rewrite -def_w2 size_map -cardE cardC1 card_Iirr_abelian ?cyclic_abelian. +have uS2: uniq S2 by apply/dinjectiveP; apply: in2W (prTIred_inj pddM). +have redS2: {subset S2 <= [predC irr M]}. + by move=> _ /imageP[j _ ->]; apply: (prTIred_not_irr pddM). +have sS2S: {subset S2 <= calS} by move=> _ /imageP[j /calSmu Smu_j ->]. +have S1'2: {subset S2 <= [predC S1]}. + by move=> xi /redS2; apply: contra (irrS1 _). +have w1_dv_p21: w1 %| p ^ 2 - 1 by rewrite (subn_sqr p 1) addn1 dvdn_mull. +have [j nz_j] := has_nonprincipal_irr ntW2. +have [Dmu2_1 Ddelta_ lt1d Nn] := FTtype345_constants. +have{lt1d} [defS szS1 Dd Ddel Dn]: + [/\ perm_eq calS (S1 ++ S2), size S1 = (p ^ 2 - 1) %/ w1, + d = p, delta = -1 & n = 2%:R]. +- pose X_ (S0 : seq 'CF(M)) := [set s | 'Ind[M, H] 'chi_s \in S0]. + pose sumX_ cS0 := \sum_(s in X_ cS0) 'chi_s 1%g ^+ 2. + have defX1: X_ S1 = Iirr_kerD H H H'. + by apply/setP=> s; rewrite !inE mem_seqInd // !inE. + have defX: X_ calS = Iirr_kerD H H 1%g. + by apply/setP=> s; rewrite !inE mem_seqInd ?normal1 //= !inE. + have sumX1: sumX_ S1 = (p ^ 2)%:R - 1. + by rewrite /sumX_ defX1 sum_Iirr_kerD_square // iH'H indexgg mul1r. + have ->: size S1 = (p ^ 2 - 1) %/ w1. + apply/eqP; rewrite eqn_div // -eqC_nat mulnC [w1](index_sdprod defM). + rewrite (size_irr_subseq_seqInd _ (subseq_refl S1)) //. + rewrite natrB ?expn_gt0 ?prime_gt0 // -sumr_const -sumX1. + apply/eqP/esym/eq_bigr => s. + by rewrite defX1 !inE -lin_irr_der1 => /and3P[_ _ /eqP->]; rewrite expr1n. + have oX2: #|X_ S2| = p.-1. + by rewrite -(size_red_subseq_seqInd_typeP MtypeP uS2 sS2S). + have sumX2: (p ^ 2 * p.-1)%:R <= sumX_ S2 ?= iff (d == p). + rewrite /sumX_ (eq_bigr (fun _ => d%:R ^+ 2)) => [|s]; last first. + rewrite inE => /imageP[j1 nz_j1 Dj1]; congr (_ ^+ 2). + apply: (mulfI (neq0CiG M H)); rewrite -cfInd1 // -(index_sdprod defM). + by rewrite Dj1 (prTIred_1 pddM) Dmu2_1. + rewrite sumr_const oX2 mulrnA (mono_lerif (ler_pmuln2r _)); last first. + by rewrite -def_w2 -(subnKC w2gt2). + rewrite natrX (mono_in_lerif ler_sqr) ?rpred_nat // eq_sym lerif_nat. + apply/leqif_eq; rewrite dvdn_leq 1?ltnW //. + have: (mu2_ 0 j 1%g %| (p ^ 3)%N)%C. + by rewrite -(cfRes1 H) cfRes_prTIirr -oH dvd_irr1_cardG. + rewrite Dmu2_1 // dvdC_nat => /dvdn_pfactor[//|[_ d1|e _ ->]]. + by rewrite d1 in lt1d. + by rewrite expnS dvdn_mulr. + pose S3 := filter [predC S1 ++ S2] calS. + have sumX3: 0 <= sumX_ S3 ?= iff nilp S3. + rewrite /sumX_; apply/lerifP. + have [-> | ] := altP nilP; first by rewrite big_pred0 // => s; rewrite !inE. + rewrite -lt0n -has_predT => /hasP[xi S3xi _]. + have /seqIndP[s _ Dxi] := mem_subseq (filter_subseq _ _) S3xi. + rewrite (bigD1 s) ?inE -?Dxi //= ltr_spaddl ?sumr_ge0 // => [|s1 _]. + by rewrite exprn_gt0 ?irr1_gt0. + by rewrite ltrW ?exprn_gt0 ?irr1_gt0. + have [_ /esym] := lerif_add sumX2 sumX3. + have /(canLR (addKr _)) <-: sumX_ calS = sumX_ S1 + (sumX_ S2 + sumX_ S3). + rewrite [sumX_ _](big_setID (X_ S1)); congr (_ + _). + by apply: eq_bigl => s; rewrite !inE andb_idl // => /sS1S. + rewrite (big_setID (X_ S2)); congr (_ + _); apply: eq_bigl => s. + by rewrite !inE andb_idl // => S2s; rewrite [~~ _]S1'2 ?sS2S. + by rewrite !inE !mem_filter /= mem_cat orbC negb_or andbA. + rewrite sumX1 /sumX_ defX sum_Iirr_kerD_square ?sub1G ?normal1 // indexgg. + rewrite addr0 mul1r indexg1 oH opprD addrACA addNr addr0 addrC. + rewrite (expnSr p 2) -[p in (_ ^ 2 * p)%:R - _]prednK ?prime_gt0 // mulnSr. + rewrite natrD addrK eqxx => /andP[/eqP Dd /nilP S3nil]. + have uS12: uniq (S1 ++ S2). + by rewrite cat_uniq seqInd_uniq uS2 andbT; apply/hasPn. + rewrite uniq_perm_eq ?seqInd_uniq {uS12}// => [|xi]; last first. + apply/idP/idP; apply: allP xi; last by rewrite all_cat !(introT allP _). + by rewrite -(canLR negbK (has_predC _ _)) has_filter -/S3 S3nil. + have: (w1 %| d%:R - delta)%C. + by rewrite unfold_in pnatr_eq0 eqn0Ngt w1_gt0 rpred_Cnat. + rewrite /n Dd def_p_w1 /delta; case: (Idelta _) => [_|/idPn[] /=]. + by rewrite opprK -(natrD _ _ 1) subnK ?muln_gt0 // natrM mulfK ?neq0CG. + rewrite mul2n -addnn -{1}(subnKC (ltnW w1gt2)) !addSn mulrSr addrK dvdC_nat. + by rewrite add0n dvdn_addl // -(subnKC w1gt2) gtnNdvd // leqW. +have scohS1 := subset_subcoherent scohS0 sS10. +have o1S1: orthonormal S1. + rewrite orthonormalE andbC; have [_ _ -> _ _] := scohS1. + by apply/allP=> xi /irrS1/irrP[t ->]; rewrite /= cfnorm_irr. +have [tau1 cohS1]: coherent S1 M^# tau. + apply: uniform_degree_coherence scohS1 _; apply: all_pred1_constant w1%:R _ _. + by rewrite all_map; apply/allP=> xi /S1w1/= ->. +have [[Itau1 Ztau1] Dtau1] := cohS1. +have o1S1tau: orthonormal (map tau1 S1) by apply: map_orthonormal. +have S1zeta: zeta \in S1. + by have:= Szeta; rewrite (perm_eq_mem defS) mem_cat => /orP[//|/redS2/negP]. +(* This is the main part of step 10.10.3; as the definition of alpha_ remains *) +(* valid we do not need to reprove alpha_on. *) +have Dalpha i (al_ij := alpha_ i j) : + al_ij^\tau = delta *: (eta_ i j - eta_ i 0) - n *: tau1 zeta. +- have [Y S1_Y [X [Dal_ij _ oXY]]] := orthogonal_split (map tau1 S1) al_ij^\tau. + have [a_ Da_ defY] := orthonormal_span o1S1tau S1_Y. + have oXS1 lam : lam \in S1 -> '[X, tau1 lam] = 0. + by move=> S1lam; rewrite (orthoPl oXY) ?map_f. + have{Da_} Da_ lam : lam \in S1 -> a_ (tau1 lam) = '[al_ij^\tau, tau1 lam]. + by move=> S1lam; rewrite Dal_ij cfdotDl oXS1 // addr0 Da_. + pose a := n + a_ (tau1 zeta); have [_ oS1S1] := orthonormalP o1S1. + have Da_z: a_ (tau1 zeta) = - n + a by rewrite addKr. + have Za: a \in Cint. + rewrite rpredD ?Dn ?rpred_nat // Da_ // Cint_cfdot_vchar ?Zalpha_tau //=. + by rewrite Ztau1 ?mem_zchar. + have Da_z' lam: lam \in S1 -> lam != zeta -> a_ (tau1 lam) = a. + move=> S1lam zeta'lam; apply: canRL (subrK _) _. + rewrite !Da_ // -cfdotBr -raddfB. + have S1dlam: lam - zeta \in 'Z[S1, M^#]. + by rewrite zcharD1E rpredB ?mem_zchar //= !cfunE !S1w1 ?subrr. + rewrite Dtau1 // Dade_isometry ?alpha_on ?tauM' //; last first. + by rewrite -zcharD1_seqInd ?(zchar_subset sS1S). + have o_mu2_lam k: '[mu2_ i k, lam] = 0 by rewrite o_mu2_irr ?sS1S ?irrS1. + rewrite !cfdotBl !cfdotZl !cfdotBr !o_mu2_lam !o_mu2_zeta !(subr0, mulr0). + by rewrite irrWnorm ?oS1S1 // eq_sym (negPf zeta'lam) !add0r mulrN1 opprK. + have lb_n2alij: (a - n) ^+ 2 + (size S1 - 1)%:R * a ^+ 2 <= '[al_ij^\tau]. + rewrite Dal_ij cfnormDd; last first. + by rewrite cfdotC (span_orthogonal oXY) ?rmorph0 // memv_span1. + rewrite ler_paddr ?cfnorm_ge0 // defY cfnorm_sum_orthonormal //. + rewrite (big_rem (tau1 zeta)) ?map_f //= ler_eqVlt; apply/predU1P; left. + congr (_ + _). + by rewrite Da_z addrC Cint_normK 1?rpredD // rpredN Dn rpred_nat. + rewrite (eq_big_seq (fun _ => a ^+ 2)) => [|tau1lam]; last first. + rewrite rem_filter ?free_uniq ?orthonormal_free // filter_map. + case/mapP=> lam; rewrite mem_filter /= andbC => /andP[S1lam]. + rewrite (inj_in_eq (Zisometry_inj Itau1)) ?mem_zchar // => zeta'lam ->. + by rewrite Da_z' // Cint_normK. + rewrite big_tnth sumr_const card_ord size_rem ?map_f // size_map. + by rewrite mulr_natl subn1. + have{lb_n2alij} ub_a2: (size S1)%:R * a ^+ 2 <= 2%:R * a * n + 2%:R. + rewrite norm_alpha // addrC sqrrB !addrA ler_add2r in lb_n2alij. + rewrite mulr_natl -mulrSr ler_subl_addl subn1 in lb_n2alij. + by rewrite -mulrA !mulr_natl; case: (S1) => // in S1zeta lb_n2alij *. + have{ub_a2} ub_8a2: 8%:R * a ^+ 2 <= 4%:R * a + 2%:R. + rewrite mulrAC Dn -natrM in ub_a2; apply: ler_trans ub_a2. + rewrite -Cint_normK // ler_wpmul2r ?exprn_ge0 ?normr_ge0 // leC_nat szS1. + rewrite (subn_sqr p 1) def_p_w1 subnK ?muln_gt0 // mulnA mulnK // mulnC. + by rewrite -subnDA -(mulnBr 2 _ 1%N) mulnA (@leq_pmul2l 4 2) ?ltn_subRL. + have Z_4a1: 4%:R * a - 1%:R \in Cint by rewrite rpredB ?rpredM ?rpred_nat. + have{ub_8a2} ub_4a1: `|4%:R * a - 1| < 3%:R. + rewrite -ltr_sqr ?rpred_nat ?qualifE ?normr_ge0 // -natrX Cint_normK //. + rewrite sqrrB1 exprMn -natrX -mulrnAl -mulrnA (natrD _ 8 1) ltr_add2r. + rewrite (natrM _ 2 4) (natrM _ 2 8) -!mulrA -mulrBr ltr_pmul2l ?ltr0n //. + by rewrite ltr_subl_addl (ler_lt_trans ub_8a2) // ltr_add2l ltr_nat. + have{ub_4a1} a0: a = 0. + apply: contraTeq ub_4a1 => a_nz; have:= norm_Cint_ge1 Za a_nz. + rewrite real_ltr_norml ?real_ler_normr ?Creal_Cint //; apply: contraL. + case/andP; rewrite ltr_subl_addr -(natrD _ 3 1) gtr_pmulr ?ltr0n //. + rewrite ltr_oppl opprB -mulrN => /ltr_le_trans/=/(_ _ (leC_nat 3 5)). + by rewrite (natrD _ 1 4) ltr_add2l gtr_pmulr ?ltr0n //; do 2!move/ltr_geF->. + apply: (def_tau_alpha cohS1 sS10 nz_j S1zeta). + by rewrite -Da_ // Da_z a0 addr0. +have o_eta__zeta i j1: '[tau1 zeta, eta_ i j1] = 0. + by rewrite (coherent_ortho_cycTIiso _ sS10 cohS1) ?mem_irr. +(* This is step (10.4), the final one. *) +have Dmu0zeta: (mu_ 0 - zeta)^\tau = \sum_i eta_ i 0 - tau1 zeta. + have A0mu0tau: mu_ 0 - zeta \in 'CF(M, 'A0(M)). + rewrite /'A0(M) defA; apply: (cfun_onS (subsetUl _ _)). + rewrite cfun_onD1 [mu_ 0](prTIred0 pddM) !cfunE zeta1w1 cfuniE // group1. + by rewrite mulr1 subrr rpredB ?rpredZnat ?cfuni_on ?(seqInd_on _ Szeta) /=. + have [chi [Dmu0 Zchi n1chi o_chi_w]] := FTtype345_Dade_bridge0 w1_lt_w2. + have dirr_chi: chi \in dirr G by rewrite dirrE Zchi n1chi /=. + have dirr_zeta: tau1 zeta \in dirr G. + by rewrite dirrE Ztau1 ?Itau1 ?mem_zchar //= irrWnorm. + have: '[(alpha_ 0 j)^\tau, (mu_ 0 - zeta)^\tau] == - delta + n. + rewrite Dade_isometry ?alpha_on // !cfdotBl !cfdotZl !cfdotBr. + rewrite !o_mu2_zeta 2!cfdot_prTIirr_red (negPf nz_j) cfdotC o_mu_zeta. + by rewrite eqxx irrWnorm // conjC0 !(subr0, add0r) mulr1 mulrN1 opprK. + rewrite Dalpha // Dmu0 !{1}(cfdotBl, cfdotZl) !cfdotBr 2!{1}(cfdotC _ chi). + rewrite !o_chi_w conjC0 !cfdot_sumr big1 => [|i]; first last. + by rewrite (cfdot_cycTIiso pddM) (negPf nz_j) andbF. + rewrite (bigD1 0) //= cfdot_cycTIiso big1 => [|i nz_i]; first last. + by rewrite cfdot_cycTIiso eq_sym (negPf nz_i). + rewrite big1 // !subr0 !add0r addr0 mulrN1 mulrN opprK (can_eq (addKr _)). + rewrite {2}Dn -mulr_natl Dn (inj_eq (mulfI _)) ?pnatr_eq0 //. + by rewrite cfdot_dirr_eq1 // => /eqP->. +have [] := uniform_prTIred_coherent pddM nz_j; rewrite -/sigma. +have ->: uniform_prTIred_seq pddM j = S2. + congr (map _ _); apply: eq_enum => k; rewrite !inE -!/(mu_ _). + by rewrite andb_idr // => nz_k; rewrite 2!{1}prTIred_1 2?Dmu2_1. +case=> _ _ ccS2 _ _ [tau2 Dtau2 cohS2]. +have{cohS2} cohS2: coherent_with S2 M^# tau tau2 by apply: cohS2. +have sS20: cfConjC_subset S2 calS0. + by split=> // xi /sS2S Sxi; have [_ ->] := sSS0. +rewrite perm_eq_sym perm_catC in defS; apply: perm_eq_coherent defS _. +suffices: (mu_ j - d%:R *: zeta)^\tau = tau2 (mu_ j) - tau1 (d%:R *: zeta). + apply: (bridge_coherent scohS0 sS20 cohS2 sS10 cohS1) => [phi|]. + by apply: contraL => /S1'2. + rewrite cfunD1E !cfunE zeta1w1 prTIred_1 mulrC Dmu2_1 // subrr. + by rewrite image_f // rpredZnat ?mem_zchar. +have sumA: \sum_i alpha_ i j = mu_ j - delta *: mu_ 0 - (d%:R - delta) *: zeta. + rewrite !sumrB sumr_const /= -scaler_sumr; congr (_ - _ - _). + rewrite card_Iirr_abelian ?cyclic_abelian // -/w1 -scaler_nat. + by rewrite scalerA mulrC divfK ?neq0CG. +rewrite scalerBl opprD opprK addrACA in sumA. +rewrite -{sumA}(canLR (addrK _) sumA) opprD opprK -scalerBr. +rewrite linearD linearZ linear_sum /= Dmu0zeta scalerBr. +rewrite (eq_bigr _ (fun i _ => Dalpha i)) sumrB sumr_const nirrW1. +rewrite -!scaler_sumr sumrB addrAC !addrA scalerBr subrK -addrA -opprD. +rewrite raddfZnat Dtau2 Ddelta_ //; congr (_ - _). +by rewrite addrC -scaler_nat scalerA mulrC divfK ?neq0CG // -scalerDl subrK. +Qed. + +End OneMaximal. + +Implicit Type M : {group gT}. + +(* This is the exported version of Peterfalvi, Theorem (10.8). *) +Theorem FTtype345_noncoherence M (M' := M^`(1)%G) (maxM : M \in 'M) : + (FTtype M > 2)%N -> ~ coherent (seqIndD M' M M' 1) M^# (FT_Dade0 maxM). +Proof. +rewrite ltnNge 2!leq_eqVlt => /norP[notMtype2 /norP[notMtype1 _]] [tau1 cohS]. +have [U W W1 W2 defW MtypeP] := FTtypeP_witness maxM notMtype1. +have [zeta [irr_zeta Szeta zeta1w1]] := FTtypeP_ref_irr maxM MtypeP. +exact: (FTtype345_noncoherence_main MtypeP _ irr_zeta Szeta zeta1w1 cohS). +Qed. + +(* This is the exported version of Peterfalvi, Theorem (10.10). *) +Theorem FTtype5_exclusion M : M \in 'M -> FTtype M != 5. +Proof. +move=> maxM; apply: wlog_neg; rewrite negbK => Mtype5. +have notMtype2: FTtype M != 2 by rewrite (eqP Mtype5). +have [U W W1 W2 defW [[MtypeP _] _]] := FTtypeP 5 maxM Mtype5. +have [zeta [irr_zeta Szeta zeta1w1]] := FTtypeP_ref_irr maxM MtypeP. +exact: (FTtype5_exclusion_main _ MtypeP _ irr_zeta). +Qed. + +(* This the first assertion of Peterfalvi (10.11). *) +Lemma FTtypeP_pair_primes S T W W1 W2 (defW : W1 \x W2 = W) : + typeP_pair S T defW -> prime #|W1| /\ prime #|W2|. +Proof. +move=> pairST; have [[_ maxS maxT] _ _ _ _] := pairST. +have type24 maxM := compl_of_typeII_IV maxM _ (FTtype5_exclusion maxM). +split; first by have [U /type24[]] := typeP_pairW pairST. +have xdefW: W2 \x W1 = W by rewrite dprodC. +by have [U /type24[]] := typeP_pairW (typeP_pair_sym xdefW pairST). +Qed. + +Corollary FTtypeP_primes M U W W1 W2 (defW : W1 \x W2 = W) : + M \in 'M -> of_typeP M U defW -> prime #|W1| /\ prime #|W2|. +Proof. +move=> maxM MtypeP; have [T pairMT _] := FTtypeP_pair_witness maxM MtypeP. +exact: FTtypeP_pair_primes pairMT. +Qed. + +(* This is the remainder of Peterfalvi (10.11). *) +Lemma FTtypeII_prime_facts M U W W1 W2 (defW : W1 \x W2 = W) (maxM : M \in 'M) : + of_typeP M U defW -> FTtype M == 2 -> + let H := M`_\F%G in let HU := M^`(1)%G in + let calS := seqIndD HU M H 1 in let tau := FT_Dade0 maxM in + let p := #|W2| in let q := #|W1| in + [/\ p.-abelem H, (#|H| = p ^ q)%N & coherent calS M^# tau]. +Proof. +move=> MtypeP Mtype2 H HU calS tau p q. +have Mnot5: FTtype M != 5 by rewrite (eqP Mtype2). +have [_ cUU _ _ _] := compl_of_typeII maxM MtypeP Mtype2. +have [q_pr p_pr]: prime q /\ prime p := FTtypeP_primes maxM MtypeP. +have:= typeII_IV_core maxM MtypeP Mnot5; rewrite Mtype2 -/p -/q => [[_ oH]]. +have [] := Ptype_Fcore_kernel_exists maxM MtypeP Mnot5. +have [_ _] := Ptype_Fcore_factor_facts maxM MtypeP Mnot5. +rewrite -/H; set H0 := Ptype_Fcore_kernel _; set Hbar := (H / H0)%G. +rewrite def_Ptype_factor_prime // -/p -/q => oHbar chiefHbar _. +have trivH0: H0 :=: 1%g. + have [/maxgroupp/andP[/andP[sH0H _] nH0M] /andP[sHM _]] := andP chiefHbar. + apply: card1_trivg; rewrite -(setIidPr sH0H) -divg_index. + by rewrite -card_quotient ?(subset_trans sHM) // oHbar -oH divnn cardG_gt0. +have abelHbar: p.-abelem Hbar. + have pHbar: p.-group Hbar by rewrite /pgroup oHbar pnat_exp pnat_id. + by rewrite -is_abelem_pgroup // (sol_chief_abelem _ chiefHbar) ?mmax_sol. +rewrite /= trivH0 -(isog_abelem (quotient1_isog _)) in abelHbar. +have:= Ptype_core_coherence maxM MtypeP Mnot5; rewrite trivH0. +set C := _ MtypeP; have sCU: C \subset U by rewrite [C]unlock subsetIl. +by rewrite (derG1P (abelianS sCU cUU)) [(1 <*> 1)%G]join1G. +Qed. + +End Ten. diff --git a/mathcomp/odd_order/PFsection11.v b/mathcomp/odd_order/PFsection11.v new file mode 100644 index 0000000..3584dbe --- /dev/null +++ b/mathcomp/odd_order/PFsection11.v @@ -0,0 +1,1193 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg poly finset center. +Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +Require Import gfunctor gproduct cyclic commutator gseries nilpotent pgroup. +Require Import sylow hall abelian maximal frobenius. +Require Import matrix mxalgebra mxrepresentation mxabelem vector. +Require Import BGsection1 BGsection3 BGsection7 BGsection15 BGsection16. +Require Import ssrnum ssrint algC classfun character inertia vcharacter. +Require Import PFsection1 PFsection2 PFsection3 PFsection4 PFsection5. +Require Import PFsection6 PFsection7 PFsection8 PFsection9 PFsection10. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 11: Maximal subgroups of Types *) +(* III and IV. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory FinRing.Theory Num.Theory. + +Section Eleven. + +(* This is Peterfalvi (11.1). *) +Lemma lbound_expn_odd_prime p q : + odd p -> odd q -> prime p -> prime q -> p != q -> 4 * q ^ 2 + 1 < p ^ q. +Proof. +move=> odd_p odd_q pr_p pr_q p_neq_q. +have{pr_p pr_q} [pgt2 qgt2] : 2 < p /\ 2 < q by rewrite !odd_prime_gt2. +have [qlt5 | qge5 {odd_q qgt2 p_neq_q}] := ltnP q 5. + have /eqP q3: q == 3 by rewrite eqn_leq qgt2 andbT -ltnS -(odd_ltn 5). + apply: leq_trans (_ : 5 ^ q <= p ^ q); first by rewrite q3. + by rewrite leq_exp2r // odd_geq // ltn_neqAle pgt2 eq_sym -q3 p_neq_q. +apply: leq_trans (_ : 3 ^ q <= p ^ q); last by rewrite -(subnKC qge5) leq_exp2r. +elim: q qge5 => // q IHq; rewrite ltnS leq_eqVlt => /predU1P[<- // | qge5]. +rewrite (expnS 3); apply: leq_trans {IHq}(leq_mul (leqnn 3) (IHq qge5)). +rewrite -!addnS mulnDr leq_add // mulnCA leq_mul // !(mul1n, mulSnr). +by rewrite -addn1 sqrnD muln1 -(subnKC qge5) !leq_add ?leq_mul. +Qed. + +Local Open Scope ring_scope. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types H K L N P Q R S T U V W : {group gT}. + +Variables M U W W1 W2 : {group gT}. +Hypotheses (maxM : M \in 'M) (defW : W1 \x W2 = W) (MtypeP : of_typeP M U defW). +Hypothesis notMtype2 : FTtype M != 2. + +Let notMtype5 : FTtype M != 5. Proof. exact: FTtype5_exclusion. Qed. +Let notMtype1 : FTtype M != 1%N. Proof. exact: FTtypeP_neq1 MtypeP. Qed. +Let Mtype34 : FTtype M \in pred2 3 4. +Proof. +by have:= FTtype_range M; rewrite -mem_iota !inE !orbA orbC 3?[_ == _](negPf _). +Qed. +Let Mtype_gt2 : (FTtype M > 2)%N. Proof. by case/pred2P: Mtype34 => ->. Qed. + +Local Notation H0 := (Ptype_Fcore_kernel MtypeP). +Local Notation "` 'H0'" := (gval H0) (at level 0, only parsing) : group_scope. +Local Notation "` 'M'" := (gval M) (at level 0, only parsing) : group_scope. +Local Notation "` 'U'" := (gval U) (at level 0, only parsing) : group_scope. +Local Notation "` 'W'" := (gval W) (at level 0, only parsing) : group_scope. +Local Notation "` 'W1'" := (gval W1) (at level 0, only parsing) : group_scope. +Local Notation "` 'W2'" := (gval W2) (at level 0, only parsing) : group_scope. +Local Notation H := `M`_\F%G. +Local Notation "` 'H'" := `M`_\F (at level 0) : group_scope. +Local Notation HU := M^`(1)%G. +Local Notation "` 'HU'" := `M^`(1)%g (at level 0) : group_scope. +Local Notation U' := U^`(1)%G. +Local Notation "` 'U''" := `U^`(1)%g (at level 0) : group_scope. +Local Notation C := 'C_U(`H)%G. +Local Notation "` 'C'" := 'C_`U(`H) (at level 0) : group_scope. +Local Notation HC := (`H <*> `C)%G. +Local Notation "` 'HC'" := (`H <*> `C) (at level 0) : group_scope. +Local Notation H0C := (`H0 <*> `C)%G. +Local Notation "` 'H0C'" := (`H0 <*> `C) (at level 0) : group_scope. +Local Notation Hbar := (`H / `H0)%g. + +Local Notation S_ := (seqIndD HU M HU). +Local Notation tau := (FT_Dade0 maxM). +Local Notation R := (FTtypeP_coh_base maxM MtypeP). +Local Notation V := (cyclicTIset defW). + +Let Mtype24 := compl_of_typeII_IV maxM MtypeP notMtype5. + +Let defMs : M`_\s = HU. Proof. exact: FTcore_type_gt2. Qed. +Let defA1 : 'A1(M) = HU^#. Proof. by rewrite /= -defMs. Qed. +Let defA : 'A(M) = HU^#. Proof. by rewrite FTsupp_eq1. Qed. +Let sHU_A0 : HU^# \subset 'A0(M). Proof. by rewrite -defA FTsupp_sub0. Qed. + +Let calS := seqIndD HU M M`_\s 1. +Let scohM : subcoherent calS tau R. Proof. exact: FTtypeP_subcoherent. Qed. +Let scoh1 : subcoherent (S_ 1) tau R. +Proof. by rewrite -{2}(group_inj defMs). Qed. + +Let p := #|W2|. +Let pr_p : prime p. Proof. by have [] := FTtypeP_primes maxM MtypeP. Qed. +Let ntW2 : W2 :!=: 1%g. Proof. by rewrite -cardG_gt1 prime_gt1. Qed. +Let cycW2 : cyclic W2. Proof. exact: prime_cyclic. Qed. +Let def_p : pdiv #|Hbar| = p. Proof. exact: typeIII_IV_core_prime. Qed. + +Let q := #|W1|. +Let pr_q : prime q. Proof. by have [] := FTtypeP_primes maxM MtypeP. Qed. +Let ntW1 : W1 :!=: 1%g. Proof. by rewrite -cardG_gt1 prime_gt1. Qed. +Let cycW1 : cyclic W1. Proof. exact: prime_cyclic. Qed. + +Let defM : HU ><| W1 = M. Proof. by have [[]] := MtypeP. Qed. +Let defHU : H ><| U = HU. Proof. by have [_ []] := MtypeP. Qed. + +Let nsHUM : HU <| M. Proof. exact: gFnormal. Qed. +Let sHUM : HU \subset M. Proof. exact: gFsub. Qed. +Let sHHU : H \subset HU. Proof. by have /mulG_sub[] := sdprodW defHU. Qed. +Let sUHU : U \subset HU. Proof. by have /mulG_sub[] := sdprodW defHU. Qed. +Let sUM : U \subset M. Proof. exact: subset_trans sUHU sHUM. Qed. + +Let coHUq : coprime #|HU| q. +Proof. by rewrite (coprime_sdprod_Hall_r defM); have [[]] := MtypeP. Qed. +Let coUq : coprime #|U| q. Proof. exact: coprimeSg coHUq. Qed. + +Let neq_pq : p != q. +Proof. +apply: contraTneq coHUq => <-; rewrite coprime_sym prime_coprime ?cardSg //. +by rewrite -(typeP_cent_compl MtypeP) subsetIl. +Qed. + +Let solHU : solvable HU. Proof. exact: solvableS sHUM (mmax_sol maxM). Qed. +Let solH : solvable H. Proof. exact: solvableS sHHU solHU. Qed. + +Let ltM''HU : M^`(2)%g \proper HU. +Proof. by rewrite (sol_der1_proper solHU) // -defMs FTcore_neq1. Qed. + +Let frobMtilde : [Frobenius M / M^`(2) = (HU / M^`(2)) ><| (W1 / M^`(2))]. +Proof. +have [[_ _ _ _] _ _ [_ _ _ sW2M'' prHUW1 ] _] := MtypeP. +by rewrite Frobenius_coprime_quotient ?gFnormal //; split=> // _ /prHUW1->. +Qed. + +Let defHC : H \x C = HC. +Proof. +by have [defHC _ _ _] := typeP_context MtypeP; rewrite /= (dprodWY defHC). +Qed. + +Let nC_UW1 : U <*> W1 \subset 'N(C). +Proof. +have /sdprodP[_ _ nHUW1 _] := Ptype_Fcore_sdprod MtypeP. +by rewrite normsI ?norms_cent // join_subG normG; have [_ []] := MtypeP. +Qed. + +Let nsCM : C <| M. +Proof. +rewrite /normal subIset ?sUM //= -{1}(sdprodW (Ptype_Fcore_sdprod MtypeP)). +by rewrite mulG_subG cents_norm // centsC subsetIr. +Qed. + +Let nsCU : C <| U. Proof. exact: normalS (subsetIl _ _) sUM nsCM. Qed. +Let nsHC_M : HC <| M. Proof. by rewrite normalY ?gFnormal. Qed. +Let sHC_HU : HC \subset HU. Proof. by rewrite join_subG sHHU subIset ?sUHU. Qed. +Let nsHC_HU : HC <| HU. Proof. exact: normalS nsHC_M. Qed. + +Let chiefH0 : chief_factor M H0 H. +Proof. by have [] := Ptype_Fcore_kernel_exists maxM MtypeP notMtype5. Qed. + +Let minHbar : minnormal Hbar (M / H0). +Proof. exact: chief_factor_minnormal. Qed. + +Let abelHbar : p.-abelem Hbar. +Proof. +have solHbar : solvable (H / H0) by rewrite quotient_sol. +have [_ _] := minnormal_solvable minHbar (subxx _) solHbar. +by rewrite /is_abelem def_Ptype_factor_prime. +Qed. + +Let sH0H : H0 \subset H. +Proof. by have/andP[/maxgroupp/andP[/proper_sub]]:= chiefH0. Qed. + +Let nH0M: M \subset 'N(H0). +Proof. by have /andP[/maxgroupp/andP[]] := chiefH0. Qed. + +Let nsH0H : H0 <| H. +Proof. by rewrite /normal sH0H (subset_trans (Fcore_sub _)). Qed. + +Let nsH0C_M : H0C <| M. +Proof. by rewrite !normalY ?gFnormal /normal ?(subset_trans sH0H) ?gFsub. Qed. + +Let defH0C : H0 \x C = H0C. +Proof. +have /dprodP[_ _ cHC tiHC] := defHC. +by rewrite dprodEY ?(centsS sH0H) //; apply/trivgP; rewrite -tiHC setSI. +Qed. + +(* Group-theoretic consequences of the coherence and non-coherence theorems *) +(* of Sections 5, 9 and 10 for maximal subgroups of type III and IV. *) + +(* This is Peterfalvi (11.3). *) +Lemma FTtype34_noncoherence : ~ coherent (S_ H0C) M^# tau. +Proof. +move=> cohH0C; suff: coherent (S_ 1) M^# tau by apply: FTtype345_noncoherence. +have /mulG_sub[_ sW1M] := sdprodW defM. +have [nsHHU _ _ nHU tiHU] := sdprod_context defHU. +have sHM: H \subset M := gFsub _ _. +have [sCM sH0C_M]: C \subset M /\ H0C \subset M by rewrite !normal_sub. +have nH0_C := subset_trans sCM nH0M. +have sH0C_HC: H0C \subset HC by apply: genS (setSU _ _). +have defF: HC :=: 'F(M) by have [/dprodWY] := typeP_context MtypeP. +have{defF} nilHC: nilpotent (HC / 1) by rewrite defF quotient_nil ?Fitting_nil. +have /bounded_seqIndD_coherent-bounded_coh1 := scoh1. +apply: bounded_coh1 nilHC cohH0C _; rewrite ?sub1G ?normal1 //. +have[_ _ /= oHbar] := Ptype_Fcore_factor_facts maxM MtypeP notMtype5. +rewrite -(index_sdprod defM) -divgS // -(dprod_card defHC) -(dprod_card defH0C). +rewrite divnMr ?cardG_gt0 // divg_normal // oHbar def_p -/q. +by rewrite lbound_expn_odd_prime ?mFT_odd. +Qed. + +(* This is Peterfalvi (11.4). *) +Lemma bounded_proper_coherent H1 : + H1 <| M -> H1 \proper HU -> coherent (S_ H1) M^# tau -> + (#|HU : H1| <= 2 * q * #|U : C| + 1)%N. +Proof. +move=> nsH1_M psH1_M' cohH1; have [nsHHU _ _ _ _] := sdprod_context defHU. +rewrite -leC_nat natrD -ler_subl_addr. +have ->: (2 * q * #|U : C|)%:R = 2%:R * #|M : HC|%:R * sqrtC #|HC : HC|%:R. + rewrite indexgg sqrtC1 mulr1 -mulnA natrM; congr (_ * _%:R). + apply/eqP; rewrite // -(eqn_pmul2l (cardG_gt0 HC)) Lagrange ?normal_sub //. + rewrite mulnCA -(dprod_card defHC) -mulnA (Lagrange (subsetIl _ _)). + by rewrite -(sdprod_card defM) -(sdprod_card defHU) mulnC. +have ns_M: [/\ H1 <| M, H0C <| M, HC <| M & HC <| M] by []. +case: (coherent_seqIndD_bound _ _ scoh1 ns_M) FTtype34_noncoherence => //. +suffices /center_idP->: abelian (HC / H0C) by rewrite genS ?setSU. +suffices /isog_abelian->: HC / H0C \isog H / H0 by apply: abelem_abelian p _ _. +by rewrite /= (joingC H0) isog_sym quotient_sdprodr_isog ?(dprodWsdC defHC). +Qed. + +(* This is Peterfalvi (11.5). *) +Lemma FTtype34_der2 : M^`(2)%g = HC. +Proof. +have [defFM [_ not_cHU] _ _] := typeP_context MtypeP. +have [_ sW1M _ _ tiHU_W1] := sdprod_context defM. +have{defFM} sM''_HC: M^`(2)%g \subset HC. + by rewrite -defHC defFM; have [_ _ []] := MtypeP. +have scohM'': subcoherent (S_ M^`(2)) tau R. + exact/(subset_subcoherent scoh1)/seqInd_conjC_subset1. +have cohM'': coherent (S_ M^`(2)) M^# tau. + apply: uniform_degree_coherence scohM'' _. + apply: all_pred1_constant #|M : HU|%:R _ _; rewrite all_map. + apply/allP=> _ /seqIndP[s /setDP[kerM'' _] ->] /=; rewrite inE in kerM''. + by rewrite cfInd1 ?gFsub // lin_char1 ?mulr1 ?lin_irr_der1. +have ubHC: (#|HC : M^`(2)| < 2 * q + 1)%N. + rewrite -(ltn_pmul2r (indexg_gt0 U C)) mulnDl mul1n. + apply: leq_ltn_trans (_ : 2 * q * #|U : C| + 1 < _)%N; last first. + by rewrite ltn_add2l indexg_gt1 subsetIidl not_cHU //; have [] := Mtype24. + have {1}->: #|U : C| = #|HU : HC| by apply: index_sdprodr (subsetIl _ _). + by rewrite mulnC (Lagrange_index sHC_HU) // bounded_proper_coherent ?gFnormal. +have regHC_W1: semiregular (HC / M^`(2)) (W1 / M^`(2)). + by apply: semiregularS (Frobenius_reg_ker frobMtilde); rewrite quotientS. +suffices /dvdnP[k Dk]: 2 * q %| #|HC : M^`(2)|.-1. + apply: contraTeq ubHC; rewrite -leqNgt eqEsubset sM''_HC -indexg_gt1 addn1. + by rewrite -[#|_:_|]prednK // {}Dk !ltnS muln_gt0 => /andP[/leq_pmull->]. +rewrite Gauss_dvd; last by rewrite coprime2n mFT_odd. +rewrite dvdn2 -subn1 odd_sub // addbT negbK subn1. +rewrite -card_quotient; last by rewrite (subset_trans sHC_HU) // (der_norm 1). +have Dq: q = #|W1 / M^`(2)|%g. + apply/card_isog/quotient_isog; first by rewrite (subset_trans sW1M) ?gFnorm. + by apply/trivgP; rewrite -tiHU_W1 setSI // (der_sub 1). +rewrite quotient_odd ?mFT_odd //= Dq regular_norm_dvd_pred ?quotient_norms //. +by rewrite (subset_trans sW1M) ?normal_norm. +Qed. +Local Notation defM'' := FTtype34_der2. + +(* This is Peterfalvi (11.6). *) +Lemma FTtype34_facts (H' := H^`(1)%g): + [/\ p.-group H, U \subset 'C(H0), H0 :=: H' & C :=: U']. +Proof. +have nilH: nilpotent H := Fcore_nil M. +have /sdprod_context[/andP[_ nHM] sUW1M _ _ _] := Ptype_Fcore_sdprod MtypeP. +have coH_UW1: coprime #|H| #|U <*> W1| := Ptype_Fcore_coprime MtypeP. +have [[_ mulHU _ tiHU] [nHU isomHU]] := (sdprodP defHU, sdprod_isom defHU). +have{sUW1M} cH0U: U \subset 'C(H0). + have frobUW1 := Ptype_compl_Frobenius maxM MtypeP notMtype5. + have nH0_UW1 := subset_trans sUW1M nH0M; have [_ nH0W1] := joing_subP nH0_UW1. + have [coH0_UW1 solH0] := (coprimeSg sH0H coH_UW1, solvableS sH0H solH). + have [_ -> //] := Frobenius_Wielandt_fixpoint frobUW1 nH0_UW1 coH0_UW1 solH0. + have ->: 'C_H0(W1) = H0 :&: 'C_H(W1) by rewrite setIA (setIidPl sH0H). + have nH0C: 'C_H(W1) \subset 'N(H0) by rewrite subIset // normal_norm. + rewrite cardMg_TI // -LagrangeMl -card_quotient {nH0C}//. + rewrite coprime_quotient_cent ?(coprimeSg sHHU) //=. + have [_ -> _] := Ptype_Fcore_factor_facts maxM MtypeP notMtype5. + by rewrite (typeP_cent_core_compl MtypeP) def_p. +have{isomHU} defC: C :=: U'. + have [injHquo defHUb] := isomP isomHU. + apply: (injm_morphim_inj injHquo); rewrite ?subsetIl ?morphim_der ?der_sub //. + rewrite defHUb morphim_restrm -quotientE setIA setIid -quotientMidl /=. + by rewrite (dprodW defHC) -defM'' -quotient_der // -mulHU mul_subG ?normG. +have{coH_UW1} defH0: H0 :=: H'. + pose Hhat := (H / H')%g; pose Uhat := (U / H')%g; pose HUhat := (HU / H')%g. + have nH'H: H \subset 'N(H') := gFnorm _ _. + have nH'U: U \subset 'N(H') := char_norm_trans (der_char _ _) nHU. + apply/eqP; rewrite eqEsubset andbC. + rewrite der1_min ?(abelem_abelian abelHbar) ?normal_norm //=. + rewrite -quotient_sub1 /= -/H'; last exact: subset_trans sH0H nH'H. + suffices <-: 'C_Hhat(Uhat) = 1%g. + by rewrite subsetI quotientS //= quotient_cents // centsC. + suffices: ~~ ('C_Hhat(Uhat)^`(1)%g \proper 'C_Hhat(Uhat)). + exact: contraNeq (sol_der1_proper (quotient_sol _ solH) (subsetIl Hhat _)). + have {2}<-: HUhat^`(1)%g :&: 'C_Hhat(Uhat) = 'C_Hhat(Uhat). + rewrite -quotient_der ?[HU^`(1)%g]defM''; last by rewrite -mulHU mul_subG. + by rewrite (setIidPr _) ?subIset // quotientS ?joing_subl. + suffices defHUhat: 'C_Hhat(Uhat) \x ([~: Hhat, Uhat] <*> Uhat) = HUhat. + rewrite -(dprod_modl (der_dprod 1 defHUhat)) ?der_sub //= -/Hhat. + rewrite [rhs in _ \x rhs](trivgP _) ?dprodg1 ?properxx //= -/Hhat. + by have [_ _ _ <-] := dprodP defHUhat; rewrite setIC setIS ?der_sub. + have coHUhat: coprime #|Hhat| #|Uhat|. + by rewrite coprime_morph ?(coprimegS _ coH_UW1) ?joing_subl. + have defHhat: 'C_Hhat(Uhat) \x [~: Hhat, Uhat] = Hhat. + by rewrite dprodC coprime_abelian_cent_dprod ?der_abelian ?quotient_norms. + rewrite /HUhat -(sdprodWY defHU) quotientY //= -(dprodWY defHhat). + have [_ _ cCRhat tiCRhat] := dprodP defHhat. + rewrite dprodEY ?joingA //; first by rewrite join_subG cCRhat centsC subsetIr. + apply/trivgP; rewrite /= joingC norm_joinEl ?commg_normr //= -/Hhat -/Uhat. + rewrite -tiCRhat !(setIAC _ 'C(_)) setSI // subsetI subsetIl /=. + by rewrite -group_modr ?commg_subl ?quotient_norms //= coprime_TIg ?mul1g. +suffices{defC defH0}: p.-group H by []. +pose R := 'O_p^'(H); have hallR: p^'.-Hall(H) R := nilpotent_pcore_Hall _ nilH. +have defRHp: R \x 'O_p(H) = H by rewrite dprodC nilpotent_pcoreC. +suffices R_1: R :=: 1%g by rewrite -defRHp R_1 dprod1g pcore_pgroup. +have /subsetIP[sRH cUR]: R \subset 'C_H(U). + have oH: #|H| = (p ^ q * #|'C_H(U)|)%N. + by have:= typeII_IV_core maxM MtypeP notMtype5 => /=; rewrite ifN => // -[]. + apply/setIidPl/eqP; rewrite eqEcard subsetIl /= (card_Hall hallR) {}oH. + rewrite (card_Hall (setI_normal_Hall _ hallR _)) ?subsetIl ?gFnormal //. + rewrite partnM ?expn_gt0 ?cardG_gt0 //= part_p'nat ?mul1n ?pnatNK //. + by rewrite pnat_exp ?pnat_id. +suffices: ~~ (R^`(1)%g \proper R) by apply: contraNeq (sol_der1_proper solH _). +have /setIidPr {2}<-: R \subset HU^`(1)%g. + by rewrite [HU^`(1)%g]defM'' -(dprodWY defHC) sub_gen ?subsetU ?sRH. +suffices defRHpU: R \x ('O_p(H) <*> U) = HU. + rewrite -(dprod_modl (der_dprod 1 defRHpU)) ?der_sub //= -/R setIC. + rewrite [rhs in _ \x rhs](trivgP _) ?dprodg1 ?properxx //= -/R. + by have /dprodP[_ _ _ <-] := defRHpU; rewrite setIS ?der_sub. +rewrite -(sdprodWY defHU) -[in RHS](dprodWY defRHp) -joingA. +have [_ _ cRHp tiRHp] := dprodP defRHp. +rewrite dprodEY //= -/R; first by rewrite join_subG cRHp centsC. +rewrite joingC (norm_joinEl (char_norm_trans (pcore_char p H) nHU)). +by rewrite -(setIidPl sRH) -setIA -group_modr ?gFsub // tiHU mul1g. +Qed. + +Let frobUW1bar : [Frobenius U <*> W1 / C = (U / C) ><| (W1 / C)]. +Proof. +have frobUW1: [Frobenius U <*> W1 = U ><| W1]. + exact: Ptype_compl_Frobenius MtypeP notMtype5. +have [defUW1 ntU _ _ _] := Frobenius_context frobUW1. +have [[_ _ _ defC] regUW1] := (FTtype34_facts, Frobenius_reg_ker frobUW1). +rewrite Frobenius_coprime_quotient // /normal ?subIset ?joing_subl //. +by split=> [|x /regUW1->]; rewrite ?sub1G //= defC (sol_der1_proper solHU). +Qed. + +(* This is Peterfalvi (11.7). *) +(* We have recast the linear algebra arguments in the original text in pure- *) +(* group-theoretic terms: the overhead of the abelem_rV correspondence is not *) +(* justifiable here, as the Ssreflect linear algebra library lacks specific *) +(* support for bilinear forms: we use D y z := [~ coset Q y, coset Q z] as *) +(* our "linear form". D is antisymmetric as D z y = (D y z)^-1, so we only *) +(* show that D is "linear" in z, that is, that D y is a group morphism with *) +(* domain H whose kernel contains H0, when y \in H, and we do not bother to *) +(* factor D to obtain a form over Hbar = H / H0. *) +(* We further rework the argument to support this change in perspective: *) +(* - We remove any reference to linear algebra in the "Galois" (9.7b) case, *) +(* where U acts irreducibly on Hbar: we revert to the proof of the *) +(* original Odd Order paper, using the fact that H / Q is extraspecial. *) +(* - In the "non-Galois" (9.7a) case, we use the W1-conjugation class of a *) +(* generator of H1 as an explicit basis of Hbar, indexed by W1, and we *) +(* use the elements xbar_ w = coset H0 (x_ w) of this basis instead of *) +(* arbitrary y in H_i, as the same argument then justifies extending *) +(* commutativity to all of Hbar. *) +(* - We construct phi as the morphism mapping ubar in Ubar to the n such *) +(* that the action of ubar on H1 is exponentiation by n. We derive a *) +(* morphism phi_ w ubar for the action of Ubar on H1 ^ w, for w in W1, by *) +(* composign with the action QV of W1 on Ubar by inverse conjugation. *) +(* - We exchange the two alternatives in the (9.7a) case; most of proof is *) +(* thus by contradiction with the C_U(Hbar) != u assertion in (9.6), *) +(* first establishing case 9.7a (as 9.7b contradicts q odd), then that D *) +(* is nontrivial for some x_ w1 and x_ w2 (as (H / Q)' = H0 / Q != 1), *) +(* whence (phi_ w1 u)(phi_ w2 u) = 1, whence (phi u)^-1 = phi u and *) +(* phi = 1, i.e., Ubar centralises Wbar. *) +(* Note finally that we simply construct U as a maximal subgroup of H0 normal *) +(* in H, as the nilpotence of H / Q implies that H0 / Q lies in its center. *) +Lemma FTtype34_Fcore_kernel_trivial : + [/\ p.-abelem H, #|H| = (p ^ q)%N & `H0 = 1%g]. +Proof. +have [[_ _ nHU tiHU] [pH cH0U defH' _]] := (sdprodP defHU, FTtype34_facts). +have [/mulG_sub[_ sW1M] nH0H] := (sdprodW defM, normal_norm nsH0H). +have nHW1: W1 \subset 'N(H) := subset_trans sW1M (gFnorm _ M). +have nUW1: W1 \subset 'N(U) by have [_ []] := MtypeP. +pose bar := coset_morphism H0; pose Hbar := (H / H0)%g; pose Ubar := (U / H0)%g. +have [Cbar_neqU _ /= oHbar] := Ptype_Fcore_factor_facts maxM MtypeP notMtype5. +rewrite -/Hbar def_p // -/q in oHbar. +have [nH0U nH0W1] := (subset_trans sUM nH0M, subset_trans sW1M nH0M). +suffices H0_1 : `H0 = 1%g. + have isoHbar: H \isog H / H0 by rewrite H0_1 quotient1_isog. + by rewrite (isog_abelem isoHbar) (card_isog isoHbar). +apply: contraNeq Cbar_neqU => ntH0; rewrite [Ptype_Fcompl_kernel _]unlock. +suffices: Hbar \subset 'C(Ubar). + by rewrite (sameP eqP setIidPl) sub_astabQ nH0U centsC. +have pH0 := pgroupS sH0H pH; have{ntH0} [_ _ [k oH0]] := pgroup_pdiv pH0 ntH0. +have{k oH0} [Q maxQ nsQH]: exists2 Q : {group gT}, maximal Q H0 & Q <| H. + have [Q [sQH0 nsQH oQ]] := normal_pgroup pH nsH0H (leq_pred _). + exists Q => //; apply: p_index_maximal => //. + by rewrite -divgS // oQ oH0 pfactorK //= expnS mulnK ?expn_gt0 ?cardG_gt0. +have nsQH0: Q <| H0 := p_maximal_normal (pgroupS sH0H pH) maxQ. +have [[sQH0 nQH0] nQH] := (andP nsQH0, normal_norm nsQH). +have nQU: U \subset 'N(Q) by rewrite cents_norm ?(centsS sQH0). +pose hat := coset_morphism Q; pose Hhat := (H / Q)%g; pose H0hat := (H0 / Q)%g. +have{maxQ} oH0hat: #|H0hat| = p by rewrite card_quotient ?(p_maximal_index pH0). +have defHhat': Hhat^`(1)%g = H0hat by rewrite -quotient_der -?defH'. +have ntH0hat: H0hat != 1%g by rewrite -cardG_gt1 oH0hat prime_gt1. +have pHhat: p.-group Hhat by apply: quotient_pgroup. +have nsH0Hhat: H0hat <| Hhat by apply: quotient_normal. +have sH0hatZ: H0hat \subset 'Z(Hhat). + by rewrite prime_meetG ?oH0hat // meet_center_nil ?(pgroup_nil pHhat). +have{pHhat} gal'M: ~~ typeP_Galois MtypeP. + have sZHhat: 'Z(Hhat) \subset Hhat := center_sub _. + have nsH0hatZ: H0hat <| 'Z(Hhat) := normalS sH0hatZ sZHhat nsH0Hhat. + have [f injf im_f] := third_isom sQH0 nsQH nsH0H. + have fHhat: f @* (Hhat / H0hat) = Hbar by rewrite im_f. + apply: contra (odd (logn p #|Hhat|)) _ _; last first. + rewrite -(divnK (cardSg (quotientS Q sH0H))) divg_normal // oH0hat. + by rewrite -(card_injm injf) // fHhat oHbar -expnSr pfactorK //= mFT_odd. + rewrite /typeP_Galois acts_irrQ // => /mingroupP[_ minUHbar]. + suffices /(card_extraspecial pHhat)[n _ ->]: extraspecial Hhat. + by rewrite pfactorK //= odd_double. + have abelH: p.-abelem (Hhat / H0hat)%g by rewrite -(injm_abelem injf) ?fHhat. + suffices{abelH} defZHhat: 'Z(Hhat) = H0hat. + do 2?split; rewrite defZHhat ?oH0hat //. + apply/eqP; rewrite eqEsubset (Phi_min pHhat) ?normal_norm //=. + by rewrite (Phi_joing pHhat) defHhat' joing_subl. + apply: contraNeq ntH0hat; rewrite eqEsubset sH0hatZ andbT => not_esHhat. + rewrite -defHhat'; apply/eqP/derG1P/center_idP/(quotient_inj nsH0hatZ)=> //. + apply: (injm_morphim_inj injf); rewrite ?quotientS //= fHhat -/Hhat -/H0hat. + rewrite minUHbar //= -/Hbar -?fHhat 1?morphim_injm_eq1 ?morphimS // -subG1. + rewrite quotient_sub1 ?(normal_norm nsH0hatZ) // not_esHhat -['Z(_)]cosetpreK. + rewrite im_f ?sub_cosetpre_quo // quotient_norms ?norm_quotient_pre //. + by rewrite (char_norm_trans (center_char _)) ?quotient_norms. +have [H1 []] := typeP_Galois_Pn maxM notMtype5 gal'M. +rewrite def_p => oH1 nH1Ubar _ /bigdprodWY-defHbar _. +have /cyclicP[xbar defH1]: cyclic H1 by rewrite prime_cyclic ?oH1. +have H1xbar: xbar \in H1 by rewrite defH1 cycle_id. +have sH1_Hbar: H1 \subset Hbar. + by rewrite -[Hbar]defHbar (bigD1 1%g) ?group1 ?conjsg1 ?joing_subl. +have{sH1_Hbar} Hxbar: xbar \in Hbar := subsetP sH1_Hbar xbar H1xbar. +have /morphimP[x nH0x Hx /= Dxbar] := Hxbar. +have{oH1} oxbar: #[xbar] = p by rewrite orderE -defH1. +have memH0: {in H &, forall y z, [~ y, z] \in H0}. + by rewrite defH'; apply: mem_commg. +have [_ /centsP-cHH0hat] := subsetIP sH0hatZ; move/subsetP in nQH. +pose D y z := [~ hat z, hat y]. +have D_H0_1 y z: y \in H -> z \in H0 -> D y z = 1%g. + by move=> Hy H0z; apply/eqP/commgP/cHH0hat; apply: mem_quotient. +have H0_D: {in H &, forall y z, D y z \in H0hat}. + by move=> y z Hy Hz; rewrite -defHhat' mem_commg ?mem_quotient. +have Dsym y z: (D y z)^-1%g = D z y by rewrite invg_comm. +have Dmul y: y \in H -> {in H &, {morph D y: z t / z * t}}%g. + move=> Hy z t Hz Ht; rewrite {1}/D morphM ?nQH // commMgR; congr (_ * _)%g. + by rewrite -{2}morphR ?nQH // -/(D t _) D_H0_1 ?memH0 // mulg1. +pose Dm y Hy : {morphism H >-> coset_of Q} := Morphism (Dmul y Hy). +have{D_H0_1} kerDmH0 y Hy: H0 \subset 'ker (Dm y Hy). + by rewrite subsetI sH0H; apply/subsetP=> z H0z; rewrite !inE /= D_H0_1. +pose x_ w := (x ^ w)%g; pose xbar_ w := (xbar ^ bar w)%g. +move/subsetP in nHW1; move/subsetP in nHU. +have Hx_ w: w \in W1 -> (x_ w \in H) * {in U, forall u, x_ w ^ u \in H}%g. + by move/nHW1=> nHw; split=> [|u /nHU-nHu]; rewrite !memJ_norm. +have Dx: {in H &, forall y z, {in W1, forall w, D (x_ w) y = 1} -> D y z = 1}%g. + move=> y z Hy Hz Dxy1; apply/(kerP (Dm y Hy) Hz); apply: subsetP z Hz. + rewrite -(quotientSGK nH0H) ?kerDmH0 // -defHbar gen_subG. + apply/bigcupsP=> _ /morphimP[w nH0w W1w ->] /=. + rewrite defH1 Dxbar -quotient_cycle -?quotientJ ?quotientS // -cycleJ. + by rewrite cycle_subG !inE /= Hx_ //= -Dsym eq_invg1 Dxy1. +pose ntrivD := [exists w in [predX W1 & W1], #[D (x_ w.1) (x_ w.2)] == p]. +have{ntrivD Dx} /exists_inP[[w1 w2] /andP/=[Ww1 Ww2] /eqP-oDx12]: ntrivD. + apply: contraR ntH0hat => Dx_1; rewrite -defHhat' -subG1 gen_subG. + apply/subsetP=> _ /imset2P[_ _ /morphimP[y ? Hy ->] /morphimP[z ? Hz ->] ->]. + apply/set1P/Dx=> // w2 Ww2; rewrite Dx ?Hx_ // => w1 Ww1. + have abelH0hat: p.-abelem H0hat by apply: prime_abelem. + apply: contraNeq Dx_1 => /(abelem_order_p abelH0hat)oDx12. + by apply/exists_inP; exists (w1, w2); rewrite ?inE ?Ww1 // oDx12 ?H0_D ?Hx_. +have /subsetP-nUW1bar: (W1 / H0)%g \subset 'N(Ubar) := quotient_norms H0 nUW1. +move/subsetP in nH0H; move/subsetP in nH0W1. +pose n (phi : {morphism Ubar >-> {unit 'F_p}}) ubar : nat := val (phi ubar). +have [phi Dphi]: {phi | {in Ubar, forall ub, xbar ^ ub =xbar ^+ n phi ub}}%g. + pose xbar_Autm := invm (injm_Zp_unitm xbar). + have /restrmP[phi [Dphi _ _ _]]: Ubar \subset 'dom (xbar_Autm \o conj_aut H1). + by rewrite -sub_morphim_pre //= im_Zp_unitm -defH1 Aut_conj_aut. + rewrite /n pdiv_id // -oxbar; exists phi => ubar /(subsetP nH1Ubar)Uubar. + transitivity (Zp_unitm (phi ubar) xbar); last by rewrite autE /= -?defH1. + by rewrite Dphi invmK ?im_Zp_unitm -?defH1 ?Aut_aut ?norm_conj_autE. +pose QV ubar w := (ubar ^ (bar w)^-1)%g. +have UbarQV: {in Ubar & W1, forall ubar w, QV ubar w \in Ubar}. + by move=> ub w Uub W1w; rewrite /= memJ_norm ?groupV ?nUW1bar ?mem_quotient. +pose phi_ w ub := phi (QV ub w); pose nphi_ w ub := n phi (QV ub w). +have xbarJ: {in W1 & Ubar, forall w ub, xbar_ w ^ ub = xbar_ w ^+ nphi_ w ub}%g. + by move=> w ubar * /=; rewrite -conjgM conjgCV conjgM Dphi ?UbarQV // conjXg. +have{oDx12} phi_w12 ubar: ubar \in Ubar -> (phi_ w1 ubar * phi_ w2 ubar = 1)%g. + pose n_u := nphi_ ^~ ubar => Uubar; have [u nH0u Uu Dubar] := morphimP Uubar. + suffices: n_u w1 * n_u w2 == 1 %[mod #[D (x_ w1) (x_ w2)]]. + by apply: contraTeq; rewrite oDx12 -!val_Fp_nat // natrM !natr_Zp. + have DXn: {in H & W1, forall y w, D y (x_ w) ^+ n_u w = D y (x_ w ^ u)}%g. + move=> y w Hy W1w; set z := x_ w; have [Hz /(_ u Uu) Hzu] := Hx_ w W1w. + rewrite -(morphX (Dm y Hy)) //; apply/rcoset_kerP; rewrite ?groupX //. + have /subsetP: H0 :* z ^ u \subset 'ker (Dm y Hy) :* z ^ u by rewrite mulSg. + apply; apply/rcoset_kercosetP; rewrite ?groupX ?nH0H //. + by rewrite morphX ?morphJ ?(nH0W1 w) // ?nH0H //= -Dubar -Dxbar xbarJ. + rewrite -eq_expg_mod_order -{1}Dsym expgM expgVn ?(DXn, Dsym) ?Hx_ //. + rewrite /D -!morphR ?nQH ?Hx_ // -conjRg (conjg_fixP _) //. + by apply/commgP/esym/(centsP cH0U); rewrite ?memH0 ?Hx_. +pose wbar := bar (w1 * w2 ^-1)%g; pose W1bar := (W1 / H0)%g. +have W1wbar: wbar \in W1bar by rewrite mem_quotient ?groupM ?groupV. +have{phi_w12} phiJ: {in Ubar, forall ubar, phi (ubar ^ wbar) = (phi ubar)^-1}%g. + move=> ubar Uubar; apply/esym/eqP; rewrite eq_invg_mul. + rewrite [wbar]morphM ?morphV ?nH0W1 ?groupV // -{1}[ubar](conjgK (bar w1)). + by rewrite conjgM phi_w12 // memJ_norm ?nUW1bar ?mem_quotient. +have coW1bar2: coprime #|W1bar| 2 by rewrite coprimen2 quotient_odd ?mFT_odd. +have coUbar2: coprime #|Ubar| 2 by rewrite coprimen2 quotient_odd ?mFT_odd. +have{wbar phiJ W1wbar} phiV: {in Ubar, forall ubar, phi ubar = (phi ubar)^-1}%g. + move=> ubar Uubar; rewrite /= -phiJ // -(expgK coW1bar2 W1wbar) -expgM mul2n. + elim: (expg_invn _ _) => [|k IHk]; first by rewrite conjg1. + by do 2!rewrite expgSr conjgM phiJ ?memJ_norm ?nUW1bar ?groupX // ?invgK. +rewrite -[Hbar]defHbar gen_subG defH1; apply/bigcupsP=> _ /morphimP[w _ Ww ->]. +rewrite -cycleJ cycle_subG -/(xbar_ _); apply/centP=> ubar Uubar; apply/commgP. +apply/conjg_fixP; rewrite xbarJ // /nphi_ -[QV _ w](expgK coUbar2) ?UbarQV //. +by rewrite /n !morphX ?groupX 1?expgS 1?{1}phiV ?UbarQV // mulVg expg1n. +Qed. + +Let defU' : C :=: U'. Proof. by have [] := FTtype34_facts. Qed. +Let H0_1 : H0 :=: 1%g. Proof. by have [] := FTtype34_Fcore_kernel_trivial. Qed. + +Lemma Ptype_Fcompl_kernel_cent : Ptype_Fcompl_kernel MtypeP :=: C. +Proof. +rewrite [Ptype_Fcompl_kernel MtypeP]unlock /= (group_inj H0_1). +by rewrite astabQ -morphpreIim -injm_cent ?injmK ?ker_coset ?norms1. +Qed. +Local Notation defC := Ptype_Fcompl_kernel_cent. + +(* Character theory proper. *) + +Let pddM := FT_prDade_hyp maxM MtypeP. +Let ptiWM : primeTI_hypothesis M HU defW := FT_primeTI_hyp MtypeP. +Let ctiWG : cyclicTI_hypothesis G defW := pddM. +Let ctiWM : cyclicTI_hypothesis M defW := prime_cycTIhyp ptiWM. + +Local Notation sigma := (cyclicTIiso ctiWG). +Local Notation w_ i j := (cyclicTIirr defW i j). +Local Notation eta_ i j := (sigma (w_ i j)). +Local Notation mu_ := (primeTIred ptiWM). +Local Notation Idelta := (primeTI_Isign ptiWM). +Local Notation delta_ j := (primeTIsign ptiWM j). +Local Notation d := (FTtype345_TIirr_degree MtypeP). +Local Notation n := (FTtype345_ratio MtypeP). +Local Notation delta := (FTtype345_TIsign MtypeP). + +Implicit Types zeta xi lambda : 'CF(M). + +Let u := #|U / C|%g. +Let mu2_ i j := primeTIirr ptiWM i j. +Let etaW := map sigma (irr W). +Let eq_proj_eta (alpha gamma : 'CF(G)) := orthogonal (alpha - gamma) etaW. +Let eta_col j := \sum_i eta_ i j. +Let bridge0 zeta := mu_ 0 - zeta. + +Let proj_col_eta j0 i j : '[eta_col j0, eta_ i j] = (j == j0)%:R. +Proof. +rewrite cfdot_suml (bigD1 i) //= cfdot_cycTIiso eqxx eq_sym. +by rewrite big1 ?addr0 // => k /negPf-i'k; rewrite cfdot_cycTIiso i'k. +Qed. + +Let nirrW1 : #|Iirr W1| = q. Proof. by rewrite card_Iirr_cyclic. Qed. +Let NirrW1 : Nirr W1 = q. Proof. by rewrite -nirrW1 card_ord. Qed. + +Let nirrW2 : #|Iirr W2| = p. Proof. by rewrite card_Iirr_cyclic. Qed. +Let NirrW2 : Nirr W2 = p. Proof. by rewrite -nirrW2 card_ord. Qed. + +Let calT := seqIndT HU M. +Let S1 := S_ HC. +Let S2 := seqIndD HU M HC C. + +Let sS10 : cfConjC_subset S1 calS. +Proof. by apply: seqInd_conjC_subset1; rewrite /= ?defMs. Qed. + +Let sS20 : cfConjC_subset S2 calS. +Proof. by apply: seqInd_conjC_subset1; rewrite /= ?defMs. Qed. + +Let scohS1 : subcoherent S1 tau R. Proof. exact: subset_subcoherent sS10. Qed. +Let scohS2 : subcoherent S2 tau R. Proof. exact: subset_subcoherent sS20. Qed. + +Let S1_1 : {in S1, forall zeta, zeta 1%g = q%:R}. +Proof. +move=> _ /seqIndP[s /setDP[kerM'' _] ->]; rewrite !inE -defM'' in kerM''. +by rewrite cfInd1 ?gFsub // -(index_sdprod defM) lin_char1 ?mulr1 ?lin_irr_der1. +Qed. + +Let cohS1 : coherent S1 M^# tau. +Proof. +apply: uniform_degree_coherence scohS1 _. +by apply/(@all_pred1_constant _ q%:R)/allP=> _ /=/mapP[zeta /S1_1<- ->]. +Qed. + +Let irrS1 : {subset S1 <= irr M}. +Proof. +move=> _ /seqIndP[s /setDP[kerHC kerHU] ->]; rewrite !inE in kerHC kerHU. +rewrite -(quo_IirrK _ kerHC) // mod_IirrE // cfIndMod // cfMod_irr //. +have /irr_induced_Frobenius_ker := FrobeniusWker frobMtilde; rewrite defM''. +by apply; rewrite quo_Iirr_eq0 // -subGcfker. +Qed. + +Let o1S1 : orthonormal S1. +Proof. exact: sub_orthonormal (seqInd_uniq _ _) (irr_orthonormal _). Qed. + +Let cfdotS1 : {in S1 &, forall zeta xi, '[zeta, xi] = (zeta == xi)%:R}. +Proof. by case/orthonormalP: o1S1. Qed. + +Let omu2S1 i j : {in S1, forall zeta, '[mu2_ i j, zeta] = 0}. +Proof. +move=> zeta S1zeta; have [s _ Dzeta] := seqIndP S1zeta. +rewrite Dzeta -cfdot_Res_l cfRes_prTIirr cfdot_irr mulrb ifN_eq //. +apply: contraNneq (prTIred_not_irr ptiWM j) => Ds. +by rewrite -cfInd_prTIres Ds -Dzeta irrS1. +Qed. + +Let Tmu j : mu_ j \in calT. Proof. by rewrite -cfInd_prTIres mem_seqIndT. Qed. + +Let omuS1 j : {in S1, forall zeta, '[mu_ j, zeta] = 0}. +Proof. +by move=> zeta S1zeta /=; rewrite cfdot_suml big1 // => i _; apply: omu2S1. +Qed. + +Let Zbridge0 : {in S1, forall zeta, bridge0 zeta \in 'Z[irr M, HU^#]}. +Proof. +have mu0_1: mu_ 0 1%g = q%:R by rewrite prTIred_1 prTIirr0_1 mulr1. +move=> zeta S1zeta; rewrite /= zcharD1 !cfunE subr_eq0 mu0_1 S1_1 // eqxx. +by rewrite rpredB ?(seqInd_vchar _ (Tmu 0)) ?(seqInd_vchar _ S1zeta). +Qed. + +Let A0bridge0 : {in S1, forall zeta, bridge0 zeta \in 'CF(M, 'A0(M))}. +Proof. by move=> zeta /Zbridge0/zchar_on/cfun_onS->. Qed. + +Let sS1S2' : {subset S1 <= [predC S2]}. +Proof. +by move=> _ /seqIndP[s /setDP[kHCs _] ->]; rewrite !inE mem_seqInd // inE kHCs. +Qed. + +Let defS2: S2 = seqIndD HU M H H0C. +Proof. by rewrite /S2 H0_1 -!joinGE join1G joinGC seqIndDY. Qed. + +Let cohS2: coherent S2 M^# tau. +Proof. +apply: subset_coherent (Ptype_core_coherence maxM MtypeP notMtype5). +by rewrite defC defS2; apply: seqIndS; rewrite Iirr_kerDS ?genS ?setUS ?der_sub. +Qed. + +Let Smu := [seq mu_ j | j in predC1 0]. +Let Sred := filter [predC irr M] (seqIndD HU M H H0). + +Let memSred : Sred =i Smu. +Proof. +have [szSred _ memSred _] := typeP_reducible_core_Ind maxM MtypeP notMtype5. +have uSred: uniq Sred by apply: filter_uniq (seqInd_uniq _ _). +suffices{uSred}: (size Smu <= size Sred)%N by case/leq_size_perm. +by rewrite szSred def_p size_map -cardE cardC1 nirrW2. +Qed. + +Let mu_1 j : j != 0 -> mu_ j 1%g = (q * u)%:R. +Proof. +move=> nzj; have Smuj: mu_ j \in Sred by rewrite memSred image_f. +have [_ _ _ /(_ _ Smuj)[]] := typeP_reducible_core_Ind maxM MtypeP notMtype5. +by rewrite defC. +Qed. + +Let memS2red : [predD S2 & irr M] =i Smu. +Proof. +move=> xi; rewrite defS2 -memSred mem_filter; apply: andb_id2l => /= red_xi. +apply/idP/idP=> [|Sxi]; first by apply: seqIndS; rewrite Iirr_kerDS ?joing_subl. +have [_ _ _ /(_ xi)] := typeP_reducible_core_Ind maxM MtypeP notMtype5. +by rewrite defC mem_filter /= red_xi; case. +Qed. + +Let i1 : Iirr W1 := inord 1. +Let nz_i1 : i1 != 0. Proof. by rewrite Iirr1_neq0. Qed. +Let j1 : Iirr W2 := inord 1. +Let nz_j1 : j1 != 0. Proof. by rewrite Iirr1_neq0. Qed. + +(* This is Peterfalvi (11.8). *) +(* We have rearranged the argument somewhat: *) +(* - Step (11.8.4) was out of sequence as it involves changing the definition *) +(* of tau2, which requires showing that (11.8.2-3) are preserved by this *) +(* change; since (11.8.4) does not use (11.8.2-3) we avoid this by proving *) +(* (11.8.4) first. *) +(* - The first part of step (11.8.3) is the last fact that needs to be proved *) +(* for an arbitrary j != 0; (11.8.1, 5-6) can all use the same fixed j != 0 *) +(* (we take j = 1), provided (11.8.3) is proved before (11.8.2), which it *) +(* doe not use. *) +(* - Steps (11.8.2) and (11.8.5) are really as combination, to provide an *) +(* expression for tau (alpha i j) for an arbitrary i. We merge their proofs *) +(* so we can use a fixed i for the whole combined step and hide some *) +(* intermediate technical facts. *) +(* - We also reorganise the contents of the superstep, proving most of *) +(* (11.8.5) between the first and last two parts of (11.8.2); this *) +(* simplifies the latter because a is then known to be even, so we can show *) +(* directly that a is 0 or 2, and then that X = eta i j - eta i 0. *) +Lemma FTtype34_not_ortho_cycTIiso zeta : + zeta \in S1 -> ~~ eq_proj_eta (tau (bridge0 zeta)) (eta_col 0). +Proof. +move=> S1zeta; set psi := tau _; apply/negP=> proj_psi_eta. +have irr_zeta: zeta \in irr M := irrS1 S1zeta. +have Szeta: zeta \in S_ 1 by apply: seqInd_sub S1zeta. +have Zzeta_S1: {in S1, forall xi, zeta - xi \in 'Z[S1, M^#]}. + by move=> * /=; rewrite zcharD1E rpredB ?mem_zchar //= !cfunE !S1_1 ?subrr. +have n1S1: {in S1, forall xi, '[xi] = 1} by move=> xi /irrS1/irrWnorm. +have Z_S1: {in S1, forall xi, xi \in 'Z[S1]} by apply: mem_zchar. +have [p_gt0 q_gt0 u_gt0]: [/\ p > 0, q > 0 & u > 0]%N by rewrite !cardG_gt0. +have q_gt2: (q > 2)%N by rewrite odd_prime_gt2 ?mFT_odd. +have mu2_1 i j: j != 0 -> mu2_ i j 1%g = d%:R. + by have [/(_ i j)] := FTtype345_constants maxM MtypeP notMtype2. +(* This is (11.8.1). *) +have [Dd delta1 Dn]: [/\ d = u, delta = 1 & n = (size S1)%:R]. + have size_S1 : (size S1 * q = u - 1)%N. + rewrite mulnC [q](index_sdprod defM). + rewrite (size_irr_subseq_seqInd _ (subseq_refl _)) //. + transitivity #|[set mod_Iirr t | t : Iirr (HU / HC) in predC1 0]|. + apply/eq_card=> s; rewrite inE mem_seqInd // !inE subGcfker. + apply/andP/imsetP=> [[nzs kHCs] | [t nzt ->]]. + by exists (quo_Iirr HC s); rewrite ?quo_IirrK // inE quo_Iirr_eq0. + by rewrite mod_Iirr_eq0 // mod_IirrE // cfker_mod. + rewrite card_imset; last exact: can_inj (mod_IirrK _). + have isoUC: U / C \isog HU / HC by apply: quotient_sdprodr_isog. + rewrite subn1 cardC1 card_Iirr_abelian -?(card_isog isoUC) //. + by rewrite -(isog_abelian isoUC) defU' der_abelian. + have Dd: d = u. + apply/eqP; rewrite -(eqn_pmul2l q_gt0) -eqC_nat -(mu_1 nz_j1). + by rewrite natrM prTIred_1 mu2_1. + suffices delta1: delta = 1. + by rewrite /n Dd delta1 -(@natrB _ _ 1) // -size_S1 natrM mulfK ?neq0CG. + have: (delta == 1 %[mod q])%C. + rewrite -(eqCmod_transl _ (prTIirr1_mod ptiWM 0 j1)) mu2_1 // -/q Dd. + by rewrite /eqCmod -(@natrB _ u 1) // dvdC_nat -size_S1 dvdn_mull. + rewrite -[1]subr0 [delta]signrE -/ptiWM eqCmodDl eqCmodN opprK. + by rewrite eqCmod0_nat; case: (Idelta j1); first rewrite gtnNdvd. +have deltaZ gamma: delta *: gamma = gamma by rewrite delta1 scale1r. +have [tau1 coh_tau1] := cohS1; pose zeta1 := tau1 zeta. +(* This is (11.8.4). *) +without loss Dpsi: tau1 coh_tau1 @zeta1 / psi = eta_col 0 - zeta1. + move=> IHtau1; have [[Itau1 Ztau1] Dtau1] := coh_tau1. + have tau1_dirr: {in S1, forall xi, tau1 xi \in dirr G}. + by move=> xi S1xi; rewrite /= dirrE Ztau1 ?Itau1 ?mem_zchar //= n1S1. + pose chi : 'CF(G) := eta_col 0 - psi. + have Dpsi: psi = eta_col 0 - chi by rewrite opprD opprK addNKr. + have chi'zeta1: chi <> zeta1. + by move=> Dchi; case: (IHtau1 tau1); rewrite -/zeta1 -?Dchi. + have dirr_chi: chi \in dirr G. + apply: dirr_norm1. + rewrite rpredB ?rpred_sum // => [i _|]; first exact: cycTIiso_vchar. + rewrite Dade_vchar // zchar_split A0bridge0 //. + by rewrite rpredB ?char_vchar ?prTIred_char ?irrWchar. + apply: (addrI q%:R); transitivity '[psi]; last first. + rewrite Dade_isometry ?A0bridge0 // (cfnormBd (omuS1 _ _)) //. + by rewrite cfnorm_prTIred n1S1. + rewrite Dpsi [RHS]cfnormDd; last first. + rewrite opprB cfdotC cfdot_sumr big1 ?conjC0 // => i _. + by rewrite (orthoPl proj_psi_eta) ?map_f ?mem_irr. + rewrite cfnormN -nirrW1 -sumr_const cfdot_sumr. + by congr (_ + _); apply: eq_bigr => i _; rewrite proj_col_eta. + have Dchi: {in S1, forall xi, xi != zeta -> chi = - tau1 xi}. + move=> xi S1xi /negPf-zeta'xi; have irr_xi := irrS1 S1xi. + suffices: '[zeta1 - tau1 xi, chi] = 1. + by case/cfdot_add_dirr_eq1; rewrite ?rpredN ?tau1_dirr. + rewrite /= cfdotBr cfdot_sumr big1 => [|i _]; last first. + have oS1eta := coherent_ortho_cycTIiso MtypeP sS10 coh_tau1. + by rewrite cfdotBl !oS1eta ?irrS1 ?subrr. + rewrite -raddfB Dtau1 ?Zzeta_S1 // Dade_isometry ?A0bridge0 //; last first. + exact: cfun_onS sHU_A0 (zcharD1_seqInd_on _ (Zzeta_S1 xi S1xi)). + rewrite cfdotBr cfdotC cfdotBr 2?omuS1 // subrr conjC0 !sub0r opprK. + by rewrite cfdotBl n1S1 // cfdotS1 // zeta'xi subr0. + have S1zetaC: zeta^*%CF \in S1 by rewrite cfAut_seqInd. + have Dchi_zetaC: chi = - tau1 zeta^*%CF. + by rewrite -Dchi ?(seqInd_conjC_neq _ _ _ S1zeta) ?mFT_odd. + suffices S1le2: (size S1 <= size [:: zeta; zeta^*%CF])%N. + case: (IHtau1 (dual_iso tau1)); last by rewrite /= -Dchi_zetaC. + exact: dual_coherence scohS1 coh_tau1 S1le2. + rewrite uniq_leq_size ?seqInd_uniq // => xi S1xi. + rewrite !inE -implyNb; apply/implyP=> zeta'xi; apply/eqP. + apply: (Zisometry_inj Itau1); rewrite ?mem_zchar ?cfAut_seqInd //. + by apply: oppr_inj; rewrite -Dchi. +have [[Itau1 Ztau1] Dtau1] := coh_tau1. +have tau1_dirr: {in S1, forall xi, tau1 xi \in dirr G}. + by move=> xi S1xi; rewrite /= dirrE Ztau1 ?Itau1 ?mem_zchar //= n1S1. +have oS1eta i j: {in S1, forall xi, '[tau1 xi, eta_ i j] = 0}. + by move=> xi S1xi /=; rewrite (coherent_ortho_cycTIiso _ _ coh_tau1) ?irrS1. +pose alpha_ := FTtype345_bridge MtypeP zeta. +have A0alpha i j : j != 0 -> alpha_ i j \in 'CF(M, 'A0(M)). + by move/supp_FTtype345_bridge->; rewrite ?S1_1. +have alpha_S1 i j: {in S1, forall xi, '[alpha_ i j, xi] = n *- (xi == zeta)}. + move=> xi S1xi; rewrite /= !cfdotBl !cfdotZl !omu2S1 // mulr0 subrr add0r. + by rewrite cfdotS1 // eq_sym mulr_natr. +pose beta_ i j := tau (alpha_ i j) - (eta_ i j - eta_ i 0) + n *: zeta1. +pose beta := beta_ 0 j1. +(* This is the first part of (11.8.3). *) +have betaE i j: j != 0 -> beta_ i j = beta. + move=> nz_j; transitivity (beta_ i j1); congr (_ + _); apply/eqP. + rewrite eq_sym -subr_eq0 [rhs in _ + rhs]opprD addrACA -opprD subr_eq0. + rewrite -linearB /= !opprB !addrA !subrK -!/(mu2_ i _). + by rewrite [Dade pddM _]prDade_sub_TIirr ?mu2_1 //= deltaZ. + rewrite -subr_eq0 !opprD addrACA -3!opprD opprK subr_eq0 addrACA addrA. + rewrite -(prDade_sub2_TIirr pddM) -!/(mu2_ _ _) !deltaZ -linearB /=. + by rewrite opprB addrA subrK !deltaZ opprD opprK addrACA addrA. +pose j := j1. (* The remainder of the proof only uses j = 1. *) +(* This is the second part of (11.8.3). *) +have Rbeta: cfReal beta. + rewrite /cfReal eq_sym -subr_eq0 rmorphD !rmorphB /= opprB 2!opprD opprB -/j. + rewrite 2![(eta_ 0 _)^*%CF]cfAut_cycTIiso -!cycTIirr_aut !aut_Iirr0 -Dade_aut. + set k := aut_Iirr conjC j; rewrite -(betaE 0 k) ?aut_Iirr_eq0 // addrACA. + rewrite addrC addr_eq0 addrCA subrK opprD opprK Dn raddfZnat -!raddfB /= -Dn. + apply/eqP; rewrite (cfConjC_Dade_coherent coh_tau1) ?mFT_odd // -raddfB. + rewrite Dtau1 ?Zzeta_S1 ?cfAut_seqInd //= -linearZ scalerBr; congr (tau _). + rewrite opprD !rmorphB !deltaZ /= -!prTIirr_aut !aut_Iirr0 addrACA subrr. + by rewrite add0r opprK addrC Dn -raddfZnat. +(* This is the consequence of Peterfalvi (11.8.2) and (11.8.5). *) +have tau_alpha i: tau (alpha_ i j) = eta_ i j - eta_ i 0 - n *: zeta1. + set phi := tau (alpha_ i j); pose sum_tau1 := \sum_(xi <- S1) tau1 xi. + have A0alpha_j k: alpha_ k j \in 'CF(M, 'A0(M)) by apply: A0alpha. + have Zphi: phi \in 'Z[irr G]. + by rewrite Dade_vchar // zchar_split vchar_FTtype345_bridge /=. + have [Y S1_Y [X [Dphi oYX oXS1]]] := orthogonal_split (map tau1 S1) phi. + (* This is the first part of 11.8.2 *) + have [a Za defY]: exists2 a, a \in Cint & Y = a *: sum_tau1 - n *: zeta1. + have [a_ Da defY] := orthonormal_span (map_orthonormal Itau1 o1S1) S1_Y. + have{Da} Da: {in S1, forall xi, a_ (tau1 xi) = '[phi, tau1 xi]}. + by move=> xi Sxi; rewrite Da Dphi cfdotDl (orthoPl oXS1) ?map_f ?addr0. + exists (a_ (tau1 zeta) + n). + by rewrite Dn rpredD ?rpred_nat // Da // Cint_cfdot_vchar ?Ztau1 ?Z_S1. + rewrite defY big_map scaler_sumr !(bigD1_seq _ S1zeta) ?seqInd_uniq //=. + rewrite addrAC scalerDl addrK !(big_seq_cond (predC1 _)) /=; congr (_ + _). + apply: eq_bigr => xi /andP[S1xi zeta'xi]; congr (_ *: _); rewrite !Da //. + apply: canRL (addNKr _) _; rewrite addrC -opprB -!raddfB Dtau1 ?Zzeta_S1//=. + rewrite Dade_isometry //; last first. + exact: cfun_onS (zcharD1_seqInd_on _ (Zzeta_S1 _ S1xi)). + by rewrite cfdotBr !alpha_S1 // !mulrb eqxx ifN_eq // !(addr0, opprK). + have psi_phi: '[psi, phi] = -1 + n. (* This is part of (11.8.5). *) + rewrite cfdotC Dade_isometry ?A0bridge0 //. + rewrite cfdotBr !cfdotBl deltaZ !cfdotZl n1S1 // mulr1. + rewrite !cfdot_prTIirr_red (negPf nz_j1) eqxx !omu2S1 //= cfdotC omuS1 //. + by rewrite conjC0 mulr0 opprB !subr0 add0r rmorphD rmorphN Dn !rmorph_nat. + have{psi_phi} col0_beta: '[eta_col 0, beta] = a. (* Also part of (11.8.5). *) + apply/(addIr (-1 + n))/(canRL (addNKr _)). + rewrite addrCA addrA addrACA -{}psi_phi Dpsi cfdotBl; congr (_ + _). + rewrite -(betaE i j) // cfdotDr !cfdotBr -/phi cfdotZr -!addrA. + apply/(canLR (addNKr _)); rewrite addNr !cfdot_suml. + rewrite big1 ?add0r ?opprK => [|k _]; last first. + by rewrite cfdot_cycTIiso andbC eq_sym (negPf nz_j1). + rewrite addrCA big1 ?mulr0 ?add0r => [|k _]; last first. + by rewrite cfdotC oS1eta ?conjC0. + rewrite addrC (bigD1 i) // cfnorm_cycTIiso /= addKr big1 // => k i'k. + by rewrite cfdot_cycTIiso (negPf i'k). + rewrite cfdotC Dphi cfdotDl (orthoPl oXS1) ?map_f // addr0. + rewrite defY cfdotBl scaler_sumr cfproj_sum_orthonormal //. + rewrite cfdotZl Itau1 ?mem_zchar ?n1S1 // mulr1 rmorphB opprD opprK. + by rewrite Dn rmorph_nat conj_Cint. + have a_even: (2 %| a)%C. (* Third internal part of (11.8.5). *) + have Zbeta: beta \in 'Z[irr G]. + rewrite -{1}(betaE i j) // rpredD ?rpredB ?Zphi ?cycTIiso_vchar //. + by rewrite Dn rpredZnat // Ztau1 ?mem_zchar. + rewrite -col0_beta cfdot_real_vchar_even ?mFT_odd //; first 1 last. + split; first by apply/rpred_sum=> k _; apply: cycTIiso_vchar. + apply/eqP; rewrite [RHS](reindex_inj (can_inj (@conjC_IirrK _ _))) /=. + rewrite rmorph_sum; apply/eq_bigr=> k _ /=. + by rewrite cfAut_cycTIiso -cycTIirr_aut aut_Iirr0. + have eta00: eta_ 0 0 = 1 by rewrite cycTIirr00 cycTIiso1. + rewrite orbC cfdotDl 2!cfdotBl cfdotZl -eta00 oS1eta // mulr0 addr0. + rewrite opprB addrC 2!{1}cfdot_cycTIiso (negPf nz_j1) subr0 /= eta00. + rewrite Dade_reciprocity // => [|x _ y _]; last by rewrite !cfun1E !inE. + rewrite cfRes_cfun1 !cfdotBl deltaZ !cfdotZl -!/(mu2_ 0 _). + rewrite -(prTIirr00 ptiWM) !cfdot_prTIirr cfdotC omu2S1 // conjC0 mulr0. + by rewrite (negPf nz_j1) add0r subr0 subrr rpred0. + have nY: '[Y] = n * a * (a - 2%:R) + n ^+ 2. (* Resuming step (11.8.2). *) + rewrite defY cfnormD cfnormN !cfnormZ cfdotNr cfdotZr. + rewrite cfnorm_map_orthonormal // -Dn Itau1 ?mem_zchar ?n1S1 // mulr1. + rewrite scaler_sumr cfproj_sum_orthonormal // rmorphN addrAC. + rewrite Dn rmorphM !Cint_normK ?rpred_nat // !rmorph_nat conj_Cint // -Dn. + by rewrite -mulr2n mulrC mulrA -mulr_natr mulNr -mulrBr. + have{a_even} Da: (a == 0) || (a == 2%:R). (* Second part of (11.8.2). *) + suffices (b := a - 1): b ^+ 2 == 1. + by rewrite -!(can_eq (subrK 1) a) add0r addrK orbC -eqf_sqr expr1n. + have S1gt0: (0 < size S1)%N by case: (S1) S1zeta. + have: n * b ^+ 2 <= n *+ 3. + have: 2%:R + n <= n *+ 3 by rewrite addrC ler_add2l ler_muln2r Dn ler1n. + apply: ler_trans; rewrite sqrrB1 -mulr_natr -mulrBr mulrDr mulrA mulr1. + rewrite ler_add2r -(ler_add2r (n ^+ 2 + '[X])) !addrA -nY -cfnormDd //. + by rewrite -Dphi norm_FTtype345_bridge ?S1_1 // ler_addl cfnorm_ge0. + have Zb: b \in Cint by rewrite rpredB ?rpred1 ?Za. + have nz_b: b != 0 by rewrite subr_eq0 (memPn _ a a_even) ?(dvdC_nat 2 1). + rewrite eqr_le sqr_Cint_ge1 {nz_b}//= andbT -Cint_normK // Dn -mulrnA. + have /CnatP[m ->] := Cnat_norm_Cint Zb; rewrite -natrX -natrM leC_nat. + by rewrite leq_pmul2l // lern1 -ltnS (ltn_sqr m 2) (leq_sqr m 1). + have{nY Da} defX: X = eta_ i j - eta_ i 0. (* Last part of (11.8.2). *) + have{nY Da} /eqP-nY: '[Y] == n ^+ 2. + by rewrite -subr_eq0 nY addrK -mulrA !mulf_eq0 !subr_eq0 Da orbT. + have coh_zeta_phi := FTtype345_bridge_coherence _ _ Szeta _ coh_tau1. + have:= Dphi; rewrite addrC => /coh_zeta_phi->; rewrite ?S1_1 ?deltaZ //. + rewrite defY scaler_sumr big_seq rpredB ?rpred_sum // => [xi Sxi|]. + by rewrite rpredZ_Cint ?mem_zchar ?map_f. + by rewrite Dn rpredZnat ?mem_zchar ?map_f. + have{col0_beta} a0: a = 0. (* This is the conclusion of (11.8.5). *) + rewrite cfdot_suml big1 // in col0_beta => k _. + rewrite -(betaE i j) // /beta_ -/phi Dphi -defX addrK defY subrK. + rewrite cfdotZr cfdot_sumr big1_seq ?mulr0 // => xi S1xi. + by rewrite cfdotC oS1eta ?conjC0. + by rewrite Dphi defY defX a0 ?inE ?eqxx // scale0r sub0r addrC. +(* This is step (11.8.6). *) +pose theta := mu_ j - d%:R *: zeta. +have /andP/=[red_muj S2muj]: mu_ j \in [predD S2 & irr M]. + by rewrite memS2red image_f. +have HUtheta: theta \in 'CF(M, HU^#). + rewrite cfun_onD1 !cfunE mu_1 ?S1_1 // Dd mulrC natrM subrr eqxx. + by rewrite rpredB ?rpredZ ?(seqInd_on _ S1zeta) ?(seqInd_on _ S2muj). +have Dtheta: theta = mu_ 0 - zeta + \sum_i alpha_ i j. + rewrite !sumrB -scaler_sumr delta1 scale1r. + rewrite [X in _ = X]addrC -!addrA -/(mu_ 0); congr (_ + _). + rewrite [X in _ = _ + X]addrC !addrA addNr add0r -opprD; congr (- _). + rewrite sumr_const nirrW1 -scaler_nat scalerA mulrC. + by rewrite divfK ?neq0CG // delta1 addrC scalerBl scale1r subrK. +have tau_theta: tau theta = eta_col j - d%:R *: zeta1. + pose psi1 i := eta_ i j1 - eta_ i 0 - n *: zeta1. + have Dpsi1 i: tau (alpha_ i j) = psi1 i by apply: tau_alpha. + rewrite Dtheta [tau _]raddfD raddf_sum (eq_bigr psi1) //= {Dpsi1}/psi1 -/psi. + rewrite Dpsi !sumrB [X in X = _]addrC -!addrA; congr (_ + _). + rewrite -opprB -opprD -opprB -/(eta_col 0) addrA addrK; congr (- _). + rewrite sumr_const nirrW1 -scaler_nat scalerA mulrC. + by rewrite divfK ?neq0CG // delta1 scalerBl scale1r subrK. +have [tau2 coh_tau2] := cohS2. +without loss tau2muj: tau2 coh_tau2 / tau2 (mu_ j) = eta_col j; last first. + case: FTtype34_noncoherence; rewrite H0_1 -joinGE join1G. + have uS12: uniq (S2 ++ S1). + by rewrite cat_uniq ?seqInd_uniq ?andbT //=; apply/hasPn. + have /perm_eq_coherent: perm_eq (S2 ++ S1) (S_ C); last apply. + apply: uniq_perm_eq; rewrite ?seqInd_uniq // => xi; rewrite mem_cat. + apply/idP/idP=> [/orP | /seqIndP[i /setDP[kCi k'HUi] ->]]. + by case; apply/seqIndS/Iirr_kerDS; rewrite ?joing_subr. + by rewrite !mem_seqInd // inE orbC inE kCi k'HUi andbT orbN. + move: tau_theta; rewrite -tau2muj // -raddfZnat. + apply: (bridge_coherent scohM) sS20 coh_tau2 sS10 coh_tau1 sS1S2' _. + by rewrite (cfun_onS _ HUtheta) ?setSD // rpredZnat ?Z_S1. +move=> IHtau2; apply: (IHtau2 tau2 coh_tau2); have [IZtau2 Dtau2] := coh_tau2. +have{IHtau2} /hasP[xi S2xi /=irr_xi]: has [mem irr M] S2. + apply/hasPn=> redS2 {tau2 coh_tau2 IZtau2 Dtau2}. + have muS2: {subset S2 <= Smu} by move=> xi S2xi; rewrite -memS2red !inE redS2. + have [_ [tau2 tau2mu coh_tau2]] := uniform_prTIred_coherent pddM nz_j1. + have S2uniform: {subset S2 <= uniform_prTIred_seq pddM j}. + move=> _ /muS2/imageP[k nz_k ->]; apply: image_f. + by rewrite !inE [_ != 0]nz_k /= !mu_1. + apply: (IHtau2 tau2); first exact: subset_coherent_with coh_tau2. + have [_ /(_ _ nz_j1) Ez _ _] := FTtype345_constants maxM MtypeP notMtype2. + by have:= tau2mu j; rewrite Ez -/delta delta1 scale1r. +suffices: '[tau2 (mu_ j), eta_col j] != 0. + have:= FTtypeP_coherent_TIred sS20 coh_tau2 irr_xi S2xi S2muj. + case=> _ -> [[-> ->] | [-> -> _] /eqP[]]; first by rewrite deltaZ. + rewrite -[cyclicTIiso _]/sigma cfdot_sumr big1 ?mulr0 // => i _. + rewrite cfdotZl proj_col_eta -(inj_eq irr_inj) conjC_IirrE eq_sym. + by rewrite odd_eq_conj_irr1 ?mFT_odd // irr_eq1 (negPf nz_j1) mulr0. +pose gamma := xi 1%g *: mu_ j - mu_ j 1%g *: xi. +have: '[tau2 gamma, tau theta] != 0. + have [Txi Tzeta] := (seqInd_subT S2xi, seqInd_subT S1zeta). + have S2gamma: gamma \in 'Z[S2, HU^#] by apply: sub_seqInd_zchar. + rewrite Dtau2 ?zcharD1_seqInd //; move/zchar_on in S2gamma. + rewrite Dade_isometry ?(cfun_onS sHU_A0) // cfdotBr cfdotZr !cfdotBl !cfdotZl. + rewrite cfnorm_prTIred omuS1 // (seqInd_ortho _ _ S2muj) ?(memPn red_muj) //. + rewrite (seqInd_ortho _ Txi) ?(memPn (sS1S2' _)) // !(mulr0, subr0) mulf_eq0. + by rewrite char1_eq0 ?irrWchar // -cfnorm_eq0 irrWnorm ?oner_eq0 ?neq0CG. +apply: contraNneq => o_muj_etaj; rewrite {}tau_theta !{gamma}raddfB subr_eq0 /=. +have /CnatP[xi1 ->]: xi 1%g \in Cnat by rewrite Cnat_char1 ?irrWchar. +rewrite mu_1 // cfdotZr !cfdotBl !raddfZnat !cfdotZl {}o_muj_etaj cfdot_sumr. +have /orthogonalP oS2_S1: orthogonal (map tau2 S2) (map tau1 S1). + exact: (coherent_ortho scohM) sS20 coh_tau2 sS10 coh_tau1 sS1S2'. +rewrite !oS2_S1 ?map_f // big1 ?(mulr0, subr0) // => k _. +exact: (coherent_ortho_cycTIiso _ _ coh_tau2). +Qed. + +(* This is Peterfalvi (11.9). *) +(* Note that in the proof of part (a), the text improperly suggests using *) +(* (5.3.b) to show then tau (zeta - zeta^alpha) is orthogonal to the eta i j. *) +(* Since alpha might not be conjugation, this is not obvious. Indeed the best *) +(* way to derive this is to use (5.5) together with the coherence of S(HC). *) +(* In part (c) we start by reducing the proof to q <= p - 1; we also don't *) +(* compute [tau (mu0 - zeta), tau2 lambda] = [chi, tau2 lambda] since this *) +(* is not needed to prove than u = a: one only needs to show that the *) +(* the left-hand side is an integer, which is in fact required to show that *) +(* the right-hand is an integer. *) +Lemma FTtype34_structure (eta0row := \sum_j eta_ 0 j) : + [/\ (*a*) {in S1, forall zeta, eq_proj_eta (tau (bridge0 zeta)) eta0row}, + (*b*) (p < q)%N + & (*c*) FTtype M == 3 /\ typeP_Galois MtypeP]. +Proof. +have sum_etaW F: \sum_(eta <- etaW) F eta = \sum_i \sum_j F (eta_ i j). + rewrite big_map big_tuple (reindex (dprod_Iirr defW)) /=. + by rewrite pair_bigA; apply: eq_bigr => -[i j]. + by exists (inv_dprod_Iirr defW) => ij; rewrite ?dprod_IirrK ?inv_dprod_IirrK. +have bridgeS1: {in S1, forall zeta, eq_proj_eta (tau (bridge0 zeta)) eta0row}. + move=> zeta S1zeta; set phi := bridge0 zeta; have irr_zeta := irrS1 S1zeta. + have [X etaX [chi [Dchi oXchi o_chi_eta]]] := orthogonal_split etaW (tau phi). + have [Isigma Zsigma] := cycTI_Zisometry ctiWG. + have{o_chi_eta} o_chi_eta i j: '[chi, eta_ i j] = 0. + by rewrite (orthoPl o_chi_eta) ?map_f ?mem_irr. + have o1etaW: orthonormal etaW by rewrite map_orthonormal ?irr_orthonormal. + have [a Da defX] := orthonormal_span o1etaW etaX; pose a_ := a (eta_ _ _). + have{Da} Da i j: a_ i j = '[tau phi, eta_ i j]. + by rewrite Dchi cfdotDl o_chi_eta addr0 /a_ Da. + have Zphi: phi \in 'Z[irr M, HU^#] by apply: Zbridge0. + have A0phi: phi \in 'CF(M, 'A0(M)) by apply: A0bridge0. + have a00_1 : a_ 0 0 = 1. + rewrite Da cycTIirr00 [sigma 1]cycTIiso1. + rewrite Dade_reciprocity // => [|x _ y _]; last by rewrite !cfun1E !inE. + rewrite rmorph1 /= -(prTIirr00 ptiWM) -/(mu2_ 0 0) cfdotC. + by rewrite cfdotBr cfdot_prTIirr_red omu2S1 // subr0 rmorph1. + have aut_phi nu: cfAut nu (tau phi) = tau phi + tau (zeta - cfAut nu zeta). + rewrite -Dade_aut !rmorphB !raddfB /= !addrA subrK. + by rewrite -prTIred_aut aut_Iirr0. + have Za i j: a_ i j \in Cint. + rewrite Da Cint_cfdot_vchar ?cycTIiso_vchar //. + by rewrite Dade_vchar ?(zchar_onS sHU_A0). + have [tau1 coh_tau1] := cohS1; have [_ Dtau1] := coh_tau1. + have o_tau1_eta := coherent_ortho_cycTIiso MtypeP sS10 coh_tau1. + have a_aut nu i j: a (cfAut nu (eta_ i j)) = a_ i j. + symmetry; transitivity '[cfAut nu (tau phi), cfAut nu (eta_ i j)]. + by rewrite cfdot_aut_vchar ?cycTIiso_vchar // -Da aut_Cint. + rewrite aut_phi cfAut_cycTIiso -cycTIirr_aut [a _]Da cfdotDl addrC. + rewrite -Dtau1 ?zcharD1_seqInd ?seqInd_sub_aut_zchar // raddfB cfdotBl. + by rewrite !o_tau1_eta ?cfAut_seqInd ?cfAut_irr // subr0 add0r. + pose a10 := a_ i1 0; pose a01 := a_ 0 j1; pose a11 := a_ i1 j1. + have Da10 i: i != 0 -> a_ i 0 = a10. + case/(cfExp_prime_transitive pr_q nz_i1) => k co_k_wi1 Dwi. + rewrite -(cforder_dprodl defW) -dprod_IirrEl in co_k_wi1. + have [[nu eta10nu] _] := cycTIiso_aut_exists ctiWG co_k_wi1. + by rewrite /a_ dprod_IirrEl Dwi rmorphX /= -dprod_IirrEl eta10nu a_aut. + have Da01 j: j != 0 -> a_ 0 j = a01. + case/(cfExp_prime_transitive pr_p nz_j1) => k co_k_wj1 Dwj. + rewrite -(cforder_dprodr defW) -dprod_IirrEr in co_k_wj1. + have [[nu eta01nu] _] := cycTIiso_aut_exists ctiWG co_k_wj1. + by rewrite /a_ dprod_IirrEr Dwj rmorphX /= -dprod_IirrEr eta01nu a_aut. + have DaB1 i j: a_ i j = a_ i 0 + a_ 0 j - a_ 0 0. + apply: (canRL (addrK _)); rewrite !Da cycTIiso_cfdot_exchange // => x Vx. + have /setDP[A0x A'x]: x \in 'A0(M) :\: 'A(M). + by rewrite (FTsupp0_typeP maxM MtypeP) // mem_class_support. + by rewrite Dade_id // (cfun_on0 (zchar_on Zphi)) // -defA. + pose p1 : algC := p.-1%:R; pose q1 : algC := q.-1%:R. + have normX: '[X] = 1 + q1 * a10 ^+ 2 + p1 * a01 ^+ 2 + p1 * q1 * a11 ^+ 2. + transitivity (\sum_i \sum_j a_ i j ^+ 2). + rewrite defX cfnorm_sum_orthonormal // sum_etaW. + by apply/eq_bigr=> i _; apply/eq_bigr=> j _; rewrite Cint_normK ?Za. + rewrite -addrA addrACA (bigD1 0) //= (bigD1 0) //= a00_1 expr1n. + rewrite -natrM !mulr_natl mulrnA -mulrnDl. + rewrite -nirrW1 -nirrW2 -!(cardC1 0) -!sumr_const. + congr (1 + _ + _); first by apply: eq_bigr => j /Da01->. + apply: eq_bigr => i /Da10-Dai0; rewrite (bigD1 0) //= Dai0; congr (_ + _). + by apply: eq_bigr => j /Da01-Da0j; rewrite DaB1 Dai0 Da0j -DaB1. + have normX_le_q: '[X] <= q%:R. + rewrite -(ler_add2r '[chi]) -cfnormDd // -Dchi -ler_subl_addl. + have ->: '[tau phi] - q%:R = 1. + rewrite Dade_isometry ?A0bridge0 // cfnormBd; last by rewrite omuS1. + by rewrite cfnorm_prTIred cfdotS1 // eqxx addrC addKr. + suffices: '[chi] != 0. + suffices /CnatP[nchi ->]: '[chi] \in Cnat by rewrite ler1n lt0n -eqC_nat. + rewrite Cnat_cfnorm_vchar // -(canLR (addKr _) Dchi) defX addrC rpredB //. + by rewrite Dade_vchar // (zchar_onS (FTsupp_sub0 M)) ?defA. + rewrite big_map big_seq rpred_sum // => _ /(cycTIirrP defW)[i [j ->]]. + by rewrite rpredZ_Cint ?Za ?cycTIiso_vchar. + pose theta := zeta - zeta^*%CF. + have Ztheta: theta \in 'Z[S1, HU^#] by apply: seqInd_sub_aut_zchar. + have: '[tau phi, tau theta] != 0. + rewrite Dade_isometry //; last by rewrite (cfun_onS _ (zchar_on Ztheta)). + rewrite cfdotBl !cfdotBr ?omuS1 ?cfAut_seqInd // subr0 add0r oppr_eq0. + rewrite irrWnorm // (seqInd_conjC_ortho _ _ _ S1zeta) ?mFT_odd //. + by rewrite subr0 oner_eq0. + rewrite cfnorm_eq0 Dchi; apply: contraNneq => ->; rewrite addr0 defX. + rewrite -Dtau1 ?zcharD1_seqInd //. + rewrite cfdot_suml big_map big1_seq // => _ /(cycTIirrP defW)[i [j ->]]. + apply/eqP; rewrite cfdotC fmorph_eq0 cfdotZr raddfB cfdotBl. + by rewrite !o_tau1_eta ?cfAut_seqInd ?irr_aut // subrr mulr0. + have a2_ge0 i j: 0 <= a_ i j ^+ 2 by rewrite -realEsqr Creal_Cint. + have a11_0: a11 = 0. + have: ('[X] < (2 * q.-1)%:R). + rewrite (ler_lt_trans normX_le_q) // ltC_nat -subn1 mulnBr ltn_subRL. + by rewrite !mul2n -!addnn ltn_add2r odd_prime_gt2 ?mFT_odd. + apply: contraTeq => nz_a11; rewrite ler_gtF // normX ler_paddl //. + by rewrite !mulr_natl ?addr_ge0 ?ler01 ?mulrn_wge0 ?a2_ge0. + rewrite -mulr_natl -natrM ?ler_pmul ?natr_ge0 ?sqr_Cint_ge1 ?Za //. + by rewrite leC_nat leq_mul // -subn1 ltn_subRL odd_prime_gt2 ?mFT_odd. + rewrite a11_0 expr0n /= mulr0 addr0 in normX. + have a10_a01: a10 + a01 = 1. + by apply/eqP; rewrite -subr_eq0 -a00_1 -DaB1 -/a11 a11_0. + have{o_chi_eta} o_chi_eta: orthogonal chi etaW. + by apply/orthoPl=> _ /mapP[_ /(cycTIirrP defW)[i [j ->]] ->]. + have a10_0: a10 = 0. + apply: contraNeq (FTtype34_not_ortho_cycTIiso S1zeta) => nz_a10. + have a01_0: a01 = 0. + apply: contraTeq normX_le_q => nz_a01; rewrite normX ltr_geF //. + rewrite ltr_spaddr 1?mulr_gt0 ?ltr0n -?subn1 ?subn_gt0 ?prime_gt1 //. + by rewrite ltr_def sqrf_eq0 nz_a01 a2_ge0. + rewrite -ler_subl_addl -(natrB _ (prime_gt0 pr_q)) subn1 -mulr_natl. + by rewrite ler_wpmul2l ?ler0n // sqr_Cint_ge1 ?Za. + suffices <-: X = eta_col 0 by rewrite Dchi /eq_proj_eta addrC addKr. + rewrite defX sum_etaW exchange_big (bigD1 0) //= addrC. + rewrite big1 ?add0r => [|j nz_j]; first apply: eq_bigr => i _; last first. + rewrite (bigD1 0) // [a _]Da01 //= a01_0 scale0r add0r big1 // => i nz_i. + by rewrite [a _]DaB1 Da10 // Da01 // a10_a01 a00_1 subrr scale0r. + have [-> | nz_i] := eqVneq i 0; first by rewrite [a _]a00_1 scale1r. + by rewrite [a _]Da10 // (canRL (addrK _) a10_a01) a01_0 subr0 scale1r. + suffices <-: X = eta0row by rewrite Dchi /eq_proj_eta addrC addKr. + rewrite defX sum_etaW (bigD1 0) //= addrC. + rewrite big1 ?add0r => [|i nz_i]; first apply: eq_bigr => j _; last first. + rewrite (bigD1 0) // [a _]Da10 //= a10_0 scale0r add0r big1 // => j nz_j. + by rewrite [a _]DaB1 Da10 // Da01 // a10_a01 a00_1 subrr scale0r. + have [-> | nz_j] := eqVneq j 0; first by rewrite [a _]a00_1 scale1r. + by rewrite [a _]Da01 // (canRL (addKr _) a10_a01) a10_0 oppr0 add0r scale1r. +have [zeta [irr_zeta Szeta zeta1]] := FTtypeP_ref_irr maxM MtypeP. +have{zeta1} [S1zeta zeta1]: zeta \in S1 /\ zeta 1%g = q%:R. + split=> //; have [k nz_k Dzeta] := seqIndC1P Szeta. + rewrite Dzeta mem_seqInd // !inE subGcfker nz_k -defM'' lin_char_der1 //. + rewrite -mulr_natl Dzeta cfInd1 //= -(index_sdprod defM) in zeta1. + by apply/andP; rewrite irr_char -(mulfI _ zeta1) ?neq0CG. +have{Szeta} ltpq: (p < q)%N. + rewrite ltn_neqAle neq_pq leqNgt /=. + apply: contra (FTtype34_not_ortho_cycTIiso S1zeta) => ltqp. + case/(FTtype345_Dade_bridge0 _ MtypeP): Szeta => // chi [-> _ _ o_chi_eta]. + rewrite /eq_proj_eta addrC addKr (orthogonal_oppl chi). + by apply/orthoPl=> _ /mapP[_ /(cycTIirrP defW)[i [j ->]] ->]. +suffices galM: typeP_Galois MtypeP. + have [_ [_ _ _ [/= cycUbar _ _]]] := typeP_Galois_P maxM notMtype5 galM. + have{cycUbar} cycUbar: cyclic (U / U') by rewrite -defU' -defC. + have nilU: nilpotent U by have [_ []] := MtypeP. + case/orP: Mtype34 => // /(compl_of_typeIV maxM MtypeP)[_ /negP[]]. + exact/cyclic_abelian/cyclic_nilpotent_quo_der1_cyclic. +apply: contraLR ltpq => gal'M; rewrite -leqNgt (leq_trans _ (leq_pred _)) //. +have [_ _ _] := typeP_nonGalois_characters maxM notMtype5 gal'M. +case: (_ gal'M) => H1 /= [_ _ nH1U _ []]; set a := #|U : _| => a_gt1. +rewrite def_p -/q -defU' -defS2 => a_dv_p1 cycUhat _. +set irr_qa := [pred lambda in irr M | lambda 1%g == (q * a)%:R] => S2_qa. +have{S2_qa}/hasP[lambda S2lambda /andP[irr_lambda /eqP-lambda1]]: has irr_qa S2. + have [a2_dv_pu] := S2_qa; rewrite has_count; apply: leq_trans. + rewrite -(@ltn_pmul2r (a ^ 2 * #|C|)); last first. + by rewrite !muln_gt0 (ltnW a_gt1) cardG_gt0. + by rewrite mul0n divnK // muln_gt0 cardG_gt0 -subn1 subn_gt0 prime_gt1. +have{nH1U cycUhat} a_dv_u: a %| u. + rewrite /u card_quotient ?normal_norm // indexgS // defU'. + rewrite der1_min ?cyclic_abelian // normsI ?normG //. + by rewrite (subset_trans nH1U) // astab_norm. +pose j := j1; pose psi := mu_ j - (u %/ a)%:R *: lambda. +have /andP/=[red_muj S2muj]: mu_ j \in [predD S2 & irr M]. + by rewrite memS2red image_f. +have S2psi: psi \in 'Z[S2, M^#]. + rewrite zcharD1E rpredB ?rpredZnat ?mem_zchar //=. + by rewrite !cfunE mu_1 // lambda1 -natrM mulnCA divnK ?subrr. +pose phi := tau (mu_ 0 - zeta). +have o_phi_psi: '[phi, tau psi] = 0. + have Apsi: psi \in 'CF(M, 'A(M)) by rewrite defA (zcharD1_seqInd_on _ S2psi). + have [Tzeta Tlambda] := (seqInd_subT S1zeta, seqInd_subT S2lambda). + rewrite Dade_isometry ?A0bridge0 ?(cfun_onS (FTsupp_sub0 M)) //. + rewrite cfdotBl !cfdotBr !cfdotZr cfdot_prTIred eq_sym (negPf nz_j1) add0r. + rewrite !(seqInd_ortho _ Tzeta) ?Tmu ?(memPnC (sS1S2' S1zeta)) // add0r. + rewrite (seqInd_ortho _ (Tmu 0)) ?(memPnC (prTIred_not_irr _ _)) // !mulr0. + by rewrite subrr. +have [tau2 coh_tau2] := cohS2; have [[_ Ztau2] Dtau2] := coh_tau2. +have ua_1: (u %/ a)%:R * `|'[phi, tau2 lambda]| == 1. + rewrite -normr_nat -normrM mulr_natl -!raddfMn -[_ *+ _](subrK (mu_ j)) /=. + rewrite -opprB addrC raddfB cfdotBr -scaler_nat (Dtau2 _ S2psi) o_phi_psi. + case: (FTtypeP_coherent_TIred _ coh_tau2 _ S2lambda S2muj) => // -[b k] -> _. + rewrite -/(eta_col k) cfdotZr rmorph_sign subr0 normrMsign. + rewrite -[phi](subrK eta0row) cfdotDl cfdot_sumr big1 => [|j' _]; last first. + by rewrite (orthoPl (bridgeS1 _ _)) ?map_f ?mem_irr. + rewrite add0r cfdotC norm_conjC cfdot_sumr (bigD1 k) //= proj_col_eta eqxx. + by rewrite big1 ?addr0 ?normr1 // => i k'i; rewrite proj_col_eta (negPf k'i). +have Du: u = a. + apply/eqP; rewrite -[a]mul1n eqn_mul ?(ltnW a_gt1) // -eqC_nat. + move: ua_1; rewrite Cnat_mul_eq1 ?rpred_nat //; first by case/andP. + rewrite Cnat_norm_Cint ?Cint_cfdot_vchar //; last by rewrite Ztau2 ?mem_zchar. + rewrite Dade_vchar // zchar_split A0bridge0 //. + by rewrite rpredB ?char_vchar ?prTIred_char ?irrWchar. +have lequ: (q <= u)%N. + have u1_gt0: (0 < u.-1)%N by rewrite -subn1 subn_gt0 Du. + rewrite (leq_trans _ (leq_pred u)) // dvdn_leq //. + suffices ->: q = #|W1 / C|%g by apply: Frobenius_dvd_ker1 frobUW1bar. + apply/card_isog/quotient_isog; first by have [] := joing_subP nC_UW1. + by rewrite setIAC (coprime_TIg coUq) setI1g. +by rewrite (leq_trans lequ) // Du dvdn_leq // -subn1 subn_gt0 prime_gt1. +Qed. + +End Eleven. diff --git a/mathcomp/odd_order/PFsection12.v b/mathcomp/odd_order/PFsection12.v new file mode 100644 index 0000000..2b3a3e1 --- /dev/null +++ b/mathcomp/odd_order/PFsection12.v @@ -0,0 +1,1371 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg finset center. +Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +Require Import gfunctor gproduct cyclic commutator gseries nilpotent pgroup. +Require Import sylow hall abelian maximal frobenius. +Require Import matrix mxalgebra mxpoly mxrepresentation mxabelem vector. +Require Import falgebra fieldext finfield. +Require Import BGsection1 BGsection2 BGsection3 BGsection4 BGsection7. +Require Import BGsection14 BGsection15 BGsection16. +Require Import ssrnum ssrint algC cyclotomic algnum. +Require Import classfun character inertia vcharacter. +Require Import PFsection1 PFsection2 PFsection3 PFsection4 PFsection5. +Require Import PFsection6 PFsection7 PFsection8 PFsection9 PFsection10. +Require Import PFsection11. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory FinRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Section PFTwelve. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types (p q : nat) (x y z : gT). +Implicit Types H K L M N P Q R S T U V W : {group gT}. + +Section Twelve2. + +(* Hypothesis 12.1 *) +Variable L : {group gT}. + +Hypotheses (maxL : L \in 'M) (Ltype1 : FTtype L == 1%N). + +Local Notation "` 'L'" := (gval L) (at level 0, only parsing) : group_scope. +Local Notation H := `L`_\F%G. +Local Notation "` 'H'" := `L`_\F (at level 0) : group_scope. + +Let nsHL : H <| L. Proof. exact: gFnormal. Qed. +Let calS := seqIndD H L H 1%G. +Let tau := FT_Dade maxL. +Let S_ (chi : 'CF(L)) := [set i in irr_constt chi]. +Let calX : {set Iirr L} := Iirr_kerD L H 1%g. +Let calI := [seq 'chi_i | i in calX]. + +(* This does not actually use the Ltype1 assumption. *) +Lemma FTtype1_ref_irr : exists2 phi, phi \in calS & phi 1%g = #|L : H|%:R. +Proof. +have [solH ntH] := (nilpotent_sol (Fcore_nil L), mmax_Fcore_neq1 maxL). +have [s Ls nzs] := solvable_has_lin_char ntH solH. +exists ('Ind 'chi_s); last by rewrite cfInd1 ?gFsub // lin_char1 ?mulr1. +by rewrite mem_seqInd ?gFnormal ?normal1 // !inE sub1G subGcfker -irr_eq1 nzs. +Qed. + +Let mem_calI i : i \in calX -> 'chi_i \in calI. +Proof. by move=> i_Iirr; apply/imageP; exists i. Qed. + +Lemma FTtype1_irrP i : + reflect (exists2 chi, chi \in calS & i \in S_ chi) (i \in calX). +Proof. +have [sHL nHL] := andP nsHL; rewrite !inE sub1G andbT. +apply/(iffP idP) => [kerH'i | [_ /seqIndC1P[t nz_t ->]]]; last first. + by rewrite inE => /sub_cfker_constt_Ind_irr <-; rewrite ?subGcfker. +have [t] := constt_cfRes_irr H i; rewrite -constt_Ind_Res => tLi. +rewrite -(sub_cfker_constt_Ind_irr tLi) // in kerH'i. +suffices: 'Ind 'chi_t \in calS by exists ('Ind 'chi_t); rewrite // inE. +by rewrite mem_seqInd ?normal1 // !inE sub1G kerH'i. +Qed. + +Lemma FTtype1_irr_partition : + partition [set Si in [seq S_ chi | chi <- calS]] calX. +Proof. +apply/and3P; split; last 1 first. +- rewrite inE; apply/mapP=> [[chi Schi /esym/setP S_0]]. + have /eqP[] := seqInd_neq0 nsHL Schi. + rewrite [chi]cfun_sum_constt big1 // => i chi_i. + by have:= S_0 i; rewrite inE chi_i inE. +- apply/eqP/setP=> i; apply/bigcupP/FTtype1_irrP=> [[S_chi] | [chi Schi Si]]. + by rewrite inE => /mapP[chi Schi ->]; exists chi. + by exists (S_ chi); rewrite // inE map_f. +apply/trivIsetP=> S_chi1 S_chi2. +rewrite !inE => /mapP[chi1 Schi1 ->] /mapP[chi2 Schi2 ->] {S_chi1 S_chi2}chi2'1. +apply/pred0P=> i; rewrite /= !inE; apply/andP=> [[chi1_i chi2_i]]. +suffices: '['chi_i] == 0 by rewrite cfnorm_irr oner_eq0. +rewrite (constt_ortho_char (seqInd_char Schi1) (seqInd_char Schi2)) //. +by rewrite (seqInd_ortho _ Schi1 Schi2) // (contraNneq _ chi2'1) // => ->. +Qed. + +(* This is Peterfalvi (12.2)(a), first part *) +Lemma FTtype1_seqInd_facts chi : + chi \in calS -> + [/\ chi = \sum_(i in S_ chi) 'chi_i, + constant [seq 'chi_i 1%g | i in S_ chi] + & {in S_ chi, forall i, 'chi_i \in 'CF(L, 1%g |: 'A(L))}]. +Proof. +move=> calS_chi; have [t nz_t Dchi] := seqIndC1P calS_chi. +pose T := 'I_L['chi_t]%g. +have sTL: T \subset L by apply: Inertia_sub. +have sHT: H \subset T by apply/sub_Inertia/gFsub. +have sHL: H \subset L by apply: normal_sub. +have hallH: Hall T H := pHall_Hall (pHall_subl sHT sTL (Fcore_Hall L)). +have [U [LtypeF _]] := FTtypeP _ maxL Ltype1. +have [[_ _ sdHU] [U1 inertU1] _] := LtypeF. +have defT: H ><| 'I_U['chi_t] = T := sdprod_modl sdHU (sub_inertia 'chi_t). +have abTbar : abelian (T / H). + have [_ _ /(_ _ _ inertU1 nz_t)sItU1] := typeF_context LtypeF. + by rewrite -(isog_abelian (sdprod_isog defT)) (abelianS sItU1); case: inertU1. +have [DtL _ X_1] := cfInd_Hall_central_Inertia nsHL abTbar hallH. +have Dchi_sum : chi = \sum_(i in S_ chi) 'chi_i. + by rewrite {1}Dchi DtL -Dchi; apply: eq_bigl => i; rewrite !inE. +have lichi : constant [seq 'chi_i 1%g | i in S_ chi]. + pose c := #|L : T|%:R * 'chi_t 1%g; apply: (@all_pred1_constant _ c). + by apply/allP=> _ /imageP[s tLs ->] /=; rewrite inE Dchi in tLs; rewrite X_1. +split=> // j Schi_j /=; apply/cfun_onP=> y A'y. +have [Ly | /cfun0->//] := boolP (y \in L). +have CHy1: 'C_H[y] = 1%g. + apply: contraNeq A'y => /trivgPn[z /setIP[Hz cyz] ntz]. + rewrite !inE -implyNb; apply/implyP=> nty; apply/bigcupP. + rewrite FTsupp1_type1 Ltype1 //=; exists z; first by rewrite !inE ntz. + by rewrite 3!inE nty Ly cent1C. +have: j \in calX by apply/FTtype1_irrP; exists chi. +by rewrite !inE => /andP[/not_in_ker_char0->]. +Qed. + +(* This is Peterfalvi (12.2)(a), second part. *) +Lemma FPtype1_irr_isometry : + {in 'Z[calI, L^#], isometry tau, to 'Z[irr G, G^#]}. +Proof. +apply: (sub_iso_to _ _ (Dade_Zisometry _)) => // phi. +rewrite zcharD1E => /andP[S_phi phi1_0]. +have /subsetD1P[_ /setU1K <-] := FTsupp_sub L; rewrite zcharD1 {}phi1_0 andbT. +apply: zchar_trans_on phi S_phi => _ /imageP[i /FTtype1_irrP[j calSj Sj_i] ->]. +by rewrite zchar_split irr_vchar; have [_ _ ->] := FTtype1_seqInd_facts calSj. +Qed. + +Lemma FPtype1_irr_subcoherent : + {R : 'CF(L) -> seq 'CF(G) | subcoherent calI tau R}. +Proof. +apply: irr_subcoherent; last exact: FPtype1_irr_isometry. + have UcalI: uniq calI by apply/dinjectiveP; apply: in2W irr_inj. + split=> // _ /imageP[i Ii ->]; rewrite !inE in Ii; first exact: mem_irr. + by apply/imageP; exists (conjC_Iirr i); rewrite ?inE conjC_IirrE ?cfker_aut. +apply/hasPn=> psi; case/imageP => i; rewrite !inE => /andP[kerH'i _] ->. +rewrite /cfReal odd_eq_conj_irr1 ?mFT_odd // irr_eq1 -subGcfker. +by apply: contra kerH'i; apply: subset_trans; apply: gFsub. +Qed. +Local Notation R1gen := FPtype1_irr_subcoherent. + +(* This is Peterfalvi (12.2)(b). *) +Lemma FPtype1_subcoherent (R1 := sval R1gen) : + {R : 'CF(L) -> seq 'CF(G) | + [/\ subcoherent calS tau R, + {in Iirr_kerD L H 1%G, forall i (phi := 'chi_i), + [/\ orthonormal (R1 phi), + size (R1 phi) = 2 + & tau (phi - phi^*%CF) = \sum_(mu <- R1 phi) mu]} + & forall chi, R chi = flatten [seq R1 'chi_i | i in S_ chi]]}. +Proof. +have nrS: ~~ has cfReal calS by apply: seqInd_notReal; rewrite ?mFT_odd. +have U_S: uniq calS by apply: seqInd_uniq. +have ccS: conjC_closed calS by apply: cfAut_seqInd. +have conjCS: cfConjC_subset calS (seqIndD H L H 1) by split. +case: R1gen @R1 => /= R1 subc1. +have [[chi_char nrI ccI] tau_iso oI h1 hortho] := subc1. +pose R chi := flatten [seq R1 'chi_i | i in S_ chi]. +have memI phi i: phi \in calS -> i \in S_ phi -> 'chi_i \in calI. + by move=> Sphi Sphi_i; apply/image_f/FTtype1_irrP; exists phi. +have aux phi psi i j mu nu: + phi \in calS -> psi \in calS -> i \in S_ phi -> j \in S_ psi -> + mu \in R1 'chi_i -> nu \in R1 'chi_j -> + orthogonal 'chi_i ('chi_j :: ('chi_j)^*%CF) -> '[mu, nu] = 0. +- move=> Sphi Spsi Sphi_i Spsi_j R1i_mu R1i_nu o_ij. + apply: orthogonalP R1i_mu R1i_nu. + by apply: hortho o_ij; [apply: memI Spsi Spsi_j | apply: memI Sphi Sphi_i]. +exists R; split => //= => [| i Ii]; last first. + have mem_i := mem_calI Ii; have{h1} [Zirr oR1 tau_im] := h1 _ mem_i. + split=> //; apply/eqP; rewrite -eqC_nat -cfnorm_orthonormal // -{}tau_im. + have ?: 'chi_i - ('chi_i)^*%CF \in 'Z[calI, L^#]. + have hchi : 'chi_i \in 'Z[calI, L] by rewrite mem_zchar_on // cfun_onG. + rewrite sub_aut_zchar ?cfAut_zchar // => _ /mapP[j _ ->]; exact: irr_vchar. + have [-> // _] := tau_iso; rewrite cfnormBd ?cfnorm_conjC ?cfnorm_irr //. + by have [_ ->] := pairwise_orthogonalP oI; rewrite ?ccI // eq_sym (hasPn nrI). +have calS_portho : pairwise_orthogonal calS by apply: seqInd_orthogonal. +have calS_char : {subset calS <= character} by apply: seqInd_char. +have calS_chi_ortho : + {in calS &, forall phi psi i j, + i \in irr_constt phi -> j \in irr_constt psi -> + '[phi, psi] = 0 -> '['chi_i, 'chi_j] = 0}. +- by move=> phi psi Sphi Spsi /= i j; apply: constt_ortho_char; apply/calS_char. +have ZisoS_tau: {in 'Z[calS, L^#], isometry tau, to 'Z[irr G, G^#]}. + apply: (sub_iso_to _ _ (Dade_Zisometry _)) => // phi. + have /subsetD1P[_ /setU1K <-] := FTsupp_sub L. + rewrite zcharD1E zcharD1 => /andP[S_phi ->]; rewrite andbT. + apply: zchar_trans_on phi S_phi => psi calS_psi. + have [Dpsi _ hCF] := FTtype1_seqInd_facts calS_psi. + by rewrite zchar_split (seqInd_vcharW calS_psi) /= Dpsi rpred_sum. +split=> {ZisoS_tau}//= [phi calS_phi | phi psi calS_phi calS_psi]. + rewrite /R /[seq _ | i in _]; set e := enum _; have: uniq e := enum_uniq _. + have: all (mem (S_ phi)) e by apply/allP=> i; rewrite mem_enum. + have ->: phi - phi^*%CF = \sum_(i <- e) ('chi_i - ('chi_i)^*%CF). + rewrite big_filter sumrB -rmorph_sum. + by have [<-] := FTtype1_seqInd_facts calS_phi. + elim: e => /= [_ _ | i e IHe /andP[Si Se] /andP[e'i Ue]]. + by rewrite !big_nil /tau linear0. + rewrite big_cons [tau _]linearD big_cat /= -/tau orthonormal_cat. + have{IHe Ue} [/allP Ze -> ->] := IHe Se Ue. + have{h1} /h1[/allP Z_R1i -> -> /=] := memI _ _ calS_phi Si. + split=> //; first by apply/allP; rewrite all_cat Z_R1i. + apply/orthogonalP=> mu nu R1i_mu /flatten_mapP[j e_j R1j_nu]. + have /= Sj := allP Se j e_j; apply: (aux phi phi i j) => //. + rewrite /orthogonal /= !andbT !cfdot_irr mulrb ifN_eqC ?(memPn e'i) ?eqxx //=. + rewrite !inE in Si Sj; rewrite -conjC_IirrE; set k := conjC_Iirr j. + rewrite (calS_chi_ortho phi phi^*%CF) ?calS_char ?ccS //. + by rewrite irr_consttE conjC_IirrE cfdot_conjC fmorph_eq0. + by rewrite (seqInd_conjC_ortho _ _ _ calS_phi) ?mFT_odd. +case/andP=> /and3P[/eqP opsi_phi /eqP opsi_phiC _] _; apply/orthogonalP. +move=> nu mu /flatten_imageP[j Spsi_j R1j_nu] /flatten_imageP[i Sphi_i R1i_mu]. +apply: (aux psi phi j i) => //; rewrite /orthogonal /= !andbT -conjC_IirrE. +rewrite !inE in Sphi_i Spsi_j; rewrite (calS_chi_ortho psi phi) ?calS_char //. +rewrite (calS_chi_ortho psi phi^*%CF) ?calS_char ?ccS ?eqxx //. +by rewrite irr_consttE conjC_IirrE cfdot_conjC fmorph_eq0. +Qed. + +End Twelve2. + +Local Notation R1gen := FPtype1_irr_subcoherent. +Local Notation Rgen := FPtype1_subcoherent. + +(* This is Peterfalvi (12.3) *) +Lemma FTtype1_seqInd_ortho L1 L2 (maxL1 : L1 \in 'M) (maxL2 : L2 \in 'M) + (L1type1 : FTtype L1 == 1%N) (L2type1 : FTtype L2 == 1%N) + (H1 := L1`_\F%G) (H2 := L2`_\F%G) + (calS1 := seqIndD H1 L1 H1 1) (calS2 := seqIndD H2 L2 H2 1) + (R1 := sval (Rgen maxL1 L1type1)) (R2 := sval (Rgen maxL2 L2type1)) : + gval L2 \notin L1 :^: G -> + {in calS1 & calS2, forall chi1 chi2, orthogonal (R1 chi1) (R2 chi2)}. +Proof. +move=> notL1G_L2; without loss{notL1G_L2} disjointA1A: + L1 L2 maxL1 maxL2 L1type1 L2type1 @H1 @H2 @calS1 @calS2 @R1 @R2 / + [disjoint 'A1~(L2) & 'A~(L1)]. +- move=> IH_L; have [_ _] := FT_Dade_support_disjoint maxL1 maxL2 notL1G_L2. + by case=> /IH_L-oS12 chi1 chi2 *; first rewrite orthogonal_sym; apply: oS12. +case: (Rgen _ _) @R1 => /= R1; set R1' := sval _ => [[subcoh1 hR1' defR1]]. +case: (Rgen _ _) @R2 => /= R2; set R2' := sval _ => [[subcoh2 hR2' defR2]]. +pose tau1 := FT_Dade maxL1; pose tau2 := FT_Dade maxL2. +move=> chi1 chi2 calS1_chi1 calS2_chi2. +have [_ _ _ /(_ chi1 calS1_chi1)[Z_R1 o1R1 dtau1_chi1] _] := subcoh1. +have{o1R1} [uR1 oR1] := orthonormalP o1R1. +apply/orthogonalP=> a b R1a R2b; pose psi2 := tau2 (chi2 - chi2^*%CF). +have Z1a: a \in dirr G by rewrite dirrE Z_R1 //= oR1 ?eqxx. +suffices{b R2b}: '[a, psi2] == 0. + apply: contraTeq => nz_ab; rewrite /psi2 /tau2. + have [_ _ _ /(_ chi2 calS2_chi2)[Z_R2 o1R2 ->] _] := subcoh2. + suffices [e ->]: {e | a = if e then - b else b}. + rewrite -scaler_sign cfdotC cfdotZr -cfdotZl scaler_sumr. + by rewrite cfproj_sum_orthonormal // conjCK signr_eq0. + have [_ oR2] := orthonormalP o1R2. + have Z1b: b \in dirr G by rewrite dirrE Z_R2 //= oR2 ?eqxx. + move/eqP: nz_ab; rewrite cfdot_dirr //. + by do 2?[case: eqP => [-> | _]]; [exists true | exists false | ]. +have [chi1D _ Achi1] := FTtype1_seqInd_facts maxL1 L1type1 calS1_chi1. +pose S_chi1 := [set i0 in irr_constt chi1]. +pose bchi i := 'chi[_ : {set gT}]_i - ('chi_i)^*%CF. +have [t S_chi1t et]: exists2 t, t \in S_chi1 & tau1 (bchi _ t) = a - a^*%CF. + suffices: ~~ [forall i in S_chi1, '[tau1 (bchi L1 i), a] <= 0]. + rewrite negb_forall_in => /exists_inP[i Si tau1i_a]; exists i => //. + case/dIrrP: Z1a tau1i_a => ia ->. + have [k ->]: exists k, tau1 (bchi _ i) = bchi G k. + exact: Dade_irr_sub_conjC (mem_irr _) (Achi1 i Si). + have {1}->: bchi G k = dchi (false, k) + dchi (true, conjC_Iirr k). + by rewrite /dchi !scaler_sign conjC_IirrE. + rewrite cfdotDl !cfdot_dchi addrACA -opprD subr_le0 -!natrD leC_nat. + do 2?case: (_ =P ia) => [<-|] _ //; first by rewrite /dchi scale1r. + by rewrite /dchi scaleN1r conjC_IirrE rmorphN /= cfConjCK opprK addrC. + have: '[tau1 (chi1 - chi1^*%CF), a] == 1. + rewrite /tau1 dtau1_chi1 (bigD1_seq a) //= cfdotDl cfdot_suml oR1 // eqxx. + by rewrite big1_seq ?addr0 // => xi /andP[/negPf a'xi ?]; rewrite oR1 ?a'xi. + apply: contraL => /forall_inP tau1a_le0. + rewrite (ltr_eqF (ler_lt_trans _ ltr01)) // chi1D rmorph_sum /= -/S_chi1. + rewrite -sumrB [tau1 _]linear_sum /= -/tau1 cfdot_suml. + by rewrite -oppr_ge0 -sumrN sumr_ge0 // => i /tau1a_le0; rewrite oppr_ge0. +clear Achi1 dtau1_chi1 uR1 defR1. +suffices: '[a, psi2] == - '[a, psi2] by rewrite -addr_eq0 (mulrn_eq0 _ 2). +have A1bchi2: chi2 - (chi2^*)%CF \in 'Z[calS2, 'A1(L2)]. + by rewrite FTsupp1_type1 // seqInd_sub_aut_zchar ?gFnormal. +have{t S_chi1t et} /eqP{2}->: '[a, psi2] == '[a^*%CF, psi2]. + move/zchar_on in A1bchi2; rewrite -subr_eq0 -cfdotBl. + rewrite [psi2]FT_DadeE ?(cfun_onS (FTsupp1_sub _)) // -FT_Dade1E // -et. + rewrite (cfdot_complement (Dade_cfunS _ _)) ?(cfun_onS _ (Dade_cfunS _ _)) //. + by rewrite FT_Dade_supportE FT_Dade1_supportE setTD -disjoints_subset. +rewrite -2!raddfN opprB /= cfdot_conjCl -Dade_conjC rmorphB /= cfConjCK -/tau2. +rewrite conj_Cint ?Cint_cfdot_vchar ?(Z_R1 a) // Dade_vchar //. +rewrite (zchar_onS (FTsupp1_sub _)) // (zchar_sub_irr _ A1bchi2) //. +exact: seqInd_vcharW. +Qed. + +Section Twelve_4_to_6. + +Variable L : {group gT}. +Hypothesis maxL : L \in 'M . + +Local Notation "` 'L'" := (gval L) (at level 0, only parsing) : group_scope. +Local Notation H := `L`_\F%G. +Local Notation "` 'H'" := `L`_\F (at level 0) : group_scope. +Local Notation H' := H^`(1)%G. +Local Notation "` 'H''" := `H^`(1) (at level 0) : group_scope. + +Let calS := seqIndD H L H 1%G. +Let tau := FT_Dade maxL. +Let rho := invDade (FT_DadeF_hyp maxL). + +Section Twelve_4_5. + +Hypothesis Ltype1 : FTtype L == 1%N. + +Let R := sval (Rgen maxL Ltype1). +Let S_ (chi : 'CF(L)) := [set i in irr_constt chi]. + +(* This is Peterfalvi (12.4). *) +Lemma FTtype1_ortho_constant (psi : 'CF(G)) x : + {in calS, forall phi, orthogonal psi (R phi)} -> x \in L :\: H -> + {in x *: H, forall y, psi y = psi x}%g. +Proof. +move=> opsiR /setDP[Lx H'x]; pose Rpsi := 'Res[L] psi. +have nsHL: H <| L := gFnormal _ _; have [sHL _] := andP nsHL. +have [U [[[_ _ sdHU] [U1 inertU1] _] _]] := FTtypeP 1 maxL Ltype1. +have /= [_ _ TIsub]:= FTtypeI_II_facts maxL Ltype1 sdHU. +pose ddL := FT_Dade_hyp maxL. +have A1Hdef : 'A1(L) = H^# by apply: FTsupp1_type1. +have dot_irr xi j : xi \in calS -> j \in S_ xi -> '['chi_j, xi] = 1. + move=> xi_calS Sj. + have -> : xi = \sum_(i <- enum (S_ xi)) 'chi_i. + by rewrite big_filter; have [] := FTtype1_seqInd_facts maxL Ltype1 xi_calS. + rewrite (bigD1_seq j) ?mem_enum ?enum_uniq //= cfdotDr cfdot_sumr cfnorm_irr. + by rewrite big1 ?addr0 // => k i'k; rewrite cfdot_irr eq_sym (negPf i'k). +have {dot_irr} supp12B y xi j1 j2 : xi \in calS -> j1 \in S_ xi -> + j2 \in S_ xi -> y \notin ('A(L) :\: H^#) -> ('chi_j1 - 'chi_j2) y = 0. +- move=> calS_xi Sj1 Sj2 yADHn. + have [xiD xi_cst sup_xi] := FTtype1_seqInd_facts maxL Ltype1 calS_xi. + have [Hy | H'y] := boolP (y \in H); last first. + suffices /cfun_on0->: y \notin 1%g |: 'A(L) by rewrite ?rpredB ?sup_xi. + by rewrite !inE negb_or negb_and (group1_contra H'y) ?H'y in yADHn *. + have [s _ xiIndD] := seqIndP calS_xi. + pose sum_sL := \sum_(xi_z <- ('chi_s ^: L)%CF) xi_z. + suffices Dxi: {in S_ xi, forall i, 'chi_i y = sum_sL y}. + by rewrite !cfunE !Dxi ?subrr. + move=> k Sk; pose phiH := 'Res[H] 'chi_k. + transitivity (phiH y); first by rewrite cfResE ?normal_sub. + have phiH_s_1: '[phiH, 'chi_s] = 1 by rewrite cfdot_Res_l -xiIndD dot_irr. + have phiH_s: s \in irr_constt phiH by rewrite irr_consttE phiH_s_1 oner_eq0. + by rewrite [phiH](Clifford_Res_sum_cfclass _ phiH_s) // phiH_s_1 scale1r. +have {supp12B} oResD xi i1 i2 : xi \in calS -> i1 \in S_ xi -> i2 \in S_ xi -> + '['Res[L] psi, 'chi_i1 - 'chi_i2] = 0. +- move=> calS_xi Si1 Si2; rewrite cfdotC Frobenius_reciprocity -cfdotC. + case: (altP (i1 =P i2))=> [-> | d12]; first by rewrite subrr linear0 cfdot0r. + have {supp12B} supp12B y: y \notin 'A(L) :\: H^# -> ('chi_i1 - 'chi_i2) y = 0. + exact: (supp12B _ xi _ _ calS_xi). + case: (FTtype1_seqInd_facts maxL Ltype1 calS_xi) => _ cst1 supA. + move/(_ _ Si1): (supA) => /cfun_onP s1; case/(constantP 0): (cst1) => [n]. + move/all_pred1P /allP => nseqD; move/(_ _ Si2): (supA) => /cfun_onP s2. + have nchi11: 'chi_i1 1%g = n by apply/eqP/nseqD/image_f. + have{nseqD} nchi12: 'chi_i2 1%g = n by apply/eqP/nseqD/image_f. + have i12_1: 'chi_i1 1%g == 'chi_i2 1%g by rewrite nchi11 nchi12. + have sH1A: H^# \subset 'A(L) by rewrite Fcore_sub_FTsupp. + have nzAH: 'A(L) :\: H^# != set0. + apply: contra d12 => /eqP tADH; apply/eqP; apply: irr_inj; apply/cfunP=> w. + apply/eqP; rewrite -subr_eq0; have := supp12B w; rewrite !cfunE => -> //. + by rewrite tADH in_set0. + have{nzAH} tiH: normedTI ('A(L) :\: H^#) G L by rewrite -A1Hdef TIsub ?A1Hdef. + have{supp12B} supp12B : 'chi_i1 - 'chi_i2 \in 'CF(L, 'A(L) :\: H^#). + by apply/cfun_onP; apply: supp12B. + have [_ /subsetIP[_ nAHL] _] := normedTI_P tiH. + pose tau1 := restr_Dade ddL (subsetDl _ _) nAHL. + have tauInd: {in 'CF(L, 'A(L) :\: H^#), tau1 =1 'Ind} := Dade_Ind _ tiH. + rewrite -{}tauInd // [tau1 _]restr_DadeE {tau1 nAHL}//. + suffices Rtau12: Dade ddL ('chi_i1 - 'chi_i2) \in 'Z[R xi]. + by rewrite (span_orthogonal (opsiR xi _)) ?memv_span1 ?(zchar_span Rtau12). + case: (Rgen _ _) @R => rR [scohS]; case: (R1gen _ _) => /= R1 scohI ? DrR. + rewrite -/calS in scohS; set calI := image _ _ in scohI. + have [Ii1 Ii2]: 'chi_i1 \in calI /\ 'chi_i2 \in calI. + by split; apply/image_f/FTtype1_irrP; exists xi. + have [calI2 [I2i1 I2i2 sI2I] []] := pair_degree_coherence scohI Ii1 Ii2 i12_1. + move=> tau2 cohI2; have [_ <-] := cohI2; last first. + by rewrite zcharD1E rpredB ?mem_zchar // 2!cfunE subr_eq0. + suffices R_I2 j: j \in S_ xi -> 'chi_j \in calI2 -> tau2 'chi_j \in 'Z[rR xi]. + by rewrite raddfB rpredB ?R_I2. + move=> Sj /(mem_coherent_sum_subseq scohI sI2I cohI2)[e R1e ->]. + rewrite DrR big_seq rpred_sum // => phi /(mem_subseq R1e) R1phi. + by apply/mem_zchar/flatten_imageP; exists j. +suffices ResL: {in x *: H, forall y, Rpsi y = Rpsi x}%g. + move=> w xHw; case/lcosetP: xHw (ResL w xHw) => h Hh -> {w}. + by rewrite !cfResE ?subsetT ?groupM // ?(subsetP sHL). +move=> _ /lcosetP[h Hh ->] /=; rewrite (cfun_sum_cfdot Rpsi). +pose calX := Iirr_kerD L H 1%g; rewrite (bigID (mem calX) xpredT) /= !cfunE. +set sumX := \sum_(i in _) _; suffices HsumX: sumX \in 'CF(L, H). + rewrite !(cfun_on0 HsumX) ?groupMr // !sum_cfunE. + rewrite !add0r; apply: eq_bigr => i;rewrite !inE sub1G andbT negbK => kerHi. + by rewrite !cfunE cfkerMr ?(subsetP kerHi). +rewrite [sumX](set_partition_big _ (FTtype1_irr_partition L)) /=. +apply: rpred_sum => A; rewrite inE => /mapP[xi calS_xi defA]. +have [-> | [j Achij]] := set_0Vmem A; first by rewrite big_set0 rpred0. +suffices ->: \sum_(i in A) '[Rpsi, 'chi_i] *: 'chi_i = '[Rpsi, 'chi_j] *: xi. + by rewrite rpredZ // (seqInd_on _ calS_xi). +have [-> _ _] := FTtype1_seqInd_facts maxL Ltype1 calS_xi; rewrite -defA. +rewrite scaler_sumr; apply: eq_bigr => i Ai; congr (_ *: _); apply/eqP. +by rewrite -subr_eq0 -cfdotBr (oResD xi) /S_ -?defA. +Qed. + +(* This is Peterfalvi (12.5) *) +Lemma FtypeI_invDade_ortho_constant (psi : 'CF(G)) : + {in calS, forall phi, orthogonal psi (R phi)} -> + {in H :\: H' &, forall x y, rho psi x = rho psi y}. +Proof. +have [nsH'H nsHL]: H' <| H /\ H <| L by rewrite !gFnormal. +have [[sH'H _] [sHL _]] := (andP nsH'H, andP nsHL). +case: (Rgen _ _) @R => /= rR [scohS _ _] opsiR; set rpsi := rho psi. +have{rR scohS opsiR} o_rpsi_S xi1 xi2: + xi1 \in calS -> xi2 \in calS -> xi1 1%g = xi2 1%g -> '[rpsi, xi1 - xi2] = 0. +- move=> Sxi1 Sxi2 /eqP deg12. + have [calS2 [S2xi1 S2xi2]] := pair_degree_coherence scohS Sxi1 Sxi2 deg12. + move=> ccsS2S [tau2 cohS2]; have [[_ Dtau2] [_ sS2S _]] := (cohS2, ccsS2S). + have{deg12} L1xi12: (xi1 - xi2) 1%g == 0 by rewrite !cfunE subr_eq0. + have{ccsS2S cohS2} tau2E := mem_coherent_sum_subseq scohS ccsS2S cohS2. + have o_psi_tau2 xi: xi \in calS2 -> '[psi, tau2 xi] = 0. + move=> S2xi; have [e /mem_subseq Re ->] := tau2E xi S2xi. + by rewrite cfdot_sumr big1_seq // => _ /Re/orthoPl->; rewrite ?opsiR ?sS2S. + have A1xi12: xi1 - xi2 \in 'CF(L, H^#). + by rewrite (@zchar_on _ _ calS) ?zcharD1 ?rpredB ?seqInd_zchar. + rewrite cfdotC -invDade_reciprocity // -cfdotC. + rewrite FT_DadeF_E -?FT_DadeE ?(cfun_onS (Fcore_sub_FTsupp maxL)) //. + rewrite -Dtau2; last by rewrite zcharD1E rpredB ?mem_zchar. + by rewrite !raddfB /= !o_psi_tau2 ?subrr. +pose P_ i : {set Iirr H} := [set j in irr_constt ('Ind[H, H'] 'chi_i)]. +pose P : {set {set Iirr H}} := [set P_ i | i : Iirr H']. +have tiP: trivIset P. + apply/trivIsetP=> _ _ /imsetP[i1 _ ->] /imsetP[i2 _ ->] chi2'1. + apply/pred0P=> j; rewrite !inE; apply: contraNF chi2'1 => /andP[i1Hj i2Hj]. + case: ifP (cfclass_Ind_cases i1 i2 nsH'H) => _; first by rewrite /P_ => ->. + have NiH i: 'Ind[H,H'] 'chi_i \is a character by rewrite cfInd_char ?irr_char. + case/(constt_ortho_char (NiH i1) (NiH i2) i1Hj i2Hj)/eqP/idPn. + by rewrite cfnorm_irr oner_eq0. +have coverP: cover P =i predT. + move=> j; apply/bigcupP; have [i jH'i] := constt_cfRes_irr H' j. + by exists (P_ i); [apply: mem_imset | rewrite inE constt_Ind_Res]. +have /(all_sig_cond 0)[lambda lambdaP] A: A \in P -> {i | A = P_ i}. + by case/imsetP/sig2_eqW=> i; exists i. +pose theta A : Iirr H := odflt 0 [pick j in A :\ 0]; pose psiH := 'Res[H] rpsi. +pose a_ A := '[psiH, 'chi_(theta A)] / '['Ind 'chi_(lambda A), 'chi_(theta A)]. +pose a := '[psiH, 1 - 'chi_(theta (pblock P 0))]. +suffices Da: {in H :\: H', rpsi =1 (fun=> a)} by move=> /= *; rewrite !Da. +suffices DpsiH: psiH = \sum_(A in P) a_ A *: 'Ind 'chi_(lambda A) + a%:A. + move=> x /setDP[Hx notH'x]; transitivity (psiH x); first by rewrite cfResE. + rewrite DpsiH !cfunE sum_cfunE cfun1E Hx mulr1 big1 ?add0r // => A _. + by rewrite cfunE (cfun_onP (cfInd_normal _ _)) ?mulr0. +apply: canRL (subrK _) _; rewrite [_ - _]cfun_sum_cfdot. +rewrite -(eq_bigl _ _ coverP) big_trivIset //=; apply: eq_bigr => A P_A. +rewrite {}/a_; set i := lambda A; set k := theta A; pose Ii := 'I_H['chi_i]%G. +have /cfInd_central_Inertia[//|e _ [DiH _ DiH_1]]: abelian (Ii / H'). + by rewrite (abelianS _ (der_abelian 0 H)) ?quotientS ?subsetIl. +have defA: A = P_ i := lambdaP A P_A. +have Ak: k \in A; last 1 [have iHk := Ak; rewrite defA inE in Ak]. + have [j iHj] := constt_cfInd_irr i sH'H. + rewrite {}/k /theta; case: pickP => [k /setDP[]//| /(_ j)/=]. + by rewrite defA !in_set iHj andbT => /negbFE/eqP <-. +have{DiH} DiH: 'Ind 'chi_i = e *: \sum_(j in A) 'chi_j. + by congr (_ = _ *: _): DiH; apply: eq_bigl => j; rewrite [in RHS]defA !inE. +rewrite {2}DiH; have{DiH} ->: e = '['Ind 'chi_i, 'chi_k]. + rewrite DiH cfdotZl cfdot_suml (bigD1 k) //= cfnorm_irr big1 ?addr0 ?mulr1 //. + by move=> j /andP[_ k'j]; rewrite cfdot_irr (negPf k'j). +rewrite scalerA scaler_sumr divfK //; apply: eq_bigr => j Aj; congr (_ *: _). +rewrite cfdotBl cfdotZl -irr0 cfdot_irr mulr_natr mulrb eq_sym. +apply/(canLR (addrK _))/(canRL (addNKr _)); rewrite addrC -cfdotBr. +have [j0 | nzj] := altP eqP; first by rewrite j0 irr0 /a -j0 (def_pblock _ P_A). +have iHj := Aj; rewrite defA inE in iHj; rewrite cfdot_Res_l linearB /=. +do [rewrite o_rpsi_S ?cfInd1 ?DiH_1 //=; apply/seqIndC1P]; first by exists j. +by exists k; rewrite // /k /theta; case: pickP => [? | /(_ j)] /setD1P[]. +Qed. + +End Twelve_4_5. + +Hypothesis frobL : [Frobenius L with kernel H]. + +Lemma FT_Frobenius_type1 : FTtype L == 1%N. +Proof. +have [E /Frobenius_of_typeF LtypeF] := existsP frobL. +by apply/idPn=> /FTtypeP_witness[]// _ _ _ _ _ /typePF_exclusion/(_ E). +Qed. +Let Ltype1 := FT_Frobenius_type1. + +Lemma FTsupp_Frobenius : 'A(L) = H^#. +Proof. +apply/eqP; rewrite eqEsubset Fcore_sub_FTsupp // andbT. +apply/bigcupsP=> y; rewrite Ltype1 FTsupp1_type1 //= => H1y. +by rewrite setSD //; have [_ _ _ ->] := Frobenius_kerP frobL. +Qed. + +(* This is Peterfalvi (12.6). *) +Lemma FT_Frobenius_coherence : {subset calS <= irr L} /\ coherent calS L^# tau. +Proof. +have irrS: {subset calS <= irr L}. + by move=> _ /seqIndC1P[s nz_s ->]; apply: irr_induced_Frobenius_ker. +split=> //; have [U [MtypeF MtypeI]] := FTtypeP 1 maxL Ltype1. +have [[ntH ntU defL] _ _] := MtypeF; have nsHL: H <| L := gFnormal _ L. +have nilH: nilpotent H := Fcore_nil L; have solH := nilpotent_sol nilH. +have frobHU: [Frobenius L = H ><| U] := set_Frobenius_compl defL frobL. +pose R := sval (Rgen maxL Ltype1). +have scohS: subcoherent calS tau R by case: (svalP (Rgen maxL Ltype1)). +have [tiH | [cHH _] | [expUdvH1 _]] := MtypeI. +- have /Sibley_coherence := And3 (mFT_odd L) nilH tiH. + case/(_ U)=> [|tau1 [IZtau1 Dtau1]]; first by left. + exists tau1; split=> // chi Schi; rewrite Dtau1 //. + by rewrite /tau Dade_Ind ?FTsupp_Frobenius ?(zcharD1_seqInd_on _ Schi). +- apply/(uniform_degree_coherence scohS)/(@all_pred1_constant _ #|L : H|%:R). + apply/allP=> _ /mapP[_ /seqIndP[s _ ->] ->] /=. + by rewrite cfInd1 ?gFsub // lin_char1 ?mulr1 //; apply/char_abelianP. +have isoHbar := quotient1_isog H. +have /(_ 1%G)[|//|[_ [p [pH _] /negP[]]]] := non_coherent_chief nsHL solH scohS. + split; rewrite ?mFT_odd ?normal1 ?sub1G -?(isog_nil isoHbar) //= joingG1. + apply/existsP; exists (U / H')%G. + rewrite Frobenius_proper_quotient ?(sol_der1_proper solH) //. + exact: char_normal_trans (der_char 1 H) nsHL. +rewrite -(isog_pgroup p isoHbar) in pH. +have [pr_p p_dv_H _] := pgroup_pdiv pH ntH. +rewrite subn1 -(index_sdprod defL). +have [-> *] := typeF_context MtypeF; last by split; rewrite ?(sdprodWY defL). +by rewrite expUdvH1 // mem_primes pr_p cardG_gt0. +Qed. + +End Twelve_4_to_6. + +Section Twelve_8_to_16. + +Variable p : nat. + +(* Equivalent reformultaion of Hypothesis (12.8), avoiding quotients. *) +Hypothesis IHp : + forall q M, (q < p)%N -> M \in 'M -> FTtype M == 1%N -> ('r_q(M) > 1)%N -> + q \in \pi(M`_\F). + +Variables M P0 : {group gT}. + +Let K := M`_\F%G. +Let K' := K^`(1)%G. +Let nsKM : K <| M. Proof. exact: gFnormal. Qed. + +Hypothesis maxM : M \in 'M. +Hypothesis Mtype1 : FTtype M == 1%N. +Hypothesis prankM : ('r_p(M) > 1)%N. +Hypothesis p'K : p^'.-group K. + +Hypothesis sylP0 : p.-Sylow(M) P0. + +(* This is Peterfalvi (12.9). *) +Lemma non_Frobenius_FTtype1_witness : + [/\ abelian P0, 'r_p(P0) = 2 + & exists2 L, L \in 'M /\ P0 \subset L`_\s + & exists2 x, x \in 'Ohm_1(P0)^# + & [/\ ~~ ('C_K[x] \subset K'), 'N(<[x]>) \subset M & ~~ ('C[x] \subset L)]]. +Proof. +have ntK: K :!=: 1%g := mmax_Fcore_neq1 maxM; have [sP0M pP0 _] := and3P sylP0. +have hallK: \pi(K).-Hall(M) K := Fcore_Hall M. +have K'p: p \notin \pi(K) by rewrite -p'groupEpi. +have K'P0: \pi(K)^'.-group P0 by rewrite (pi_pgroup pP0). +have [U hallU sP0U] := Hall_superset (mmax_sol maxM) sP0M K'P0. +have sylP0_U: p.-Sylow(U) P0 := pHall_subl sP0U (pHall_sub hallU) sylP0. +have{hallU} defM: K ><| U = M by apply/(sdprod_normal_p'HallP nsKM hallU). +have{K'P0} coKP0: coprime #|K| #|P0| by rewrite coprime_pi'. +have [/(_ _ _ sylP0_U)[abP0 rankP0] uCK _] := FTtypeI_II_facts maxM Mtype1 defM. +have{rankP0} /eqP prankP0: 'r_p(P0) == 2. + by rewrite eqn_leq -{1}rank_pgroup // rankP0 (p_rank_Sylow sylP0). +have piP0p: p \in \pi(P0) by rewrite -p_rank_gt0 prankP0. +have [L maxL sP0Ls]: exists2 L, L \in 'M & P0 \subset L`_\s. + have [DpiG _ _ _] := FT_Dade_support_partition gT. + have:= piSg (subsetT P0) piP0p; rewrite DpiG => /exists_inP[L maxL piLs_p]. + have [_ /Hall_pi hallLs _] := FTcore_facts maxL. + have [P sylP] := Sylow_exists p L`_\s; have [sPLs _] := andP sylP. + have sylP_G: p.-Sylow(G) P := subHall_Sylow hallLs piLs_p sylP. + have [y _ sP0_Py] := Sylow_subJ sylP_G (subsetT P0) pP0. + by exists (L :^ y)%G; rewrite ?mmaxJ // FTcoreJ (subset_trans sP0_Py) ?conjSg. +split=> //; exists L => //; set P1 := 'Ohm_1(P0). +have abelP1: p.-abelem P1 := Ohm1_abelem pP0 abP0. +have [pP1 abP1 _] := and3P abelP1. +have sP10: P1 \subset P0 := Ohm_sub 1 P0; have sP1M := subset_trans sP10 sP0M. +have nKP1: P1 \subset 'N(K) by rewrite (subset_trans sP1M) ?gFnorm. +have nK'P1: P1 \subset 'N(K') := char_norm_trans (der_char 1 K) nKP1. +have{coKP0} coKP1: coprime #|K| #|P1| := coprimegS sP10 coKP0. +have solK: solvable K := nilpotent_sol (Fcore_nil M). +have isoP1: P1 \isog P1 / K'. + by rewrite quotient_isog // coprime_TIg ?(coprimeSg (der_sub 1 K)). +have{ntK} ntKK': (K / K' != 1)%g. + by rewrite -subG1 quotient_sub1 ?gFnorm ?proper_subn ?(sol_der1_proper solK). +have defKK': (<<\bigcup_(xbar in (P1 / K')^#) 'C_(K / K')[xbar]>> = K / K')%g. + rewrite coprime_abelian_gen_cent1 ?coprime_morph ?quotient_norms //. + by rewrite quotient_abelian. + rewrite -(isog_cyclic isoP1) (abelem_cyclic abelP1). + by rewrite -(p_rank_abelem abelP1) p_rank_Ohm1 prankP0. +have [xb P1xb ntCKxb]: {xb | xb \in (P1 / K')^# & 'C_(K / K')[xb] != 1}%g. + apply/sig2W/exists_inP; rewrite -negb_forall_in. + apply: contra ntKK' => /eqfun_inP regKP1bar. + by rewrite -subG1 /= -defKK' gen_subG; apply/bigcupsP=> xb /regKP1bar->. +have [ntxb /morphimP[x nK'x P1x Dxb]] := setD1P P1xb. +have ntx: x != 1%g by apply: contraNneq ntxb => x1; rewrite Dxb x1 morph1. +have ntCKx: ~~ ('C_K[x] \subset K'). + rewrite -quotient_sub1 ?subIset ?gFnorm // -cent_cycle subG1 /=. + have sXP1: <[x]> \subset P1 by rewrite cycle_subG. + rewrite coprime_quotient_cent ?(coprimegS sXP1) ?(subset_trans sXP1) ?gFsub//. + by rewrite quotient_cycle ?(subsetP nK'P1) // -Dxb cent_cycle. +have{uCK} UCx: 'M('C[x]) = [set M]. + rewrite -cent_set1 uCK -?card_gt0 ?cards1 // ?sub1set ?cent_set1. + by rewrite !inE ntx (subsetP sP0U) ?(subsetP sP10). + by apply: contraNneq ntCKx => ->; rewrite sub1G. +exists x; [by rewrite !inE ntx | split=> //]. + rewrite (sub_uniq_mmax UCx) /= -?cent_cycle ?cent_sub //. + rewrite mFT_norm_proper ?cycle_eq1 //. + by rewrite mFT_sol_proper abelian_sol ?cycle_abelian. +apply: contraL (leqW (p_rankS p sP0Ls)) => /(eq_uniq_mmax UCx)-> //. +by rewrite prankP0 FTcore_type1 //= ltnS p_rank_gt0. +Qed. + +Variables (L : {group gT}) (x : gT). +Hypotheses (abP0 : abelian P0) (prankP0 : 'r_p(P0) = 2). +Hypotheses (maxL : L \in 'M) (sP0_Ls : P0 \subset L`_\s). +Hypotheses (P0_1s_x : x \in 'Ohm_1(P0)^#) (not_sCxK' : ~~ ('C_K[x] \subset K')). +Hypotheses (sNxM : 'N(<[x]>) \subset M) (not_sCxL : ~~ ('C[x] \subset L)). + +Let H := L`_\F%G. +Let nsHL : H <| L. Proof. exact: gFnormal. Qed. + +(* This is Peterfalvi (12.10). *) +Let frobL : [Frobenius L with kernel H]. +Proof. +have [sP0M pP0 _] := and3P sylP0. +have [ntx /(subsetP (Ohm_sub 1 _))P0x] := setD1P P0_1s_x. +have [Ltype1 | notLtype1] := boolP (FTtype L == 1)%N; last first. + have [U W W1 W2 defW LtypeP] := FTtypeP_witness maxL notLtype1. + suffices sP0H: P0 \subset H. + have [Hx notLtype5] := (subsetP sP0H x P0x, FTtype5_exclusion maxL). + have [_ _ _ tiFL] := compl_of_typeII_IV maxL LtypeP notLtype5. + have Fx: x \in 'F(L)^# by rewrite !inE ntx (subsetP (Fcore_sub_Fitting L)). + by have /idPn[] := cent1_normedTI tiFL Fx; rewrite setTI. + have [/=/FTcore_type2<- // | notLtype2] := boolP (FTtype L == 2). + have [_ _ [Ltype3 galL]] := FTtype34_structure maxL LtypeP notLtype2. + have cycU: cyclic U. + suffices regHU: Ptype_Fcompl_kernel LtypeP :=: 1%g. + rewrite (isog_cyclic (quotient1_isog U)) -regHU. + by have [|_ _ [//]] := typeP_Galois_P maxL _ galL; rewrite (eqP Ltype3). + rewrite /Ptype_Fcompl_kernel unlock /= astabQ /=. + have [_ _ ->] := FTtype34_Fcore_kernel_trivial maxL LtypeP notLtype2. + rewrite -morphpreIim -injm_cent ?injmK ?ker_coset ?norms1 //. + have [_ _ _ ->] := FTtype34_facts maxL LtypeP notLtype2. + by apply/derG1P; have [] := compl_of_typeIII maxL LtypeP Ltype3. + have sP0L': P0 \subset L^`(1) by rewrite -FTcore_type_gt2 ?(eqP Ltype3). + have [_ [_ _ _ defL'] _ _ _] := LtypeP. + have [nsHL' _ /mulG_sub[sHL' _] _ _] := sdprod_context defL'. + have hallH := pHall_subl sHL' (der_sub 1 L) (Fcore_Hall L). + have hallU: \pi(H)^'.-Hall(L^`(1)) U. + by rewrite -(compl_pHall U hallH) sdprod_compl. + rewrite (sub_normal_Hall hallH) // (pi_pgroup pP0) //. + have: ~~ cyclic P0; last apply: contraR => piK'p. + by rewrite abelian_rank1_cyclic // (rank_pgroup pP0) prankP0. + by have [|y _ /cyclicS->] := Hall_psubJ hallU piK'p _ pP0; rewrite ?cyclicJ. +have sP0H: P0 \subset H by rewrite /= -FTcore_type1. +have [U [LtypeF /= LtypeI]] := FTtypeP 1 maxL Ltype1. +have [[_ _ defL] _ _] := LtypeF; have [_ sUL _ nHU _] := sdprod_context defL. +have not_tiH: ~ normedTI H^# G L. + have H1x: x \in H^# by rewrite !inE ntx (subsetP sP0H). + by case/cent1_normedTI/(_ x H1x)/idPn; rewrite setTI. +apply/existsP; exists U; have [_ -> _] := typeF_context LtypeF. +apply/forall_inP=> Q /SylowP[q pr_q sylQ]; have [sQU qQ _] := and3P sylQ. +rewrite (odd_pgroup_rank1_cyclic qQ) ?mFT_odd //. +apply: wlog_neg; rewrite -ltnNge => /ltnW; rewrite p_rank_gt0 => piQq. +have hallU: \pi(H)^'.-Hall(L) U. + by rewrite -(compl_pHall U (Fcore_Hall L)) sdprod_compl. +have H'q := pnatPpi (pHall_pgroup hallU) (piSg sQU piQq). +rewrite leqNgt; apply: contra (H'q) => qrankQ; apply: IHp => //; last first. + by rewrite (leq_trans qrankQ) ?p_rankS ?(subset_trans sQU). +have piHp: p \in \pi(H) by rewrite (piSg sP0H) // -p_rank_gt0 prankP0. +have pr_p: prime p by have:= piHp; rewrite mem_primes => /andP[]. +have piUq: q \in \pi(exponent U) by rewrite pi_of_exponent (piSg sQU). +have [odd_p odd_q]: odd p /\ odd q. + rewrite !odd_2'nat !pnatE //. + by rewrite (pnatPpi _ piHp) ?(pnatPpi _ piQq) -?odd_2'nat ?mFT_odd. +have pgt2 := odd_prime_gt2 odd_p pr_p. +suffices [b dv_q_bp]: exists b : bool, q %| (b.*2 + p).-1. + rewrite -ltn_double (@leq_ltn_trans (p + b.*2).-1) //; last first. + by rewrite -!addnn -(subnKC pgt2) prednK // leq_add2l; case: (b). + rewrite -(subnKC pgt2) dvdn_leq // -mul2n Gauss_dvd ?coprime2n // -{1}subn1. + by rewrite dvdn2 odd_sub // subnKC // odd_add odd_p odd_double addnC. +have [// | [cHH rankH] | [/(_ p piHp)Udvp1 _]] := LtypeI; last first. + exists false; apply: dvdn_trans Udvp1. + by have:= piUq; rewrite mem_primes => /and3P[]. +suffices: q %| p ^ 2 - 1 ^ 2. + rewrite subn_sqr addn1 subn1 Euclid_dvdM //. + by case/orP; [exists false | exists true]. +pose P := 'O_p(H); pose P1 := 'Ohm_1(P). +have chP1H: P1 \char H := char_trans (Ohm_char 1 _) (pcore_char p H). +have sylP: p.-Sylow(H) P := nilpotent_pcore_Hall p (Fcore_nil L). +have [sPH pP _] := and3P sylP. +have abelP1: p.-abelem P1 by rewrite Ohm1_abelem ?(abelianS sPH). +have [pP1 _] := andP abelP1. +have prankP1: 'r_p(P1) = 2. + apply/eqP; rewrite p_rank_Ohm1 eqn_leq -{1}rank_pgroup // -{1}rankH rankS //=. + by rewrite -prankP0 (p_rank_Sylow sylP) p_rankS. +have ntP1: P1 != 1%g by rewrite -rank_gt0 (rank_pgroup pP1) prankP1. +have [_ _ [U0 [sU0U expU0 frobHU0]]] := LtypeF. +have nP1U0: U0 \subset 'N(P1). + by rewrite (char_norm_trans chP1H) ?(subset_trans sU0U). +rewrite subn1 -prankP1 p_rank_abelem // -card_pgroup //. +have frobP1U0 := Frobenius_subl ntP1 (char_sub chP1H) nP1U0 frobHU0. +apply: dvdn_trans (Frobenius_dvd_ker1 frobP1U0). +by have:= piUq; rewrite -expU0 pi_of_exponent mem_primes => /and3P[]. +Qed. + +Let Ltype1 : FTtype L == 1%N. Proof. exact: FT_Frobenius_type1 frobL. Qed. +Let defAL : 'A(L) = H^#. Proof. exact: FTsupp_Frobenius frobL. Qed. +Let sP0H : P0 \subset H. Proof. by rewrite /= -FTcore_type1. Qed. + +(* This is the first part of Peterfalvi (12.11). *) +Let defM : K ><| (M :&: L) = M. +Proof. +have [ntx /(subsetP (Ohm_sub 1 _))P0x] := setD1P P0_1s_x. +have Dx: x \in [set y in 'A0(L) | ~~ ('C[y] \subset L)]. + by rewrite inE FTsupp0_type1 // defAL !inE ntx (subsetP sP0H). +have [_ [_ /(_ x Dx)uCx] /(_ x Dx)[[defM _] _ _ _]] := FTsupport_facts maxL. +rewrite /K /= setIC (eq_uniq_mmax uCx maxM) //= -cent_cycle. +exact: subset_trans (cent_sub <[x]>) sNxM. +Qed. + +(* This is the second part of Peterfalvi (12.11). *) +Let sML_H : M :&: L \subset H. +Proof. +have [sP0M pP0 _] := and3P sylP0. +rewrite (sub_normal_Hall (Fcore_Hall L)) ?subsetIr //. +apply/pgroupP=> q pr_q /Cauchy[]// z /setIP[Mz Lz] oz; pose A := <[z]>%G. +have z_gt1: (#[z] > 1)%N by rewrite oz prime_gt1. +have sylP0_HM: p.-Sylow(H :&: M) P0. + by rewrite (pHall_subl _ _ sylP0) ?subsetIr // subsetI sP0H. +have nP0A: A \subset 'N(P0). + have sylHp: p.-Sylow(H) 'O_p(H) := nilpotent_pcore_Hall p (Fcore_nil L). + have sP0Hp: P0 \subset 'O_p(H) by rewrite sub_Hall_pcore. + have <-: 'O_p(H) :&: M = P0. + rewrite [_ :&: _](sub_pHall sylP0_HM) ?setSI ?pcore_sub //. + by rewrite (pgroupS (subsetIl _ _)) ?pcore_pgroup. + by rewrite subsetI sP0Hp. + have chHpL: 'O_p(H) \char L := char_trans (pcore_char p H) (Fcore_char L). + by rewrite normsI ?(char_norm_trans chHpL) ?normsG // cycle_subG. +apply: wlog_neg => piH'q. +have coHQ: coprime #|H| #|A| by rewrite -orderE coprime_pi' // oz pnatE. +have frobP0A: [Frobenius P0 <*> A = P0 ><| A]. + have defHA: H ><| A = H <*> A. + by rewrite sdprodEY ?coprime_TIg // cycle_subG (subsetP (gFnorm _ _)). + have ltH_HA: H \proper H <*> A. + by rewrite /proper joing_subl -indexg_gt1 -(index_sdprod defHA). + have: [Frobenius H <*> A = H ><| A]. + apply: set_Frobenius_compl defHA _. + by apply: Frobenius_kerS frobL; rewrite // join_subG gFsub cycle_subG. + by apply: Frobenius_subl => //; rewrite -rank_gt0 (rank_pgroup pP0) prankP0. +have sP0A_M: P0 <*> A \subset M by rewrite join_subG sP0M cycle_subG. +have nKP0a: P0 <*> A \subset 'N(K) := subset_trans sP0A_M (gFnorm _ _). +have solK: solvable K := nilpotent_sol (Fcore_nil M). +have [_ [/(compl_of_typeF defM) MtypeF _]] := FTtypeP 1 maxM Mtype1. +have nreg_KA: 'C_K(A) != 1%g. + have [Kq | K'q] := boolP (q \in \pi(K)). + apply/trivgPn; exists z; rewrite -?order_gt1 //= cent_cycle inE cent1id. + by rewrite andbT (mem_normal_Hall (Fcore_Hall M)) // /p_elt oz pnatE. + have [defP0A ntP0 _ _ _] := Frobenius_context frobP0A. + have coK_P0A: coprime #|K| #|P0 <*> A|. + rewrite -(sdprod_card defP0A) coprime_mulr (p'nat_coprime p'K) //=. + by rewrite -orderE coprime_pi' // oz pnatE. + have: ~~ (P0 \subset 'C(K)); last apply: contraNneq. + have [[ntK _ _] _ [U0 [sU0ML expU0 frobKU0]]] := MtypeF. + have [P1 /pnElemP[sP1U0 abelP1 dimP1]] := p_rank_witness p U0. + have ntP1: P1 :!=: 1%g. + rewrite -rank_gt0 (rank_abelem abelP1) dimP1 p_rank_gt0 -pi_of_exponent. + rewrite expU0 pi_of_exponent (piSg (setIS M (Fcore_sub L))) //=. + by rewrite setIC -p_rank_gt0 -(p_rank_Sylow sylP0_HM) prankP0. + have frobKP1: [Frobenius K <*> P1 = K ><| P1]. + exact: Frobenius_subr ntP1 sP1U0 frobKU0. + have sP1M: P1 \subset M. + by rewrite (subset_trans (subset_trans sP1U0 sU0ML)) ?subsetIl. + have [y My sP1yP0] := Sylow_Jsub sylP0 sP1M (abelem_pgroup abelP1). + apply: contra ntK => cP0K; rewrite -(Frobenius_trivg_cent frobKP1). + rewrite (setIidPl _) // -(conjSg _ _ y) (normsP _ y My) ?gFnorm //. + by rewrite -centJ centsC (subset_trans sP1yP0). + by have [] := Frobenius_Wielandt_fixpoint frobP0A nKP0a coK_P0A solK. +have [_ [U1 [_ abU1 sCK_U1]] _] := MtypeF. +have [ntx /(subsetP (Ohm_sub 1 _))P0x] := setD1P P0_1s_x. +have cAx: A \subset 'C[x]. + rewrite -cent_set1 (sub_abelian_cent2 abU1) //. + have [y /setIP[Ky cAy] nty] := trivgPn _ nreg_KA. + apply: subset_trans (sCK_U1 y _); last by rewrite !inE nty. + by rewrite subsetI sub_cent1 cAy cycle_subG !inE Mz Lz. + have [y /setIP[Ky cxy] notK'y] := subsetPn not_sCxK'. + apply: subset_trans (sCK_U1 y _); last by rewrite !inE (group1_contra notK'y). + rewrite sub1set inE cent1C cxy (subsetP _ x P0x) //. + by rewrite subsetI sP0M (subset_trans sP0H) ?gFsub. +have [_ _ _ regHL] := Frobenius_kerP frobL. +rewrite (piSg (regHL x _)) //; first by rewrite !inE ntx (subsetP sP0H). +by rewrite mem_primes pr_q cardG_gt0 -oz cardSg // subsetI cycle_subG Lz. +Qed. + +Let E := sval (sigW (existsP frobL)). +Let e := #|E|. + +Let defL : H ><| E = L. +Proof. by rewrite /E; case: (sigW _) => E0 /=/Frobenius_context[]. Qed. + +Let Ecyclic_le_p : cyclic E /\ (e %| p.-1) || (e %| p.+1). +Proof. +pose P := 'O_p(H)%G; pose T := 'Ohm_1('Z(P))%G. +have sylP: p.-Sylow(H) P := nilpotent_pcore_Hall p (Fcore_nil L). +have [[sPH pP _] [sP0M pP0 _]] := (and3P sylP, and3P sylP0). +have sP0P: P0 \subset P by rewrite (sub_normal_Hall sylP) ?pcore_normal. +have defP0: P :&: M = P0. + rewrite [P :&: M](sub_pHall sylP0 (pgroupS _ pP)) ?subsetIl ?subsetIr //. + by rewrite subsetI sP0P. +have [ntx P01x] := setD1P P0_1s_x; have P0x := subsetP (Ohm_sub 1 P0) x P01x. +have sZP0: 'Z(P) \subset P0. + apply: subset_trans (_ : 'C_P[x] \subset P0). + by rewrite -cent_set1 setIS ?centS // sub1set (subsetP sP0P). + by rewrite -defP0 setIS // (subset_trans _ sNxM) // cents_norm ?cent_cycle. +have ntT: T :!=: 1%g. + rewrite Ohm1_eq1 center_nil_eq1 ?(pgroup_nil pP) // (subG1_contra sP0P) //. + by apply/trivgPn; exists x. +have [_ sEL _ nHE tiHE] := sdprod_context defL. +have charTP: T \char P := char_trans (Ohm_char 1 _) (center_char P). +have{ntT} [V minV sVT]: {V : {group gT} | minnormal V E & V \subset T}. + apply: mingroup_exists; rewrite ntT (char_norm_trans charTP) //. + exact: char_norm_trans (pcore_char p H) nHE. +have abelT: p.-abelem T by rewrite Ohm1_abelem ?center_abelian ?(pgroupS sZP0). +have sTP := subset_trans (Ohm_sub 1 _) sZP0. +have rankT: ('r_p(T) <= 2)%N by rewrite -prankP0 p_rankS. +have [abelV /andP[ntV nVE]] := (abelemS sVT abelT, mingroupp minV). +have pV := abelem_pgroup abelV; have [pr_p _ [n oV]] := pgroup_pdiv pV ntV. +have frobHE: [Frobenius L = H ><| E] by rewrite /E; case: (sigW _). +have: ('r_p(V) <= 2)%N by rewrite (leq_trans (p_rankS p sVT)). +rewrite (p_rank_abelem abelV) // oV pfactorK // ltnS leq_eqVlt ltnS leqn0 orbC. +have sVH := subset_trans sVT (subset_trans (char_sub charTP) sPH). +have regVE: 'C_E(V) = 1%g. + exact: cent_semiregular (Frobenius_reg_compl frobHE) sVH ntV. +case/pred2P=> dimV; rewrite {n}dimV in oV. + pose f := [morphism of restrm nVE (conj_aut V)]. + have injf: 'injm f by rewrite ker_restrm ker_conj_aut regVE. + rewrite /e -(injm_cyclic injf) // -(card_injm injf) //. + have AutE: f @* E \subset Aut V by rewrite im_restrm Aut_conj_aut. + rewrite (cyclicS AutE) ?Aut_prime_cyclic ?oV // (dvdn_trans (cardSg AutE)) //. + by rewrite card_Aut_cyclic ?prime_cyclic ?oV // totient_pfactor ?muln1. +have defV: V :=: 'Ohm_1(P0). + apply/eqP; rewrite eqEcard (subset_trans sVT) ?OhmS //= oV -prankP0. + by rewrite p_rank_abelian // -card_pgroup ?(pgroupS (Ohm_sub 1 _)). +pose rE := abelem_repr abelV ntV nVE. +have ffulE: mx_faithful rE by apply: abelem_mx_faithful. +have p'E: [char 'F_p]^'.-group E. + rewrite (eq_p'group _ (charf_eq (char_Fp pr_p))) (coprime_p'group _ pV) //. + by rewrite coprime_sym (coprimeSg sVH) ?(Frobenius_coprime frobHE). +have dimV: 'dim V = 2 by rewrite (dim_abelemE abelV) // oV pfactorK. +have cEE: abelian E. + by rewrite dimV in (rE) ffulE; apply: charf'_GL2_abelian (mFT_odd E) ffulE _. +have Enonscalar y: y \in E -> y != 1%g -> ~~ is_scalar_mx (rE y). + move=> Ey; apply: contra => /is_scalar_mxP[a rEy]; simpl in a. + have nXy: y \in 'N(<[x]>). + rewrite !inE -cycleJ cycle_subG; apply/cycleP; exists a. + have [Vx nVy]: x \in V /\ y \in 'N(V) by rewrite (subsetP nVE) ?defV. + apply: (@abelem_rV_inj p _ V); rewrite ?groupX ?memJ_norm ?morphX //=. + by rewrite zmodXgE -scaler_nat natr_Zp -mul_mx_scalar -rEy -abelem_rV_J. + rewrite -in_set1 -set1gE -tiHE inE (subsetP sML_H) //. + by rewrite inE (subsetP sEL) // (subsetP sNxM). +have /trivgPn[y nty Ey]: E != 1%G by have [] := Frobenius_context frobHE. +have cErEy: centgmx rE (rE y). + by apply/centgmxP=> z Ez; rewrite -!repr_mxM // (centsP cEE). +have irrE: mx_irreducible rE by apply/abelem_mx_irrP. +have charFp2: p \in [char MatrixGenField.gen_finFieldType irrE cErEy]. + apply: (rmorph_char (MatrixGenField.gen_rmorphism irrE cErEy)). + exact: char_Fp. +pose Fp2 := primeChar_finFieldType charFp2. +pose n1 := MatrixGenField.gen_dim (rE y). +pose rEp2 : mx_representation Fp2 E n1 := MatrixGenField.gen_repr irrE cErEy. +have n1_gt0: (0 < n1)%N := MatrixGenField.gen_dim_gt0 irrE cErEy. +have n1_eq1: n1 = 1%N. + pose d := degree_mxminpoly (rE y). + have dgt0: (0 < d)%N := mxminpoly_nonconstant _. + apply/eqP; rewrite eqn_leq n1_gt0 andbT -(leq_pmul2r dgt0). + rewrite (MatrixGenField.gen_dim_factor irrE cErEy) mul1n dimV. + by rewrite ltnNge mxminpoly_linear_is_scalar Enonscalar. +have oFp2: #|Fp2| = (p ^ 2)%N. + rewrite card_sub card_matrix card_Fp // -{1}n1_eq1. + by rewrite (MatrixGenField.gen_dim_factor irrE cErEy) dimV. +have [f rfK fK]: bijective (@scalar_mx Fp2 n1). + rewrite n1_eq1. + by exists (fun A : 'M_1 => A 0 0) => ?; rewrite ?mxE -?mx11_scalar. +pose g z : {unit Fp2} := insubd (1%g : {unit Fp2}) (f (rEp2 z)). +have val_g z : z \in E -> (val (g z))%:M = rEp2 z. + move=> Ez; rewrite insubdK ?fK //; have:= repr_mx_unit rEp2 Ez. + by rewrite -{1}[rEp2 z]fK unitmxE det_scalar !unitfE expf_eq0 n1_gt0. +have ffulEp2: mx_faithful rEp2 by rewrite MatrixGenField.gen_mx_faithful. +have gM: {in E &, {morph g: z1 z2 / z1 * z2}}%g. + move=> z1 z2 Ez1 Ez2 /=; apply/val_inj/(can_inj rfK). + rewrite {1}(val_g _ (groupM Ez1 Ez2)) scalar_mxM. + by rewrite {1}(val_g _ Ez1) (val_g _ Ez2) repr_mxM. +have inj_g: 'injm (Morphism gM). + apply/injmP=> z1 z2 Ez1 Ez2 /(congr1 (@scalar_mx _ n1 \o val)). + by rewrite /= {1}(val_g _ Ez1) (val_g _ Ez2); apply: mx_faithful_inj. +split; first by rewrite -(injm_cyclic inj_g) ?field_unit_group_cyclic. +have: e %| #|[set: {unit Fp2}]|. + by rewrite /e -(card_injm inj_g) ?cardSg ?subsetT. +rewrite card_finField_unit oFp2 -!subn1 (subn_sqr p 1) addn1. +rewrite orbC Gauss_dvdr; first by move->. +rewrite coprime_sym coprime_has_primes ?subn_gt0 ?prime_gt1 ?cardG_gt0 //. +apply/hasPn=> r; rewrite /= !mem_primes subn_gt0 prime_gt1 ?cardG_gt0 //=. +case/andP=> pr_r /Cauchy[//|z Ez oz]; rewrite pr_r /= subn1. +apply: contra (Enonscalar z Ez _); last by rewrite -order_gt1 oz prime_gt1. +rewrite -oz -(order_injm inj_g) // order_dvdn -val_eqE => /eqP gz_p1_eq1. +have /vlineP[a Dgz]: val (g z) \in 1%VS. + rewrite Fermat's_little_theorem dimv1 card_Fp //=. + by rewrite -[(p ^ 1)%N]prednK ?prime_gt0 // exprS -val_unitX gz_p1_eq1 mulr1. +apply/is_scalar_mxP; exists a; apply/row_matrixP=> i. +apply: (can_inj ((MatrixGenField.in_genK irrE cErEy) _)). +rewrite !rowE mul_mx_scalar MatrixGenField.in_genZ MatrixGenField.in_genJ //. +rewrite -val_g // Dgz mul_mx_scalar; congr (_ *: _). +rewrite -(natr_Zp a) scaler_nat. +by rewrite -(rmorph_nat (MatrixGenField.gen_rmorphism irrE cErEy)). +Qed. + +Let calS := seqIndD H L H 1. +Notation tauL := (FT_Dade maxL). +Notation tauL_H := (FT_DadeF maxL). +Notation rhoL := (invDade (FT_DadeF_hyp maxL)). + +Section Twelve_13_to_16. + +Variables (tau1 : {additive 'CF(L) -> 'CF(G)}) (chi : 'CF(L)). +Hypothesis cohS : coherent_with calS L^# tauL tau1. +Hypotheses (Schi : chi \in calS) (chi1 : chi 1%g = e%:R). +Let psi := tau1 chi. + +Let cohS_H : coherent_with calS L^# tauL_H tau1. +Proof. +have [? Dtau] := cohS; split=> // xi Sxi; have /zcharD1_seqInd_on Hxi := Sxi. +by rewrite Dtau // FT_DadeF_E ?FT_DadeE ?(cfun_onS (Fcore_sub_FTsupp _)) ?Hxi. +Qed. + +(* This is Peterfalvi (12.14). *) +Let rhoL_psi : {in K, forall g, psi (x * g)%g = chi x} /\ rhoL psi x = chi x. +Proof. +have not_LGM: gval M \notin L :^: G. + apply: contraL p'K => /= /imsetP[z _ ->]; rewrite FcoreJ pgroupJ. + by rewrite p'groupEpi (piSg sP0H) // -p_rank_gt0 prankP0. +pose rmR := sval (Rgen maxL Ltype1). +have Zpsi: psi \in 'Z[rmR chi]. + case: (Rgen _ _) @rmR => /= rmR []; rewrite -/calS => scohS _ _. + have sSS: cfConjC_subset calS calS by apply: seqInd_conjC_subset1. + have [B /mem_subseq sBR Dpsi] := mem_coherent_sum_subseq scohS sSS cohS Schi. + by rewrite [psi]Dpsi big_seq rpred_sum // => xi /sBR/mem_zchar->. +have [ntx /(subsetP (Ohm_sub 1 P0))P0x] := setD1P P0_1s_x. +have Mx: x \in M by rewrite (subsetP sNxM) // -cycle_subG normG. +have psi_xK: {in K, forall g, psi (x * g)%g = psi x}. + move=> g Kg; have{Kg}: (x * g \in x *: K)%g by rewrite mem_lcoset mulKg. + apply: FTtype1_ortho_constant => [phi calMphi|]. + apply/orthoPl=> nu /memv_span; apply: {nu}span_orthogonal (zchar_span Zpsi). + exact: FTtype1_seqInd_ortho. + rewrite inE -/K (contra _ ntx) // => Kx. + rewrite -(consttC p x) !(constt1P _) ?mulg1 ?(mem_p_elt p'K) //. + by rewrite p_eltNK (mem_p_elt (pHall_pgroup sylP0)). +have H1x: x \in H^# by rewrite !inE ntx (subsetP sP0H). +have rhoL_psi_x: rhoL psi x = psi x. + rewrite cfunElock mulrb def_FTsignalizerF H1x //=. + apply: canLR (mulKf (neq0CG _)) _; rewrite mulr_natl -sumr_const /=. + apply: eq_bigr => g; rewrite /'R_L (negPf not_sCxL) /= setIC => /setIP[cxz]. + have Dx: x \in [set y in 'A0(L) | ~~ ('C[y] \subset L)]. + by rewrite inE (subsetP (Fcore_sub_FTsupp0 _)). + have [_ [_ /(_ x Dx)defNx] _] := FTsupport_facts maxL. + rewrite (cent1P cxz) -(eq_uniq_mmax defNx maxM) => [/psi_xK//|]. + by rewrite /= -cent_cycle (subset_trans (cent_sub _)). +suffices <-: rhoL psi x = chi x by split=> // g /psi_xK->. +have irrS: {subset calS <= irr L} by have [] := FT_Frobenius_coherence maxL. +have irr_chi := irrS _ Schi. +have Sgt1: (1 < size calS)%N by apply: seqInd_nontrivial Schi; rewrite ?mFT_odd. +have De: #|L : H| = e by rewrite -(index_sdprod defL). +have [] := Dade_Ind1_sub_lin cohS_H Sgt1 irr_chi Schi; rewrite ?De //. +rewrite -/tauL_H -/calS -/psi /=; set alpha := 'Ind 1 - chi. +case=> o_tau_1 tau_alpha_1 _ [Gamma [o_tau1_Ga _ [a Za tau_alpha] _] _]. +have [[Itau1 _] Dtau1] := cohS_H. +have o1calS: orthonormal calS. + by rewrite (sub_orthonormal irrS) ?seqInd_uniq ?irr_orthonormal. +have norm_alpha: '[tauL_H alpha] = e%:R + 1. + rewrite Dade_isometry ?(cfInd1_sub_lin_on _ Schi) ?De //. + rewrite cfnormBd; last by rewrite cfdotC (seqInd_ortho_Ind1 _ _ Schi) ?conjC0. + by rewrite cfnorm_Ind_cfun1 // De irrWnorm. +pose h := #|H|; have ub_a: a ^+ 2 * ((h%:R - 1) / e%:R) - 2%:R * a <= e%:R - 1. + rewrite -[h%:R - 1](mulKf (neq0CiG L H)) -sum_seqIndC1_square // De -/calS. + rewrite -[lhs in lhs - 1](addrK 1) -norm_alpha -[tauL_H _](subrK 1). + rewrite cfnormDd; last by rewrite cfdotBl tau_alpha_1 cfnorm1 subrr. + rewrite cfnorm1 addrK [in '[_]]addrC {}tau_alpha -!addrA addKr addrCA addrA. + rewrite ler_subr_addr cfnormDd ?ler_paddr ?cfnorm_ge0 //; last first. + rewrite cfdotBl cfdotZl cfdot_suml (orthoPr o_tau1_Ga) ?map_f // subr0. + rewrite big1_seq ?mulr0 // => xi Sxi; rewrite cfdotZl. + by rewrite (orthoPr o_tau1_Ga) ?map_f ?mulr0. + rewrite cfnormB cfnormZ Cint_normK // cfdotZl cfproj_sum_orthonormal //. + rewrite cfnorm_sum_orthonormal // Itau1 ?mem_zchar // irrWnorm ?irrS // divr1. + rewrite chi1 divff ?neq0CG // mulr1 conj_Cint // addrAC mulr_natl. + rewrite !ler_add2r !(mulr_suml, mulr_sumr) !big_seq ler_sum // => xi Sxi. + rewrite irrWnorm ?irrS // !divr1 (mulrAC _^-1) -expr2 -!exprMn (mulrC _^-1). + by rewrite normf_div normr_nat norm_Cnat // (Cnat_seqInd1 Sxi). +have [pr_p p_dv_M]: prime p /\ p %| #|M|. + have: p \in \pi(M) by rewrite -p_rank_gt0 ltnW. + by rewrite mem_primes => /and3P[]. +have odd_p: odd p by rewrite (dvdn_odd p_dv_M) ?mFT_odd. +have pgt2: (2 < p)%N := odd_prime_gt2 odd_p pr_p. +have ub_e: e%:R <= (p%:R + 1) / 2%:R :> algC. + rewrite ler_pdivl_mulr ?ltr0n // -natrM -mulrSr leC_nat muln2. + have [b e_dv_pb]: exists b : bool, e %| (b.*2 + p).-1. + by have [_ /orP[]] := Ecyclic_le_p; [exists false | exists true]. + rewrite -ltnS (@leq_trans (b.*2 + p)) //; last first. + by rewrite (leq_add2r p _ 2) (leq_double _ 1) leq_b1. + rewrite dvdn_double_ltn ?mFT_odd //; first by rewrite odd_add odd_double. + by rewrite -(subnKC pgt2) !addnS. +have lb_h: p%:R ^+ 2 <= h%:R :> algC. + rewrite -natrX leC_nat dvdn_leq ?pfactor_dvdn ?cardG_gt0 //. + by rewrite -prankP0 (leq_trans (p_rankS p sP0H)) ?p_rank_le_logn. +have{ub_a ub_e} ub_a: p.-1.*2%:R * a ^+ 2 - 2%:R * a <= p.-1%:R / 2%:R :> algC. + apply: ler_trans (ler_trans ub_a _); last first. + rewrite -subn1 -subSS natrB ?ltnS ?prime_gt0 // mulrSr mulrBl. + by rewrite divff ?pnatr_eq0 ?ler_add2r. + rewrite ler_add2r mulrC -Cint_normK // -!mulrA !ler_wpmul2l ?normr_ge0 //. + rewrite ler_pdivl_mulr ?gt0CG // ler_subr_addr (ler_trans _ lb_h) //. + rewrite -muln2 natrM -mulrA -ler_subr_addr subr_sqr_1. + rewrite -(natrB _ (prime_gt0 pr_p)) subn1 ler_wpmul2l ?ler0n //. + by rewrite mulrC -ler_pdivl_mulr ?ltr0n. +have a0: a = 0. + apply: contraTeq ub_a => nz_a; rewrite ltr_geF // ltr_pdivr_mulr ?ltr0n //. + rewrite mulrC -{1}mulr_natl -muln2 natrM -mulrA mulrBr mulrCA ltr_subr_addl. + rewrite -ltr_subr_addr -mulrBr mulr_natl mulrA -expr2 -exprMn. + apply: ltr_le_trans (_ : 2%:R * ((a *+ 2) ^+ 2 - 1) <= _); last first. + rewrite (mulr_natl a 2) ler_wpmul2r // ?subr_ge0. + by rewrite sqr_Cint_ge1 ?rpredMn // mulrn_eq0. + by rewrite leC_nat -subn1 ltn_subRL. + rewrite -(@ltr_pmul2l _ 2%:R) ?ltr0n // !mulrA -expr2 mulrBr -exprMn mulr1. + rewrite -natrX 2!mulrnAr -[in rhs in _ < rhs]mulrnAl -mulrnA. + rewrite ltr_subr_addl -ltr_subr_addr -(ltr_add2r 1) -mulrSr -sqrrB1. + rewrite -Cint_normK ?rpredB ?rpredM ?rpred_nat ?rpred1 //. + rewrite (@ltr_le_trans _ (3 ^ 2)%:R) ?ltC_nat // natrX. + rewrite ler_sqr ?qualifE ?ler0n ?normr_ge0 //. + rewrite (ler_trans _ (ler_sub_dist _ _)) // normr1 normrM normr_nat. + by rewrite ler_subr_addl -mulrS mulr_natl ler_pmuln2r ?norm_Cint_ge1. +pose chi0 := 'Ind[L, H] 1. +have defS1: perm_eq (seqIndT H L) (chi0 :: calS). + by rewrite [calS]seqIndC1_rem // perm_to_rem ?seqIndT_Ind1. +have [c _ -> // _] := invDade_seqInd_sum (FT_DadeF_hyp maxL) psi defS1. +have psi_alpha_1: '[psi, tauL_H alpha] = -1. + rewrite tau_alpha a0 scale0r addr0 addrC addrA cfdotBr cfdotDr. + rewrite (orthoPr o_tau_1) ?(orthoPr o_tau1_Ga) ?map_f // !add0r. + by rewrite Itau1 ?mem_zchar ?map_f // irrWnorm ?irrS. +rewrite (bigD1_seq chi) ?seqInd_uniq //= big1_seq => [|xi /andP[chi'xi Sxi]]. + rewrite addr0 -cfdotC chi1 cfInd1 ?gFsub // cfun11 mulr1 De divff ?neq0CG //. + rewrite scale1r -opprB linearN cfdotNr psi_alpha_1 opprK. + by rewrite irrWnorm ?irrS // divr1 mul1r. +rewrite -cfdotC cfInd1 ?gFsub // cfun11 mulr1. +rewrite /chi0 -(canLR (subrK _) (erefl alpha)) scalerDr opprD addrCA -scaleNr. +rewrite linearD linearZ /= cfdotDr cfdotZr psi_alpha_1 mulrN1 rmorphN opprK. +rewrite -/tauL_H -Dtau1 ?zcharD1_seqInd ?(seqInd_sub_lin_vchar _ Schi) ?De //. +have [_ ooS] := orthonormalP o1calS. +rewrite raddfB cfdotBr Itau1 ?mem_zchar // ooS // mulrb ifN_eqC // add0r. +rewrite -De raddfZ_Cnat ?(dvd_index_seqInd1 _ Sxi) // De cfdotZr. +by rewrite Itau1 ?mem_zchar ?ooS // eqxx mulr1 subrr !mul0r. +Qed. + +Let rhoM := invDade (FT_DadeF_hyp maxM). + +Let rhoM_psi : + [/\ {in K^#, rhoM psi =1 psi}, + {in K :\: K' &, forall g1 g2, psi g1 = psi g2} + & {in K :\: K', forall g, psi g \in Cint}]. +Proof. +have pr_p: prime p. + by have:= ltnW prankM; rewrite p_rank_gt0 mem_primes => /andP[]. +have [sP0M pP0 _] := and3P sylP0; have abelP01 := Ohm1_abelem pP0 abP0. +have not_frobM: ~~ [Frobenius M with kernel K]. + apply: contraL prankM => /(set_Frobenius_compl defM)frobM. + rewrite -leqNgt -(p_rank_Sylow sylP0) -p_rank_Ohm1 p_rank_abelem //. + rewrite -abelem_cyclic // (cyclicS (Ohm_sub _ _)) //. + have sP0ML: P0 \subset M :&: L. + by rewrite subsetI sP0M (subset_trans sP0H) ?gFsub. + rewrite nil_Zgroup_cyclic ?(pgroup_nil pP0) // (ZgroupS sP0ML) //. + have [U [MtypeF _]] := FTtypeP 1 maxM Mtype1. + by have{MtypeF} /typeF_context[_ <- _] := compl_of_typeF defM MtypeF. +pose rmR := sval (Rgen maxL Ltype1). +have Zpsi: psi \in 'Z[rmR chi]. + case: (Rgen _ _) @rmR => /= rmR []; rewrite -/calS => scohS _ _. + have sSS: cfConjC_subset calS calS by apply: seqInd_conjC_subset1. + have [B /mem_subseq sBR Dpsi] := mem_coherent_sum_subseq scohS sSS cohS Schi. + by rewrite [psi]Dpsi big_seq rpred_sum // => xi /sBR/mem_zchar->. +have part1: {in K^#, rhoM psi =1 psi}. + move=> g K1g; rewrite /= cfunElock mulrb def_FTsignalizerF K1g //= /'R_M. + have [_ | sCg'M] := ifPn; first by rewrite cards1 big_set1 invr1 mul1r mul1g. + have Dg: g \in [set z in 'A0(M) | ~~ ('C[z] \subset M)]. + by rewrite inE (subsetP (Fcore_sub_FTsupp0 _)). + have [_ [_ /(_ g Dg)maxN] /(_ g Dg)[_ _ ANg Ntype12]] := FTsupport_facts maxM. + have{maxN} [maxN sCgN] := mem_uniq_mmax maxN. + have{Ntype12} Ntype1: FTtype 'N[g] == 1%N. + have [] := Ntype12; rewrite -(mem_iota 1 2) !inE => /orP[// | Ntype2] frobM. + by have /negP[] := not_frobM; apply/frobM/Ntype2. + have not_frobN: ~~ [Frobenius 'N[g] with kernel 'N[g]`_\F]. + apply/Frobenius_kerP=> [[_ _ _ regFN]]. + have [/bigcupP[y]] := setDP ANg; rewrite FTsupp1_type1 Ntype1 //. + by move=> /regFN sCyF /setD1P[ntg cNy_g]; rewrite 2!inE ntg (subsetP sCyF). + have LG'N: gval 'N[g] \notin L :^: G. + by apply: contra not_frobN => /imsetP[y _ ->]; rewrite FcoreJ FrobeniusJker. + suff /(eq_bigr _)->: {in 'C_('N[g]`_\F)[g], forall z, psi (z * g)%g = psi g}. + by rewrite sumr_const -[psi g *+ _]mulr_natl mulKf ?neq0CG. + move=> z /setIP[Fz /cent1P cgz]. + have{Fz cgz}: (z * g \in g *: 'N[g]`_\F)%g by rewrite cgz mem_lcoset mulKg. + apply: FTtype1_ortho_constant => [phi calMphi|]. + apply/orthoPl=> nu /memv_span; apply: span_orthogonal (zchar_span Zpsi). + exact: FTtype1_seqInd_ortho. + have [/(subsetP (FTsupp_sub _))/setD1P[ntg Ng]] := setDP ANg. + by rewrite FTsupp1_type1 //= !inE ntg Ng andbT. +have part2: {in K :\: K' &, forall g1 g2, psi g1 = psi g2}. + have /subsetP sK1_K: K :\: K' \subset K^# by rewrite setDS ?sub1G. + have LG'M: gval M \notin L :^: G. + apply: contra not_frobM => /imsetP[y _ /= ->]. + by rewrite FcoreJ FrobeniusJker. + move=> g1 g2 Kg1 Kg2; rewrite /= -!part1 ?sK1_K //. + apply: FtypeI_invDade_ortho_constant => // phi calMphi. + apply/orthoPl=> nu /memv_span; apply: span_orthogonal (zchar_span Zpsi). + exact: FTtype1_seqInd_ortho. +split=> // g KK'g; pose nKK' : algC := #|K :\: K'|%:R. +pose nK : algC := #|K|%:R; pose nK' : algC := #|K'|%:R. +have nzKK': nKK' != 0 by rewrite pnatr_eq0 cards_eq0; apply/set0Pn; exists g. +have Dpsi_g: nK * '['Res[K] psi, 1] = nK' * '['Res[K'] psi, 1] + nKK' * psi g. + rewrite !mulVKf ?neq0CG // (big_setID K') (setIidPr (gFsub _ _)) /=. + rewrite mulr_natl -sumr_const; congr (_ + _); apply: eq_bigr => z K'z. + by rewrite !cfun1E !cfResE ?subsetT ?(subsetP (der_sub 1 K)) ?K'z. + have [Kz _] := setDP K'z; rewrite cfun1E Kz conjC1 mulr1 cfResE ?subsetT //. + exact: part2. +have{Zpsi} Zpsi: psi \in 'Z[irr G] by have [[_ ->//]] := cohS; apply: mem_zchar. +have Qpsi1 R: '['Res[R] psi, 1] \in Crat. + by rewrite rpred_Cint ?Cint_cfdot_vchar ?rpred1 ?cfRes_vchar. +apply: Cint_rat_Aint (Aint_vchar g Zpsi). +rewrite -[psi g](mulKf nzKK') -(canLR (addKr _) Dpsi_g) addrC mulrC. +by rewrite rpred_div ?rpredB 1?rpredM ?rpred_nat ?Qpsi1. +Qed. + +(* This is the main part of Peterfalvi (12.16). *) +Lemma FTtype1_nonFrobenius_witness_contradiction : False. +Proof. +have pr_p: prime p. + by have:= ltnW prankM; rewrite p_rank_gt0 mem_primes => /andP[]. +have [sP0M pP0 _] := and3P sylP0; have abelP01 := Ohm1_abelem pP0 abP0. +have [ntx P01x] := setD1P P0_1s_x. +have ox: #[x] = p := abelem_order_p abelP01 P01x ntx. +have odd_p: odd p by rewrite -ox mFT_odd. +have pgt2 := odd_prime_gt2 odd_p pr_p. +have Zpsi: psi \in 'Z[irr G] by have [[_ ->//]] := cohS; apply: mem_zchar. +have lb_psiM: '[rhoM psi] >= #|K :\: K'|%:R / #|M|%:R * e.-1%:R ^+ 2. + have [g /setIP[Kg cxg] notK'g] := subsetPn not_sCxK'. + have KK'g: g \in K :\: K' by rewrite !inE notK'g. + have [rhoMid /(_ _ g _ KK'g)psiKK'_id /(_ g KK'g)Zpsig] := rhoM_psi. + rewrite -mulrA mulrCA ler_pmul2l ?invr_gt0 ?gt0CG // mulr_natl. + rewrite (big_setID (K :\: K')) (setIidPr _) ?subDset ?subsetU ?gFsub ?orbT //. + rewrite ler_paddr ?sumr_ge0 // => [z _|]; first exact: mul_conjC_ge0. + rewrite -sumr_const ler_sum // => z KK'z. + rewrite {}rhoMid ?(subsetP _ z KK'z) ?setDS ?sub1G // {}psiKK'_id {z KK'z}//. + rewrite -normCK ler_sqr ?qualifE ?ler0n ?normr_ge0 //. + have [eps prim_eps] := C_prim_root_exists (prime_gt0 pr_p). + have psi_xg: (psi (x * g)%g == e%:R %[mod 1 - eps])%A. + have [-> // _] := rhoL_psi; rewrite -[x]mulg1 -chi1. + rewrite (vchar_ker_mod_prim prim_eps) ?group1 ?(seqInd_vcharW Schi) //. + rewrite (subsetP _ _ P01x) // (subset_trans (Ohm_sub 1 _)) //. + by rewrite (subset_trans sP0H) ?gFsub. + have{psi_xg} /dvdCP[a Za /(canRL (subrK _))->]: (p %| psi g - e%:R)%C. + rewrite (int_eqAmod_prime_prim prim_eps) ?rpredB ?rpred_nat // eqAmod0. + apply: eqAmod_trans psi_xg; rewrite eqAmod_sym. + by rewrite (vchar_ker_mod_prim prim_eps) ?in_setT. + have [-> | nz_a] := eqVneq a 0. + by rewrite mul0r add0r normr_nat leC_nat leq_pred. + rewrite -[e%:R]opprK (ler_trans _ (ler_sub_dist _ _)) // normrN normrM. + rewrite ler_subr_addl !normr_nat -natrD. + apply: ler_trans (_ : 1 * p%:R <= _); last first. + by rewrite ler_wpmul2r ?ler0n ?norm_Cint_ge1. + rewrite mul1r leC_nat -subn1 addnBA ?cardG_gt0 // leq_subLR addnn -ltnS. + have [b e_dv_pb]: exists b : bool, e %| (b.*2 + p).-1. + by have [_ /orP[]] := Ecyclic_le_p; [exists false | exists true]. + apply: (@leq_trans (b.*2 + p)); last first. + by rewrite (leq_add2r p _ 2) (leq_double b 1) leq_b1. + rewrite dvdn_double_ltn ?odd_add ?mFT_odd ?odd_double //. + by rewrite addnC -(subnKC pgt2). +have irrS: {subset calS <= irr L} by have [] := FT_Frobenius_coherence maxL. +have lb_psiL: '[rhoL psi] >= 1 - e%:R / #|H|%:R. + have irr_chi := irrS _ Schi. + have Sgt1: (1 < size calS)%N by apply: seqInd_nontrivial (mFT_odd _) _ _ Schi. + have De: #|L : H| = e by rewrite -(index_sdprod defL). + have [|_] := Dade_Ind1_sub_lin cohS_H Sgt1 irr_chi Schi; rewrite De //=. + by rewrite -De odd_Frobenius_index_ler ?mFT_odd // => -[_ _ []//]. +have tiA1_LM: [disjoint 'A1~(L) & 'A1~(M)]. + apply: FT_Dade1_support_disjoint => //. + apply: contraL p'K => /= /imsetP[z _ ->]; rewrite FcoreJ pgroupJ. + by rewrite p'groupEpi (piSg sP0H) // -p_rank_gt0 prankP0. +have{tiA1_LM} ub_rhoML: '[rhoM psi] + '[rhoL psi] < 1. + have [[Itau1 Ztau1] _] := cohS. + have n1psi: '[psi] = 1 by rewrite Itau1 ?mem_zchar ?irrWnorm ?irrS. + rewrite -n1psi (cfnormE (cfun_onG psi)) (big_setD1 1%g) ?group1 //=. + rewrite mulrDr ltr_spaddl 1?mulr_gt0 ?invr_gt0 ?gt0CG ?exprn_gt0 //. + have /dirrP[s [i ->]]: psi \in dirr G. + by rewrite dirrE Ztau1 ?mem_zchar ?n1psi /=. + by rewrite cfunE normrMsign normr_gt0 irr1_neq0. + rewrite (big_setID 'A1~(M)) mulrDr ler_add //=. + rewrite FTsupp1_type1 // -FT_DadeF_supportE. + by rewrite (setIidPr _) ?Dade_support_subD1 ?leC_cfnorm_invDade_support. + rewrite (big_setID 'A1~(L)) mulrDr ler_paddr //=. + rewrite mulr_ge0 ?invr_ge0 ?ler0n ?sumr_ge0 // => z _. + by rewrite exprn_ge0 ?normr_ge0. + rewrite (setIidPr _); last first. + by rewrite subsetD tiA1_LM -FT_Dade1_supportE Dade_support_subD1. + by rewrite FTsupp1_type1 // -FT_DadeF_supportE leC_cfnorm_invDade_support. +have ubM: (#|M| <= #|K| * #|H|)%N. + by rewrite -(sdprod_card defM) leq_mul // subset_leq_card. +have{lb_psiM lb_psiL ub_rhoML ubM} ubK: (#|K / K'|%g < 4)%N. + rewrite card_quotient ?gFnorm -?ltC_nat //. + rewrite -ltf_pinv ?qualifE ?gt0CiG ?ltr0n // natf_indexg ?gFsub //. + rewrite invfM invrK mulrC -(subrK #|K|%:R #|K'|%:R) mulrDl divff ?neq0CG //. + rewrite -opprB mulNr addrC ltr_subr_addl -ltr_subr_addr. + have /Frobenius_context[_ _ ntE _ _] := set_Frobenius_compl defL frobL. + have egt2: (2 < e)%N by rewrite odd_geq ?mFT_odd ?cardG_gt1. + have e1_gt0: 0 < e.-1%:R :> algC by rewrite ltr0n -(subnKC egt2). + apply: ltr_le_trans (_ : e%:R / e.-1%:R ^+ 2 <= _). + rewrite ltr_pdivl_mulr ?exprn_gt0 //. + rewrite -(@ltr_pmul2r _ #|H|%:R^-1) ?invr_gt0 ?gt0CG // mulrAC. + rewrite -(ltr_add2r 1) -ltr_subl_addl -addrA. + apply: ler_lt_trans ub_rhoML; rewrite ler_add //. + apply: ler_trans lb_psiM; rewrite -natrX ler_wpmul2r ?ler0n //. + rewrite cardsD (setIidPr _) ?gFsub // -natrB ?subset_leq_card ?gFsub //. + rewrite -mulrA ler_wpmul2l ?ler0n //. + rewrite ler_pdivr_mulr ?gt0CG // ler_pdivl_mull ?gt0CG //. + by rewrite ler_pdivr_mulr ?gt0CG // mulrC -natrM leC_nat. + rewrite -(ler_pmul2l (gt0CG E)) -/e mulrA -expr2 invfM -exprMn. + apply: ler_trans (_ : (1 + 2%:R^-1) ^+ 2 <= _). + rewrite ler_sqr ?rpred_div ?rpredD ?rpred1 ?rpredV ?rpred_nat //. + rewrite -{1}(ltn_predK egt2) mulrSr mulrDl divff ?gtr_eqF // ler_add2l. + rewrite ler_pdivr_mulr // ler_pdivl_mull ?ltr0n //. + by rewrite mulr1 leC_nat -(subnKC egt2). + rewrite -(@ler_pmul2r _ (2 ^ 2)%:R) ?ltr0n // {1}natrX -exprMn -mulrA. + rewrite mulrDl mulrBl !mul1r !mulVf ?pnatr_eq0 // (mulrSr _ 3) addrK. + by rewrite -mulrSr ler_wpmul2r ?ler0n ?ler_nat. +have [U [MtypeF _]] := FTtypeP 1 maxM Mtype1. +have{U MtypeF} [_ _ [U0 [sU0ML expU0 frobU0]]] := compl_of_typeF defM MtypeF. +have [/sdprodP[_ _ nKU0 tiKU0] ntK _ _ _] := Frobenius_context frobU0. +have nK'U0: U0 \subset 'N(K') := char_norm_trans (der_char 1 K) nKU0. +have frobU0K': [Frobenius K <*> U0 / K' = (K / K') ><| (U0 / K')]%g. + have solK: solvable K by rewrite ?nilpotent_sol ?Fcore_nil. + rewrite Frobenius_proper_quotient ?(sol_der1_proper solK) // /(_ <| _). + by rewrite (subset_trans (der_sub 1 _)) ?joing_subl // join_subG gFnorm. +have isoU0: U0 \isog U0 / K'. + by rewrite quotient_isog //; apply/trivgP; rewrite -tiKU0 setSI ?gFsub. +have piU0p: p \in \pi(U0 / K')%g. + rewrite /= -(card_isog isoU0) -pi_of_exponent expU0 pi_of_exponent. + rewrite mem_primes pr_p cardG_gt0 /= -ox order_dvdG // (subsetP _ _ P01x) //. + rewrite (subset_trans (Ohm_sub 1 _)) // subsetI sP0M. + by rewrite (subset_trans sP0H) ?gFsub. +have /(Cauchy pr_p)[z U0z oz]: p %| #|U0 / K'|%g. + by rewrite mem_primes in piU0p; case/and3P: piU0p. +have frobKz: [Frobenius (K / K') <*> <[z]> = (K / K') ><| <[z]>]%g. + rewrite (Frobenius_subr _ _ frobU0K') ?cycle_subG //. + by rewrite cycle_eq1 -order_gt1 oz ltnW. +have: p %| #|K / K'|%g.-1 by rewrite -oz (Frobenius_dvd_ker1 frobKz) //. +have [_ ntKK' _ _ _] := Frobenius_context frobKz. +rewrite -subn1 gtnNdvd ?subn_gt0 ?cardG_gt1 // subn1 prednK ?cardG_gt0 //. +by rewrite -ltnS (leq_trans ubK). +Qed. + +End Twelve_13_to_16. + +Lemma FTtype1_nonFrobenius_contradiction : False. +Proof. +have [_ [tau1 cohS]] := FT_Frobenius_coherence maxL frobL. +have [chi] := FTtype1_ref_irr maxL; rewrite -(index_sdprod defL). +exact: FTtype1_nonFrobenius_witness_contradiction cohS. +Qed. + +End Twelve_8_to_16. + +(* This is Peterfalvi, Theorem (12.7). *) +Theorem FTtype1_Frobenius M : + M \in 'M -> FTtype M == 1%N -> [Frobenius M with kernel M`_\F]. +Proof. +set K := M`_\F => maxM Mtype1; have [U [MtypeF _]] := FTtypeP 1 maxM Mtype1. +have hallU: \pi(K)^'.-Hall(M) U. + by rewrite -(compl_pHall U (Fcore_Hall M)) sdprod_compl; have [[]] := MtypeF. +apply: FrobeniusWker (U) _ _; have{MtypeF} [_ -> _] := typeF_context MtypeF. +apply/forall_inP=> P0 /SylowP[p _ sylP0]. +rewrite (odd_pgroup_rank1_cyclic (pHall_pgroup sylP0)) ?mFT_odd // leqNgt. +apply/negP=> prankP0. +have piUp: p \in \pi(U) by rewrite -p_rank_gt0 -(p_rank_Sylow sylP0) ltnW. +have{piUp} K'p: p \in \pi(K)^' := pnatPpi (pHall_pgroup hallU) piUp. +have{U hallU sylP0} sylP0: p.-Sylow(M) P0 := subHall_Sylow hallU K'p sylP0. +have{P0 sylP0 prankP0} prankM: (1 < 'r_p(M))%N by rewrite -(p_rank_Sylow sylP0). +case/negP: K'p => /=. +elim: {p}_.+1 {-2}p M @K (ltnSn p) maxM Mtype1 prankM => // p IHp q M K. +rewrite ltnS leq_eqVlt => /predU1P[->{q} | /(IHp q M)//] maxM Mtype1 prankM. +apply/idPn; rewrite -p'groupEpi /= -/K => p'K. +have [P0 sylP0] := Sylow_exists p M. +have [] := non_Frobenius_FTtype1_witness maxM Mtype1 prankM p'K sylP0. +move=> abP0 prankP0 [L [maxL sP0_Ls [x P01s_x []]]]. +exact: (FTtype1_nonFrobenius_contradiction IHp) P01s_x. +Qed. + +(* This is Peterfalvi, Theorem (12.17). *) +Theorem not_all_FTtype1 : ~~ all_FTtype1 gT. +Proof. +apply/negP=> allT1; pose k := #|'M^G|. +have [partGpi coA1 _ [injA1 /(_ allT1)partG _]] := FT_Dade_support_partition gT. +move/forall_inP in allT1. +have [/subsetP maxMG _ injMG exMG] := mmax_transversalP gT. +have{partGpi exMG} kge2: (k >= 2)%N. + have [L MG_L]: exists L, L \in 'M^G. + by have [L maxL] := any_mmax gT; have [x] := exMG L maxL; exists (L :^ x)%G. + have maxL := maxMG L MG_L; have Ltype1 := allT1 L maxL. + have /Frobenius_kerP[_ ltHL nsHL _] := FTtype1_Frobenius maxL Ltype1. + rewrite ltnNge; apply: contra (proper_subn ltHL) => leK1. + rewrite (sub_normal_Hall (Fcore_Hall L)) // (pgroupS (subsetT L)) //=. + apply: sub_pgroup (pgroup_pi _) => p; rewrite partGpi => /exists_inP[M maxM]. + have /eqP defMG: [set L] == 'M^G by rewrite eqEcard sub1set MG_L cards1. + have [x] := exMG M maxM; rewrite -defMG => /set1P/(canRL (actK 'JG _))-> /=. + by rewrite FTcoreJ cardJg FTcore_type1. +pose L (i : 'I_k) : {group gT} := enum_val i; pose H i := (L i)`_\F%G. +have MG_L i: L i \in 'M^G by apply: enum_valP. +have maxL i: L i \in 'M by apply: maxMG. +have defH i: (L i)`_\s = H i by rewrite FTcore_type1 ?allT1. +pose frobL_P i E := [Frobenius L i = H i ><| gval E]. +have /fin_all_exists[E frobHE] i: exists E, frobL_P i E. + by apply/existsP/FTtype1_Frobenius; rewrite ?allT1. +have frobL i: [/\ L i \subset G, solvable (L i) & frobL_P i (E i)]. + by rewrite subsetT mmax_sol. +have{coA1} coH_ i j: i != j -> coprime #|H i| #|H j|. + move=> j'i; rewrite -!defH coA1 //; apply: contra j'i => /imsetP[x Gx defLj]. + apply/eqP/enum_val_inj; rewrite -/(L i) -/(L j); apply: injMG => //. + by rewrite defLj; apply/esym/orbit_act. +have tiH i: normedTI (H i)^# G (L i). + have ntA: (H i)^# != set0 by rewrite setD_eq0 subG1 mmax_Fcore_neq1. + apply/normedTI_memJ_P=> //=; rewrite subsetT; split=> // x z H1x Gz. + apply/idP/idP=> [H1xz | Lz]; last first. + by rewrite memJ_norm // (subsetP _ z Lz) // normD1 gFnorm. + have /subsetP sH1A0: (H i)^# \subset 'A0(L i) by apply: Fcore_sub_FTsupp0. + have [/(sub_in2 sH1A0)wccH1 [_ maxN] Nfacts] := FTsupport_facts (maxL i). + suffices{z Gz H1xz wccH1} sCxLi: 'C[x] \subset L i. + have /imsetP[y Ly defxz] := wccH1 _ _ H1x H1xz (mem_imset _ Gz). + rewrite -[z](mulgKV y) groupMr // (subsetP sCxLi) // !inE conjg_set1. + by rewrite conjgM defxz conjgK. + apply/idPn=> not_sCxM; pose D := [set y in 'A0(L i) | ~~ ('C[y] \subset L i)]. + have Dx: x \in D by rewrite inE sH1A0. + have{maxN} /mem_uniq_mmax[maxN sCxN] := maxN x Dx. + have Ntype1 := allT1 _ maxN. + have [_ _ /setDP[/bigcupP[y NFy /setD1P[ntx cxy]] /negP[]]] := Nfacts x Dx. + rewrite FTsupp1_type1 Ntype1 // in NFy cxy *. + have /Frobenius_kerP[_ _ _ regFN] := FTtype1_Frobenius maxN Ntype1. + by rewrite 2!inE ntx (subsetP (regFN y NFy)). +have /negP[] := no_coherent_Frobenius_partition (mFT_odd _) kge2 frobL tiH coH_. +rewrite eqEsubset sub1set !inE andbT; apply/andP; split; last first. + apply/bigcupP=> [[i _ /imset2P[x y /setD1P[ntx _] _ Dxy]]]. + by rewrite -(conjg_eq1 x y) -Dxy eqxx in ntx. +rewrite subDset setUC -subDset -(cover_partition partG). +apply/bigcupsP=> _ /imsetP[Li MG_Li ->]; pose i := enum_rank_in MG_Li Li. +rewrite (bigcup_max i) //=; have ->: Li = L i by rewrite /L enum_rankK_in. +rewrite -FT_Dade1_supportE //; apply/bigcupsP=> x A1x; apply: imset2S => //. +move: (FT_Dade1_hyp _) (tiH i); rewrite -defH => _ /Dade_normedTI_P[_ -> //]. +by rewrite mul1g sub1set -/(H i) -defH. +Qed. + +End PFTwelve. diff --git a/mathcomp/odd_order/PFsection13.v b/mathcomp/odd_order/PFsection13.v new file mode 100644 index 0000000..58e0142 --- /dev/null +++ b/mathcomp/odd_order/PFsection13.v @@ -0,0 +1,2185 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime binomial ssralg poly finset. +Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +Require Import gfunctor gproduct center cyclic commutator gseries nilpotent. +Require Import pgroup sylow hall abelian maximal frobenius. +Require Import matrix mxalgebra mxrepresentation mxabelem vector. +Require Import BGsection1 BGsection3 BGsection7. +Require Import BGsection14 BGsection15 BGsection16. +Require Import ssrnum rat algC cyclotomic algnum. +Require Import classfun character integral_char inertia vcharacter. +Require Import PFsection1 PFsection2 PFsection3 PFsection4. +Require Import PFsection5 PFsection6 PFsection7 PFsection8 PFsection9. +Require Import PFsection10 PFsection11 PFsection12. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 13: The Subgroups S and T. *) +(* The following definitions will be used in PFsection14: *) +(* FTtypeP_bridge StypeP j == a virtual character of S that mixes characters *) +(* (locally) beta_ j, betaS that do and do not contain P = S`_\F in their *) +(* kernels, for StypeP : of_typeP S U defW. *) +(* := 'Ind[S, P <*> W1] 1 - mu2_ 0 j. *) +(* FTtypeP_bridge_gap StypeP == the difference between the image of beta_ j *) +(* (locally) Gamma, GammaS under the Dade isometry for S, and its natural *) +(* value, 1 - eta_ 0 j (this does not actually *) +(* depend on j != 0). *) +(* The following definitions are only used locally across sections: *) +(* #1 == the irreducible index 1 (i.e., inord 1). *) +(* irr_Ind_Fittinq S chi <=> chi is an irreducible character of S induced *) +(* (locally) irrIndH from an irreducible character of 'F(S) (which *) +(* will be linear here, as 'F(S) is abelian). *) +(* typeP_TIred_coherent StypeP tau1 <=> tau1 maps the reducible induced *) +(* characters mu_ j of a type P group S, which are *) +(* the image under the cyclic TI isometry to S of *) +(* row sums of irreducibles of W = W1 x W2, to *) +(* the image of that sum under the cyclic TI *) +(* isometry to G (except maybe for a sign change *) +(* if p = #|W2| = 3). *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory FinRing.Theory Num.Theory. + +Section Thirteen. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types (p q : nat) (x y z : gT). +Implicit Types H K L N P Q R S T U W : {group gT}. + +Definition irr_Ind_Fitting S := [predI irr S & seqIndT 'F(S) S]. + +Local Notation irrIndH := (irr_Ind_Fitting _). +Local Notation "#1" := (inord 1) (at level 0). + +Section Thirteen_2_3_5_to_9. + +(* These assumptions correspond to the part of Peterfalvi, Hypothesis (13.1) *) +(* that is used to prove (13.2-3) and (13.5-9). Because of the shortcomings *) +(* of Coq's Section and Module features we will need to repeat most of these *) +(* assumptions twice down this file to exploit the symmetry between S and T. *) +(* We anticipate the use of the letter 'H' to designate the Fitting group *) +(* of S, which Peterfalvi does only locally in (13.5-9), in order not to *) +(* conflict with (13.17-19), where H denotes the F-core of a Frobenius group. *) +(* This is not a problem for us, since these lemmas will only appear in the *) +(* last section of this file, and we will have no use for H at that point *) +(* since we will have shown in (13.12) that H coincides with P = S`_\F. *) + +Variables S U W W1 W2 : {group gT}. +Hypotheses (maxS : S \in 'M) (defW : W1 \x W2 = W). +Hypotheses (StypeP : of_typeP S U defW). + +Local Notation "` 'W1'" := (gval W1) (at level 0, only parsing) : group_scope. +Local Notation "` 'W2'" := (gval W2) (at level 0, only parsing) : group_scope. +Local Notation "` 'W'" := (gval W) (at level 0, only parsing) : group_scope. +Local Notation V := (cyclicTIset defW). + +Local Notation "` 'S'" := (gval S) (at level 0, only parsing) : group_scope. +Local Notation P := `S`_\F%G. +Local Notation "` 'P'" := `S`_\F (at level 0) : group_scope. +Local Notation PU := S^`(1)%G. +Local Notation "` 'PU'" := `S^`(1) (at level 0) : group_scope. +Local Notation "` 'U'" := (gval U) (at level 0, only parsing) : group_scope. +Local Notation C := 'C_U(`P)%G. +Local Notation "` 'C'" := 'C_`U(`P) (at level 0) : group_scope. +Local Notation H := 'F(S)%G. +Local Notation "` 'H'" := 'F(`S) (at level 0) : group_scope. + +Let defS : PU ><| W1 = S. Proof. by have [[]] := StypeP. Qed. +Let defPU : P ><| U = PU. Proof. by have [_ []] := StypeP. Qed. +Let defH : P \x C = H. Proof. by have [] := typeP_context StypeP. Qed. + +Let notStype1 : FTtype S != 1%N. Proof. exact: FTtypeP_neq1 StypeP. Qed. +Let notStype5 : FTtype S != 5%N. Proof. exact: FTtype5_exclusion maxS. Qed. + +Let pddS := FT_prDade_hypF maxS StypeP. +Let ptiWS : primeTI_hypothesis S PU defW := FT_primeTI_hyp StypeP. +Let ctiWG : cyclicTI_hypothesis G defW := pddS. + +Let ntW1 : W1 :!=: 1. Proof. by have [[]] := StypeP. Qed. +Let ntW2 : W2 :!=: 1. Proof. by have [_ _ _ []] := StypeP. Qed. +Let cycW1 : cyclic W1. Proof. by have [[]] := StypeP. Qed. +Let cycW2 : cyclic W2. Proof. by have [_ _ _ []] := StypeP. Qed. + +Let p := #|W2|. +Let q := #|W1|. +Let c := #|C|. +Let u := #|U : C|. + +Let oU : #|U| = (u * c)%N. Proof. by rewrite mulnC Lagrange ?subsetIl. Qed. + +Let pr_p : prime p. Proof. by have [] := FTtypeP_primes maxS StypeP. Qed. +Let pr_q : prime q. Proof. by have [] := FTtypeP_primes maxS StypeP. Qed. + +Let qgt2 : q > 2. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. +Let pgt2 : p > 2. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. + +Let coPUq : coprime #|PU| q. +Proof. by rewrite (coprime_sdprod_Hall_r defS); have [[]] := StypeP. Qed. + +Let nirrW1 : #|Iirr W1| = q. Proof. by rewrite card_Iirr_cyclic. Qed. +Let nirrW2 : #|Iirr W2| = p. Proof. by rewrite card_Iirr_cyclic. Qed. +Let NirrW1 : Nirr W1 = q. Proof. by rewrite -nirrW1 card_ord. Qed. +Let NirrW2 : Nirr W2 = p. Proof. by rewrite -nirrW2 card_ord. Qed. + +Local Open Scope ring_scope. + +Let sigma := (cyclicTIiso ctiWG). +Let w_ i j := (cyclicTIirr defW i j). +Local Notation eta_ i j := (sigma (w_ i j)). + +Local Notation Imu2 := (primeTI_Iirr ptiWS). +Let mu2_ i j := primeTIirr ptiWS i j. +Let mu_ := primeTIred ptiWS. +Local Notation chi_ j := (primeTIres ptiWS j). + +Local Notation Idelta := (primeTI_Isign ptiWS). +Local Notation delta_ j := (primeTIsign ptiWS j). + +Local Notation tau := (FT_Dade0 maxS). +Local Notation "chi ^\tau" := (tau chi). + +Let calS0 := seqIndD PU S S`_\s 1. +Let rmR := FTtypeP_coh_base maxS StypeP. +Let scohS0 : subcoherent calS0 tau rmR. +Proof. exact: FTtypeP_subcoherent StypeP. Qed. + +Let calS := seqIndD PU S P 1. +Let sSS0 : cfConjC_subset calS calS0. +Proof. exact/seqInd_conjC_subset1/Fcore_sub_FTcore. Qed. + +Local Notation type34ker1 := (FTtype34_Fcore_kernel_trivial maxS StypeP). +Local Notation type34facts := (FTtype34_structure maxS StypeP). +Local Notation type2facts := (FTtypeII_prime_facts maxS StypeP). +Local Notation compl2facts := (compl_of_typeII maxS StypeP). +Local Notation compl3facts := (compl_of_typeIII maxS StypeP). + +Local Notation H0 := (Ptype_Fcore_kernel StypeP). + +Lemma Ptype_factor_prime : pdiv #|P / H0|%g = p. +Proof. exact: def_Ptype_factor_prime. Qed. +Local Notation pHbar_p := Ptype_factor_prime. + +Lemma Ptype_Fcore_kernel_trivial : H0 :=: 1%g. +Proof. +have [/type2facts[_ oP _]| /type34ker1[]//] := boolP (FTtype S == 2). +have [/and3P[]] := Ptype_Fcore_kernel_exists maxS StypeP notStype5. +case/maxgroupp/andP=> /proper_sub sH0P nH0S /subset_trans/(_ nH0S)nH0P _ _. +apply: card1_trivg; rewrite -(divg_indexS sH0P) -card_quotient //. +have [_ _ ->] := Ptype_Fcore_factor_facts maxS StypeP notStype5. +by rewrite pHbar_p -{}oP divnn ?cardG_gt0. +Qed. +Local Notation H0_1 := Ptype_Fcore_kernel_trivial. + +Lemma Ptype_Fcompl_kernel_cent : Ptype_Fcompl_kernel StypeP :=: C. +Proof. +rewrite [Ptype_Fcompl_kernel StypeP]unlock /= (group_inj H0_1). +by rewrite astabQ -morphpreIim -injm_cent ?injmK ?ker_coset ?norms1. +Qed. +Local Notation CHbar_C := Ptype_Fcompl_kernel_cent. + +(* This is Peterfalvi (13.2). *) +Lemma FTtypeP_facts : + [/\ (*a*) [/\ pred2 2 3 (FTtype S), q < p -> FTtype S == 2, + [Frobenius U <*> W1 = U ><| W1] & abelian U], + (*b*) p.-abelem P /\ #|P| = p ^ q, + (*c*) u <= (p ^ q).-1 %/ p.-1, + (*d*) coherent calS S^# tau + & (*e*) normedTI 'A0(S) G S /\ {in 'CF(S, 'A0(S)), tau =1 'Ind}]%N. +Proof. +have type23: pred2 2 3 (FTtype S). + by rewrite /= -implyNb; apply/implyP=> /type34facts[_ _ [->]]. +have [_ ntU _ tiFS] := compl_of_typeII_IV maxS StypeP notStype5. +have [_ /mulG_sub[_ sUPU] nPU tiPU] := sdprodP defPU. +have cUU: abelian U by case/orP: type23 => [/compl2facts | /compl3facts] [_ ->]. +split. +- split=> //; last exact: Ptype_compl_Frobenius StypeP _. + by rewrite ltnNge; apply: contraR => /type34facts[_ /ltnW]. +- by have [/type2facts[] | /type34ker1[]] := boolP (FTtype S == 2). +- have ->: u = #|U / C|%g by rewrite card_quotient ?normsI ?normG ?norms_cent. + have p1gt0: (0 < p.-1)%N by rewrite -(subnKC pgt2). + have [/typeP_Galois_P[]| /typeP_Galois_Pn[]]// := boolP (typeP_Galois StypeP). + move=> _ _ [_ _]; rewrite pHbar_p CHbar_C // -/u -/q; apply: dvdn_leq. + by rewrite divn_gt0 // -!subn1 leq_sub2r // (leq_exp2l 1) ltnW // ltnW. + rewrite -/q CHbar_C pHbar_p => H1 [_ _ _ _ [agt1 a_dv_p1 _ [V /card_isog->]]]. + apply: leq_trans (_ : p.-1 ^ q.-1 <= _)%N; last first. + have ltp1q: (p.-1 ^ q < p ^ q)%N by rewrite ltn_exp2r ?prednK // 2?ltnW. + by rewrite leq_divRL // -expnSr (ltn_predK qgt2) -ltnS (ltn_predK ltp1q). + rewrite dvdn_leq ?expn_gt0 ?p1gt0 // (dvdn_trans (cardSg (subsetT V))) //. + by rewrite cardsT card_matrix mul1n dvdn_exp2r //= card_ord Zp_cast. +- have:= Ptype_core_coherence maxS StypeP notStype5; rewrite H0_1 CHbar_C. + by rewrite (derG1P (abelianS _ cUU)) ?subsetIl ?(group_inj (joing1G _)). +have ntA0: 'A0(S) != set0 := FTsupp0_neq0 maxS. +suffices tiA0: normedTI 'A0(S) G S by split=> //; apply: Dade_Ind. +apply/normedTI_memJ_P=> //; rewrite subsetT; split=> // x g A0x Gg. +apply/idP/idP=> [A0xg | /(subsetP (FTsupp0_norm S))/memJ_norm->//]. +apply/idPn=> S'g; have Dx: x \in [set y in 'A0(S) | ~~ ('C[y] \subset S)]. + rewrite inE A0x; have [_ _ [_ _ _ wccA0 _] _] := pddS. + have /imsetP[y Sy Dxy]: x ^ g \in x ^: S by rewrite wccA0 // mem_orbit. + apply/subsetPn; exists (g * y^-1)%g; last by rewrite groupMr ?groupV. + by rewrite !inE conjg_set1 conjgM Dxy conjgK. +have [_ [_ /(_ x Dx) defL] /(_ x Dx)[_ _]] := FTsupport_facts maxS. +have{defL} [maxL _] := mem_uniq_mmax defL; set L := 'N[x] in maxL *. +rewrite -mem_iota !inE => ALx [/orP[Ltype1 _ | Ltype2]]; last first. + by case/(_ _)/existsP=> // ? /Frobenius_of_typeF/(typePF_exclusion StypeP). +have /Frobenius_kerP[_ _ _ regLF_L] := FTtype1_Frobenius maxL Ltype1. +case/andP: ALx => A1'x /bigcupP[z A1z /setD1P[ntx cLz_z]]; case/negP: A1'x. +rewrite ntx /'A1(L) -(Fcore_eq_FTcore _ _) ?(eqP Ltype1) //= in cLz_z A1z *. +exact: subsetP (regLF_L z A1z) _ cLz_z. +Qed. + +Lemma FTseqInd_TIred j : j != 0 -> mu_ j \in calS. +Proof. +move=> nz_j; rewrite -[mu_ j]cfInd_prTIres mem_seqInd ?gFnormal ?normal1 //=. +by rewrite !inE sub1G (cfker_prTIres pddS). +Qed. + +Lemma FTtypeP_Fitting_abelian : abelian H. +Proof. +rewrite -(dprodW defH) abelianM subsetIr. +have [[_ _ _ cUU] [/abelem_abelian-> _] _ _ _] := FTtypeP_facts. +by rewrite (abelianS _ cUU) ?subsetIl. +Qed. +Hint Resolve FTtypeP_Fitting_abelian. + +Local Notation calH := (seqIndT H S). + +Lemma FTtypeP_Ind_Fitting_1 lambda : lambda \in calH -> lambda 1%g = (u * q)%:R. +Proof. +case/seqIndP=> i _ ->; rewrite cfInd1 -?divgS ?gFsub //; set theta := 'chi_i. +have Ltheta: theta \is a linear_char by apply/char_abelianP. +rewrite -(sdprod_card defS) -(sdprod_card defPU) -/q -(dprod_card defH) oU. +by rewrite -mulnA divnMl // mulnAC mulnK ?cardG_gt0 // lin_char1 ?mulr1. +Qed. +Local Notation calHuq := FTtypeP_Ind_Fitting_1. + +(* This is Peterfalvi (13.3)(a). *) +Lemma FTprTIred_Ind_Fitting j : j != 0 -> mu_ j \in calH. +Proof. +move=> nz_j; have [//|_ _ _] := typeP_reducible_core_Ind maxS StypeP. +rewrite (group_inj H0_1) CHbar_C -/q /= (dprodWY defH) -/calS => /(_ (mu_ j)). +case=> [|_ _ [_ /lin_char_irr/irrP[r ->] ->]]; last exact: mem_seqIndT. +by rewrite mem_filter /= prTIred_not_irr FTseqInd_TIred. +Qed. +Local Notation Hmu := FTprTIred_Ind_Fitting. + +Lemma FTprTIred1 j : j != 0 -> mu_ j 1%g = (u * q)%:R. +Proof. by move/Hmu/calHuq. Qed. +Local Notation mu1uq := FTprTIred1. + +(* This is the first assertion of Peterfalvi (13.3)(c). *) +Lemma FTprTIsign j : delta_ j = 1. +Proof. +have [[_ _ frobUW1 cUU] _ _ cohS _] := FTtypeP_facts. +have [-> | nz_j] := eqVneq j 0; first exact: prTIsign0. +suffices: (1 == delta_ j %[mod q])%C. + rewrite signrE /eqCmod addrC opprB subrK dvdC_nat. + by case: (Idelta j); rewrite ?subr0 // gtnNdvd. +apply: eqCmod_trans (prTIirr1_mod ptiWS 0 j); rewrite -/(mu2_ 0 j) -/q. +have ->: mu2_ 0 j 1%g = u%:R. + by apply: (mulfI (neq0CG W1)); rewrite -prTIred_1 -/mu_ mu1uq // mulnC natrM. +rewrite eqCmod_sym /eqCmod -(@natrB _ u 1) ?indexg_gt0 // subn1 dvdC_nat. +have nC_UW1: U <*> W1 \subset 'N(C). + have /sdprodP[_ _ nPUW1 _] := Ptype_Fcore_sdprod StypeP. + by rewrite normsI ?norms_cent // join_subG normG; have [_ []] := StypeP. +have coUq: coprime #|U| q by have /sdprod_context[_ /coprimeSg->] := defPU. +have /Frobenius_dvd_ker1: [Frobenius U <*> W1 / C = (U / C) ><| (W1 / C)]. + have [defUW1 _ _ _ _] := Frobenius_context frobUW1. + rewrite Frobenius_coprime_quotient // /normal ?subIset ?joing_subl //. + split=> [|x /(Frobenius_reg_ker frobUW1)->]; last exact: sub1G. + rewrite properEneq subsetIl -CHbar_C andbT. + by have [] := Ptype_Fcore_factor_facts maxS StypeP notStype5. +have [nCU nCW1] := joing_subP nC_UW1; rewrite !card_quotient // -/u. +by rewrite -indexgI setIC setIAC (coprime_TIg coUq) setI1g indexg1. +Qed. +Local Notation delta1 := FTprTIsign. + +(* This is Peterfalvi (13.3)(b). *) +Lemma FTtypeP_no_Ind_Fitting_facts : + ~~ has irrIndH calS -> + [/\ typeP_Galois StypeP, `C = 1%g & u = (p ^ q).-1 %/ p.-1]. +Proof. +move=> noIndH; have [[_ _ _ cUU] _ _ _ _] := FTtypeP_facts. +have [[t []] | [->]] := typeP_reducible_core_cases maxS StypeP notStype5. + rewrite CHbar_C H0_1 (derG1P (abelianS _ cUU)) ?subsetIl //=. + rewrite (group_inj (joing1G 1)) -/calS /= (dprodWY defH) => calSt _. + case=> _ /lin_char_irr/irrP[r ->] Dt; case/hasP: noIndH. + by exists 'chi_t; rewrite //= mem_irr; apply/seqIndP; exists r; rewrite ?inE. +rewrite /= pHbar_p H0_1 oU /c => frobPU _ <- _ /=. +suffices /eqP->: C :==: 1%g by rewrite cards1 muln1. +suffices: 'C_(U / 1)(P / 1) == 1%g. + by rewrite -injm_subcent ?morphim_injm_eq1 ?norms1 ?ker_coset. +have [_ ntP _ _ _] := Frobenius_context frobPU. +by rewrite (cent_semiregular (Frobenius_reg_compl frobPU)). +Qed. + +(* Helper function for (13.3)(c). *) +Let signW2 (b : bool) := iter b (@conjC_Iirr _ W2). + +Let signW2K b : involutive (signW2 b). +Proof. by case: b => //; apply: conjC_IirrK. Qed. + +Let signW2_eq0 b : {mono signW2 b: j / j == 0}. +Proof. by case: b => //; apply: conjC_Iirr_eq0. Qed. + +(* This is a reformulation of the definition condition part of (13.3)(c) that *) +(* better fits its actual use in (13.7), (13.8) and (13.9) (note however that *) +(* the p = 3 part will in fact not be used). *) +Definition typeP_TIred_coherent tau1 := + exists2 b : bool, b -> p = 3 + & forall j, j != 0 -> tau1 (mu_ j) = (-1) ^+ b *: \sum_i eta_ i (signW2 b j). + +(* This is the main part of Peterfalvi (13.3)(c), using the definition above. *) +(* Note that the text glosses over the quantifier inversion in the second use *) +(* of (5.8) in the p = 3 case. We must rule out tau1 (mu_ k) = - tau1 (mu_ j) *) +(* by using the isometry property of tau1 (alternatively, we could use (4.8) *) +(* to compute tau1 (mu_ k) = tau (mu_ k - mu_ j) + tau1 (mu_ j) directly). *) +Lemma FTtypeP_coherence : + exists2 tau1 : {additive 'CF(S) -> 'CF(G)}, + coherent_with calS S^# tau tau1 & typeP_TIred_coherent tau1. +Proof. +have [redS|] := altP (@allP _ [predC irr S] calS). + have [k nz_k] := has_nonprincipal_irr ntW2. + have [_ [tau1 Dtau1]] := uniform_prTIred_coherent pddS nz_k. + set calT := uniform_prTIred_seq pddS k => cohT. + exists tau1; last by exists false => // j _; rewrite /= Dtau1 delta1. + apply: subset_coherent_with cohT => xi Sxi. + have [_ _ /(_ xi)] := typeP_reducible_core_Ind maxS StypeP notStype5. + rewrite (group_inj H0_1) mem_filter redS // => /(_ Sxi)/imageP[j nz_j ->] _. + by rewrite image_f // inE -/mu_ [~~ _]nz_j /= !mu1uq. +rewrite all_predC negbK => /hasP[xi Sxi irr_xi]. +have [_ _ _ [tau1 cohS] _] := FTtypeP_facts; exists tau1 => //. +have [|] := boolP [forall (j | j != 0), tau1 (mu_ j) == \sum_i eta_ i j]. + by move/eqfun_inP=> Dtau1; exists false => // j /Dtau1; rewrite scale1r. +rewrite negb_forall_in => /exists_inP[j nz_j /eqP tau1muj_neq_etaj]. +have:= FTtypeP_coherent_TIred sSS0 cohS irr_xi Sxi (FTseqInd_TIred _). +rewrite -/mu_ -/sigma -/ptiWS => tau1mu; have [dk tau1muj Ddk] := tau1mu j nz_j. +case: Ddk tau1muj => [][-> ->]{dk}; rewrite ?signrN delta1 ?scaleNr scale1r //. +set k := conjC_Iirr j => Dmu tau1muj. +have{Dmu} defIW2 l: l != 0 -> pred2 j k l. + by move=> nz_l; rewrite Dmu ?FTseqInd_TIred ?mu1uq. +have [nz_k j'k]: k != 0 /\ k != j. + rewrite conjC_Iirr_eq0 nz_j -(inj_eq irr_inj) conjC_IirrE. + by rewrite odd_eq_conj_irr1 ?mFT_odd ?irr_eq1. +have /eqP p3: p == 3. + rewrite -nirrW2 (cardD1 0) (cardD1 j) (cardD1 k) !inE nz_j nz_k j'k !eqSS. + by apply/pred0Pn=> [[l /and4P[k'l j'l /defIW2/norP[]]]]. +exists true => // _ /defIW2/pred2P[]->; first by rewrite scaler_sign. +have [[[Itau1 _] _] [d t1muk Dd]] := (cohS, tau1mu k nz_k); move: Dd t1muk. +case=> [][-> ->] => [|_]; rewrite ?signrN delta1 // scale1r. +case/(congr1 (cfdotr (tau1 (mu_ j)) \o -%R))/eqP/idPn => /=. +rewrite -tau1muj cfdotNl eq_sym !Itau1 ?mem_zchar ?FTseqInd_TIred //. +by rewrite !cfdot_prTIred (negPf j'k) eqxx mul1n oppr0 neq0CG. +Qed. + +(* We skip over (13.4), whose proof uses (13.2) and (13.3) for both groups of *) +(* a type P pair. *) + +Let calS1 := seqIndD H S P 1. + +(* Some facts about calS1 used implicitly throughout (13.5-8). *) +Let S1mu j : j != 0 -> mu_ j \in calS1. +Proof. +move=> nz_j; have /seqIndP[s _ Ds] := Hmu nz_j. +rewrite Ds mem_seqInd ?gFnormal ?normal1 // !inE sub1G andbT. +rewrite -(sub_cfker_Ind_irr s (gFsub _ _) (gFnorm _ _)) -Ds /=. +rewrite -[mu_ j](cfInd_prTIres (FT_prDade_hypF maxS StypeP)). +by rewrite sub_cfker_Ind_irr ?cfker_prTIres ?gFsub ?gFnorm. +Qed. + +Let calSirr := [seq phi <- calS | phi \in irr S]. +Let S1cases zeta : + zeta \in calS1 -> {j | j != 0 & zeta = mu_ j} + (zeta \in 'Z[calSirr]). +Proof. +move=> S1zeta; have /sig2_eqW[t /setDP[_ kerP't] Dzeta] := seqIndP S1zeta. +rewrite inE in kerP't; have /mulG_sub[sPH _] := dprodW defH. +have [/andP[sPPU nPPU] sUPU _ _ _] := sdprod_context defPU. +have sHPU: H \subset PU by rewrite /= -(dprodWC defH) mulG_subG subIset ?sUPU. +have [/eqfunP mu'zeta|] := boolP [forall j, '['Ind 'chi_t, chi_ j] == 0]. + right; rewrite Dzeta -(cfIndInd _ _ sHPU) ?gFsub //. + rewrite ['Ind 'chi_t]cfun_sum_constt linear_sum /= rpred_sum // => s tPUs. + rewrite linearZ rpredZ_Cnat ?Cnat_cfdot_char ?cfInd_char ?irr_char //=. + have [[j Ds] | [irr_zeta _]] := prTIres_irr_cases ptiWS s. + by case/eqP: tPUs; rewrite Ds mu'zeta. + rewrite mem_zchar // mem_filter irr_zeta mem_seqInd ?gFnormal ?normal1 //=. + by rewrite !inE sub1G andbT -(sub_cfker_constt_Ind_irr tPUs). +rewrite negb_forall => /existsP/sigW[j]. +rewrite -irr_consttE constt_Ind_Res => jHt. +have nz_j: j != 0; last do [left; exists j => //]. + apply: contraTneq jHt => ->; rewrite prTIres0 rmorph1 -irr0 constt_irr. + by apply: contraNneq kerP't => ->; rewrite irr0 cfker_cfun1. +have /pairwise_orthogonalP[_ ooS1]: pairwise_orthogonal calS1. + by rewrite seqInd_orthogonal ?gFnormal. +rewrite -(cfRes_prTIirr _ 0) cfResRes ?gFsub //= in jHt. +have muj_mu0j: Imu2 (0, j) \in irr_constt (mu_ j). + by rewrite irr_consttE cfdotC cfdot_prTIirr_red eqxx conjC1 oner_eq0. +apply: contraNeq (constt_Res_trans (prTIred_char _ _) muj_mu0j jHt). +by rewrite cfdot_Res_l /= -Dzeta eq_sym => /ooS1-> //; rewrite S1mu. +Qed. + +Let sS1S : {subset calS1 <= 'Z[calS]}. +Proof. +move=> zeta /S1cases[[j nz_j ->]|]; first by rewrite mem_zchar ?FTseqInd_TIred. +by apply: zchar_subset; apply: mem_subseq (filter_subseq _ _). +Qed. + +(* This is Peterfalvi (13.5). *) +(* We have adapted the statement to its actual use by replacing the Dade *) +(* (partial) isometry by a (total) coherent isometry, and strengthening the *) +(* orthogonality condition. This simplifies the assumptions as zeta0 is no *) +(* longer needed. Note that this lemma is only used to establish various *) +(* inequalities (13.6-8) that contribute to (13.10), so it does not need to *) +(* be exported from this section. *) +Let calS1_split1 (tau1 : {additive _}) zeta1 chi : + coherent_with calS S^# tau tau1 -> zeta1 \in calS1 -> chi \in 'Z[irr G] -> + {in calS1, forall zeta, zeta != zeta1 -> '[tau1 zeta, chi] = 0} -> + let a := '[tau1 zeta1, chi] in + exists2 alpha, + alpha \in 'Z[irr H] /\ {subset irr_constt alpha <= Iirr_ker H P} & + [/\ (*a*) {in H^#, forall x, chi x = a / '[zeta1] * zeta1 x + alpha x}, + (*b*) + \sum_(x in H^#) `|chi x| ^+ 2 = + a ^+ 2 / '[zeta1] * (#|S|%:R - zeta1 1%g ^+ 2 / '[zeta1]) + - 2%:R * a * (zeta1 1%g * alpha 1%g / '[zeta1]) + + (\sum_(x in H^#) `|alpha x| ^+ 2) + & (*c*) + \sum_(x in H^#) `|alpha x| ^+ 2 >= #|P|.-1%:R * alpha 1%g ^+ 2]. +Proof. +case=> _ Dtau1 S1zeta1 Zchi o_tau1S_chi a. +have sW2P: W2 \subset P by have [_ _ _ []] := StypeP. +have /mulG_sub[sPH _] := dprodW defH. +have ntH: H :!=: 1%g by apply: subG1_contra ntW2; apply: subset_trans sPH. +have sH1S: H^# \subset G^# by rewrite setSD ?subsetT. +have[nsHS nsPS ns1S]: [/\ H <| S, P <| S & 1 <| S] by rewrite !gFnormal normal1. +have [[sHS nHS] [sPS nPS]] := (andP nsHS, andP nsPS). +have tiH: normedTI H^# G S by have [] := compl_of_typeII_IV maxS StypeP. +have ddH := normedTI_Dade tiH sH1S; have [_ ddH_1] := Dade_normedTI_P ddH tiH. +pose tauH := Dade ddH. +have DtauH: {in 'CF(S, H^#), tauH =1 'Ind} := Dade_Ind ddH tiH. +have sS1H: {subset calS1 <= calH} by apply: seqInd_subT. +pose zeta0 := zeta1^*%CF. +have S1zeta0: zeta0 \in calS1 by rewrite cfAut_seqInd. +have zeta1'0: zeta0 != zeta1. + by rewrite (hasPn (seqInd_notReal _ _ _ _) _ S1zeta1) ?gFnormal ?mFT_odd. +have Hzeta0 := sS1H _ S1zeta0. +have dH_1 zeta: zeta \in calH -> (zeta - zeta0) 1%g == 0. + by move=> Tzeta; rewrite 2!cfunE !calHuq // subrr eqxx. +have H1dzeta zeta: zeta \in calH -> zeta - zeta0 \in 'CF(S, H^#). + have HonH: {subset calH <= 'CF(S, H)} by exact: seqInd_on. + by move=> Hzeta; rewrite cfun_onD1 rpredB ?HonH ?dH_1. +pose calH1 := rem zeta1 (rem zeta0 (filter [mem calS1] calH)). +pose calH2 := filter [predC calS1] calH. +have DcalH: perm_eq calH (zeta0 :: zeta1 :: calH1 ++ calH2). + rewrite -(perm_filterC [mem calS1]) -!cat_cons perm_cat2r. + rewrite (perm_eqlP (@perm_to_rem _ zeta0 _ _)) ?mem_filter /= ?S1zeta0 //. + rewrite perm_cons perm_to_rem // mem_rem_uniq ?filter_uniq ?seqInd_uniq //. + by rewrite !inE mem_filter /= eq_sym zeta1'0 S1zeta1 sS1H. +have{DcalH} [a_ _ Dchi _] := invDade_seqInd_sum ddH chi DcalH. +have Da_ zeta: zeta \in calH -> a_ zeta = '['Ind (zeta - zeta0), chi]. + move=> Tzeta; rewrite /a_ !calHuq // divff ?scale1r; last first. + by rewrite pnatr_eq0 -lt0n muln_gt0 indexg_gt0 cardG_gt0. + by rewrite [Dade _ _]DtauH ?H1dzeta. +have Za_ zeta: zeta \in calH -> a_ zeta \in Cint. + move=> Hzeta; rewrite Da_ // Cint_cfdot_vchar ?cfInd_vchar //. + by rewrite rpredB ?char_vchar ?(seqInd_char Hzeta) ?(seqInd_char Hzeta0). +have{Da_} Da_ zeta: zeta \in calS1 -> a_ zeta = '[tau1 zeta, chi]. + move=> S1zeta; have Hzeta := sS1H _ S1zeta. + rewrite Da_ //; have [_ _ _ _ [_ <-]] := FTtypeP_facts. + rewrite -Dtau1; last by rewrite zcharD1E rpredB ?sS1S ?dH_1. + by rewrite raddfB cfdotBl (o_tau1S_chi zeta0) ?subr0. + by rewrite (cfun_onS (Fitting_sub_FTsupp0 maxS)) ?H1dzeta. +pose alpha := 'Res[H] (\sum_(zeta <- calH2) (a_ zeta)^* / '[zeta] *: zeta). +have{Dchi} Dchi: {in H^#, forall x, chi x = a / '[zeta1] * zeta1 x + alpha x}. + move=> x H1x; have [_ Hx] := setD1P H1x. + transitivity (invDade ddH chi x). + by rewrite cfunElock ddH_1 // big_set1 H1x mul1g cards1 invr1 mul1r. + rewrite cfResE ?gFsub ?Dchi // big_cons conj_Cint ?Za_ ?Da_ ?sS1H //= -/a. + congr (_ + _); rewrite big_cat /= sum_cfunE big1_seq ?add0r //= => [|zeta]. + by apply: eq_bigr => zeta; rewrite cfunE. + rewrite ?(mem_rem_uniq, inE) ?rem_uniq ?filter_uniq ?seqInd_uniq //=. + rewrite mem_filter => /and4P[/= zeta1'z _ S1zeta _]. + by rewrite Da_ ?o_tau1S_chi // conjC0 !mul0r. +have kerHalpha: {subset irr_constt alpha <= Iirr_ker H P}. + move=> s; apply: contraR => kerP's; rewrite [alpha]rmorph_sum cfdot_suml. + rewrite big1_seq // => psi; rewrite mem_filter /= andbC => /andP[]. + case/seqIndP=> r _ ->; rewrite mem_seqInd // !inE sub1G andbT negbK => kerPr. + rewrite cfdot_Res_l cfdotZl mulrC cfdot_sum_irr big1 ?mul0r // => t _. + apply: contraNeq kerP's; rewrite mulf_eq0 fmorph_eq0 inE => /norP[rSt sSt]. + by rewrite (sub_cfker_constt_Ind_irr sSt) -?(sub_cfker_constt_Ind_irr rSt). +have Zalpha: alpha \in 'Z[irr H]. + rewrite [alpha]rmorph_sum big_seq rpred_sum // => zeta; rewrite mem_filter /=. + case/andP=> S1'zeta Tzeta; rewrite linearZ /= -scalerA. + rewrite rpredZ_Cint ?conj_Cint ?Za_ //; have [s _ ->] := seqIndP Tzeta. + rewrite cfResInd_sum_cfclass ?reindex_cfclass -?cfnorm_Ind_irr //=. + rewrite scalerK ?cfnorm_eq0 ?cfInd_eq0 ?irr_neq0 ?irr_char ?gFsub //. + by apply: rpred_sum => i _; apply: irr_vchar. +have{Da_ Za_} Za: a \in Cint by rewrite -[a]Da_ ?Za_ ?sS1H. +exists alpha => //; split=> //. + set a1 := a / _ in Dchi; pose phi := a1 *: 'Res zeta1 + alpha. + transitivity (#|H|%:R * '[phi] - `|phi 1%g| ^+ 2). + rewrite (cfnormE (cfun_onG phi)) mulVKf ?neq0CG // addrC. + rewrite (big_setD1 _ (group1 H)) addKr; apply: eq_bigr => x H1x. + by have [_ Hx] := setD1P H1x; rewrite !cfunE cfResE // Dchi. + have Qa1: a1 \in Creal. + apply: rpred_div; first by rewrite rpred_Cint. + by rewrite rpred_Cnat // Cnat_cfdot_char ?(seqInd_char S1zeta1). + rewrite cfnormDd; last first. + rewrite [alpha]cfun_sum_constt cfdotZl cfdot_sumr big1 ?mulr0 // => s. + move/kerHalpha; rewrite inE cfdotZr mulrC cfdot_Res_l => kerPs. + have [r kerP'r ->] := seqIndP S1zeta1; rewrite cfdot_sum_irr big1 ?mul0r //. + move=> t _; apply: contraTeq kerP'r; rewrite !inE sub1G andbT negbK. + rewrite mulf_eq0 fmorph_eq0 => /norP[]; rewrite -!irr_consttE. + by move=> /sub_cfker_constt_Ind_irr-> // /sub_cfker_constt_Ind_irr <-. + rewrite cfnormZ 2!cfunE cfRes1 2?real_normK //; last first. + rewrite rpredD 1?rpredM // Creal_Cint ?Cint_vchar1 // ?char_vchar //. + by rewrite (seqInd_char S1zeta1). + rewrite mulrDr mulrCA sqrrD opprD addrACA; congr (_ + _); last first. + rewrite (cfnormE (cfun_onG _)) mulVKf ?neq0CG //. + by rewrite (big_setD1 1%g) // Cint_normK ?Cint_vchar1 // addrC addKr. + rewrite opprD addrA; congr (_ - _); last first. + rewrite -[_ * a * _]mulrA -mulr_natl; congr (_ * _). + by rewrite -[a1 * _]mulrA -(mulrA a); congr (_ * _); rewrite -mulrA mulrC. + rewrite mulrBr; congr (_ - _); last first. + by rewrite mulrACA -expr2 -!exprMn mulrAC. + rewrite -mulrA exprMn -mulrA; congr (_ * _); rewrite expr2 -mulrA. + congr (_ * _); apply: canLR (mulKf (cfnorm_seqInd_neq0 nsHS S1zeta1)) _. + rewrite (cfnormE (cfun_onG _)) mulVKf ?neq0CG // mulrC. + rewrite (cfnormE (seqInd_on nsHS S1zeta1)) mulVKf ?neq0CG //. + by apply: eq_bigr => x Hx; rewrite cfResE. +rewrite -subn1 natrB // -Cint_normK ?Cint_vchar1 // mulrBl mul1r ler_subl_addl. +apply: ler_trans (_ : \sum_(x in H) `|alpha x| ^+ 2 <= _); last first. + by rewrite (big_setD1 1%g). +rewrite (big_setID P) /= (setIidPr sPH) ler_paddr ?sumr_ge0 // => [x _|]. + by rewrite mulr_ge0 ?normr_ge0. +rewrite mulr_natl -sumr_const ler_sum // => y Py. +suffices ->: alpha y = alpha 1%g by apply: lerr. +rewrite [alpha]cfun_sum_constt !sum_cfunE; apply: eq_bigr => i. +by rewrite !cfunE => /kerHalpha; rewrite inE => /subsetP/(_ y Py)/cfker1->. +Qed. + +Local Notation eta10 := (eta_ #1 0). +Local Notation eta01 := (eta_ 0 #1). + +Let o_tau1_eta (tau1 : {additive _}) i j: + coherent_with calS S^# tau tau1 -> + {in 'Z[calSirr], forall zeta, '[tau1 zeta, eta_ i j] = 0}. +Proof. +move=> cohS _ /zchar_expansion[|z Zz ->]. + by rewrite filter_uniq ?seqInd_uniq. +rewrite raddf_sum cfdot_suml big1_seq //= => phi; rewrite mem_filter. +case/andP=> irr_phi /(coherent_ortho_cycTIiso StypeP sSS0 cohS) o_phi_eta. +by rewrite raddfZ_Cint {Zz}//= cfdotZl o_phi_eta ?mulr0. +Qed. + +Let P1_int2_lb b : b \in Cint -> 2%:R * u%:R * b <= #|P|.-1%:R * b ^+ 2. +Proof. +move=> Zb; rewrite -natrM; apply: ler_trans (_ : (2 * u)%:R * b ^+ 2 <= _). + by rewrite ler_wpmul2l ?ler0n ?Cint_ler_sqr. +rewrite ler_wpmul2r -?realEsqr ?Creal_Cint // leC_nat mulnC -leq_divRL //. +have [_ [_ ->] /leq_trans-> //] := FTtypeP_facts. +by rewrite leq_div2l // -subn1 ltn_subRL. +Qed. + +(* This is Peterfalvi (13.6). *) +Lemma FTtypeP_sum_Ind_Fitting_lb (tau1 : {additive _}) lambda : + coherent_with calS S^# tau tau1 -> lambda \in irrIndH -> lambda \in calS -> + \sum_(x in H^#) `|tau1 lambda x| ^+ 2 >= #|S|%:R - lambda 1%g ^+ 2. +Proof. +move=> cohS /andP[Ilam Hlam] Slam; have [[Itau1 Ztau1] _] := cohS. +have Zlam1: tau1 lambda \in 'Z[irr G] by rewrite Ztau1 ?mem_zchar. +have S1lam: lambda \in calS1. + have [[s kerP's Ds] [r _ Dr]] := (seqIndP Slam, seqIndP Hlam). + rewrite Dr mem_seqInd ?gFnormal ?normal1 // !inE !sub1G !andbT in kerP's *. + rewrite -(sub_cfker_Ind_irr r (gFsub _ _) (gFnorm _ _)) /= -Dr. + by rewrite Ds sub_cfker_Ind_irr ?gFsub ?gFnorm. +have [|alpha [Zalpha kerPalpha]] := calS1_split1 cohS S1lam Zlam1. + move=> zeta S1zeta lam'zeta; rewrite Itau1 ?sS1S //. + suffices: pairwise_orthogonal calS1 by case/pairwise_orthogonalP=> _ ->. + by rewrite seqInd_orthogonal ?gFnormal. +rewrite Itau1 ?mem_zchar // irrWnorm // expr1n !divr1 mul1r => [[Dlam ->]]. +rewrite mulr1 -ler_subl_addl addrC opprB subrK calHuq //; apply: ler_trans. +have [[x W2x ntx] [y W1y nty]] := (trivgPn _ ntW2, trivgPn _ ntW1). +have [_ _ _ [_ _ sW2P _ _] _] := StypeP; have Px := subsetP sW2P x W2x. +have [eps pr_eps] := C_prim_root_exists (prime_gt0 pr_q). +have{y W1y W2x nty} lamAmod: (tau1 lambda x == lambda x %[mod 1 - eps])%A. + have [_ /mulG_sub[_ sW1S] _ tiPUW1] := sdprodP defS. + have [_ /mulG_sub[sW1W sW2W] cW12 _] := dprodP defW. + have /mulG_sub[sPPU _] := sdprodW defPU. + have [o_y cxy]: #[y] = q /\ x \in 'C[y]. + split; last by apply/cent1P; red; rewrite (centsP cW12). + by apply: nt_prime_order => //; apply/eqP; rewrite -order_dvdn order_dvdG. + have lam1yx: (tau1 lambda (y * x)%g == tau1 lambda x %[mod 1 - eps])%A. + by rewrite (vchar_ker_mod_prim pr_eps) ?in_setT. + have [Sx Sy] := (subsetP (gFsub _ _) x Px, subsetP sW1S y W1y). + have PUx := subsetP sPPU x Px. + have lam_yx: (lambda (y * x)%g == lambda x %[mod 1 - eps])%A. + by rewrite (vchar_ker_mod_prim pr_eps) ?char_vchar ?(seqInd_char Slam). + apply: eqAmod_trans lam_yx; rewrite eqAmod_sym; apply: eqAmod_trans lam1yx. + have PUlam: lambda \in 'CF(S, PU) by rewrite (seqInd_on _ Slam) ?gFnormal. + have PU'yx: (y * x)%g \notin PU. + by rewrite groupMr //= -[y \in PU]andbT -W1y -in_setI tiPUW1 !inE. + rewrite (cfun_on0 PUlam PU'yx) (ortho_cycTIiso_vanish pddS) //. + apply/orthoPl=> _ /mapP[_ /(cycTIirrP defW)[i [j ->]] ->]. + by rewrite (coherent_ortho_cycTIiso StypeP sSS0). + rewrite !inE (groupMl x (subsetP sW1W y _)) // (subsetP sW2W) // andbT. + rewrite groupMl // -[x \in _]andTb -PUx -in_setI tiPUW1 !inE negb_or ntx /=. + by rewrite (contra _ PU'yx) // => /(subsetP sW2P)/(subsetP sPPU). +have{x ntx Px lamAmod} alphaAmod: (alpha 1%g == 0 %[mod 1 - eps])%A. + have Hx: x \in H by have/mulG_sub[/subsetP->] := dprodW defH. + have:= lamAmod; rewrite -[lambda x]addr0 Dlam ?inE ?ntx // mul1r eqAmodDl. + rewrite cfker1 // [alpha]cfun_sum_constt (subsetP (cfker_sum _ _ _)) //. + rewrite !inE Hx (subsetP _ x Px) //; apply/bigcapsP=> i /kerPalpha. + by rewrite !inE => /subset_trans-> //; apply: cfker_scale. +have /dvdCP[b Zb ->]: (q %| alpha 1%g)%C. + by rewrite (int_eqAmod_prime_prim pr_eps) // Cint_vchar1. +rewrite natrM mulrACA exprMn !mulrA 2?ler_pmul2r ?gt0CG //. +by rewrite -[_ * b * b]mulrA P1_int2_lb. +Qed. + +(* This is Peterfalvi (13.7). *) +Lemma FTtypeP_sum_cycTIiso10_lb : \sum_(x in H^#) `|eta10 x| ^+ 2 >= #|H^#|%:R. +Proof. +pose mu1 := mu_ #1; have S1mu1: mu1 \in calS1 by rewrite S1mu ?Iirr1_neq0. +have Zeta10: eta10 \in 'Z[irr G] by rewrite cycTIiso_vchar. +have [tau1 cohS [b _ Dtau1]] := FTtypeP_coherence. +have{b Dtau1} oS1eta10: {in calS1, forall zeta, '[tau1 zeta, eta10] = 0}. + move=> zeta /S1cases[[j nz_j ->] | /o_tau1_eta-> //]. + rewrite Dtau1 // cfdotZl cfdot_suml big1 ?mulr0 // => i _. + by rewrite cfdot_cycTIiso signW2_eq0 (negPf nz_j) andbF. +have [_ /oS1eta10//|alpha [Zalpha kerPalpha]] := calS1_split1 cohS S1mu1 Zeta10. +rewrite {}oS1eta10 // expr0n mulr0 !mul0r subrr add0r => [[Deta10 -> ub_alpha]]. +have{Deta10} Deta10: {in H^#, eta10 =1 alpha}. + by move=> x /Deta10; rewrite !mul0r add0r. +set a1_2 := alpha 1%g ^+ 2 in ub_alpha. +have Dsum_alpha: \sum_(x in H^#) `|alpha x| ^+ 2 = #|H|%:R * '[alpha] - a1_2. + rewrite (cfnormE (cfun_onG _)) mulVKf ?neq0CG // (big_setD1 _ (group1 H)) /=. + by rewrite addrC Cint_normK ?addKr ?Cint_vchar1. +have [/mulG_sub[sPH _] [_ _ _ [_ _ sW2P _ _] _]] := (dprodW defH, StypeP). +have nz_alpha: alpha != 0. + have [[x W2x ntx] [y W1y nty]] := (trivgPn _ ntW2, trivgPn _ ntW1). + have [eps pr_eps] := C_prim_root_exists (prime_gt0 pr_q). + have [_ mulW12 cW12 tiW12] := dprodP defW. + have [sW1W sW2W] := mulG_sub mulW12. + have [o_y cxy]: #[y] = q /\ x \in 'C[y]. + split; last by apply/cent1P; red; rewrite (centsP cW12). + by apply: nt_prime_order => //; apply/eqP; rewrite -order_dvdn order_dvdG. + have eta10x: (eta10 x == eta10 (y * x)%g %[mod 1 - eps])%A. + by rewrite eqAmod_sym (vchar_ker_mod_prim pr_eps) ?in_setT. + have eta10xy: (eta10 (y * x)%g == 1 %[mod 1 - eps])%A. + rewrite cycTIiso_restrict; last first. + rewrite !inE -mulW12 mem_mulg // andbT groupMl ?groupMr // -[_ || _]andTb. + by rewrite andb_orr -{1}W2x -W1y andbC -!in_setI tiW12 !inE (negPf ntx). + have {2}<-: w_ #1 0 x = 1. + rewrite -[x]mul1g /w_ dprod_IirrE cfDprodE // irr0 cfun1E W2x mulr1. + by rewrite lin_char1 ?irr_cyclic_lin. + rewrite (vchar_ker_mod_prim pr_eps) ?(subsetP sW1W y) ?(subsetP sW2W) //. + by rewrite irr_vchar. + have: (alpha x == 1 %[mod 1 - eps])%A. + rewrite -Deta10; last by rewrite !inE ntx (subsetP sPH) ?(subsetP sW2P). + exact: eqAmod_trans eta10x eta10xy. + apply: contraTneq => ->; rewrite cfunE eqAmod_sym. + apply/negP=> /(int_eqAmod_prime_prim pr_eps pr_q (rpred1 _))/idPn[]. + by rewrite (dvdC_nat q 1) -(subnKC qgt2). +apply: wlog_neg => suma_lt_H. +suffices{ub_alpha} lb_a1_2: a1_2 >= #|H^#|%:R. + have Pgt2: (2 < #|P|)%N by apply: leq_trans (subset_leq_card sW2P). + apply: ler_trans (ler_trans lb_a1_2 _) ub_alpha. + rewrite ler_pmull ?(ltr_le_trans _ lb_a1_2) ?ler1n ?ltr0n //. + by rewrite -(subnKC Pgt2). + have:= leq_trans (ltnW Pgt2) (subset_leq_card sPH). + by rewrite (cardsD1 1%g) group1. +have /CnatP[n Dn]: '[alpha] \in Cnat by rewrite Cnat_cfnorm_vchar. +have /CnatP[m Dm]: a1_2 \in Cnat by rewrite Cnat_exp_even ?Cint_vchar1. +rewrite Dm leC_nat leqNgt; apply: contra suma_lt_H => a1_2_lt_H. +rewrite {1}Dsum_alpha Dn Dm -natrM ler_subr_addl (cardsD1 1%g H) group1 /=. +case Dn1: n => [|[|n1]]; first by rewrite -cfnorm_eq0 Dn Dn1 eqxx in nz_alpha. + have /dirrP[b [i Dalpha]]: alpha \in dirr H by rewrite dirrE Zalpha Dn Dn1 /=. + rewrite -Dm /a1_2 Dalpha cfunE exprMn sqrr_sign mul1r muln1 mulrS ler_add2r. + by rewrite lin_char1 ?expr1n //; apply/char_abelianP. +rewrite -natrD leC_nat -add2n mulnDr (addnC 1%N) mulnDl -addnA. +by apply: leq_trans (leq_addr _ _); rewrite muln2 -addnn leq_add2r ltnW. +Qed. + +(* This is Peterfalvi (13.8). *) +(* We have filled a logical gap in the textbook, which quotes (13.3.c) to get *) +(* a j such that eta_01 is a component of mu_j^tau1, then asserts that the *) +(* (orthogonality) assumptions of (13.5) have been checked, apparently *) +(* implying that because for zeta in calS1 \ mu_j, zeta^tau1 is orthogonal to *) +(* mu_j^tau1, as per the proof of (13.6), zeta^tau1 must be orthogonal to *) +(* eta_01. This is wrong, because zeta^tau1, mu_j^tau1 and eta_01 are not *) +(* characters, but virtual characters. We need to use a more careful line of *) +(* reasoning, using the more precise characterization of calS1 in the lemma *) +(* S1cases above (which does use the orthogonal-constituent argument, but *) +(* for chi_j and Res_H zeta), and the decomposition given in (13.3.c) for all *) +(* the mu_k. *) +Lemma FTtypeP_sum_cycTIiso01_lb : + \sum_(x in H^#) `|eta01 x| ^+ 2 >= #|PU|%:R - (u ^ 2)%:R. +Proof. +have [tau1 cohS [b _ Dtau1]] := FTtypeP_coherence. +have Zeta01: eta01 \in 'Z[irr G] by rewrite cycTIiso_vchar. +pose j1 := signW2 b #1; pose d : algC := (-1) ^+ b; pose mu1 := mu_ j1. +have nzj1: j1 != 0 by [rewrite signW2_eq0 ?Iirr1_neq0]; have S1mu1 := S1mu nzj1. +have o_mu_eta01 j: j != 0 -> '[tau1 (mu_ j), eta01] = d *+ (j == j1). + move/Dtau1->; rewrite -/d cfdotZl cfdot_suml big_ord_recl /=. + rewrite cfdot_cycTIiso andTb (inv_eq (signW2K b)). + by rewrite big1 ?addr0 ?mulr_natr // => i _; rewrite cfdot_cycTIiso. +have [zeta | alpha [Zalpha kerPalpha [_]]] := calS1_split1 cohS S1mu1 Zeta01. + case/S1cases=> [[j nz_j ->] | /o_tau1_eta-> //]. + by rewrite o_mu_eta01 // (inj_eq (prTIred_inj _)) => /negPf->. +rewrite o_mu_eta01 // eqxx mulrb => -> lb_alpha. +rewrite -ler_subl_addl cfnorm_prTIred -/q mulrAC sqrr_sign mul1r. +rewrite mu1uq // natrM exprMn (mulrAC _ q%:R) (mulrA _ q%:R) !mulfK ?neq0CG //. +rewrite natrX -(sdprod_card defS) natrM -mulrBl mulfK ?neq0CG //. +rewrite addrC opprB subrK mulrACA; apply: ler_trans lb_alpha. +apply: ler_trans (P1_int2_lb _) _; first by rewrite rpredMsign Cint_vchar1. +by rewrite exprMn sqrr_sign mul1r lerr. +Qed. + +(* These are the assumptions for (13.9); K will be set to 'F(T) in the only *) +(* application of this lemma, in the proof of (13.10). *) + +Variable K : {group gT}. +Let G0 := ~: (class_support H G :|: class_support K G). + +Variables (tau1 : {additive 'CF(S) -> 'CF(G)}) (lambda : 'CF(S)). + +Hypothesis cohS : coherent_with calS S^# tau tau1. +Hypothesis cohSmu : typeP_TIred_coherent tau1. + +Hypotheses (Slam : lambda \in calS) (irrHlam : irrIndH lambda). + +(* This is Peterfalvi (13.9)(a). *) +(* As this part is only used to establish (13.9.b) it can be Section-local. *) +Let cover_G0 : {in G0, forall x, tau1 lambda x != 0 \/ eta_ #1 0 x != 0}. +Proof. +have [[b _ Dtau1_mu] [/= Ilam Hlam]] := (cohSmu, andP irrHlam). +pose sum_eta1 := (-1) ^+ b *: \sum_i eta_ i #1. +have{Dtau1_mu} [j nz_j tau1muj]: exists2 j, j != 0 & tau1 (mu_ j) = sum_eta1. + pose j := signW2 b #1; have nz: j != 0 by rewrite signW2_eq0 Iirr1_neq0. + by exists j; rewrite // Dtau1_mu // signW2K. +move=> x; rewrite !inE => /norP[H'x _]. +have{tau1muj} ->: tau1 lambda x = sum_eta1 x. + rewrite -[lambda](subrK (mu_ j)) raddfD cfunE tau1muj. + rewrite [tau1 _ x](cfun_on0 _ H'x) ?add0r {x H'x}//=. + have Hmuj: mu_ j \in calH := Hmu nz_j. + have dmu1: (lambda - mu_ j) 1%g == 0 by rewrite !cfunE !calHuq ?subrr. + have H1dmu: lambda - mu_ j \in 'CF(S, H^#). + by rewrite cfun_onD1 rpredB ?((seqInd_on (gFnormal _ _)) setT). + have [_ ->] := cohS; last first. + by rewrite zcharD1E ?rpredB ?mem_zchar ?FTseqInd_TIred /=. + have A0dmu := cfun_onS (Fitting_sub_FTsupp0 maxS) H1dmu. + have [_ _ _ _ [_ -> //]] := FTtypeP_facts. + by rewrite cfInd_on ?subsetT // (cfun_onS _ H1dmu) ?imset2Sl ?subsetDl. +apply/nandP/andP=> [[/eqP sum_eta1x_0 /eqP eta1x_0]]. +have cycW: cyclic W by have [] := ctiWG. +have W'x: x \notin class_support (cyclicTIset defW) G. + apply: contra_eqN eta1x_0 => /imset2P[{x H'x sum_eta1x_0}x g Wx Gg ->]. + rewrite cfunJ {g Gg}// cycTIiso_restrict //. + by rewrite lin_char_neq0 ?irr_cyclic_lin //; case/setDP: Wx. +have nz_i1 : #1 != 0 :> Iirr W1 by rewrite Iirr1_neq0. +have eta_x_0 i: i != 0 -> eta_ i 0 x = 0. + rewrite /w_ dprod_IirrEl => /(cfExp_prime_transitive pr_q nz_i1)[k co_k_p ->]. + have: coprime k #[w_ #1 0]%CF by rewrite /w_ dprod_IirrEl cforder_sdprod. + rewrite rmorphX /= -dprod_IirrEl => /(cycTIiso_aut_exists ctiWG)[[uu ->] _]. + by rewrite cfunE /= -/sigma eta1x_0 rmorph0. +have eta_i1 i: i != 0 -> eta_ i #1 x = eta_ 0 #1 x - 1. + move=> nz_i; apply/eqP; pose alpha := cfCyclicTIset defW i #1. + have Walpha: alpha \in 'CF(W, cyclicTIset defW). + by rewrite (cfCycTI_on ctiWG) ?Iirr1_neq0. + have: sigma alpha x == 0. + by rewrite cycTIiso_Ind // (cfun_on0 _ W'x) ?cfInd_on ?subsetT. + rewrite [alpha]cfCycTI_E linearD !linearB /= !cfunE cycTIiso1 cfun1E inE. + by rewrite {1}eta_x_0 //= subr0 addrC addr_eq0 opprB. +have eta11x: eta_ #1 #1 x = - (q%:R)^-1. + rewrite -mulN1r; apply: canRL (mulfK (neq0CG W1)) _. + transitivity ((-1) ^+ b * sum_eta1 x - 1); last first. + by rewrite sum_eta1x_0 mulr0 add0r. + rewrite cfunE signrMK mulr_natr -/q -nirrW1 -sumr_const sum_cfunE. + by rewrite !(bigD1 0 isT) /= addrAC eta_i1 // (eq_bigr _ eta_i1). +have: - eta_ #1 #1 x \in Cint. + rewrite rpredN Cint_rat_Aint ?Aint_vchar ?cycTIiso_vchar //. + by rewrite eta11x rpredN rpredV rpred_nat. +case/norm_Cint_ge1/implyP/idPn; rewrite eta11x opprK invr_eq0 neq0CG /=. +by rewrite normfV normr_nat invf_ge1 ?gt0CG // lern1 -ltnNge ltnW. +Qed. + +(* This is Peterfalvi (13.9)(b). *) +Lemma FTtypeP_sum_nonFitting_lb : + \sum_(x in G0) (`|tau1 lambda x| ^+ 2 + `|eta_ #1 0 x| ^+ 2) >= #|G0|%:R. +Proof. +pose A (xi : 'CF(G)) := [set x in G0 | xi x != 0]. +suffices A_ub xi: xi \in dirr G -> #|A xi|%:R <= \sum_(x in G0) `|xi x| ^+ 2. + apply: ler_trans (_ : (#|A (tau1 lambda)| + #|A (eta_ #1 0)|)%:R <= _). + rewrite leC_nat -cardsUI /A !setIdE -setIUr (leq_trans _ (leq_addr _ _)) //. + rewrite subset_leq_card // subsetIidl. + by apply/subsetP=> x /cover_G0/orP; rewrite !inE. + rewrite natrD big_split ler_add ?A_ub ?cycTIiso_dirr //. + have [[[Itau1 Ztau1] _] [Ilam _]] := (cohS, andP irrHlam). + by rewrite dirrE Ztau1 ?Itau1 ?mem_zchar //= irrWnorm. +case/dirrP=> d [t Dxi]; rewrite (big_setID [set x | xi x != 0]) /= addrC. +rewrite -setIdE -/(A _) big1 ?add0r => [|x]; last first. + by rewrite !inE negbK => /andP[/eqP-> _]; rewrite normr0 expr0n. +rewrite -sum1_card !(partition_big_imset (@cycle _)) /= natr_sum. +apply: ler_sum => _ /imsetP[x Ax ->]. +pose B := [pred y | generator <[x]> y]; pose phi := 'Res[<[x]>] 'chi_t. +have defA: [pred y in A xi | <[y]> == <[x]>] =i B. + move=> y; rewrite inE /= eq_sym andb_idl // !inE => eq_xy. + have LGxy L (LG := class_support L G): x \notin LG -> y \notin LG. + rewrite /LG class_supportEr; apply: contra => /bigcupP[g Gg Lg_y]. + apply/bigcupP; exists g => //; move: Lg_y. + by rewrite -!cycle_subG (eqP eq_xy). + move: Ax; rewrite !inE !negb_or -andbA => /and3P[/LGxy-> /LGxy->]. + apply: contraNneq => chi_y_0. + have [k co_k_y ->]: exists2 k, coprime k #[y] & x = (y ^+ k)%g. + have Yx: generator <[y]> x by rewrite [generator _ _]eq_sym. + have /cycleP[k Dx] := cycle_generator Yx; exists k => //. + by rewrite coprime_sym -generator_coprime -Dx. + have Zxi: xi \in 'Z[irr G] by rewrite Dxi rpredZsign irr_vchar. + have [uu <- // _] := make_pi_cfAut [group of G] co_k_y. + by rewrite cfunE chi_y_0 rmorph0. +have resB: {in B, forall y, `|xi y| ^+ 2 = `|phi y| ^+ 2}. + move=> y /cycle_generator Xy. + by rewrite Dxi cfunE normrMsign cfResE ?subsetT. +rewrite !(eq_bigl _ _ defA) sum1_card (eq_bigr _ resB). +apply: sum_norm2_char_generators => [|y By]. + by rewrite cfRes_char ?irr_char. +rewrite -normr_eq0 -sqrf_eq0 -resB // sqrf_eq0 normr_eq0. +by move: By; rewrite -defA !inE -andbA => /and3P[]. +Qed. + +End Thirteen_2_3_5_to_9. + +Section Thirteen_4_10_to_16. + +(* These assumptions correspond to Peterfalvi, Hypothesis (13.1), most of *) +(* which gets used to prove (13.4) and (13.9-13). *) + +Variables S U W W1 W2 : {group gT}. +Hypotheses (maxS : S \in 'M) (defW : W1 \x W2 = W). +Hypotheses (StypeP : of_typeP S U defW). + +Local Notation "` 'W1'" := (gval W1) (at level 0, only parsing) : group_scope. +Local Notation "` 'W2'" := (gval W2) (at level 0, only parsing) : group_scope. +Local Notation "` 'W'" := (gval W) (at level 0, only parsing) : group_scope. +Local Notation V := (cyclicTIset defW). + +Local Notation "` 'S'" := (gval S) (at level 0, only parsing) : group_scope. +Local Notation P := `S`_\F%G. +Local Notation "` 'P'" := `S`_\F (at level 0) : group_scope. +Local Notation PU := S^`(1)%G. +Local Notation "` 'PU'" := `S^`(1) (at level 0) : group_scope. +Local Notation "` 'U'" := (gval U) (at level 0, only parsing) : group_scope. +Local Notation C := 'C_U(`P)%G. +Local Notation "` 'C'" := 'C_`U(`P) (at level 0) : group_scope. +Local Notation H := 'F(S)%G. +Local Notation "` 'H'" := 'F(`S) (at level 0) : group_scope. + +Let defS : PU ><| W1 = S. Proof. by have [[]] := StypeP. Qed. +Let defPU : P ><| U = PU. Proof. by have [_ []] := StypeP. Qed. +Let defH : P \x C = H. Proof. by have [] := typeP_context StypeP. Qed. + +Let notStype1 : FTtype S != 1%N. Proof. exact: FTtypeP_neq1 StypeP. Qed. +Let notStype5 : FTtype S != 5%N. Proof. exact: FTtype5_exclusion maxS. Qed. + +Let pddS := FT_prDade_hypF maxS StypeP. +Let ptiWS : primeTI_hypothesis S PU defW := FT_primeTI_hyp StypeP. +Let ctiWG : cyclicTI_hypothesis G defW := pddS. +Local Notation Sfacts := (FTtypeP_facts maxS StypeP). + +Let ntW1 : W1 :!=: 1. Proof. by have [[]] := StypeP. Qed. +Let ntW2 : W2 :!=: 1. Proof. by have [_ _ _ []] := StypeP. Qed. +Let cycW1 : cyclic W1. Proof. by have [[]] := StypeP. Qed. +Let cycW2 : cyclic W2. Proof. by have [_ _ _ []] := StypeP. Qed. + +Let p := #|W2|. +Let q := #|W1|. + +Let pr_p : prime p. Proof. by have [] := FTtypeP_primes maxS StypeP. Qed. +Let pr_q : prime q. Proof. by have [] := FTtypeP_primes maxS StypeP. Qed. + +Let qgt2 : q > 2. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. +Let pgt2 : p > 2. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. + +Let coPUq : coprime #|PU| q. +Proof. by rewrite (coprime_sdprod_Hall_r defS); have [[]] := StypeP. Qed. + +Let sW2P: W2 \subset P. Proof. by have [_ _ _ []] := StypeP. Qed. + +Let p'q : q != p. +Proof. +by rewrite -dvdn_prime2 -?prime_coprime -?(cyclic_dprod defW) //; case: ctiWG. +Qed. + +Let nirrW1 : #|Iirr W1| = q. Proof. by rewrite card_Iirr_cyclic. Qed. +Let nirrW2 : #|Iirr W2| = p. Proof. by rewrite card_Iirr_cyclic. Qed. +Let NirrW1 : Nirr W1 = q. Proof. by rewrite -nirrW1 card_ord. Qed. +Let NirrW2 : Nirr W2 = p. Proof. by rewrite -nirrW2 card_ord. Qed. + +Local Open Scope ring_scope. + +Let sigma := (cyclicTIiso ctiWG). +Let w_ i j := (cyclicTIirr defW i j). +Local Notation eta_ i j := (sigma (w_ i j)). + +Let mu_ := primeTIred ptiWS. +Local Notation tau := (FT_Dade0 maxS). + +Let calS0 := seqIndD PU S S`_\s 1. +Let rmR := FTtypeP_coh_base maxS StypeP. +Let scohS0 : subcoherent calS0 tau rmR. +Proof. exact: FTtypeP_subcoherent StypeP. Qed. + +Let calS := seqIndD PU S P 1. +Let sSS0 : cfConjC_subset calS calS0. +Proof. exact/seqInd_conjC_subset1/Fcore_sub_FTcore. Qed. + +Local Notation calH := (seqIndT H S). +Local Notation calHuq := (FTtypeP_Ind_Fitting_1 maxS StypeP). + +Section Thirteen_10_to_13_15. + +(* This section factors the assumption that S contains an irreducible induced *) +(* from a linear character of H. It does not actually export (13.4) and *) +(* and (4.11) but instead uses them to carry out the bulk of the proofs of *) +(* (4.12), (4.13) and (4.15). The combinatorial bound m is also local to this *) +(* Section, but (4.10) has to be exported from an inner Section that factors *) +(* facts about T, the typeP pair associate of S. *) +(* Note that u and c are bound locally to this section; we will set u = #|U| *) +(* after this section. *) + +Variable lambda : 'CF(S). +Hypotheses (Slam : lambda \in calS) (irrHlam : irrIndH lambda). +Let Hlam : lambda \in calH. Proof. by have [] := andP irrHlam. Qed. +Let Ilam : lambda \in irr S. Proof. by have [] := andP irrHlam. Qed. + +Let c := #|C|. +Let u := #|U : C|. +Let oU : #|U| = (u * c)%N. Proof. by rewrite mulnC Lagrange ?subsetIl. Qed. + +Let m : algC := 1 - q.-1%:R^-1 - q.-1%:R / (q ^ p)%:R + (q.-1 * q ^ p)%:R^-1. + +Section Thirteen_4_10. + +(* This Section factors assumptions and facts about T, including Lemma (13.4) *) +(* is local to this Section. *) + +Variables T V : {group gT}. +Hypotheses (maxT : T \in 'M) (xdefW : W2 \x W1 = W). +Hypothesis TtypeP : of_typeP T V xdefW. + +Local Notation Q := (gval T)`_\F. +Local Notation D := 'C_(gval V)(Q). +Local Notation K := 'F(gval T). +Let v := #|V : D|. + +Local Notation calT := (seqIndD T^`(1) T (gval T)`_\F 1). + +(* This part of the proof of (13.4) is reused in (13.10). *) +Let tiHK: class_support H^# G :&: class_support K^# G = set0. +Proof. +apply/eqP/set0Pn => [[_ /setIP[/imset2P[x g1 H1x _ ->] /imset2P[xg g2]]]]. +pose g := (g2 * g1^-1)%g => /setD1P[_ Kxg] _ Dxg. +have{Kxg Dxg} Kgx: x \in K :^ g by rewrite conjsgM mem_conjgV Dxg memJ_conjg. +have{Kgx} cxQg: Q :^ g \subset 'C[x]. + rewrite sub_cent1 (subsetP _ _ Kgx) // centJ conjSg centsC. + have [/dprodW/mulG_sub[/subset_trans-> //=]] := typeP_context TtypeP. + exact: FTtypeP_Fitting_abelian TtypeP. +have{cxQg} sQgS: Q :^ g \subset S. + have sH1A0 := subset_trans (Fitting_sub_FTsupp maxS) (FTsupp_sub0 S). + have{sH1A0} A0x: x \in 'A0(S) := subsetP sH1A0 x H1x. + have [_ _ _ _ [tiA0 _]] := Sfacts. + by have:= cent1_normedTI tiA0 A0x; rewrite setTI; apply: subset_trans. +have /pHallP[_ eq_Sq_q]: q.-Hall(S) W1. + have qW1: q.-group W1 by rewrite /pgroup pnat_id. + have [|//] := coprime_mulGp_Hall (sdprodW defS) _ qW1. + by rewrite /pgroup p'natE // -prime_coprime // coprime_sym. +have:= partn_dvd q (cardG_gt0 _) (cardSg sQgS). +rewrite cardJg /= -eq_Sq_q => /(dvdn_leq_log q (cardG_gt0 _))/idPn[]. +have [_ [_ ->] _ _ _] := FTtypeP_facts maxT TtypeP. +by rewrite -ltnNge p_part !pfactorK // logn_prime // eqxx ltnW. +Qed. + +(* This is Peterfalvi (13.4). *) +Let T_Galois : [/\ typeP_Galois TtypeP, D = 1%g & v = (q ^ p).-1 %/ q.-1]. +Proof. +apply: FTtypeP_no_Ind_Fitting_facts => //; apply/hasPn=> theta Ttheta. +apply/andP=> [[/= irr_theta Ktheta]]; set calK := seqIndT _ T in Ktheta. +have [tau1S cohS [bS _ Dtau1Smu]] := FTtypeP_coherence maxS StypeP. +have [tau1T cohT [bT _ Dtau1Tnu]] := FTtypeP_coherence maxT TtypeP. +have [[[Itau1S Ztau1S] Dtau1S] [[Itau1T Ztau1T] Dtau1T]] := (cohS, cohT). +have onF0 := cfun_onS (Fitting_sub_FTsupp0 _). +pose HG := class_support H^# G; pose KG := class_support K^# G. +have Hdlambda xi: + xi \in calH -> xi \in calS -> tau1S (lambda - xi) \in 'CF(G, HG). +- move=> Hxi Sxi; have H1dxi: lambda - xi \in 'CF(S, H^#). + rewrite cfun_onD1 rpredB ?((seqInd_on (gFnormal _ _)) setT) //=. + by rewrite !cfunE !calHuq ?subrr. + rewrite Dtau1S ?zcharD1E ?rpredB ?mem_zchar ?(cfun_on0 H1dxi) ?inE ?eqxx //=. + by have [_ _ _ _ [_ ->]] := Sfacts; rewrite ?onF0 // cfInd_on ?subsetT. +have Kdtheta xi: + xi \in calK -> xi \in calT -> tau1T (theta - xi) \in 'CF(G, KG). +- move=> Kxi Txi; have K1dxi: theta - xi \in 'CF(T, K^#). + rewrite cfun_onD1 rpredB ?((seqInd_on (gFnormal _ _)) setT) //=. + by rewrite !cfunE !(FTtypeP_Ind_Fitting_1 _ TtypeP) ?subrr. + rewrite Dtau1T ?zcharD1E ?rpredB ?mem_zchar ?(cfun_on0 K1dxi) ?inE ?eqxx //=. + have [_ _ _ _ [_ ->]] := FTtypeP_facts maxT TtypeP; last exact: onF0. + by rewrite cfInd_on ?subsetT. +have oHK alpha beta: + alpha \in 'CF(G, HG) -> beta \in 'CF(G, KG) -> '[alpha, beta] = 0. +- by move=> Halpha Kbeta; rewrite (cfdotElr Halpha Kbeta) tiHK big_set0 mulr0. +have o_lambda_theta: '[tau1S lambda, tau1T theta] = 0. + pose S1 := lambda :: lambda^*%CF; pose T1 := theta :: theta^*%CF. + have sS1S: {subset S1 <= calS} by apply/allP; rewrite /= Slam cfAut_seqInd. + have sT1T: {subset T1 <= calT} by apply/allP; rewrite /= Ttheta cfAut_seqInd. + have ooS1: orthonormal (map tau1S S1). + rewrite map_orthonormal //; first exact: (sub_in2 (zchar_subset sS1S)). + apply: seqInd_conjC_ortho2 Slam; rewrite ?gFnormal ?mFT_odd //. + by have /mulG_sub[] := sdprodW defPU. + have ooT1: orthonormal (map tau1T T1). + rewrite map_orthonormal //; first exact: (sub_in2 (zchar_subset sT1T)). + apply: seqInd_conjC_ortho2 Ttheta; rewrite ?gFnormal ?mFT_odd //. + by have [_ [_ _ _ /sdprodW/mulG_sub[]]] := TtypeP. + have /andP/orthonormal_vchar_diff_ortho := conj ooS1 ooT1; apply. + by split; apply/allP; rewrite /= ?Ztau1S ?Ztau1T ?mem_zchar ?cfAut_seqInd. + have on1'G M beta: beta \in 'CF(G, class_support M^# G) -> beta 1%g = 0. + move/cfun_on0->; rewrite // class_supportEr -cover_imset -class_supportD1. + by rewrite !inE eqxx. + rewrite -!raddfB; set alpha := tau1S _; set beta := tau1T _. + have [Halpha Kbeta]: alpha \in 'CF(G, HG) /\ beta \in 'CF(G, KG). + by rewrite Hdlambda ?Kdtheta ?cfAut_seqInd ?cfAut_seqIndT. + by rewrite oHK // {1}(on1'G _ _ Halpha) (on1'G _ _ Kbeta) !eqxx. +pose ptiWT := FT_primeTI_hyp TtypeP; pose nu_ := primeTIred ptiWT. +have etaC := cycTIisoC (FT_cyclicTI_hyp StypeP) (FT_cyclicTI_hyp TtypeP). +have /idPn[]: '[tau1S (lambda - mu_ #1), tau1T (theta - nu_ #1)] == 0. + rewrite oHK //. + by rewrite Hdlambda ?FTseqInd_TIred ?FTprTIred_Ind_Fitting ?Iirr1_neq0. + by rewrite Kdtheta ?FTseqInd_TIred ?FTprTIred_Ind_Fitting ?Iirr1_neq0. +rewrite !raddfB /= !cfdotBl o_lambda_theta Dtau1Smu ?Dtau1Tnu ?Iirr1_neq0 //. +rewrite !cfdotZl !cfdotZr rmorph_sign !cfdot_suml big1 => [|i _]; last first. + rewrite cfdotC etaC (coherent_ortho_cycTIiso TtypeP _ cohT) ?conjC0 //. + by apply: seqInd_conjC_subset1; apply: Fcore_sub_FTcore. +rewrite cfdot_sumr big1 ?mulr0 ?subr0 ?add0r ?opprK => [|j _]; last first. + by rewrite -etaC (coherent_ortho_cycTIiso StypeP _ cohS). +set i1 := iter bT _ #1; set j1 := iter bS _ #1. +rewrite !mulf_eq0 !signr_eq0 (bigD1 i1) //= addrC big1 => [|i i1'i]; last first. + rewrite etaC cfdot_sumr big1 // => j _; rewrite cfdot_cycTIiso. + by rewrite (negPf i1'i) andbF. +rewrite etaC cfdot_sumr (bigD1 j1) //= cfdot_cycTIiso !eqxx addrCA. +rewrite big1 ?addr0 ?oner_eq0 // => j j1'j; rewrite cfdot_cycTIiso. +by rewrite eq_sym (negPf j1'j). +Qed. + +(* This is Peterfalvi (13.10). *) +Lemma FTtypeP_compl_ker_ratio_lb : m * (p ^ q.-1)%:R / q%:R < u%:R / c%:R. +Proof. +have [tau1 cohS cohSmu] := FTtypeP_coherence maxS StypeP. +pose lam1 := tau1 lambda; pose eta10 := eta_ #1 0. +pose H1G := class_support H^# G; pose K1G := class_support K^# G. +pose G0 := ~: (class_support H G :|: class_support K G). +pose invJ (f : gT -> algC) := forall y x, f (x ^ y) = f x. +pose nm2 (chi : 'CF(G)) x := `|chi x| ^+ 2; pose g : algC := #|G|%:R. +have injJnm2 chi: invJ (nm2 chi) by move=> y x; rewrite /nm2 cfunJ ?inE. +have nm2_dirr chi: chi \in dirr G -> g^-1 <= nm2 chi 1%g / g. + case/dIrrP=> d ->; rewrite -{1}[g^-1]mul1r ler_pmul2r ?invr_gt0 ?gt0CG //. + rewrite expr_ge1 ?normr_ge0 // cfunE normrMsign. + by rewrite irr1_degree normr_nat ler1n irr_degree_gt0. +pose mean (F M : {set gT}) (f : gT -> algC) := (\sum_(x in F) f x) / #|M|%:R. +have meanTI M (F := 'F(gval M)^#) (FG := class_support F G) f: + M \in 'M -> normedTI F G M -> invJ f -> mean FG G f = mean F M f. +- move=> maxM /and3P[ntF tiF /eqP defN] fJ; apply: canLR (mulfK (neq0CG _)) _. + rewrite (set_partition_big _ (partition_class_support ntF tiF)) /=. + rewrite mulrAC -mulrA -natf_indexg ?subsetT //=. + have ->: #|G : M| = #|F :^: G| by rewrite card_conjugates defN. + rewrite mulr_natr -sumr_const; apply: eq_bigr => _ /imsetP[y _ ->]. + by rewrite (big_imset _ (in2W (conjg_inj _))) (eq_bigr _ (in1W (fJ y))). +have{meanTI} meanG f : + invJ f -> mean G G f = f 1%g / g + mean H^# S f + mean K^# T f + mean G0 G f. +- have type24 maxM := compl_of_typeII_IV maxM _ (FTtype5_exclusion maxM). + have tiH: normedTI H^# G S by have/type24[] := StypeP. + have{type24} tiK: normedTI K^# G T by have/type24[] := TtypeP. + move=> fJ; rewrite -!meanTI // {1}/mean (big_setD1 1%g) // (big_setID H1G) /=. + rewrite [in rhs in _ + (_ + rhs)](big_setID K1G) /= -/g -!mulrDl !addrA. + congr ((_ + _ + _ + _) / g); rewrite ?(setIidPr _) // /H1G /K1G. + + by rewrite class_supportEr -cover_imset -class_supportD1 setSD ?subsetT. + + rewrite subsetD -setI_eq0 setIC tiHK eqxx andbT. + by rewrite class_supportEr -cover_imset -class_supportD1 setSD ?subsetT. + rewrite !class_supportEr -!cover_imset -!class_supportD1. + apply: eq_bigl => x; rewrite !inE andbT -!negb_or orbCA orbA orbC. + by case: (x =P 1%g) => //= ->; rewrite mem_class_support ?group1. +have lam1_ub: mean G0 G (nm2 lam1) <= lambda 1%g ^+ 2 / #|S|%:R - g^-1. + have [[Itau1 Ztau1] _] := cohS. + have{Itau1} n1lam1: '[lam1] = 1 by rewrite Itau1 ?mem_zchar ?irrWnorm. + have{Ztau1} Zlam1: lam1 \in 'Z[irr G] by rewrite Ztau1 ?mem_zchar. + rewrite -ler_opp2 opprB -(ler_add2l '[lam1]) {1}n1lam1 addrCA. + rewrite (cfnormE (cfun_onG _)) (mulrC g^-1) [_ / g](meanG (nm2 _)) // addrK. + rewrite -addrA ler_add ?nm2_dirr //; first by rewrite dirrE Zlam1 n1lam1 /=. + rewrite ler_paddr ?divr_ge0 ?ler0n //. + by apply: sumr_ge0 => x _; rewrite exprn_ge0 ?normr_ge0. + rewrite ler_pdivl_mulr ?gt0CG // mulrBl mul1r divfK ?neq0CG //. + by rewrite (FTtypeP_sum_Ind_Fitting_lb StypeP). +pose ub_lam1 : algC := (#|T^`(1)%g|%:R - (v ^ 2)%:R - #|Q|.-1%:R) / #|T|%:R. +have [_ D_1 Dv] := T_Galois. +have defK : K = Q by have [<-] := typeP_context TtypeP; rewrite D_1 dprodg1. +have eta10_ub: mean G0 G (nm2 (eta_ #1 0)) <= #|G0|%:R / g - ub_lam1. + rewrite -ler_opp2 opprB -(ler_add2l '[eta_ #1 0]) {2}(cfnormE (cfun_onG _)). + rewrite (mulrC g^-1) [_ / g in rhs in _ <= rhs](meanG (nm2 _)) // addrK. + have ->: '[eta_ #1 0] = mean G G (fun _ => 1). + by rewrite /mean sumr_const cfdot_cycTIiso eqxx divff ?neq0CG. + rewrite meanG // [in lhs in lhs <= _]/mean !sumr_const addrACA subrr addr0. + rewrite [lhs in lhs <= _]addrAC -addrA -mulrDl (cardsD1 1%g Q) group1 -defK. + rewrite mul1r subrK ?ler_add ?ler_pmul2r ?invr_gt0 ?gt0CG //. + - by rewrite nm2_dirr ?cycTIiso_dirr. + - exact: (FTtypeP_sum_cycTIiso10_lb _ StypeP). + congr (_ <= _): (FTtypeP_sum_cycTIiso01_lb maxT TtypeP). + by apply: eq_bigr => x _; congr (nm2 _ x); apply: cycTIisoC. +have: ub_lam1 < lambda 1%g ^+ 2 / #|S|%:R. + rewrite -[_ / _](subrK g^-1) ltr_spaddr ?invr_gt0 ?gt0CG //. + rewrite -(ler_add2r (#|G0|%:R / g)) -ler_subr_addl -addrA. + apply: ler_trans (ler_add lam1_ub eta10_ub). + rewrite -mulrDl -big_split /= ler_pmul2r ?invr_gt0 ?gt0CG //. + exact: FTtypeP_sum_nonFitting_lb. +rewrite calHuq // -/u -(sdprod_card defS) -/q -(sdprod_card defPU) oU mulnC. +rewrite mulnCA mulnAC !natrM !invfM expr2 !mulrA !mulfK ?neq0CG ?neq0CiG //. +rewrite mulrAC ltr_pdivl_mulr ?ltr_pdivr_mulr ?gt0CG //. +congr (_ < _); last by rewrite -mulrA mulrC. +have [_ [_ ->] _ _ _] := Sfacts; rewrite -/p -/q. +rewrite -{1}(ltn_predK qgt2) expnS natrM mulrA; congr (_ * _). +have /sdprod_card oT: T^`(1) ><| W2 = T by have [[]] := TtypeP. +rewrite /ub_lam1 -{}oT natrM invfM mulrA divfK ?mulrBl ?divff ?neq0CG //. +have /sdprod_card <-: Q ><| V = T^`(1)%g by have [_ []] := TtypeP. +have ->: #|V| = v by rewrite /v D_1 indexg1. +rewrite mulnC !natrM invfM mulrA mulfK ?neq0CiG //. +have [_ [_ oQ] _ _ _] := FTtypeP_facts maxT TtypeP; rewrite -/p -/q /= in oQ. +rewrite Dv natf_div ?dvdn_pred_predX // oQ. +rewrite invfM invrK -mulrA -subn1 mulVKf ?gtr_eqF ?ltr0n //; last first. + by rewrite subn_gt0 -(exp1n p) ltn_exp2r ltnW // ltnW. +rewrite -oQ natrB ?cardG_gt0 // !mulrBl mul1r mulrC mulKf ?neq0CG // -invfM. +by rewrite -natrM oQ opprD opprK addrA addrAC. +Qed. + +End Thirteen_4_10. + +(* This is (13.10) without the dependency on T. *) +Let gen_lb_uc : m * (p ^ q.-1)%:R / q%:R < u%:R / c%:R. +Proof. +have [T pairST [xdefW [V TtypeP]]] := FTtypeP_pair_witness maxS StypeP. +by apply: FTtypeP_compl_ker_ratio_lb TtypeP; have [[]] := pairST. +Qed. + +Import ssrint. +(* This is Peterfalvi (13.11). *) +Let lb_m_cases : + [/\ (*a*) (q >= 7)%N -> m > 8%:R / 10%:R, + (*b*) (q >= 5)%N -> m > 7%:R / 10%:R + & (*c*) q = 3 -> + m > 49%:R / 100 %:R /\ u%:R / c%:R > (p ^ 2).-1%:R / 6%:R :> algC]. +Proof. +pose mkrat b d := fracq (b, d%:Z). +pose test r b d := 1 - mkrat 1 r.-1 - mkrat 1 (r ^ 2)%N > mkrat b%:Z d. +have lb_m r b d: test r.+2 b d -> (q >= r.+2)%N -> m > b%:R / d%:R. + rewrite /test /mkrat !fracqE !CratrE /= => ub_bd le_r_q. + apply: ltr_le_trans ub_bd _; rewrite ler_paddr ?invr_ge0 ?ler0n //. + rewrite -!addrA ler_add2l -!opprD ler_opp2 ler_add //. + rewrite mul1r lef_pinv ?qualifE ?ltr0n //; last by rewrite -(subnKC qgt2). + by rewrite leC_nat -ltnS (ltn_predK qgt2). + rewrite -(ltn_predK pgt2) expnSr natrM invfM mulrA. + rewrite ler_pdivr_mulr ?gt0CG // mulrAC mul1r -subn1. + rewrite ler_pmul ?invr_ge0 ?ler0n ?leC_nat ?leq_subr //. + rewrite lef_pinv ?qualifE ?ltr0n ?leC_nat ?expn_gt0 ?(prime_gt0 pr_q) //. + apply: leq_trans (_ : q ^ 2 <= _)%N; first by rewrite leq_exp2r. + by rewrite -(subnKC qgt2) leq_pexp2l // -subn1 ltn_subRL. +split=> [||q3]; try by apply: lb_m; compute. +pose d r : algC := (3 ^ r.-1)%:R^-1; pose f r := (r ^ 2)%:R * d r. +have Dm: m = (1 - d p) / 2%:R. + rewrite mulrBl mul1r -mulrN mulrC /m q3 /= addrAC -addrA natrM invfM -mulrBl. + rewrite -{1}(ltn_predK pgt2) expnS natrM invfM mulrA. + by congr (_ + _ / _); apply/eqP; rewrite -!CratrE; compute. +split; last apply: ler_lt_trans gen_lb_uc. + apply: ltr_le_trans (_ : (1 - d 5) / 2%:R <= _). + by rewrite /d -!CratrE; compute. + rewrite Dm ler_pmul2r ?invr_gt0 ?ltr0n // ler_add2l ler_opp2. + rewrite lef_pinv ?qualifE ?ltr0n ?expn_gt0 // leC_nat leq_pexp2l //=. + by rewrite -subn1 ltn_subRL odd_geq ?mFT_odd //= ltn_neqAle pgt2 andbT -q3. +rewrite -mulrA mulrCA Dm -mulrA -invfM -natrM mulrA q3 mulrBr mulr1. +rewrite ler_pmul2r ?invr_gt0 ?ltr0n //= -subn1 natrB ?expn_gt0 ?prime_gt0 //. +rewrite ler_add2l ler_opp2 -/(f p) -(subnKC pgt2). +elim: (p - 3)%N => [|r]; first by rewrite /f /d -!CratrE; compute. +apply: ler_trans; rewrite addnS /f /d; set x := (3 + r)%N. +rewrite ler_pdivr_mulr ?ltr0n ?expn_gt0 // mulrAC (expnS 3) (natrM _ 3). +rewrite mulrA mulfK ?gtr_eqF ?ltr0n ?expn_gt0 //. +rewrite -ler_pdivr_mull ?ltr0n // !natrX -exprVn -exprMn. +rewrite mulrS mulrDr mulr1 mulVf ?pnatr_eq0 //. +apply: ler_trans (_ : (3%:R^-1 + 1) ^+ 2 <= _); last by rewrite -!CratrE. +rewrite ler_sqr ?rpredD ?rpred1 ?rpredV ?rpred_nat // ler_add2r. +by rewrite lef_pinv ?qualifE ?ltr0n ?leC_nat. +Qed. + +(* This corollary of (13.11) is used in both (13.12) and (13.15). *) +Let small_m_q3 : m < (q * p)%:R / (q.*2.+1 * p.-1)%:R -> q = 3 /\ (p >= 5)%N. +Proof. +move=> ub_m; have [lb7_m lb5_m _] := lb_m_cases. +have [p3 | p_neq3] := eqVneq p 3. + have ub7_m: ~~ (8%:R / 10%:R < m). + rewrite ltr_gtF // (ltr_le_trans ub_m) // p3 /=. + apply: ler_trans (_ : 3%:R / 4%:R <= _); last first. + by rewrite -!CratrE; compute. + rewrite ler_pdivl_mulr ?ltr0n // mulrAC ler_pdivr_mulr ?ltr0n ?muln_gt0 //. + by rewrite -!natrM leC_nat mulnCA mulSn -muln2 -!mulnA leq_addl. + have{ub7_m} q5: q = 5. + apply: contraNeq ub7_m; rewrite neq_ltn odd_ltn ?mFT_odd //= ltnS leqNgt. + by rewrite ltn_neqAle qgt2 -{1}p3 eq_sym p'q -(odd_geq 7) ?mFT_odd. + have /implyP := ltr_trans (lb5_m _) ub_m. + by rewrite q5 p3 -!CratrE; compute. +have pge5: (5 <= p)%N by rewrite odd_geq ?mFT_odd // ltn_neqAle eq_sym p_neq3. +have ub5_m: ~~ (7%:R / 10%:R < m). + rewrite ltr_gtF // (ltr_le_trans ub_m) //. + apply: ler_trans (_ : 2%:R^-1 * (1 + 4%:R^-1) <= _); last first. + by rewrite -!CratrE; compute. + rewrite !natrM invfM mulrACA ler_pmul ?divr_ge0 ?ler0n //. + rewrite ler_pdivr_mulr ?ler_pdivl_mull ?ltr0n // -natrM mul2n leC_nat. + by rewrite ltnW. + rewrite -(subnKC pge5) [_%:R]mulrSr mulrDl divff ?pnatr_eq0 // ler_add2l. + by rewrite mul1r lef_pinv ?qualifE ?ltr0n // leC_nat. +split=> //; apply: contraNeq ub5_m. +by rewrite neq_ltn ltnNge qgt2 -(odd_geq 5) ?mFT_odd. +Qed. + +(* A more usable form for (13.10). *) +Let gen_ub_m : m < (q * u)%:R / (c * p ^ q.-1)%:R. +Proof. +rewrite !natrM invfM mulrA ltr_pdivl_mulr ?ltr0n ?expn_gt0 ?cardG_gt0 //. +by rewrite -mulrA -ltr_pdivr_mull ?gt0CG // mulrC. +Qed. + +(* This is the bulk of the proof of Peterfalvi (13.12). *) +Lemma FTtypeP_Ind_Fitting_reg_Fcore : c = 1%N. +Proof. +apply/eqP/wlog_neg; rewrite eqn_leq cardG_gt0 andbT -ltnNge => c_gt1. +have ub_m: m < (q * (p ^ q).-1)%:R / (c * p ^ q.-1 * p.-1)%:R. + rewrite 2!natrM invfM mulrACA mulrAC -natf_div ?dvdn_pred_predX // -natrM. + rewrite (ltr_le_trans gen_ub_m) // ler_pmul ?invr_ge0 ?ler0n // leC_nat. + by rewrite leq_mul //; case: Sfacts. +have regCW1: semiregular C W1. + have [[_ _ /Frobenius_reg_ker regUW1 _] _ _ _] := FTtypeP_facts maxS StypeP. + by move=> _ y /regUW1 regUx; rewrite setIAC regUx setI1g. +have{regCW1} dv_2q_c1: q.*2 %| c.-1. + rewrite -(subnKC c_gt1) -mul2n Gauss_dvd ?coprime2n ?dvdn2 ?mFT_odd //=. + rewrite odd_sub ?mFT_odd -?subSn // subn2 regular_norm_dvd_pred //. + have /mulG_sub[_ sW1S] := sdprodW defS. + apply: normsI; first by have [_ []] := StypeP. + by rewrite (subset_trans sW1S) ?norms_cent ?gFnorm. +have [q3 pge5]: q = 3 /\ (p >= 5)%N. + apply: small_m_q3; apply: (ltr_le_trans ub_m). + rewrite !natrM -!mulrA ler_pmul2l ?gt0CG //. + rewrite !invfM !mulrA -(subnKC pgt2) ler_pmul2r ?invr_gt0 ?ltr0n //. + rewrite ler_pdivr_mulr ?ltr0n ?expn_gt0 // mulrAC -natrM -expnS. + rewrite prednK ?cardG_gt0 // ler_pmul ?invr_ge0 ?ler0n ?leC_nat ?leq_pred //. + rewrite lef_pinv ?qualifE ?gt0CG ?ltr0n // leC_nat. + by rewrite -(subnKC c_gt1) ltnS dvdn_leq //= -subSn ?subn2. +have [_ _ [//|lb_m lb_uc]] := lb_m_cases. +pose sum3 r : algC := (r.+1 ^ 2)%:R^-1 + r.+1%:R^-1 + 1. +have [b Dc1] := dvdnP dv_2q_c1; rewrite q3 in Dc1. +have [b0 | b_gt0] := posnP b; first by rewrite b0 -(subnKC c_gt1) in Dc1. +have ub3_m r a: (r < p)%N -> (a <= b)%N -> m < 3%:R / (a * 6).+1%:R * sum3 r. + move=> lb_p lb_b; apply: ltr_le_trans ub_m _. + rewrite !natrM !invfM mulrACA -!mulrA q3 ler_pmul2l ?ltr0n //. + rewrite -(ltn_predK c_gt1) Dc1 ler_pmul ?mulr_ge0 ?invr_ge0 ?ler0n //. + by rewrite lef_pinv ?qualifE ?ltr0n // leC_nat ltnS leq_mul. + rewrite predn_exp mulnC natrM 2!big_ord_recl big_ord1 /= /bump /= expn1. + rewrite -(subnKC (ltnW pgt2)) add2n in lb_p *. + rewrite mulfK ?pnatr_eq0 // addnA 2!natrD 2!mulrDr mulr1 {-1}natrM invfM. + rewrite mulrA divfK ?mulVf ?pnatr_eq0 // ler_add2r. + by rewrite ler_add ?lef_pinv ?qualifE ?ltr0n ?leC_nat ?leq_sqr. +have beq1: b = 1%N. + apply: contraTeq lb_m; rewrite neq_ltn ltnNge b_gt0 => /(ub3_m 4) ub41. + by rewrite ltr_gtF // (ltr_trans (ub41 _)) // /sum3 -!CratrE; compute. +have c7: c = 7 by rewrite -(ltn_predK c_gt1) Dc1 beq1. +have plt11: (p < 11)%N. + rewrite ltnNge; apply: contraL lb_m => /ub3_m/(_ b_gt0) ub100. + by rewrite ltr_gtF // (ltr_trans ub100) // /sum3 -!CratrE; compute. +have{plt11} p5: p = 5. + suffices: p \in [seq r <- iota q.+1 7 | prime r & coprime r c]. + by rewrite c7 q3 inE => /eqP. + rewrite mem_filter mem_iota ltn_neqAle p'q q3 pgt2 pr_p (coprimeSg sW2P) //. + by rewrite (coprimegS _ (Ptype_Fcore_coprime StypeP)) ?subIset ?joing_subl. +have [galS | gal'S] := boolP (typeP_Galois StypeP); last first. + have [H1 [_ _ _ _ []]] := typeP_Galois_Pn maxS notStype5 gal'S. + case/pdivP=> r pr_r r_dv_a /(dvdn_trans r_dv_a)/idPn[]. + rewrite Ptype_factor_prime // -/p p5 (Euclid_dvdM 2 2) // gtnNdvd //. + rewrite odd_prime_gt2 ?(dvdn_odd (dvdn_trans r_dv_a (dvdn_indexg _ _))) //. + by rewrite mFT_odd. +have{galS} u_dv_31: u %| 31. + have [_ _ [_ _]] := typeP_Galois_P maxS notStype5 galS. + rewrite Ptype_factor_prime ?Ptype_Fcompl_kernel_cent // -/p -/q p5 q3. + rewrite card_quotient // normsI ?normG ?norms_cent //. + by have [] := sdprodP defPU. +have hallH: Hall S H. + rewrite /Hall -divgS ?gFsub //= -(sdprod_card defS) -(sdprod_card defPU). + rewrite -(dprod_card defH) -mulnA divnMl ?cardG_gt0 // -/c oU mulnAC c7. + have [_ [_ ->] _ _ _] := FTtypeP_facts maxS StypeP. + by rewrite mulnK // -/q -/p q3 p5 coprime_mulr (coprime_dvdr u_dv_31). +rewrite -(leq_pmul2l (cardG_gt0 P)) muln1 (dprod_card defH) subset_leq_card //. +by rewrite (Fcore_max (Hall_pi hallH)) ?gFnormal ?Fitting_nil. +Qed. +Local Notation c1 := FTtypeP_Ind_Fitting_reg_Fcore. + +(* This is the main part of the proof of Peterfalvi (13.13). *) +Lemma FTtypeP_Ind_Fitting_nonGalois_facts : + ~~ typeP_Galois StypeP -> q = 3 /\ #|U| = (p.-1./2 ^ 2)%N. +Proof. +have even_p1: 2 %| p.-1 by rewrite -subn1 -subSS dvdn_sub ?dvdn2 //= mFT_odd. +move=> gal'S; have{gal'S} u_dv_p2q: u %| p.-1./2 ^ q.-1. + have [H1 [_ _ _ _ []]] := typeP_Galois_Pn maxS notStype5 gal'S. + rewrite Ptype_factor_prime ?Ptype_Fcompl_kernel_cent // -/p -/q. + set a := #|U : _| => a_gt1 a_dv_p1 _ [Uhat isoUhat]. + have a_odd: odd a by rewrite (dvdn_odd (dvdn_indexg _ _)) ?mFT_odd. + have [_ _ nPU _] := sdprodP defPU. + rewrite /u -card_quotient ?normsI ?normG ?norms_cent // (card_isog isoUhat). + apply: dvdn_trans (cardSg (subsetT _)) _; rewrite cardsT card_matrix mul1n. + rewrite card_ord Zp_cast ?dvdn_exp2r // -(@Gauss_dvdl a _ 2) ?coprimen2 //. + by rewrite -divn2 divnK. +have [_ lb5_m lb3_m] := lb_m_cases. +pose f r : algC := r%:R / (2 ^ r.-1)%:R. +have ub_m: m < f q. + apply: ltr_le_trans gen_ub_m _; rewrite c1 mul1n. + rewrite natrM ler_pdivr_mulr ?ltr0n ?expn_gt0 ?cardG_gt0 // -mulrA. + rewrite ler_wpmul2l ?ler0n // mulrC !natrX -expr_div_n. + apply: ler_trans (_ : (p.-1 %/ 2)%:R ^+ q.-1 <= _). + by rewrite -natrX leC_nat divn2 dvdn_leq // expn_gt0 -(subnKC pgt2). + rewrite -(subnKC qgt2) ler_pexpn2r ?rpred_div ?rpred_nat // natf_div //. + by rewrite ler_wpmul2r ?invr_ge0 ?ler0n // leC_nat leq_pred. +have{ub_m} q3: q = 3. + apply: contraTeq ub_m; rewrite neq_ltn ltnNge qgt2 -(odd_geq 5) ?mFT_odd //=. + move=> qge5; rewrite ltr_gtF // -(subnKC qge5). + elim: (q - 5)%N => [|r]; last apply: ler_lt_trans. + by apply: ltr_trans (lb5_m qge5); rewrite /f -!CratrE; compute. + rewrite addnS ler_pdivr_mulr ?ltr0n ?expn_gt0 // natrM mulrACA mulrA. + by rewrite divfK ?pnatr_eq0 ?expn_eq0 // mulr_natr mulrS ler_add2r ler1n. +have [[]] := dvdnP u_dv_p2q; rewrite q3; first by rewrite -(subnKC pgt2). +case=> [|b] Du; first by rewrite oU c1 Du muln1 mul1n. +have [_ /idPn[]] := lb3_m q3; rewrite c1 divr1 ler_gtF //. +apply: ler_trans (_ : (p.-1 ^ 2)%:R / 8%:R <= _). + rewrite (natrX _ 2 3) exprSr invfM mulrA natrX -expr_div_n -natf_div // divn2. + by rewrite -natrX Du ler_pdivl_mulr ?ltr0n // mulrC -natrM leC_nat leq_mul. +rewrite -!subn1 (subn_sqr p 1) !natrM -!mulrA ler_wpmul2l ?ler0n //. +rewrite ler_pdivr_mulr 1?mulrAC ?ler_pdivl_mulr ?ltr0n // -!natrM leC_nat. +rewrite (mulnA _ 3 2) (mulnA _ 4 2) leq_mul // mulnBl mulnDl leq_subLR. +by rewrite addnCA (mulnSr p 3) -addnA leq_addr. +Qed. + +(* This is the bulk of the proof of Peterfalvi (13.15). *) +(* We improve slightly on the end of the argument by maing better use of the *) +(* bound on u to get p = 5 directly. *) +Lemma FTtypeP_Ind_Fitting_Galois_ub b : + (p ^ q).-1 %/ p.-1 = (b * u)%N -> (b <= q.*2)%N. +Proof. +move=> Dbu; have: U :!=: 1%g by have [[_ _ /Frobenius_context[]]] := Sfacts. +rewrite trivg_card1 oU c1 muln1 leqNgt; apply: contra => bgt2q. +have [|q3 pge5] := small_m_q3. + apply: ltr_le_trans gen_ub_m _; rewrite c1 mul1n !natrM -!mulrA. + rewrite ler_wpmul2l ?ler0n // ler_pdivr_mulr ?ltr0n ?expn_gt0 ?cardG_gt0 //. + rewrite mulrAC invfM -natrM -expnS prednK ?cardG_gt0 // mulrCA. + rewrite ler_pdivl_mull ?ltr0n // -natrM. + apply: ler_trans (_ : (b * u)%:R <= _); first by rewrite leC_nat leq_mul. + rewrite -Dbu natf_div ?dvdn_pred_predX // ler_wpmul2r ?invr_ge0 ?ler0n //. + by rewrite leC_nat leq_pred. +have ub_p: ((p - 3) ^ 2 < 4 ^ 2)%N. + have [_ _ [] // _] := lb_m_cases; rewrite c1 divr1 ltr_pdivr_mulr ?ltr0n //. + rewrite -natrM ltC_nat prednK ?expn_gt0 ?cardG_gt0 // => /(leq_mul bgt2q). + rewrite mulnC mulnA -Dbu q3 predn_exp mulKn; last by rewrite -(subnKC pgt2). + rewrite 2!big_ord_recl big_ord1 /= /bump /= !mulnDl expn0 expn1. + rewrite addnA mulnS leq_add2r -(leq_add2r 9) (mulnCA p 2 3) -addnA addnCA. + by rewrite -leq_subLR -(sqrn_sub pgt2). +have{ub_p pge5} p5: p = 5. + apply/eqP; rewrite eqn_leq pge5 andbT. + by rewrite ltn_sqr ltnS leq_subLR -ltnS odd_ltn ?mFT_odd in ub_p. +have bgt1: (1 < b)%N by rewrite -(subnKC bgt2q) q3. +rewrite -(eqn_pmul2l (ltnW bgt1)) muln1 eq_sym. +by apply/eqP/prime_nt_dvdP; rewrite ?dvdn_mulr ?gtn_eqF // -Dbu q3 p5. +Qed. + +End Thirteen_10_to_13_15. + +(* This is Peterfalvi (13.12). *) +Lemma FTtypeP_reg_Fcore : C :=: 1%g. +Proof. +have [] := boolP (has irrIndH calS); last first. + by case/(FTtypeP_no_Ind_Fitting_facts maxS StypeP). +by case/hasP=> lambda Slam /FTtypeP_Ind_Fitting_reg_Fcore/card1_trivg->. +Qed. + +Lemma Ptype_Fcompl_kernel_trivial : Ptype_Fcompl_kernel StypeP :=: 1%g. +Proof. by rewrite Ptype_Fcompl_kernel_cent ?FTtypeP_reg_Fcore. Qed. + +(* Since C is trivial, from here on u will denote #|U|. *) +Let u := #|U|. +Let ustar := (p ^ q).-1 %/ p.-1. + +(* This is Peterfalvi (13.13). *) +Lemma FTtypeP_nonGalois_facts : + ~~ typeP_Galois StypeP -> q = 3 /\ u = (p.-1./2 ^ 2)%N. +Proof. +move=> gal'S; have: has irrIndH calS. + by apply: contraR gal'S => /(FTtypeP_no_Ind_Fitting_facts maxS StypeP)[]. +by case/hasP=> lambda Slam /FTtypeP_Ind_Fitting_nonGalois_facts; apply. +Qed. + +Import FinRing.Theory. + +(* This is Peterfalvi (13.14). *) +Lemma FTtypeP_primes_mod_cases : + [/\ odd ustar, + p == 1 %[mod q] -> q %| ustar + & p != 1 %[mod q] -> + [/\ coprime ustar p.-1, ustar == 1 %[mod q] + & forall b, b %| ustar -> b == 1 %[mod q]]]. +Proof. +have ustar_mod r: p = 1 %[mod r] -> ustar = q %[mod r]. + move=> pr1; rewrite -[q]card_ord -sum1_card /ustar predn_exp //. + rewrite -(subnKC pgt2) mulKn // subnKC //. + elim/big_rec2: _ => // i s1 s2 _ eq_s12. + by rewrite -modnDm -modnXm pr1 eq_s12 modnXm modnDm exp1n. +have ustar_odd: odd ustar. + by apply: (can_inj oddb); rewrite -modn2 ustar_mod ?modn2 ?mFT_odd. +split=> // [p1_q|p'1_q]; first by rewrite /dvdn ustar_mod ?modnn //; apply/eqP. +have ustar_gt0: (ustar > 0)%N by rewrite odd_geq. +have [p1_gt0 p_gt0]: (p.-1 > 0 /\ p > 0)%N by rewrite -(subnKC pgt2). +have co_ustar_p1: coprime ustar p.-1. + rewrite coprime_pi' //; apply/pnatP=> //= r pr_r. + rewrite inE -subn1 -eqn_mod_dvd //= mem_primes pr_r ustar_gt0 => /eqP rp1. + rewrite /dvdn ustar_mod // [_ == _]dvdn_prime2 //. + by apply: contraNneq p'1_q => <-; apply/eqP. +suffices ustar_mod_q b: b %| ustar -> b == 1 %[mod q]. + by split; rewrite // ustar_mod_q. +move=> b_dv_ustar; have b_gt0 := dvdn_gt0 ustar_gt0 b_dv_ustar. +rewrite (prod_prime_decomp b_gt0) prime_decompE big_map /= big_seq. +elim/big_rec: _ => // r s /(pi_of_dvd b_dv_ustar ustar_gt0). +rewrite mem_primes -modnMml -modnXm => /and3P[pr_r _ r_dv_ustar]. +suffices{s} ->: r = 1 %[mod q] by rewrite modnXm modnMml exp1n mul1n. +apply/eqP; rewrite eqn_mod_dvd ?prime_gt0 // subn1. +have ->: r.-1 = #|[set: {unit 'F_r}]|. + rewrite card_units_Zp ?prime_gt0 ?pdiv_id //. + by rewrite -[r]expn1 totient_pfactor ?muln1. +have pq_r: p%:R ^+ q == 1 :> 'F_r. + rewrite -subr_eq0 -natrX -(@natrB _ _ 1) ?expn_gt0 ?cardG_gt0 // subn1. + rewrite -(divnK (dvdn_pred_predX p q)) -Fp_nat_mod //. + by rewrite -modnMml (eqnP r_dv_ustar) mod0n. +have Up_r: (p%:R : 'F_r) \is a GRing.unit. + by rewrite -(unitrX_pos _ (prime_gt0 pr_q)) (eqP pq_r) unitr1. +congr (_ %| _): (order_dvdG (in_setT (FinRing.unit 'F_r Up_r))). +apply/prime_nt_dvdP=> //; last by rewrite order_dvdn -val_eqE val_unitX. +rewrite -dvdn1 order_dvdn -val_eqE /= -subr_eq0 -val_eqE -(@natrB _ p 1) //=. +rewrite subn1 val_Fp_nat //; apply: contraFN (esym (mem_primes r 1)). +by rewrite pr_r /= -(eqnP co_ustar_p1) dvdn_gcd r_dv_ustar. +Qed. + +(* This is Peterfalvi (13.15). *) +Lemma card_FTtypeP_Galois_compl : + typeP_Galois StypeP -> u = (if p == 1 %[mod q] then ustar %/ q else ustar). +Proof. +case/typeP_Galois_P=> //= _ _ [_ _ /dvdnP[b]]; rewrite Ptype_factor_prime //. +rewrite -/ustar Ptype_Fcompl_kernel_trivial -(card_isog (quotient1_isog _)) -/u. +move=> Dbu; have ub_b: (b <= q.*2)%N. + have [[lambda Slam irrHlam]| ] := altP (@hasP _ irrIndH calS). + apply: (FTtypeP_Ind_Fitting_Galois_ub Slam irrHlam). + by rewrite FTtypeP_reg_Fcore indexg1. + case/(FTtypeP_no_Ind_Fitting_facts maxS StypeP) => _ /= ->. + rewrite indexg1 -/ustar -(leq_pmul2r (cardG_gt0 U)) -/u => Du. + by rewrite -Dbu -Du -(subnKC qgt2) leq_pmull. +have [ustar_odd p1_q p'1_q] := FTtypeP_primes_mod_cases. +have b_odd: odd b by rewrite Dbu odd_mul mFT_odd andbT in ustar_odd. +case: ifPn => [/p1_q q_dv_ustar | /p'1_q[_ _ /(_ b)]]. + have /dvdnP[c Db]: q %| b. + rewrite Dbu Gauss_dvdl // coprime_sym in q_dv_ustar. + by apply: coprimeSg coPUq; have /mulG_sub[_ sUPU] := sdprodW defPU. + have c_odd: odd c by rewrite Db odd_mul mFT_odd andbT in b_odd. + suffices /eqP c1: c == 1%N by rewrite Dbu Db c1 mul1n mulKn ?prime_gt0. + rewrite eqn_leq odd_gt0 // andbT -ltnS -(odd_ltn 3) // ltnS. + by rewrite -(leq_pmul2r (ltnW (ltnW qgt2))) -Db mul2n. +have Db: b = (b - 1).+1 by rewrite subn1 prednK ?odd_gt0. +rewrite Dbu dvdn_mulr // eqn_mod_dvd Db // -Db => /(_ isT)/dvdnP[c Db1]. +have c_even: ~~ odd c by rewrite Db Db1 /= odd_mul mFT_odd andbT in b_odd. +suffices /eqP->: b == 1%N by rewrite mul1n. +have:= ub_b; rewrite Db Db1 -mul2n ltn_pmul2r ?cardG_gt0 //. +by rewrite -ltnS odd_ltn //= !ltnS leqn0 => /eqP->. +Qed. + +(* This is Peterfalvi (13.16). *) +(* We have transposed T and Q here so that the lemma does not require *) +(* assumptions on the associate group. *) +Lemma FTtypeP_norm_cent_compl : P ><| W1 = 'N(W2) /\ P ><| W1 = 'C(W2). +Proof. +have [/mulG_sub[_ sW1S] /mulG_sub[sPPU sUPU]] := (sdprodW defS, sdprodW defPU). +have nPW1: W1 \subset 'N(P) by rewrite (subset_trans sW1S) ?gFnorm. +have [[_ _ frobUW1 cUU] [abelP _] _ _ _] := Sfacts. +have [pP cPP _] := and3P abelP; have [_ _ cW12 tiW12] := dprodP defW. +have cW2P: P \subset 'C(W2) by rewrite sub_abelian_cent. +suffices sNPW2: 'N(W2) \subset P <*> W1. + have cW2PW1: P <*> W1 \subset 'C(W2) by rewrite join_subG cW2P centsC. + rewrite sdprodEY ?coprime_TIg ?(coprimeSg sPPU) //. + split; apply/eqP; rewrite eqEsubset ?(subset_trans cW2PW1) ?cent_sub //. + by rewrite (subset_trans (cent_sub _)). +have tiP: normedTI P^# G S. + have [_ _ _] := compl_of_typeII_IV maxS StypeP notStype5. + by rewrite -defH FTtypeP_reg_Fcore dprodg1. +have ->: 'N(W2) = 'N_S(W2). + apply/esym/setIidPr/subsetP=> y nW2y; have [x W2x ntx] := trivgPn _ ntW2. + have [_ _ tiP_J] := normedTI_memJ_P tiP. + by rewrite -(tiP_J x) ?inE ?conjg_eq1 // ntx (subsetP sW2P) ?memJ_norm. +rewrite -{1}(sdprodW defS) setIC -group_modr ?cents_norm 1?centsC //=. +rewrite mulG_subG joing_subr /= -(sdprodW defPU) setIC. +rewrite -group_modl ?cents_norm //= mulG_subG joing_subl /= andbT. +set K := 'N_U(W2). +have nPKW1: K <*> W1 \subset 'N(P). + rewrite (subset_trans _ (gFnorm _ _)) // -(sdprodWY defS) genS ?setSU //. + by rewrite subIset ?sUPU. +have nW2KW1: K <*> W1 \subset 'N(W2). + by rewrite join_subG subsetIr cents_norm // centsC. +have coPKW1: coprime #|P| #|K <*> W1|. + by rewrite (coprimegS _ (Ptype_Fcore_coprime StypeP)) ?genS ?setSU ?subsetIl. +have p'KW1: p^'.-group (K <*> W1). + by rewrite /pgroup p'natE // -prime_coprime ?(coprimeSg sW2P). +have [Q1 defP nQ1KW1] := Maschke_abelem abelP p'KW1 sW2P nPKW1 nW2KW1. +have [-> | ntK] := eqVneq K 1%g; first by rewrite sub1G. +have frobKW1: [Frobenius K <*> W1 = K ><| W1]. + apply: Frobenius_subl frobUW1; rewrite ?subsetIl //. + rewrite normsI ?norms_norm //; first by have [_ []] := StypeP. + by rewrite cents_norm // centsC. +have regQ1W1: 'C_Q1(W1) = 1%g. + have [_ /mulG_sub[_ /setIidPl defQ1] _ tiW2Q1] := dprodP defP. + by rewrite -defQ1 -setIA (typeP_cent_core_compl StypeP) setIC. +have cQ1K: K \subset 'C(Q1). + have /mulG_sub[_ sQ1P] := dprodW defP; have coQ1KW1 := coprimeSg sQ1P coPKW1. + have solQ1 := solvableS sQ1P (abelian_sol cPP). + by have [_ ->] := Frobenius_Wielandt_fixpoint frobKW1 nQ1KW1 coQ1KW1 solQ1. +have /subsetIP[_ cW1K]: K \subset 'C_(K <*> W1)(W2). + have cCW1: W1 \subset 'C_(K <*> W1)(W2) by rewrite subsetI joing_subr centsC. + apply: contraR ntW1 => /(Frobenius_normal_proper_ker frobKW1) ltCK. + rewrite -subG1; have [/eqP/sdprodP[_ _ _ <-] _] := andP frobKW1. + rewrite subsetIidr (subset_trans cCW1) // proper_sub //. + rewrite ltCK //; last by rewrite norm_normalI ?norms_cent. + by rewrite (solvableS _ (abelian_sol cUU)) ?subsetIl. +case/negP: ntK; rewrite -subG1 -FTtypeP_reg_Fcore subsetI subsetIl /=. +by rewrite -(dprodW defP) centM subsetI cW1K. +Qed. + +End Thirteen_4_10_to_16. + +Section Thirteen_17_to_19. + +(* These assumptions repeat the part of Peterfalvi, Hypothesis (13.1) used *) +(* to prove (13.17-19). *) + +Variables S U W W1 W2 : {group gT}. +Hypotheses (maxS : S \in 'M) (defW : W1 \x W2 = W). +Hypotheses (StypeP : of_typeP S U defW). + +Local Notation "` 'W1'" := (gval W1) (at level 0, only parsing) : group_scope. +Local Notation "` 'W2'" := (gval W2) (at level 0, only parsing) : group_scope. +Local Notation "` 'W'" := (gval W) (at level 0, only parsing) : group_scope. +Local Notation V := (cyclicTIset defW). + +Local Notation "` 'S'" := (gval S) (at level 0, only parsing) : group_scope. +Local Notation P := `S`_\F%G. +Local Notation "` 'P'" := `S`_\F (at level 0) : group_scope. +Local Notation PU := S^`(1)%G. +Local Notation "` 'PU'" := `S^`(1) (at level 0) : group_scope. +Local Notation "` 'U'" := (gval U) (at level 0, only parsing) : group_scope. + +Let defS : PU ><| W1 = S. Proof. by have [[]] := StypeP. Qed. +Let defPU : P ><| U = PU. Proof. by have [_ []] := StypeP. Qed. + +Let notStype1 : FTtype S != 1%N. Proof. exact: FTtypeP_neq1 StypeP. Qed. +Let notStype5 : FTtype S != 5%N. Proof. exact: FTtype5_exclusion maxS. Qed. + +Let pddS := FT_prDade_hypF maxS StypeP. +Let ptiWS : primeTI_hypothesis S PU defW := FT_primeTI_hyp StypeP. +Let ctiWG : cyclicTI_hypothesis G defW := pddS. +Local Notation Sfacts := (FTtypeP_facts maxS StypeP). + +Let ntW1 : W1 :!=: 1. Proof. by have [[]] := StypeP. Qed. +Let ntW2 : W2 :!=: 1. Proof. by have [_ _ _ []] := StypeP. Qed. +Let cycW1 : cyclic W1. Proof. by have [[]] := StypeP. Qed. +Let cycW2 : cyclic W2. Proof. by have [_ _ _ []] := StypeP. Qed. + +Let p := #|W2|. +Let q := #|W1|. + +Let pr_p : prime p. Proof. by have [] := FTtypeP_primes maxS StypeP. Qed. +Let pr_q : prime q. Proof. by have [] := FTtypeP_primes maxS StypeP. Qed. + +Let qgt2 : q > 2. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. +Let pgt2 : p > 2. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. + +Let coPUq : coprime #|PU| q. +Proof. by rewrite (coprime_sdprod_Hall_r defS); have [[]] := StypeP. Qed. + +Let sW2P: W2 \subset P. Proof. by have [_ _ _ []] := StypeP. Qed. + +Let p'q : q != p. +Proof. +by rewrite -dvdn_prime2 -?prime_coprime -?(cyclic_dprod defW) //; case: ctiWG. +Qed. + +Let nirrW1 : #|Iirr W1| = q. Proof. by rewrite card_Iirr_cyclic. Qed. +Let nirrW2 : #|Iirr W2| = p. Proof. by rewrite card_Iirr_cyclic. Qed. +Let NirrW1 : Nirr W1 = q. Proof. by rewrite -nirrW1 card_ord. Qed. +Let NirrW2 : Nirr W2 = p. Proof. by rewrite -nirrW2 card_ord. Qed. + +Local Open Scope ring_scope. + +Let sigma := (cyclicTIiso ctiWG). +Let w_ i j := (cyclicTIirr defW i j). +Local Notation eta_ i j := (sigma (w_ i j)). + +Let mu_ := primeTIred ptiWS. +Local Notation tau := (FT_Dade0 maxS). + +Let calS0 := seqIndD PU S S`_\s 1. +Let rmR := FTtypeP_coh_base maxS StypeP. +Let scohS0 : subcoherent calS0 tau rmR. +Proof. exact: FTtypeP_subcoherent StypeP. Qed. + +Let calS := seqIndD PU S P 1. +Let sSS0 : cfConjC_subset calS calS0. +Proof. exact/seqInd_conjC_subset1/Fcore_sub_FTcore. Qed. + +(* This is Peterfalvi (13.17). *) +Lemma FTtypeII_support_facts T L (Q := T`_\F) (H := L`_\F) : + FTtype S == 2 -> typeP_pair S T defW -> L \in 'M('N(U)) -> + [/\ (*a*) [Frobenius L with kernel H], + (*b*) U \subset H + & (*c*) H ><| W1 = L \/ (exists2 y, y \in Q & H ><| (W1 <*> W2 :^ y) = L)]. +Proof. +move=> Stype2 pairST /setIdP[maxL sNU_L]. +have [pgt0 qgt0] := (ltnW (ltnW pgt2), ltnW (ltnW qgt2)). +have [[_ _ maxT] _ _ _ allST] := pairST. +have [[_ ntU _ _] _ not_sNU_S _ _] := compl_of_typeII maxS StypeP Stype2. +have [[_ _ frobUW1 cUU] _ _ _ _] := Sfacts. +have xdefW: W2 \x W1 = W by rewrite dprodC. +have [V TtypeP] := typeP_pairW (typeP_pair_sym xdefW pairST). +have [abelQ oQ]: q.-abelem Q /\ #|Q| = (q ^ p)%N. + by have [] := FTtypeP_facts maxT TtypeP. +have sUL: U \subset L := subset_trans (normG U) sNU_L. +have [/mulG_sub[sPPU sUPU] sPUS] := (sdprodW defPU, der_sub 1 S). +have nUW1: W1 \subset 'N(U) by have [_ []] := StypeP. +have sW1L := subset_trans nUW1 sNU_L. +have Ltype1: FTtype L == 1%N. + apply: contraR not_sNU_S => /allST/setUP[]// /imsetP[y _ defL]. + have hallU: \pi(U).-Hall(S) U. + have /Hall_pi/(subHall_Hall _ (piSg sUPU)): Hall PU U. + have /pHall_Hall:= pHall_subl sPPU sPUS (Fcore_Hall S). + by rewrite (sdprod_Hall defPU). + by apply; rewrite Hall_pi // -(coprime_sdprod_Hall_l defS). + have hallUy: \pi(U).-Hall(S) (U :^ y^-1). + by rewrite pHallE sub_conjgV -defL sUL /= cardJg -(card_Hall hallU). + have [x /conjGid <- ->] := Hall_trans (mmax_sol maxS) hallU hallUy. + by rewrite !normJ conjSg sub_conjgV -defL. + have oH: #|H| = (q ^ p)%N by rewrite /H defL FcoreJ cardJg. + have sW1H: W1 \subset H. + rewrite (sub_normal_Hall (Fcore_Hall L)) ?gFnormal //=. + by rewrite oH pi_of_exp ?prime_gt0 // pgroup_pi. + have regUW1: 'C_U(W1) = 1%g := Frobenius_trivg_cent frobUW1. + have /negP[] := ntU; rewrite -subG1 -regUW1 subsetIidl (sameP commG1P trivgP). + have /coprime_TIg <-: coprime #|U| #|H|. + by rewrite oH coprime_pexpr ?(coprimeSg sUPU). + rewrite commg_subI //; last by rewrite subsetI sW1H. + by rewrite subsetIidl (subset_trans sUL) ?gFnorm. +have frobL := FTtype1_Frobenius maxL Ltype1. +have solH: solvable H by rewrite nilpotent_sol ?Fcore_nil. +have coHW1: coprime #|H| #|W1|. + rewrite -(coprime_pexpr _ _ pgt0) -oQ. + apply/(coprimegS (Fcore_sub_FTcore maxT))/(coprimeSg (Fcore_sub_FTcore maxL)). + have [_ -> //] := FT_Dade_support_partition gT. + have: FTtype T != 1%N := FTtypeP_neq1 maxT TtypeP. + by apply: contra => /imsetP[y _ ->] /=; rewrite FTtypeJ. +have tiHW1: H :&: W1 = 1%g := coprime_TIg coHW1. +have sUH: U \subset H; last split=> //. + have [ntH _ /andP[sHL nHL] regHL] := Frobenius_kerP frobL. + have regHE E: gval E != 1%g -> E \subset L -> H :&: E = 1%g -> 'C_H(E) = 1%g. + move=> ntE sEL tiHE; apply: contraNeq ntE => /trivgPn[x /setIP[Hx cEx] ntx]. + rewrite -subG1 -tiHE subsetIidr (subset_trans _ (regHL x _)) ?inE ?ntx //. + by rewrite subsetI sEL sub_cent1. + suffices /trivgPn[x /setIP[Hx Ux] ntx]: H :&: U != 1%g. + apply: subset_trans (regHL x _); last by rewrite !inE ntx. + by rewrite subsetI sUL sub_cent1 (subsetP cUU). + apply: contraNneq (ntH) => tiHU; rewrite trivg_card1. + have [nHU nHW1] := (subset_trans sUL nHL, subset_trans sW1L nHL). + have nHUW1: U <*> W1 \subset 'N(H) by rewrite join_subG nHU. + have coHUW1: coprime #|H| #|U <*> W1|. + have [/eqP defUW1 _] := andP frobUW1. + rewrite (sdprodWY defUW1) -(sdprod_card defUW1) coprime_mulr coHW1 andbT. + have defHU: H ><| U = H <*> U by rewrite sdprodEY. + rewrite (coprime_sdprod_Hall_l defHU). + apply: pHall_Hall (pHall_subl (joing_subl _ _) _ (Fcore_Hall L)). + by rewrite join_subG sHL. + have [_ _] := Frobenius_Wielandt_fixpoint frobUW1 nHUW1 coHUW1 solH. + by move->; rewrite regHE // cards1 exp1n. +have [E sW1E frobHE]: exists2 E, W1 \subset gval E & [Frobenius L = H ><| E]. + have [E frobHE] := existsP frobL; have [/eqP defL _] := andP frobHE. + have hallE: \pi(H)^'.-Hall(L) E. + by rewrite -(compl_pHall E (Fcore_Hall L)) sdprod_compl. + have [|x Lx sW1Ex] := Hall_subJ (mmax_sol maxL) hallE sW1L. + by rewrite /pgroup -coprime_pi' ?cardG_gt0. + rewrite -(FrobeniusJ x) conjGid // (normsP (gFnorm _ _)) // in frobHE. + by exists (E :^ x)%G. +have [defL ntH ntE _ _] := Frobenius_context frobHE. +have [_ sEL _ nHE _] := sdprod_context defL. +have solE := solvableS sEL (mmax_sol maxL). +have [regHE regEH] := (Frobenius_reg_ker frobHE, Frobenius_reg_compl frobHE). +have qW1: q.-group W1 by apply: pnat_id. +have cycEr (r : nat) R: r.-group R -> R \subset E -> cyclic R. + move=> rR sRE; have nHR := subset_trans sRE nHE. + apply: odd_regular_pgroup_cyclic rR (mFT_odd _) ntH nHR _. + by move=> y /setD1P[nty Ry]; rewrite regHE // !inE nty (subsetP sRE). +have /normal_norm nW1E: W1 <| E. + exact: prime_odd_regular_normal (mFT_odd E) _ _ _ (Frobenius_reg_ker frobHE). +have defNW1: Q ><| W2 = 'N(W1). + by have [] := FTtypeP_norm_cent_compl maxT TtypeP. +have [nsQN sW2N _ _ _] := sdprod_context defNW1. +have sylQ: q.-Sylow('N(W1)) Q. + rewrite /pHall normal_sub // abelem_pgroup //=. + by rewrite -(index_sdprod defNW1) pnatE //= !inE eq_sym. +have hallW2: q^'.-Hall('N(W1)) W2 by rewrite -(compl_pHall _ sylQ) sdprod_compl. +pose Q1 := Q :&: E; have sylQ1: q.-Sylow(E) Q1 by apply: setI_normal_Hall nW1E. +have defQ1: Q1 = W1. + have abelQ1: q.-abelem Q1 := abelemS (subsetIl Q E) abelQ. + have sW1Q: W1 \subset Q by have [_ _ _ []] := TtypeP. + have sW1Q1: W1 \subset Q1 by apply/subsetIP. + have ntQ1: Q1 != 1%g by apply: subG1_contra ntW1. + apply/esym/eqP; rewrite eqEcard sW1Q1 (cyclic_abelem_prime abelQ1) //=. + by rewrite (cycEr q) ?(pHall_pgroup sylQ1) ?subsetIr. +have [P2 hallP2] := Hall_exists q^' solE; have [sP2E q'P2 _] := and3P hallP2. +have defE: W1 ><| P2 = E. + apply/(sdprod_normal_p'HallP _ hallP2); rewrite /= -defQ1 //. + by rewrite /Q1 setIC norm_normalI // (subset_trans nW1E) ?normal_norm. +have [P2_1 | ntP2] := eqsVneq P2 1%g. + by left; rewrite -defE P2_1 sdprodg1 in defL. +have solNW1: solvable 'N(W1). + by rewrite mFT_sol ?mFT_norm_proper // mFT_sol_proper (solvableS sW1E). +have [zy /=] := Hall_subJ solNW1 hallW2 (subset_trans sP2E nW1E) q'P2. +rewrite -{1}(sdprodWC defNW1) => /mulsgP[z y W2z Qy ->{zy}]. +rewrite conjsgM (conjGid W2z) {z W2z} => sP2W2y. +right; exists y => //; congr (_ ><| _ = _): defL. +rewrite -(sdprodWY defE); congr (W1 <*> _). +by apply/eqP; rewrite eqEsubset sP2W2y prime_meetG ?cardJg ?(setIidPr _). +Qed. + +Local Notation Imu2 := (primeTI_Iirr ptiWS). +Local Notation mu2_ i j := (primeTIirr ptiWS i j). + +Definition FTtypeP_bridge j := 'Ind[S, P <*> W1] 1 - mu2_ 0 j. +Local Notation beta_ := FTtypeP_bridge. +Definition FTtypeP_bridge_gap := tau (beta_ #1) - 1 + eta_ 0 #1. +Local Notation Gamma := FTtypeP_bridge_gap. + +Let u := #|U|. + +(* This is Peterfalvi (13.18). *) +(* Part (d) is stated with a slightly weaker hypothesis that fits better with *) +(* the usage pattern in (13.19) and (14.9). *) +Lemma FTtypeP_bridge_facts (V_S := class_support (cyclicTIset defW) S) : + [/\ (*a*) [/\ forall j, j != 0 -> beta_ j \in 'CF(S, 'A0(S)) + & forall j, j != 0 -> beta_ j \in 'CF(S, P^# :|: V_S)], + (*b*) forall j, j != 0 -> '[beta_ j] = (u.-1 %/ q + 2)%:R, + (*c*) [/\ forall j, j != 0 -> tau (beta_ j) - 1 + eta_ 0 j = Gamma, + '[Gamma, 1] = 0 & cfReal Gamma], + (*d*) forall X Y : 'CF(G), + Gamma = X + Y -> '[X, Y] = 0 -> + orthogonal Y (map sigma (irr W)) -> + '[Y] <= (u.-1 %/ q)%:R + & q %| u.-1]. +Proof. +have [_ sW1S _ nPUW1 tiPUW1] := sdprod_context defS. +have /mulG_sub[sPPU sUPU] := sdprodW defPU. +have sPW1S: P <*> W1 \subset S by rewrite join_subG gFsub. +have /= defS_P := Ptype_Fcore_sdprod StypeP; have nsPS: P <| S := gFnormal _ _. +have defPW1: P ><| W1 = P <*> W1 := sdprod_subr defS_P (joing_subr U W1). +pose W1bar := (W1 / P)%g; pose Sbar := (S / P)%g; pose Ubar := (U / P)%g. +pose gamma := 'Ind[Sbar, W1bar] 1. +have Dgamma: 'Ind[S, P <*> W1] 1 = (gamma %% P)%CF. + rewrite -(rmorph1 _ : 1 %% P = 1)%CF cfIndMod ?joing_subl //. + by rewrite quotientYidl //; have [] := sdprodP defPW1. +have gamma1: gamma 1%g = u%:R. + rewrite -cfMod1 -Dgamma cfInd1 // cfun11 -divgS // -(sdprod_card defPW1). + by rewrite mulr1 -(sdprod_card defS) -(sdprod_card defPU) divnMr // mulKn. +have frobUW1: [Frobenius U <*> W1 = U ><| W1] by have [[]] := Sfacts. +have q_dv_u1: q %| u.-1 := Frobenius_dvd_ker1 frobUW1. +have [nP_UW1 /isomP[/=]] := sdprod_isom defS_P; set h := restrm _ _ => injh hS. +have /joing_sub[sUUW1 sW1UW1] := erefl (U <*> W1). +have [hU hW1]: h @* U = Ubar /\ h @* W1 = W1bar. + by rewrite !morphim_restrm /= !(setIidPr _). +have{hS} frobSbar: [Frobenius Sbar = Ubar ><| W1bar]. + by rewrite -[Sbar]hS -hU -hW1 injm_Frobenius. +have tiW1bar: normedTI W1bar^# Sbar W1bar by have /and3P[] := frobSbar. +have gammaW1 xbar: xbar \in W1bar^# -> gamma xbar = 1. + move=> W1xbar; have [ntxbar _] := setD1P W1xbar. + rewrite cfIndE ?quotientS //; apply: canLR (mulKf (neq0CG _)) _. + have ->: #|W1bar| = #|Sbar :&: W1bar| by rewrite (setIidPr _) ?quotientS. + rewrite mulr1 cardsE -sumr_const big_mkcondr; apply: eq_bigr => zbar Szbar. + have [_ _ W1bar_xJ] := normedTI_memJ_P tiW1bar. + by rewrite -mulrb -(W1bar_xJ xbar) // !inE conjg_eq1 ntxbar cfun1E. +have PVSbeta j: j != 0 -> beta_ j \in 'CF(S, P^# :|: V_S). + move=> nzj; apply/cfun_onP=> z; rewrite !inE => /norP[P'z VS'z]. + have [Sz | /cfun0->//] := boolP (z \in S); apply/eqP; rewrite !cfunE subr_eq0. + have [[_ mulW12 _ tiW12] C1] := (dprodP defW, FTtypeP_reg_Fcore maxS StypeP). + have [PUz {VS'z} | PU'z {P'z}] := boolP (z \in PU). + rewrite eq_sym -(cfResE _ _ PUz) ?gFsub // -['Res _](scalerK (neq0CG W1)). + rewrite cfRes_prTIirr -cfRes_prTIred -/q cfunE cfResE ?gFsub // mulrC. + case/nandP: P'z => [/negbNE/eqP-> | P'z]. + rewrite Dgamma cfModE // morph1 gamma1 FTprTIred1 // C1 indexg1. + by rewrite natrM mulfK ?neq0CG. + have:= seqInd_on (Fitting_normal S) (FTprTIred_Ind_Fitting maxS StypeP nzj). + have [/= <- _ _ _] := typeP_context StypeP; rewrite C1 dprodg1 -/(mu_ j). + move/cfun_on0->; rewrite // mul0r (cfun_on0 (cfInd_on _ (cfun_onG _))) //. + rewrite -(sdprodW defPW1); apply: contra P'z => /imset2P[x t PW1x St Dz]. + rewrite Dz !memJ_norm ?(subsetP (gFnorm _ _)) // in PUz *. + by rewrite -(mulg1 P) -tiPUW1 setIC group_modl // inE PW1x. + have /imset2P[x t /setD1P[ntx W1x] St ->]: z \in class_support W1^# S. + have /bigcupP[_ /rcosetsP[x W1x ->]]: z \in cover (rcosets PU W1). + by rewrite (cover_partition (rcosets_partition_mul _ _)) (sdprodW defS). + have [-> | ntx] := eqVneq x 1%g; first by rewrite mulg1 => /idPn[]. + have nPUx: x \in 'N(PU) by rewrite (subsetP nPUW1). + have coPUx: coprime #|PU| #[x] by rewrite (coprime_dvdr (order_dvdG W1x)). + have [/cover_partition <- _] := partition_cent_rcoset nPUx coPUx. + have [_ _ _ [_ _ _ _ prPUW1] _] := StypeP; rewrite {}prPUW1 ?inE ?ntx //. + rewrite cover_imset => /bigcupP[t PUt /imsetP[_ /rcosetP[y W2y ->] Dz]]. + have{PUt} St: t \in S by rewrite (subsetP _ _ PUt) ?der_sub. + have [y1 | nty] := eqVneq y 1%g. + by rewrite Dz y1 mul1g memJ_class_support // !inE ntx. + rewrite Dz memJ_class_support // !inE groupMr // groupMl // in VS'z. + rewrite -(dprodWC defW) mem_mulg // andbT; apply/norP. + by rewrite -!in_set1 -set1gE -tiW12 !inE W1x W2y andbT in ntx nty. + rewrite !cfunJ // Dgamma cfModE ?(subsetP sW1S) // gammaW1; last first. + by rewrite !inE (morph_injm_eq1 injh) ?(subsetP sW1UW1) ?ntx ?mem_quotient. + rewrite prTIirr_id ?FTprTIsign // ?scale1r ?dprod_IirrEr; last first. + rewrite -in_set1 -set1gE -tiW12 inE W1x /= in ntx. + by rewrite inE ntx -mulW12 (subsetP (mulG_subl W2 W1)). + by rewrite -[x]mulg1 cfDprodEr ?lin_char1 ?irr_prime_lin. +have A0beta j: j != 0 -> beta_ j \in 'CF(S, 'A0(S)). + move/PVSbeta; apply: cfun_onS; rewrite (FTtypeP_supp0_def _ StypeP) //. + by rewrite setSU ?(subset_trans _ (FTsupp1_sub _)) ?setSD ?Fcore_sub_FTcore. +have norm_beta j: j != 0 -> '[beta_ j] = (u.-1 %/ q + 2)%:R. + move=> nzj; rewrite cfnormBd ?Dgamma; last first. + apply: contraNeq (cfker_prTIres pddS nzj); rewrite -irr_consttE => S1_mu0j. + rewrite -(cfRes_prTIirr _ 0) sub_cfker_Res //. + rewrite (subset_trans _ (cfker_constt _ S1_mu0j)) ?cfker_mod //. + by rewrite -Dgamma cfInd_char ?rpred1. + have [[/eqP defUW1 _] [/eqP defSbar _]] := (andP frobUW1, andP frobSbar). + rewrite cfnorm_irr cfMod_iso //. + rewrite (cfnormE (cfInd_on _ (cfun_onG _))) ?quotientS // -/gamma. + rewrite card_quotient ?gFnorm // -(index_sdprod defS_P) -(sdprod_card defUW1). + rewrite -/u -/q (big_setD1 1%g) ?mem_class_support ?group1 //=. + have{tiW1bar} [_ tiW1bar /eqP defNW1bar] := and3P tiW1bar. + rewrite gamma1 normr_nat class_supportD1 big_trivIset //=. + rewrite (eq_bigr (fun xbar => #|W1bar|.-1%:R)) ?sumr_const; last first. + rewrite (cardsD1 1%g) group1 /= => _ /imsetP[tbar Stbar ->]. + rewrite -sumr_const big_imset /=; last exact: in2W (conjg_inj tbar). + by apply: eq_bigr => xbar W1xbar; rewrite cfunJ ?gammaW1 // normr1 expr1n. + rewrite card_conjugates -divgS ?subsetIl //= -(sdprod_card defSbar) defNW1bar. + rewrite mulnK ?cardG_gt0 // -hU -hW1 ?card_injm // -/q -/u natrM invfM mulrC. + rewrite -[rhs in _ ^+ 2 + rhs]mulr_natr -mulrDl mulrA mulfK ?neq0CG //. + rewrite -subn1 natrB ?cardG_gt0 // addrCA mulrDl divff ?neq0CG //. + by rewrite -natrB ?cardG_gt0 // subn1 -natf_div // addrAC addrC natrD. +have nzj1: #1 != 0 :> Iirr W2 by apply: Iirr1_neq0. +have [_ _ _ _ [_ Dtau]] := Sfacts; pose eta01 := eta_ 0 #1. +have oeta01_1: '[eta01, 1] = 0. + by rewrite -(cycTIiso1 ctiWG) -(cycTIirr00 defW) cfdot_cycTIiso (negPf nzj1). +have Deta01s: eta01^*%CF = eta_ 0 (conjC_Iirr #1). + by rewrite cfAut_cycTIiso /w_ !dprod_IirrEr cfAutDprodr aut_IirrE. +have oGamma1: '[Gamma, 1] = 0. + rewrite cfdotDl cfdotBl cfnorm1 oeta01_1 addr0 Dtau ?A0beta //. + rewrite -cfdot_Res_r rmorph1 cfdotBl -cfdot_Res_r rmorph1 cfnorm1. + by rewrite -(prTIirr00 ptiWS) cfdot_prTIirr (negPf nzj1) subr0 subrr. +have defGamma j: j != 0 -> tau (beta_ j) - 1 + eta_ 0 j = Gamma. + move=> nzj; apply/eqP; rewrite -subr_eq0 opprD addrACA opprB !addrA subrK. + rewrite -linearB opprD addrACA subrr add0r -opprD linearN /=. + move/prDade_sub_TIirr: pddS => -> //; last first. + by apply: (mulfI (neq0CG W1)); rewrite -!prTIred_1 !FTprTIred1. + by rewrite -/sigma FTprTIsign // scale1r -addrA addNr. +have GammaReal: cfReal Gamma. + rewrite /cfReal rmorphD rmorphB rmorph1 /= Deta01s Dtau ?A0beta // cfAutInd. + rewrite rmorphB /= cfAutInd rmorph1 -prTIirr_aut aut_Iirr0 -/(beta_ _). + by rewrite -Dtau ?A0beta ?defGamma ?aut_Iirr_eq0. +split=> // X Y defXY oXY oYeta; pose a := '[Gamma, eta01]. +have Za: a \in Cint. + rewrite Cint_cfdot_vchar ?(rpredB, rpredD, rpred1, cycTIiso_vchar) //. + by rewrite Dtau ?A0beta // !(cfInd_vchar, rpredB) ?rpred1 ?irr_vchar. +have{oYeta} oYeta j: '[Y, eta_ 0 j] = 0. + by rewrite (orthoPl oYeta) ?map_f ?mem_irr. +have o_eta1s1: '[eta01^*, eta01] = 0. + rewrite Deta01s cfdot_cycTIiso /= -(inj_eq irr_inj) aut_IirrE. + by rewrite odd_eq_conj_irr1 ?mFT_odd // irr_eq1 (negPf nzj1). +rewrite -(ler_add2r 2%:R) -natrD -(norm_beta #1) //. +have ->: '[beta_ #1] = '[Gamma - eta01 + 1]. + by rewrite addrK subrK Dade_isometry ?A0beta. +rewrite addrA cfnormDd ?cfnorm1 ?ler_add2r; last first. + by rewrite cfdotBl oeta01_1 oGamma1 subrr. +rewrite defXY addrAC addrC cfnormDd ?ler_add2r; last first. + by rewrite cfdotBl oXY cfdotC oYeta conjC0 subrr. +have oXeta j: '[X, eta_ 0 j] = '[Gamma, eta_ 0 j]. + by rewrite defXY cfdotDl oYeta addr0. +pose X1 := X - a *: eta01 - a *: eta01^*%CF. +have ->: X - eta01 = X1 + a *: eta01^*%CF + (a - 1) *: eta01. + by rewrite scalerBl scale1r addrA !subrK. +rewrite cfnormDd; last first. + rewrite cfdotZr subrK cfdotBl oXeta -/a cfdotZl cfnorm_cycTIiso mulr1. + by rewrite subrr mulr0. +rewrite cfnormDd; last first. + rewrite cfdotZr !cfdotBl !cfdotZl Deta01s cfnorm_cycTIiso oXeta -Deta01s. + rewrite !cfdot_conjCr o_eta1s1 conjC0 mulr0 ((_ =P Gamma) GammaReal) -/a. + by rewrite conj_Cint // mulr1 subr0 subrr mulr0. +rewrite -addrA ler_paddl ?cfnorm_ge0 // !cfnormZ Deta01s !cfnorm_cycTIiso. +rewrite !mulr1 !Cint_normK ?rpredB ?rpred1 // sqrrB1 !addrA -mulr2n. +by rewrite -subr_ge0 addrK subr_ge0 ler_pmuln2r ?Cint_ler_sqr. +Qed. + +(* The assumptions of Peterfalvi (13.19). *) +(* We do not need to put these in a subsection as this is the last Lemma. *) +Variable L : {group gT}. +Hypotheses (maxL : L \in 'M) (Ltype1 : FTtype L == 1%N). + +Local Notation "` 'L'" := (gval L) (at level 0, only parsing) : group_scope. +Local Notation H := `L`_\F%G. +Local Notation "` 'H'" := `L`_\F (at level 0) : group_scope. + +Let e := #|L : H|. +Let tauL := FT_DadeF maxL. +Let calL := seqIndD H L H 1. + +Let frobL : [Frobenius L with kernel H]. Proof. exact: FTtype1_Frobenius. Qed. + +(* The coherence part of the preamble of (13.19). *) +Lemma FTtype1_coherence : coherent calL L^# tauL. +Proof. +have [_ [tau1 [IZtau1 Dtau1]]] := FT_Frobenius_coherence maxL frobL. +exists tau1; split=> // phi Sphi; rewrite ?Dtau1 //. +move/(zcharD1_seqInd_on (Fcore_normal _)) in Sphi. +by rewrite /tauL FT_DadeF_E ?FT_DadeE ?(cfun_onS (Fcore_sub_FTsupp _)). +Qed. + +Lemma FTtype1_Ind_irr : {subset calL <= irr L}. +Proof. by case: (FT_Frobenius_coherence maxL frobL). Qed. +Let irrL := FTtype1_Ind_irr. + +(* We re-quantify over the witnesses so that the main part of the lemma can *) +(* be used for Section variables in the very last part of Section 14. *) +Variables (tau1 : {additive 'CF(L) -> 'CF(G)}) (phi : 'CF(L)). +Hypothesis cohL : coherent_with calL L^# tauL tau1. +Hypotheses (Lphi : phi \in calL) (phi1e : phi 1%g = e%:R). + +Let betaL := 'Ind[L, H] 1 - phi. +Let betaS := beta_ #1. +Let eta01 := eta_ 0 #1. + +(* This is Peterfalvi (13.19). *) +Lemma FTtypeI_bridge_facts : + [/\ (*a*) 'A~(L) :&: (class_support P G :|: class_support W G) = set0, + (*b*) orthogonal (map tau1 calL) (map sigma (irr W)), + (*c*) forall j, j != 0 -> '[tauL betaL, eta_ 0 j] = '[tauL betaL, eta01] + & (*c1*) ('[tau betaS, tau1 phi] == 1 %[mod 2])%C + /\ #|H|.-1%:R / e%:R <= (u.-1 %/ q)%:R :> algC + \/ (*c2*) ('[tauL betaL, eta01] == 1 %[mod 2])%C /\ (p <= e)%N]. +Proof. +have nsHL: H <| L := gFnormal _ L; have [sHL nHL] := andP nsHL. +have coHr T r: T \in 'M -> FTtype T != 1%N -> r.-abelem T`_\F -> coprime #|H| r. + move=> maxT notTtype1 /andP[rR _]. + have [_ _ [n oR]] := pgroup_pdiv rR (mmax_Fcore_neq1 maxT). + rewrite -(coprime_pexpr _ r (ltn0Sn n)) -oR /= -FTcore_type1 //. + apply: coprimegS (Fcore_sub_FTcore maxT) _. + have [_ -> //] := FT_Dade_support_partition gT. + by apply: contra notTtype1 => /imsetP[y _ ->] /=; rewrite FTtypeJ. +have coHp: coprime #|H| p by apply: (coHr S) => //; have [_ []] := Sfacts. +have{coHr} coHq: coprime #|H| q. + have [T pairST [xdefW [V TtypeP]]] := FTtypeP_pair_witness maxS StypeP. + have [[_ _ maxT] _ _ _ _] := pairST; have Ttype'1 := FTtypeP_neq1 maxT TtypeP. + by rewrite (coHr T) ?Ttype'1 //; have [_ []] := FTtypeP_facts maxT TtypeP. +have defA: 'A(L) = H^# := FTsupp_Frobenius maxL frobL. +set PWG := class_support P G :|: class_support W G. +have tiA_PWG: 'A~(L) :&: PWG = set0. + apply/setP=> x; rewrite !inE; apply/andP=> [[Ax PWGx]]. + suffices{Ax}: \pi(H)^'.-elt x. + have [y Ay /imset2P[_ t /rcosetP[z Rz ->] _ ->]] := bigcupP Ax => H'zyt. + do [rewrite -def_FTsignalizer //; set ddL := FT_Dade_hyp maxL] in Rz. + have /setD1P[nty Hy]: y \in H^# by rewrite -defA. + have /idPn[]: (z * y).`_\pi('C_H[y]) == 1%g. + rewrite (constt1P _) // -(p_eltJ _ _ t); apply: sub_in_pnat H'zyt => r _. + by apply: contra; apply: piSg; apply: subsetIl. + rewrite consttM; last first. + exact: cent1P (subsetP (Dade_signalizer_cent _ y) z Rz). + rewrite (constt1P (mem_p_elt _ Rz)) ?mul1g; last first. + rewrite /pgroup -coprime_pi' ?cardG_gt0 // coprime_sym. + by rewrite (coprimegS _ (Dade_coprime _ Ay Ay)) ?setSI. + by rewrite (constt_p_elt (mem_p_elt (pgroup_pi _) _)) // inE Hy cent1id. + suffices /pnat_dvd: #[x] %| #|P| * #|W|. + have [_ [_ ->] _ _ _] := Sfacts; rewrite -(dprod_card defW) -/p -/q. + by apply; rewrite !pnat_mul pnat_exp -!coprime_pi' ?cardG_gt0 ?coHp ?coHq. + case/orP: PWGx => /imset2P[y z PWy _ ->]; rewrite {z}orderJ. + by rewrite dvdn_mulr ?order_dvdG. + by rewrite dvdn_mull ?order_dvdG. +have ZsubL psi: psi \in calL -> psi - psi^*%CF \in 'Z[calL, L^#]. + have ZcalL: {subset calL <= 'Z[irr L]} by apply: seqInd_vcharW. + by move=> Lpsi; rewrite sub_aut_zchar ?zchar_onG ?mem_zchar ?cfAut_seqInd. +have mem_eta j: eta_ 0 j \in map sigma (irr W) by rewrite map_f ?mem_irr. +have otau1eta: orthogonal (map tau1 calL) (map sigma (irr W)). + apply/orthogonalP=> _ _ /mapP[psi Lpsi ->] /mapP[w irr_w ->]. + have{w irr_w} [i [j ->]] := cycTIirrP defW irr_w; rewrite -/(w_ i j). + pose Psi := tau1 (psi - psi^*%CF); pose NC := cyclicTI_NC ctiWG. + have [[Itau1 Ztau1] Dtau1] := cohL. + have Lpsis: psi^*%CF \in calL by rewrite cfAut_seqInd. + have Z1dpsi := ZsubL _ Lpsi; have Zdpsi := zcharW Z1dpsi. + have{Dtau1} PsiV0: {in V, Psi =1 \0}. + move=> x /setDP[Wx _]; rewrite /Psi Dtau1 ?(cfun_on0 (Dade_cfunS _ _)) //. + rewrite FT_DadeF_supportE -defA; apply: contra_eqN tiA_PWG => Ax. + by apply/set0Pn; exists x; rewrite !inE Ax orbC mem_class_support. + have opsi: '[psi, psi^*] = 0 by apply: seqInd_conjC_ortho (mFT_odd _) _ Lpsi. + have n2Psi: '[Psi] = 2%:R. + by rewrite Itau1 ?cfnormBd // cfnorm_conjC ?irrWnorm ?irrL. + have NC_Psi: (NC Psi < minn q p)%N. + by rewrite (@leq_ltn_trans 2) ?leq_min ?qgt2 // cycTI_NC_norm ?Ztau1 ?n2Psi. + apply: contraTeq (NC_Psi) => t1psi_eta; rewrite -leqNgt cycTI_NC_minn //. + rewrite mul2n -addnn (leq_trans NC_Psi) ?leq_addl // andbT card_gt0. + suffices [b Deta]: exists b : bool, eta_ i j = (-1) ^+ b *: tau1 psi. + apply/set0Pn; exists (i, j); rewrite !inE /= /Psi raddfB cfdotBl {2}Deta. + by rewrite cfdotZr Itau1 ?mem_zchar // cfdot_conjCl opsi conjC0 mulr0 subr0. + exists (tau1 psi == - eta_ i j); apply: (canRL (signrZK _)). + move/eqP: t1psi_eta; rewrite cfdot_dirr ?cycTIiso_dirr //; last first. + by rewrite dirrE Itau1 ?Ztau1 ?mem_zchar //= irrWnorm ?irrL. + by rewrite scaler_sign; do 2!case: eqP => //. +have [[A0beta PVbeta] n2beta [defGa Ga1 R_Ga] ubGa dvu] := FTtypeP_bridge_facts. +have [_ _ _ _ [_ Dtau]] := Sfacts. +have o_tauL_S zeta j: j != 0 -> '[tauL zeta, tau (beta_ j)] = 0. + move=> nzj; pose ABS := class_support (P^# :|: class_support V S) G. + have ABSbeta: tau (beta_ j) \in 'CF(G, ABS). + by rewrite Dtau ?A0beta // cfInd_on ?subsetT ?PVbeta. + have{ABSbeta} PWGbeta: tau (beta_ j) \in 'CF(G, PWG). + apply: cfun_onS ABSbeta; apply/subsetP=> _ /imset2P[x t PVSx _ ->]. + case/setUP: PVSx => [/setD1P[_ Px] | /imset2P[y z /setDP[Wy _] _ ->]]. + by rewrite inE memJ_class_support ?inE. + by rewrite -conjgM inE orbC memJ_class_support ?inE. + rewrite (cfdotElr (Dade_cfunS _ _) PWGbeta) big_pred0 ?mulr0 // => x. + by rewrite FT_DadeF_supportE -defA tiA_PWG inE. +have betaLeta j: j != 0 -> '[tauL betaL, eta_ 0 j] = '[tauL betaL, eta01]. + move=> nzj; apply/eqP; rewrite -subr_eq0 -cfdotBr. + rewrite (canRL (addKr _) (defGa j nzj)) !addrA addrK -addrA addrCA. + by rewrite opprD subrK cfdotBr !o_tauL_S ?subrr ?Iirr1_neq0. +split=> //; have [[[Itau1 Ztau1] Dtau1] irr_phi] := (cohL, irrL Lphi). +pose GammaL := tauL betaL - (1 - tau1 phi). +have DbetaL: tauL betaL = 1 - tau1 phi + GammaL by rewrite addrC subrK. +have RealGammaL: cfReal GammaL. + rewrite /cfReal -subr_eq0 !rmorphB rmorph1 /= !opprB !addrA subrK addrC. + rewrite -addrA addrCA addrA addr_eq0 opprB -Dade_aut -linearB /= -/tauL. + rewrite rmorphB /= cfAutInd rmorph1 addrC opprB addrA subrK. + by rewrite (cfConjC_Dade_coherent cohL) ?mFT_odd // -raddfB Dtau1 // ZsubL. +have:= Dade_Ind1_sub_lin cohL _ irr_phi Lphi; rewrite -/betaL -/tauL -/calL. +rewrite (seqInd_nontrivial _ _ _ irr_phi) ?odd_Frobenius_index_ler ?mFT_odd //. +case=> // -[o_tauL_1 o_betaL_1 ZbetaL] ub_betaL _. +have{o_tauL_1 o_betaL_1} o_GaL_1: '[GammaL, 1] = 0. + by rewrite !cfdotBl cfnorm1 o_betaL_1 (orthoPr o_tauL_1) ?map_f ?subr0 ?subrr. +have Zt1phi: tau1 phi \in 'Z[irr G] by rewrite Ztau1 ?mem_zchar. +have Zeta01: eta01 \in 'Z[irr G] by apply: cycTIiso_vchar. +have ZbetaS: tau betaS \in 'Z[irr G]. + rewrite Dade_vchar // zchar_split A0beta ?Iirr1_neq0 //. + by rewrite rpredB ?irr_vchar ?cfInd_vchar ?rpred1. +have Z_Ga: Gamma \in 'Z[irr G] by rewrite rpredD ?rpredB ?rpred1. +have Z_GaL: GammaL \in 'Z[irr G] by rewrite !rpredB ?rpred1. +have{RealGammaL} Gamma_even: (2 %| '[GammaL, Gamma])%C. + by rewrite cfdot_real_vchar_even ?mFT_odd // o_GaL_1 (dvdC_nat 2 0). +set bSphi := '[tau betaS, tau1 phi]; set bLeta := '[tauL betaL, eta01]. +have [ZbSphi ZbLeta]: bSphi \in Cint /\ bLeta \in Cint. + by rewrite !Cint_cfdot_vchar. +have{Gamma_even} odd_bSphi_bLeta: (bSphi + bLeta == 1 %[mod 2])%C. + rewrite -(conj_Cint ZbSphi) -cfdotC /bLeta DbetaL cfdotDl cfdotBl. + have: '[tauL betaL, tau betaS] == 0 by rewrite o_tauL_S ?Iirr1_neq0. + have ->: tau betaS = 1 - eta01 + Gamma by rewrite addrCA !addrA !subrK. + rewrite !['[tau1 _, _]]cfdotDr 2!cfdotDr !cfdotNr DbetaL. + rewrite 2!cfdotDl 2!['[_, eta01]]cfdotDl 2!['[_, Gamma]]cfdotDl !cfdotNl. + rewrite cfnorm1 o_GaL_1 ['[1, Gamma]]cfdotC Ga1 conjC0 addr0 add0r. + have ->: 1 = eta_ 0 0 by rewrite /w_ cycTIirr00 cycTIiso1. + rewrite cfdot_cycTIiso mulrb ifN_eqC ?Iirr1_neq0 // add0r. + rewrite 2?(orthogonalP otau1eta _ _ (map_f _ _) (mem_eta _)) // oppr0 !add0r. + by rewrite addr0 addrA addrC addr_eq0 !opprB addrA /eqCmod => /eqP <-. +have abs_mod2 a: a \in Cint -> {b : bool | a == b%:R %[mod 2]}%C. + move=> Za; pose n := truncC `|a|; exists (odd n). + apply: eqCmod_trans (eqCmod_addl_mul _ (rpred_nat _ n./2) _). + rewrite addrC -natrM -natrD muln2 odd_double_half truncCK ?Cnat_norm_Cint //. + rewrite -{1}[a]mul1r -(canLR (signrMK _) (CintEsign Za)) eqCmodMr // signrE. + by rewrite /eqCmod opprB addrC subrK dvdC_nat dvdn2 odd_double. +have [[bL DbL] [bS DbS]] := (abs_mod2 _ ZbLeta, abs_mod2 _ ZbSphi). +have{odd_bSphi_bLeta} xor_bS_bL: bS (+) bL. + rewrite eqCmod_sym in odd_bSphi_bLeta. + have:= eqCmod_trans odd_bSphi_bLeta (eqCmodD DbS DbL). + rewrite -natrD eqCmod_sym -(eqCmodDr _ 1) -mulrSr => xor_bS_bL. + have:= eqCmod_trans xor_bS_bL (eqCmodm0 _); rewrite /eqCmod subr0. + by rewrite (dvdC_nat 2 _.+1) dvdn2 /= negbK odd_add !oddb; case: (_ (+) _). +have ?: (0 != 1 %[mod 2])%C by rewrite eqCmod_sym /eqCmod subr0 (dvdC_nat 2 1). +case is_c1: bS; [left | right]. + rewrite is_c1 in DbS; split=> //. + pose a_ (psi : 'CF(L)) := psi 1%g / e%:R. + have Na_ psi: psi \in calL -> a_ psi \in Cnat by apply: dvd_index_seqInd1. + have [X tau1X [D [dGa oXD oDtau1]]] := orthogonal_split (map tau1 calL) Gamma. + have oo_L: orthonormal calL. + by apply: sub_orthonormal (irr_orthonormal L); rewrite ?seqInd_uniq. + have oo_tau1L: orthonormal (map tau1 calL) by apply: map_orthonormal. + have defX: X = bSphi *: (\sum_(psi <- calL) a_ psi *: tau1 psi). + have [_ -> defX] := orthonormal_span oo_tau1L tau1X. + rewrite defX big_map scaler_sumr; apply: eq_big_seq => psi Lpsi. + rewrite scalerA; congr (_ *: _); apply/eqP; rewrite -subr_eq0 mulrC. + rewrite -[X](addrK D) -dGa cfdotBl (orthoPl oDtau1) ?map_f // subr0. + rewrite cfdotC cfdotDr cfdotBr -/betaS -/eta01. + have ->: 1 = eta_ 0 0 by rewrite /w_ cycTIirr00 cycTIiso1. + rewrite 2?(orthogonalP otau1eta _ _ (map_f _ _) (mem_eta _)) // subrK. + rewrite -cfdotC -(conj_Cnat (Na_ _ Lpsi)) -cfdotZr -cfdotBr. + rewrite -raddfZ_Cnat ?Na_ // -raddfB cfdotC. + rewrite Dtau1; last by rewrite zcharD1_seqInd ?seqInd_sub_lin_vchar. + by rewrite o_tauL_S ?Iirr1_neq0 ?conjC0. + have nz_bSphi: bSphi != 0 by apply: contraTneq DbS => ->. + have ub_a: \sum_(psi <- calL) a_ psi ^+ 2 <= (u.-1 %/ q)%:R. + apply: ler_trans (ubGa D X _ _ _); first 1 last; first by rewrite addrC. + - by rewrite cfdotC oXD conjC0. + - apply/orthoPl=> eta Weta; rewrite (span_orthogonal otau1eta) //. + exact: memv_span. + rewrite defX cfnormZ cfnorm_sum_orthonormal // mulr_sumr !big_seq. + apply: ler_sum => psi Lpsi; rewrite -{1}(norm_Cnat (Na_ _ _)) //. + by rewrite ler_pemull ?exprn_ge0 ?normr_ge0 // Cint_normK // sqr_Cint_ge1. + congr (_ <= _): ub_a; do 2!apply: (mulIf (neq0CiG L H)); rewrite -/e. + rewrite divfK ?neq0CiG // -mulrA -expr2 mulr_suml. + rewrite -subn1 natrB ?neq0CG // -indexg1 mulrC. + rewrite -(sum_seqIndD_square nsHL) ?normal1 ?sub1G // -/calL. + apply: eq_big_seq => psi Lpsi; rewrite irrWnorm ?irrL // divr1. + by rewrite -exprMn divfK ?neq0CiG. +rewrite is_c1 /= in xor_bS_bL; rewrite xor_bS_bL in DbL; split=> //. +have nz_bL: bLeta != 0 by apply: contraTneq DbL => ->. +have{ub_betaL} [X [otau1X oX1 [a Za defX]] [//|_ ubX]] := ub_betaL. +rewrite -/e in defX; rewrite -leC_nat -(ler_add2r (-1)); apply: ler_trans ubX. +pose calX0 := [seq w_ 0 j | j in predC1 0]. +have ooX0: orthonormal calX0. + apply: sub_orthonormal (irr_orthonormal W). + by move=> _ /imageP[j _ ->]; apply: mem_irr. + by apply/dinjectiveP=> j1 j2 _ _ /irr_inj/dprod_Iirr_inj[]. +have Isigma: {in 'Z[calX0] &, isometry sigma}. + by apply: in2W; apply: cycTIisometry. +rewrite -[X](subrK (bLeta *: (\sum_(xi <- calX0) sigma xi))). +rewrite cfnormDd ?ler_paddl ?cfnorm_ge0 //; last first. + rewrite cfdotZr cfdot_sumr big1_seq ?mulr0 // => xi X0xi. + apply/eqP; rewrite cfdotBl scaler_sumr cfproj_sum_orthonormal // subr_eq0. + have {xi X0xi}[j nzj ->] := imageP X0xi; rewrite inE /= in nzj. + rewrite -[bLeta](betaLeta j nzj) defX cfdotDl -addrA cfdotDl. + have ->: 1 = eta_ 0 0 by rewrite /w_ cycTIirr00 cycTIiso1. + rewrite cfdot_cycTIiso mulrb (ifN_eqC _ _ nzj) add0r eq_sym -subr_eq0 addrK. + rewrite (span_orthogonal otau1eta) //; last by rewrite memv_span ?mem_eta. + rewrite big_seq rpredD ?(rpredN, rpredZ, rpred_sum) ?memv_span ?map_f //. + by move=> xi Lxi; rewrite rpredZ ?memv_span ?map_f. +rewrite cfnormZ cfnorm_map_orthonormal // size_image cardC1 nirrW2. +rewrite -(natrB _ (prime_gt0 pr_p)) Cint_normK // subn1. +by rewrite ler_pemull ?ler0n ?sqr_Cint_ge1. +Qed. + +End Thirteen_17_to_19. + +End Thirteen. + diff --git a/mathcomp/odd_order/PFsection14.v b/mathcomp/odd_order/PFsection14.v new file mode 100644 index 0000000..bd7ae60 --- /dev/null +++ b/mathcomp/odd_order/PFsection14.v @@ -0,0 +1,1257 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime binomial ssralg poly finset. +Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +Require Import gfunctor gproduct center cyclic commutator gseries nilpotent. +Require Import pgroup sylow hall abelian maximal frobenius. +Require Import matrix mxalgebra mxrepresentation mxabelem vector. +Require Import BGsection1 BGsection3 BGsection7. +Require Import BGsection14 BGsection15 BGsection16 BGappendixC. +Require Import ssrnum rat algC cyclotomic algnum. +Require Import classfun character integral_char inertia vcharacter. +Require Import PFsection1 PFsection2 PFsection3 PFsection4. +Require Import PFsection5 PFsection6 PFsection7 PFsection8 PFsection9. +Require Import PFsection10 PFsection11 PFsection12 PFsection13. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 14: Non_existence of G. *) +(* It completes the proof of the Odd Order theorem. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory FinRing.Theory Num.Theory. + +Section Fourteen. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types (p q : nat) (x y z : gT). +Implicit Types H K L N P Q R S T U W : {group gT}. + +Local Notation "#1" := (inord 1) (at level 0). + +(* Supplementary results that apply to both S and T, but that are not *) +(* formally stated as such; T, V, L, tau1L and phi are only used at the end *) +(* of this section, to state and prove FTtype2_support_coherence. *) +Section MoreSTlemmas. + +Local Open Scope ring_scope. +Variables W W1 W2 S T U V L : {group gT}. +Variables (tau1L : {additive 'CF(L) -> 'CF(G)}) (phi : 'CF(L)). + +(* Implicit (dependent) forward assuptions. *) +Hypotheses (defW : W1 \x W2 = W) (xdefW : W2 \x W1 = W) (maxL : L \in 'M). + +Local Notation "` 'S'" := (gval S) (at level 0, only parsing) : group_scope. +Local Notation P := `S`_\F%G. +Local Notation "` 'P'" := `S`_\F (at level 0) : group_scope. +Local Notation PU := S^`(1)%G. +Local Notation "` 'PU'" := `S^`(1)%g (at level 0) : group_scope. +Local Notation "` 'L'" := (gval L) (at level 0, only parsing). +Local Notation H := `L`_\F%G. +Local Notation "` 'H'" := `L`_\F%g (at level 0, format "` 'H'") : group_scope. + +Let p := #|W2|. +Let q := #|W1|. +Let u := #|U|. +Let v := #|V|. +Let h := #|H|. +Let e := #|L : H|. +Let ccG A := class_support A G. + +Let calL := seqIndD H L H 1. +Let betaL := 'Ind[L, H] 1 - phi. +Local Notation tauL := (FT_DadeF maxL). + +(* Explicit (non-dependent) forward assumptions. *) +Hypotheses (StypeP : of_typeP S U defW) (TtypeP : of_typeP T V xdefW). +Hypothesis (cohL : coherent_with calL L^# tauL tau1L) (Lphi : phi \in calL). + +(* The remaining assumptions can be generated as backchaining gools. *) +Hypotheses (maxS : S \in 'M) (maxT : T \in 'M). + +Let pddS := FT_prDade_hypF maxS StypeP. +Let pddT := FT_prDade_hypF maxT TtypeP. +Let ctiWG : cyclicTI_hypothesis G defW := pddS. +Let sigma := cyclicTIiso ctiWG. +Let w_ i j := cyclicTIirr defW i j. + +(* An inequality used in the proof of (14.11.4), at the bottom of page 90, to *) +(* show that 1/uq and 1/vp are less that 1/2q^2 (so Wn is either W1 or W2). *) +Lemma FTtypeP_complV_ltr (Wn : {group gT}) : + (#|Wn| <= q)%N -> (u * q)%:R^-1 < (2 * #|Wn| ^ 2)%:R^-1 :> algC. +Proof. +move=> leWn_q; rewrite !natrM ltf_pinv ?rpredM ?qualifE ?gt0CG ?ltr0n //. +rewrite -!natrM ltr_nat (@leq_ltn_trans (2 * q ^ 2)) ?mulnA ?leq_mul // mul2n. +have: [Frobenius U <*> W1 = U ><| W1] by have [[]] := FTtypeP_facts maxS StypeP. +by move/ltn_odd_Frobenius_ker/implyP; rewrite mFT_odd ltn_pmul2r ?cardG_gt0. +Qed. + +(* This formalizes the loose symmetry used in (14.11.3) to show that #[g] is *) +(* coprime to pq. *) +Lemma coprime_typeP_Galois_core g : + typeP_Galois StypeP -> g \notin ccG W^# -> g \notin ccG P^# -> coprime #[g] p. +Proof. +move=> galS W'g; apply: contraR => p_g. +have ntg: g != 1%g by apply: contraNneq p_g => ->; rewrite order1 coprime1n. +have [pr_q pr_p]: prime q /\ prime p := FTtypeP_primes maxS StypeP. +have [[_ hallW1 _ defS] [_ _ _ defPU] _ [_ _ sW2P _ regPUW1] _] := StypeP. +have coPUq: coprime #|PU| q by rewrite (coprime_sdprod_Hall_r defS). +have [[_ _ nPUW1 _] [_ _ nPU _]] := (sdprodP defS, sdprodP defPU). +have ntP: P :!=: 1%g := mmax_Fcore_neq1 maxS. +have frobPU: [Frobenius PU = P ><| U]. + have notS5 := FTtype5_exclusion maxS. + have inN1 x: x \in 'N(1) by rewrite norm1 inE. + have [_ ntU _ _] := compl_of_typeII_IV maxS StypeP notS5. + have [] := typeP_Galois_P maxS notS5 galS; rewrite Ptype_factor_prime //. + rewrite (group_inj (Ptype_Fcore_kernel_trivial _ _)) // => F [fP [fU _]]. + rewrite Ptype_Fcompl_kernel_trivial //. + case=> /trivgP injfU fJ [_ /isomP[injfP _] _] _. + apply/Frobenius_semiregularP=> // y /setD1P[nty Uy]. + apply/trivgP/subsetP=> x /setIP[Px /cent1P-cxy]; apply: contraR nty. + rewrite -(morph_injm_eq1 injfU) // -val_eqE -(cosetpre1 1) !(inN1, inE) /=. + rewrite -(morph_injm_eq1 injfP) ?mem_quotient //= => /mulfI/inj_eq <-. + rewrite mulr1 -[_ * _]fJ ?mem_quotient //= qactE ?dom_qactJ //=. + by rewrite conjgE cxy mulKg. +have pP: p.-group P by have [_ [/andP[]]] := FTtypeP_facts _ StypeP. +have{p_g}[y [a P1a cagy]]: exists y, exists2 a, a \in P^# & g ^ y \in 'C[a]. + have sylP: p.-Sylow(G) P. + have [/Hall_pi/= hallP _ _] := FTcore_facts maxS; apply: etrans hallP. + have [_ _ [n ->]] := pgroup_pdiv pP (mmax_Fcore_neq1 maxS). + by apply/eq_pHall => r1; rewrite pi_of_exp ?pi_of_prime. + have [y _ Pa] := Sylow_Jsub sylP (subsetT _) (p_elt_constt p g). + pose a := g.`_p ^ y; have{Pa} Pa: a \in P by rewrite -cycle_subG cycleJ. + exists y, a; last by rewrite cent1C /a conjXg groupX ?cent1id. + rewrite !inE conjg_eq1 (contraNneq _ p_g) // => /constt1P/p'nat_coprime-> //. + exact: pnat_id. +have /(mem_sdprod defS)[x [w [PUx W1w Dgy _]]]: g ^ y \in S. + have A0a: a \in 'A0(S) := subsetP (Fcore_sub_FTsupp0 maxS) a P1a. + have [_ _ _ _ [tiA0 _]] := FTtypeP_facts _ StypeP. + by rewrite (subsetP (cent1_normedTI tiA0 A0a)) // 2!inE. +suffices w_eq1: w = 1%g. + have sCaP: 'C_PU[a] \subset P := Frobenius_cent1_ker frobPU P1a. + rewrite -[g](conjgK y) mem_imset2 ?inE //= conjg_eq1 ntg /=. + by rewrite (subsetP sCaP) // inE cagy Dgy w_eq1 mulg1 PUx. +apply: contraNeq W'g => ntw; have nPUw := subsetP nPUW1 w W1w. +have{x PUx Dgy} /imset2P[x z W2w_x _ Dgy]: g ^ y \in class_support (W2 :* w) PU. + rewrite -(regPUW1 w) ?inE ?ntw // class_supportEr -cover_imset. + have coPUw := coprime_dvdr (order_dvdG W1w) coPUq. + have [/cover_partition-> _] := partition_cent_rcoset nPUw coPUw. + by rewrite Dgy mem_rcoset mulgK. +rewrite -[g](conjgK (y * z^-1)%g) mem_imset2 ?inE //= conjg_eq1 ntg /= conjgM. +by rewrite Dgy conjgK -(dprodWC defW) -[x](mulgKV w) mem_mulg -?mem_rcoset. +Qed. + +Hypothesis Stype2 : FTtype S == 2. + +(* This is used to bound #|ccG P^#| and #|ccG Q^#| in the proof of (14.11.4). *) +Lemma FTtype2_cc_core_ler : #|G|%:R^-1 * #|ccG P^#|%:R <= (u * q)%:R^-1 :> algC. +Proof. +have ->: (u * q)%:R^-1 = #|S|%:R^-1 * #|P|%:R :> algC. + have [[_ _ _ /sdprod_card <-] [_ _ _ /sdprod_card <-] _ _ _] := StypeP. + by rewrite mulrC -mulnA [in RHS]natrM invfM mulVKf ?neq0CG. +have [_ _] := FTtypeII_ker_TI maxS Stype2; rewrite FTsupp1_type2 // => tiP1. +rewrite {tiP1}(card_support_normedTI tiP1) natrM natf_indexg ?subsetT //. +rewrite mulrCA mulKf ?neq0CG // mulrC ler_pmul2l ?invr_gt0 ?gt0CG // leC_nat. +by rewrite cardsDS ?sub1G ?leq_subr. +Qed. + +Hypotheses (maxNU_L : L \in 'M('N(U))) (phi1 : phi 1%g = e%:R). + +(* This is Peterfalvi (14.11.2), stated for S and L rather than T and M; it *) +(* is loosely used in this form at the very end of the proof of (14.16). *) +Lemma FTtype2_support_coherence : + (u.-1 %/ q < h.-1 %/ e)%N -> (v.-1 %/ p < h.-1 %/ e)%N -> + [/\ e = (p * q)%N + & exists nb, exists2 chi, chi = tau1L phi \/ chi = - tau1L phi^*%CF + & tauL betaL = \sum_ij (-1)^+ nb ij *: sigma 'chi_ij - chi]. +Proof. +move=> ub_u ub_v; have nsHL : H <| L := gFnormal _ _. +have pairST := of_typeP_pair maxS StypeP maxT TtypeP. +have [//|frobL sUH defL] := FTtypeII_support_facts maxS StypeP _ pairST maxNU_L. +have Ltype1 := FT_Frobenius_type1 maxL frobL. +have irr_phi: phi \in irr L by apply: FTtype1_Ind_irr Lphi. +have betaL_P := FTtypeI_bridge_facts _ _ Ltype1 cohL Lphi phi1. +have e_dv_h1: e %| h.-1 by apply: Frobenius_ker_dvd_ker1. +pose a i j := '[tauL betaL, sigma (w_ i j)]. +have a0j j: j != 0 -> (a 0 j == 1 %[mod 2])%C. + rewrite /a => nz_j; case/betaL_P: StypeP => _ _ -> //. + by case=> [[_ /idPn[]] | [//]]; rewrite -natf_div // leC_nat -ltnNge. +have ai0 i: i != 0 -> (a i 0 == 1 %[mod 2])%C. + rewrite /a (cycTIisoC _ pddT) => nz_i; case/betaL_P: TtypeP => _ _ -> //. + by case=> [[_ /idPn[]] | [//]]; rewrite -natf_div // leC_nat -ltnNge. +have HbetaL: betaL \in 'CF(L, H^#) by apply: cfInd1_sub_lin_on Lphi phi1. +have betaL_W_0: {in cyclicTIset defW, tauL betaL =1 \0}. + move=> z; case/betaL_P: StypeP => tiAM_W _ _ _. + rewrite !inE -(setCK W) inE => /andP[_]; apply: cfun_onP z. + apply: cfun_onS (Dade_cfunS _ _); rewrite FT_DadeF_supportE -disjoints_subset. + rewrite -FTsupp_Frobenius // -setI_eq0 -subset0 -tiAM_W setIS //. + by rewrite setUC subsetU ?sub_class_support. +have calL_gt1: (1 < size calL)%N. + by apply: seqInd_nontrivial Lphi; rewrite ?mFT_odd. +have [] := Dade_Ind1_sub_lin cohL calL_gt1 irr_phi Lphi phi1; rewrite -/betaL. +rewrite -/calL odd_Frobenius_index_ler ?mFT_odd //= -/e -/h. +case=> _ a00 ZbetaL [Gamma [o_tau1_Ga o_1_Ga [aa Zaa Dbeta] []// _ ubGa _]]. +have{a00} a00: a 0 0 = 1 by rewrite /a /w_ cycTIirr00 cycTIiso1. +have{a0j ai0} a_odd i j: (a i j == 1 %[mod 2])%C. + have [[-> | /ai0 ai01] [-> | /a0j a0j1] //] := (eqVneq i 0, eqVneq j 0). + by rewrite a00 (eqCmod_nat 2 1 1). + by rewrite -(eqCmodDr _ 1) -{1}a00 cycTIiso_cfdot_exchange // eqCmodD. +have [_ o_tauLeta _ _] := FTtypeI_bridge_facts _ StypeP Ltype1 cohL Lphi phi1. +pose etaW := map sigma (irr W). +have o1eta: orthonormal etaW := cycTIiso_orthonormal _. +have [X etaX [Y [defGa oXY oYeta]]] := orthogonal_split etaW (Gamma + 1). +have lbY: 0 <= '[Y] ?= iff (Y == 0). + by split; rewrite ?cfnorm_ge0 // eq_sym cfnorm_eq0. +have [b Db defX] := orthonormal_span o1eta etaX. +do [rewrite addrC !addrA addrAC -addrA; set Z := _ - _] in Dbeta. +have oZeta: orthogonal Z etaW. + apply/orthoPl=> xi /memv_span; apply: {xi}(span_orthogonal o_tauLeta). + rewrite rpredB ?rpredZ ?big_seq ?rpred_sum ?memv_span ?map_f // => xi Lxi. + by rewrite rpredZ ?memv_span ?map_f. +have lb_b ij (b_ij := b (sigma 'chi_ij)): + 1 <= `|b_ij| ^+ 2 ?= iff [exists n : bool, b_ij == (-1) ^+ n]. +- have /codomP[[i j] Dij] := dprod_Iirr_onto defW ij. + have{b_ij} ->: b_ij = a i j. + rewrite /a /w_ -Dij Dbeta defGa 2!cfdotDl. + have ->: '[X, sigma 'chi_ij] = b_ij by rewrite /b_ij Db. + by rewrite (orthoPl oYeta) ?(orthoPl oZeta) ?map_f ?mem_irr // !addr0. + have Zaij: a i j \in Cint by rewrite Cint_cfdot_vchar ?cycTIiso_vchar. + rewrite Cint_normK //; split. + rewrite sqr_Cint_ge1 //; apply: contraTneq (a_odd i j) => ->. + by rewrite (eqCmod_nat 2 0 1). + apply/eqP/exists_eqP=> [a2_1|[n ->]]; last by rewrite sqrr_sign. + rewrite (CintEsign Zaij) normC_def conj_Cint // -expr2 -a2_1 sqrtC1 mulr1. + by exists (a i j < 0). +have ub_e: e%:R <= #|Iirr W|%:R ?= iff (e == p * q)%N :> algC. + rewrite lerif_nat card_Iirr_cyclic //; last by have [] := ctiWG. + rewrite -(dprod_card xdefW); apply: leqif_eq. + case: defL => [|[y Qy]] defL; rewrite /e -(index_sdprod defL). + by rewrite leq_pmull ?cardG_gt0. + suffices /normP <-: y \in 'N(W1). + by rewrite -conjYg !cardJg (dprodWY defW) -(dprod_card xdefW). + have cQQ: abelian T`_\F by have [_ [/and3P[]]] := FTtypeP_facts maxT TtypeP. + have sW1Q: W1 \subset T`_\F by have [_ _ _ []] := TtypeP. + by rewrite (subsetP _ y Qy) // sub_abelian_norm. +have /(_ predT) := lerif_add (lerif_sum (in1W lb_b)) lbY. +rewrite sumr_const addr0 => /(lerif_trans ub_e)/ger_lerif/esym. +have ->: \sum_i `|b (sigma 'chi_i)| ^+ 2 = '[X]. + rewrite defX cfnorm_sum_orthonormal // big_map (big_nth 0) big_mkord. + by rewrite size_tuple; apply: eq_bigr => ij _; rewrite -tnth_nth. +rewrite -cfnormDd // -defGa cfnormDd // cfnorm1 -ler_subr_addr ubGa. +case/and3P=> /eqP-De /'forall_exists_eqP/fin_all_exists[/= n Dn] /eqP-Y0. +pose chi := X - tauL betaL; split=> //; exists n, chi; last first. + apply: canRL (addrK _) _; rewrite addrC subrK defX big_map (big_nth 0). + by rewrite big_mkord size_tuple; apply: eq_bigr => ij; rewrite -tnth_nth Dn. +have Z1chi: chi \in dirr G. + rewrite dirrE rpredB //=; last first. + rewrite defX big_map (big_nth 0) big_mkord size_tuple rpred_sum //= => ij. + have [_ Zsigma] := cycTI_Zisometry ctiWG. + by rewrite -tnth_nth Dn rpredZsign ?Zsigma ?irr_vchar. + apply/eqP/(addIr '[X]); rewrite -cfnormBd; last first. + rewrite /chi Dbeta defGa Y0 addr0 opprD addNKr cfdotNl. + by rewrite (span_orthogonal oZeta) ?oppr0 // memv_span ?mem_head. + rewrite addrAC subrr add0r cfnormN Dade_isometry // cfnormBd; last first. + by rewrite cfdotC (seqInd_ortho_Ind1 _ _ Lphi) ?conjC0. + rewrite cfnorm_Ind_cfun1 // -/e irrWnorm // addrC; congr (1 + _). + rewrite defX cfnorm_sum_orthonormal // big_map big_tuple. + rewrite De (dprod_card xdefW) -card_Iirr_cyclic //; last by have[]:= ctiWG. + by rewrite -sumr_const; apply: eq_bigr => ij _; rewrite Dn normr_sign expr1n. +have [[Itau1 Ztau1] Dtau1] := cohL. +suffices /cfdot_add_dirr_eq1: '[tau1L phi - tau1L phi^*%CF, chi] = 1. + rewrite -(cfConjC_Dade_coherent cohL) ?mFT_odd // rpredN dirr_aut. + by apply; rewrite // dirrE Ztau1 ?Itau1 ?mem_zchar ?irrWnorm /=. +rewrite cfdotBr (span_orthogonal o_tauLeta) ?add0r //; last first. + by rewrite rpredB ?memv_span ?map_f ?cfAut_seqInd. +have Zdphi := seqInd_sub_aut_zchar nsHL conjC Lphi. +rewrite -raddfB Dtau1 ?zcharD1_seqInd // Dade_isometry ?(zchar_on Zdphi) //. +rewrite cfdotBr !cfdotBl cfdot_conjCl cfAutInd rmorph1 irrWnorm //. +rewrite (seqInd_ortho_Ind1 _ _ Lphi) // conjC0 subrr add0r opprK. +by rewrite cfdot_conjCl (seqInd_conjC_ortho _ _ _ Lphi) ?mFT_odd ?conjC0 ?subr0. +Qed. + +End MoreSTlemmas. + +Section NonconjType1. +(* Properties of non-conjugate type I groups, used symmetrically for L and M *) +(* in the proofs of (14.14) and (14.16). *) + +Local Open Scope ring_scope. +Variables (M L : {group gT}) (phi : 'CF(L)) (psi : 'CF(M)). +Variable (tau1L : {additive 'CF(L) -> 'CF(G)}). +Variable (tau1M : {additive 'CF(M) -> 'CF(G)}). +Hypotheses (maxL : L \in 'M) (maxM : M \in 'M). +Let ddL := FT_DadeF_hyp maxL. +Let ddM := FT_DadeF_hyp maxM. +Let tauL := Dade ddL. +Let tauM := Dade ddM. +Let H := L`_\F%G. +Let K := M`_\F%G. +Let calL := seqIndD H L H 1. +Let calM := seqIndD K M K 1. +Let u : algC := #|L : H|%:R. +Let v : algC := #|M : K|%:R. +Let betaL := 'Ind[L, H] 1 - phi. +Let a := '[tauL betaL, tau1M psi]. + +Hypothesis (cohL : coherent_with calL L^# tauL tau1L). +Hypothesis (cohM : coherent_with calM M^# tauM tau1M). +Hypotheses (Lphi : phi \in calL) (Mpsi : psi \in calM). +Hypotheses (phi1 : phi 1%g = u) (psi1 : psi 1%g = v). +Hypotheses (Ltype1 : FTtype L == 1%N) (Mtype1 : FTtype M == 1%N). +Hypothesis not_MG_L : gval L \notin M :^: G. + +Let irrL := FTtype1_Ind_irr maxL Ltype1. +Let irrM := FTtype1_Ind_irr maxM Mtype1. + +Lemma disjoint_Dade_FTtype1 : [disjoint Dade_support ddM & Dade_support ddL]. +Proof. +by rewrite !FT_DadeF_supportE -!FTsupp1_type1 ?FT_Dade1_support_disjoint. +Qed. +Let TItauML := disjoint_Dade_FTtype1. + +Lemma coherent_FTtype1_ortho : orthogonal (map tau1M calM) (map tau1L calL). +Proof. +apply/orthogonalP=> _ _ /mapP[xiM Mxi ->] /mapP[xiL Lxi ->]. +have [irrLxi irrMxi] := (irrL Lxi, irrM Mxi). +exact: (disjoint_coherent_ortho (mFT_odd _) _ cohM cohL). +Qed. +Let oML := coherent_FTtype1_ortho. + +(* This is the inequality used in both branches of (14.14). *) +Lemma coherent_FTtype1_core_ltr : a != 0 -> #|K|.-1%:R / v <= u - 1. +Proof. +have [nsHL nsKM]: H <| L /\ K <| M by rewrite !gFnormal. +have [irr_phi irr_psi] := (irrL Lphi, irrM Mpsi). +have frobL: [Frobenius L with kernel H] := FTtype1_Frobenius maxL Ltype1. +have [[Itau1 Ztau1] Dtau1] := cohM. +have o1M: orthonormal (map tau1M calM). + apply: map_orthonormal Itau1 _. + exact: sub_orthonormal (undup_uniq _) (irr_orthonormal M). +have Lgt1: (1 < size calL)%N by apply: seqInd_nontrivial (mFT_odd _ ) _ _ Lphi. +have [[_ _]] := Dade_Ind1_sub_lin cohL Lgt1 irr_phi Lphi phi1. +rewrite -/tauL -/betaL -/calL => ZbetaL [Gamma [_ _ [b _ Dbeta]]]. +rewrite odd_Frobenius_index_ler ?mFT_odd // -/u => -[]// [_ ub_Ga] _ nz_a. +have Za: a \in Cint by rewrite Cint_cfdot_vchar // ?Ztau1 ?mem_zchar. +have [X M_X [Del [defGa oXD oDM]]] := orthogonal_split (map tau1M calM) Gamma. +apply: ler_trans ub_Ga; rewrite defGa cfnormDd // ler_paddr ?cfnorm_ge0 //. +suffices ->: '[X] = (a / v) ^+ 2 * (\sum_(xi <- calM) xi 1%g ^+ 2 / '[xi]). + rewrite sum_seqIndC1_square // -(natrB _ (cardG_gt0 K)) subn1. + rewrite exprMn !mulrA divfK ?neq0CiG // mulrAC -mulrA. + by rewrite ler_pemull ?sqr_Cint_ge1 // divr_ge0 ?ler0n. +have [_ -> defX] := orthonormal_span o1M M_X. +have Mgt1: (1 < size calM)%N by apply: seqInd_nontrivial (mFT_odd _ ) _ _ Mpsi. +have [[oM1 _ _] _ _] := Dade_Ind1_sub_lin cohM Mgt1 irr_psi Mpsi psi1. +rewrite exprMn -(Cint_normK Za) -[v]normr_nat -normfV -/v mulr_sumr. +rewrite defX cfnorm_sum_orthonormal // big_map; apply: eq_big_seq => xi Mxi. +have Zxi1 := Cint_seqInd1 Mxi; rewrite -(Cint_normK Zxi1) -(conj_Cint Zxi1). +rewrite irrWnorm ?irrM // divr1 -!exprMn -!normrM; congr (`|_| ^+ 2). +rewrite -mulrA mulrC -mulrA; apply: canRL (mulKf (neq0CiG _ _)) _. +rewrite -(canLR (addrK _) defGa) cfdotBl (orthoPl oDM) ?map_f // subr0. +rewrite -(canLR (addKr _) Dbeta) cfdotDl cfdotNl cfdotC cfdotDr cfdotBr. +rewrite (orthoPr oM1) ?map_f // (orthogonalP oML) ?map_f // subrr add0r. +rewrite cfdotZr cfdot_sumr big1_seq ?mulr0 ?oppr0 => [|nu Mnu]; last first. + by rewrite cfdotZr (orthogonalP oML) ?map_f ?mulr0. +apply/eqP; rewrite conjC0 oppr0 add0r -subr_eq0 -conjC_nat -!cfdotZr. +rewrite -raddfZnat -raddfZ_Cint // -cfdotBr -raddfB -/v -psi1. +rewrite Dtau1 ?zcharD1_seqInd ?sub_seqInd_zchar //. +rewrite (cfdotElr (Dade_cfunS _ _) (Dade_cfunS _ _)) setIC. +by have:= TItauML; rewrite -setI_eq0 => /eqP->; rewrite big_set0 mulr0. +Qed. + +End NonconjType1. + +(* This is the context associated with Hypothesis (13.1). *) +Variables S T U V W W1 W2 : {group gT}. +Hypotheses (defW : W1 \x W2 = W) (xdefW : W2 \x W1 = W). +Hypotheses (pairST : typeP_pair S T defW) (maxS : S \in 'M) (maxT : T \in 'M). +Hypotheses (StypeP : of_typeP S U defW) (TtypeP : of_typeP T V xdefW). + +Local Notation "` 'W1'" := (gval W1) (at level 0, only parsing) : group_scope. +Local Notation "` 'W2'" := (gval W2) (at level 0, only parsing) : group_scope. +Local Notation "` 'W'" := (gval W) (at level 0, only parsing) : group_scope. +Local Notation What := (cyclicTIset defW). + +Local Notation "` 'S'" := (gval S) (at level 0, only parsing) : group_scope. +Local Notation P := `S`_\F%G. +Local Notation "` 'P'" := `S`_\F (at level 0) : group_scope. +Local Notation PU := S^`(1)%G. +Local Notation "` 'PU'" := `S^`(1) (at level 0) : group_scope. +Local Notation "` 'U'" := (gval U) (at level 0, only parsing) : group_scope. + +Local Notation "` 'T'" := (gval T) (at level 0, only parsing) : group_scope. +Local Notation Q := `T`_\F%G. +Local Notation "` 'Q'" := `T`_\F (at level 0) : group_scope. +Local Notation QV := T^`(1)%G. +Local Notation "` 'QV'" := `T^`(1) (at level 0) : group_scope. +Local Notation "` 'V'" := (gval V) (at level 0, only parsing) : group_scope. + +Let defS : PU ><| W1 = S. Proof. by have [[]] := StypeP. Qed. +Let defPU : P ><| U = PU. Proof. by have [_ []] := StypeP. Qed. + +Let defT : QV ><| W2 = T. Proof. by have [[]] := TtypeP. Qed. +Let defQV : Q ><| V = QV. Proof. by have [_ []] := TtypeP. Qed. + +Let notStype1 : FTtype S != 1%N. Proof. exact: FTtypeP_neq1 StypeP. Qed. +Let notStype5 : FTtype S != 5%N. Proof. exact: FTtype5_exclusion maxS. Qed. + +Let pddS := FT_prDade_hypF maxS StypeP. +Let ptiWS : primeTI_hypothesis S PU defW := FT_primeTI_hyp StypeP. +Let ctiWG : cyclicTI_hypothesis G defW := pddS. + +Let pddT := FT_prDade_hypF maxT TtypeP. +Let ptiWT : primeTI_hypothesis T QV xdefW := FT_primeTI_hyp TtypeP. + +Let ntW1 : W1 :!=: 1. Proof. by have [[]] := StypeP. Qed. +Let ntW2 : W2 :!=: 1. Proof. by have [_ _ _ []] := StypeP. Qed. +Let cycW1 : cyclic W1. Proof. by have [[]] := StypeP. Qed. +Let cycW2 : cyclic W2. Proof. by have [_ _ _ []] := StypeP. Qed. + +Let p := #|W2|. +Let q := #|W1|. +Let u := #|U|. +Let v := #|V|. +Let nU := (p ^ q).-1 %/ p.-1. +Let nV := (q ^ p).-1 %/ q.-1. + +Let pr_p : prime p. Proof. by have [] := FTtypeP_primes maxS StypeP. Qed. +Let pr_q : prime q. Proof. by have [] := FTtypeP_primes maxS StypeP. Qed. + +Local Open Scope ring_scope. + +Let qgt2 : (q > 2)%N. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. +Let pgt2 : (p > 2)%N. Proof. by rewrite odd_gt2 ?mFT_odd ?cardG_gt1. Qed. + +Let coPUq : coprime #|PU| q. +Proof. by rewrite (coprime_sdprod_Hall_r defS); have [[]] := StypeP. Qed. + +Let nirrW1 : #|Iirr W1| = q. Proof. by rewrite card_Iirr_cyclic. Qed. +Let nirrW2 : #|Iirr W2| = p. Proof. by rewrite card_Iirr_cyclic. Qed. +Let NirrW1 : Nirr W1 = q. Proof. by rewrite -nirrW1 card_ord. Qed. +Let NirrW2 : Nirr W2 = p. Proof. by rewrite -nirrW2 card_ord. Qed. + +Let sigma := (cyclicTIiso ctiWG). +Let w_ i j := (cyclicTIirr defW i j). +Local Notation eta_ i j := (sigma (w_ i j)). + +Local Notation Imu2 := (primeTI_Iirr ptiWS). +Let mu2_ i j := primeTIirr ptiWS i j. +Let mu_ := primeTIred ptiWS. +Local Notation chi_ j := (primeTIres ptiWS j). + +Local Notation Inu2 := (primeTI_Iirr ptiWT). +Let nu2_ i j := primeTIirr ptiWT j i. +Let nu_ := primeTIred ptiWT. + +Local Notation tauS := (FT_Dade0 maxS). +Local Notation tauT := (FT_Dade0 maxT). + +Let calS0 := seqIndD PU S S`_\s 1. +Let rmR_S := FTtypeP_coh_base maxS StypeP. +Let scohS0 : subcoherent calS0 tauS rmR_S. +Proof. exact: FTtypeP_subcoherent StypeP. Qed. + +Let calS := seqIndD PU S P 1. +Let sSS0 : cfConjC_subset calS calS0. +Proof. exact/seqInd_conjC_subset1/Fcore_sub_FTcore. Qed. + +Let calT := seqIndD QV T Q 1. + +(* This is Hypothesis (14.1). *) +Hypothesis ltqp: (q < p)%N. + +(* This corresponds to Peterfalvi, Theorem (14.2). *) +(* As we import the conclusion of BGappendixC, which covers Appendix C of the *) +(* Bender and Glauberman text, we can state this theorem negatively. This *) +(* will avoid having to repeat its statement thoughout the proof : we will *) +(* simply end each nested set of assumptions (corresponding to (14.3) and *) +(* (14.10)) with a contradiction. *) +Theorem no_full_FT_Galois_structure : + ~ [/\ (*a*) exists Fpq : finFieldImage P W2 U, + [/\ #|P| = (p ^ q)%N, #|U| = nU & coprime nU p.-1] + & (*b*) [/\ q.-abelem Q, W2 \subset 'N(Q) + & exists2 y, y \in Q & W2 :^ y \subset 'N(U)]]. +Proof. +case=> [[Fpq [oP oU coUp1]] [abelQ nQW2 nU_W2Q]]. +have /idPn[] := ltqp; rewrite -leqNgt. +exact: (prime_dim_normed_finField _ _ _ defPU) nU_W2Q. +Qed. + +(* Justification for Hypothesis (14.3). *) +Fact FTtypeP_max_typeII : FTtype S == 2. +Proof. by have [[_ ->]] := FTtypeP_facts maxS StypeP. Qed. +Let Stype2 := FTtypeP_max_typeII. + +(* These correspond to Peterfalvi, Hypothesis (14.3). *) +Variables (L : {group gT}) (tau1L : {additive 'CF(L) -> 'CF(G)}) (phi : 'CF(L)). +Local Notation "` 'L'" := (gval L) (at level 0, only parsing). +Local Notation H := `L`_\F%G. +Local Notation "` 'H'" := `L`_\F%g (at level 0, format "` 'H'") : group_scope. + +Hypothesis maxNU_L : L \in 'M('N(U)). + +(* Consequences of the above. *) +Hypotheses (maxL : L \in 'M) (sNUL : 'N(U) \subset L) (sUH : U \subset H). +Hypotheses (frobL : [Frobenius L with kernel H]) (Ltype1 : FTtype L == 1%N). + +Let calL := seqIndD H L H 1. +Local Notation tauL := (FT_DadeF maxL). +Let nsHL : H <| L. Proof. exact: gFnormal. Qed. +Let irrL : {subset calL <= irr L}. Proof. exact: FTtype1_Ind_irr. Qed. + +Hypothesis cohL : coherent_with calL L^# tauL tau1L. +Hypotheses (Lphi : phi \in calL) (phi1 : phi 1%g = #|L : H|%:R). + +Let betaS := FTtypeP_bridge StypeP #1. +Let betaT := FTtypeP_bridge TtypeP #1. +Let betaL := 'Ind[L, H] 1 - phi. + +(* This is the first assertion of Peterfalvi (14.4). *) +Let galT : typeP_Galois TtypeP. +Proof. +apply: contraLR ltqp => /(FTtypeP_nonGalois_facts maxT)[]. +by rewrite -/p -leqNgt => ->. +Qed. + +(* This is the second assertion of Peterfalvi (14.4). *) +Let oV : v = nV. +Proof. +rewrite /v (card_FTtypeP_Galois_compl maxT galT) -/nV. +by rewrite !modn_small ?gtn_eqF // ltnW. +Qed. + +(* This is Peterfalvi (14.5). *) +Let defL : exists2 y, y \in Q & H ><| (W1 <*> W2 :^ y) = L. +Proof. +have [//|_ _ []// defL] := FTtypeII_support_facts maxS StypeP _ pairST maxNU_L. +have [_ _ /negP[]] := compl_of_typeII maxS StypeP Stype2. +have [_ _ _] := FTtypeI_bridge_facts maxS StypeP Ltype1 cohL Lphi phi1. +case=> [[_ ubH] | [_ /idPn[]]]; last by rewrite -(index_sdprod defL) -ltnNge. +have{ubH} /eqP defH: `H == U. + rewrite eq_sym eqEcard sUH /= -(prednK (cardG_gt0 U)) -add1n -leq_subLR subn1. + have [_ _ _ _ /divnK <-] := FTtypeP_bridge_facts maxS StypeP. + by rewrite -leC_nat natrM -ler_pdivr_mulr ?gt0CG // {1}(index_sdprod defL). +rewrite (subset_trans sNUL) // -(sdprodW defL) -(sdprodW defS) mulSg //. +by rewrite -(sdprodW defPU) defH mulG_subr. +Qed. + +Let indexLH : #|L : H| = (p * q)%N. +Proof. +have [y Qy /index_sdprod <-] := defL; rewrite (dprod_card xdefW). +suffices /normP <-: y \in 'N(W1) by rewrite -conjYg cardJg (dprodWY defW). +have cQQ: abelian Q by have [_ [/and3P[]]] := FTtypeP_facts _ TtypeP. +by apply: (subsetP (sub_abelian_norm cQQ _)) => //; have [_ _ _ []] := TtypeP. +Qed. + +(* This is Peterfalvi (14.6). *) +Let galS : typeP_Galois StypeP. +Proof. +apply/idPn=> gal'S; have [q3 oU] := FTtypeP_nonGalois_facts maxS gal'S. +have [H1 [_ _ _ _]] := typeP_Galois_Pn maxS (FTtype5_exclusion maxS) gal'S. +rewrite def_Ptype_factor_prime // Ptype_Fcompl_kernel_trivial // -/p q3 /=. +set a := #|U : _| => [] [a_gt1 a_dv_p1 _ [U1 isoU1]]. +have{isoU1} isoU: U \isog U1 := isog_trans (quotient1_isog U) isoU1. +have{a_gt1 a_dv_p1} defU1: U1 :=: [set: 'rV_2]. + apply/eqP; rewrite eqEcard subsetT -(card_isog isoU) oU. + rewrite cardsT card_matrix card_ord Zp_cast // leq_sqr -/p. + apply: dvdn_leq; first by rewrite -(subnKC pgt2). + rewrite -divn2 -(@Gauss_dvdl a _ 2) ?divnK //. + by rewrite dvdn2 -subn1 odd_sub ?odd_gt0 ?mFT_odd. + by rewrite coprimen2 (dvdn_odd (dvdn_indexg U _)) ?mFT_odd. +have [r pr_r r_r_U] := rank_witness U. +have [R0 sylR0] := Sylow_exists r U; have [sR0U rR0 _] := and3P sylR0. +have [R sylR sR0R] := Sylow_superset (subset_trans sR0U sUH) rR0. +have [sRH rR _] := and3P sylR. +have cUU: abelian U by have [[]] := FTtypeP_facts maxS StypeP. +have tiA0: normedTI 'A0(S) G S by have [_ _ _ _ []] := FTtypeP_facts _ StypeP. +have [_ sUPU _ nPU _] := sdprod_context defPU. +have coPU := coprimegS (joing_subl U W1) (Ptype_Fcore_coprime StypeP). +have abR0: abelian R0 := abelianS sR0U cUU. +have{a U1 defU1 isoU r_r_U} rR0_2: 'r(R0) = 2. + by rewrite (rank_Sylow sylR0) -r_r_U (isog_rank isoU) defU1 rank_mx_group. +have piUr: r \in \pi(U) by rewrite -p_rank_gt0 -(rank_Sylow sylR0) rR0_2. +have /exists_inP[x /setD1P[ntx R0x] ntCPx]: [exists x in R0^#, 'C_P[x] != 1%g]. + have ncycR0: ~~ cyclic R0 by rewrite abelian_rank1_cyclic ?rR0_2. + have coPR0: coprime #|P| #|R0| := coprimegS sR0U coPU. + rewrite -negb_forall_in; apply: contra (mmax_Fcore_neq1 maxS) => regR0P. + rewrite -subG1 -(coprime_abelian_gen_cent1 abR0 _ (subset_trans sR0U nPU)) //. + by rewrite gen_subG; apply/bigcupsP=> x /(eqfun_inP regR0P)->. +have{x ntx R0x ntCPx} sZR_R0: 'Z(R) \subset R0. + have A0x: x \in 'A0(S). + have [z /setIP[Pz cyz] ntz] := trivgPn _ ntCPx. + apply/setUP; left; apply/bigcupP; exists z. + by rewrite !inE ntz (subsetP (Fcore_sub_FTcore maxS)). + by rewrite (eqP Stype2) 3!inE ntx cent1C (subsetP sUPU) ?(subsetP sR0U). + have sCxS: 'C[x] \subset S by rewrite -['C[x]]setTI (cent1_normedTI tiA0). + suffices <-: 'C_R[x] = R0. + by rewrite -cent_set1 setIS ?centS // sub1set (subsetP sR0R). + have /Hall_pi hallU: Hall PU U by rewrite -(coprime_sdprod_Hall_r defPU). + have /Hall_pi hallPU: Hall S PU by rewrite -(coprime_sdprod_Hall_l defS). + have sylR0_S: r.-Sylow(S) R0. + by apply: subHall_Sylow piUr sylR0; apply: subHall_Hall (piSg sUPU) hallU. + rewrite ['C_R[x]](sub_pHall sylR0_S) ?(pgroupS _ rR) ?subsetIl //. + by rewrite subsetI sR0R sub_cent1 (subsetP abR0). + by rewrite subIset ?sCxS ?orbT. +pose R1 := 'Ohm_1('Z(R))%G; pose m := logn r #|R1|. +have sR10: R1 \subset R0 by rewrite (subset_trans (Ohm_sub 1 _)). +have oR1: #|R1| = (r ^ m)%N by rewrite -card_pgroup ?(pgroupS sR10). +have{sZR_R0 rR0_2} m12: pred2 1%N 2 m. + transitivity (0 < m < 1 + 2)%N; first by rewrite -mem_iota !inE. + rewrite -[m]p_rank_abelian ?center_abelian -?rank_pgroup ?(pgroupS sZR_R0) //. + rewrite rank_gt0 ltnS -rR0_2 rankS // center_nil_eq1 ?(pgroup_nil rR) //. + by rewrite (subG1_contra sR0R) // -rank_gt0 rR0_2. +have [y Qy defLy] := defL; have [_ _ /joing_subP[_ nHW2y] _] := sdprodP defLy. +have chR1H: R1 \char H. + apply: char_trans (char_trans (Ohm_char 1 _) (center_char R)) _. + by rewrite (nilpotent_Hall_pcore (Fcore_nil L) sylR) gFchar. +have nR1W2y: W2 :^ y \subset 'N(R1) := char_norm_trans chR1H nHW2y. +have regR1W2y: semiregular R1 (W2 :^ y). + have /Frobenius_reg_ker regHW12y := set_Frobenius_compl defLy frobL. + exact: semiregularS (char_sub chR1H) (joing_subr _ _) regHW12y. +have /idPn[]: r %| p.-1./2. + have:= piUr; rewrite mem_primes => /and3P[_ _ /=]. + by rewrite oU Euclid_dvdX ?andbT. +rewrite gtnNdvd //; first by rewrite -(subnKC pgt2). +apply: leq_trans (_ : p.-1 <= r)%N. + by rewrite -divn2 ltn_divLR // -{1}[p.-1]muln1 -(subnKC pgt2) ltn_pmul2l. +have: p %| (r ^ m).-1. + by have:= regular_norm_dvd_pred nR1W2y regR1W2y; rewrite cardJg oR1. +rewrite -[p.-1]subn1 leq_subLR predn_exp Euclid_dvdM // => /orP[]/dvdn_leq. + by rewrite -(subnKC (prime_gt1 pr_r)) => /implyP/leq_trans->; rewrite 2?ltnW. +move/implyP; case/pred2P: m12 => ->; rewrite !big_ord_recl big_ord0 ?addn0 //=. +by rewrite -(subnKC pgt2). +Qed. + +(* This is Peterfalvi (14.7). *) +Let not_charUH : ~~ (U \char H). +Proof. +have [y Qy defLy] := defL; have [_ _ /joing_subP[_ nHW2y] _] := sdprodP defLy. +apply/negP=> chUH; have nUW2y := char_norm_trans chUH nHW2y. +case: no_full_FT_Galois_structure; split; last first. + split; [by have [_ []] := FTtypeP_facts _ TtypeP | | by exists y]. + by have /sdprodP[_ _ /joing_subP[]] := Ptype_Fcore_sdprod TtypeP. +have <-: #|U| = nU. + have regUW2y: semiregular U (W2 :^ y). + have /Frobenius_reg_ker regHW12y := set_Frobenius_compl defLy frobL. + exact: semiregularS (char_sub chUH) (joing_subr _ _) regHW12y. + case: ifP (card_FTtypeP_Galois_compl maxS galS) => //. + rewrite -/p -/q -/nU => p_modq_1 oU. + have{p_modq_1 oU} oU: (#|U| * q)%N = nU. + by rewrite oU divnK //; have [|_ ->] := FTtypeP_primes_mod_cases _ StypeP. + have /eqP Umodp: #|U| == 1 %[mod p]. + have:= regular_norm_dvd_pred nUW2y regUW2y. + by rewrite cardJg -/p -subn1 eqn_mod_dvd. + have: nU == 1 %[mod p]. + rewrite /nU predn_exp mulKn; last by rewrite -(subnKC pgt2). + rewrite -(ltn_predK qgt2) big_ord_recl addnC -modnDml -modn_summ modnDml. + by rewrite big1 // => i _; rewrite expnS modnMr. + by rewrite -oU -modnMml Umodp modnMml mul1n !modn_small ?gtn_eqF ?prime_gt1. +have [F []] := typeP_Galois_P maxS (FTtype5_exclusion maxS) galS. +rewrite Ptype_factor_prime ?(group_inj (Ptype_Fcore_kernel_trivial _ _)) //. +rewrite Ptype_Fcompl_kernel_trivial // => psiP [psiU _ [/trivgP inj_psiU psiJ]]. +rewrite /= -injm_subcent ?coset1_injm ?norms1 // -morphim_comp -/p. +rewrite (typeP_cent_core_compl StypeP) => [[_ /isomP[inj_psiP im_psiP] psiW2]]. +rewrite -(card_isog (quotient1_isog U)) => [[_ coUp1 _]]. +suffices FPU: finFieldImage P W2 U. + by exists FPU; have [_ []] := FTtypeP_facts maxS StypeP. +have /domP[sig [Dsig Ksig _ im_sig]]: 'dom (psiP \o coset 1) = P. + by apply: injmK; rewrite ?coset1_injm ?norms1. +have{Ksig} inj_sig: 'injm sig by rewrite Ksig injm_comp ?coset1_injm. +exists F sig; first by apply/isomP; rewrite im_sig morphim_comp. + by rewrite -psiW2 -im_sig injmK // -(typeP_cent_core_compl StypeP) subsetIl. +exists psiU => // z x Pz Ux /=; have inN1 x1: x1 \in 'N(1) by rewrite norm1 inE. +by rewrite !Dsig -psiJ ?mem_morphim //= qactE ?dom_qactJ. +Qed. + +(* This is Peterfalvi (14.8)(a). *) +(* In order to avoid the use of real analysis and logarithms we bound the *) +(* binomial expansion of n.+1 ^ q.+1 directly. *) +Let qp1_gt_pq1 : (q ^ p.+1 > p ^ q.+1)%N. +Proof. +have: (4 < p)%N by rewrite odd_geq ?mFT_odd ?(leq_trans _ ltqp). +elim: p ltqp => // n IHn; rewrite !ltnS => ngeq. +rewrite leq_eqVlt => /predU1P[/esym n4 | ngt4]. + suffices /eqP <-: 3 == q by rewrite n4. + by rewrite eqn_leq qgt2 -ltnS -(odd_ltn 5) ?mFT_odd // -n4. +apply: leq_trans (_ : q * n ^ q.+1 <= _)%N; last first. + rewrite (expnS q) leq_mul //. + by move: ngeq; rewrite leq_eqVlt => /predU1P[-> | /IHn/(_ ngt4)/ltnW]. +apply: leq_trans (_ : (2 * q.+1 + n) * n ^ q <= _)%N; last first. + rewrite expnS mulnA leq_mul // addnC. + move: ngeq; rewrite leq_eqVlt => /predU1P[-> | n_gtq]. + apply: leq_trans (_ : 4 * n <= _)%N; last by rewrite leq_mul // ltnW. + by rewrite mulnSr addnA -mulSn (mulSnr 3) leq_add2l 3?ltnW. + by rewrite -{2}(subnKC qgt2) addSn (mulSn _ n) leq_add2l leq_mul. +rewrite mulnDl -expnS -[n.+1]add1n expnDn big_ord_recr binn subnn !mul1n /=. +rewrite ltn_add2r -(@ltn_pmul2l (2 ^ q)) ?expn_gt0 // !mulnA -expnSr. +apply: leq_ltn_trans (_ : (2 ^ q.+1).-1 * q.+1 * n ^ q < _)%N; last first. + by rewrite -(subnKC ngt4) !ltn_pmul2r ?prednK ?expn_gt0. +rewrite -mulnA predn_exp mul1n big_distrr big_distrl leq_sum // => [[i]] /=. +rewrite ltnS exp1n mul1n => leiq _; rewrite -{1 4}(subnKC leiq) !expnD. +rewrite -mulnA leq_mul // mulnA mulnCA mulnC leq_mul // -bin_sub ?leqW //. +rewrite -(leq_pmul2r (fact_gt0 (q.+1 - i))) -mulnA bin_ffact mulnC subSn //. +rewrite ffactnS /= -!mulnA leq_mul //=; elim: {i leiq}(q - i)%N => //= i IHi. +rewrite ffactnSr expnSr mulnACA expnS factS (mulnACA n) mulnC leq_mul //. +by rewrite leq_mul // (leq_trans (leq_subr _ _)). +Qed. + +(* This is Peterfalvi (14.8)(b). *) +Let v1p_gt_u1q : (v.-1 %/ p > u.-1 %/ q)%N. +Proof. +have ub_u: (u.-1 <= nU - 1)%N. + rewrite -subn1 leq_sub2r //; have [_ _] := FTtypeP_facts maxS StypeP. + by rewrite (FTtypeP_reg_Fcore maxS StypeP) indexg1. +rewrite ltn_divLR ?prime_gt0 // {ub_u}(leq_ltn_trans ub_u) //. +have p_dv_v1: p %| v.-1 by have [] := FTtypeP_bridge_facts maxT TtypeP. +rewrite divn_mulAC // ltn_divRL ?dvdn_mulr // oV -subn1. +rewrite -(@ltn_pmul2l q.-1) ?(mulnCA q.-1); last by rewrite -(subnKC qgt2). +rewrite !mulnA -(@ltn_pmul2l p.-1); last by rewrite -(subnKC pgt2). +rewrite -mulnA mulnCA mulnA !(mulnBl _ _ _.-1) !divnK ?dvdn_pred_predX //. +rewrite !mul1n mulnCA -!subn1 ltn_mul ?ltn_sub2r ?prime_gt1 //. +rewrite -!subnDA !subnKC ?prime_gt0 // !mulnBl -!expnSr !mulnn. +by rewrite -subSn ?leq_exp2l ?leqW ?prime_gt1 ?leq_sub ?leq_exp2r // ltnW. +Qed. + +Let calT0 := seqIndD QV T T`_\s 1. +Let rmR_T := FTtypeP_coh_base maxT TtypeP. +Let scohT0 : subcoherent calT0 tauT rmR_T. +Proof. exact: FTtypeP_subcoherent. Qed. + +Let sTT0 : cfConjC_subset calS calS0. +Proof. exact/seqInd_conjC_subset1/Fcore_sub_FTcore. Qed. + +(* This is Peterfalvi (14.9). *) +Lemma FTtypeP_min_typeII : FTtype T == 2. +Proof. +apply: contraLR v1p_gt_u1q => notTtype2; rewrite -leqNgt -leC_nat. +have [o_betaT0_eta _ [Ttype3 _]] := FTtype34_structure maxT TtypeP notTtype2. +have Ttype_gt2: (2 < FTtype T)%N by rewrite (eqP Ttype3). +have [[_ _ frobVW2 cVV] _ _ _ _] := FTtypeP_facts _ TtypeP. +pose calT1 := seqIndD QV T QV Q; have sT10: cfConjC_subset calT1 calT0. + by apply/seqInd_conjC_subset1; rewrite /= FTcore_eq_der1. +rewrite (FTtypeP_reg_Fcore maxT TtypeP) (group_inj (joingG1 _)) in o_betaT0_eta. +do [rewrite -/calT1; set eta_0 := \sum_j _] in o_betaT0_eta. +have scohT1: subcoherent calT1 tauT rmR_T := subset_subcoherent scohT0 sT10. +have [nsQQV sVQV _ _ _] := sdprod_context defQV. +have nsQVT: QV <| T := der_normal 1 T. +have calT1_1p zeta: zeta \in calT1 -> zeta 1%g = p%:R. + case/seqIndP=> s /setDP[kerQs _] -> /=; rewrite inE in kerQs. + rewrite cfInd1 ?gFsub // -(index_sdprod defT) lin_char1 ?mulr1 //. + rewrite lin_irr_der1 (subset_trans _ kerQs) // der1_min ?normal_norm //. + by rewrite -(isog_abelian (sdprod_isog defQV)). +have [tau1T cohT1]: coherent calT1 T^# tauT. + apply/(uniform_degree_coherence scohT1)/(@all_pred1_constant _ p%:R). + by apply/allP=> _ /mapP[zeta T1zeta ->]; rewrite /= calT1_1p. +have irrT1: {subset calT1 <= irr T}. + move=> _ /seqIndP[s /setDP[kerQs nz_s] ->]; rewrite inE in kerQs. + rewrite inE subGcfker in nz_s; rewrite -(quo_IirrK nsQQV kerQs) mod_IirrE //. + rewrite cfIndMod ?normal_sub ?cfMod_irr ?gFnormal //. + rewrite irr_induced_Frobenius_ker ?quo_Iirr_eq0 //=. + have /sdprod_isom[nQ_VW1 /isomP[injQ <-]] := Ptype_Fcore_sdprod TtypeP. + have ->: (QV / Q)%g = (V / Q)%g by rewrite -(sdprodW defQV) quotientMidl. + have ->: (V / Q)%g = restrm nQ_VW1 (coset Q) @* V. + by rewrite morphim_restrm (setIidPr _) // joing_subl. + by rewrite injm_Frobenius_ker // (FrobeniusWker frobVW2). +have [[A0betaS PVbetaS] _ [_]] := FTtypeP_bridge_facts maxS StypeP. +rewrite -/q -/u; set Gamma := FTtypeP_bridge_gap _ _ => oGa1 R_Ga lb_Ga _. +have oT1eta: orthogonal (map tau1T calT1) (map sigma (irr W)). + apply/orthogonalP=> _ _ /mapP[zeta T1zeta ->] /mapP[omega Womega ->]. + have{omega Womega} [i [j ->]] := cycTIirrP defW Womega. + by rewrite (cycTIisoC _ pddT) (coherent_ortho_cycTIiso _ sT10 cohT1) ?irrT1. +have [[Itau1T Ztau1T] Dtau1T] := cohT1. +have nzT1_Ga zeta: zeta \in calT1 -> `|'[Gamma, tau1T zeta]| ^+ 2 >= 1. + have Z_Ga: Gamma \in 'Z[irr G]. + rewrite rpredD ?cycTIiso_vchar // rpredB ?rpred1 ?Dade_vchar // zchar_split. + by rewrite A0betaS ?Iirr1_neq0 // rpredB ?cfInd_vchar ?rpred1 ?irr_vchar. + move=> T1zeta; rewrite expr_ge1 ?normr_ge0 // norm_Cint_ge1 //. + by rewrite Cint_cfdot_vchar ?Ztau1T ?(mem_zchar T1zeta). + suffices: ('[Gamma, tau1T zeta] == 1 %[mod 2])%C. + by apply: contraTneq => ->; rewrite (eqCmod_nat 2 0 1). + pose betaT0 := nu_ 0 - zeta. + have{o_betaT0_eta} o_eta0_betaT0 j: '[eta_ 0 j, tauT betaT0] = (j == 0)%:R. + transitivity '[eta_ 0 j, eta_0]; rewrite (cycTIisoC _ pddT). + apply/eqP; rewrite -subr_eq0 -cfdotBr cfdotC. + by rewrite (orthoPl (o_betaT0_eta _ _)) ?conjC0 // map_f ?mem_irr. + rewrite cfdot_sumr (bigD1 0) //= cfdot_cycTIiso andbT big1 ?addr0 //. + by move=> i /negPf nz_i; rewrite cfdot_cycTIiso andbC eq_sym nz_i. + have QVbetaT0: betaT0 \in 'CF(T, QV^#). + rewrite cfun_onD1 rpredB ?(seqInd_on _ T1zeta) //=; last first. + by rewrite /nu_ -cfInd_prTIres cfInd_normal. + by rewrite !cfunE prTIred_1 prTIirr0_1 mulr1 calT1_1p ?subrr. + have A0betaT0: betaT0 \in 'CF(T, 'A0(T)). + by rewrite (cfun_onS (FTsupp1_sub0 _)) // /'A1(T) ?FTcore_eq_der1. + have ZbetaT0: betaT0 \in 'Z[irr T]. + by rewrite rpredB ?char_vchar ?(seqInd_char T1zeta) ?prTIred_char. + pose Delta := tauT betaT0 - 1 + tau1T zeta. + have nz_i1: #1 != 0 := Iirr1_neq0 ntW2. + rewrite -(canLR (addKr _) (erefl Delta)) opprB cfdotDr cfdotBr oGa1 add0r. + rewrite cfdotDl cfdotBl -/betaS o_eta0_betaT0 (negPf nz_i1) // addr0 opprB. + rewrite -(cycTIiso1 pddS) -(cycTIirr00 defW) {}o_eta0_betaT0 mulr1n. + have QV'betaS: tauS betaS \in 'CF(G, ~: class_support QV^# G). + have [_ [pP _] _ _ [_ ->]] := FTtypeP_facts _ StypeP; rewrite ?A0betaS //. + apply: cfun_onS (cfInd_on (subsetT S) (PVbetaS _ nz_i1)). + apply/subsetP=> x PWx; rewrite inE. + have{PWx}: p \in \pi(#[x]). + case/imset2P: PWx => {x}x y PWx _ ->; rewrite {y}orderJ. + case/setUP: PWx => [/setD1P[ntx Px] | /imset2P[{x}x y Wx _ ->]]. + rewrite -p_rank_gt0 -rank_pgroup ?rank_gt0 ?cycle_eq1 //. + exact: mem_p_elt (abelem_pgroup pP) Px. + case/setDP: Wx; rewrite {y}orderJ; have [_ <- cW12 _] := dprodP defW. + case/mulsgP=> {x}x y W1x W2y ->; have cyx := centsP cW12 _ W2y _ W1x. + have [-> | nty _] := eqVneq y 1%g; first by rewrite inE mulg1 W1x. + have p'x: p^'.-elt x. + by rewrite (mem_p_elt _ W1x) /pgroup ?pnatE ?inE ?ltn_eqF. + have p_y: p.-elt y by rewrite (mem_p_elt (pnat_id _)). + rewrite -cyx orderM ?(pnat_coprime p_y) // pi_ofM // inE /=. + by rewrite -p_rank_gt0 -rank_pgroup // rank_gt0 cycle_eq1 nty. + apply: contraL => /imset2P[z y /setD1P[_ QVz] _ ->]; rewrite {x y}orderJ. + rewrite -p'natEpi // [_.-nat _](mem_p_elt _ QVz) // /pgroup ?p'natE //. + rewrite -prime_coprime // coprime_sym (coprime_sdprod_Hall_r defT). + by have [[]] := TtypeP. + have [_ _ _ _ [_ -> //]] := FTtypeP_facts _ TtypeP. + rewrite (cfdotElr QV'betaS (cfInd_on _ QVbetaT0)) ?subsetT //=. + rewrite setIC setICr big_set0 mulr0 subr0 addrC /eqCmod addrK. + rewrite cfdot_real_vchar_even ?mFT_odd ?oGa1 ?rpred0 //; split. + rewrite rpredD ?Ztau1T ?(mem_zchar T1zeta) // rpredB ?rpred1 //. + by rewrite Dade_vchar // zchar_split ZbetaT0. + rewrite /cfReal -subr_eq0 opprD opprB rmorphD rmorphB rmorph1 /= addrACA. + rewrite !addrA subrK -Dade_aut -linearB /= -/tauT rmorphB opprB /=. + rewrite -prTIred_aut aut_Iirr0 -/nu_ [sum in tauT sum]addrC addrA subrK. + rewrite -Dtau1T; last first. + by rewrite (zchar_onS _ (seqInd_sub_aut_zchar _ _ _)) // setSD ?der_sub. + rewrite raddfB -addrA addrC addrA subrK subr_eq0. + by rewrite (cfConjC_Dade_coherent cohT1) ?mFT_odd ?irrT1. +have [Y T1_Y [X [defGa oYX oXT1]]] := orthogonal_split (map tau1T calT1) Gamma. +apply: ler_trans (lb_Ga X Y _ _ _); first 1 last; rewrite 1?addrC //. +- by rewrite cfdotC oYX conjC0. +- by apply/orthoPl=> eta Weta; rewrite (span_orthogonal oT1eta) // memv_span. +have ->: v.-1 = (p * size calT1)%N; last rewrite mulKn ?prime_gt0 //. + rewrite [p](index_sdprod defT); have isoV := sdprod_isog defQV. + rewrite [v](card_isog isoV) -card_Iirr_abelian -?(isog_abelian isoV) //. + rewrite -(card_imset _ (can_inj (mod_IirrK nsQQV))) (cardD1 0) /=. + rewrite -{1}(mod_Iirr0 QV Q) mem_imset //=. + rewrite (size_irr_subseq_seqInd _ (subseq_refl _)) //=. + apply: eq_card => s; rewrite !inE mem_seqInd ?gFnormal // !inE subGcfker. + congr (_ && _); apply/idP/idP=> [/imsetP[r _ ->] | kerQs]. + by rewrite mod_IirrE ?cfker_mod. + by rewrite -(quo_IirrK nsQQV kerQs) mem_imset. +have o1T1: orthonormal (map tau1T calT1). + rewrite map_orthonormal ?(sub_orthonormal irrT1) ?seqInd_uniq //. + exact: irr_orthonormal. +have [_ -> ->] := orthonormal_span o1T1 T1_Y. +rewrite cfnorm_sum_orthonormal // big_map -sum1_size natr_sum !big_seq. +apply: ler_sum => // zeta T1zeta; rewrite -(canLR (addrK X) defGa). +by rewrite cfdotBl (orthoPl oXT1) ?subr0 ?nzT1_Ga ?map_f. +Qed. +Let Ttype2 := FTtypeP_min_typeII. + +(* These declarations correspond to Hypothesis (14.10). *) +Variables (M : {group gT}) (tau1M : {additive 'CF(M) -> 'CF(G)}) (psi : 'CF(M)). +Hypothesis maxNV_M : M \in 'M('N(V)). + +Local Notation "` 'M'" := (gval M) (at level 0, only parsing). +Local Notation K := `M`_\F%G. +Local Notation "` 'K'" := `M`_\F%g (at level 0, format "` 'K'") : group_scope. + +(* Consequences of the above. *) +Hypotheses (maxM : M \in 'M) (sNVM : 'N(V) \subset M). +Hypotheses (frobM : [Frobenius M with kernel K]) (Mtype1 : FTtype M == 1%N). + +Let calM := seqIndD K M K 1. +Local Notation tauM := (FT_DadeF maxM). +Let nsKM : K <| M. Proof. exact: gFnormal. Qed. +Let irrM : {subset calM <= irr M}. Proof. exact: FTtype1_Ind_irr. Qed. + +Hypothesis cohM : coherent_with calM M^# tauM tau1M. +Hypotheses (Mpsi : psi \in calM) (psi1 : psi 1%g = #|M : K|%:R). + +Let betaM := 'Ind[M, K] 1 - psi. + +Let pairTS : typeP_pair T S xdefW. Proof. exact: typeP_pair_sym pairST. Qed. + +Let pq : algC := (p * q)%:R. +Let h := #|H|. + +(* This is the first (and main) part of Peterfalvi (14.11). *) +Let defK : `K = V. +Proof. +pose e := #|M : K|; pose k := #|K|; apply: contraTeq isT => notKV. +have [_ sVK defM] := FTtypeII_support_facts maxT TtypeP Ttype2 pairTS maxNV_M. +have ltVK: V \proper K by rewrite properEneq eq_sym notKV. +have e_dv_k1: e %| k.-1 := Frobenius_ker_dvd_ker1 frobM. +have [e_lepq regKW2]: (e <= p * q)%N /\ semiregular K W2. + case: defM => [|[y Py]] defM; rewrite /e -(index_sdprod defM). + have /Frobenius_reg_ker regHW1 := set_Frobenius_compl defM frobM. + by rewrite leq_pmulr ?cardG_gt0. + have /Frobenius_reg_ker regHW21y := set_Frobenius_compl defM frobM. + split; last exact: semiregularS (joing_subl _ _) regHW21y. + suffices /normP <-: y \in 'N(W2). + by rewrite -conjYg cardJg (dprodWY xdefW) -(dprod_card xdefW). + have cPP: abelian P by have [_ [/and3P[]]] := FTtypeP_facts maxS StypeP. + have sW2P: W2 \subset P by have [_ _ _ []] := StypeP. + by rewrite (subsetP _ y Py) // sub_abelian_norm. +(* This is (14.11.1). *) +have{regKW2} [lb_k lb_k1e_v]: (2 * p * v < k /\ v.-1 %/ p < k.-1 %/ e)%N. + have /dvdnP[x Dk]: v %| k := cardSg sVK. + have lb_x: (p.*2 < x)%N. + have x_gt1: (1 < x)%N. + by rewrite -(ltn_pmul2r (cardG_gt0 V)) -Dk mul1n proper_card. + have x_gt0 := ltnW x_gt1; rewrite -(prednK x_gt0) ltnS -subn1. + rewrite dvdn_leq ?subn_gt0 // -mul2n Gauss_dvd ?coprime2n ?mFT_odd //. + rewrite dvdn2 odd_sub // (dvdn_odd _ (mFT_odd K)) -/k ?Dk ?dvdn_mulr //=. + rewrite -eqn_mod_dvd // -[x]muln1 -modnMmr. + have nVW2: W2 \subset 'N(V) by have [_ []] := TtypeP. + have /eqP{1} <-: (v == 1 %[mod p]). + rewrite eqn_mod_dvd ?cardG_gt0 // subn1 regular_norm_dvd_pred //. + exact: semiregularS regKW2. + rewrite modnMmr -Dk /k eqn_mod_dvd // subn1 regular_norm_dvd_pred //. + by rewrite (subset_trans (subset_trans nVW2 sNVM)) ?gFnorm. + have lb_k: (2 * p * v < k)%N by rewrite mul2n Dk ltn_pmul2r ?cardG_gt0. + split=> //; rewrite ltn_divLR ?cardG_gt0 // divn_mulAC ?prednK ?cardG_gt0 //. + rewrite leq_divRL ?indexg_gt0 // (leq_trans (leq_mul (leqnn v) e_lepq)) //. + rewrite mulnA mulnAC leq_mul // -ltnS prednK ?cardG_gt0 //. + apply: leq_ltn_trans lb_k; rewrite mulnC leq_mul // ltnW ?(leq_trans ltqp) //. + by rewrite mul2n -addnn leq_addl. +have lb_k1e_u := ltn_trans v1p_gt_u1q lb_k1e_v; have irr_psi := irrM Mpsi. +have Mgt1: (1 < size calM)%N by apply: seqInd_nontrivial Mpsi; rewrite ?mFT_odd. +(* This is (14.11.2). *) +have [] // := FTtype2_support_coherence TtypeP StypeP cohM Mpsi. +rewrite -/e -/p -/q mulnC /= => De [nb [chi Dchi]]. +rewrite cycTIiso_irrelC -/sigma -/betaM => DbetaM. +pose ddMK := FT_DadeF_hyp maxM; pose AM := Dade_support ddMK. +have defAM: AM = 'A~(M) by rewrite FTsupp_Frobenius -?FT_DadeF_supportE. +pose ccG A := class_support A G. +pose G0 := ~: ('A~(M) :|: ccG What :|: ccG P^# :|: ccG Q^#). +have sW2P: W2 \subset P by have [_ _ _ []] := StypeP. +have sW1Q: W1 \subset Q by have [_ _ _ []] := TtypeP. +(* This is (14.11.3). *) +have lbG0 g: g \in G0 -> 1 <= `|tau1M psi g| ^+ 2. + rewrite !inE ?expr_ge1 ?normr_ge0 // => /norP[/norP[/norP[AM'g W'g P'g Q'g]]]. + have{W'g} /coprime_typeP_Galois_core-co_p_g: g \notin ccG W^#. + apply: contra W'g => /imset2P[x y /setD1P[ntx Wx] Gy Dg]. + rewrite Dg mem_imset2 // !inE Wx andbT; apply/norP; split. + by apply: contra Q'g => /(subsetP sW1Q)?; rewrite Dg mem_imset2 ?inE ?ntx. + by apply: contra P'g => /(subsetP sW2P)Px; rewrite Dg mem_imset2 ?inE ?ntx. + have{AM'g} betaMg0: tauM betaM g = 0. + by apply: cfun_on0 AM'g; rewrite -defAM Dade_cfunS. + suffices{betaMg0}: 1 <= `|(\sum_ij (-1) ^+ nb ij *: sigma 'chi_ij) g|. + rewrite -[\sum_i _](subrK chi) -DbetaM !cfunE betaMg0 add0r. + case: Dchi => -> //; rewrite cfunE normrN. + by rewrite -(cfConjC_Dade_coherent cohM) ?mFT_odd ?cfunE ?norm_conjC. + have{co_p_g} Zeta_g ij: sigma 'chi_ij g \in Cint. + apply/Cint_cycTIiso_coprime/(coprime_dvdr (cforder_lin_char_dvdG _)). + by apply: irr_cyclic_lin; have [] := ctiWG. + rewrite -(dprod_card defW) coprime_mulr. + by apply/andP; split; [apply: co_p_g galT _ | apply: co_p_g galS _]. + rewrite sum_cfunE norm_Cint_ge1 ?rpred_sum // => [ij _|]. + by rewrite cfunE rpredMsign. + set a := \sum_i _; suffices: (a == 1 %[mod 2])%C. + by apply: contraTneq=> ->; rewrite (eqCmod_nat 2 0 1). + have signCmod2 n ij (b := sigma 'chi_ij g): ((-1) ^+ n * b == b %[mod 2])%C. + rewrite -signr_odd mulr_sign eqCmod_sym; case: ifP => // _. + by rewrite -(eqCmodDl _ b) subrr -[b + b](mulr_natr b 2) eqCmodMl0 /b. + rewrite -[1]addr0 [a](bigD1 0) {a}//= cfunE eqCmodD //. + by rewrite (eqCmod_trans (signCmod2 _ _)) // irr0 cycTIiso1 cfun1E inE. + rewrite (partition_big_imset (fun ij => [set ij; conjC_Iirr ij])) /= eqCmod0. + apply: rpred_sum => _ /=/imsetP[ij /negPf nz_ij ->]. + rewrite (bigD1 ij) /=; last by rewrite unfold_in nz_ij eqxx. + rewrite (big_pred1 (conjC_Iirr ij)) => [|ij1 /=]; last first. + rewrite unfold_in eqEsubset !subUset !sub1set !inE !(eq_sym ij). + rewrite !(can_eq (@conjC_IirrK _ _)) (canF_eq (@conjC_IirrK _ _)). + rewrite -!(eq_sym ij1) -!(orbC (_ == ij)) !andbb andbAC -andbA. + rewrite andb_orr andNb andbA andb_idl // => /eqP-> {ij1}. + rewrite conjC_Iirr_eq0 nz_ij -(inj_eq irr_inj) conjC_IirrE. + by rewrite odd_eq_conj_irr1 ?mFT_odd // irr_eq1 nz_ij. + rewrite -signr_odd -[odd _]negbK signrN !cfunE mulNr addrC. + apply: eqCmod_trans (signCmod2 _ _) _. + by rewrite eqCmod_sym conjC_IirrE -cfAut_cycTIiso cfunE conj_Cint. +have cardG_D1 R: #|R^#| = #|R|.-1 by rewrite [#|R|](cardsD1 1%g) group1. +pose rho := invDade ddMK; pose nG : algC := #|G|%:R. +pose sumG0 := \sum_(g in G0) `|tau1M psi g| ^+ 2. +pose sumG0_diff := sumG0 - (#|G0| + #|ccG What| + #|ccG P^#| + #|ccG Q^#|)%:R. +have ub_rho: '[rho (tau1M psi)] <= k.-1%:R / #|M|%:R - nG^-1 * sumG0_diff. + have NtauMpsi: '[tau1M psi] = 1. + by have [[Itau1 _] _] := cohM; rewrite Itau1 ?mem_zchar //= irrWnorm. + rewrite ler_subr_addl -subr_le0 -addrA. + have ddM_ i j: i != j :> 'I_1 -> [disjoint AM & AM] by rewrite !ord1. + apply: ler_trans (Dade_cover_inequality ddM_ NtauMpsi); rewrite -/nG -/AM. + rewrite !big_ord1 cardG_D1 ler_add2r ler_pmul2l ?invr_gt0 ?gt0CG //= defAM. + rewrite setTD ler_add ?ler_opp2 ?leC_nat //; last first. + do 3!rewrite -?addnA -cardsUI ?addnA (leq_trans _ (leq_addr _ _)) //. + by rewrite subset_leq_card // -setCD setCS -!setUA subDset setUC. + rewrite (big_setID G0) /= (setIidPr _) ?setCS -?setUA ?subsetUl // ler_addl. + by apply: sumr_ge0 => g _; rewrite ?exprn_ge0 ?normr_ge0. +have lb_rho: 1 - pq / k%:R <= '[rho (tau1M psi)]. + have [_] := Dade_Ind1_sub_lin cohM Mgt1 irr_psi Mpsi psi1; rewrite -/e -/k. + rewrite odd_Frobenius_index_ler ?mFT_odd // => -[_ _ [|/(ler_trans _)->] //]. + by rewrite ler_add2l ler_opp2 ler_pmul2r ?invr_gt0 ?gt0CG // leC_nat. +have{rho sumG0 sumG0_diff ub_rho lb_rho} []: + ~ pq / k%:R + 2%:R / pq + (u * q)%:R^-1 + (v * p)%:R^-1 < p%:R^-1 + q%:R^-1. +- rewrite ler_gtF // -!addrA -ler_subl_addl -ler_subr_addl -(ler_add2l 1). + apply: ler_trans {ub_rho lb_rho}(ler_trans lb_rho ub_rho) _. + rewrite /sumG0_diff -!addnA natrD opprD addrA mulrBr opprB addrA. + rewrite ler_subl_addr ler_paddr //. + by rewrite mulr_ge0 ?invr_ge0 ?ler0n // subr_ge0 -sumr_const ler_sum. + rewrite mulrDl -!addrA addrCA [1 + _]addrA [_ + (_ - _)]addrA ler_add //. + rewrite -(Lagrange (normal_sub nsKM)) natrM invfM mulrA -/k -/e /pq -De. + rewrite ler_pmul2r ?invr_gt0 ?gt0CiG // ler_pdivr_mulr ?gt0CG //. + by rewrite mul1r leC_nat leq_pred. + rewrite [1 + _ + _]addrA addrAC !natrD !mulrDr !ler_add //; first 1 last. + + exact: (FTtype2_cc_core_ler StypeP). + + exact: (FTtype2_cc_core_ler TtypeP). + have [_ _ /card_support_normedTI->] := ctiWG. + rewrite natrM natf_indexg ?subsetT // mulrCA mulKf ?neq0CG // card_cycTIset. + rewrite mulnC -(dprod_card xdefW) /pq !natrM -!subn1 !natrB // -/p -/q invfM. + rewrite mulrACA !mulrBl ?divff ?neq0CG // !mul1r mulrBr mulr1 opprB. + by rewrite addrACA -opprB opprK. +rewrite -!addrA ler_lt_add //; last first. + pose q2 : algC := (q ^ 2)%:R. + apply: ltr_le_trans (_ : 2%:R / q2 + (2%:R * q2)^-1 *+ 2 <= _); last first. + rewrite addrC -[_ *+ 2]mulr_natl invfM mulVKf ?pnatr_eq0 //. + rewrite mulr_natl -mulrS -mulr_natl [q2]natrM. + by rewrite ler_pdivr_mulr ?mulr_gt0 ?gt0CG // mulKf ?neq0CG ?leC_nat. + rewrite -natrM !addrA ltr_add ?(FTtypeP_complV_ltr TtypeP) 1?ltnW //. + rewrite ltr_add ?(FTtypeP_complV_ltr StypeP) // /pq mulnC /q2 !natrM !invfM. + by rewrite !ltr_pmul2l ?ltf_pinv ?invr_gt0 ?qualifE ?gt0CG ?ltr0n ?ltr_nat. +rewrite ler_pdivr_mulr ?ler_pdivl_mull ?gt0CG // -natrM leC_nat. +apply: leq_trans lb_k; rewrite leqW // mulnAC mulnC leq_mul //. +have [[_ _ frobVW2 _] _ _ _ _] := FTtypeP_facts maxT TtypeP. +rewrite -[(p * q)%N]mul1n leq_mul // (leq_trans _ (leq_pred v)) // dvdn_leq //. + by rewrite -subn1 subn_gt0 cardG_gt1; have[] := Frobenius_context frobVW2. +rewrite Gauss_dvd ?prime_coprime ?(dvdn_prime2 pr_p pr_q) ?gtn_eqF //. +rewrite (Frobenius_dvd_ker1 frobVW2) /= oV /nV predn_exp. +rewrite -(subnKC qgt2) -(ltn_predK pgt2) mulKn // subnKC //. +by rewrite big_ord_recl dvdn_sum // => i _; rewrite expnS dvdn_mulr. +Qed. + +(* This is the first part of Peterfalvi (14.11). *) +Let indexMK : #|M : K| = (p * q)%N. +Proof. +have [_ _ [defM|]] := FTtypeII_support_facts maxT TtypeP Ttype2 pairTS maxNV_M. + have:= Ttype2; rewrite (mmax_max maxM (mmax_proper maxT)) ?(eqP Mtype1) //. + rewrite -(sdprodW (Ptype_Fcore_sdprod TtypeP)) -defK (sdprodWY defM). + exact: mulG_subr. +case=> y Py /index_sdprod <-; rewrite (dprod_card xdefW) -(dprodWY xdefW). +suffices /normP {1}<-: y \in 'N(W2) by rewrite -conjYg cardJg. +have cPP: abelian P by have [_ [/and3P[]]] := FTtypeP_facts maxS StypeP. +by rewrite (subsetP (sub_abelian_norm cPP _)) //; have [_ _ _ []] := StypeP. +Qed. + +(* This is Peterfalvi (14.12), and also (14.13) since we have already proved *) +(* the negation of Theorem (14.2). *) +Let not_MG_L : (L : {set gT}) \notin M :^: G. +Proof. +rewrite orbit_sym; apply: contra not_charUH => /imsetP[z _ /= defLz]. +rewrite sub_cyclic_char // -(cyclicJ _ z) -FcoreJ -defLz defK. +have [_ _ [cycV _ _]] := typeP_Galois_P maxT (FTtype5_exclusion maxT) galT. +rewrite Ptype_Fcompl_kernel_trivial // in cycV. +by rewrite -(isog_cyclic (quotient1_isog V)) in cycV. +Qed. + +(* This is Peterfalvi (14.14). *) +Let LM_cases : + '[tauM betaM, tau1L phi] != 0 /\ h.-1%:R / pq <= pq - 1 + \/ '[tauL betaL, tau1M psi] != 0 /\ q = 3 /\ p = 5. +Proof. +have [irr_phi irr_psi] := (irrL Lphi, irrM Mpsi). +have:= Dade_sub_lin_nonorthogonal (mFT_odd _) _ cohM cohL _ Mpsi _ _ Lphi. +rewrite -/betaL -/betaM disjoint_Dade_FTtype1 //. +case=> //; set a := '[_, _] => nz_a; [left | right]; split=> //. + rewrite {1}/pq -indexLH /pq -indexMK. + by rewrite (coherent_FTtype1_core_ltr cohM cohL Mpsi Lphi) // orbit_sym. +have ub_v: v.-1%:R / pq <= pq - 1. + rewrite {1}/pq -indexMK /pq -indexLH /v -defK. + exact: (coherent_FTtype1_core_ltr cohL cohM Lphi Mpsi). +have{ub_v} ub_qp: (q ^ (p - 3) < p ^ 2)%N. + rewrite -(@ltn_pmul2l (q ^ 3)) ?expn_gt0 ?cardG_gt0 // -expnD subnKC //. + have: v.-1%:R < pq ^+ 2. + rewrite -ltr_pdivr_mulr ?ltr0n ?muln_gt0 ?cardG_gt0 //. + by rewrite (ler_lt_trans ub_v) // ltr_subl_addl -mulrS ltC_nat. + rewrite -natrX ltC_nat prednK ?cardG_gt0 // mulnC expnMn oV. + rewrite leq_divLR ?dvdn_pred_predX // mulnC -subn1 leq_subLR. + move/leq_ltn_trans->; rewrite // -addSn addnC -(leq_add2r (q ^ 2 * p ^ 2)). + rewrite addnAC -mulSnr prednK ?cardG_gt0 // mulnA leq_add2l -expnMn. + by rewrite (ltn_sqr 1) (@ltn_mul 1 1) ?prime_gt1. +have q3: q = 3. + apply/eqP; rewrite eqn_leq qgt2 -ltnS -(odd_ltn 5) ?mFT_odd // -ltnS. + rewrite -(ltn_exp2l _ _ (ltnW pgt2)) (leq_trans qp1_gt_pq1) // ltnW //. + by rewrite -{1}(subnK pgt2) -addnS expnD (expnD p 2 4) ltn_mul ?ltn_exp2r. +split=> //; apply/eqP; rewrite eqn_leq -ltnS andbC. +rewrite (odd_geq 5) -1?(odd_ltn 7) ?mFT_odd //= doubleS -{1}q3 ltqp /=. +move: ub_qp; rewrite 2!ltnNge q3; apply: contra. +elim: p => // x IHx; rewrite ltnS leq_eqVlt => /predU1P[<- // | xgt6]. +apply: (@leq_trans (3 * x ^ 2)); last first. + rewrite subSn ?(leq_trans _ xgt6) //. + by rewrite [rhs in (_ <= rhs)%N]expnS leq_mul ?IHx. +rewrite -addn1 sqrnD -addnA (mulSn 2) leq_add2l muln1. +rewrite (@leq_trans (2 * (x * 7))) ?leq_mul //. +by rewrite mulnCA (mulnDr x 12 2) mulnC leq_add2r -(subnKC xgt6). +Qed. + +(* This is Peterfalvi (14.15). *) +Let oU : u = nU. +Proof. +case: ifP (card_FTtypeP_Galois_compl maxS galS) => // p1modq oU. +pose x := #|H : U|; rewrite -/u -/nU -/p -/q in p1modq oU. +have DnU: (q * u)%N = nU. + rewrite mulnC oU divnK //. + by have [_ ->] := FTtypeP_primes_mod_cases maxS StypeP. +have oH: h = (u * x)%N by rewrite Lagrange. +have xmodp: x = q %[mod p]. + have hmodp: h = 1 %[mod p]. + apply/eqP; rewrite eqn_mod_dvd ?cardG_gt0 // subn1. + apply: dvdn_trans (Frobenius_ker_dvd_ker1 frobL). + have [y _ /index_sdprod <-] := defL. + by rewrite -[p](cardJg _ y) cardSg ?joing_subr. + rewrite -[q]muln1 -modnMmr -hmodp modnMmr oH mulnA DnU -modnMml. + suffices ->: nU = 1 %[mod p] by rewrite modnMml mul1n. + rewrite /nU predn_exp mulKn; last by rewrite -(subnKC pgt2). + apply/eqP; rewrite -(ltn_predK qgt2) big_ord_recl eqn_mod_dvd ?subn1 //=. + by apply: dvdn_sum => i _; rewrite expnS dvdn_mulr. +have{xmodp} [n Dx]: {n | x = q + n * p}%N. + by exists (x %/ p); rewrite -(modn_small ltqp) addnC -xmodp -divn_eq. +have nmodq: n = 1 %[mod q]. + have [y _ defLy] := defL; have [_ _ /joing_subP[nHW1 _] _] := sdprodP defLy. + have regHW1: semiregular H W1. + have /Frobenius_reg_ker := set_Frobenius_compl defLy frobL. + by apply: semiregularS; rewrite ?joing_subl. + have hmodq: h = 1 %[mod q]. + apply/eqP; rewrite eqn_mod_dvd ?cardG_gt0 // subn1. + exact: regular_norm_dvd_pred regHW1. + have umodq: u = 1 %[mod q]. + apply/eqP; rewrite eqn_mod_dvd ?cardG_gt0 // subn1. + apply: regular_norm_dvd_pred; first by have [_ []] := StypeP. + exact: semiregularS regHW1. + rewrite -hmodq oH -modnMml umodq modnMml mul1n Dx modnDl. + by rewrite -modnMmr (eqP p1modq) modnMmr muln1. +have{n nmodq Dx} lb_x: (q + q.+1 * p <= x)%N. + rewrite (divn_eq n q) nmodq (modn_small (ltnW qgt2)) addn1 in Dx. + rewrite Dx leq_add2l leq_mul // ltnS leq_pmull // lt0n. + have: odd x by rewrite (dvdn_odd (dvdn_indexg _ _)) ?mFT_odd. + by rewrite Dx odd_add odd_mul !mFT_odd; apply: contraNneq => ->. +have lb_h: (p ^ q < h)%N. + rewrite (@leq_trans (p * nU)) //; last first. + rewrite -DnU oH mulnA mulnC leq_mul // (leq_trans _ lb_x) //. + by rewrite mulSn addnA mulnC leq_addl. + rewrite /nU predn_exp mulKn; last by rewrite -(subnKC pgt2). + rewrite -(subnKC (ltnW qgt2)) subn2 big_ord_recr big_ord_recl /=. + by rewrite -add1n !mulnDr -!expnS -addnA leq_add ?leq_addl // cardG_gt0. +have ub_h: (h <= p ^ 2 * q ^ 2)%N. + have [[_ ub_h] | [_ [q3 p5]]] := LM_cases; last by rewrite q3 p5 in p1modq. + rewrite -expnMn -(ltn_predK lb_h) -ltC_nat natrM -/pq. + rewrite -ltr_pdivr_mulr ?ltr0n ?muln_gt0 ?cardG_gt0 //. + by rewrite (ler_lt_trans ub_h) // ltr_subl_addl -mulrS ltC_nat. +have{lb_h} lb_q2: (p ^ q.-2 < q ^ 2)%N. + rewrite -(@ltn_pmul2l (p ^ 2)) ?expn_gt0 ?cardG_gt0 // (leq_trans _ ub_h) //. + by rewrite -subn2 -expnD subnKC // ltnW. +have q3: q = 3. + apply/eqP; rewrite eqn_leq qgt2 -(subnKC (ltnW qgt2)) subn2 ltnS. + by rewrite -(ltn_exp2l _ _ (ltnW pgt2)) (ltn_trans lb_q2) ?ltn_exp2r. +have{lb_q2 p1modq} p7: p = 7. + suff: p \in [seq n <- iota 4 5 | prime n & n == 1 %[mod 3]] by case/predU1P. + by rewrite mem_filter pr_p mem_iota -q3 p1modq ltqp; rewrite q3 in lb_q2 *. +rewrite oH mulnC oU /nU q3 p7 -leq_divRL //= in ub_h lb_x. +by have:= leq_trans lb_x ub_h. +Qed. + +(* This is Peterfalvi (14.16), the last step towards the final contradiction. *) +Let defH : `H = U. +Proof. +pose x := #|H : U|; have oH: h = (u * x)%N by rewrite Lagrange. +apply/eqP/idPn; rewrite eqEsubset sUH andbT -indexg_gt1 -/x => xgt1. +have hmodpq: h = 1 %[mod p * q]. + apply/eqP; rewrite eqn_mod_dvd ?cardG_gt0 // -indexLH subn1. + exact: Frobenius_ker_dvd_ker1. +have [[_ _ frobUW1 _] _ _ _ _] := FTtypeP_facts maxS StypeP. +have /eqP umodpq: u == 1 %[mod p * q]. + rewrite chinese_remainder ?prime_coprime ?dvdn_prime2 ?(gtn_eqF ltqp) //. + rewrite !eqn_mod_dvd ?cardG_gt0 // subn1 (Frobenius_dvd_ker1 frobUW1). + rewrite oU /nU predn_exp mulKn; last by rewrite -(subnKC pgt2). + by rewrite -(ltn_predK qgt2) big_ord_recl dvdn_sum //= => i; rewrite dvdn_exp. +have{hmodpq} lb_x: (p * q < x)%N. + rewrite -(subnKC (ltnW xgt1)) ltnS dvdn_leq ?subn_gt0 //. + by rewrite -eqn_mod_dvd 1?ltnW // -hmodpq oH -modnMml umodpq modnMml mul1n. +have [[_ ub_h] | [nz_a [q3 p5]]] := LM_cases. + have /idPn[]: (p * q < u)%N. + have ugt1: (1 < u)%N. + by rewrite cardG_gt1; have [] := Frobenius_context frobUW1. + rewrite -(subnKC (ltnW ugt1)) ltnS dvdn_leq ?subn_gt0 //. + by rewrite -eqn_mod_dvd ?umodpq 1?ltnW. + rewrite -leqNgt -(leq_pmul2r (indexg_gt0 L H)) indexLH. + apply: (@leq_trans h.-1). + by rewrite -ltnS prednK ?cardG_gt0 // oH ltn_pmul2l ?cardG_gt0. + rewrite -indexLH -leC_nat natrM -ler_pdivr_mulr ?gt0CiG // indexLH -/pq. + by rewrite (ler_trans ub_h) // ler_subl_addl -mulrS leC_nat ltnW. +have lb_h1e_v: (v.-1 %/ p < h.-1 %/ #|L : H|)%N. + rewrite -(@ltn_pmul2l u) ?cardG_gt0 // -oH oU /nU q3 p5 /= in lb_x. + rewrite -(ltn_subRL 1) /= subn1 in lb_x. + by rewrite leq_divRL ?indexG_gt0 // oV /nV indexLH q3 p5 (leq_trans _ lb_x). +have oLM: orthogonal (map tau1L calL) (map tau1M calM). + by rewrite orthogonal_sym coherent_FTtype1_ortho. +case/eqP: nz_a; have lb_h1e_u := ltn_trans v1p_gt_u1q lb_h1e_v. +have [] // := FTtype2_support_coherence StypeP TtypeP cohL Lphi. +rewrite -/tauL -/sigma => _ [nb [chi Dchi ->]]. +rewrite cfdotBl cfdot_suml big1 => [|ij _]; last first. + have [_ o_tauMeta _ _] := FTtypeI_bridge_facts _ StypeP Mtype1 cohM Mpsi psi1. + rewrite cfdotZl cfdotC (orthogonalP o_tauMeta) ?map_f ?mem_irr //. + by rewrite conjC0 mulr0. +case: Dchi => ->; first by rewrite (orthogonalP oLM) ?map_f // subr0. +by rewrite cfdotNl opprK add0r (orthogonalP oLM) ?map_f // cfAut_seqInd. +Qed. + +Lemma FTtype2_exclusion : False. +Proof. by have /negP[] := not_charUH; rewrite /= defH char_refl. Qed. + +End Fourteen. + +Lemma no_minSimple_odd_group (gT : minSimpleOddGroupType) : False. +Proof. +have [/forall_inP | [S [T [_ W W1 W2 defW pairST]]]] := FTtypeP_pair_cases gT. + exact/negP/not_all_FTtype1. +have xdefW: W2 \x W1 = W by rewrite dprodC. +have pairTS := typeP_pair_sym xdefW pairST. +pose p := #|W2|; pose q := #|W1|. +have p'q: q != p. + have [[[ctiW _ _] _ _ _ _] /mulG_sub[sW1W sW2W]] := (pairST, dprodW defW). + have [cycW _ _] := ctiW; apply: contraTneq (cycW) => eq_pq. + rewrite (cyclic_dprod defW) ?(cyclicS _ cycW) // -/q eq_pq. + by rewrite /coprime gcdnn -trivg_card1; have [] := cycTI_nontrivial ctiW. +without loss{p'q} ltqp: S T W1 W2 defW xdefW pairST pairTS @p @q / q < p. + move=> IH_ST; rewrite neq_ltn in p'q. + by case/orP: p'q; [apply: (IH_ST S T) | apply: (IH_ST T S)]. +have [[_ maxS maxT] _ _ _ _] := pairST. +have [[U StypeP] [V TtypeP]] := (typeP_pairW pairST, typeP_pairW pairTS). +have Stype2: FTtype S == 2 := FTtypeP_max_typeII maxS StypeP ltqp. +have Ttype2: FTtype T == 2 := FTtypeP_min_typeII maxS maxT StypeP TtypeP ltqp. +have /mmax_exists[L maxNU_L]: 'N(U) \proper setT. + have [[_ ntU _ _] cUU _ _ _] := compl_of_typeII maxS StypeP Stype2. + by rewrite mFT_norm_proper // mFT_sol_proper abelian_sol. +have /mmax_exists[M maxNV_M]: 'N(V) \proper setT. + have [[_ ntV _ _] cVV _ _ _] := compl_of_typeII maxT TtypeP Ttype2. + by rewrite mFT_norm_proper // mFT_sol_proper abelian_sol. +have [[maxL sNU_L] [maxM sNV_M]] := (setIdP maxNU_L, setIdP maxNV_M). +have [frobL sUH _] := FTtypeII_support_facts maxS StypeP Stype2 pairST maxNU_L. +have [frobM _ _] := FTtypeII_support_facts maxT TtypeP Ttype2 pairTS maxNV_M. +have Ltype1 := FT_Frobenius_type1 maxL frobL. +have Mtype1 := FT_Frobenius_type1 maxM frobM. +have [tau1L cohL] := FTtype1_coherence maxL Ltype1. +have [tau1M cohM] := FTtype1_coherence maxM Mtype1. +have [phi Lphi phi1] := FTtype1_ref_irr maxL. +have [psi Mpsi psi1] := FTtype1_ref_irr maxM. +exact: (FTtype2_exclusion pairST maxS maxT StypeP TtypeP ltqp + maxNU_L sNU_L sUH frobL Ltype1 cohL Lphi phi1 + maxNV_M sNV_M frobM Mtype1 cohM Mpsi psi1). +Qed. + +Theorem Feit_Thompson (gT : finGroupType) (G : {group gT}) : + odd #|G| -> solvable G. +Proof. exact: (minSimpleOdd_ind no_minSimple_odd_group). Qed. + +Theorem simple_odd_group_prime (gT : finGroupType) (G : {group gT}) : + odd #|G| -> simple G -> prime #|G|. +Proof. exact: (minSimpleOdd_prime no_minSimple_odd_group). Qed. + + diff --git a/mathcomp/odd_order/PFsection2.v b/mathcomp/odd_order/PFsection2.v new file mode 100644 index 0000000..9eef9e8 --- /dev/null +++ b/mathcomp/odd_order/PFsection2.v @@ -0,0 +1,822 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg poly finset center. +Require Import fingroup morphism perm automorphism quotient action zmodp. +Require Import gfunctor gproduct cyclic pgroup frobenius ssrnum. +Require Import matrix mxalgebra mxrepresentation vector algC classfun character. +Require Import inertia vcharacter PFsection1. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 2: the Dade isometry *) +(* Defined here: *) +(* Dade_hypothesis G L A <-> G, L, and A satisfy the hypotheses under which *) +(* which the Dade isometry relative to G, L and *) +(* A is well-defined. *) +(* For ddA : Dade_hypothesis G L A, we also define *) +(* Dade ddA == the Dade isometry relative to G, L and A. *) +(* Dade_signalizer ddA a == the normal complement to 'C_L[a] in 'C_G[a] for *) +(* a in A (this is usually denoted by H a). *) +(* Dade_support1 ddA a == the set of elements identified with a by the Dade *) +(* isometry. *) +(* Dade_support ddA == the natural support of the Dade isometry. *) +(* The following are used locally in expansion of the Dade isometry as a sum *) +(* of induced characters: *) +(* Dade_transversal ddA == a transversal of the L-conjugacy classes *) +(* of non empty subsets of A. *) +(* Dade_set_signalizer ddA B == the generalization of H to B \subset A, *) +(* denoted 'H(B) below. *) +(* Dade_set_normalizer ddA B == the generalization of 'C_G[a] to B. *) +(* denoted 'M(B) = 'H(B) ><| 'N_L(B) below. *) +(* Dade_cfun_restriction ddA B aa == the composition of aa \in 'CF(L, A) *) +(* with the projection of 'M(B) onto 'N_L(B), *) +(* parallel to 'H(B). *) +(* In addition, if sA1A : A1 \subset A and nA1L : L \subset 'N(A1), we have *) +(* restr_Dade_hyp ddA sA1A nA1L : Dade_hypothesis G L A1 H *) +(* restr_Dade ddA sA1A nA1L == the restriction of the Dade isometry to *) +(* 'CF(L, A1). *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Reserved Notation "alpha ^\tau" (at level 2, format "alpha ^\tau"). + +Section Two. + +Variable gT : finGroupType. + +(* This is Peterfalvi (2.1). *) +Lemma partition_cent_rcoset (H : {group gT}) g (C := 'C_H[g]) (Cg := C :* g) : + g \in 'N(H) -> coprime #|H| #[g] -> + partition (Cg :^: H) (H :* g) /\ #|Cg :^: H| = #|H : C|. +Proof. +move=> nHg coHg; pose pi := \pi(#[g]). +have notCg0: Cg != set0 by apply/set0Pn; exists g; exact: rcoset_refl. +have id_pi: {in Cg, forall u, u.`_ pi = g}. + move=> _ /rcosetP[u /setIP[Hu cgu] ->]; rewrite consttM; last exact/cent1P. + rewrite (constt_p_elt (pgroup_pi _)) (constt1P _) ?mul1g //. + by rewrite (mem_p_elt _ Hu) // /pgroup -coprime_pi' // coprime_sym. +have{id_pi} /and3P[_ tiCg /eqP defC]: normedTI Cg H C. + apply/normedTI_P; rewrite subsetI subsetIl normsM ?normG ?subsetIr //. + split=> // x Hx /pred0Pn[u /andP[/= Cu Cxu]]; rewrite !inE Hx /= conjg_set1. + by rewrite -{2}(id_pi _ Cu) -(conjgKV x u) consttJ id_pi -?mem_conjg. +have{tiCg} partCg := partition_class_support notCg0 tiCg. +have{defC} oCgH: #|Cg :^: H| = #|H : C| by rewrite -defC -astab1Js -card_orbit. +split=> //; congr (partition _ _): (partCg); apply/eqP. +rewrite eqEcard card_rcoset {1}class_supportEr; apply/andP; split. + apply/bigcupsP=> x Hx; rewrite conjsgE -rcosetM conjgCV rcosetM mulgA. + by rewrite mulSg ?mul_subG ?subsetIl // sub1set ?memJ_norm ?groupV. +have oCg Cx: Cx \in Cg :^: H -> #|Cx| = #|C|. + by case/imsetP=> x _ ->; rewrite cardJg card_rcoset. +by rewrite (card_uniform_partition oCg partCg) oCgH mulnC Lagrange ?subsetIl. +Qed. + +Definition is_Dade_signalizer (G L A : {set gT}) (H : gT -> {group gT}) := + {in A, forall a, H a ><| 'C_L[a] = 'C_G[a]}. + +(* This is Peterfalvi Definition (2.2). *) +Definition Dade_hypothesis (G L A : {set gT}) := + [/\ A <| L, L \subset G, 1%g \notin A, + (*a*) {in A &, forall x, {subset x ^: G <= x ^: L}} + & (*b*) exists2 H, is_Dade_signalizer G L A H + & (*c*) {in A &, forall a b, coprime #|H a| #|'C_L[b]| }]. + +Variables (G L : {group gT}) (A : {set gT}). + +Let pi := [pred p | [exists a in A, p \in \pi('C_L[a])]]. + +Let piCL a : a \in A -> pi.-group 'C_L[a]. +Proof. +move=> Aa; apply: sub_pgroup (pgroup_pi _) => p cLa_p. +by apply/exists_inP; exists a. +Qed. + +Fact Dade_signalizer_key : unit. Proof. by []. Qed. +Definition Dade_signalizer_def a := 'O_pi^'('C_G[a])%G. +Definition Dade_signalizer of Dade_hypothesis G L A := + locked_with Dade_signalizer_key Dade_signalizer_def. + +Hypothesis ddA : Dade_hypothesis G L A. +Notation H := (Dade_signalizer ddA). +Canonical Dade_signalizer_unlockable := [unlockable fun H]. + +Let pi'H a : pi^'.-group (H a). Proof. by rewrite unlock pcore_pgroup. Qed. +Let nsHC a : H a <| 'C_G[a]. Proof. by rewrite unlock pcore_normal. Qed. + +Lemma Dade_signalizer_sub a : H a \subset G. +Proof. by have /andP[/subsetIP[]] := nsHC a. Qed. + +Lemma Dade_signalizer_cent a : H a \subset 'C[a]. +Proof. by have /andP[/subsetIP[]] := nsHC a. Qed. + +Let nsAL : A <| L. Proof. by have [->] := ddA. Qed. +Let sAL : A \subset L. Proof. exact: normal_sub nsAL. Qed. +Let nAL : L \subset 'N(A). Proof. exact: normal_norm nsAL. Qed. +Let sLG : L \subset G. Proof. by have [_ ->] := ddA. Qed. +Let notA1 : 1%g \notin A. Proof. by have [_ _ ->] := ddA. Qed. +Let conjAG : {in A &, forall x, {subset x ^: G <= x ^: L}}. +Proof. by have [_ _ _ ? _] := ddA. Qed. +Let sHG := Dade_signalizer_sub. +Let cHA := Dade_signalizer_cent. +Let notHa0 a : H a :* a :!=: set0. +Proof. by rewrite -cards_eq0 -lt0n card_rcoset cardG_gt0. Qed. + +Let HallCL a : a \in A -> pi.-Hall('C_G[a]) 'C_L[a]. +Proof. +move=> Aa; have [_ _ _ _ [H1 /(_ a Aa)/sdprodP[_ defCa _ _] coH1L]] := ddA. +have [|//] := coprime_mulGp_Hall defCa _ (piCL Aa). +apply: sub_pgroup (pgroup_pi _) => p; apply: contraL => /exists_inP[b Ab]. +by apply: (@pnatPpi \pi(_)^'); rewrite -coprime_pi' ?cardG_gt0 ?coH1L. +Qed. + +Lemma def_Dade_signalizer H1 : is_Dade_signalizer G L A H1 -> {in A, H =1 H1}. +Proof. +move=> defH1 a Aa; apply/val_inj; rewrite unlock /=; have defCa := defH1 a Aa. +have /sdprod_context[nsH1Ca _ _ _ _] := defCa. +by apply/normal_Hall_pcore=> //; apply/(sdprod_normal_pHallP _ (HallCL Aa)). +Qed. + +Lemma Dade_sdprod : is_Dade_signalizer G L A H. +Proof. +move=> a Aa; have [_ _ _ _ [H1 defH1 _]] := ddA. +by rewrite (def_Dade_signalizer defH1) ?defH1. +Qed. +Let defCA := Dade_sdprod. + +Lemma Dade_coprime : {in A &, forall a b, coprime #|H a| #|'C_L[b]| }. +Proof. by move=> a b _ Ab; apply: p'nat_coprime (pi'H a) (piCL Ab). Qed. +Let coHL := Dade_coprime. + +Definition Dade_support1 a := class_support (H a :* a) G. +Local Notation dd1 := Dade_support1. + +Lemma mem_Dade_support1 a x : a \in A -> x \in H a -> (x * a)%g \in dd1 a. +Proof. by move=> Aa Hx; rewrite -(conjg1 (x * a)) !mem_imset2 ?set11. Qed. + +(* This is Peterfalvi (2.3), except for the existence part, which is covered *) +(* below in the NormedTI section. *) +Lemma Dade_normedTI_P : + reflect (A != set0 /\ {in A, forall a, H a = 1%G}) (normedTI A G L). +Proof. +apply: (iffP idP) => [tiAG | [nzA trivH]]. + split=> [|a Aa]; first by have [] := andP tiAG. + apply/trivGP; rewrite -(coprime_TIg (coHL Aa Aa)) subsetIidl subsetI cHA. + by rewrite (subset_trans (normal_sub (nsHC a))) ?(cent1_normedTI tiAG). +apply/normedTI_memJ_P; split=> // a g Aa Gg. +apply/idP/idP=> [Aag | Lg]; last by rewrite memJ_norm ?(subsetP nAL). +have /imsetP[k Lk def_ag] := conjAG Aa Aag (mem_imset _ Gg). +suffices: (g * k^-1)%g \in 'C_G[a]. + by rewrite -Dade_sdprod ?trivH // sdprod1g inE groupMr ?groupV // => /andP[]. +rewrite !inE groupM ?groupV // ?(subsetP sLG) //=. +by rewrite conjg_set1 conjgM def_ag conjgK. +Qed. + +(* This is Peterfalvi (2.4)(a) (extended to all a thanks to our choice of H). *) +Lemma DadeJ a x : x \in L -> H (a ^ x) :=: H a :^ x. +Proof. +by move/(subsetP sLG)=> Gx; rewrite unlock -pcoreJ conjIg -cent1J conjGid. +Qed. + +Lemma Dade_support1_id a x : x \in L -> dd1 (a ^ x) = dd1 a. +Proof. +move=> Lx; rewrite {1}/dd1 DadeJ // -conjg_set1 -conjsMg. +by rewrite class_supportGidl ?(subsetP sLG). +Qed. + +Let piHA a u : a \in A -> u \in H a :* a -> u.`_pi = a. +Proof. +move=> Aa /rcosetP[{u}u Hu ->]; have pi'u: pi^'.-elt u by apply: mem_p_elt Hu. +rewrite (consttM _ (cent1P (subsetP (cHA a) u Hu))). +suffices pi_a: pi.-elt a by rewrite (constt1P pi'u) (constt_p_elt _) ?mul1g. +by rewrite (mem_p_elt (piCL Aa)) // inE cent1id (subsetP sAL). +Qed. + +(* This is Peterfalvi (2.4)(b). *) +Lemma Dade_support1_TI : {in A &, forall a b, + ~~ [disjoint dd1 a & dd1 b] -> exists2 x, x \in L & b = a ^ x}. +Proof. +move=> a b Aa Ab /= /pred0Pn[_ /andP[/imset2P[x u /(piHA Aa) def_x Gu ->]]] /=. +case/imset2P=> y v /(piHA Ab) def_y Gv /(canLR (conjgK v)) def_xuv. +have def_b: a ^ (u * v^-1) = b by rewrite -def_x -consttJ conjgM def_xuv def_y. +by apply/imsetP/conjAG; rewrite // -def_b mem_imset ?groupM ?groupV. +Qed. + +(* This is an essential strengthening of Peterfalvi (2.4)(c). *) +Lemma Dade_cover_TI : {in A, forall a, normedTI (H a :* a) G 'C_G[a]}. +Proof. +move=> a Aa; apply/normedTI_P; split=> // [|g Gg]. + by rewrite subsetI subsetIl normsM ?subsetIr ?normal_norm ?nsHC. +rewrite disjoint_sym => /pred0Pn[_ /andP[/imsetP[u Ha_u ->] Ha_ug]]. +by rewrite !inE Gg /= conjg_set1 -{1}(piHA Aa Ha_u) -consttJ (piHA Aa). +Qed. + +(* This is Peterfalvi (2.4)(c). *) +Lemma norm_Dade_cover : {in A, forall a, 'N_G(H a :* a) = 'C_G[a]}. +Proof. by move=> a /Dade_cover_TI /and3P[_ _ /eqP]. Qed. + +Definition Dade_support := \bigcup_(a in A) dd1 a. +Local Notation Atau := Dade_support. + +Lemma not_support_Dade_1 : 1%g \notin Atau. +Proof. +apply: contra notA1 => /bigcupP[a Aa /imset2P[u x Ha_u _ ux1]]. +suffices /set1P <-: a \in [1] by []. +have [_ _ _ <-] := sdprodP (defCA Aa). +rewrite 2!inE cent1id (subsetP sAL) // !andbT. +by rewrite -groupV -(mul1g a^-1)%g -mem_rcoset -(conj1g x^-1) ux1 conjgK. +Qed. + +Lemma Dade_support_sub : Atau \subset G. +Proof. +apply/bigcupsP=> a Aa; rewrite class_support_subG // mul_subG ?sHG //. +by rewrite sub1set (subsetP sLG) ?(subsetP sAL). +Qed. + +Lemma Dade_support_norm : G \subset 'N(Atau). +Proof. +by rewrite norms_bigcup //; apply/bigcapsP=> a _; exact: class_support_norm. +Qed. + +Lemma Dade_support_normal : Atau <| G. +Proof. by rewrite /normal Dade_support_sub Dade_support_norm. Qed. + +Lemma Dade_support_subD1 : Atau \subset G^#. +Proof. by rewrite subsetD1 Dade_support_sub not_support_Dade_1. Qed. + +(* This is Peterfalvi Definition (2.5). *) +Fact Dade_subproof (alpha : 'CF(L)) : + is_class_fun <<G>> [ffun x => oapp alpha 0 [pick a in A | x \in dd1 a]]. +Proof. +rewrite genGid; apply: intro_class_fun => [x y Gx Gy | x notGx]. + congr (oapp _ _); apply: eq_pick => a; rewrite memJ_norm //. + apply: subsetP Gy; exact: class_support_norm. +case: pickP => // a /andP[Aa Ha_u]. +by rewrite (subsetP Dade_support_sub) // in notGx; apply/bigcupP; exists a. +Qed. +Definition Dade alpha := Cfun 1 (Dade_subproof alpha). + +Lemma Dade_is_linear : linear Dade. +Proof. +move=> mu alpha beta; apply/cfunP=> x; rewrite !cfunElock. +by case: pickP => [a _ | _] /=; rewrite ?mulr0 ?addr0 ?cfunE. +Qed. +Canonical Dade_additive := Additive Dade_is_linear. +Canonical Dade_linear := Linear Dade_is_linear. + +Local Notation "alpha ^\tau" := (Dade alpha). + +(* This is the validity of Peterfalvi, Definition (2.5) *) +Lemma DadeE alpha a u : a \in A -> u \in dd1 a -> alpha^\tau u = alpha a. +Proof. +move=> Aa Ha_u; rewrite cfunElock. +have [b /= /andP[Ab Hb_u] | ] := pickP; last by move/(_ a); rewrite Aa Ha_u. +have [|x Lx ->] := Dade_support1_TI Aa Ab; last by rewrite cfunJ. +by apply/pred0Pn; exists u; rewrite /= Ha_u. +Qed. + +Lemma Dade_id alpha : {in A, forall a, alpha^\tau a = alpha a}. +Proof. +by move=> a Aa; rewrite /= -{1}[a]mul1g (DadeE _ Aa) ?mem_Dade_support1. +Qed. + +Lemma Dade_cfunS alpha : alpha^\tau \in 'CF(G, Atau). +Proof. +apply/cfun_onP=> x; rewrite cfunElock. +by case: pickP => [a /andP[Aa Ha_x] /bigcupP[] | //]; exists a. +Qed. + +Lemma Dade_cfun alpha : alpha^\tau \in 'CF(G, G^#). +Proof. by rewrite (cfun_onS Dade_support_subD1) ?Dade_cfunS. Qed. + +Lemma Dade1 alpha : alpha^\tau 1%g = 0. +Proof. by rewrite (cfun_on0 (Dade_cfun _)) // !inE eqxx. Qed. + +Lemma Dade_id1 : + {in 'CF(L, A) & 1%g |: A, forall alpha a, alpha^\tau a = alpha a}. +Proof. +move=> alpha a Aalpha; case/setU1P=> [-> |]; last exact: Dade_id. +by rewrite Dade1 (cfun_on0 Aalpha). +Qed. + +Section AutomorphismCFun. + +Variable u : {rmorphism algC -> algC}. +Local Notation "alpha ^u" := (cfAut u alpha). + +Lemma Dade_aut alpha : (alpha^u)^\tau = (alpha^\tau)^u. +Proof. +apply/cfunP => g; rewrite cfunE. +have [/bigcupP[a Aa A1g] | notAtau_g] := boolP (g \in Atau). + by rewrite !(DadeE _ Aa A1g) cfunE. +by rewrite !(cfun_on0 _ notAtau_g) ?rmorph0 ?Dade_cfunS. +Qed. + +End AutomorphismCFun. + +Lemma Dade_conjC alpha : (alpha^*)^\tau = ((alpha^\tau)^*)%CF. +Proof. exact: Dade_aut. Qed. + +(* This is Peterfalvi (2.7), main part *) +Lemma general_Dade_reciprocity alpha (phi : 'CF(G)) (psi : 'CF(L)) : + alpha \in 'CF(L, A) -> + {in A, forall a, psi a = #|H a|%:R ^-1 * (\sum_(x in H a) phi (x * a)%g)} -> + '[alpha^\tau, phi] = '[alpha, psi]. +Proof. +move=> CFalpha psiA; rewrite (cfdotEl _ (Dade_cfunS _)). +pose T := [set repr (a ^: L) | a in A]. +have sTA: {subset T <= A}. + move=> _ /imsetP[a Aa ->]; have [x Lx ->] := repr_class L a. + by rewrite memJ_norm ?(subsetP nAL). +pose P_G := [set dd1 x | x in T]. +have dd1_id: {in A, forall a, dd1 (repr (a ^: L)) = dd1 a}. + by move=> a Aa /=; have [x Lx ->] := repr_class L a; apply: Dade_support1_id. +have ->: Atau = cover P_G. + apply/setP=> u; apply/bigcupP/bigcupP=> [[a Aa Fa_u] | [Fa]]; last first. + by case/imsetP=> a /sTA Aa -> Fa_u; exists a. + by exists (dd1 a) => //; rewrite -dd1_id //; do 2!apply: mem_imset. +have [tiP_G inj_dd1]: trivIset P_G /\ {in T &, injective dd1}. + apply: trivIimset => [_ _ /imsetP[a Aa ->] /imsetP[b Ab ->] |]; last first. + apply/imsetP=> [[a]]; move/sTA=> Aa; move/esym; move/eqP; case/set0Pn. + by exists (a ^ 1)%g; apply: mem_imset2; rewrite ?group1 ?rcoset_refl. + rewrite !dd1_id //; apply: contraR. + by case/Dade_support1_TI=> // x Lx ->; rewrite classGidl. +rewrite big_trivIset //= big_imset {P_G tiP_G inj_dd1}//=. +symmetry; rewrite (cfdotEl _ CFalpha). +pose P_A := [set a ^: L | a in T]. +have rLid x: repr (x ^: L) ^: L = x ^: L. + by have [y Ly ->] := repr_class L x; rewrite classGidl. +have {1}<-: cover P_A = A. + apply/setP=> a; apply/bigcupP/idP=> [[_ /imsetP[d /sTA Ab ->]] | Aa]. + by case/imsetP=> x Lx ->; rewrite memJ_norm ?(subsetP nAL). + by exists (a ^: L); rewrite ?class_refl // -rLid; do 2!apply: mem_imset. +have [tiP_A injFA]: trivIset P_A /\ {in T &, injective (class^~ L)}. + apply: trivIimset => [_ _ /imsetP[a Aa ->] /imsetP[b Ab ->] |]; last first. + by apply/imsetP=> [[a _ /esym/eqP/set0Pn[]]]; exists a; exact: class_refl. + rewrite !rLid; apply: contraR => /pred0Pn[c /andP[/=]]. + by do 2!move/class_transr <-. +rewrite big_trivIset //= big_imset {P_A tiP_A injFA}//=. +apply: canRL (mulKf (neq0CG G)) _; rewrite mulrA big_distrr /=. +apply: eq_bigr => a /sTA=> {T sTA}Aa. +have [La def_Ca] := (subsetP sAL a Aa, defCA Aa). +rewrite (eq_bigr (fun _ => alpha a * (psi a)^*)) => [|ax]; last first. + by case/imsetP=> x Lx ->{ax}; rewrite !cfunJ. +rewrite sumr_const -index_cent1 mulrC -mulr_natr -!mulrA. +rewrite (eq_bigr (fun xa => alpha a * (phi xa)^*)) => [|xa Fa_xa]; last first. + by rewrite (DadeE _ Aa). +rewrite -big_distrr /= -rmorph_sum; congr (_ * _). +rewrite mulrC mulrA -natrM mulnC -(Lagrange (subsetIl G 'C[a])). +rewrite -mulnA mulnCA -(sdprod_card def_Ca) -mulnA Lagrange ?subsetIl //. +rewrite mulnA natrM mulfK ?neq0CG // -conjC_nat -rmorphM; congr (_ ^*). +have /and3P[_ tiHa _] := Dade_cover_TI Aa. +rewrite (set_partition_big _ (partition_class_support _ _)) //=. +rewrite (eq_bigr (fun _ => \sum_(x in H a) phi (x * a)%g)); last first. + move=> _ /imsetP[x Gx ->]; rewrite -rcosetE. + rewrite (big_imset _ (in2W (conjg_inj x))) (big_imset _ (in2W (mulIg a))) /=. + by apply: eq_bigr => u Hu; rewrite cfunJ ?groupM ?(subsetP sLG a). +rewrite sumr_const card_orbit astab1Js norm_Dade_cover //. +by rewrite natrM -mulrA mulr_natl psiA // mulVKf ?neq0CG. +Qed. + +(* This is Peterfalvi (2.7), second part. *) +Lemma Dade_reciprocity alpha (phi : 'CF(G)) : + alpha \in 'CF(L, A) -> + {in A, forall a, {in H a, forall u, phi (u * a)%g = phi a}} -> + '[alpha^\tau, phi] = '[alpha, 'Res[L] phi]. +Proof. +move=> CFalpha phiH; apply: general_Dade_reciprocity => // a Aa. +rewrite cfResE ?(subsetP sAL) //; apply: canRL (mulKf (neq0CG _)) _. +by rewrite mulr_natl -sumr_const; apply: eq_bigr => x Hx; rewrite phiH. +Qed. + +(* This is Peterfalvi (2.6)(a). *) +Lemma Dade_isometry : {in 'CF(L, A) &, isometry Dade}. +Proof. +move=> alpha beta CFalpha CFbeta /=. +rewrite Dade_reciprocity ?Dade_cfun => // [|a Aa u Hu]; last first. + by rewrite (DadeE _ Aa) ?mem_Dade_support1 ?Dade_id. +rewrite !(cfdotEl _ CFalpha); congr (_ * _); apply: eq_bigr => x Ax. +by rewrite cfResE ?(subsetP sAL) // Dade_id. +Qed. + +(* Supplement to Peterfalvi (2.3)/(2.6)(a); implies Isaacs Lemma 7.7. *) +Lemma Dade_Ind : normedTI A G L -> {in 'CF(L, A), Dade =1 'Ind}. +Proof. +case/Dade_normedTI_P=> _ trivH alpha Aalpha. +rewrite [alpha^\tau]cfun_sum_cfdot ['Ind _]cfun_sum_cfdot. +apply: eq_bigr => i _; rewrite -cfdot_Res_r -Dade_reciprocity // => a Aa /= u. +by rewrite trivH // => /set1P->; rewrite mul1g. +Qed. + +Definition Dade_set_signalizer (B : {set gT}) := \bigcap_(a in B) H a. +Local Notation "''H' ( B )" := (Dade_set_signalizer B) + (at level 8, format "''H' ( B )") : group_scope. +Canonical Dade_set_signalizer_group B := [group of 'H(B)]. +Definition Dade_set_normalizer B := 'H(B) <*> 'N_L(B). +Local Notation "''M' ( B )" := (Dade_set_normalizer B) + (at level 8, format "''M' ( B )") : group_scope. +Canonical Dade_set_normalizer_group B := [group of 'M(B)]. + +Let calP := [set B : {set gT} | B \subset A & B != set0]. + +(* This is Peterfalvi (2.8). *) +Lemma Dade_set_sdprod : {in calP, forall B, 'H(B) ><| 'N_L(B) = 'M(B)}. +Proof. +move=> B /setIdP[sBA notB0]; apply: sdprodEY => /=. + apply/subsetP=> x /setIP[Lx nBx]; rewrite inE. + apply/bigcapsP=> a Ba; have Aa := subsetP sBA a Ba. + by rewrite sub_conjg -DadeJ ?groupV // bigcap_inf // memJ_norm ?groupV. +have /set0Pn[a Ba] := notB0; have Aa := subsetP sBA a Ba. +have [_ /mulG_sub[sHaC _] _ tiHaL] := sdprodP (defCA Aa). +rewrite -(setIidPl sLG) -setIA setICA (setIidPl sHaC) in tiHaL. +by rewrite setICA ['H(B)](bigD1 a) //= !setIA tiHaL !setI1g. +Qed. + +Section DadeExpansion. + +Variable aa : 'CF(L). +Hypothesis CFaa : aa \in 'CF(L, A). + +Definition Dade_restrm B := + if B \in calP then remgr 'H(B) 'N_L(B) else trivm 'M(B). +Fact Dade_restrM B : {in 'M(B) &, {morph Dade_restrm B : x y / x * y}%g}. +Proof. +rewrite /Dade_restrm; case: ifP => calP_B; last exact: morphM. +have defM := Dade_set_sdprod calP_B; have [nsHM _ _ _ _] := sdprod_context defM. +by apply: remgrM; first exact: sdprod_compl. +Qed. +Canonical Dade_restr_morphism B := Morphism (@Dade_restrM B). +Definition Dade_cfun_restriction B := + cfMorph ('Res[Dade_restrm B @* 'M(B)] aa). + +Local Notation "''aa_' B" := (Dade_cfun_restriction B) + (at level 3, B at level 2, format "''aa_' B") : ring_scope. + +Definition Dade_transversal := [set repr (B :^: L) | B in calP]. +Local Notation calB := Dade_transversal. + +Lemma Dade_restrictionE B x : + B \in calP -> 'aa_B x = aa (remgr 'H(B) 'N_L(B) x) *+ (x \in 'M(B)). +Proof. +move=> calP_B; have /sdprodP[_ /= defM _ _] := Dade_set_sdprod calP_B. +have [Mx | /cfun0-> //] := boolP (x \in 'M(B)). +rewrite mulrb cfMorphE // morphimEdom /= /Dade_restrm calP_B. +rewrite cfResE ?mem_imset {x Mx}//= -defM. +by apply/subsetP=> _ /imsetP[x /mem_remgr/setIP[Lx _] ->]. +Qed. +Local Notation rDadeE := Dade_restrictionE. + +Lemma Dade_restriction_vchar B : aa \in 'Z[irr L] -> 'aa_B \in 'Z[irr 'M(B)]. +Proof. +rewrite /'aa_B => /vcharP[a1 Na1 [a2 Na2 ->]]. +by rewrite !linearB /= rpredB // char_vchar ?cfMorph_char ?cfRes_char. +Qed. + +Let sMG B : B \in calP -> 'M(B) \subset G. +Proof. +case/setIdP=> /subsetP sBA /set0Pn[a Ba]. +by rewrite join_subG ['H(B)](bigD1 a Ba) !subIset ?sLG ?sHG ?sBA. +Qed. + +(* This is Peterfalvi (2.10.1) *) +Lemma Dade_Ind_restr_J : + {in L & calP, forall x B, 'Ind[G] 'aa_(B :^ x) = 'Ind[G] 'aa_B}. +Proof. +move=> x B Lx dB; have [defMB [sBA _]] := (Dade_set_sdprod dB, setIdP dB). +have dBx: B :^ x \in calP. + by rewrite !inE -{2}(normsP nAL x Lx) conjSg -!cards_eq0 cardJg in dB *. +have defHBx: 'H(B :^ x) = 'H(B) :^ x. + rewrite /'H(_) (big_imset _ (in2W (conjg_inj x))) -bigcapJ /=. + by apply: eq_bigr => a Ba; rewrite DadeJ ?(subsetP sBA). +have defNBx: 'N_L(B :^ x) = 'N_L(B) :^ x by rewrite conjIg -normJ (conjGid Lx). +have [_ mulHNB _ tiHNB] := sdprodP defMB. +have defMBx: 'M(B :^ x) = 'M(B) :^ x. + rewrite -mulHNB conjsMg -defHBx -defNBx. + by case/sdprodP: (Dade_set_sdprod dBx). +have def_aa_x y: 'aa_(B :^ x) (y ^ x) = 'aa_B y. + rewrite !rDadeE // defMBx memJ_conjg !mulrb -mulHNB defHBx defNBx. + have [[h z Hh Nz ->] | // ] := mulsgP. + by rewrite conjMg !remgrMid ?cfunJ ?memJ_conjg // -conjIg tiHNB conjs1g. +apply/cfunP=> y; have Gx := subsetP sLG x Lx. +rewrite [eq]lock !cfIndE ?sMG //= {1}defMBx cardJg -lock; congr (_ * _). +rewrite (reindex_astabs 'R x) ?astabsR //=. +by apply: eq_bigr => z _; rewrite conjgM def_aa_x. +Qed. + +(* This is Peterfalvi (2.10.2) *) +Lemma Dade_setU1 : {in calP & A, forall B a, 'H(a |: B) = 'C_('H(B))[a]}. +Proof. +move=> B a dB Aa; rewrite /'H(_) bigcap_setU big_set1 -/'H(B). +apply/eqP; rewrite setIC eqEsubset setIS // subsetI subsetIl /=. +have [sHBG pi'HB]: 'H(B) \subset G /\ pi^'.-group 'H(B). + have [sBA /set0Pn[b Bb]] := setIdP dB; have Ab := subsetP sBA b Bb. + have sHBb: 'H(B) \subset H b by rewrite ['H(B)](bigD1 b) ?subsetIl. + by rewrite (pgroupS sHBb) ?pi'H ?(subset_trans sHBb) ?sHG. +have [nsHa _ defCa _ _] := sdprod_context (defCA Aa). +have [hallHa _] := coprime_mulGp_Hall defCa (pi'H a) (piCL Aa). +by rewrite (sub_normal_Hall hallHa) ?(pgroupS (subsetIl _ _)) ?setSI. +Qed. + +Let calA g (X : {set gT}) := [set x in G | g ^ x \in X]%g. + +(* This is Peterfalvi (2.10.3) *) +Lemma Dade_Ind_expansion B g : + B \in calP -> + [/\ g \notin Atau -> ('Ind[G, 'M(B)] 'aa_B) g = 0 + & {in A, forall a, g \in dd1 a -> + ('Ind[G, 'M(B)] 'aa_B) g = + (aa a / #|'M(B)|%:R) * + \sum_(b in 'N_L(B) :&: a ^: L) #|calA g ('H(B) :* b)|%:R}]. +Proof. +move=> dB; set LHS := 'Ind _ g. +have defMB := Dade_set_sdprod dB; have [_ mulHNB nHNB tiHNB] := sdprodP defMB. +have [sHMB sNMB] := mulG_sub mulHNB. +have{LHS} ->: LHS = #|'M(B)|%:R^-1 * \sum_(x in calA g 'M(B)) 'aa_B (g ^ x). + rewrite {}/LHS cfIndE ?sMG //; congr (_ * _). + rewrite (bigID [pred x | g ^ x \in 'M(B)]) /= addrC big1 ?add0r => [|x]. + by apply: eq_bigl => x; rewrite inE. + by case/andP=> _ notMgx; rewrite cfun0. +pose fBg x := remgr 'H(B) 'N_L(B) (g ^ x). +pose supp_aBg := [pred b in A | g \in dd1 b]. +have supp_aBgP: {in calA g 'M(B), forall x, + ~~ supp_aBg (fBg x) -> 'aa_B (g ^ x)%g = 0}. +- move=> x /setIdP[]; set b := fBg x => Gx MBgx notHGx; rewrite rDadeE // MBgx. + have Nb: b \in 'N_L(B) by rewrite mem_remgr ?mulHNB. + have Cb: b \in 'C_L[b] by rewrite inE cent1id; have [-> _] := setIP Nb. + rewrite (cfun_on0 CFaa) // -/(fBg x) -/b; apply: contra notHGx => Ab. + have nHb: b \in 'N('H(B)) by rewrite (subsetP nHNB). + have [sBA /set0Pn[a Ba]] := setIdP dB; have Aa := subsetP sBA a Ba. + have [|/= partHBb _] := partition_cent_rcoset nHb. + rewrite (coprime_dvdr (order_dvdG Cb)) //= ['H(B)](bigD1 a) //=. + by rewrite (coprimeSg (subsetIl _ _)) ?coHL. + have Hb_gx: g ^ x \in 'H(B) :* b by rewrite mem_rcoset mem_divgr ?mulHNB. + have [defHBb _ _] := and3P partHBb; rewrite -(eqP defHBb) in Hb_gx. + case/bigcupP: Hb_gx => Cy; case/imsetP=> y HBy ->{Cy} Cby_gx. + have sHBa: 'H(B) \subset H a by rewrite bigcap_inf. + have sHBG: 'H(B) \subset G := subset_trans sHBa (sHG a). + rewrite Ab -(memJ_conjg _ x) class_supportGidr // -(conjgKV y (g ^ x)). + rewrite mem_imset2 // ?(subsetP sHBG) {HBy}// -mem_conjg. + apply: subsetP Cby_gx; rewrite {y}conjSg mulSg //. + have [nsHb _ defCb _ _] := sdprod_context (defCA Ab). + have [hallHb _] := coprime_mulGp_Hall defCb (pi'H b) (piCL Ab). + rewrite (sub_normal_Hall hallHb) ?setSI // (pgroupS _ (pi'H a)) //=. + by rewrite subIset ?sHBa. +split=> [notHGg | a Aa Hag]. + rewrite big1 ?mulr0 // => x; move/supp_aBgP; apply; set b := fBg x. + by apply: contra notHGg; case/andP=> Ab Hb_x; apply/bigcupP; exists b. +rewrite -mulrA mulrCA; congr (_ * _); rewrite big_distrr /=. +set nBaL := _ :&: _; rewrite (bigID [pred x | fBg x \in nBaL]) /= addrC. +rewrite big1 ?add0r => [|x /andP[calAx not_nBaLx]]; last first. + apply: supp_aBgP => //; apply: contra not_nBaLx. + set b := fBg x => /andP[Ab Hb_g]; have [Gx MBx] := setIdP calAx. + rewrite inE mem_remgr ?mulHNB //; apply/imsetP/Dade_support1_TI => //. + by apply/pred0Pn; exists g; exact/andP. +rewrite (partition_big fBg (mem nBaL)) /= => [|x]; last by case/andP. +apply: eq_bigr => b; case/setIP=> Nb aLb; rewrite mulr_natr -sumr_const. +apply: eq_big => x; rewrite ![x \in _]inE -!andbA. + apply: andb_id2l=> Gx; apply/and3P/idP=> [[Mgx _] /eqP <- | HBb_gx]. + by rewrite mem_rcoset mem_divgr ?mulHNB. + suffices ->: fBg x = b. + by rewrite inE Nb (subsetP _ _ HBb_gx) // -mulHNB mulgS ?sub1set. + by rewrite /fBg; have [h Hh ->] := rcosetP HBb_gx; exact: remgrMid. +move/and4P=> [_ Mgx _ /eqP def_fx]. +rewrite rDadeE // Mgx -/(fBg x) def_fx; case/imsetP: aLb => y Ly ->. +by rewrite cfunJ // (subsetP sAL). +Qed. + +(* This is Peterfalvi (2.10) *) +Lemma Dade_expansion : + aa^\tau = - \sum_(B in calB) (- 1) ^+ #|B| *: 'Ind[G, 'M(B)] 'aa_B. +Proof. +apply/cfunP=> g; rewrite !cfunElock sum_cfunE. +pose n1 (B : {set gT}) : algC := (-1) ^+ #|B| / #|L : 'N_L(B)|%:R. +pose aa1 B := ('Ind[G, 'M(B)] 'aa_B) g. +have dBJ: {acts L, on calP | 'Js}. + move=> x Lx /= B; rewrite !inE -!cards_eq0 cardJg. + by rewrite -{1}(normsP nAL x Lx) conjSg. +transitivity (- (\sum_(B in calP) n1 B * aa1 B)); last first. + congr (- _); rewrite {1}(partition_big_imset (fun B => repr (B :^: L))) /=. + apply: eq_bigr => B /imsetP[B1 dB1 defB]. + have B1L_B: B \in B1 :^: L by rewrite defB (mem_repr B1) ?orbit_refl. + have{dB1} dB1L: {subset B1 :^: L <= calP}. + by move=> _ /imsetP[x Lx ->]; rewrite dBJ. + have dB: B \in calP := dB1L B B1L_B. + rewrite (eq_bigl (mem (B :^: L))) => [|B2 /=]; last first. + apply/andP/idP=> [[_ /eqP <-] | /(orbit_trans B1L_B) B1L_B2]. + by rewrite orbit_sym (mem_repr B2) ?orbit_refl. + by rewrite [B2 :^: L](orbit_transl B1L_B2) -defB dB1L. + rewrite (eq_bigr (fun _ => n1 B * aa1 B)) => [|_ /imsetP[x Lx ->]]. + rewrite cfunE sumr_const -mulr_natr mulrAC card_orbit astab1Js divfK //. + by rewrite pnatr_eq0 -lt0n indexg_gt0. + rewrite /aa1 Dade_Ind_restr_J //; congr (_ * _). + by rewrite /n1 cardJg -{1 2}(conjGid Lx) normJ -conjIg indexJg. +case: pickP => /= [a /andP[Aa Ha_g] | notHAg]; last first. + rewrite big1 ?oppr0 // /aa1 => B dB. + have [->] := Dade_Ind_expansion g dB; first by rewrite mulr0. + by apply/bigcupP=> [[a Aa Ha_g]]; case/andP: (notHAg a). +pose P_ b := [set B in calP | b \in 'N_L(B)]. +pose aa2 B b : algC := #|calA g ('H(B) :* b)|%:R. +pose nn2 (B : {set gT}) : algC := (-1) ^+ #|B| / #|'H(B)|%:R. +pose sumB b := \sum_(B in P_ b) nn2 B * aa2 B b. +transitivity (- aa a / #|L|%:R * \sum_(b in a ^: L) sumB b); last first. + rewrite !mulNr; congr (- _). + rewrite (exchange_big_dep (mem calP)) => [|b B _] /=; last by case/setIdP. + rewrite big_distrr /aa1; apply: eq_bigr => B dB; rewrite -big_distrr /=. + have [_ /(_ a) -> //] := Dade_Ind_expansion g dB; rewrite !mulrA. + congr (_ * _); last by apply: eq_bigl => b; rewrite inE dB /= andbC -in_setI. + rewrite -mulrA mulrCA -!mulrA; congr (_ * _). + rewrite -invfM mulrCA -invfM -!natrM; congr (_ / _%:R). + rewrite -(sdprod_card (Dade_set_sdprod dB)) mulnA mulnAC; congr (_ * _)%N. + by rewrite mulnC Lagrange ?subsetIl. +rewrite (eq_bigr (fun _ => sumB a)) /= => [|_ /imsetP[x Lx ->]]; last first. + rewrite {1}/sumB (reindex_inj (@conjsg_inj _ x)) /=. + symmetry; apply: eq_big => B. + rewrite ![_ \in P_ _]inE dBJ //. + by rewrite -{2}(conjGid Lx) normJ -conjIg memJ_conjg. + case/setIdP=> dB Na; have [sBA _] := setIdP dB. + have defHBx: 'H(B :^ x) = 'H(B) :^ x. + rewrite /'H(_) (big_imset _ (in2W (conjg_inj x))) -bigcapJ /=. + by apply: eq_bigr => b Bb; rewrite DadeJ ?(subsetP sBA). + rewrite /nn2 /aa2 defHBx !cardJg; congr (_ * _%:R). + rewrite -(card_rcoset _ x); apply: eq_card => y. + rewrite !(inE, mem_rcoset, mem_conjg) conjMg conjVg conjgK -conjgM. + by rewrite groupMr // groupV (subsetP sLG). +rewrite sumr_const mulrC [sumB a](bigD1 [set a]) /=; last first. + by rewrite 3!inE cent1id sub1set Aa -cards_eq0 cards1 (subsetP sAL). +rewrite -[_ *+ _]mulr_natr -mulrA mulrDl -!mulrA ['H(_)]big_set1 cards1. +have ->: aa2 [set a] a = #|'C_G[a]|%:R. + have [u x Ha_ux Gx def_g] := imset2P Ha_g. + rewrite -(card_lcoset _ x^-1); congr _%:R; apply: eq_card => y. + rewrite ['H(_)]big_set1 mem_lcoset invgK inE def_g -conjgM. + rewrite -(groupMl y Gx) inE; apply: andb_id2l => Gxy. + by have [_ _ -> //] := normedTI_memJ_P (Dade_cover_TI Aa); rewrite inE Gxy. +rewrite mulN1r mulrC mulrA -natrM -(sdprod_card (defCA Aa)). +rewrite -mulnA card_orbit astab1J Lagrange ?subsetIl // mulnC natrM. +rewrite mulrAC mulfK ?neq0CG // mulrC divfK ?neq0CG // opprK. +rewrite (bigID [pred B : {set gT} | a \in B]) /= mulrDl addrA. +apply: canRL (subrK _) _; rewrite -mulNr -sumrN; congr (_ + _ * _). +symmetry. +rewrite (reindex_onto (fun B => a |: B) (fun B => B :\ a)) /=; last first. + by move=> B; case/andP=> _; exact: setD1K. +symmetry; apply: eq_big => B. + rewrite setU11 andbT -!andbA; apply/and3P/and3P; case. + do 2![case/setIdP] => sBA ntB /setIP[La nBa] _ notBa. + rewrite 3!inE subUset sub1set Aa sBA La setU1K // -cards_eq0 cardsU1 notBa. + rewrite -sub1set normsU ?sub1set ?cent1id //= eq_sym eqEcard subsetUl /=. + by rewrite cards1 cardsU1 notBa ltnS leqn0 cards_eq0. + do 2![case/setIdP] => /subUsetP[_ sBA] _ /setIP[La]. + rewrite inE conjUg (normP (cent1id a)) => /subUsetP[_ sBa_aB]. + rewrite eq_sym eqEcard subsetUl cards1 (cardsD1 a) setU11 ltnS leqn0 /=. + rewrite cards_eq0 => notB0 /eqP defB. + have notBa: a \notin B by rewrite -defB setD11. + split=> //; last by apply: contraNneq notBa => ->; exact: set11. + rewrite !inE sBA La -{1 3}defB notB0 subsetD1 sBa_aB. + by rewrite mem_conjg /(a ^ _) invgK mulgA mulgK. +do 2![case/andP] => /setIdP[dB Na] _ notBa. +suffices ->: aa2 B a = #|'H(B) : 'H(a |: B)|%:R * aa2 (a |: B) a. + rewrite /nn2 cardsU1 notBa exprS mulN1r !mulNr; congr (- _). + rewrite !mulrA; congr (_ * _); rewrite -!mulrA; congr (_ * _). + apply: canLR (mulKf (neq0CG _)) _; apply: canRL (mulfK (neq0CG _)) _ => /=. + by rewrite -natrM mulnC Lagrange //= Dade_setU1 ?subsetIl. +rewrite /aa2 Dade_setU1 //= -natrM; congr _%:R. +have defMB := Dade_set_sdprod dB; have [_ mulHNB nHNB tiHNB] := sdprodP defMB. +have [sHMB sNMB] := mulG_sub mulHNB; have [La nBa] := setIP Na. +have nHa: a \in 'N('H(B)) by rewrite (subsetP nHNB). +have Ca: a \in 'C_L[a] by rewrite inE cent1id La. +have [|/= partHBa nbHBa] := partition_cent_rcoset nHa. + have [sBA] := setIdP dB; case/set0Pn=> b Bb; have Ab := subsetP sBA b Bb. + rewrite (coprime_dvdr (order_dvdG Ca)) //= ['H(B)](bigD1 b) //=. + by rewrite (coprimeSg (subsetIl _ _)) ?coHL. +pose pHBa := mem ('H(B) :* a). +rewrite -sum1_card (partition_big (fun x => g ^ x) pHBa) /= => [|x]; last first. + by case/setIdP=> _ ->. +rewrite (set_partition_big _ partHBa) /= -nbHBa -sum_nat_const. +apply: eq_bigr => _ /imsetP[x Hx ->]. +rewrite (big_imset _ (in2W (conjg_inj x))) /=. +rewrite -(card_rcoset _ x) -sum1_card; symmetry; set HBaa := 'C_(_)[a] :* a. +rewrite (partition_big (fun y => g ^ (y * x^-1)) (mem HBaa)); last first. + by move=> y; rewrite mem_rcoset => /setIdP[]. +apply: eq_bigr => /= u Ca_u; apply: eq_bigl => z. +rewrite -(canF_eq (conjgKV x)) -conjgM; apply: andb_id2r; move/eqP=> def_u. +have sHBG: 'H(B) \subset G. + have [sBA /set0Pn[b Bb]] := setIdP dB; have Ab := subsetP sBA b Bb. + by rewrite (bigcap_min b) ?sHG. +rewrite mem_rcoset !inE groupMr ?groupV ?(subsetP sHBG x Hx) //=. +congr (_ && _); have [/eqP defHBa _ _] := and3P partHBa. +symmetry; rewrite def_u Ca_u -defHBa -(mulgKV x z) conjgM def_u -/HBaa. +by rewrite cover_imset -class_supportEr mem_imset2. +Qed. + +End DadeExpansion. + +(* This is Peterfalvi (2.6)(b) *) +Lemma Dade_vchar alpha : alpha \in 'Z[irr L, A] -> alpha^\tau \in 'Z[irr G]. +Proof. +rewrite [alpha \in _]zchar_split => /andP[Zaa CFaa]. +rewrite Dade_expansion // rpredN rpred_sum // => B dB. +suffices calP_B: B \in calP. + by rewrite rpredZsign cfInd_vchar // Dade_restriction_vchar. +case/imsetP: dB => B0 /setIdP[sB0A notB00] defB. +have [x Lx ->]: exists2 x, x \in L & B = B0 :^ x. + by apply/imsetP; rewrite defB (mem_repr B0) ?orbit_refl. +by rewrite inE -cards_eq0 cardJg cards_eq0 -(normsP nAL x Lx) conjSg sB0A. +Qed. + +(* This summarizes Peterfalvi (2.6). *) +Lemma Dade_Zisometry : {in 'Z[irr L, A], isometry Dade, to 'Z[irr G, G^#]}. +Proof. +split; first by apply: sub_in2 Dade_isometry; exact: zchar_on. +by move=> phi Zphi; rewrite /= zchar_split Dade_vchar ?Dade_cfun. +Qed. + +End Two. + +Section RestrDade. + +Variables (gT : finGroupType) (G L : {group gT}) (A A1 : {set gT}). +Hypothesis ddA : Dade_hypothesis G L A. +Hypotheses (sA1A : A1 \subset A) (nA1L : L \subset 'N(A1)). +Let ssA1A := subsetP sA1A. + +(* This is Peterfalvi (2.11), first part. *) +Lemma restr_Dade_hyp : Dade_hypothesis G L A1. +Proof. +have [/andP[sAL nAL] notA_1 sLG conjAG [H defCa coHL]] := ddA. +have nsA1L: A1 <| L by rewrite /normal (subset_trans sA1A). +split; rewrite ?(contra (@ssA1A _)) //; first exact: sub_in2 conjAG. +by exists H; [exact: sub_in1 defCa | exact: sub_in2 coHL]. +Qed. +Local Notation ddA1 := restr_Dade_hyp. + +Local Notation H dd := (Dade_signalizer dd). +Lemma restr_Dade_signalizer H1 : {in A, H ddA =1 H1} -> {in A1, H ddA1 =1 H1}. +Proof. +move=> defH1; apply: def_Dade_signalizer => a /ssA1A Aa. +by rewrite -defH1 ?Dade_sdprod. +Qed. + +Lemma restr_Dade_support1 : {in A1, Dade_support1 ddA1 =1 Dade_support1 ddA}. +Proof. +by move=> a A1a; rewrite /Dade_support1 (@restr_Dade_signalizer (H ddA)). +Qed. + +Lemma restr_Dade_support : + Dade_support ddA1 = \bigcup_(a in A1) Dade_support1 ddA a. +Proof. by rewrite -(eq_bigr _ restr_Dade_support1). Qed. + +Definition restr_Dade := Dade ddA1. + +(* This is Peterfalvi (2.11), second part. *) +Lemma restr_DadeE : {in 'CF(L, A1), restr_Dade =1 Dade ddA}. +Proof. +move=> aa CF1aa; apply/cfunP=> g; rewrite cfunElock. +have CFaa: aa \in 'CF(L, A) := cfun_onS sA1A CF1aa. +have [a /= /andP[A1a Ha_g] | no_a /=] := pickP. + by apply/esym/DadeE; rewrite -1?restr_Dade_support1 ?ssA1A. +rewrite cfunElock; case: pickP => //= a /andP[_ Ha_g]. +rewrite (cfun_on0 CF1aa) //; apply: contraFN (no_a a) => A1a. +by rewrite A1a restr_Dade_support1. +Qed. + +End RestrDade. + +Section NormedTI. + +Variables (gT : finGroupType) (G L : {group gT}) (A : {set gT}). +Hypotheses (tiAG : normedTI A G L) (sAG1 : A \subset G^#). + +(* This is the existence part of Peterfalvi (2.3). *) +Lemma normedTI_Dade : Dade_hypothesis G L A. +Proof. +have [[sAG notA1] [_ _ /eqP defL]] := (subsetD1P sAG1, and3P tiAG). +have [_ sLG tiAG_L] := normedTI_memJ_P tiAG. +split=> // [|a b Aa Ab /imsetP[x Gx def_b]|]. +- rewrite /(A <| L) -{2}defL subsetIr andbT; apply/subsetP=> a Aa. + by rewrite -(tiAG_L a) ?(subsetP sAG) // conjgE mulKg. +- by rewrite def_b mem_imset // -(tiAG_L a) -?def_b. +exists (fun _ => 1%G) => [a Aa | a b _ _]; last by rewrite cards1 coprime1n. +by rewrite sdprod1g -(setIidPl sLG) -setIA (setIidPr (cent1_normedTI tiAG Aa)). +Qed. + +Let def_ddA := Dade_Ind normedTI_Dade tiAG. + +(* This is the identity part of Isaacs, Lemma 7.7. *) +Lemma normedTI_Ind_id1 : + {in 'CF(L, A) & 1%g |: A, forall alpha, 'Ind[G] alpha =1 alpha}. +Proof. by move=> aa a CFaa A1a; rewrite /= -def_ddA // Dade_id1. Qed. + +(* A more restricted, but more useful form. *) +Lemma normedTI_Ind_id : + {in 'CF(L, A) & A, forall alpha, 'Ind[G] alpha =1 alpha}. +Proof. by apply: sub_in11 normedTI_Ind_id1 => //; apply/subsetP/subsetUr. Qed. + +(* This is the isometry part of Isaacs, Lemma 7.7. *) +(* The statement in Isaacs is slightly more general in that it allows for *) +(* beta \in 'CF(L, 1%g |: A); this appears to be more cumbersome than useful. *) +Lemma normedTI_isometry : {in 'CF(L, A) &, isometry 'Ind[G]}. +Proof. by move=> aa bb CFaa CFbb; rewrite /= -!def_ddA // Dade_isometry. Qed. + +End NormedTI.
\ No newline at end of file diff --git a/mathcomp/odd_order/PFsection3.v b/mathcomp/odd_order/PFsection3.v new file mode 100644 index 0000000..063aacc --- /dev/null +++ b/mathcomp/odd_order/PFsection3.v @@ -0,0 +1,1854 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg matrix poly finset. +Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +Require Import gfunctor center gproduct cyclic pgroup abelian frobenius. +Require Import mxalgebra mxrepresentation vector falgebra fieldext galois. +Require Import ssrnum rat algC algnum classfun character. +Require Import integral_char inertia vcharacter. +Require Import PFsection1 PFsection2. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 3: TI-Subsets with Cyclic Normalizers *) +(******************************************************************************) +(* Given a direct product decomposition defW : W1 \x W2 = W, we define here: *) +(* cyclicTIset defW == the set complement of W1 and W2 in W; this *) +(* (locally) V definition is usually Let-bound to V. *) +(* := W :\: (W1 :|: W2). *) +(* cyclicTI_hypothesis G defW <-> W is a cyclic of odd order that is the *) +(* normaliser in G of its non-empty TI subset *) +(* V = cyclicTIset defW = W :\: (W1 :|: W2). *) +(* -> This is Peterfalvi, Hypothesis (3.1), or Feit-Thompson (13.1). *) +(* cyclicTIirr defW i j == the irreducible character of W coinciding with *) +(* (locally) w_ i j chi_i and 'chi_j on W1 and W2, respectively. *) +(* This notation is usually Let-bound to w_ i j. *) +(* := 'chi_(dprod_Iirr defW (i, j)). *) +(* cfCyclicTIset defW i j == the virtual character of 'Z[irr W, V] coinciding *) +(* (locally) alpha_ i j with 1 - chi_i and 1 - 'chi_j on W1 and W2, *) +(* respectively. This definition is denoted by *) +(* alpha_ i j in this file, and is only used in the *) +(* proof if Peterfalvi (13.9) in the sequel. *) +(* := cfDprod defW (1 - 'chi_i) (1 - 'chi_j). *) +(* = 1 - w_ i 0 - w_ 0 j + w_ i j. *) +(* cfCyclicTIsetBase defW := the tuple of all the alpha_ i j, for i, j != 0. *) +(* (locally) cfWVbase This is a basis of 'CF(W, V); this definition is *) +(* not used outside this file. *) +(* For ctiW : cyclicTI_hypothesis defW G we also define *) +(* cyclicTIiso ctiW == a linear isometry from 'CF(W) to 'CF(G) that *) +(* (locally) sigma that extends induction on 'CF(W, V), maps the *) +(* w_ i j to virtual characters, and w_ 0 0 to 1. *) +(* This definition is usually Let-bound to sigma, *) +(* and only depends extensionally on W, V and G. *) +(* (locally) eta_ i j := sigma (w_ i j), as in sections 13 and 14 of *) +(* tha Peterfalv text. *) +(* cyclicTI_NC ctiW phi == the number of eta_ i j constituents of phi. *) +(* (locally) NC phi := #|[set ij | '[phi, eta_ ij .1 ij.2] != 0]|. *) +(* The construction of sigma involves a large combinatorial proof, for which *) +(* it is worthwhile to use reflection techniques to automate mundane and *) +(* repetitive arguments. We isolate the necessary boilerplate in a separate *) +(* CyclicTIisoReflexion module. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Section Definitions. + +Variables (gT : finGroupType) (G W W1 W2 : {set gT}). + +Definition cyclicTIset of W1 \x W2 = W := W :\: (W1 :|: W2). + +Definition cyclicTI_hypothesis (defW : W1 \x W2 = W) := + [/\ cyclic W, odd #|W| & normedTI (cyclicTIset defW) G W]. + +End Definitions. + +(* These is defined as a Notation which clients can bind with a Section Let *) +(* that can be folded easily. *) +Notation cyclicTIirr defW i j := 'chi_(dprod_Iirr defW (i, j)). + +Module CyclicTIisoReflexion. + +(******************************************************************************) +(* Support for carrying out the combinatorial parts of the proof of Theorem *) +(* (3.5) by reflection. Specifically, we need to show that in a rectangular *) +(* array of virtual characters of norm 3, of even dimensions, and such that *) +(* the dot product of two entries is 1 if they are on the same row or column, *) +(* the entries of each column contain a "pivot" normal virtual character *) +(* orthogonal to all other columns. The proof never needs to consider more *) +(* than a 4 x 2 rectangle, but frequently renumbers lines, columns and *) +(* orthonormal components in order to do so. *) +(* We want to use reflection to automate this renumbering; we also want to *) +(* automate the evaluation of the dot product constaints for partially *) +(* described entries of the matrix. *) +(* To do so we define a "theory" data structure to store a reifed form of *) +(* such partial descriptions: a set of "clauses", each consisting in an index *) +(* (i, j) into the array, and a collection of "literals" (k, v) representing *) +(* constraints '[b_(i, j), x`_k] = v%:~R, with v = 0, 1 or -1. A clause with *) +(* exactly three nonzero literals defines b_(i, j) exactly. *) +(* We define special notation for the concrete instances that appear in *) +(* reflected proofs; for example *) +(* |= & b11 = -x1 + x2 + x3 & x1, ~x2 in b12 & ? in b31 *) +(* denotes the "theory" of arrays whose two left entries decomposes into *) +(* x1 + x2 + x3 for some orthonormal x1, x2, and x3, such that the second top *) +(* entry has x1 is a signed component but is orthogonal to x2, and which have *) +(* an (unconstrained) first entry in the third column. (The concrete encoding *) +(* shifts indices to start at 0.) *) +(* The "models" in which such theories are interpreted supply the dimensions *) +(* of the array, which must be even, nonequal and at least 2, the function *) +(* mapping indices to array entries, which must be virtual characters with *) +(* the requisite norms and dot products, and an orthonormal sequence of *) +(* virtual characters that will be used to interpret the xij; a model coerces *) +(* to any of these three components. *) +(* We are primarily interested in two predicates: *) +(* sat m th <=> the interpretation of th in m is well-defined (no out of *) +(* bound indices) and valid (all constraints true). *) +(* unsat th <-> forall m, ~ sat m th *) +(* While the main theorem of this section, column_pivot, can be seen as an *) +(* instance of "sat", all the principal combinatorial lemmas use "unsat", *) +(* whose universal quantifier allows symmetry reductions. We present the set *) +(* of lemmas implementing reflection-assisted proofs of "unsat th" as a small *) +(* domain-specific proof language consisting of the following tactics: *) +(* consider bij ::= add a clause for bij, which must not appear in th, *) +(* changing the goal to unsat th & ? in bij. *) +(* bij must be within a 4 x 2 bounding box, and th *) +(* must be symmetric if bij "breaks" the 2 x 2 box. *) +(* fill bij ::= add an x(k.+1) literal to the bij clause in th, *) +(* where x1, ..., xk are all the normal characters *) +(* appearing in th, and the clause for bij exists and *) +(* contains assumptions for all of x1, ..., xk, at *) +(* two of which are nonzero. *) +(* uwlog Dcl: cl [by tac] ::= add the clause cl to th, replacing an existing *) +(* clause for the same matrix entry. This produces a *) +(* side goal of unsat th, but with an additional *) +(* assumption Dcl : unsat th+cl, which can be resolved *) +(* with the optional "by tac". *) +(* uhave lit in bij as T(ij, kl) ::= adds the literal lit (one of xk, -xk, or *) +(* ~xk) to an existing clause for bij in th, using the *) +(* reflection lemma T(ij, kl) to rule out the other *) +(* possibilities for xk. Here T can be either O *) +(* (general dot product evaluation) or L (specific *) +(* line/column constraints following from (3.5.2)). *) +(* uhave lit, lit' in bij as T(ij, kl) ::= adds both lit and lit'. *) +(* uhave lit | lit' in bij as T(ij, kl) ::= produces two subgoals, where lit *) +(* (resp. lit') is added to the ... in bij clause in *) +(* th, using T(ij, kl) to eliminate the third literal. *) +(* (lit and lit' must constrain the same component). *) +(* uhave lit | lit' | lit'' in bij ::= produces three subgoals, where lit *) +(* (resp. lit', lit'') is added to the bij clause in *) +(* th; lit, lit', lit'' should be a permutation of xk, *) +(* -xk, and ~xk for some k. *) +(* uwlog Ebij: lit | lit' in bij as T(ij, kl) ::= adds lit to the bij clause *) +(* in th, but produces a side goal where lit' has been *) +(* added instead, with an additional assumption *) +(* Ebij: th + (lit in bij); T(ij, kl) is used to rule *) +(* out the third value. *) +(* counter to T(ij, kl) ::= use T(ij, kl) to conclude that unsat th. *) +(* uexact Hth' ::= use Hth' : unsat th', where th' is a subset of th *) +(* (with the same order of literals) to conclude. *) +(* symmetric to Hth' ::= use Hth' : unsat th', where th' is a permutation *) +(* of a subset of th (preserving columns, and with at *) +(* most one row exchange) to conclude. *) +(******************************************************************************) + +Import ssrint. + +(* Clause left-hand side, a reference to a value of beta; in the reference *) +(* model m, (i, j) stands for beta_ (inord i.+1) (inord j.+1). *) +Definition ref := (nat * nat)%type. +Implicit Type ij : ref. +Definition Ref b_ij : ref := edivn (b_ij - 11) 10. (* Ref 21 = (1, 0). *) +Notation "''b' ij" := (Ref ij) (at level 0, ij at level 0, format "''b' ij"). +Notation b11 := 'b11. Notation b12 := 'b12. +Notation b21 := 'b21. Notation b22 := 'b22. +Notation b31 := 'b31. Notation b32 := 'b32. +Notation b41 := 'b41. Notation b42 := 'b42. + +Definition bbox := (nat * nat)%type. (* bounding box for refs. *) +Implicit Type bb : bbox. +Identity Coercion pair_of_bbox : bbox >-> prod. + +Definition sub_bbox bb1 bb2 := (bb1.1 <= bb2.1)%N && (bb1.2 <= bb2.2)%N. +Definition wf_ref bb := [pred ij : ref | (ij.1 < bb.1)%N && (ij.2 < bb.2)%N]. +Definition dot_ref ij1 ij2 := ((ij1.1 == ij2.1).+1 * (ij1.2 == ij2.2).+1 - 1)%N. + +Lemma bbox_refl bb : sub_bbox bb bb. Proof. exact/andP. Qed. + +(* Clause right-hand side litteral, denoting the projection of the left-hand *) +(* side on an irreducible character of G: in a valid model m, (k, v) stands *) +(* for the component m`_k *~ v = (model_xi m)`_k, and for the projection *) +(* constraint '[m i j, m`_k] == v%:~R. *) +Definition lit := (nat * int)%type. (* +x1 = (0,1) ~x2 = (1,0) -x3 = (2, -1) *) +Implicit Types (kv : lit) (kvs : seq lit). +Definition Lit k1 v : lit := if (0 + k1)%N is k.+1 then (k, v) else (k1, v). +Notation "+x k" := (Lit k 1) (at level 0, k at level 0, format "+x k"). +Notation "-x k" := (Lit k (-1)) (at level 0, k at level 0, format "-x k"). +Notation "~x k" := (Lit k 0) (at level 0, k at level 0, format "~x k"). +Notation x1 := +x1. Notation x2 := +x2. +Notation x3 := +x3. Notation x4 := +x4. +Notation x5 := +x5. Notation x6 := +x6. +Notation x7 := +x7. Notation x8 := +x8. + +Definition AndLit kvs kv := kv :: kvs. +Definition AddLit := AndLit. +Notation "(*dummy*)" := (Prop Prop) (at level 0) : defclause_scope. +Arguments Scope AddLit [defclause_scope _]. +Infix "+" := AddLit : defclause_scope. +Definition SubLit kvs kv := AddLit kvs (kv.1, - kv.2). +Arguments Scope SubLit [defclause_scope _]. +Infix "-" := SubLit : defclause_scope. +Coercion LastLit kv := [:: kv]. + +Fixpoint norm_cl kvs : nat := + (if kvs is (_, v) :: kvs1 then `|v| ^ 2 + norm_cl kvs1 else 0)%N. + +Definition clause := (ref * seq lit)%type. +Implicit Type cl : clause. +Definition Clause ij kvs : clause := (ij, kvs). +Notation "& kv1 , .. , kvn 'in' ij" := + (Clause ij (AndLit .. (AndLit nil kv1) .. kvn)) + (at level 200, ij, kv1, kvn at level 0, + format "& kv1 , .. , kvn 'in' ij"). +Notation "& ? 'in' ij" := (Clause ij nil) + (at level 200, ij at level 0, format "& ? 'in' ij"). +Definition DefClause := Clause. +Arguments Scope DefClause [_ defclause_scope]. +Notation "& ij = kvs" := (DefClause ij kvs) + (at level 200, ij at level 0, format "& ij = kvs"). + +Definition theory := seq clause. +Implicit Type th : theory. +Definition AddClause th cl : theory := cl :: th. +Notation "|= cl1 .. cln" := (AddClause .. (AddClause nil cl1) .. cln) + (at level 8, cl1, cln at level 200, + format "|= '[hv' cl1 '/' .. '/' cln ']'"). + +(* Transpose (W1 / W2 symmetry). *) + +Definition tr (ij : nat * nat) : ref := (ij.2, ij.1). +Definition tr_th th : theory := [seq (tr cl.1, cl.2) | cl <- th]. + +Lemma trK : involutive tr. Proof. by case. Qed. +Lemma tr_thK : involutive tr_th. Proof. by apply: mapK => [[[i j] kvs]]. Qed. + +(* Index range of a theory. *) + +Fixpoint th_bbox th : bbox := + if th is (i, j, _) :: th1 then + let: (ri, rj) := th_bbox th1 in (maxn i.+1 ri, maxn j.+1 rj) + else (0, 0)%N. + +Lemma th_bboxP th bb : + reflect {in th, forall cl, cl.1 \in wf_ref bb} (sub_bbox (th_bbox th) bb). +Proof. +pose in_bb := [pred cl : clause | cl.1 \in wf_ref bb]. +suffices ->: sub_bbox (th_bbox th) bb = all in_bb th by apply: allP. +elim: th => [|[[i j] _] th] //=; case: (th_bbox th) => ri rj /=. +by rewrite /sub_bbox /= !geq_max andbACA => ->. +Qed. +Implicit Arguments th_bboxP [th bb]. + +Fixpoint th_dim th : nat := + if th is (_, kvs) :: th1 then + foldr (fun kv => maxn kv.1.+1) (th_dim th1) kvs + else 0%N. + +Lemma th_dimP th bk : + reflect {in th, forall cl, {in cl.2, forall kv, kv.1 < bk}}%N + (th_dim th <= bk)%N. +Proof. +pose in_bk := [pred cl : clause | all (fun kv => kv.1 < bk)%N cl.2]. +suffices ->: (th_dim th <= bk)%N = all in_bk th. + by apply: (iffP allP) => bk_th cl /bk_th/allP. +elim: th => // [[_ kvs] th /= <-]; elim: kvs => //= kv kvs. +by rewrite -andbA geq_max => ->. +Qed. +Implicit Arguments th_dimP [th bk]. + +(* Theory and clause lookup. *) + +CoInductive get_spec T (P : T -> Prop) (Q : Prop) : option T -> Prop := + | GetSome x of P x : get_spec P Q (Some x) + | GetNone of Q : get_spec P Q None. + +Fixpoint get_cl ij (th : theory) : option clause := + if th is cl :: th1 then if cl.1 == ij then Some cl else get_cl ij th1 + else None. + +Lemma get_clP ij (th : theory) : + get_spec (fun cl : clause => cl \in th /\ cl.1 = ij) True (get_cl ij th). +Proof. +elim: th => /= [|cl th IHth]; first by right. +case: eqP => [Dij | _]; first by left; rewrite ?mem_head. +by case: IHth => [cl1 [th_cl1 Dij]|]; constructor; rewrite // mem_behead. +Qed. + +Fixpoint get_lit (k0 : nat) kvs : option int := + if kvs is (k, v) :: kvs1 then if k == k0 then Some v else get_lit k0 kvs1 + else None. + +Lemma get_litP k0 kvs : + get_spec (fun v => (k0, v) \in kvs) (k0 \notin unzip1 kvs) (get_lit k0 kvs). +Proof. +elim: kvs => [|[k v] kvs IHkvs /=]; [by right | rewrite inE eq_sym]. +have [-> | k'0] := altP eqP; first by left; rewrite ?mem_head. +by have [v0 kvs_k0v | kvs'k0] := IHkvs; constructor; rewrite // mem_behead. +Qed. + +(* Theory extension. *) + +Fixpoint set_cl cl2 th : wrapped theory := + if th is cl :: th1 then + let: Wrap th2 := set_cl cl2 th1 in + if cl.1 == cl2.1 then Wrap (AddClause th2 cl2) else Wrap (AddClause th2 cl) + else Wrap nil. + +Definition ext_cl th cl k v := + let: (ij, kvs) := cl in set_cl (Clause ij (AndLit kvs (Lit k.+1 v))) th. + +Definition wf_ext_cl cl k rk := (k \notin unzip1 cl.2) && (k < rk)%N. + +Definition wf_fill k kvs := (size kvs == k) && (norm_cl kvs < 3)%N. + +Lemma ext_clP cl1 th k v (cl1k := (cl1.1, (k, v) :: cl1.2)) : + cl1 \in th -> + exists2 th1, ext_cl th cl1 k v = Wrap th1 + & cl1k \in th1 + /\ th1 =i [pred cl | if cl.1 == cl1.1 then cl == cl1k else cl \in th]. +Proof. +case: cl1 => ij kvs /= in cl1k * => th_cl1; set th1p := [pred cl | _]. +pose th1 := [seq if cl.1 == ij then cl1k else cl | cl <- th]. +exists th1; first by elim: (th) @th1 => //= cl th' ->; rewrite -2!fun_if. +suffices Dth1: th1 =i th1p by rewrite Dth1 !inE !eqxx. +move=> cl; rewrite inE; apply/mapP/idP=> [[{cl}cl th_cl ->] | ]. + by case cl_ij: (cl.1 == ij); rewrite ?eqxx ?cl_ij. +case: ifP => [_ /eqP-> | cl'ij th_cl]; last by exists cl; rewrite ?cl'ij. +by exists (ij, kvs); rewrite ?eqxx. +Qed. + +(* Satisfiability tests. *) + +Definition sat_test (rO : rel clause) ij12 th := + if get_cl (Ref ij12.1) th is Some cl1 then + oapp (rO cl1) true (get_cl (Ref ij12.2) th) + else true. + +(* This reflects the application of (3.5.1) for an arbitrary pair of entries. *) +Definition Otest cl1 cl2 := + let: (ij1, kvs1) := cl1 in let: (ij2, kvs2) := cl2 in + let fix loop s1 s2 kvs2 := + if kvs2 is (k, v2) :: kvs2 then + if get_lit k kvs1 is Some v1 then loop (v1 * v2 + s1) s2 kvs2 else + loop s1 s2.+1 kvs2 + else (s1, if norm_cl kvs1 == 3 then 0%N else s2) in + let: (s1, s2) := loop 0 0%N kvs2 in + (norm_cl kvs2 == 3) ==> (`|s1 - dot_ref ij1 ij2| <= s2)%N. + +(* Matching up to a permutation of the rows, columns, and base vectors. *) + +Definition sub_match th1 th2 := + let match_cl cl1 cl2 := + if cl2.1 == cl1.1 then subseq cl1.2 cl2.2 else false in + all [pred cl1 | has (match_cl cl1) th2] th1. + +Definition wf_consider ij th (ri := (th_bbox th).1) := + (ij.1 < 2 + ((2 < ri) || sub_match th (tr_th th)).*2)%N && (ij.2 < 2)%N. + +CoInductive sym := Sym (si : seq nat) (sj : seq nat) (sk : seq nat). + +Definition sym_match s th1 th2 := + let: Sym si sj sk := s in let: (ri, rj, rk) := (th_bbox th1, th_dim th1) in + let is_sym r s := uniq s && all (gtn r) s in + let match_cl cl2 := + let: (i2, j2, kvs2) := cl2 in let ij := (nth ri si i2, nth rj sj j2) in + let match_lit kvs1 kv := (nth rk sk kv.1, kv.2) \in kvs1 in + let match_cl1 cl1 := + let: (ij1, kvs1) := cl1 in (ij1 == ij) && all (match_lit kvs1) kvs2 in + uniq (unzip1 kvs2) && has match_cl1 th1 in + [&& is_sym ri si, is_sym rj sj, is_sym rk sk & all match_cl th2]. + +(* Try to compute the base vector permutation for a given row and column *) +(* permutation. We assume each base vector is determined by the entries of *) +(* which it is a proper constituent, and that there are at most two columns. *) +Definition find_sym_k th1 th2 (si sj : seq nat) := + let store_lit c kv ksig := + let: (k, v) := kv in if v == 0 then ksig else let cv := (c, v) in + let fix insert_in (cvs : seq (nat * int)) := + if cvs is cv' :: cvs' then + if (c < cv'.1)%N then cv :: cvs else cv' :: insert_in cvs' + else [:: cv] in + set_nth nil ksig k (insert_in (nth nil ksig k)) in + let fix read_lit ksig1 ksig2 := + if ksig1 is cvs :: ksig1' then + let k := index cvs ksig2 in + k :: read_lit ksig1' (set_nth nil ksig2 k nil) + else nil in + let fix store2 ksig1 ksig2 cls1 := + if cls1 is (i1, j1, kvs1) :: cls1' then + if get_cl (nth 0 si i1, nth 0 sj j1)%N th2 is Some (_, kvs2) then + let st_kvs := foldr (store_lit (i1.*2 + j1)%N) in (* assume j1 <= 1 *) + store2 (st_kvs ksig1 kvs1) (st_kvs ksig2 kvs2) cls1' + else None + else + let sk := read_lit ksig1 ksig2 in + if all (gtn (size ksig2)) sk then Some (Sym si sj sk) else None in + store2 nil nil th1. + +(* Try to find a symmetry that maps th1 to th2, assuming the same number of *) +(* rows and columns, and considering at most one row exchange. *) +Definition find_sym th1 th2 := + let: (ri, rj) := th_bbox th2 in let si := iota 0 ri in let sj := iota 0 rj in + if find_sym_k th1 th2 si sj is Some _ as s then s else + let fix loop m := + if m is i.+1 then + let fix inner_loop m' := + if m' is i'.+1 then + let si' := (set_nth 0 (set_nth 0 si i i') i' i)%N in + if find_sym_k th1 th2 si' sj is Some _ as s then s else inner_loop i' + else None in + if inner_loop i is Some _ as s then s else loop i + else None in + loop ri. + +Section Interpretation. + +Variables (gT : finGroupType) (G : {group gT}). + +Definition is_Lmodel bb b := + [/\ [/\ odd bb.1.+1, odd bb.2.+1, bb.1 > 1, bb.2 > 1 & bb.1 != bb.2]%N, + forall ij, b ij \in 'Z[irr G] + & {in wf_ref bb &, forall ij1 ij2, '[b ij1, b ij2] = (dot_ref ij1 ij2)%:R}]. + +Definition is_Rmodel X := orthonormal X /\ {subset X <= 'Z[irr G]}. + +Inductive model := Model bb f X of is_Lmodel bb f & is_Rmodel X. + +Coercion model_bbox m := let: Model d _ _ _ _ := m in d. +Definition model_entry m := let: Model _ f _ _ _ := m in f. +Coercion model_entry : model >-> Funclass. +Coercion model_basis m := let: Model _ _ X _ _ := m in X. +Lemma LmodelP (m : model) : is_Lmodel m m. Proof. by case: m. Qed. +Lemma RmodelP (m : model) : is_Rmodel m. Proof. by case: m. Qed. + +Fact nil_RmodelP : is_Rmodel nil. Proof. by []. Qed. + +Definition eval_cl (m : model) kvs := \sum_(kv <- kvs) m`_kv.1 *~ kv.2. + +Definition sat_lit (m : model) ij kv := '[m ij, m`_kv.1] == kv.2%:~R. +Definition sat_cl m cl := uniq (unzip1 cl.2) && all (sat_lit m cl.1) cl.2. + +Definition sat (m : model) th := + [&& sub_bbox (th_bbox th) m, th_dim th <= size m & all (sat_cl m) th]%N. +Definition unsat th := forall m, ~ sat m th. + +Lemma satP (m : model) th : + reflect {in th, forall cl, + [/\ cl.1 \in wf_ref m, uniq (unzip1 cl.2) + & {in cl.2, forall kv, kv.1 < size m /\ sat_lit m cl.1 kv}%N]} + (sat m th). +Proof. +apply: (iffP and3P) => [[/th_bboxP thbP /th_dimP thdP /allP thP] cl th_cl |thP]. + have /andP[-> clP] := thP _ th_cl; split=> // [|kv cl_kv]; first exact: thbP. + by rewrite (thdP _ th_cl) ?(allP clP). +split; first by apply/th_bboxP=> cl /thP[]. + by apply/th_dimP=> cl /thP[_ _ clP] kv /clP[]. +by apply/allP=> cl /thP[_ Ucl clP]; rewrite /sat_cl Ucl; apply/allP=> kv /clP[]. +Qed. +Implicit Arguments satP [m th]. + +(* Reflexion of the dot product. *) + +Lemma norm_clP m th cl : + sat m th -> cl \in th -> + let norm := norm_cl cl.2 in let beta := m cl.1 in + [/\ (norm <= 3)%N, norm == 3 -> beta = eval_cl m cl.2 + & (norm < 3)%N -> size cl.2 == size m -> + exists2 dk, dk \in dirr_constt beta & orthogonal (dchi dk) m]. +Proof. +case: cl => ij kvs /satP thP /thP[wf_ij Uks clP] norm beta. +have [[_ ZmL Dm] [o1m ZmR]] := (LmodelP m, RmodelP m). +set ks := unzip1 kvs in Uks; pose Aij := [seq m`_k | k <- ks]. +have lt_ks k: k \in ks -> (k < size m)%N by case/mapP=> kv /clP[ltk _] ->. +have sAm: {subset Aij <= (m : seq _)} + by move=> _ /mapP[k /lt_ks ltk ->]; rewrite mem_nth. +have o1Aij: orthonormal Aij. + have [Um _] := orthonormalP o1m; apply: sub_orthonormal o1m => //. + rewrite map_inj_in_uniq // => k1 k2 /lt_ks ltk1 /lt_ks ltk2 /eqP. + by apply: contraTeq; rewrite nth_uniq. +have [X AijX [Y [defXY oXY oYij]]] := orthogonal_split Aij beta. +have{AijX} defX: X = \sum_(xi <- Aij) '[beta, xi] *: xi. + have [_ -> ->] := orthonormal_span o1Aij AijX; apply: eq_big_seq => xi CFxi. + by rewrite defXY cfdotDl (orthoPl oYij) ?addr0. +have ->: eval_cl m kvs = X. + rewrite {}defX !big_map; apply: eq_big_seq => kv /clP[_ /eqP->]. + by rewrite scaler_int. +rewrite -leC_nat -ltC_nat -eqC_nat /=. +have <-: '[beta] = 3%:R by rewrite Dm // /dot_ref !eqxx. +have <-: '[X] = norm%:R. + rewrite {}defX {}/norm cfnorm_sum_orthonormal // {o1Aij oYij sAm}/Aij. + transitivity (\sum_(kv <- kvs) `|kv.2%:~R| ^+ 2 : algC). + by rewrite !big_map; apply: eq_big_seq => kv /clP[_ /eqP->]. + rewrite unlock /=; elim: (kvs) => //= [[k v] kvs' ->]. + by rewrite -intr_norm -natrX -natrD. +rewrite defXY cfnormDd //; split; first by rewrite ler_paddr ?cfnorm_ge0. + by rewrite eq_sym addrC -subr_eq0 addrK cfnorm_eq0 => /eqP->; rewrite addr0. +have{ZmL} Zbeta: beta \in 'Z[irr G] by apply: ZmL. +have Z_X: X \in 'Z[irr G]. + rewrite defX big_seq rpred_sum // => xi /sAm/ZmR Zxi. + by rewrite rpredZ_Cint ?Cint_cfdot_vchar. +rewrite -ltr_subl_addl subrr cnorm_dconstt; last first. + by rewrite -[Y](addKr X) -defXY addrC rpredB. +have [-> | [dk Ydk] _ /eqP sz_kvs] := set_0Vmem (dirr_constt Y). + by rewrite big_set0 ltrr. +have Dks: ks =i iota 0 (size m). + have: {subset ks <= iota 0 (size m)} by move=> k /lt_ks; rewrite mem_iota. + by case/leq_size_perm=> //; rewrite size_iota size_map sz_kvs. +suffices o_dk_m: orthogonal (dchi dk) m. + exists dk; rewrite // dirr_consttE defX cfdotDl cfdot_suml. + rewrite big1_seq ?add0r -?dirr_consttE // => xi /sAm CFxi. + by rewrite cfdotC cfdotZr (orthoPl o_dk_m) // mulr0 conjC0. +apply/orthoPl=> _ /(nthP 0)[k ltk <-]; have [Um o_m] := orthonormalP o1m. +have Z1k: m`_k \in dirr G by rewrite dirrE ZmR ?o_m ?eqxx ?mem_nth. +apply: contraTeq Ydk => /eqP; rewrite dirr_consttE cfdot_dirr ?dirr_dchi //. +have oYm: '[Y, m`_k] = 0 by rewrite (orthoPl oYij) ?map_f // Dks mem_iota. +by do 2?case: eqP => [-> | _]; rewrite // ?cfdotNr oYm ?oppr0 ltrr. +Qed. + +Lemma norm_cl_eq3 m th cl : + sat m th -> cl \in th -> norm_cl cl.2 == 3 -> m cl.1 = eval_cl m cl.2. +Proof. by move=> m_th /(norm_clP m_th)[]. Qed. + +Lemma norm_lit m th cl kv : + sat m th -> cl \in th -> kv \in cl.2 -> (`|kv.2| <= 1)%N. +Proof. +move=> m_th /(norm_clP m_th)[cl_le3 _ _]. +elim: cl.2 => //= [[k v] kvs IHkvs] in cl_le3 * => /predU1P[-> | /IHkvs->//]. + by apply: contraLR cl_le3; rewrite -ltnNge -leq_sqr => /subnKC <-. +exact: leq_trans (leq_addl _ _) cl_le3. +Qed. + +(* Decision procedure framework (in which we will define O and L). *) + +Definition is_sat_test (tO : pred theory) := forall m th, sat m th -> tO th. + +Lemma sat_testP (rO : rel clause) ij12 : + (forall m th cl1 cl2, sat m th -> cl1 \in th -> cl2 \in th -> rO cl1 cl2) -> + is_sat_test (sat_test rO ij12). +Proof. +rewrite /sat_test => O m th /O O_th; case: get_clP => // cl1 [th_cl1 _]. +by case: get_clP => // cl2 [th_cl2 _]; apply: O_th. +Qed. + +(* Case analysis on the value of a specific projection. *) + +Definition lit_vals : seq int := [:: 0; 1; -1]. + +Lemma sat_cases (m : model) th k cl : + sat m th -> cl \in th -> wf_ext_cl cl k (size m) -> + exists2 v, v \in lit_vals & sat m (unwrap (ext_cl th cl k v)). +Proof. +case: cl => ij kvs /satP thP th_cl /andP[cl'k ltkm]. +have [[_ ZmL _] [o1m ZmR]] := (LmodelP m, RmodelP m). +have [m_ij Uij clP] := thP _ th_cl. +have /CintP[v Dv]: '[m ij, m`_k] \in Cint. + by rewrite Cint_cfdot_vchar ?ZmL ?ZmR ?mem_nth. +have [/= th1 Dthx [th1_cl Dth1]] := ext_clP k v th_cl. +suffices{Dthx} m_th1: sat m th1. + exists v; last by rewrite /ext_cl Dthx. + by case: (v) (norm_lit m_th1 th1_cl (mem_head _ _)); do 2?case. +apply/satP=> cl1; rewrite Dth1 inE; case: ifP => [_ /eqP-> | _ /thP] //=. +by rewrite cl'k; split=> // kv /predU1P[-> | /clP//]; rewrite /sat_lit Dv. +Qed. +Implicit Arguments sat_cases [m th cl]. + +Definition unsat_cases_hyp th0 kvs tO cl := + let: (k, _) := head (2, 0) kvs in let thk_ := ext_cl th0 cl k in + let th's := [seq unwrap (thk_ v) | v <- lit_vals & v \notin unzip2 kvs] in + let add hyp kv := + let: (_, v) := kv in let: Wrap th := thk_ v in hyp /\ unsat th in + foldl add (wf_ext_cl cl k (th_dim th0) && all (predC tO) th's) kvs. + +Lemma unsat_cases th ij kvs tO : + is_sat_test tO -> oapp (unsat_cases_hyp th kvs tO) False (get_cl ij th) -> + unsat th. +Proof. +case: get_clP => //= cl [th_cl _] O; rewrite /unsat_cases_hyp. +case: head => k _; set thk_ := ext_cl th cl k; set add := fun _ _ => _. +set wf_kvs := _ && _; rewrite -[kvs]revK foldl_rev => Ukvs m m_th. +have{Ukvs}: all (fun kv => ~~ sat m (unwrap (thk_ kv.2))) (rev kvs) && wf_kvs. + elim: rev Ukvs => // [[_ v] /= kvs' IH]; case Dthk: (thk_ v) => [thv] [/IH]. + by rewrite -andbA => -> Uthk; rewrite andbT; apply/negP; apply: Uthk. +case/and3P=> /allP Uthkvs /andP[cl'k ltkr] /allP Uthkv's. +have [|{cl'k ltkr} v lit_v m_thv] := sat_cases k m_th th_cl. + by rewrite /wf_ext_cl cl'k (leq_trans ltkr) //; have [] := and3P m_th. +have /idPn[] := O _ _ m_thv; apply: Uthkv's; apply: map_f. +rewrite mem_filter lit_v andbT -mem_rev -map_rev. +by apply: contraL m_thv => /mapP[kv /Uthkvs m'thkv ->]. +Qed. + +(* Dot product reflection. *) + +Lemma O ij12 : is_sat_test (sat_test Otest ij12). +Proof. +apply: sat_testP => m th [ij1 kvs1] [ij2 kvs2] /= m_th th_cl1 th_cl2. +set cl1eq := _ == 3; set cl2eq := _ == 3; have [_ _ Dm] := LmodelP m. +pose goal s1 s2 := cl2eq ==> (`|s1 - (dot_ref ij1 ij2)%:~R| <= s2%:R :> algC). +set kvs := kvs2; set s1 := 0; set s2 := {2}0%N; have thP := satP m_th. +have{thP} [[wf_cl1 _ cl1P] [wf_cl2 _ cl2P]] := (thP _ th_cl1, thP _ th_cl2). +have: goal (s1%:~R + '[m ij1, eval_cl m kvs]) (if cl1eq then 0%N else s2). + apply/implyP=> /(norm_cl_eq3 m_th th_cl2) <-. + by rewrite if_same Dm // addrK normr0. +have /allP: {subset kvs <= kvs2} by []. +rewrite cfdot_sumr unlock; elim: kvs s1 s2 => [|[k v2] kvs IHkvs] s1 s2 /=. + by rewrite addr0 /goal -rmorphB pmulrn -!CintrE. +case/andP=> kvs2_v /IHkvs{IHkvs}IHkvs; have{cl2P} [ltk _] := cl2P _ kvs2_v. +have [v1 /cl1P[_ /eqP/=Dv1] | kvs1'k] := get_litP. + rewrite addrA => gl12; apply: IHkvs; congr (goal (_ + _) _): gl12. + by rewrite raddfMz addrC /= Dv1 -mulrzA -rmorphD. +move=> gl12; apply: IHkvs; case: ifP gl12 => [/(norm_cl_eq3 m_th th_cl1)->|_]. + rewrite cfdot_suml big1_seq ?add0r //= => kv1 kvs1_kv1. + have [[ltk1 _] [/orthonormalP[Um oom] _]] := (cl1P _ kvs1_kv1, RmodelP m). + rewrite -!scaler_int cfdotZl cfdotZr oom ?mem_nth ?nth_uniq // mulrb. + by rewrite ifN ?mulr0 //; apply: contraNneq kvs1'k => <-; apply: map_f. +rewrite /goal -(ler_add2r 1) -mulrSr; case: (cl2eq) => //; apply: ler_trans. +set s := '[_, _]; rewrite -[_ + _](addrK s) (ler_trans (ler_norm_sub _ _)) //. +rewrite 2![_ + s]addrAC addrA ler_add2l {}/s -scaler_int cfdotZr rmorph_int. +have [|v1 _] := sat_cases k m_th th_cl1; first exact/andP. +have [th1 -> /= [th1_cl1 _] m_th1] := ext_clP k v1 th_cl1. +have [_ _ /(_ _ (mem_head _ _))[_ /eqP->]] := satP m_th1 _ th1_cl1. +have ubv1: (`|v1| <= 1)%N := norm_lit m_th1 th1_cl1 (mem_head _ _). +have ubv2: (`|v2| <= 1)%N := norm_lit m_th th_cl2 kvs2_v. +by rewrite -rmorphM -intr_norm lern1 abszM /= (leq_mul ubv2 ubv1). +Qed. + +(* "Without loss" cut rules. *) + +Lemma unsat_wlog cl th : + (let: Wrap th1 := set_cl cl th in (unsat th1 -> unsat th) /\ unsat th1) -> + unsat th. +Proof. by case: set_cl => th1 [Uth /Uth]. Qed. + +Lemma unsat_wlog_cases th1 th2 : + (unsat th1 -> unsat th2) -> unsat th1 -> (true /\ unsat th1) /\ unsat th2. +Proof. by move=> Uth2 Uth1; split; last exact: Uth2. Qed. + +(* Extend the orthonormal basis *) + +Lemma sat_fill m th cl (k := th_dim th) : + sat m th -> cl \in th -> wf_fill k cl.2 -> + exists mr : {CFk | is_Rmodel CFk}, + sat (Model (LmodelP m) (svalP mr)) (unwrap (ext_cl th cl k 1)). +Proof. +move=> m_th th_cl /andP[/eqP sz_kvs n3cl]. +wlog sz_m: m m_th / size m = k. + have lekm: (k <= size m)%N by have [] := and3P m_th. + have mrP: is_Rmodel (take k m). + have [] := RmodelP m; rewrite -{1 2}(cat_take_drop k m) orthonormal_cat /=. + by case/andP=> o1mr _ /allP; rewrite all_cat => /andP[/allP]. + move/(_ (Model (LmodelP m) mrP)); apply; rewrite ?size_takel //. + congr (_ && _): m_th; rewrite lekm size_takel ?leqnn //=. + apply: eq_in_all => cl1 /th_dimP lt_cl1; congr (_ && _). + by apply: eq_in_all => kv1 /lt_cl1 lt_kv1; rewrite /sat_lit nth_take ?lt_kv1. +have [_ _ [//||dk cl_dk o_dk_m]] := norm_clP m_th th_cl. + by rewrite sz_kvs sz_m. +have CFkP: is_Rmodel (rcons m (dchi dk)). + have [o1m /allP Zm] := RmodelP m. + split; last by apply/allP; rewrite all_rcons /= dchi_vchar. + rewrite -cats1 orthonormal_cat o1m orthogonal_sym o_dk_m. + by rewrite /orthonormal /= cfnorm_dchi eqxx. +exists (exist _ _ CFkP); set mk := Model _ _. +have{m_th} mk_th: sat mk th. + congr (_ && _): m_th; rewrite size_rcons sz_m leqnn ltnW //=. + apply: eq_in_all => cl1 /th_dimP lt_cl1; congr (_ && _). + apply: eq_in_all => kv1 /lt_cl1 lt_kv1; congr ('[_, _] == _). + by rewrite nth_rcons sz_m lt_kv1. +have [|{mk_th}v ub_v m_th] := sat_cases k mk_th th_cl. + rewrite /wf_ext_cl size_rcons sz_m (contraFN _ (ltnn k)) //=. + by case/mapP=> kv kv_cl {1}->; rewrite (th_dimP _ _ th_cl). +suffices: 0 < v by case/or4P: ub_v m_th => // /eqP->. +case: (ext_clP k v th_cl) m_th => th1 -> [th1_cl1 _] /and3P[_ _]. +case/allP/(_ _ th1_cl1)/and3P=> _ /eqP/=. +by rewrite nth_rcons sz_m ltnn eqxx CintrE => <- _; rewrite -dirr_consttE. +Qed. + +Lemma unsat_fill ij th : + let fill_cl cl := + if (th_dim th).+1 %/ 1 is k.+1 then + let: Wrap thk := ext_cl th cl k 1 in wf_fill k cl.2 /\ unsat thk + else True in + oapp fill_cl False (get_cl ij th) -> unsat th. +Proof. +rewrite divn1; case: get_clP => //= cl [th_cl _]. +case Dthk: ext_cl => [th1] [wf_thk Uth1] m m_th. +by have [mk] := sat_fill m_th th_cl wf_thk; rewrite Dthk => /Uth1. +Qed. + +(* Matching an assumption exactly. *) + +Lemma sat_exact m th1 th2 : sub_match th1 th2 -> sat m th2 -> sat m th1. +Proof. +move/allP=> s_th12 /satP th2P; apply/satP => cl1 /s_th12/hasP[cl2 th_cl2]. +case: eqP => // <- s_cl12; have [wf_ij2 Ucl2 cl2P] := th2P _ th_cl2. +split=> // [|kv /(mem_subseq s_cl12)/cl2P//]. +by rewrite (subseq_uniq _ Ucl2) ?map_subseq. +Qed. + +Lemma unsat_exact th1 th2 : sub_match th1 th2 -> unsat th1 -> unsat th2. +Proof. by move=> sth21 Uth1 m /(sat_exact sth21)/Uth1. Qed. + +(* Transpose (W1 / W2 symmetry). *) + +Fact tr_Lmodel_subproof (m : model) : is_Lmodel (tr m) (fun ij => m (tr ij)). +Proof. +case: m => /= d f _ [[odd_d1 odd_d2 d1gt1 d2gt1 neq_d12] Zf fP] _. +split=> // [|[j1 i1] [j2 i2]]; first by rewrite eq_sym. +by rewrite ![_ \in _]andbC /= => wf_ij1 wf_ij2; rewrite fP // /dot_ref mulnC. +Qed. + +Definition tr_model m := Model (tr_Lmodel_subproof m) (RmodelP m). + +Lemma sat_tr m th : sat m th -> sat (tr_model m) (tr_th th). +Proof. +move/satP=> thP; apply/satP=> _ /mapP[[[i j] kvs] /thP[m_ij Uks kvsP] ->]. +by rewrite inE /= andbC. +Qed. + +(* Extend the theory (add a new empty clause). *) + +Lemma unsat_consider ij th : + wf_consider ij th -> unsat (AddClause th (& ? in ij)) -> unsat th. +Proof. +case: ij => i j; case/andP; set sym_t := sub_match _ _ => lti ltj Uthij m m_th. +wlog le_m21: m m_th / sym_t -> (m.2 <= m.1)%N. + move=> IH; apply: (IH m m_th) => sym_th. + rewrite leqNgt; apply/negP=> /leqW le_m1_m2. + by have /(sat_exact sym_th)/IH[] := sat_tr m_th. +apply: (Uthij m); congr (_ && _): (m_th) => /=; case: (th_bbox th) => ri rj /=. +have [[odd_m1 _ m1gt1 m2gt1 neq_m12] _ _] := LmodelP m. +rewrite /sub_bbox !geq_max (leq_trans ltj) ?(leq_trans lti) //; case: orP => //. +rewrite -(ltnS 4) (odd_geq _ odd_m1) ltnS. +case=> [/leq_trans-> // | /le_m21]; first by have [/andP[]] := and3P m_th. +by rewrite leq_eqVlt eq_sym (negPf neq_m12); apply: leq_trans. +Qed. + +(* Matching up to a permutation of the rows, columns, and base vectors. *) + +Lemma unsat_match s th1 th2 : sym_match s th1 th2 -> unsat th2 -> unsat th1. +Proof. +pose I_ si mi := si ++ filter [predC si] (iota 0 mi). +have SsP mi si ri (Ii := I_ si mi): + uniq si && all (gtn ri) si -> (ri <= mi)%N -> + [/\ {in Ii, forall i, i < mi}%N, uniq Ii & size Ii = mi]. +- case/andP=> Usi /allP/=ltsi le_ri_mi; have uIm := iota_uniq 0 mi. + have uIi: uniq Ii by rewrite cat_uniq Usi -all_predC filter_all filter_uniq. + have defIi: Ii =i iota 0 mi. + move=> i; rewrite mem_cat mem_filter orb_andr orbN mem_iota. + by apply: orb_idl => /ltsi/leq_trans->. + split=> // [i|]; first by rewrite defIi mem_iota. + by rewrite (perm_eq_size (uniq_perm_eq _ _ defIi)) ?size_iota. +have lt_nth ri si i: (nth ri si i < ri)%N -> (i < size si)%N. + by rewrite !ltnNge; apply: contra => le_si; rewrite nth_default. +case: s => [si sj sk] /= sym12 Uth2 m m_th1; case/and3P: (m_th1) sym12. +case: th_bbox (th_bboxP (bbox_refl (th_bbox th1))) => ri rj rijP. +case/andP=> /= leri lerj lerk _ /and4P[Ssi Ssj /andP[Usk /allP/=lesrk] sym12]. +have{Ssi} /SsP/(_ leri)[ltIi uIi szIi] := Ssi. +have{Ssj SsP} /SsP/(_ lerj)[ltIj uIj szIj] := Ssj. +pose smL ij := m (nth ri (I_ si m.1) ij.1, nth rj (I_ sj m.2) ij.2)%N. +pose smR := [seq m`_k | k <- sk]. +have [[lb_m ZmL Dm] [o1m ZmR]] := (LmodelP m, RmodelP m). +have{lb_m} smLP: is_Lmodel m smL. + split=> // [ij | ij1 ij2 /andP[lti1 ltj1] /andP[lti2 ltj2]]; first exact: ZmL. + by rewrite Dm ?inE /dot_ref/= ?nth_uniq ?ltIi ?ltIj ?mem_nth ?szIi ?szIj. +have{lesrk} ubk k: k \in sk -> (k < size m)%N by move=> /lesrk/leq_trans->. +have smRP: is_Rmodel smR. + have ssmR: {subset smR <= (m : seq _)}. + by move=> _ /mapP[k s_k ->]; rewrite mem_nth ?ubk. + split=> [|xi /ssmR/ZmR//]; have [Um _] := orthonormalP o1m. + apply: sub_orthonormal o1m; rewrite ?map_inj_in_uniq //. + by apply: can_in_inj (index^~ m) _ => k s_k; rewrite /= index_uniq ?ubk. +apply: (Uth2 (Model smLP smRP)); apply/satP=> [][[i2 j2] kvs2] /(allP sym12). +case/andP=> -> /hasP[[[i1 j1] kvs1] th1_cl1 /andP[/eqP[Di1 Dj1] /allP s_kv12]]. +have:= rijP _ th1_cl1; rewrite Di1 Dj1 => /andP[/lt_nth lti1 /lt_nth ltj1]. +rewrite !inE -szIi -szIj !size_cat !(leq_trans _ (leq_addr _ _)) //. +split=> // kv /s_kv12 kvs1_kv1; rewrite size_map /sat_lit /=. +have /lt_nth ltk := th_dimP (leqnn _) _ th1_cl1 _ kvs1_kv1; split=> //. +rewrite (nth_map (th_dim th1)) // /smL !nth_cat lti1 ltj1 -Di1 -Dj1. +by have [_ _ /(_ _ kvs1_kv1)[]] := satP m_th1 _ th1_cl1. +Qed. + +Lemma unsat_sym th1 th2 : + (if find_sym th1 th2 is Some s then sym_match s th2 th1 else false) -> + unsat th1 -> unsat th2. +Proof. by case: find_sym => // s; apply: unsat_match. Qed. + +End Interpretation. + +Implicit Arguments satP [gT G m th]. +Implicit Arguments unsat [gT G]. +Implicit Arguments sat_cases [gT G m th cl]. +Implicit Arguments unsat_cases [gT G th tO]. +Implicit Arguments unsat_wlog [gT G]. +Implicit Arguments unsat_fill [gT G]. +Implicit Arguments unsat_consider [gT G]. +Implicit Arguments unsat_match [gT G th1 th2]. + +(* The domain-specific tactic language. *) + +Tactic Notation "consider" constr(ij) := + apply: (unsat_consider ij); first exact isT. + +(* Note that "split" here would be significantly less efficient, because it *) +(* would evaluate the reflected assumption four times. *) +Tactic Notation "fill" constr(ij) := + apply: (unsat_fill ij); apply: (conj isT _). + +Tactic Notation "uwlog" simple_intropattern(IH) ":" constr(cl) := + apply: (unsat_wlog cl); split=> [IH | ]. + +Tactic Notation "uwlog" simple_intropattern(IH) ":" constr(cl) + "by" tactic4(tac) := + apply: (unsat_wlog cl); split=> [IH | ]; first by [tac]. + +Tactic Notation "uhave" constr(kv) "in" constr(ij) + "as" constr(T) constr(ij12) := + apply: (unsat_cases ij [:: kv] (T ij12)); apply: (conj isT _). + +Tactic Notation "uhave" constr(kv1) "," constr(kv2) "in" constr(ij) + "as" constr(T) constr(ij12) := + uhave kv1 in ij as T ij12; uhave kv2 in ij as T ij12. + +Tactic Notation "uhave" constr(kv1) "|" constr(kv2) "in" constr(ij) + "as" constr(T) constr(ij12) := + apply: (unsat_cases ij [:: kv1; kv2] (T ij12)); apply: (conj (conj isT _) _). + +Tactic Notation "uhave" constr(kv1) "|" constr(kv2) "|" constr(kv3) + "in" constr(ij) := + apply: (unsat_cases ij [:: kv1; kv2; kv3] (fun _ _ _ => isT)); + apply: (conj (conj (conj isT _) _) _). + +Tactic Notation "uwlog" simple_intropattern(IH) ":" + constr(kv1) "|" constr(kv2) "in" constr(ij) + "as" constr(T) constr(ij12) := + apply: (unsat_cases ij [:: kv1; kv2] (T ij12)); + apply: unsat_wlog_cases => [IH | ]. + +Tactic Notation "counter" "to" constr(T) constr(ij12) := by move=> ? /(T ij12). + +Tactic Notation "uexact" constr(IH) := apply: unsat_exact IH; exact isT. + +Tactic Notation "symmetric" "to" constr(IH) := apply: unsat_sym (IH); exact isT. + +End CyclicTIisoReflexion. + +Section Three. + +Variables (gT : finGroupType) (G W W1 W2 : {group gT}). +Hypothesis defW : W1 \x W2 = W. + +Let V := cyclicTIset defW. +Let w_ i j := cyclicTIirr defW i j. +Let w1 := #|W1|. +Let w2 := #|W2|. + +Lemma cyclicTIirrC (xdefW : W2 \x W1 = W) i j : + cyclicTIirr xdefW j i = w_ i j. +Proof. by rewrite (dprod_IirrC xdefW defW). Qed. + +Lemma cycTIirrP chi : chi \in irr W -> {i : Iirr W1 & {j | chi = w_ i j}}. +Proof. +case/irrP/sig_eqW=> k ->{chi}. +by have /codomP/sig_eqW[[i j] ->] := dprod_Iirr_onto defW k; exists i, j. +Qed. + +Lemma cycTIirr_aut u i j : w_ (aut_Iirr u i) (aut_Iirr u j) = cfAut u (w_ i j). +Proof. by rewrite /w_ !dprod_IirrE cfAutDprod !aut_IirrE. Qed. + +Let sW1W : W1 \subset W. Proof. by have /mulG_sub[] := dprodW defW. Qed. +Let sW2W : W2 \subset W. Proof. by have /mulG_sub[] := dprodW defW. Qed. + +Lemma card_cycTIset : #|V| = (w1.-1 * w2.-1)%N. +Proof. +have [_ _ _ tiW12] := dprodP defW. +rewrite cardsD (setIidPr _) ?subUset ?sW1W // cardsU {}tiW12 cards1. +rewrite -(dprod_card defW) -addnBA // -!subn1 -/w1 -/w2 subnDA. +by rewrite mulnBl mulnBr mul1n muln1. +Qed. + +Definition cfCyclicTIset i j := cfDprod defW (1 - 'chi_i) (1 - 'chi_j). +Local Notation alpha_ := cfCyclicTIset. + +Lemma cycTIirr00 : w_ 0 0 = 1. Proof. by rewrite /w_ dprod_Iirr0 irr0. Qed. +Local Notation w_00 := cycTIirr00. + +Lemma cycTIirr_split i j : w_ i j = w_ i 0 * w_ 0 j. +Proof. by rewrite /w_ !dprod_IirrE !irr0 cfDprod_split. Qed. + +Lemma cfker_cycTIl j : W1 \subset cfker (w_ 0 j). +Proof. by rewrite /w_ dprod_IirrE irr0 cfDprod_cfun1l cfker_sdprod. Qed. + +Lemma cfker_cycTIr i : W2 \subset cfker (w_ i 0). +Proof. by rewrite /w_ dprod_IirrE irr0 cfDprod_cfun1r cfker_sdprod. Qed. + +Let cfdot_w i1 j1 i2 j2 : '[w_ i1 j1, w_ i2 j2] = ((i1 == i2) && (j1 == j2))%:R. +Proof. exact: cfdot_dprod_irr. Qed. + +Lemma cfCycTI_E i j : alpha_ i j = 1 - w_ i 0 - w_ 0 j + w_ i j. +Proof. +rewrite -w_00 -[w_ i j]opprK /w_ !dprod_IirrE !irr0 -addrA -opprD -!mulrBl. +by rewrite -mulrBr -!rmorphB. +Qed. +Local Notation alphaE := cfCycTI_E. + +Lemma cfCycTI_vchar i j : alpha_ i j \in 'Z[irr W]. +Proof. by rewrite alphaE rpredD ?rpredB ?rpred1 ?irr_vchar. Qed. + +Definition cfCyclicTIsetBase := + [seq alpha_ ij.1 ij.2 | ij in setX [set~ 0] [set~ 0]]. +Local Notation cfWVbase := cfCyclicTIsetBase. + +Let cfdot_alpha_w i1 j1 i2 j2 : + i2 != 0 -> j2 != 0 -> '[alpha_ i1 j1, w_ i2 j2] = [&& i1 == i2 & j1 == j2]%:R. +Proof. +move=> nzi2 nzj2; rewrite alphaE -w_00 !cfdotDl !cfdotNl !cfdot_w. +by rewrite !(eq_sym 0) (negPf nzi2) (negPf nzj2) /= andbF !subr0 add0r. +Qed. + +Let cfdot_alpha_1 i j : i != 0 -> j != 0 -> '[alpha_ i j, 1] = 1. +Proof. +move=> nzi nzj; rewrite alphaE -w_00 !cfdotDl !cfdotNl !cfdot_w. +by rewrite !eqxx andbT /= (negPf nzi) (negPf nzj) addr0 !subr0. +Qed. + +Let cfnorm_alpha i j : i != 0 -> j != 0 -> '[alpha_ i j] = 4%:R. +Proof. +move=> nzi nzj; rewrite -[4]/(size [:: 1; - w_ i 0; - w_ 0 j; w_ i j]). +rewrite -cfnorm_orthonormal 3?big_cons ?big_seq1 ?addrA -?alphaE //. +rewrite /orthonormal -w_00 /= !cfdotNl !cfdotNr !opprK !oppr_eq0 !cfnorm_irr. +by rewrite !cfdot_w !eqxx /= !(eq_sym 0) (negPf nzi) (negPf nzj) !eqxx. +Qed. + +Lemma cfCycTIbase_free : free cfWVbase. +Proof. +apply/freeP=> s /= s_alpha_0 ij; case def_ij: (enum_val ij) => [i j]. +have /andP[nzi nzj]: (i != 0) && (j != 0). + by rewrite -!in_setC1 -in_setX -def_ij enum_valP. +have:= congr1 (cfdotr (w_ i j)) s_alpha_0; rewrite raddf_sum raddf0 => <-. +rewrite (bigD1 ij) //= nth_image def_ij cfdotZl cfdot_alpha_w // !eqxx mulr1. +rewrite big1 ?addr0 // => ij1; rewrite nth_image -(inj_eq enum_val_inj) def_ij. +case: (enum_val ij1) => i1 j1 /= => ne_ij1_ij. +by rewrite cfdotZl cfdot_alpha_w // mulr_natr mulrb ifN. +Qed. + +(* Further results on alpha_ depend on the assumption that W is cyclic. *) + +Hypothesis ctiW : cyclicTI_hypothesis G defW. + +Let cycW : cyclic W. Proof. by case: ctiW. Qed. +Let oddW : odd #|W|. Proof. by case: ctiW. Qed. +Let tiV : normedTI V G W. Proof. by case: ctiW. Qed. +Let ntV : V != set0. Proof. by case/andP: tiV. Qed. + +Lemma cyclicTIhyp_sym (xdefW : W2 \x W1 = W) : cyclicTI_hypothesis G xdefW. +Proof. by split; rewrite // /cyclicTIset setUC. Qed. + +Let cycW1 : cyclic W1. Proof. exact: cyclicS cycW. Qed. +Let cycW2 : cyclic W2. Proof. exact: cyclicS cycW. Qed. +Let coW12 : coprime w1 w2. Proof. by rewrite -(cyclic_dprod defW). Qed. + +Let Wlin k : 'chi[W]_k \is a linear_char. Proof. exact/irr_cyclic_lin. Qed. +Let W1lin i : 'chi[W1]_i \is a linear_char. Proof. exact/irr_cyclic_lin. Qed. +Let W2lin i : 'chi[W2]_i \is a linear_char. Proof. exact/irr_cyclic_lin. Qed. +Let w_lin i j : w_ i j \is a linear_char. Proof. exact: Wlin. Qed. + +Let nirrW1 : #|Iirr W1| = w1. Proof. exact: card_Iirr_cyclic. Qed. +Let nirrW2 : #|Iirr W2| = w2. Proof. exact: card_Iirr_cyclic. Qed. +Let NirrW1 : Nirr W1 = w1. Proof. by rewrite -nirrW1 card_ord. Qed. +Let NirrW2 : Nirr W2 = w2. Proof. by rewrite -nirrW2 card_ord. Qed. + +Lemma cycTI_nontrivial : W1 :!=: 1%g /\ W2 :!=: 1%g. +Proof. +apply/andP; rewrite -!cardG_gt1 -!(subn_gt0 1) !subn1 -muln_gt0. +by rewrite -card_cycTIset card_gt0. +Qed. + +Let ntW1 : W1 :!=: 1%g. Proof. by case: cycTI_nontrivial. Qed. +Let ntW2 : W2 :!=: 1%g. Proof. by case: cycTI_nontrivial. Qed. +Let oddW1 : odd w1. Proof. exact: oddSg oddW. Qed. +Let oddW2 : odd w2. Proof. exact: oddSg oddW. Qed. +Let w1gt2 : (2 < w1)%N. Proof. by rewrite odd_gt2 ?cardG_gt1. Qed. +Let w2gt2 : (2 < w2)%N. Proof. by rewrite odd_gt2 ?cardG_gt1. Qed. + +Let neq_w12 : w1 != w2. +Proof. +by apply: contraTneq coW12 => ->; rewrite /coprime gcdnn -(subnKC w2gt2). +Qed. + +Let cWW : abelian W. Proof. exact: cyclic_abelian. Qed. +Let nsVW : V <| W. Proof. by rewrite -sub_abelian_normal ?subsetDl. Qed. +Let sWG : W \subset G. Proof. by have [_ /subsetIP[]] := normedTI_P tiV. Qed. +Let sVG : V \subset G^#. Proof. by rewrite setDSS ?subsetU ?sub1G. Qed. + +Let alpha1 i j : alpha_ i j 1%g = 0. +Proof. by rewrite cfDprod1 !cfunE cfun11 lin_char1 // subrr mul0r. Qed. + +(* This first part of Peterfalvi (3.4) will be used in (4.10) and (13.9). *) +Lemma cfCycTI_on i j : alpha_ i j \in 'CF(W, V). +Proof. +apply/cfun_onP=> x; rewrite !inE negb_and negbK orbC. +case/or3P => [/cfun0->// | W1x | W2x]. + by rewrite -[x]mulg1 cfDprodE // !cfunE cfun11 lin_char1 ?subrr ?mulr0. +by rewrite -[x]mul1g cfDprodE // !cfunE cfun11 lin_char1 ?subrr ?mul0r. +Qed. + +(* This is Peterfalvi (3.4). *) +Lemma cfCycTIbase_basis : basis_of 'CF(W, V) cfWVbase. +Proof. +rewrite basisEfree cfCycTIbase_free /=. +have ->: \dim 'CF(W, V) = #|V| by rewrite dim_cfun_on_abelian ?subsetDl. +rewrite size_tuple cardsX !cardsC1 nirrW1 nirrW2 -card_cycTIset leqnn andbT. +by apply/span_subvP=> _ /imageP[[i j] _ ->]; apply: cfCycTI_on. +Qed. +Local Notation cfWVbasis := cfCycTIbase_basis. + +Section CyclicTIisoBasis. + +Import CyclicTIisoReflexion ssrint. + +Local Notation unsat := (@unsat gT G). +Local Notation O := (@O gT G). +Local Notation "#1" := (inord 1). + +(* This is the combinatorial core of Peterfalvi (3.5.2). *) +(* Peterfalvi uses evaluation at 1%g to conclude after the second step; since *) +(* this is not covered by our model, we have used the dot product constraints *) +(* between b12 and b11, b21 instead. *) +Let unsat_J : unsat |= & x1 in b11 & -x1 in b21. +Proof. +uwlog b11x1: (& b11 = x1 + x2 + x3) by do 2!fill b11. +uwlog b21x1: (& b21 = -x1 + x2 + x3) by uhave x2, x3 in b21 as O(21, 11). +consider b12; uhave -x2 | x2 | ~x2 in b12. +- by uhave x1 in b12 as O(12, 11); counter to O(12, 21). +- uhave x1 | ~x1 in b12 as O(12, 21). + by uhave ~x3 in b12 as O(12, 21); counter to O(12, 11). + by uhave ~x3 in b12 as O(12, 11); counter to O(12, 21). +uhave x3 | ~x3 in b12 as O(12, 11). + by uhave x1 in b12 as O(12, 21); counter to O(12, 11). +by uhave x1 in b12 as O(12, 11); counter to O(12, 21). +Qed. + +Let unsat_II: unsat |= & x1, x2 in b11 & x1, x2 in b21. +Proof. by fill b11; uhave -x3 in b21 as O(21, 11); symmetric to unsat_J. Qed. + +(* This reflects the application of (3.5.2), but only to rule out nonzero *) +(* components of the first entry that conflict with positive components of *) +(* the second entry; Otest covers all the other uses of (3.5.2) in the proof. *) +Let Ltest (cl1 cl2 : clause) := + let: (i1, j1, kvs1) := cl1 in let: (i2, j2, kvs2) := cl2 in + let fix loop mm kvs2' := + if kvs2' is (k, v2) :: kvs2'' then + let v1 := odflt 0 (get_lit k kvs1) in + if (v2 != 1) || (v1 == 0) then loop mm kvs2'' else + if (v1 != 1) || mm then false else loop true kvs2'' + else true in + (i1 == i2) (+) (j1 == j2) ==> loop false kvs2. + +Let L ij12 : is_sat_test G (sat_test Ltest ij12). +Proof. +apply: sat_testP => m th [[i1 j1] kvs1] [[i2 j2] kvs2] m_th th_cl1 th_cl2. +wlog eq_j: m th i1 i2 j1 j2 m_th th_cl1 th_cl2 / j1 == j2. + move=> IH; case eq_j: (j1 == j2); first exact: IH m_th th_cl1 th_cl2 eq_j. + case eq_i: (i1 == i2); last by rewrite /= eq_i eq_j. + have /(_ (_, _, _)) mem_trt: _ \in tr_th th := map_f _ _. + by rewrite /= addbC; apply: IH (sat_tr m_th) _ _ eq_i; rewrite ?mem_trt. +apply/implyP; rewrite eq_j addbT => neq_i. +rewrite -[f in f _ kvs2]/(idfun _); set f := idfun _; rewrite /= in f *. +have [/= _ Ukvs2 kvsP] := satP m_th _ th_cl2. +move: Ukvs2; set kvs2' := kvs2; set mm := false. +have /allP: {subset kvs2' <= kvs2} by []. +pose lit12 k := (k, 1) \in kvs1 /\ (k, 1) \in kvs2. +have: mm -> {k | lit12 k & k \notin unzip1 kvs2'} by []. +elim: kvs2' mm => [|[k v2] kvs2' IH] //= mm mmP /andP[kvs2k /IH{IH}IHkvs]. +case/andP=> kvs2'k /IHkvs{IHkvs}IHkvs; case: ifP => [_ | /norP[]]. + by apply/IHkvs=> /mmP[kv kvs12kv /norP[]]; exists kv. +have [v1 /= kvs1k | //] := get_litP; case: eqP => // -> in kvs2k * => _ nz_v1. +case Dbb: (th_bbox th) (th_bboxP (bbox_refl (th_bbox th))) => [ri rj] rijP. +have [/andP[/=lti1r ltj1r] /andP[/=lti2r _]] := (rijP _ th_cl1, rijP _ th_cl2). +have rkP := th_dimP (leqnn _) _ th_cl1; have /= ltkr := rkP _ kvs1k. +have symP := unsat_match (Sym [:: i2; i1] [:: j1] _) _ _ m m_th. +rewrite /= Dbb lti1r lti2r ltj1r inE eq_sym neq_i /= in symP. +have [Dv1 | v1_neq1] /= := altP eqP; first rewrite Dv1 in kvs1k. + case: ifP => [/mmP[k0 [kvs1k0 kvs2k0]] | _]; last by apply: IHkvs; exists k. + case/norP=> k'k0; have [/=] := symP [:: k0; k] _ _ unsat_II. + rewrite inE k'k0 ltkr (rkP _ kvs1k0) /= andbT; apply/andP; split; apply/hasP. + by exists (i1, j1, kvs1) => //=; rewrite eqxx kvs1k kvs1k0. + by exists (i2, j2, kvs2) => //=; rewrite (eqP eq_j) eqxx kvs2k kvs2k0. +have{nz_v1 v1_neq1} Dv1: v1 = -1; last rewrite Dv1 in kvs1k. + by case: (v1) nz_v1 v1_neq1 (norm_lit m_th th_cl1 kvs1k) => [[|[]] | []]. +have[] := symP [:: k] _ _ unsat_J; rewrite /= ltkr !andbT /=; apply/andP; split. + by apply/hasP; exists (i1, j1, kvs1); rewrite //= eqxx kvs1k. +by apply/hasP; exists (i2, j2, kvs2); rewrite //= (eqP eq_j) eqxx kvs2k. +Qed. + +(* This is the combinatorial core of Peterfalvi (3.5.4). *) +(* We have made a few simplifications to the combinatorial analysis in the *) +(* text: we omit the (unused) step (3.5.4.4) entirely, which lets us inline *) +(* step (3.5.4.1) in the proof of (3.5.4.2); we clear the assumptions on b31 *) +(* and b32 before the final step (3.5.4.5), exposing a hidden symmetry. *) +Let unsat_Ii : unsat |= & x1 in b11 & x1 in b21 & ~x1 in b31. +Proof. +uwlog Db11: (& b11 = x1 + x2 + x3) by do 2!fill b11. +uwlog Db21: (& b21 = x1 + x4 + x5). + by uhave ~x2, ~x3 in b21 as L(21, 11); do 2!fill b21; uexact Db21. +uwlog Db31: (& b31 = x2 + x4 + x6). + uwlog b31x2: x2 | ~x2 in b31 as L(31, 11). + by uhave x3 in b31 as O(31, 11); symmetric to b31x2. + uwlog b31x4: x4 | ~x4 in b31 as L(31, 21). + by uhave x5 in b31 as O(31, 21); symmetric to b31x4. + uhave ~x3 in b31 as O(31, 11); uhave ~x5 in b31 as L(31, 21). + by fill b31; uexact Db31. +consider b41; uwlog b41x1: x1 | ~x1 in b41 as L(41, 11). + uwlog Db41: (& b41 = x3 + x5 + x6) => [|{b41x1}]. + uhave ~x2 | x2 in b41 as L(41, 11); last symmetric to b41x1. + uhave ~x4 | x4 in b41 as L(41, 21); last symmetric to b41x1. + uhave x3 in b41 as O(41, 11); uhave x5 in b41 as O(41, 21). + by uhave x6 in b41 as O(41, 31); uexact Db41. + consider b12; uwlog b12x1: x1 | ~x1 in b12 as L(12, 11). + uhave ~x2 | x2 in b12 as L(12, 11); last symmetric to b12x1. + by uhave x3 in b12 as O(12, 11); symmetric to b12x1. + uwlog b12x4: -x4 | ~x4 in b12 as O(12, 21). + by uhave -x5 in b12 as O(12, 21); symmetric to b12x4. + uhave ~x2, ~x3 in b12 as L(12, 11); uhave ~x5 in b12 as O(12, 21). + by uhave x6 in b12 as O(12, 31); counter to O(12, 41). +uwlog Db41: (& b41 = x1 + x6 + x7). + uhave ~x2, ~x3 in b41 as L(41, 11); uhave ~x4, ~x5 in b41 as L(41, 21). + by uhave x6 in b41 as O(41, 31); fill b41; uexact Db41. +consider b32; uwlog Db32: (& b32 = x6 - x7 + x8). + uwlog b32x6: x6 | ~x6 in b32 as L(32, 31). + uhave ~x2 | x2 in b32 as L(32, 31); last symmetric to b32x6. + by uhave x4 in b32 as O(32, 31); symmetric to b32x6. + uhave ~x2, ~x4 in b32 as L(32, 31). + uhave -x7 | ~x7 in b32 as O(32, 41). + uhave ~x1 in b32 as O(32, 41); uhave ~x3 in b32 as O(32, 11). + by uhave ~x5 in b32 as O(32, 21); fill b32; uexact Db32. + uhave -x1 in b32 as O(32, 41). + by uhave x3 in b32 as O(32, 11); counter to O(32, 21). +consider b42; uwlog Db42: (& b42 = x6 - x4 + x5). + uhave ~x6 | x6 in b42 as L(42, 41). + uhave ~x7 | x7 in b42 as L(42, 41); last counter to O(42, 32). + uhave x1 in b42 as O(42, 41); uhave x8 in b42 as O(42, 32). + uhave ~x2 | -x2 in b42 as O(42, 11); last counter to O(42, 21). + by uhave -x3 in b42 as O(42, 11); counter to O(42, 21). + uwlog b42x4: -x4 | ~x4 in b42 as O(42, 31). + by uhave -x2 in b42 as O(42, 31); symmetric to b42x4. + by uhave ~x1 in b42 as L(42, 41); uhave x5 in b42 as O(42, 21); uexact Db42. +uwlog Db32: (& ? in b32); first uexact Db32. +uwlog Db41: (& ? in b41); first uexact Db41. +consider b12; uwlog b12x5: x5 | ~x5 in b12 as L(12, 42). + uhave ~x6 | x6 in b12 as L(12, 42); last by consider b22; symmetric to b12x5. + uhave -x4 in b12 as O(12, 42); uhave x1 in b12 as O(12, 21). + by uhave ~x2 in b12 as L(12, 11); counter to O(12, 31). +uhave ~x6 in b12 as L(12, 42); uhave ~x4 in b12 as O(12, 42). +uhave ~x2 in b12 as O(12, 31). +by uhave -x1 in b12 as O(12, 21); counter to L(12, 11). +Qed. + +Let unsat_C : unsat |= & x1 in b11 & x1 in b21 & x1 in b12. +Proof. +consider b31; uwlog Db21: (& b21 = x1 + x2 + x3) by do 2!fill b21. +uwlog Db12: (& b12 = x1 - x2 + x4). + uwlog b21x2: -x2 | ~x2 in b12 as O(12, 21). + by uhave -x3 in b12 as O(12, 21); symmetric to b21x2. + by uhave ~x3 in b12 as O(12, 21); fill b12; uexact Db12. +uwlog Db31: (& b31 = x1 - x4 + x5). + uhave x1 | ~x1 in b31 as L(31, 21); last uexact unsat_Ii. + uhave ~x2, ~x3 in b31 as L(31, 21). + by uhave -x4 in b31 as O(31, 12); fill b31; uexact Db31. +consider b41; uhave x1 | ~x1 in b41 as L(41, 21); last symmetric to unsat_Ii. +uhave ~x5 in b41 as L(41, 31); uhave ~x4 in b41 as O(41, 31). +by uhave ~x2 in b41 as L(41, 21); counter to O(41, 12). +Qed. + +(* This refinement of Peterfalvi (3.5.4) is the essential part of (3.5.5). *) +Let column_pivot (m : model G) (j0 : 'I_m.2.+1) : + exists dk, forall (i : 'I_m.1.+1) (j : 'I_m.2.+1), + j0 != 0 -> i != 0 -> j != 0 -> '[m (i.-1, j.-1), dchi dk] = (j == j0)%:R. +Proof. +pose t_i (i0 i1 : nat) := [eta id with i0 |-> i1, i1 |-> i0]. +pose t_ij i0 i1 ij : ref := (t_i i0 i1 ij.1, ij.2). +have t_iK i0 i1: involutive (t_i i0 i1). + move=> i /=; have [-> | i0'i] := altP (i =P i0). + by rewrite eqxx; case: eqP. + by have [-> | /negPf->] := altP (i =P i1); rewrite ?eqxx // ifN. +have lt_t_i i0 i1 ri i: (i0 <= i1 < ri)%N -> (t_i i0 i1 i < ri)%N = (i < ri)%N. + case/andP=> le_i01 lti1 /=. + by do 2?case: eqP => [-> | _] //; rewrite ?(leq_trans _ lti1). +have t_mP i0 i1 (m0 : model G): + (i0 <= i1 < m0.1)%N -> is_Lmodel m0 (m0 \o t_ij i0 i1). +- have [lbm0 Zm0 Dm0] := LmodelP m0; split=> //= ij1 ij2 wf_ij1 wf_ij2. + by rewrite Dm0 /dot_ref ?(can_eq (t_iK _ _)) // !inE ?lt_t_i. +pose t_m i0 i1 m0 lti01 := Model (t_mP i0 i1 m0 lti01) (RmodelP m0). +without loss suffices{j0 lt_t_i} IHm: m / + exists dk, {in wf_ref m, forall ij, '[m ij, dchi dk] = (ij.2 == 0%N)%:R}. +- have [_ | nzj0] := altP eqP; first by exists (dirr1 G). + have ltj0: (j0.-1 < m.2)%N by rewrite prednK ?lt0n ?leq_ord. + have{IHm} [dk Ddk] := IHm (tr_model (t_m 0%N j0.-1 (tr_model m) ltj0)). + exists dk => i j _ nzi nzj; rewrite -[j.-1](t_iK 0%N j0.-1). + rewrite (Ddk (_, _)) ?inE ?lt_t_i // ?prednK ?lt0n ?leq_ord //. + by rewrite (inv_eq (t_iK _ _)) -eqSS !prednK ?lt0n. +pose cl11 := & b11 = x1 + x2 + x3. +without loss m_th: m / sat m |= cl11 & ? in b21. + move=> IHm; suffices{IHm}: sat m |= & ? in b11 & ? in b21. + have fill_b11 := sat_fill _ (mem_nth cl11 (_ : 1 < _))%N. + by do 3![case/fill_b11=> // ?]; apply: IHm. + have [[_ _ m1gt2 /ltnW m2gt0 _] _ _] := LmodelP m. + by rewrite /sat /= -!andbA /= m2gt0 -(subnKC m1gt2). +without loss{m_th} m_th: m / sat m |= & x1 in b11 & x1 in b21. + pose sat123P := @allP _ (fun k => sat_lit m _ (k, _)) (rev (iota 0 3)). + have [m123 | ] := altP (sat123P b21 0). + suffices: sat m |= cl11 & ~x1, ~x2, ~x3 in b21 by move/(O(21, 11)). + by rewrite /sat /= {1}/sat_cl /= !m123. + case/allPn=> k k012 /negP nz_m21 IHm; rewrite -[sat_lit _ _ _]andbT in nz_m21. + have ltk3: (k < 3)%N by rewrite mem_rev mem_iota in k012. + have [[/andP[/allP/=n1m _] Zm] [_ /= m_gt2 _]] := (RmodelP m, and3P m_th). + have ltk := leq_trans ltk3 m_gt2. + have{n1m Zm} mkP: is_Rmodel [:: m`_k]. + by split=> [|_ /predU1P[->|//]]; rewrite /orthonormal /= ?n1m ?Zm ?mem_nth. + pose mk := Model (LmodelP m) mkP; apply: {IHm}(IHm mk). + have{m_th} [v lit_v m_th] := sat_cases k m_th (mem_head _ _) ltk. + suffices: sat mk |= & x1 in b11 & (Lit 1 v) in b21. + by case/or4P: lit_v m_th => // /eqP-> => [/and4P[] | | _ /(L(21,11))]. + have [m_bb _ m_b21 /sat123P m_b11 _] := and5P m_th. + by apply/and5P; split; rewrite // /sat_cl /= [sat_lit _ _ _]m_b11. +have /dIrrP[dk Ddk]: m`_0 \in dirr G. + have [[/andP[/allP n1m _] Zm] [_ m_gt0 _]] := (RmodelP m, and3P m_th). + by rewrite dirrE Zm ?[_ == 1]n1m ?mem_nth. +exists dk => [][i j] /andP[/= lti ltj]; apply/eqP. +suffices{dk Ddk}: sat_cl m (& (Lit 1 (j == 0))%N in (i, j)). + by rewrite /sat_cl /= andbT /sat_lit Ddk. +without loss{i lti} ->: m i ltj m_th / i = 0%N. + have [bb21_m m_gt0 m11_x1 m21_x1 _] := and5P m_th. + move=> IHi; suffices{IHi} m_i1_x1: sat_lit m (i, 0)%N x1 && true. + apply: (IHi (t_m 0%N i m lti) 0%N); rewrite /sat /sat_cl //= bb21_m m_gt0. + by rewrite /= m_i1_x1 /sat_lit /= andbT /t_ij /=; case: ifP. + case i_gt1: (1 < i)%N; last by case: (i) i_gt1 => [|[|[]]]. + have itv_i: (1 < i < m.1)%N by [apply/andP]; pose m2 := t_m 2 i m itv_i. + have m2_th: sat m2 |= & x1 in b11 & x1 in b21 & ? in b31. + rewrite /sat m_gt0 -andbA (leq_trans _ lti) ?(leq_trans _ ltj) /sat_cl //=. + by rewrite /sat_lit /= -(subnKC i_gt1); have [_ _] := and3P m_th. + have [v] := sat_cases _ m2_th (mem_head _ _) m_gt0; rewrite !inE. + by case/or3P=> /eqP-> => [/unsat_Ii | /and4P[] | /(L(31,11))]. +have [-> | nzj] := posnP j; first by case/and5P: m_th. +without loss{ltj nzj} ->: m j m_th / j = 1%N. + have itv_j: (0 < j < m.2)%N by rewrite nzj. + move/(_ (tr_model (t_m _ j (tr_model m) itv_j)) _ _ (erefl _)) => /=. + by rewrite /sat /sat_cl /sat_lit /= -(prednK nzj) => ->. +have{m_th}[/= _ m_gt0 m_x1] := and3P m_th. +have{m_x1} m_th: sat m |= & x1 in b11 & x1 in b21 & ? in b12. + by rewrite /sat m_gt0 /sub_bbox; have [[_ _ -> ->]] := LmodelP m. +have [v] := sat_cases 0%N m_th (mem_head _ _) m_gt0; rewrite !inE. +by case/or3P=> /eqP-> => [/and4P[] | /unsat_C | /(L(12,11))]. +Qed. + +(* This is Peterfalvi (3.5). *) +(* We have inlined part of the proof of (3.5.5) in this main proof, replacing *) +(* some combinatorial arguments with direct computation of the dot product, *) +(* this avoids the duplicate case analysis required to exploit (3.5.5) as it *) +(* is stated in the text. *) +Lemma cyclicTIiso_basis_exists : + exists xi_ : Iirr W1 -> Iirr W2 -> 'CF(G), + [/\ xi_ 0 0 = 1, forall i j, xi_ i j \in 'Z[irr G], + forall i j, i != 0 -> j != 0 -> + 'Ind (alpha_ i j) = 1 - xi_ i 0 - xi_ 0 j + xi_ i j + & forall i1 j1 i2 j2, '[xi_ i1 j1, xi_ i2 j2] = ((i1, j1) == (i2, j2))%:R]. +Proof. +(* Instantiate the abstract theory vertically and horizontally. *) +pose beta i j : 'CF(G) := 'Ind[G] (alpha_ i j) - 1. +have Zbeta i j: beta i j \in 'Z[irr G]. + by rewrite rpredB ?rpred1 ?cfInd_vchar ?cfCycTI_vchar. +have o_alphaG_1 i j: i != 0 -> j != 0 -> '['Ind[G] (alpha_ i j), 1] = 1. + by move=> nz_i nz_j; rewrite -cfdot_Res_r rmorph1 cfdot_alpha_1. +have o_beta_1 i j: i != 0 -> j != 0 -> '[beta i j, 1] = 0. + by move=> nzi nzj; rewrite cfdotBl o_alphaG_1 // cfnorm1 subrr. +have o_beta i1 j1 i2 j2 : i1 != 0 -> j1 != 0 -> i2 != 0 -> j2 != 0 -> + '[beta i1 j1, beta i2 j2] = ((i1 == i2).+1 * (j1 == j2).+1 - 1)%:R. +- move=> nzi1 nzj1 nzi2 nzj2; rewrite mulSnr addnS mulnSr /=. + rewrite cfdotBr o_beta_1 // subr0 cfdotBl (cfdotC 1) o_alphaG_1 //. + rewrite (normedTI_isometry tiV) ?cfCycTI_on // rmorph1 addrC. + rewrite (alphaE i2) cfdotDr !cfdotBr cfdot_alpha_1 // -!addrA addKr addrA. + rewrite addrC cfdot_alpha_w // subn1 -addnA !natrD mulnb; congr (_ + _). + rewrite alphaE -w_00 !(cfdotBl, cfdotDl) !cfdot_w !eqxx !(eq_sym 0). + rewrite (negPf nzi1) (negPf nzj1) (negPf nzi2) (negPf nzj2) /= !andbF !andbT. + by rewrite !addr0 !subr0 !opprB !subr0. +pose beta_fun := [fun ij => beta (inord ij.1.+1) (inord ij.2.+1)]. +have beta_modelP: is_Lmodel ((Nirr W1).-1, (Nirr W2).-1) beta_fun. + split=> [ | //= | ij1 ij2 /=/andP[lti1 ltj1] /andP[lti2 ltj2]]. + by rewrite -!(ltnS 2) -eqSS NirrW1 NirrW2. + by rewrite o_beta -?val_eqE /= ?inordK. +pose beta_model := Model beta_modelP (nil_RmodelP G). +have betaE i j: i != 0 -> j != 0 -> beta i j = beta_fun (i.-1, j.-1). + by move=> nzi nzj /=; rewrite !prednK ?lt0n ?inord_val. +have /fin_all_exists [dXi0 betaXi0] i0: exists dX, i0 != 0 -> + forall i j, i != 0 -> j != 0 -> '[beta i j, dchi dX] = (i == i0)%:R. +- have [/= dX DdX] := @column_pivot (tr_model beta_model) i0. + by exists dX => nzi0 i j nzi nzj; rewrite betaE ?DdX. +have /fin_all_exists [dX0j betaX0j] j0: exists dX, j0 != 0 -> + forall i j, i != 0 -> j != 0 -> '[beta i j, dchi dX] = (j == j0)%:R. +- have [dX DdX] := @column_pivot beta_model j0. + by exists dX => nzj0 i j nzi nzj; rewrite betaE ?DdX. +pose Xi0 j := dchi (dXi0 j); pose X0j i := dchi (dX0j i). +(* Construct the orthonormal family xi_ i j. *) +pose xi_ i j := if i == 0 then if j == 0 then 1 else - X0j j else + if j == 0 then - Xi0 i else beta i j - Xi0 i - X0j j. +exists xi_; split=> [| i j | i j nzi nzj | i1 j1 i2 j2]. +- by rewrite /xi_ !eqxx. +- rewrite /xi_; do 2!case: ifP => _; rewrite ?rpred1 ?rpredN ?dchi_vchar //. + by rewrite 2?rpredB ?dchi_vchar. +- by rewrite /xi_ /= !ifN // addrCA subrK addrACA subrK addrA addrK. +have o_dchi i j dk1 dk2 (phi := beta i j): + '[phi, dchi dk1] = 1 -> '[phi, dchi dk2] = 0 -> '[dchi dk1, dchi dk2] = 0. +- move=> phi1 phi0; have /eqP: 1 != 0 :> algC := oner_neq0 _. + rewrite -phi1 cfdot_dchi; do 2!case: eqP => [->|_]; rewrite ?subrr //. + by rewrite dchi_ndirrE cfdotNr phi0 oppr0. +have [nzi01 nzj01] := (Iirr1_neq0 ntW1, Iirr1_neq0 ntW2). +have X0j_1 j: j != 0 -> '[X0j j, 1] = 0. + by move=> nzj; rewrite -dchi1 (o_dchi #1 j) ?betaX0j ?eqxx ?dchi1 ?o_beta_1. +have Xi0_1 i: i != 0 -> '[Xi0 i, 1] = 0. + by move=> nzi; rewrite -dchi1 (o_dchi i #1) ?betaXi0 ?eqxx ?dchi1 ?o_beta_1. +have Xi0_X0j i j: i != 0 -> j != 0 -> '[Xi0 i, X0j j] = 0. + move=> nzi nzj; pose j' := conjC_Iirr j. + apply: (o_dchi i j'); rewrite (betaX0j, betaXi0) ?conjC_Iirr_eq0 ?eqxx //. + by rewrite -(inj_eq irr_inj) conjC_IirrE mulrb ifN ?odd_eq_conj_irr1 ?irr_eq1. +have X0j_X0j j j0: j != 0 -> j0 != 0 -> '[X0j j, X0j j0] = (j == j0)%:R. + move=> nzj nzj0; case: (altP eqP) => [-> | j0'j]; first exact: cfnorm_dchi. + by apply: (o_dchi #1 j); rewrite ?betaX0j ?eqxx ?(negPf j0'j). +have Xi0_Xi0 i i0: i != 0 -> i0 != 0 -> '[Xi0 i, Xi0 i0] = (i == i0)%:R. + move=> nzi nzi0; case: (altP eqP) => [-> | i0'i]; first exact: cfnorm_dchi. + by apply: (o_dchi i #1); rewrite ?betaXi0 ?eqxx ?(negPf i0'i). +have oxi_00 i j: '[xi_ i j, xi_ 0 0] = ((i == 0) && (j == 0))%:R. + rewrite /xi_; case: ifPn => [_ | nzi]. + by case: ifPn => [_ | nzj]; rewrite ?cfnorm1 // cfdotNl X0j_1 ?oppr0. + case: ifPn => [_ | nzj]; first by rewrite cfdotNl Xi0_1 ?oppr0. + by rewrite 2!cfdotBl o_beta_1 ?X0j_1 ?Xi0_1 ?subr0. +have oxi_0j i j j0: '[xi_ i j, xi_ 0 j0] = ((i == 0) && (j == j0))%:R. + rewrite /xi_; have [-> | nzj0] := altP (j0 =P 0); first exact: oxi_00. + rewrite cfdotNr; case: ifPn => [_ | nzi]. + have [-> | nzj] := altP eqP; last by rewrite cfdotNl opprK X0j_X0j. + by rewrite cfdotC X0j_1 // conjC0 oppr0 mulrb ifN_eqC. + have [_ | nzj] := ifPn; first by rewrite cfdotNl Xi0_X0j ?oppr0. + by rewrite 2!cfdotBl Xi0_X0j // subr0 betaX0j ?X0j_X0j // subrr oppr0. +have{oxi_00} oxi_i0 i j i0: '[xi_ i j, xi_ i0 0] = ((i == i0) && (j == 0))%:R. + rewrite /xi_; have [-> | nzi0] := altP (i0 =P 0); first exact: oxi_00. + rewrite cfdotNr andbC; have [_ | nzj] := boolP (j == 0). + have [-> | nzi] := altP eqP; last by rewrite cfdotNl opprK Xi0_Xi0. + by rewrite cfdotC Xi0_1 // conjC0 oppr0 mulrb ifN_eqC. + have [_ | nzi] := ifPn; first by rewrite cfdotNl opprK cfdotC Xi0_X0j ?conjC0. + rewrite 2!cfdotBl betaXi0 ?Xi0_Xi0 // subrr add0r opprK. + by rewrite cfdotC Xi0_X0j // conjC0. +have [-> | nzi2] := altP (i2 =P 0); first exact: oxi_0j. +have [-> | nzj2] := altP (j2 =P 0); first exact: oxi_i0. +rewrite cfdotC eq_sym; apply: canLR conjCK _; rewrite rmorph_nat. +have [-> | nzi1] := altP (i1 =P 0); first exact: oxi_0j. +have [-> | nzj1] := altP (j1 =P 0); first exact: oxi_i0. +have ->: xi_ i1 j1 = beta i1 j1 + xi_ i1 0 + xi_ 0 j1 by rewrite /xi_ !ifN. +rewrite 2!cfdotDr oxi_i0 oxi_0j andbC /xi_ (negPf nzi2) (negPf nzj2) !addr0. +rewrite eq_sym xpair_eqE cfdotC 2!cfdotBr o_beta // betaXi0 ?betaX0j //. +by rewrite -!CintrE /= rmorph_int; do 2!case: (_ == _). +Qed. + +End CyclicTIisoBasis. + +(* This is PeterFalvi, Theorem (3.2)(a, b, c). *) +Theorem cyclicTIiso_exists : + {sigma : 'Hom(cfun_vectType W, cfun_vectType G) | + [/\ {in 'Z[irr W], isometry sigma, to 'Z[irr G]}, sigma 1 = 1 + & {in 'CF(W, V), forall phi : 'CF(W), sigma phi = 'Ind[G] phi}]}. +Proof. +pose sigmaVP f := ('CF(W, V) <= lker (linfun f - linfun 'Ind[G]))%VS. +pose sigmaP f := [&& orthonormal (map f (irr W)), f 1 == 1 & sigmaVP f]. +pose sigma_base f := [seq (dchi (f k) : 'CF(G)) | k : Iirr W]. +pose sigma_spec f := sigmaP (sval (linear_of_free (irr W) (sigma_base f))). +suffices /sigW[f /and3P[]]: exists f : {ffun _}, sigma_spec f. + case: linear_of_free => /=sigma Dsigma o1sigma /eqP sigma1 /eqlfun_inP sigmaV. + exists (linfun sigma); split=> [|| phi /sigmaV]; try by rewrite !lfunE. + do [rewrite size_map !size_tuple => /(_ (irr_free W) (card_ord _))] in Dsigma. + have [inj_sigma dot_sigma] := orthonormalP o1sigma. + rewrite -(map_tnth_enum (irr W)) -map_comp in Dsigma inj_sigma. + move/eq_in_map in Dsigma; move/injectiveP in inj_sigma. + split=> [|_ /zchar_tuple_expansion[z Zz ->]]. + apply: isometry_in_zchar=> _ _ /irrP[k1 ->] /irrP[k2 ->] /=. + by rewrite !lfunE dot_sigma ?map_f ?mem_irr // cfdot_irr (inj_eq inj_sigma). + rewrite linear_sum rpred_sum // => k _; rewrite linearZ rpredZ_Cint //=. + by rewrite -tnth_nth lfunE [sigma _]Dsigma ?mem_enum ?dchi_vchar. +have [xi_ [xi00 Zxi Dxi o1xi]] := cyclicTIiso_basis_exists. +pose f := [ffun k => dirr_dIirr (prod_curry xi_) (inv_dprod_Iirr defW k)]. +exists f; apply/and3P; case: linear_of_free => /= sigma Dsigma. +have{f Dsigma} Deta i j: sigma (w_ i j) = xi_ i j. + rewrite /w_ -tnth_map /= (tnth_nth 0) /=. + rewrite Dsigma ?irr_free //; last by rewrite !size_tuple card_ord. + rewrite nth_mktuple ffunE dprod_IirrK dirr_dIirrE // => {i j} [[i j]] /=. + by rewrite dirrE Zxi o1xi !eqxx. +have sigma1: sigma 1 = 1 by rewrite -w_00 Deta. +rewrite sigma1 /sigmaVP -(span_basis cfWVbasis); split=> //. + rewrite map_orthonormal ?irr_orthonormal //; apply: isometry_in_zchar. + move=> _ _ /cycTIirrP[i1 [j1 ->]] /cycTIirrP[i2 [j2 ->]] /=. + by rewrite !Deta o1xi cfdot_w. +apply/span_subvP=> _ /imageP[[i j] /setXP[nzi nzj] ->]; rewrite !inE in nzi nzj. +rewrite memv_ker !lfun_simp /= subr_eq0 Dxi //. +by rewrite alphaE linearD !linearB sigma1 !Deta. +Qed. + +Fact cyclicTIiso_key : unit. Proof. by []. Qed. +Definition cyclicTIiso := + locked_with cyclicTIiso_key (lfun_linear (sval cyclicTIiso_exists)). +Local Notation sigma := cyclicTIiso. +Let im_sigma := map sigma (irr W). +Let eta_ i j := sigma (w_ i j). + +Lemma cycTI_Zisometry : {in 'Z[irr W], isometry sigma, to 'Z[irr G]}. +Proof. by rewrite [sigma]unlock; case: cyclicTIiso_exists => ? []. Qed. + +Let Isigma : {in 'Z[irr W] &, isometry sigma}. +Proof. by case: cycTI_Zisometry. Qed. +Let Zsigma : {in 'Z[irr W], forall phi, sigma phi \in 'Z[irr G]}. +Proof. by case: cycTI_Zisometry. Qed. + +Lemma cycTIisometry : isometry sigma. +Proof. +move=> phi psi; have [[a ->] [b ->]] := (cfun_irr_sum phi, cfun_irr_sum psi). +rewrite !linear_sum !cfdot_suml; apply: eq_bigr => i _. +rewrite !cfdot_sumr; apply: eq_bigr => j _. +by rewrite !linearZ !cfdotZl !cfdotZr /= Isigma ?irr_vchar. +Qed. + +Lemma cycTIiso_vchar i j : eta_ i j \in 'Z[irr G]. +Proof. by rewrite Zsigma ?irr_vchar. Qed. + +Lemma cfdot_cycTIiso i1 i2 j1 j2 : + '[eta_ i1 j1, eta_ i2 j2] = ((i1 == i2) && (j1 == j2))%:R. +Proof. by rewrite cycTIisometry. Qed. + +Lemma cfnorm_cycTIiso i j : '[eta_ i j] = 1. +Proof. by rewrite cycTIisometry cfnorm_irr. Qed. + +Lemma cycTIiso_dirr i j : eta_ i j \in dirr G. +Proof. by rewrite dirrE cycTIiso_vchar /= cfnorm_cycTIiso. Qed. + +Lemma cycTIiso_orthonormal : orthonormal im_sigma. +Proof. by rewrite map_orthonormal ?irr_orthonormal. Qed. + +Lemma cycTIiso_eqE i1 i2 j1 j2 : + (eta_ i1 j1 == eta_ i2 j2) = ((i1 == i2) && (j1 == j2)). +Proof. +have /inj_in_eq-> := Zisometry_inj Isigma; try exact: irr_vchar. +by rewrite (inj_eq irr_inj) (inj_eq (dprod_Iirr_inj _)). +Qed. + +Lemma cycTIiso_neqN i1 i2 j1 j2 : (eta_ i1 j1 == - eta_ i2 j2) = false. +Proof. +rewrite -addr_eq0; apply/eqP=> /(congr1 (cfdot (eta_ i1 j1)))/eqP. +by rewrite cfdot0r cfdotDr !cfdot_cycTIiso !eqxx -mulrS pnatr_eq0. +Qed. + +Lemma cycTIiso1 : sigma 1 = 1. +Proof. by rewrite [sigma]unlock; case: cyclicTIiso_exists => ? []. Qed. + +Lemma cycTIiso_Ind : {in 'CF(W, V), forall phi, sigma phi = 'Ind[G, W] phi}. +Proof. by rewrite [sigma]unlock; case: cyclicTIiso_exists => ? []. Qed. + +Let sigma_Res_V : + [/\ forall phi, {in V, sigma phi =1 phi} + & forall psi : 'CF(G), orthogonal psi im_sigma -> {in V, psi =1 \0}]. +Proof. +have sigW i j : '[sigma 'chi_i, sigma 'chi_j] = (i == j)%:R. + by rewrite cycTIisometry cfdot_irr. +have [j | sigmaV sigma'V] := equiv_restrict_compl_ortho sWG nsVW cfWVbasis sigW. + rewrite /= -/cfWVbase -(eq_bigr _ (fun _ _ => linearZ _ _)) /= -linear_sum. + rewrite -cfun_sum_cfdot cycTIiso_Ind //. + by rewrite (basis_mem cfWVbasis) ?mem_nth ?size_image. +split=> [phi v Vv | psi /orthoPl o_psi_sigma]. + rewrite [phi]cfun_sum_cfdot linear_sum !sum_cfunE. + by apply: eq_bigr => k _; rewrite linearZ !cfunE sigmaV. +by apply: sigma'V => k; rewrite o_psi_sigma ?map_f ?mem_irr. +Qed. + +(* This is Peterfalvi, Theorem (3.2)(d). *) +Theorem cycTIiso_restrict phi : {in V, sigma phi =1 phi}. +Proof. by case: sigma_Res_V. Qed. + +(* This is Peterfalvi, Theorem (3.2)(e). *) +Theorem ortho_cycTIiso_vanish (psi : 'CF(G)) : + orthogonal psi im_sigma -> {in V, forall x, psi x = 0}. +Proof. by case: sigma_Res_V psi. Qed. + +(* This is PeterFalvi (3.7). *) +Lemma cycTIiso_cfdot_exchange (psi : 'CF(G)) i1 i2 j1 j2 : + {in V, forall x, psi x = 0} -> + '[psi, eta_ i1 j1] + '[psi, eta_ i2 j2] + = '[psi, eta_ i1 j2] + '[psi, eta_ i2 j1]. +Proof. +move=> psiV_0; pose phi : 'CF(W) := w_ i1 j1 + w_ i2 j2 - w_ i1 j2 - w_ i2 j1. +have Vphi: phi \in 'CF(W, V). + apply/cfun_onP=> g; rewrite inE negb_and negbK !inE orbC. + case/or3P=> [/cfun0-> // | W1g | W2g]; apply/eqP; rewrite !cfunE subr_eq0. + by rewrite addrC -[g]mulg1 /w_ !dprod_IirrE !cfDprodE ?lin_char1 ?addKr. + by rewrite -[g]mul1g /w_ !dprod_IirrE !cfDprodE ?lin_char1 ?addrK. +suffices: '[psi, 'Ind[G] phi] == 0. + rewrite -!cycTIiso_Ind // !linearB !linearD !cfdotBr !cfdotDr. + by rewrite -addrA -opprD subr_eq0 => /eqP. +rewrite (cfdotEr _ (cfInd_on sWG Vphi)) big1 ?mulr0 //. +by move=> _ /imset2P[x y Vx Gy ->]; rewrite cfunJ ?psiV_0 ?mul0r. +Qed. + +(* This is NC as defined in PeterFalvi (3.6). *) +Definition cyclicTI_NC phi := #|[set ij | '[phi, eta_ ij.1 ij.2] != 0]|. +Local Notation NC := cyclicTI_NC. + +Lemma cycTI_NC_opp (phi : 'CF(G)) : (NC (- phi)%R = NC phi)%N. +Proof. by apply: eq_card=> [[i j]]; rewrite !inE cfdotNl oppr_eq0. Qed. + +Lemma cycTI_NC_sign (phi : 'CF(G)) n : (NC ((-1) ^+ n *: phi)%R = NC phi)%N. +Proof. +elim: n=> [|n IH]; rewrite ?(expr0,scale1r) //. +by rewrite exprS -scalerA scaleN1r cycTI_NC_opp. +Qed. + +Lemma cycTI_NC_iso i j : NC (eta_ i j) = 1%N. +Proof. +rewrite -(cards1 (i, j)); apply: eq_card => [[i1 j1]]; rewrite !inE /=. +rewrite cfdot_cycTIiso //= pnatr_eq0 (can_eq oddb _ false) eqbF_neg negbK. +by rewrite -xpair_eqE eq_sym. +Qed. + +Lemma cycTI_NC_irr i : (NC 'chi_i <= 1)%N. +Proof. +apply: wlog_neg; rewrite -ltnNge => /ltnW/card_gt0P[[i1 j1]]. +rewrite inE cfdot_dirr ?(irr_dirr, cycTIiso_dirr) //=. +case: ('chi_i =P _) => [-> | _]; first by rewrite cycTI_NC_opp cycTI_NC_iso. +by case: ('chi_i =P _)=> [-> | _]; rewrite (cycTI_NC_iso, eqxx). +Qed. + +Lemma cycTI_NC_dirr f : f \in dirr G -> (NC f <= 1)%N. +Proof. by case/dirrP=> b [i ->]; rewrite cycTI_NC_sign cycTI_NC_irr. Qed. + +Lemma cycTI_NC_dchi di : (NC (dchi di) <= 1)%N. +Proof. by rewrite cycTI_NC_dirr ?dirr_dchi. Qed. + +Lemma cycTI_NC_0 : NC 0 = 0%N. +Proof. by apply: eq_card0 => ij; rewrite !inE cfdot0l eqxx. Qed. + +Lemma cycTI_NC_add n1 n2 phi1 phi2 : + (NC phi1 <= n1 -> NC phi2 <= n2 -> NC (phi1 + phi2)%R <= n1 + n2)%N. +Proof. +move=> ub1 ub2; apply: leq_trans {ub1 ub2}(leq_add ub1 ub2). +rewrite -cardsUI -[NC _]addn0 leq_add // subset_leq_card //. +apply/subsetP=> [[i j]]; rewrite !inE /= -negb_and cfdotDl. +by apply: contra => /andP[/eqP-> /eqP->]; rewrite addr0. +Qed. + +Lemma cycTI_NC_sub n1 n2 phi1 phi2 : + (NC phi1 <= n1 -> NC phi2 <= n2 -> NC (phi1 - phi2)%R <= n1 + n2)%N. +Proof. by move=> ub1 ub2; rewrite cycTI_NC_add ?cycTI_NC_opp. Qed. + +Lemma cycTI_NC_scale_nz a phi : a != 0 -> NC (a *: phi) = NC phi. +Proof. +move=> nz_a; apply: eq_card => ij. +by rewrite !inE cfdotZl mulf_eq0 negb_or nz_a. +Qed. + +Lemma cycTI_NC_scale a phi n : (NC phi <= n -> NC (a *: phi) <= n)%N. +Proof. +have [-> _ | /cycTI_NC_scale_nz-> //] := eqVneq a 0. +by rewrite scale0r cycTI_NC_0. +Qed. + +Lemma cycTI_NC_norm phi n : + phi \in 'Z[irr G] -> '[phi] <= n%:R -> (NC phi <= n)%N. +Proof. +move=> Zphi ub_phi; apply: leq_trans (_ : #|dirr_constt phi| <= n)%N. + rewrite {1}[phi]cfun_sum_dconstt // -sum1_card. + elim/big_rec2: _ => [|/= i n1 phi1 _]; first by rewrite cycTI_NC_0. + by apply: cycTI_NC_add; rewrite cycTI_NC_scale ?cycTI_NC_dchi. +rewrite -leC_nat (ler_trans _ ub_phi) ?cnorm_dconstt // -sumr_const. +apply: ler_sum => i phi_i; rewrite sqr_Cint_ge1 ?Cint_Cnat ?Cnat_dirr //. +by rewrite gtr_eqF -?dirr_consttE. +Qed. + +(* This is PeterFalvi (3.8). *) +Lemma small_cycTI_NC phi i0 j0 (a0 := '[phi, eta_ i0 j0]) : + {in V, forall x, phi x = 0} -> (NC phi < 2 * minn w1 w2)%N -> a0 != 0 -> + (forall i j, '[phi, eta_ i j] = (j == j0)%:R * a0) + \/ (forall i j, '[phi, eta_ i j] = (i == i0)%:R * a0). +Proof. +pose a i j := '[phi, eta_ i j]; pose A := [set ij | a ij.1 ij.2 != 0]. +rewrite -[NC phi]/#|A| ltnNge => phiV_0 ubA nz_a0. +have{phiV_0} Da i2 j2 i1 j1 : a i1 j1 = a i1 j2 + a i2 j1 - a i2 j2. + by rewrite cycTIiso_cfdot_exchange ?addrK. +have ubA2: ~~ (w2 + w1 <= #|A| + 2)%N. + rewrite addnC addn2 -ltnS (contra _ ubA) //; apply: (@leq_trans _ _.+3). + rewrite odd_geq /= ?odd_add ?oddW1 ?oddW2 // mul2n -addn_min_max -addnn. + by rewrite uphalf_double leq_add2l gtn_min !leq_max !ltnn orbF -neq_ltn. +(* This is step (3.8.1). *) +have Za i1 i2 j1 j2 : a i1 j2 == 0 -> a i2 j1 == 0 -> a i1 j1 == 0. + have [-> // | /negPf i2'1 /eqP Za12 /eqP Za21] := eqVneq i1 i2. + apply: contraR ubA2 => nz_a11. + pose L := [set (if a i1 j == 0 then i2 else i1, j) | j : Iirr W2]. + pose C := [set (i, if a i j1 == 0 then j2 else j1) | i : Iirr W1]. + have [<- <-]: #|L| = w2 /\ #|C| = w1 by rewrite !card_imset // => ? ? []. + have <-: #|[set (i1, j1); (i2, j2)]| = 2 by rewrite cards2 xpair_eqE i2'1. + rewrite -cardsUI leq_add ?subset_leq_card //; last first. + apply/subsetP=> _ /setIP[/imsetP[j _ ->] /imsetP[i _ []]]. + by case: ifP => _ <- ->; rewrite !inE ?Za21 ?(negPf nz_a11) !eqxx ?orbT. + apply/subsetP=> ij /setUP[] /imsetP[] => [j | i] _ {ij}->; rewrite inE. + by case: ifPn => // /eqP Za1j; rewrite (Da i1 j1) Za21 Za1j !add0r oppr_eq0. + by case: ifPn => // /eqP Zai1; rewrite (Da i1 j1) Za12 Zai1 !add0r oppr_eq0. +pose L i := [set ij | ij.1 == i] :&: A; pose C j := [set ij | ij.2 == j] :&: A. +have{ubA2} ubLC i j: (#|L i| + #|C j| != w2 + w1)%N. + apply: contraNneq ubA2 => <-; rewrite addnS leqW // -cardsUI -setIUl -setIIl. + rewrite -(card1 (i, j)) leq_add ?subset_leq_card ?subsetIr //. + by apply/subsetP=> ij /setIP[]; rewrite !inE. +have lbA L1 L2: L1 :&: L2 =i set0 -> (#|L1 :&: A| + #|L2 :&: A| <= #|A|)%N. + rewrite -cardsUI -setIUl -setIIl => /setP->. + by rewrite set0I cards0 addn0 subset_leq_card ?subsetIr. +have oL i1: ~~ [exists j, a i1 j == 0] -> #|L i1| = w2. + rewrite negb_exists => /forallP nz_a1. + transitivity #|predX (pred1 i1) (Iirr W2)|; last by rewrite cardX card1 mul1n. + by apply/eq_card=> ij; rewrite !inE andbT andb_idr // => /eqP->. +have oC i1 j1 j2 : a i1 j1 != 0 -> a i1 j2 == 0 -> #|C j1| = w1. + move=> nz_a11 /(Za i1)/contra/(_ nz_a11) nz_a1. + transitivity #|predX (Iirr W1) (pred1 j1)|; last by rewrite cardX card1 muln1. + by apply/eq_card=> ij; rewrite !inE andb_idr // => /eqP->. +(* This is step (3.8.2). *) +have [/existsP[j1 Za01] | /oL oL0] := boolP [exists j, a i0 j == 0]. + have j0'1 : j1 != j0 by apply: contraTneq Za01 => ->. + have oC0: #|C j0| = w1 by apply: oC nz_a0 Za01. + suffices Za0 i j: j != j0 -> a i j = 0. + left=> i j; rewrite -/(a i j) mulr_natl mulrb; have [->|/Za0//] := altP eqP. + by rewrite (Da i0 j1) !(Za0 _ j1) // subr0 add0r. + move=> j0'j; apply: contraNeq (ubLC i j0) => nz_aij; rewrite oC0 oL //. + apply: contra ubA => /existsP[_ /Za/contra/(_ nz_aij) nz_a_j]. + rewrite minn_mulr geq_min mul2n -addnn -{2}oC0 -(oC i0 j j1) ?lbA // => ij. + by rewrite !inE; apply/andP=> [[/eqP-> /idPn]]. +(* This is step (3.8.3). *) +suffices Za0 i j: i != i0 -> a i j = 0. + right=> i j; rewrite -/(a i j) mulr_natl mulrb; have [->|/Za0//] := altP eqP. + have /card_gt0P[i1 i0'i]: (0 < #|predC1 i0|)%N. + by rewrite cardC1 nirrW1 -(subnKC w1gt2). + by rewrite (Da i1 j0) !(Za0 i1) // subr0 addr0. +move=> i0'i; suffices /existsP[j1 Zai1]: [exists j, a i j == 0]. + by apply: contraNeq (ubLC i0 j) => /oC/(_ Zai1)->; rewrite oL0. +apply: contraR ubA; rewrite minn_mulr geq_min orbC mul2n -addnn => /oL{1}<-. +by rewrite -oL0 lbA // => ij; rewrite !inE; apply/andP=> [[/eqP-> /idPn]]. +Qed. + +(* A weaker version of PeterFalvi (3.8). *) +Lemma cycTI_NC_minn (phi : 'CF(G)) : + {in V, forall x, phi x = 0} -> (0 < NC phi < 2 * minn w1 w2)%N -> + (minn w1 w2 <= NC phi)%N. +Proof. +move=> phiV_0 /andP[/card_gt0P[[i0 j0]]]; rewrite inE /= => nz_a0 ubNC. +pose L := [seq (i0, j) | j : Iirr W2]; pose C := [seq (i, j0) | i : Iirr W1]. +have [oL oC]: #|L| = w2 /\ #|C| = w1 by rewrite !card_image // => i j []. +have [Da | Da] := small_cycTI_NC phiV_0 ubNC nz_a0. + rewrite geq_min -oC subset_leq_card //. + by apply/subsetP=> _ /codomP[i ->]; rewrite !inE /= Da eqxx mul1r. +rewrite geq_min orbC -oL subset_leq_card //. +by apply/subsetP=> _ /codomP[j ->]; rewrite !inE /= Da eqxx mul1r. +Qed. + +(* Another consequence of (3.8), used in (4.8), (10.5), (10.10) and (11.8). *) +Lemma eq_signed_sub_cTIiso phi e i j1 j2 : + let rho := (-1) ^+ e *: (eta_ i j1 - eta_ i j2) in + phi \in 'Z[irr G] -> '[phi] = 2%:R -> j1 != j2 -> + {in V, phi =1 rho} -> phi = rho. +Proof. +set rho := _ - _; move: phi => phi0 /= Zphi0 n2phi0 neq_j12 eq_phi_rho. +pose phi := (-1) ^+ e *: phi0; pose psi := phi - rho. +have{eq_phi_rho} psiV0 z: z \in V -> psi z = 0. + by move=> Vz; rewrite !cfunE eq_phi_rho // !cfunE signrMK subrr. +have{Zphi0} Zphi: phi \in 'Z[irr G] by rewrite rpredZsign. +have{n2phi0} n2phi: '[phi] = 2%:R by rewrite cfnorm_sign. +have Zrho: rho \in 'Z[irr G] by rewrite rpredB ?cycTIiso_vchar. +have n2rho: '[rho] = 2%:R. + by rewrite cfnormBd !cfdot_cycTIiso ?eqxx ?(negPf neq_j12) ?andbF. +have [oIphi _ Dphi] := dirr_small_norm Zphi n2phi isT. +have [oIrho _ Drho] := dirr_small_norm Zrho n2rho isT. +set Iphi := dirr_constt _ in oIphi Dphi. +set Irho := dirr_constt _ in oIrho Drho. +suffices /eqP eqIrho: Irho == Iphi by rewrite Drho eqIrho -Dphi signrZK. +have psi_phi'_lt0 di: di \in Irho :\: Iphi -> '[psi, dchi di] < 0. + case/setDP=> rho_di phi'di; rewrite cfdotBl subr_lt0. + move: rho_di; rewrite dirr_consttE; apply: ler_lt_trans. + rewrite real_lerNgt -?dirr_consttE ?real0 ?Creal_Cint //. + by rewrite Cint_cfdot_vchar ?dchi_vchar. +have NCpsi: (NC psi < 2 * minn w1 w2)%N. + suffices NCpsi4: (NC psi <= 2 + 2)%N. + by rewrite (leq_ltn_trans NCpsi4) // !addnn mul2n ltn_double leq_min w1gt2. + by rewrite cycTI_NC_sub // cycTI_NC_norm ?n2phi ?n2rho. +pose rhoId := dirr_dIirr (fun sk => (-1) ^+ (sk.1 : bool) *: eta_ i sk.2). +have rhoIdE s k: dchi (rhoId (s, k)) = (-1) ^+ s *: eta_ i k. + by apply: dirr_dIirrE => sk; rewrite rpredZsign cycTIiso_dirr. +rewrite eqEcard oIrho oIphi andbT -setD_eq0; apply/set0Pn=> [[dk1 phi'dk1]]. +have [[rho_dk1 _] psi_k1_lt0] := (setDP phi'dk1, psi_phi'_lt0 _ phi'dk1). +have dot_dk1: '[rho, dchi dk1] = 1. + rewrite Drho cfdot_suml (big_setD1 dk1) //= cfnorm_dchi big1 ?addr0 //. + move=> dk2 /setD1P[/negPf dk1'2 /dirr_constt_oppl]; rewrite cfdot_dchi dk1'2. + by case: eqP => [-> /negP[] | _ _]; rewrite ?subrr ?ndirrK. +have dot_dk2: 0 < '[rho, rho - dchi dk1]. + by rewrite cfdotBr dot_dk1 n2rho addrK ltr01. +have{dot_dk1 dot_dk2} [s [k Dk1 rho_k2]]: + exists s, exists2 k, rhoId (s, k.1) = dk1 & rhoId (~~ s, k.2) \in Irho. +- move/cfdot_add_dirr_eq1: dot_dk1. + rewrite dirr_dchi rpredN !cycTIiso_dirr //. + case=> // Dk1; [exists false, (j1, j2) | exists true, (j2, j1)]; + try apply: dirr_inj; rewrite ?dirr_consttE rhoIdE scaler_sign //=. + + by rewrite addrC Dk1 addKr in dot_dk2. + by rewrite Dk1 addrK in dot_dk2. +rewrite -Dk1 rhoIdE cfdotZr rmorph_sign in psi_k1_lt0. +have psi_k1_neq0: '[psi, eta_ i k.1] != 0. + by rewrite -(can_eq (signrMK s)) mulr0 ltr_eqF. +set dk2 := rhoId _ in rho_k2. +have NCk2'_le1 (dI : {set _}): + dk2 \in dI -> #|dI| = 2%N -> (NC (\sum_(dk in dI :\ dk2) dchi dk)%R <= 1)%N. +- rewrite (cardsD1 dk2) => -> /eqP/cards1P[dk ->]. + by rewrite big_set1 cycTI_NC_dirr ?dirr_dchi. +suffices /psi_phi'_lt0/ltr_geF/idP[]: dk2 \in Irho :\: Iphi. + rewrite rhoIdE cfdotZr signrN rmorphN mulNr oppr_ge0 rmorph_sign. + have := small_cycTI_NC psiV0 NCpsi psi_k1_neq0. + by case=> // ->; rewrite mulrCA nmulr_lle0 ?ler0n. +have: (1 + 1 < NC psi)%N. + apply (@leq_trans (minn w1 w2)); first by rewrite leq_min w1gt2. + apply: cycTI_NC_minn => //; rewrite NCpsi /NC. + by rewrite (cardsD1 (i, k.1)) inE /= psi_k1_neq0. +rewrite inE rho_k2 andbT ltnNge; apply: contra => phi_k2. +rewrite /psi Drho (big_setD1 dk2) //= Dphi (big_setD1 dk2) //=. +by rewrite addrAC opprD addNKr addrC cycTI_NC_sub ?NCk2'_le1. +Qed. + +(* This is PeterFalvi (3.9)(a). *) +Lemma eq_in_cycTIiso (i : Iirr W) (phi : 'CF(G)) : + phi \in dirr G -> {in V, phi =1 'chi_i} -> phi = sigma 'chi_i. +Proof. +move=> Dphi; rewrite -(inv_dprod_IirrK defW i). +case: (inv_dprod_Iirr _)=> /= i1 j1 EphiC. +pose psi : 'CF(G) := eta_ i1 j1 - phi. +have ZpsiV: {in V, forall g, psi g = 0}=> [g GiV|]. + by rewrite /psi !cfunE cycTIiso_restrict // -(EphiC _ GiV) subrr. +pose a i j := '[psi, eta_ i j]; pose S := [set ij | a ij.1 ij.2 != 0]. +case: (boolP ((i1, j1) \in S))=> [I1J1iS|]; last first. + rewrite inE negbK /a cfdotBl cfdot_cycTIiso !eqxx /=. + rewrite cfdot_dirr ?(irr_dirr, cycTIiso_dirr) //. + case: (boolP (phi == _))=> [|_]. + by rewrite opprK -(natrD _ 1 1) pnatr_eq0. + case: (boolP (phi == _))=> [/eqP //|]. + by rewrite subr0 oner_eq0. +have SPos : (0 < #|S|)%N by rewrite (cardD1 (i1,j1)) I1J1iS. +have SLt: (#|S| <= 2)%N. + by rewrite -[2]add1n cycTI_NC_sub // !cycTI_NC_dirr // cycTIiso_dirr. +have: (0 < #|S| < 2 * minn w1 w2)%N. + rewrite SPos; apply: leq_ltn_trans SLt _. + by rewrite -{1}[2%N]muln1 ltn_mul2l /= leq_min ![(1 < _)%N]ltnW. +move/(cycTI_NC_minn ZpsiV); rewrite leqNgt; case/negP. +by apply: leq_ltn_trans SLt _; rewrite leq_min w1gt2. +Qed. + +(* This is the second part of Peterfalvi (3.9)(a). *) +Lemma cfAut_cycTIiso u phi : cfAut u (sigma phi) = sigma (cfAut u phi). +Proof. +rewrite [phi]cfun_sum_cfdot !raddf_sum; apply: eq_bigr => ij _. +rewrite /= !(linearZ, cfAutZ) /= -aut_IirrE; congr (_ *: _) => {phi}. +apply: eq_in_cycTIiso => [|x Vx /=]. + by have /cycTIirrP[i [j ->]] := mem_irr ij; rewrite dirr_aut cycTIiso_dirr. +by rewrite cfunE cycTIiso_restrict // aut_IirrE cfunE. +Qed. + +Section AutCyclicTI. + +Variable iw : Iirr W. +Let w := 'chi_iw. +Let a := #[w]%CF. + +Let Zsigw : sigma w \in 'Z[irr G]. +Proof. by have [_ -> //] := cycTI_Zisometry; apply: irr_vchar. Qed. + +Let lin_w: w \is a linear_char := Wlin iw. + +(* This is Peterfalvi (3.9)(b). *) +Lemma cycTIiso_aut_exists k : + coprime k a -> + [/\ exists u, sigma (w ^+ k) = cfAut u (sigma w) + & forall x, coprime #[x] a -> sigma (w ^+ k) x = sigma w x]. +Proof. +case/(make_pi_cfAut G)=> u Du_a Du_a'. +suffices Dwk: sigma (w ^+ k) = cfAut u (sigma w). + by split=> [|x co_x_a]; [exists u | rewrite Dwk Du_a']. +rewrite cfAut_cycTIiso; congr (sigma _); apply/cfun_inP=> x Wx. +have Wxbar: coset _ x \in (W / cfker w)%G by rewrite mem_quotient. +rewrite exp_cfunE // cfunE -cfQuoEker //. +rewrite -lin_charX ?cfQuo_lin_char ?cfker_normal // -Du_a ?cfunE //. + by rewrite char_vchar ?cfQuo_char ?irr_char. +by rewrite [a]cforder_lin_char // dvdn_exponent. +Qed. + +(* This is Peterfalvi (3.9)(c). *) +Lemma Cint_cycTIiso_coprime x : coprime #[x] a -> sigma w x \in Cint. +Proof. +move=> co_x_a; apply: Cint_rat_Aint (Aint_vchar _ Zsigw). +have [Qb galQb [QbC AutQbC [w_b genQb memQb]]] := group_num_field_exists <[x]>. +have{memQb} [wx Dwx]: exists wx, sigma w x = QbC wx. + have /memQb Qbx := dvdnn #[x]. + have [sw1 /Qbx[wx1 Dwx1] [sw2 /Qbx[wx2 Dwx2] ->]] := vcharP _ Zsigw. + by exists (wx1 - wx2); rewrite rmorphB !cfunE Dwx1 Dwx2. +suffices: wx \in fixedField 'Gal({:Qb} / 1). + rewrite Dwx (galois_fixedField galQb) ?subvf // => /vlineP[z ->]. + by rewrite -in_algE fmorph_eq_rat fmorph_rat Crat_rat. +apply/fixedFieldP=> [|v_b _]; first exact: memvf. +have [v Dv] := AutQbC v_b; apply: (fmorph_inj QbC); rewrite Dv -Dwx. +have [u uQb uQb'] := dvd_restrict_cfAut (W / cfker w) #[x] v. +transitivity (sigma (cfAut u w) x); first by rewrite -cfAut_cycTIiso cfunE -uQb. +congr (sigma _ _); apply/cfun_inP=> y Wy; rewrite cfunE -cfQuoEker //. +rewrite uQb' ?char_vchar ?cfQuo_char ?irr_char // coprime_sym. +apply: coprime_dvdr co_x_a; rewrite [a]cforder_lin_char //. +by rewrite dvdn_exponent ?mem_quotient. +Qed. + +End AutCyclicTI. + +End Three. + +Implicit Arguments ortho_cycTIiso_vanish [gT G W W1 W2 defW psi]. + +Section ThreeSymmetry. + +Variables (gT : finGroupType) (G W W1 W2 : {group gT}). +Implicit Types (defW : W1 \x W2 = W) (xdefW : W2 \x W1 = W). +Local Notation sigma_ := (@cyclicTIiso gT G W _ _). +Local Notation w_ defW i j := (cyclicTIirr defW i j). + +Lemma cycTIisoC defW xdefW ctiW xctiW i j : + @sigma_ defW ctiW (w_ defW i j) = @sigma_ xdefW xctiW (w_ xdefW j i). +Proof. +apply: eq_in_cycTIiso; first exact: cycTIiso_dirr. +by rewrite /cyclicTIset setUC cyclicTIirrC; apply: cycTIiso_restrict. +Qed. + +Lemma cycTIiso_irrelC defW xdefW ctiW xctiW : + @sigma_ defW ctiW = @sigma_ xdefW xctiW. +Proof. +suffices: sigma_ ctiW =1 sigma_ xctiW by rewrite ![sigma_ _]unlock => /lfunP->. +move=> phi; have [z_ ->] := cfun_irr_sum phi; rewrite !linear_sum. +apply/eq_bigr=> ij _; have [i [j ->]] := cycTIirrP defW (mem_irr ij). +by rewrite !linearZ /= {1}cycTIisoC cyclicTIirrC. +Qed. + +Lemma cycTIiso_irrel defW defW' ctiW ctiW' : + @sigma_ defW ctiW = @sigma_ defW' ctiW'. +Proof. +have xdefW: W2 \x W1 = W by rewrite dprodC. +by rewrite !(cycTIiso_irrelC _ (cyclicTIhyp_sym ctiW xdefW)). +Qed. + +End ThreeSymmetry. diff --git a/mathcomp/odd_order/PFsection4.v b/mathcomp/odd_order/PFsection4.v new file mode 100644 index 0000000..816ac05 --- /dev/null +++ b/mathcomp/odd_order/PFsection4.v @@ -0,0 +1,987 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg poly finset fingroup. +Require Import morphism perm automorphism quotient action gfunctor gproduct. +Require Import center commutator zmodp cyclic pgroup nilpotent hall frobenius. +Require Import matrix mxalgebra mxrepresentation vector ssrnum algC classfun. +Require Import character inertia vcharacter PFsection1 PFsection2 PFsection3. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 4: The Dade isometry of a certain *) +(* type of subgroup. *) +(* Given defW : W1 \x W2 = W, we define here: *) +(* primeTI_hypothesis L K defW <-> *) +(* L = K ><| W1, where W1 acts in a prime manner on K (see *) +(* semiprime in frobenius.v), and both W1 and W2 = 'C_K(W1) *) +(* are nontrivial and cyclic of odd order; these conditions *) +(* imply that cyclicTI_hypothesis L defW holds. *) +(* -> This is Peterfalvi, Hypothesis (4.2), or Feit-Thompson (13.2). *) +(* prime_Dade_definition L K H A A0 defW <-> *) +(* A0 = A :|: class_support (cyclicTIset defW) L where A is *) +(* an L-invariant subset of K^# containing all the elements *) +(* of K that do not act freely on H <| L; in addition *) +(* W2 \subset H \subset K. *) +(* prime_Dade_hypothesis G L K H A A0 defW <-> *) +(* The four assumptions primeTI_hypothesis L K defW, *) +(* cyclicTI_hypothesis G defW, Dade_hypothesis G L A0 and *) +(* prime_Dade_definition L K H A A0 defW hold jointly. *) +(* -> This is Peterfalvi, Hypothesis (4.6), or Feit-Thompson (13.3) (except *) +(* that H is not required to be nilpotent, and the "supporting groups" *) +(* assumptions have been replaced by Dade hypothesis). *) +(* -> This hypothesis is one of the alternatives under which Sibley's *) +(* coherence theorem holds (see PFsection6.v), and is verified by all *) +(* maximal subgroups of type P in a minimal simple odd group. *) +(* -> prime_Dade_hypothesis coerces to Dade_hypothesis, cyclicTI_hypothesis, *) +(* primeTI_hypothesis and prime_Dade_definition. *) +(* For ptiW : primeTI_hypothesis L K defW we also define: *) +(* prime_cycTIhyp ptiW :: cyclicTI_hypothesis L defW (though NOT a coercion) *) +(* primeTIirr ptiW i j == the (unique) irreducible constituent of the image *) +(* (locally) mu2_ i j in 'CF(L) of w_ i j = cyclicTIirr defW i j under *) +(* the sigma = cyclicTIiso (prime_cycTIhyp ptiW). *) +(* primeTI_Iirr ptiW ij == the index of mu2_ ij.1 ij.2; indeed mu2_ i j is *) +(* just notation for 'chi_(primeTI_Iirr ptiW (i, j)). *) +(* primeTIsign ptiW j == the sign of mu2_ i j in sigma (w_ i j), which does *) +(* (locally) delta_ j not depend on i. *) +(* primeTI_Isign ptiW j == the boolean b such that delta_ j := (-1) ^+ b. *) +(* primeTIres ptiW j == the restriction to K of mu2_ i j, which is an *) +(* (locally) chi_ j irreducible character that does not depend on i. *) +(* primeTI_Ires ptiW j == the index of chi_ j := 'chi_(primeTI_Ires ptiW j). *) +(* primeTIred ptiW j == the (reducible) character equal to the sum of all *) +(* (locally) mu_ j the mu2_ i j, and also to 'Ind (chi_ j). *) +(* uniform_prTIred_seq ptiW k == the sequence of all the mu_ j, j != 0, with *) +(* the same degree as mu_ k (s.t. mu_ j 1 = mu_ k 1). *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Section Four_1_to_2. + +(* This is Peterfalvi (4.1). *) + +Variable gT : finGroupType. + +Lemma vchar_pairs_orthonormal (X : {group gT}) (a b c d : 'CF(X)) u v : + {subset (a :: b) <= 'Z[irr X]} /\ {subset (c :: d) <= 'Z[irr X]} -> + orthonormal (a :: b) && orthonormal (c :: d) -> + [&& u \is Creal, v \is Creal, u != 0 & v != 0] -> + [&& '[a - b, u *: c - v *: d] == 0, + (a - b) 1%g == 0 & (u *: c - v *: d) 1%g == 0] -> + orthonormal [:: a; b; c; d]. +Proof. +have osym2 (e f : 'CF(X)) : orthonormal (e :: f) -> orthonormal (f :: e). + by rewrite !(orthonormal_cat [::_] [::_]) orthogonal_sym andbCA. +have def_o f S: orthonormal (f :: S) -> '[f : 'CF(X)] = 1. + by case/andP=> /andP[/eqP]. +case=> /allP/and3P[Za Zb _] /allP/and3P[Zc Zd _] /andP[o_ab o_cd]. +rewrite (orthonormal_cat (a :: b)) o_ab o_cd /=. +case/and4P=> r_u r_v nz_u nz_v /and3P[o_abcd ab1 cd1]. +wlog suff: a b c d u v Za Zb Zc Zd o_ab o_cd r_u r_v nz_u nz_v o_abcd ab1 cd1 / + '[a, c]_X == 0. +- move=> IH; rewrite /orthogonal /= !andbT (IH a b c d u v) //=. + have vc_sym (e f : 'CF(X)) : ((e - f) 1%g == 0) = ((f - e) 1%g == 0). + by rewrite -opprB cfunE oppr_eq0. + have ab_sym e: ('[b - a, e] == 0) = ('[a - b, e] == 0). + by rewrite -opprB cfdotNl oppr_eq0. + rewrite (IH b a c d u v) // 1?osym2 1?vc_sym ?ab_sym //=. + rewrite -oppr_eq0 -cfdotNr opprB in o_abcd. + by rewrite (IH a b d c v u) ?(IH b a d c v u) // 1?osym2 1?vc_sym ?ab_sym. +apply: contraLR cd1 => nz_ac. +have [/orthonormal2P[ab0 a1 b1] /orthonormal2P[cd0 c1 d1]] := (o_ab, o_cd). +have [ea [ia def_a]] := vchar_norm1P Za a1. +have{nz_ac} [e defc]: exists e : bool, c = (-1) ^+ e *: a. + have [ec [ic def_c]] := vchar_norm1P Zc c1; exists (ec (+) ea). + move: nz_ac; rewrite def_a def_c scalerA; rewrite -signr_addb addbK. + rewrite cfdotZl cfdotZr cfdot_irr mulrA mulrC mulf_eq0. + by have [-> // | _]:= ia =P ic; rewrite eqxx. +have def_vbd: v * '[b, d]_X = - ((-1) ^+ e * u). + apply/eqP; have:= o_abcd; rewrite cfdotDl cfdotNl !raddfB /=. + rewrite defc !cfdotZr a1 (cfdotC b) ab0 rmorph0 mulr1. + rewrite -[a]scale1r -{2}[1]/((-1) ^+ false) -(addbb e) signr_addb -scalerA. + rewrite -defc cfdotZl cd0 !mulr0 opprK addrA !subr0 mulrC addrC addr_eq0. + by rewrite rmorph_sign !conj_Creal. +have nz_bd: '[b, d] != 0. + move/esym/eqP: def_vbd; apply: contraTneq => ->. + by rewrite mulr0 oppr_eq0 mulf_eq0 signr_eq0. +have{nz_bd} defd: d = '[b, d] *: b. + move: nz_bd; have [eb [ib ->]] := vchar_norm1P Zb b1. + have [ed [id ->]] := vchar_norm1P Zd d1. + rewrite scalerA cfdotZl cfdotZr rmorph_sign mulrA cfdot_irr. + have [-> _ | _] := ib =P id; last by rewrite !mulr0 eqxx. + by rewrite mulr1 mulrAC -!signr_addb addbb. +rewrite defd scalerA def_vbd scaleNr opprK defc scalerA mulrC -raddfD cfunE. +rewrite !mulf_neq0 ?signr_eq0 // -(subrK a b) -opprB addrCA 2!cfunE. +rewrite (eqP ab1) oppr0 add0r cfunE -mulr2n -mulr_natl mulf_eq0 pnatr_eq0. +by rewrite /= def_a cfunE mulf_eq0 signr_eq0 /= irr1_neq0. +Qed. + +Corollary orthonormal_vchar_diff_ortho (X : {group gT}) (a b c d : 'CF(X)) : + {subset a :: b <= 'Z[irr X]} /\ {subset c :: d <= 'Z[irr X]} -> + orthonormal (a :: b) && orthonormal (c :: d) -> + [&& '[a - b, c - d] == 0, (a - b) 1%g == 0 & (c - d) 1%g == 0] -> + '[a, c] = 0. +Proof. +move=> Zabcd Oabcd; rewrite -[c - d]scale1r scalerBr. +move/(vchar_pairs_orthonormal Zabcd Oabcd) => /implyP. +rewrite rpred1 oner_eq0 (orthonormal_cat (a :: b)) /=. +by case/and3P=> _ _ /andP[] /andP[] /eqP. +Qed. + +(* This is Peterfalvi, Hypothesis (4.2), with explicit parameters. *) +Definition primeTI_hypothesis (L K W W1 W2 : {set gT}) of W1 \x W2 = W := + [/\ (*a*) [/\ K ><| W1 = L, W1 != 1, Hall L W1 & cyclic W1], + (*b*) [/\ W2 != 1, W2 \subset K & cyclic W2], + {in W1^#, forall x, 'C_K[x] = W2} + & (*c*) odd #|W|]%g. + +End Four_1_to_2. + +Arguments Scope primeTI_hypothesis + [_ group_scope group_scope group_scope _ group_scope group_scope]. + +Section Four_3_to_5. + +Variables (gT : finGroupType) (L K W W1 W2 : {group gT}) (defW : W1 \x W2 = W). +Hypothesis ptiWL : primeTI_hypothesis L K defW. + +Let V := cyclicTIset defW. +Let w1 := #|W1|. +Let w2 := #|W2|. + +Let defL : K ><| W1 = L. Proof. by have [[]] := ptiWL. Qed. +Let ntW1 : W1 :!=: 1%g. Proof. by have [[]] := ptiWL. Qed. +Let cycW1 : cyclic W1. Proof. by have [[]] := ptiWL. Qed. +Let hallW1 : Hall L W1. Proof. by have [[]] := ptiWL. Qed. + +Let ntW2 : W2 :!=: 1%g. Proof. by have [_ []] := ptiWL. Qed. +Let sW2K : W2 \subset K. Proof. by have [_ []] := ptiWL. Qed. +Let cycW2 : cyclic W2. Proof. by have [_ []] := ptiWL. Qed. +Let prKW1 : {in W1^#, forall x, 'C_K[x] = W2}. Proof. by have [] := ptiWL. Qed. + +Let oddW : odd #|W|. Proof. by have [] := ptiWL. Qed. + +Let nsKL : K <| L. Proof. by case/sdprod_context: defL. Qed. +Let sKL : K \subset L. Proof. by case/andP: nsKL. Qed. +Let sW1L : W1 \subset L. Proof. by case/sdprod_context: defL. Qed. +Let sWL : W \subset L. +Proof. by rewrite -(dprodWC defW) -(sdprodW defL) mulgSS. Qed. +Let sW1W : W1 \subset W. Proof. by have /mulG_sub[] := dprodW defW. Qed. +Let sW2W : W2 \subset W. Proof. by have /mulG_sub[] := dprodW defW. Qed. + +Let coKW1 : coprime #|K| #|W1|. +Proof. by rewrite (coprime_sdprod_Hall_r defL). Qed. +Let coW12 : coprime #|W1| #|W2|. +Proof. by rewrite coprime_sym (coprimeSg sW2K). Qed. + +Let cycW : cyclic W. Proof. by rewrite (cyclic_dprod defW). Qed. +Let cWW : abelian W. Proof. exact: cyclic_abelian. Qed. +Let oddW1 : odd w1. Proof. exact: oddSg oddW. Qed. +Let oddW2 : odd w2. Proof. exact: oddSg oddW. Qed. + +Let ntV : V != set0. +Proof. +by rewrite -card_gt0 card_cycTIset muln_gt0 -!subn1 !subn_gt0 !cardG_gt1 ntW1. +Qed. + +Let sV_V2 : V \subset W :\: W2. Proof. by rewrite setDS ?subsetUr. Qed. + +Lemma primeTIhyp_quotient (M : {group gT}) : + (W2 / M != 1)%g -> M \subset K -> M <| L -> + {defWbar : (W1 / M) \x (W2 / M) = W / M + & primeTI_hypothesis (L / M) (K / M) defWbar}%g. +Proof. +move=> ntW2bar sMK /andP[_ nML]. +have coMW1: coprime #|M| #|W1| by rewrite (coprimeSg sMK). +have [nMW1 nMW] := (subset_trans sW1L nML, subset_trans sWL nML). +have defWbar: (W1 / M) \x (W2 / M) = (W / M)%g. + by rewrite (quotient_coprime_dprod nMW) ?quotient_odd. +exists defWbar; split; rewrite ?quotient_odd ?quotient_cyclic ?quotientS //. + have isoW1: W1 \isog W1 / M by rewrite quotient_isog ?coprime_TIg. + by rewrite -(isog_eq1 isoW1) ?morphim_Hall // (quotient_coprime_sdprod nML). +move=> Kx /setD1P[ntKx /morphimP[x nKx W1x defKx]] /=. +rewrite -cent_cycle -cycle_eq1 {Kx}defKx -quotient_cycle // in ntKx *. +rewrite -strongest_coprime_quotient_cent ?cycle_subG //; first 1 last. +- by rewrite subIset ?sMK. +- by rewrite (coprimeSg (subsetIl M _)) // (coprimegS _ coMW1) ?cycle_subG. +- by rewrite orbC abelian_sol ?cycle_abelian. +rewrite cent_cycle prKW1 // !inE W1x (contraNneq _ ntKx) // => ->. +by rewrite cycle1 quotient1. +Qed. + +(* This is the first part of PeterFalvi, Theorem (4.3)(a). *) +Theorem normedTI_prTIset : normedTI (W :\: W2) L W. +Proof. +have [[_ _ cW12 _] [_ _ nKW1 tiKW1]] := (dprodP defW, sdprodP defL). +have nV2W: W \subset 'N(W :\: W2) by rewrite sub_abelian_norm ?subsetDl. +have piW1_W: {in W1 & W2, forall x y, (x * y).`_\pi(W1) = x}. + move=> x y W1x W2y /=; rewrite consttM /commute ?(centsP cW12 y) //. + rewrite constt_p_elt ?(mem_p_elt _ W1x) ?pgroup_pi // (constt1P _) ?mulg1 //. + by rewrite /p_elt -coprime_pi' // (coprimegS _ coW12) ?cycle_subG. +have nzV2W: W :\: W2 != set0 by apply: contraNneq ntV; rewrite -subset0 => <-. +apply/normedTI_memJ_P; split=> // xy g V2xy Lg. +apply/idP/idP=> [| /(subsetP nV2W)/memJ_norm->//]. +have{xy V2xy} [/(mem_dprod defW)[x [y [W1x W2y -> _]]] W2'xy] := setDP V2xy. +have{W2'xy} ntx: x != 1%g by have:= W2'xy; rewrite groupMr // => /group1_contra. +have{g Lg} [k [w [Kk /(subsetP sW1W)Ww -> _]]] := mem_sdprod defL Lg. +rewrite conjgM memJ_norm ?(subsetP nV2W) ?(groupMr k) // => /setDP[Wxyk _]. +have{Wxyk piW1_W} W1xk: x ^ k \in W1. + have [xk [yk [W1xk W2yk Dxyk _]]] := mem_dprod defW Wxyk. + by rewrite -(piW1_W x y) // -consttJ Dxyk piW1_W. +rewrite (subsetP sW2W) // -(@prKW1 x) ?in_setD1 ?ntx // inE Kk /=. +rewrite cent1C (sameP cent1P commgP) -in_set1 -set1gE -tiKW1 inE. +by rewrite (subsetP _ _ (mem_commg W1x Kk)) ?commg_subr // groupM ?groupV. +Qed. + +(* Second part of PeterFalvi, Theorem (4.3)(a). *) +Theorem prime_cycTIhyp : cyclicTI_hypothesis L defW. +Proof. +have nVW: W \subset 'N(V) by rewrite sub_abelian_norm ?subsetDl. +by split=> //; apply: normedTI_S normedTI_prTIset. +Qed. +Local Notation ctiW := prime_cycTIhyp. +Let sigma := cyclicTIiso ctiW. +Let w_ i j := cyclicTIirr defW i j. + +Let Wlin k : 'chi[W]_k \is a linear_char. Proof. exact/irr_cyclic_lin. Qed. +Let W1lin i : 'chi[W1]_i \is a linear_char. Proof. exact/irr_cyclic_lin. Qed. +Let W2lin i : 'chi[W2]_i \is a linear_char. Proof. exact/irr_cyclic_lin. Qed. +Let w_lin i j : w_ i j \is a linear_char. Proof. exact: Wlin. Qed. + +Let nirrW1 : #|Iirr W1| = w1. Proof. exact: card_Iirr_cyclic. Qed. +Let nirrW2 : #|Iirr W2| = w2. Proof. exact: card_Iirr_cyclic. Qed. +Let NirrW1 : Nirr W1 = w1. Proof. by rewrite -nirrW1 card_ord. Qed. +Let NirrW2 : Nirr W2 = w2. Proof. by rewrite -nirrW2 card_ord. Qed. +Let w1gt1 : (1 < w1)%N. Proof. by rewrite cardG_gt1. Qed. + +Let cfdot_w i1 j1 i2 j2 : '[w_ i1 j1, w_ i2 j2] = ((i1 == i2) && (j1 == j2))%:R. +Proof. exact: cfdot_dprod_irr. Qed. + +(* Witnesses for Theorem (4.3)(b). *) +Fact primeTIdIirr_key : unit. Proof. by []. Qed. +Definition primeTIdIirr_def := dirr_dIirr (sigma \o prod_curry w_). +Definition primeTIdIirr := locked_with primeTIdIirr_key primeTIdIirr_def. +Definition primeTI_Iirr ij := (primeTIdIirr ij).2. +Definition primeTI_Isign j := (primeTIdIirr (0, j)).1. +Local Notation Imu2 := primeTI_Iirr. +Local Notation mu2_ i j := 'chi_(primeTI_Iirr (i, j)). +Local Notation delta_ j := (GRing.sign algCring (primeTI_Isign j)). + +Let ew_ i j := w_ i j - w_ 0 j. +Let V2ew i j : ew_ i j \in 'CF(W, W :\: W2). +Proof. +apply/cfun_onP=> x; rewrite !inE negb_and negbK => /orP[W2x | /cfun0->//]. +by rewrite -[x]mul1g !cfunE /w_ !dprod_IirrE !cfDprodE ?lin_char1 ?subrr. +Qed. + +(* This is Peterfalvi, Theorem (4.3)(b, c). *) +Theorem primeTIirr_spec : + [/\ (*b*) injective Imu2, + forall i j, 'Ind (ew_ i j) = delta_ j *: (mu2_ i j - mu2_ 0 j), + forall i j, sigma (w_ i j) = delta_ j *: mu2_ i j, + (*c*) forall i j, {in W :\: W2, mu2_ i j =1 delta_ j *: w_ i j} + & forall k, k \notin codom Imu2 -> {in W :\: W2, 'chi_k =1 \0}]. +Proof. +have isoV2 := normedTI_isometry normedTI_prTIset (setDSS sWL (sub1G W2)). +have /fin_all_exists2[dmu injl_mu Ddmu] j: + exists2 dmu : bool * {ffun Iirr W1 -> Iirr L}, injective dmu.2 + & forall i, 'Ind (ew_ i j) = dchi (dmu.1, dmu.2 i) - dchi (dmu.1, dmu.2 0). +- pose Sj := [tuple w_ i j | i < Nirr W1]. + have Sj0: Sj`_0 = w_ 0 j by rewrite (nth_mktuple _ 0 0). + have irrSj: {subset Sj <= irr W} by move=> ? /mapP[i _ ->]; apply: mem_irr. + have: {in 'Z[Sj, W :\: W2], isometry 'Ind, to 'Z[irr L, L^#]}. + split=> [|phi]; first by apply: sub_in2 isoV2; apply: zchar_on. + move/(zchar_subset irrSj)/(zchar_onS (setDS W (sub1G W2))). + by rewrite !zcharD1E cfInd1 // mulf_eq0 orbC => /andP[/cfInd_vchar-> // ->]. + case/vchar_isometry_base=> // [|||i|mu Umu [d Ddmu]]; first by rewrite NirrW1. + + rewrite orthonormal_free // (sub_orthonormal irrSj) ?irr_orthonormal //. + by apply/injectiveP=> i1 i2 /irr_inj/dprod_Iirr_inj[]. + + by move=> _ /mapP[i _ ->]; rewrite Sj0 !lin_char1. + + by rewrite nth_mktuple Sj0 V2ew. + exists (d, [ffun i => tnth mu i]) => [|i]. + apply/injectiveP; congr (uniq _): Umu. + by rewrite (eq_map (ffunE _)) map_tnth_enum. + by rewrite -scalerBr /= !ffunE !(tnth_nth 0 mu) -Ddmu nth_mktuple Sj0. +pose Imu ij := (dmu ij.2).2 ij.1; pose mu i j := 'chi_(Imu (i, j)). +pose d j : algC := (-1) ^+ (dmu j).1. +have{Ddmu} Ddmu i j: 'Ind (ew_ i j) = d j *: (mu i j - mu 0 j). + by rewrite Ddmu scalerBr. +have{injl_mu} inj_Imu: injective Imu. + move=> [i1 j1] [i2 j2]; rewrite /Imu /=; pose S i j k := mu i j :: mu k j. + have [-> /injl_mu-> // | j2'1 /eqP/negPf[] /=] := eqVneq j1 j2. + apply/(can_inj oddb)/eqP; rewrite -eqC_nat -cfdot_irr -!/(mu _ _) mulr0n. + have oIew_j12 i k: '['Ind[L] (ew_ i j1), 'Ind[L] (ew_ k j2)] = 0. + by rewrite isoV2 // cfdotBl !cfdotBr !cfdot_w (negPf j2'1) !andbF !subr0. + have defSd i j k: mu i j - mu k j = d j *: ('Ind (ew_ i j) - 'Ind (ew_ k j)). + by rewrite !Ddmu -scalerBr signrZK opprB addrA subrK. + have Sd1 i j k: (mu i j - mu k j) 1%g == 0. + by rewrite defSd !(cfunE, cfInd1) ?lin_char1 // !subrr mulr0. + have exS i j: {k | {subset S i j k <= 'Z[irr L]} & orthonormal (S i j k)}. + have:= w1gt1; rewrite -nirrW1 (cardD1 i) => /card_gt0P/sigW[k /andP[i'k _]]. + exists k; first by apply/allP; rewrite /= !irr_vchar. + apply/andP; rewrite /= !cfdot_irr !eqxx !andbT /=. + by rewrite (inj_eq (injl_mu j)) mulrb ifN_eqC. + have [[k1 ZS1 o1S1] [k2 ZS2 o1S2]] := (exS i1 j1, exS i2 j2). + rewrite (orthonormal_vchar_diff_ortho (conj ZS1 ZS2)) ?o1S1 ?Sd1 ?andbT //. + by rewrite !defSd cfdotZl cfdotZr cfdotBl !cfdotBr !oIew_j12 !subrr !mulr0. +pose V2base := [tuple of [seq ew_ ij.1 ij.2 | ij in predX (predC1 0) predT]]. +have V2basis: basis_of 'CF(W, W :\: W2) V2base. + suffices V2free: free V2base. + rewrite basisEfree V2free size_image /= cardX cardC1 nirrW1 nirrW2 -subn1. + rewrite mulnBl mul1n dim_cfun_on_abelian ?subsetDl //. + rewrite cardsD (setIidPr _) // (dprod_card defW) leqnn andbT. + by apply/span_subvP=> _ /mapP[ij _ ->]. + apply/freeP=> /= z zV2e0 k. + move Dk: (enum_val k) (enum_valP k) => [i j] /andP[/= nz_i _]. + rewrite -(cfdot0l (w_ i j)) -{}zV2e0 cfdot_suml (bigD1 k) //= cfdotZl. + rewrite nth_image Dk cfdotBl !cfdot_w !eqxx eq_sym (negPf nz_i) subr0 mulr1. + rewrite big1 ?addr0 // => k1; rewrite -(inj_eq enum_val_inj) {}Dk nth_image. + case: (enum_val k1) => /= i1 j1 ij'ij1. + rewrite cfdotZl cfdotBl !cfdot_dprod_irr [_ && _](negPf ij'ij1). + by rewrite eq_sym (negPf nz_i) subr0 mulr0. +have nsV2W: W :\: W2 <| W by rewrite -sub_abelian_normal ?subsetDl. +pose muW k := let: ij := inv_dprod_Iirr defW k in d ij.2 *: mu ij.1 ij.2. +have inW := codomP (dprod_Iirr_onto defW _). +have ImuW k1 k2: '[muW k1, muW k2] = (k1 == k2)%:R. + have [[[i1 j1] -> {k1}] [[i2 j2] -> {k2}]] := (inW k1, inW k2). + rewrite cfdotZl cfdotZr !dprod_IirrK (can_eq (dprod_IirrK _)) /= rmorph_sign. + rewrite cfdot_irr (inj_eq inj_Imu (_, _) (_, _)) -/(d _). + by case: eqP => [[_ ->] | _]; rewrite ?signrMK ?mulr0. +have [k|muV2 mu'V2] := equiv_restrict_compl_ortho sWL nsV2W V2basis ImuW. + rewrite nth_image; case: (enum_val k) (enum_valP k) => /= i j /andP[/= nzi _]. + pose inWj i1 := dprod_Iirr defW (i1, j); rewrite (bigD1 (inWj 0)) //=. + rewrite (bigD1 (inWj i)) ?(can_eq (dprod_IirrK _)) ?xpair_eqE ?(negPf nzi) //. + rewrite /= big1 ?addr0 => [|k1 /andP[]]; last first. + rewrite !(eq_sym k1); have [[i1 j1] -> {k1}] := inW k1. + rewrite !(can_eq (dprod_IirrK _)) => ij1'i ij1'0. + by rewrite cfdotBl !cfdot_w !mulrb !ifN // subrr scale0r. + rewrite /muW !dprod_IirrK /= addrC !cfdotBl !cfdot_w !eqxx /= !andbT. + by rewrite eq_sym (negPf nzi) subr0 add0r scaleNr !scale1r -scalerBr. +have Dsigma i j: sigma (w_ i j) = d j *: mu i j. + apply/esym/eq_in_cycTIiso=> [|x Vx]; first exact: (dirr_dchi (_, _)). + by rewrite -muV2 ?(subsetP sV_V2) // /muW dprod_IirrK. +have /all_and2[Dd Dmu] j: d j = delta_ j /\ forall i, Imu (i, j) = Imu2 (i, j). + suffices DprTI i: primeTIdIirr (i, j) = ((dmu j).1, (dmu j).2 i). + by split=> [|i]; rewrite /primeTI_Isign /Imu2 DprTI. + apply: dirr_inj; rewrite /primeTIdIirr unlock_with dirr_dIirrE /= ?Dsigma //. + by case=> i1 j1; apply: cycTIiso_dirr. +split=> [[i1 j1] [i2 j2] | i j | i j | i j x V2x | k mu2p'k]. +- by rewrite -!Dmu => /inj_Imu. +- by rewrite -!Dmu -Dd -Ddmu. +- by rewrite -Dmu -Dd -Dsigma. +- by rewrite cfunE -muV2 // /muW dprod_IirrK Dd cfunE signrMK -Dmu. +apply: mu'V2 => k1; have [[i j] ->{k1}] := inW k1. +apply: contraNeq mu2p'k; rewrite cfdotZr rmorph_sign mulf_eq0 signr_eq0 /=. +rewrite /mu Dmu dprod_IirrK -irr_consttE constt_irr inE /= => /eqP <-. +exact: codom_f. +Qed. + +(* These lemmas restate the various parts of Theorem (4.3)(b, c) separately. *) +Lemma prTIirr_inj : injective Imu2. Proof. by case: primeTIirr_spec. Qed. + +Corollary cfdot_prTIirr i1 j1 i2 j2 : + '[mu2_ i1 j1, mu2_ i2 j2] = ((i1 == i2) && (j1 == j2))%:R. +Proof. by rewrite cfdot_irr (inj_eq prTIirr_inj). Qed. + +Lemma cfInd_sub_prTIirr i j : + 'Ind[L] (w_ i j - w_ 0 j) = delta_ j *: (mu2_ i j - mu2_ 0 j). +Proof. by case: primeTIirr_spec i j. Qed. + +Lemma cycTIiso_prTIirr i j : sigma (w_ i j) = delta_ j *: mu2_ i j. +Proof. by case: primeTIirr_spec. Qed. + +Lemma prTIirr_id i j : {in W :\: W2, mu2_ i j =1 delta_ j *: w_ i j}. +Proof. by case: primeTIirr_spec. Qed. + +Lemma not_prTIirr_vanish k : k \notin codom Imu2 -> {in W :\: W2, 'chi_k =1 \0}. +Proof. by case: primeTIirr_spec k. Qed. + +(* This is Peterfalvi, Theorem (4.3)(d). *) +Theorem prTIirr1_mod i j : (mu2_ i j 1%g == delta_ j %[mod w1])%C. +Proof. +rewrite -(cfRes1 W1) -['Res _](subrK ('Res (delta_ j *: w_ i j))) cfunE. +set phi := _ - _; pose a := '[phi, 1]. +have phi_on_1: phi \in 'CF(W1, 1%g). + apply/cfun_onP=> g; have [W1g | /cfun0-> //] := boolP (g \in W1). + rewrite -(coprime_TIg coW12) inE W1g !cfunE !cfResE //= => W2'g. + by rewrite prTIirr_id ?cfunE ?subrr // inE W2'g (subsetP sW1W). +have{phi_on_1} ->: phi 1%g = a * w1%:R. + rewrite mulrC /a (cfdotEl _ phi_on_1) mulVKf ?neq0CG //. + by rewrite big_set1 cfun11 conjC1 mulr1. +rewrite cfResE // cfunE lin_char1 // mulr1 eqCmod_addl_mul //. +by rewrite Cint_cfdot_vchar ?rpred1 ?rpredB ?cfRes_vchar ?rpredZsign ?irr_vchar. +Qed. + +Lemma prTIsign_aut u j : delta_ (aut_Iirr u j) = delta_ j. +Proof. +have /eqP := cfAut_cycTIiso ctiW u (w_ 0 j). +rewrite -cycTIirr_aut aut_Iirr0 -/sigma !cycTIiso_prTIirr raddfZsign /=. +by rewrite -aut_IirrE eq_scaled_irr => /andP[/eqP]. +Qed. + +Lemma prTIirr_aut u i j : + mu2_ (aut_Iirr u i) (aut_Iirr u j) = cfAut u (mu2_ i j). +Proof. +rewrite -!(canLR (signrZK _) (cycTIiso_prTIirr _ _)) -!/(delta_ _). +by rewrite prTIsign_aut raddfZsign /= cfAut_cycTIiso -cycTIirr_aut. +Qed. + +(* The (reducible) column sums of the prime TI irreducibles. *) +Definition primeTIred j : 'CF(L) := \sum_i mu2_ i j. +Local Notation mu_ := primeTIred. + +Definition uniform_prTIred_seq j0 := + image mu_ [pred j | j != 0 & mu_ j 1%g == mu_ j0 1%g]. + +Lemma prTIred_aut u j : mu_ (aut_Iirr u j) = cfAut u (mu_ j). +Proof. +rewrite raddf_sum [mu_ _](reindex_inj (aut_Iirr_inj u)). +by apply: eq_bigr => i _; rewrite /= prTIirr_aut. +Qed. + +Lemma cfdot_prTIirr_red i j k : '[mu2_ i j, mu_ k] = (j == k)%:R. +Proof. +rewrite cfdot_sumr (bigD1 i) // cfdot_prTIirr eqxx /=. +rewrite big1 ?addr0 // => i1 neq_i1i. +by rewrite cfdot_prTIirr eq_sym (negPf neq_i1i). +Qed. + +Lemma cfdot_prTIred j1 j2 : '[mu_ j1, mu_ j2] = ((j1 == j2) * w1)%:R. +Proof. +rewrite cfdot_suml (eq_bigr _ (fun i _ => cfdot_prTIirr_red i _ _)) sumr_const. +by rewrite mulrnA card_Iirr_cyclic. +Qed. + +Lemma cfnorm_prTIred j : '[mu_ j] = w1%:R. +Proof. by rewrite cfdot_prTIred eqxx mul1n. Qed. + +Lemma prTIred_neq0 j : mu_ j != 0. +Proof. by rewrite -cfnorm_eq0 cfnorm_prTIred neq0CG. Qed. + +Lemma prTIred_char j : mu_ j \is a character. +Proof. by apply: rpred_sum => i _; apply: irr_char. Qed. + +Lemma prTIred_1_gt0 j : 0 < mu_ j 1%g. +Proof. by rewrite char1_gt0 ?prTIred_neq0 ?prTIred_char. Qed. + +Lemma prTIred_1_neq0 i : mu_ i 1%g != 0. +Proof. by rewrite char1_eq0 ?prTIred_neq0 ?prTIred_char. Qed. + +Lemma prTIred_inj : injective mu_. +Proof. +move=> j1 j2 /(congr1 (cfdot (mu_ j1)))/esym/eqP; rewrite !cfdot_prTIred. +by rewrite eqC_nat eqn_pmul2r ?cardG_gt0 // eqxx; case: (j1 =P j2). +Qed. + +Lemma prTIred_not_real j : j != 0 -> ~~ cfReal (mu_ j). +Proof. +apply: contraNneq; rewrite -prTIred_aut -irr_eq1 -odd_eq_conj_irr1 //. +by rewrite -aut_IirrE => /prTIred_inj->. +Qed. + +Lemma prTIsign0 : delta_ 0 = 1. +Proof. +have /esym/eqP := cycTIiso_prTIirr 0 0; rewrite -[sigma _]scale1r. +by rewrite /w_ /sigma cycTIirr00 cycTIiso1 -irr0 eq_scaled_irr => /andP[/eqP]. +Qed. + +Lemma prTIirr00 : mu2_ 0 0 = 1. +Proof. +have:= cycTIiso_prTIirr 0 0; rewrite prTIsign0 scale1r. +by rewrite /w_ /sigma cycTIirr00 cycTIiso1. +Qed. + +(* This is PeterFalvi (4.4). *) +Lemma prTIirr0P k : + reflect (exists i, 'chi_k = mu2_ i 0) (K \subset cfker 'chi_k). +Proof. +suff{k}: [set k | K \subset cfker 'chi_k] == [set Imu2 (i, 0) | i : Iirr W1]. + move/eqP/setP/(_ k); rewrite inE => ->. + by apply: (iffP imsetP) => [[i _]|[i /irr_inj]] ->; exists i. +have [isoW1 abW1] := (sdprod_isog defL, cyclic_abelian cycW1). +have abLbar: abelian (L / K) by rewrite -(isog_abelian isoW1). +rewrite eqEcard andbC card_imset ?nirrW1 => [| i1 i2 /prTIirr_inj[] //]. +rewrite [w1](card_isog isoW1) -card_Iirr_abelian //. +rewrite -(card_image (can_inj (mod_IirrK nsKL))) subset_leq_card; last first. + by apply/subsetP=> _ /imageP[k1 _ ->]; rewrite inE mod_IirrE ?cfker_mod. +apply/subsetP=> k; rewrite inE => kerKk. +have /irrP[ij DkW]: 'Res 'chi_k \in irr W. + rewrite lin_char_irr ?cfRes_lin_char // lin_irr_der1. + by apply: subset_trans kerKk; rewrite der1_min ?normal_norm. +have{ij DkW} [i DkW]: exists i, 'Res 'chi_k = w_ i 0. + have /codomP[[i j] Dij] := dprod_Iirr_onto defW ij; exists i. + rewrite DkW Dij; congr (w_ i _); apply/eqP; rewrite -subGcfker. + rewrite -['chi_j](cfDprodKr_abelian defW i) // -dprod_IirrE -{}Dij -{}DkW. + by rewrite cfResRes // sub_cfker_Res // (subset_trans sW2K kerKk). +apply/imsetP; exists i => //=; apply/irr_inj. +suffices ->: 'chi_k = delta_ 0 *: mu2_ i 0 by rewrite prTIsign0 scale1r. +rewrite -cycTIiso_prTIirr -(eq_in_cycTIiso _ (irr_dirr k)) // => x /setDP[Wx _]. +by rewrite -/(w_ i 0) -DkW cfResE. +Qed. + +(* This is the first part of PeterFalvi, Theorem (4.5)(a). *) +Theorem cfRes_prTIirr_eq0 i j : 'Res[K] (mu2_ i j) = 'Res (mu2_ 0 j). +Proof. +apply/eqP; rewrite -subr_eq0 -rmorphB /=; apply/eqP/cfun_inP=> x0 Kx0. +rewrite -(canLR (signrZK _) (cfInd_sub_prTIirr i j)) -/(delta_ j). +rewrite cfResE // !cfunE (cfun_on0 (cfInd_on _ (V2ew i j))) ?mulr0 //. +apply: contraL Kx0 => /imset2P[x y /setDP[Wx W2'x] Ly ->] {x0}. +rewrite memJ_norm ?(subsetP (normal_norm nsKL)) //; apply: contra W2'x => Kx. +by rewrite -(mul1g W2) -(coprime_TIg coKW1) group_modr // inE Kx (dprodW defW). +Qed. + +Lemma prTIirr_1 i j : mu2_ i j 1%g = mu2_ 0 j 1%g. +Proof. by rewrite -!(@cfRes1 _ K L) cfRes_prTIirr_eq0. Qed. + +Lemma prTIirr0_1 i : mu2_ i 0 1%g = 1. +Proof. by rewrite prTIirr_1 prTIirr00 cfun11. Qed. + +Lemma prTIirr0_linear i : mu2_ i 0 \is a linear_char. +Proof. by rewrite qualifE irr_char /= prTIirr0_1. Qed. + +Lemma prTIred_1 j : mu_ j 1%g = w1%:R * mu2_ 0 j 1%g. +Proof. +rewrite mulr_natl -nirrW1 sum_cfunE. +by rewrite -sumr_const; apply: eq_bigr => i _; rewrite prTIirr_1. +Qed. + +Definition primeTI_Ires j : Iirr K := cfIirr ('Res[K] (mu2_ 0 j)). +Local Notation Ichi := primeTI_Ires. +Local Notation chi_ j := 'chi_(Ichi j). + +(* This is the rest of PeterFalvi, Theorem (4.5)(a). *) +Theorem prTIres_spec j : chi_ j = 'Res (mu2_ 0 j) /\ mu_ j = 'Ind (chi_ j). +Proof. +rewrite /Ichi; set chi_j := 'Res _. +have [k chi_j_k]: {k | k \in irr_constt chi_j} := constt_cfRes_irr K _. +have Nchi_j: chi_j \is a character by rewrite cfRes_char ?irr_char. +have lb_mu_1: w1%:R * 'chi_k 1%g <= mu_ j 1%g ?= iff (chi_j == 'chi_k). + have [chi' Nchi' Dchi_j] := constt_charP _ Nchi_j chi_j_k. + rewrite prTIred_1 (mono_lerif (ler_pmul2l (gt0CG W1))). + rewrite -subr_eq0 Dchi_j addrC addKr -(canLR (addrK _) Dchi_j) !cfunE. + rewrite lerif_subLR addrC -lerif_subLR cfRes1 subrr -char1_eq0 // eq_sym. + by apply: lerif_eq; rewrite char1_ge0. +pose psi := 'Ind 'chi_k - mu_ j; have Npsi: psi \is a character. + apply/forallP=> l; rewrite coord_cfdot cfdotBl; set a := '['Ind _, _]. + have Na: a \in Cnat by rewrite Cnat_cfdot_char_irr ?cfInd_char ?irr_char. + have [[i /eqP Dl] | ] := altP (@existsP _ (fun i => 'chi_l == mu2_ i j)). + have [n Da] := CnatP a Na; rewrite Da cfdotC Dl cfdot_prTIirr_red. + rewrite rmorph_nat -natrB ?Cnat_nat // eqxx lt0n -eqC_nat -Da. + by rewrite -irr_consttE constt_Ind_Res Dl cfRes_prTIirr_eq0. + rewrite negb_exists => /forallP muj'l. + rewrite cfdot_suml big1 ?subr0 // => i _. + rewrite cfdot_irr -(inj_eq irr_inj) mulrb ifN_eqC ?muj'l //. +have ub_mu_1: mu_ j 1%g <= 'Ind[L] 'chi_k 1%g ?= iff ('Ind 'chi_k == mu_ j). + rewrite -subr_eq0 -/psi (canRL (subrK _) (erefl psi)) cfunE -lerif_subLR. + by rewrite subrr -char1_eq0 // eq_sym; apply: lerif_eq; rewrite char1_ge0. +have [_ /esym] := lerif_trans lb_mu_1 ub_mu_1; rewrite cfInd1 //. +by rewrite -(index_sdprod defL) eqxx => /andP[/eqP-> /eqP <-]; rewrite irrK. +Qed. + +Lemma cfRes_prTIirr i j : 'Res[K] (mu2_ i j) = chi_ j. +Proof. by rewrite cfRes_prTIirr_eq0; case: (prTIres_spec j). Qed. + +Lemma cfInd_prTIres j : 'Ind[L] (chi_ j) = mu_ j. +Proof. by have [] := prTIres_spec j. Qed. + +Lemma cfRes_prTIred j : 'Res[K] (mu_ j) = w1%:R *: chi_ j. +Proof. +rewrite -nirrW1 scaler_nat -sumr_const linear_sum /=; apply: eq_bigr => i _. +exact: cfRes_prTIirr. +Qed. + +Lemma prTIres_aut u j : chi_ (aut_Iirr u j) = cfAut u (chi_ j). +Proof. +by rewrite -(cfRes_prTIirr (aut_Iirr u 0)) prTIirr_aut -cfAutRes cfRes_prTIirr. +Qed. + +Lemma prTIres0 : chi_ 0 = 1. +Proof. by rewrite -(cfRes_prTIirr 0) prTIirr00 cfRes_cfun1. Qed. + +Lemma prTIred0 : mu_ 0 = w1%:R *: '1_K. +Proof. +by rewrite -cfInd_prTIres prTIres0 cfInd_cfun1 // -(index_sdprod defL). +Qed. + +Lemma prTIres_inj : injective Ichi. +Proof. by move=> j1 j2 Dj; apply: prTIred_inj; rewrite -!cfInd_prTIres Dj. Qed. + +(* This is the first assertion of Peterfalvi (4.5)(b). *) +Theorem prTIres_irr_cases k (theta := 'chi_k) (phi := 'Ind theta) : + {j | theta = chi_ j} + {phi \in irr L /\ (forall i j, phi != mu2_ i j)}. +Proof. +pose imIchi := [set Ichi j | j : Iirr W2]. +have [/imsetP/sig2_eqW[j _] | imIchi'k] := boolP (k \in imIchi). + by rewrite /theta => ->; left; exists j. +suffices{phi} theta_inv: 'I_L[theta] = K. + have irr_phi: phi \in irr L by apply: inertia_Ind_irr; rewrite ?theta_inv. + right; split=> // i j; apply: contraNneq imIchi'k => Dphi; apply/imsetP. + exists j => //; apply/eqP; rewrite -[k == _]constt_irr -(cfRes_prTIirr i). + by rewrite -constt_Ind_Res -/phi Dphi irr_consttE cfnorm_irr oner_eq0. +rewrite -(sdprodW (sdprod_modl defL (sub_inertia _))); apply/mulGidPl. +apply/subsetP=> z /setIP[W1z Itheta_z]; apply: contraR imIchi'k => K'z. +have{K'z} [Lz ntz] := (subsetP sW1L z W1z, group1_contra K'z : z != 1%g). +have [p p_pr p_z]: {p | prime p & p %| #[z]} by apply/pdivP; rewrite order_gt1. +have coKp := coprime_dvdr (dvdn_trans p_z (order_dvdG W1z)) coKW1. +wlog{p_z} p_z: z W1z Lz Itheta_z ntz / p.-elt z. + move/(_ z.`_p)->; rewrite ?groupX ?p_elt_constt //. + by rewrite (sameP eqP constt1P) /p_elt p'natE ?negbK. +have JirrP: is_action L (@conjg_Iirr gT K); last pose Jirr := Action JirrP. + split=> [y k1 k2 eq_k12 | k1 y1 y2 Gy1 Gy2]; apply/irr_inj. + by apply/(can_inj (cfConjgK y)); rewrite -!conjg_IirrE eq_k12. + by rewrite !conjg_IirrE (cfConjgM _ nsKL). +have [[_ nKL] [nKz _]] := (andP nsKL, setIdP Itheta_z). +suffices{k theta Itheta_z} /eqP->: imIchi == 'Fix_Jirr[z]. + by apply/afix1P/irr_inj; rewrite conjg_IirrE inertiaJ. +rewrite eqEcard; apply/andP; split. + apply/subsetP=> _ /imsetP[j _ ->]; apply/afix1P/irr_inj. + by rewrite conjg_IirrE -(cfRes_prTIirr 0) (cfConjgRes _ _ nsKL) ?cfConjg_id. +have ->: #|imIchi| = w2 by rewrite card_imset //; apply: prTIres_inj. +have actsL_KK: [acts L, on classes K | 'Js \ subsetT L]. + rewrite astabs_ract subsetIidl; apply/subsetP=> y Ly; rewrite !inE /=. + apply/subsetP=> _ /imsetP[x Kx ->]; rewrite !inE /= -class_rcoset. + by rewrite norm_rlcoset ?class_lcoset ?mem_classes ?memJ_norm ?(subsetP nKL). +rewrite (card_afix_irr_classes Lz actsL_KK) => [|k x y Kx /=]; last first. + by case/imsetP=> _ /imsetP[t Kt ->] ->; rewrite conjg_IirrE cfConjgEJ ?cfunJ. +apply: leq_trans (subset_leq_card _) (leq_imset_card (class^~ K) _). +apply/subsetP=> _ /setIP[/imsetP[x Kx ->] /afix1P/normP nxKz]. +suffices{Kx} /pred0Pn[t /setIP[xKt czt]]: #|'C_(x ^: K)[z]| != 0%N. + rewrite -(class_transr xKt); apply: mem_imset; have [y Ky Dt] := imsetP xKt. + by rewrite -(@prKW1 z) ?(czt, inE) ?ntz // Dt groupJ. +have{coKp}: ~~ (p %| #|K|) by rewrite -prime_coprime // coprime_sym. +apply: contraNneq => /(congr1 (modn^~ p))/eqP; rewrite mod0n. +rewrite -cent_cycle -afixJ -sylow.pgroup_fix_mod ?astabsJ ?cycle_subG //. +by move/dvdn_trans; apply; rewrite -index_cent1 dvdn_indexg. +Qed. + +(* Implicit elementary converse to the above. *) +Lemma prTIred_not_irr j : mu_ j \notin irr L. +Proof. by rewrite irrEchar cfnorm_prTIred pnatr_eq1 gtn_eqF ?andbF. Qed. + +(* This is the second assertion of Peterfalvi (4.5)(b). *) +Theorem prTIind_irr_cases ell (phi := 'chi_ell) : + {i : _ & {j | phi = mu2_ i j}} + + {k | k \notin codom Ichi & phi = 'Ind 'chi_k}. +Proof. +have [k] := constt_cfRes_irr K ell; rewrite -constt_Ind_Res => kLell. +have [[j Dk] | [/irrP/sig_eqW[l1 DkL] chi'k]] := prTIres_irr_cases k. + have [i /=/eqP <- | mu2j'l] := pickP (fun i => mu2_ i j == phi). + by left; exists i, j. + case/eqP: kLell; rewrite Dk cfInd_prTIres cfdot_suml big1 // => i _. + by rewrite cfdot_irr -(inj_eq irr_inj) mu2j'l. +right; exists k; last by move: kLell; rewrite DkL constt_irr inE => /eqP <-. +apply/codomP=> [[j Dk]]; have/negP[] := prTIred_not_irr j. +by rewrite -cfInd_prTIres -Dk DkL mem_irr. +Qed. + +End Four_3_to_5. + +Notation primeTIsign ptiW j := + (GRing.sign algCring (primeTI_Isign ptiW j)) (only parsing). +Notation primeTIirr ptiW i j := 'chi_(primeTI_Iirr ptiW (i, j)) (only parsing). +Notation primeTIres ptiW j := 'chi_(primeTI_Ires ptiW j) (only parsing). + +Implicit Arguments prTIirr_inj [gT L K W W1 W2 defW x1 x2]. +Implicit Arguments prTIred_inj [gT L K W W1 W2 defW x1 x2]. +Implicit Arguments prTIres_inj [gT L K W W1 W2 defW x1 x2]. +Implicit Arguments not_prTIirr_vanish [gT L K W W1 W2 defW k]. + +Section Four_6_t0_10. + +Variables (gT : finGroupType) (G L K H : {group gT}) (A A0 : {set gT}). +Variables (W W1 W2 : {group gT}) (defW : W1 \x W2 = W). + +Local Notation V := (cyclicTIset defW). + +(* These correspond to Peterfalvi, Hypothesis (4.6). *) +Definition prime_Dade_definition := + [/\ (*c*) [/\ H <| L, W2 \subset H & H \subset K], + (*d*) [/\ A <| L, \bigcup_(h in H^#) 'C_K[h]^# \subset A & A \subset K^#] + & (*e*) A0 = A :|: class_support V L]. + +Record prime_Dade_hypothesis : Prop := PrimeDadeHypothesis { + prDade_cycTI :> cyclicTI_hypothesis G defW; + prDade_prTI :> primeTI_hypothesis L K defW; + prDade_hyp :> Dade_hypothesis G L A0; + prDade_def :> prime_Dade_definition +}. + +Hypothesis prDadeHyp : prime_Dade_hypothesis. + +Let ctiWG : cyclicTI_hypothesis G defW := prDadeHyp. +Let ptiWL : primeTI_hypothesis L K defW := prDadeHyp. +Let ctiWL : cyclicTI_hypothesis L defW := prime_cycTIhyp ptiWL. +Let ddA0 : Dade_hypothesis G L A0 := prDadeHyp. +Local Notation ddA0def := (prDade_def prDadeHyp). + +Local Notation w_ i j := (cyclicTIirr defW i j). +Local Notation sigma := (cyclicTIiso ctiWG). +Local Notation eta_ i j := (sigma (w_ i j)). +Local Notation mu2_ i j := (primeTIirr ptiWL i j). +Local Notation delta_ j := (primeTIsign ptiWL j). +Local Notation chi_ j := (primeTIres ptiWL j). +Local Notation mu_ := (primeTIred ptiWL). +Local Notation tau := (Dade ddA0). + +Let defA0 : A0 = A :|: class_support V L. Proof. by have [] := ddA0def. Qed. +Let nsAL : A <| L. Proof. by have [_ []] := ddA0def. Qed. +Let sAA0 : A \subset A0. Proof. by rewrite defA0 subsetUl. Qed. + +Let nsHL : H <| L. Proof. by have [[]] := ddA0def. Qed. +Let sHK : H \subset K. Proof. by have [[]] := ddA0def. Qed. +Let defL : K ><| W1 = L. Proof. by have [[]] := ptiWL. Qed. +Let sKL : K \subset L. Proof. by have /mulG_sub[] := sdprodW defL. Qed. +Let coKW1 : coprime #|K| #|W1|. +Proof. by rewrite (coprime_sdprod_Hall_r defL); have [[]] := ptiWL. Qed. + +Let sIH_A : \bigcup_(h in H^#) 'C_K[h]^# \subset A. +Proof. by have [_ []] := ddA0def. Qed. + +Let sW2H : W2 \subset H. Proof. by have [[]] := ddA0def. Qed. +Let ntW1 : W1 :!=: 1%g. Proof. by have [[]] := ptiWL. Qed. +Let ntW2 : W2 :!=: 1%g. Proof. by have [_ []] := ptiWL. Qed. + +Let oddW : odd #|W|. Proof. by have [] := ctiWL. Qed. +Let sW1W : W1 \subset W. Proof. by have /mulG_sub[] := dprodW defW. Qed. +Let sW2W : W2 \subset W. Proof. by have /mulG_sub[] := dprodW defW. Qed. +Let tiW12 : W1 :&: W2 = 1%g. Proof. by have [] := dprodP defW. Qed. + +Let cycW : cyclic W. Proof. by have [] := ctiWG. Qed. +Let cycW1 : cyclic W1. Proof. by have [[]] := ptiWL. Qed. +Let cycW2 : cyclic W2. Proof. by have [_ []] := ptiWL. Qed. +Let sLG : L \subset G. Proof. by case: ddA0. Qed. +Let sW2K : W2 \subset K. Proof. by have [_ []] := ptiWL. Qed. + +Let sWL : W \subset L. +Proof. by rewrite -(dprodWC defW) -(sdprodW defL) mulgSS. Qed. +Let sWG : W \subset G. Proof. exact: subset_trans sWL sLG. Qed. + +Let oddW1 : odd #|W1|. Proof. exact: oddSg oddW. Qed. +Let oddW2 : odd #|W2|. Proof. exact: oddSg oddW. Qed. + +Let w1gt1 : (2 < #|W1|)%N. Proof. by rewrite odd_gt2 ?cardG_gt1. Qed. +Let w2gt2 : (2 < #|W2|)%N. Proof. by rewrite odd_gt2 ?cardG_gt1. Qed. + +Let nirrW1 : #|Iirr W1| = #|W1|. Proof. exact: card_Iirr_cyclic. Qed. +Let nirrW2 : #|Iirr W2| = #|W2|. Proof. exact: card_Iirr_cyclic. Qed. +Let W1lin i : 'chi[W1]_i \is a linear_char. Proof. exact/irr_cyclic_lin. Qed. +Let W2lin i : 'chi[W2]_i \is a linear_char. Proof. exact/irr_cyclic_lin. Qed. + +(* This is the first part of Peterfalvi (4.7). *) +Lemma prDade_irr_on k : + ~~ (H \subset cfker 'chi[K]_k) -> 'chi_k \in 'CF(K, 1%g |: A). +Proof. +move=> kerH'i; apply/cfun_onP=> g; rewrite !inE => /norP[ntg A'g]. +have [Kg | /cfun0-> //] := boolP (g \in K). +apply: not_in_ker_char0 (normalS _ _ nsHL) kerH'i _ => //. +apply/trivgP/subsetP=> h /setIP[Hh cgh]; apply: contraR A'g => nth. +apply/(subsetP sIH_A)/bigcupP; exists h; first exact/setDP. +by rewrite 3!inE ntg Kg cent1C. +Qed. + +(* This is the second part of Peterfalvi (4.7). *) +Lemma prDade_Ind_irr_on k : + ~~ (H \subset cfker 'chi[K]_k) -> 'Ind[L] 'chi_k \in 'CF(L, 1%g |: A). +Proof. +move/prDade_irr_on/(cfInd_on sKL); apply: cfun_onS; rewrite class_supportEr. +by apply/bigcupsP=> _ /normsP-> //; rewrite normsU ?norms1 ?normal_norm. +Qed. + +(* Third part of Peterfalvi (4.7). *) +Lemma cfker_prTIres j : j != 0 -> ~~ (H \subset cfker (chi_ j)). +Proof. +rewrite -(cfRes_prTIirr _ 0) cfker_Res ?irr_char // subsetI sHK /=. +apply: contra => kerHmu0j; rewrite -irr_eq1; apply/eqP/cfun_inP=> y W2y. +have [[x W1x ntx] mulW] := (trivgPn _ ntW1, dprodW defW). +rewrite cfun1E W2y -(cfDprodEr defW _ W1x W2y) -dprodr_IirrE -dprod_Iirr0l. +have{ntx} W2'x: x \notin W2 by rewrite -[x \in W2]andTb -W1x -in_setI tiW12 inE. +have V2xy: (x * y)%g \in W :\: W2 by rewrite inE -mulW mem_mulg ?groupMr ?W2'x. +rewrite -[w_ 0 j](signrZK (primeTI_Isign ptiWL j)) cfunE -prTIirr_id //. +have V2x: x \in W :\: W2 by rewrite inE W2'x (subsetP sW1W). +rewrite cfkerMr ?(subsetP (subset_trans sW2H kerHmu0j)) ?prTIirr_id // cfunE. +by rewrite signrMK -[x]mulg1 dprod_Iirr0l dprodr_IirrE cfDprodEr ?lin_char1. +Qed. + +(* Fourth part of Peterfalvi (4.7). *) +Lemma prDade_TIres_on j : j != 0 -> chi_ j \in 'CF(K, 1%g |: A). +Proof. by move/cfker_prTIres/prDade_irr_on. Qed. + +(* Last part of Peterfalvi (4.7). *) +Lemma prDade_TIred_on j : j != 0 -> mu_ j \in 'CF(L, 1%g |: A). +Proof. by move/cfker_prTIres/prDade_Ind_irr_on; rewrite cfInd_prTIres. Qed. + +Import ssrint. + +(* Second part of PeterFalvi (4.8). *) +Lemma prDade_TIsign_eq i j k : + mu2_ i j 1%g = mu2_ i k 1%g -> delta_ j = delta_ k. +Proof. +move=> eqjk; have{eqjk}: (delta_ j == delta_ k %[mod #|W1|])%C. + apply: eqCmod_trans (prTIirr1_mod ptiWL i k). + by rewrite eqCmod_sym -eqjk (prTIirr1_mod ptiWL). +have /negP: ~~ (#|W1| %| 2) by rewrite gtnNdvd. +rewrite /eqCmod -![delta_ _]intr_sign -rmorphB dvdC_int ?Cint_int //= intCK. +by do 2!case: (primeTI_Isign _ _). +Qed. + +(* First part of PeterFalvi (4.8). *) +Lemma prDade_sub_TIirr_on i j k : + j != 0 -> k != 0 -> mu2_ i j 1%g = mu2_ i k 1%g -> + mu2_ i j - mu2_ i k \in 'CF(L, A0). +Proof. +move=> nzj nzk eq_mu1. +apply/cfun_onP=> g; rewrite defA0 !inE negb_or !cfunE => /andP[A'g V'g]. +have [Lg | L'g] := boolP (g \in L); last by rewrite !cfun0 ?subrr. +have{Lg} /bigcupP[_ /rcosetsP[x W1x ->] Kx_g]: g \in cover (rcosets K W1). + by rewrite (cover_partition (rcosets_partition_mul W1 K)) (sdprodW defL). +have [x1 | ntx] := eqVneq x 1%g. + have [-> | ntg] := eqVneq g 1%g; first by rewrite eq_mu1 subrr. + have{A'g} A1'g: g \notin 1%g |: A by rewrite !inE negb_or ntg. + rewrite x1 mulg1 in Kx_g; rewrite -!(cfResE (mu2_ i _) sKL) ?cfRes_prTIirr //. + by rewrite !(cfun_onP (prDade_TIres_on _)) ?subrr. +have coKx: coprime #|K| #[x] by rewrite (coprime_dvdr (order_dvdG W1x)). +have nKx: x \in 'N(K) by have [_ _ /subsetP->] := sdprodP defL. +have [/cover_partition defKx _] := partition_cent_rcoset nKx coKx. +have def_cKx: 'C_K[x] = W2 by have [_ _ -> //] := ptiWL; rewrite !inE ntx. +move: Kx_g; rewrite -defKx def_cKx cover_imset => /bigcupP[z /(subsetP sKL)Lz]. +case/imsetP=> _ /rcosetP[y W2y ->] Dg; rewrite Dg !cfunJ //. +have V2yx: (y * x)%g \in W :\: W2. + rewrite inE -(dprodWC defW) mem_mulg // andbT groupMl //. + by rewrite -[x \in W2]andTb -W1x -in_setI tiW12 inE. +rewrite 2?{1}prTIirr_id //. +have /set1P->: y \in [1]. + rewrite -tiW12 inE W2y andbT; apply: contraR V'g => W1'y. + by rewrite Dg mem_imset2 // !inE negb_or -andbA -in_setD groupMr ?W1'y. +rewrite -commute1 (prDade_TIsign_eq eq_mu1) !cfunE -mulrBr. +by rewrite !dprod_IirrE !cfDprodE // !lin_char1 // subrr mulr0. +Qed. + +(* This is last part of PeterFalvi (4.8). *) +Lemma prDade_sub_TIirr i j k : + j != 0 -> k != 0 -> mu2_ i j 1%g = mu2_ i k 1%g -> + tau (mu2_ i j - mu2_ i k) = delta_ j *: (eta_ i j - eta_ i k). +Proof. +move=> nz_j nz_k eq_mu2jk_1. +have [-> | k'j] := eqVneq j k; first by rewrite !subrr !raddf0. +have [[Itau Ztau] [_ Zsigma]] := (Dade_Zisometry ddA0, cycTI_Zisometry ctiWL). +set dmu2 := _ - _; set dsw := _ - _; have Dmu2 := prTIirr_id ptiWL. +have Zmu2: dmu2 \in 'Z[irr L, A0]. + by rewrite zchar_split rpredB ?irr_vchar ?prDade_sub_TIirr_on. +apply: eq_signed_sub_cTIiso => // [||x Vx]. +- exact: zcharW (Ztau _ Zmu2). +- rewrite Itau // cfnormBd ?cfnorm_irr // (cfdot_prTIirr ptiWL). + by rewrite (negPf k'j) andbF. +have V2x: x \in W :\: W2 by rewrite (subsetP _ x Vx) // setDS ?subsetUr. +rewrite !(cfunE, Dade_id) ?(cycTIiso_restrict _ _ Vx) //; last first. + by rewrite defA0 inE orbC mem_class_support. +by rewrite !Dmu2 // (prDade_TIsign_eq eq_mu2jk_1) !cfunE -mulrBr. +Qed. + +Lemma prDade_supp_disjoint : V \subset ~: K. +Proof. +rewrite subDset setUC -subDset setDE setCK setIC -(dprod_modr defW sW2K). +by rewrite coprime_TIg // dprod1g subsetUr. +Qed. + +(* This is Peterfalvi (4.9). *) +(* We have added the "obvious" fact that calT is pairwise orthogonal, since *) +(* we require this to prove membership in 'Z[calT], we encapsulate the *) +(* construction of tau1, and state its conformance to tau on the "larger" *) +(* domain 'Z[calT, L^#], so that clients can avoid using the domain equation *) +(* in part (a). *) +Theorem uniform_prTIred_coherent k (calT := uniform_prTIred_seq ptiWL k) : + k != 0 -> + (*a*) [/\ pairwise_orthogonal calT, ~~ has cfReal calT, conjC_closed calT, + 'Z[calT, L^#] =i 'Z[calT, A] + & exists2 psi, psi != 0 & psi \in 'Z[calT, A]] + (*b*) /\ (exists2 tau1 : {linear 'CF(L) -> 'CF(G)}, + forall j, tau1 (mu_ j) = delta_ k *: (\sum_i sigma (w_ i j)) + & {in 'Z[calT], isometry tau1, to 'Z[irr G]} + /\ {in 'Z[calT, L^#], tau1 =1 tau}). +Proof. +have uniqT: uniq calT by apply/dinjectiveP; apply: in2W; apply: prTIred_inj. +have sTmu: {subset calT <= codom mu_} by exact: image_codom. +have oo_mu: pairwise_orthogonal (codom mu_). + apply/pairwise_orthogonalP; split=> [|_ _ /codomP[j1 ->] /codomP[j2 ->]]. + apply/andP; split; last by apply/injectiveP; apply: prTIred_inj. + by apply/codomP=> [[i /esym/eqP/idPn[]]]; apply: prTIred_neq0. + by rewrite cfdot_prTIred; case: (j1 =P j2) => // -> /eqP. +have real'T: ~~ has cfReal calT. + by apply/hasPn=> _ /imageP[j /andP[nzj _] ->]; apply: prTIred_not_real. +have ccT: conjC_closed calT. + move=> _ /imageP[j Tj ->]; rewrite -prTIred_aut image_f // inE aut_Iirr_eq0. + by rewrite prTIred_aut cfunE conj_Cnat ?Cnat_char1 ?prTIred_char. +have TonA: 'Z[calT, L^#] =i 'Z[calT, A]. + have A'1: 1%g \notin A by apply: contra (subsetP sAA0 _) _; have [] := ddA0. + move => psi; rewrite zcharD1E -(setU1K A'1) zcharD1; congr (_ && _). + apply/idP/idP; [apply: zchar_trans_on psi => psi Tpsi | exact: zcharW]. + have [j /andP[nz_j _] Dpsi] := imageP Tpsi. + by rewrite zchar_split mem_zchar // Dpsi prDade_TIred_on. +move=> nzk; split. + split=> //; first exact: sub_pairwise_orthogonal oo_mu. + have Tmuk: mu_ k \in calT by rewrite image_f // inE nzk /=. + exists ((mu_ k)^*%CF - mu_ k); first by rewrite subr_eq0 (hasPn real'T). + rewrite -TonA -rpredN opprB sub_aut_zchar ?zchar_onG ?mem_zchar ?ccT //. + by move=> _ /mapP[j _ ->]; rewrite char_vchar ?prTIred_char. +pose f0 j := delta_ k *: (\sum_i eta_ i j); have in_mu := codom_f mu_. +pose f1 psi := f0 (iinv (valP (insigd (in_mu k) psi))). +have f1mu j: f1 (mu_ j) = f0 j. + have in_muj := in_mu j. + rewrite /f1 /insigd /insubd /= insubT /=; [idtac]. + by rewrite iinv_f //; apply: prTIred_inj. +have iso_f1: {in codom mu_, isometry f1, to 'Z[irr G]}. + split=> [_ _ /codomP[j1 ->] /codomP[j2 ->] | _ /codomP[j ->]]; last first. + by rewrite f1mu rpredZsign rpred_sum // => i _; apply: cycTIiso_vchar. + rewrite !f1mu cfdotZl cfdotZr rmorph_sign signrMK !cfdot_suml. + apply: eq_bigr => i1 _; rewrite !cfdot_sumr; apply: eq_bigr => i2 _. + by rewrite cfdot_cycTIiso cfdot_prTIirr. +have [tau1 Dtau1 Itau1] := Zisometry_of_iso oo_mu iso_f1. +exists tau1 => [j|]; first by rewrite Dtau1 ?codom_f ?f1mu. +split=> [|psi]; first by apply: sub_iso_to Itau1 => //; apply: zchar_subset. +rewrite zcharD1E => /andP[/zchar_expansion[//|z _ Dpsi] /eqP psi1_0]. +rewrite -[psi]subr0 -(scale0r (mu_ k)) -(mul0r (mu_ k 1%g)^-1) -{}psi1_0. +rewrite {psi}Dpsi sum_cfunE mulr_suml scaler_suml -sumrB !raddf_sum /=. +apply: eq_big_seq => _ /imageP[j /andP[nzj /eqP eq_mujk_1] ->]. +rewrite cfunE eq_mujk_1 mulfK ?prTIred_1_neq0 // -scalerBr !linearZ /=. +congr (_ *: _); rewrite {z}linearB !Dtau1 ?codom_f // !f1mu -scalerBr -!sumrB. +rewrite !linear_sum; apply: eq_bigr => i _ /=. +have{eq_mujk_1} eq_mu2ijk_1: mu2_ i j 1%g = mu2_ i k 1%g. + by apply: (mulfI (neq0CG W1)); rewrite !prTIirr_1 -!prTIred_1. +by rewrite -(prDade_TIsign_eq eq_mu2ijk_1) prDade_sub_TIirr. +Qed. + +(* This is Peterfalvi (4.10). *) +Lemma prDade_sub2_TIirr i j : + tau (delta_ j *: mu2_ i j - delta_ j *: mu2_ 0 j - mu2_ i 0 + mu2_ 0 0) + = eta_ i j - eta_ 0 j - eta_ i 0 + eta_ 0 0. +Proof. +pose V0 := class_support V L; have sVV0: V \subset V0 := sub_class_support L V. +have sV0A0: V0 \subset A0 by rewrite defA0 subsetUr. +have nV0L: L \subset 'N(V0) := class_support_norm V L. +have [_ _ /normedTI_memJ_P[ntV _ tiV]] := ctiWG. +have [/andP[sA0L _] _ A0'1 _ _] := ddA0. +have{sA0L A0'1} sV0G: V0 \subset G^#. + by rewrite (subset_trans sV0A0) // subsetD1 A0'1 (subset_trans sA0L). +have{sVV0} ntV0: V0 != set0 by apply: contraNneq ntV; rewrite -subset0 => <-. +have{ntV} tiV0: normedTI V0 G L. + apply/normedTI_memJ_P; split=> // _ z /imset2P[u y Vu Ly ->] Gz. + apply/idP/idP=> [/imset2P[u1 y1 Vu1 Ly1 Duyz] | Lz]; last first. + by rewrite -conjgM mem_imset2 ?groupM. + rewrite -[z](mulgKV y1) groupMr // -(groupMl _ Ly) (subsetP sWL) //. + by rewrite -(tiV u) ?groupM ?groupV // ?(subsetP sLG) // !conjgM Duyz conjgK. +have{ntV0 sV0A0 nV0L tiV0} DtauV0: {in 'CF(L, V0), tau =1 'Ind}. + by move=> beta V0beta; rewrite /= -(restr_DadeE _ sV0A0) //; apply: Dade_Ind. +pose alpha := cfCyclicTIset defW i j; set beta := _ *: mu2_ i j - _ - _ + _. +have Valpha: alpha \in 'CF(W, V) := cfCycTI_on ctiWL i j. +have Dalpha: alpha = w_ i j - w_ 0 j - w_ i 0 + w_ 0 0. + by rewrite addrC {1}cycTIirr00 addrA addrAC addrA addrAC -cfCycTI_E. +rewrite -!(linearB sigma) -linearD -Dalpha cycTIiso_Ind //. +suffices ->: beta = 'Ind[L] alpha by rewrite DtauV0 ?cfInd_on ?cfIndInd. +rewrite Dalpha -addrA -[w_ 0 0]opprK -opprD linearB /= /beta -scalerBr. +by rewrite !(cfInd_sub_prTIirr ptiWL) prTIsign0 scale1r opprD opprK addrA. +Qed. + +End Four_6_t0_10. diff --git a/mathcomp/odd_order/PFsection5.v b/mathcomp/odd_order/PFsection5.v new file mode 100644 index 0000000..0c3b1eb --- /dev/null +++ b/mathcomp/odd_order/PFsection5.v @@ -0,0 +1,1607 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg poly finset center. +Require Import fingroup morphism perm automorphism quotient action zmodp. +Require Import gfunctor gproduct cyclic pgroup frobenius. +Require Import matrix mxalgebra mxrepresentation vector ssrint. +Require Import ssrnum algC classfun character inertia vcharacter. +Require Import PFsection1 PFsection2 PFsection3 PFsection4. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 5: Coherence. *) +(* Defined here: *) +(* coherent_with S A tau tau1 <-> tau1 is a Z-linear isometry from 'Z[S] to *) +(* 'Z[irr G] that coincides with tau on 'Z[S, A]. *) +(* coherent S A tau <-> (S, A, tau) is coherent, i.e., there is a Z-linear *) +(* isometry tau1 s.t. coherent_with S A tau tau1. *) +(* subcoherent S tau R <-> S : seq 'cfun(L) is non empty, pairwise orthogonal *) +(* and closed under complex conjugation, tau is an *) +(* isometry from 'Z[S, L^#] to virtual characters in *) +(* G that maps the difference chi - chi^*, for each *) +(* chi \in S, to the sum of an orthonormal family *) +(* R chi of virtual characters of G; also, R chi and *) +(* R phi are orthogonal unless phi \in chi :: chi^*. *) +(* dual_iso nu == the Z-linear (additive) mapping phi |-> - nu phi^* *) +(* for nu : {additive 'CF(L) -> 'CF(G)}. If nu is an *) +(* isometry extending a subcoherent tau on 'Z[S] with *) +(* size S = 2, then so is dual_iso nu. *) +(* We provide a set of definitions that cover the various \cal S notations *) +(* introduces in Peterfalvi sections 5, 6, 7, and 9 to 14. *) +(* Iirr_ker K A == the set of all i : Iirr K such that the kernel of *) +(* 'chi_i contains A. *) +(* Iirr_kerD K B A == the set of all i : Iirr K such that the kernel of *) +(* 'chi_i contains A but not B. *) +(* seqInd L calX == the duplicate-free sequence of characters of L *) +(* induced from K by the 'chi_i for i in calX. *) +(* seqIndT K L == the duplicate-free sequence of all characters of L *) +(* induced by irreducible characters of K. *) +(* seqIndD K L H M == the duplicate-free sequence of characters of L *) +(* induced by irreducible characters of K that have M *) +(* in their kernel, but not H. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +(* Results about the set of induced irreducible characters *) +Section InducedIrrs. + +Variables (gT : finGroupType) (K L : {group gT}). +Implicit Types (A B : {set gT}) (H M : {group gT}). +Implicit Type u : {rmorphism algC -> algC}. + +Section KerIirr. + +Definition Iirr_ker A := [set i | A \subset cfker 'chi[K]_i]. + +Lemma Iirr_kerS A B : B \subset A -> Iirr_ker A \subset Iirr_ker B. +Proof. by move/subset_trans=> sBA; apply/subsetP=> i; rewrite !inE => /sBA. Qed. + +Lemma sum_Iirr_ker_square H : + H <| K -> \sum_(i in Iirr_ker H) 'chi_i 1%g ^+ 2 = #|K : H|%:R. +Proof. +move=> nsHK; rewrite -card_quotient ?normal_norm // -irr_sum_square. +rewrite (eq_bigl _ _ (in_set _)) (reindex _ (mod_Iirr_bij nsHK)) /=. +by apply: eq_big => [i | i _]; rewrite mod_IirrE ?cfker_mod ?cfMod1. +Qed. + +Definition Iirr_kerD B A := Iirr_ker A :\: Iirr_ker B. + +Lemma sum_Iirr_kerD_square H M : + H <| K -> M <| K -> M \subset H -> + \sum_(i in Iirr_kerD H M) 'chi_i 1%g ^+ 2 = #|K : H|%:R * (#|H : M|%:R - 1). +Proof. +move=> nsHK nsMK sMH; have [sHK _] := andP nsHK. +rewrite mulrBr mulr1 -natrM Lagrange_index // -!sum_Iirr_ker_square //. +apply/esym/(canLR (addrK _)); rewrite /= addrC (big_setID (Iirr_ker H)). +by rewrite (setIidPr _) ?Iirr_kerS //. +Qed. + +Lemma Iirr_ker_aut u A i : (aut_Iirr u i \in Iirr_ker A) = (i \in Iirr_ker A). +Proof. by rewrite !inE aut_IirrE cfker_aut. Qed. + +Lemma Iirr_ker_conjg A i x : + x \in 'N(A) -> (conjg_Iirr i x \in Iirr_ker A) = (i \in Iirr_ker A). +Proof. +move=> nAx; rewrite !inE conjg_IirrE. +have [nKx | /cfConjgEout-> //] := boolP (x \in 'N(K)). +by rewrite cfker_conjg // -{1}(normP nAx) conjSg. +Qed. + +Lemma Iirr_kerDS A1 A2 B1 B2 : + A2 \subset A1 -> B1 \subset B2 -> Iirr_kerD B1 A1 \subset Iirr_kerD B2 A2. +Proof. by move=> sA12 sB21; rewrite setDSS ?Iirr_kerS. Qed. + +Lemma Iirr_kerDY B A : Iirr_kerD (A <*> B) A = Iirr_kerD B A. +Proof. by apply/setP=> i; rewrite !inE join_subG; apply: andb_id2r => ->. Qed. + +Lemma mem_Iirr_ker1 i : (i \in Iirr_kerD K 1%g) = (i != 0). +Proof. by rewrite !inE sub1G andbT subGcfker. Qed. + +End KerIirr. + +Hypothesis nsKL : K <| L. +Let sKL := normal_sub nsKL. +Let nKL := normal_norm nsKL. +Let e := #|L : K|%:R : algC. +Let nze : e != 0 := neq0CiG _ _. + +Section SeqInd. + +Variable calX : {set (Iirr K)}. + +(* The set of characters induced from the irreducibles in calX. *) +Definition seqInd := undup [seq 'Ind[L] 'chi_i | i in calX]. +Local Notation S := seqInd. + +Lemma seqInd_uniq : uniq S. Proof. exact: undup_uniq. Qed. + +Lemma seqIndP phi : + reflect (exists2 i, i \in calX & phi = 'Ind[L] 'chi_i) (phi \in S). +Proof. by rewrite mem_undup; exact: imageP. Qed. + +Lemma seqInd_on : {subset S <= 'CF(L, K)}. +Proof. by move=> _ /seqIndP[i _ ->]; exact: cfInd_normal. Qed. + +Lemma seqInd_char : {subset S <= character}. +Proof. by move=> _ /seqIndP[i _ ->]; rewrite cfInd_char ?irr_char. Qed. + +Lemma Cnat_seqInd1 phi : phi \in S -> phi 1%g \in Cnat. +Proof. by move/seqInd_char/Cnat_char1. Qed. + +Lemma Cint_seqInd1 phi : phi \in S -> phi 1%g \in Cint. +Proof. by rewrite CintE; move/Cnat_seqInd1->. Qed. + +Lemma seqInd_neq0 psi : psi \in S -> psi != 0. +Proof. by move=> /seqIndP[i _ ->]; exact: Ind_irr_neq0. Qed. + +Lemma seqInd1_neq0 psi : psi \in S -> psi 1%g != 0. +Proof. by move=> Spsi; rewrite char1_eq0 ?seqInd_char ?seqInd_neq0. Qed. + +Lemma cfnorm_seqInd_neq0 psi : psi \in S -> '[psi] != 0. +Proof. by move/seqInd_neq0; rewrite cfnorm_eq0. Qed. + +Lemma seqInd_ortho : {in S &, forall phi psi, phi != psi -> '[phi, psi] = 0}. +Proof. +move=> _ _ /seqIndP[i _ ->] /seqIndP[j _ ->]. +by case: ifP (cfclass_Ind_cases i j nsKL) => // _ -> /eqP. +Qed. + +Lemma seqInd_orthogonal : pairwise_orthogonal S. +Proof. +apply/pairwise_orthogonalP; split; last exact: seqInd_ortho. +by rewrite /= undup_uniq andbT; move/memPn: seqInd_neq0. +Qed. + +Lemma seqInd_free : free S. +Proof. exact: (orthogonal_free seqInd_orthogonal). Qed. + +Lemma seqInd_zcharW : {subset S <= 'Z[S]}. +Proof. by move=> phi Sphi; rewrite mem_zchar ?seqInd_free. Qed. + +Lemma seqInd_zchar : {subset S <= 'Z[S, K]}. +Proof. by move=> phi Sphi; rewrite zchar_split seqInd_zcharW ?seqInd_on. Qed. + +Lemma seqInd_vcharW : {subset S <= 'Z[irr L]}. +Proof. by move=> phi Sphi; rewrite char_vchar ?seqInd_char. Qed. + +Lemma seqInd_vchar : {subset S <= 'Z[irr L, K]}. +Proof. by move=> phi Sphi; rewrite zchar_split seqInd_vcharW ?seqInd_on. Qed. + +Lemma zcharD1_seqInd : 'Z[S, L^#] =i 'Z[S, K^#]. +Proof. +move=> phi; rewrite zcharD1E (zchar_split _ K^#) cfun_onD1. +by apply: andb_id2l => /(zchar_trans_on seqInd_zchar)/zchar_on->. +Qed. + +Lemma zcharD1_seqInd_on : {subset 'Z[S, L^#] <= 'CF(L, K^#)}. +Proof. by move=> phi; rewrite zcharD1_seqInd => /zchar_on. Qed. + +Lemma zcharD1_seqInd_Dade A : + 1%g \notin A -> {subset S <= 'CF(L, 1%g |: A)} -> 'Z[S, L^#] =i 'Z[S, A]. +Proof. +move=> notA1 A_S phi; rewrite zcharD1E (zchar_split _ A). +apply/andb_id2l=> ZSphi; apply/idP/idP=> [phi10 | /cfun_on0-> //]. +rewrite -(setU1K notA1) cfun_onD1 {}phi10 andbT. +have{phi ZSphi} [c -> _] := free_span seqInd_free (zchar_span ZSphi). +by rewrite big_seq memv_suml // => xi /A_S/memvZ. +Qed. + +Lemma dvd_index_seqInd1 phi : phi \in S -> phi 1%g / e \in Cnat. +Proof. +by case/seqIndP=> i _ ->; rewrite cfInd1 // mulrC mulKf ?Cnat_irr1. +Qed. + +Lemma sub_seqInd_zchar phi psi : + phi \in S -> psi \in S -> psi 1%g *: phi - phi 1%g *: psi \in 'Z[S, K^#]. +Proof. +move=> Sphi Spsi; rewrite zcharD1 !cfunE mulrC subrr eqxx. +by rewrite rpredB ?scale_zchar ?Cint_seqInd1 ?seqInd_zchar. +Qed. + +Lemma sub_seqInd_on phi psi : + phi \in S -> psi \in S -> psi 1%g *: phi - phi 1%g *: psi \in 'CF(L, K^#). +Proof. by move=> Sphi Spsi; exact: zchar_on (sub_seqInd_zchar Sphi Spsi). Qed. + +Lemma size_irr_subseq_seqInd S1 : + subseq S1 S -> {subset S1 <= irr L} -> + (#|L : K| * size S1 = #|[set i | 'Ind 'chi[K]_i \in S1]|)%N. +Proof. +move=> sS1S irrS1; rewrite (card_imset_Ind_irr nsKL) => [|i|i y]; first 1 last. +- by rewrite inE => /irrS1. +- rewrite !inE => S1iG Ly; congr (_ \in S1): S1iG. + by apply: cfclass_Ind => //; apply/cfclassP; exists y; rewrite ?conjg_IirrE. +congr (_ * _)%N; rewrite -(size_map (@cfIirr _ _)) -(card_uniqP _); last first. + rewrite map_inj_in_uniq ?(subseq_uniq sS1S) ?seqInd_uniq //. + by apply: (@can_in_inj _ _ _ _ (tnth (irr L))) => phi /irrS1/cfIirrE. +apply: eq_card => s; apply/mapP/imsetP=> [[phi S1phi ->] | [i]]. + have /seqIndP[i _ Dphi] := mem_subseq sS1S S1phi. + by exists i; rewrite ?inE -Dphi. +by rewrite inE => S1iG ->; exists ('Ind 'chi_i). +Qed. + +Section Beta. + +Variable xi : 'CF(L). +Hypotheses (Sxi : xi \in S) (xi1 : xi 1%g = e). + +Lemma cfInd1_sub_lin_vchar : 'Ind[L, K] 1 - xi \in 'Z[irr L, K^#]. +Proof. +rewrite zcharD1 !cfunE xi1 cfInd1 // cfun11 mulr1 subrr eqxx andbT. +rewrite rpredB ?(seqInd_vchar Sxi) // zchar_split cfInd_normal ?char_vchar //. +by rewrite cfInd_char ?cfun1_char. +Qed. + +Lemma cfInd1_sub_lin_on : 'Ind[L, K] 1 - xi \in 'CF(L, K^#). +Proof. exact: zchar_on cfInd1_sub_lin_vchar. Qed. + +Lemma seqInd_sub_lin_vchar : + {in S, forall phi : 'CF(L), phi - (phi 1%g / e) *: xi \in 'Z[S, K^#]}. +Proof. +move=> phi Sphi; rewrite /= zcharD1 !cfunE xi1 divfK // subrr eqxx. +by rewrite rpredB ?scale_zchar ?seqInd_zchar // CintE dvd_index_seqInd1. +Qed. + +Lemma seqInd_sub_lin_on : + {in S, forall phi : 'CF(L), phi - (phi 1%g / e) *: xi \in 'CF(L, K^#)}. +Proof. by move=> phi /seqInd_sub_lin_vchar/zchar_on. Qed. + +End Beta. + +End SeqInd. + +Implicit Arguments seqIndP [calX phi]. + +Lemma seqIndS (calX calY : {set Iirr K}) : + calX \subset calY -> {subset seqInd calX <= seqInd calY}. +Proof. +by move=> sXY _ /seqIndP[i /(subsetP sXY)Yi ->]; apply/seqIndP; exists i. +Qed. + +Definition seqIndT := seqInd setT. + +Lemma seqInd_subT calX : {subset seqInd calX <= seqIndT}. +Proof. exact: seqIndS (subsetT calX). Qed. + +Lemma mem_seqIndT i : 'Ind[L, K] 'chi_i \in seqIndT. +Proof. by apply/seqIndP; exists i; rewrite ?inE. Qed. + +Lemma seqIndT_Ind1 : 'Ind[L, K] 1 \in seqIndT. +Proof. by rewrite -irr0 mem_seqIndT. Qed. + +Lemma cfAut_seqIndT u : cfAut_closed u seqIndT. +Proof. +by move=> _ /seqIndP[i _ ->]; rewrite cfAutInd -aut_IirrE mem_seqIndT. +Qed. + +Definition seqIndD H M := seqInd (Iirr_kerD H M). + +Lemma seqIndDY H M : seqIndD (M <*> H) M = seqIndD H M. +Proof. by rewrite /seqIndD Iirr_kerDY. Qed. + +Lemma mem_seqInd H M i : + H <| L -> M <| L -> ('Ind 'chi_i \in seqIndD H M) = (i \in Iirr_kerD H M). +Proof. +move=> nsHL nsML; apply/seqIndP/idP=> [[j Xj] | Xi]; last by exists i. +case/cfclass_Ind_irrP/cfclassP=> // y Ly; rewrite -conjg_IirrE => /irr_inj->. +by rewrite inE !Iirr_ker_conjg -?in_setD ?(subsetP _ y Ly) ?normal_norm. +Qed. + +Lemma seqIndC1P phi : + reflect (exists2 i, i != 0 & phi = 'Ind 'chi[K]_i) (phi \in seqIndD K 1). +Proof. +by apply: (iffP seqIndP) => [] [i nzi ->]; + exists i; rewrite // mem_Iirr_ker1 in nzi *. +Qed. + +Lemma seqIndC1_filter : seqIndD K 1 = filter (predC1 ('Ind[L, K] 1)) seqIndT. +Proof. +rewrite filter_undup filter_map (eq_enum (in_set _)) enumT. +congr (undup (map _ _)); apply: eq_filter => i /=. +by rewrite mem_Iirr_ker1 cfInd_irr_eq1. +Qed. + +Lemma seqIndC1_rem : seqIndD K 1 = rem ('Ind[L, K] 1) seqIndT. +Proof. by rewrite rem_filter ?seqIndC1_filter ?undup_uniq. Qed. + +Section SeqIndD. + +Variables H0 H M : {group gT}. + +Local Notation S := (seqIndD H M). + +Lemma cfAut_seqInd u : cfAut_closed u S. +Proof. +move=> _ /seqIndP[i /setDP[kMi not_kHi] ->]; rewrite cfAutInd -aut_IirrE. +by apply/seqIndP; exists (aut_Iirr u i); rewrite // inE !Iirr_ker_aut not_kHi. +Qed. + +Lemma seqInd_conjC_subset1 : H \subset H0 -> cfConjC_subset S (seqIndD H0 1). +Proof. +move=> sHH0; split; [exact: seqInd_uniq | apply: seqIndS | exact: cfAut_seqInd]. +by rewrite Iirr_kerDS ?sub1G. +Qed. + +Lemma seqInd_sub_aut_zchar u : + {in S, forall phi, phi - cfAut u phi \in 'Z[S, K^#]}. +Proof. +move=> phi Sphi /=; rewrite sub_aut_zchar ?seqInd_zchar ?cfAut_seqInd //. +exact: seqInd_vcharW. +Qed. + +Hypothesis sHK : H \subset K. + +Lemma seqInd_sub : {subset S <= seqIndD K 1}. +Proof. by apply: seqIndS; exact: Iirr_kerDS (sub1G M) sHK. Qed. + +Lemma seqInd_ortho_Ind1 : {in S, forall phi, '[phi, 'Ind[L, K] 1] = 0}. +Proof. +move=> _ /seqInd_sub/seqIndC1P[i nzi ->]. +by rewrite -irr0 not_cfclass_Ind_ortho // irr0 cfclass1 // inE irr_eq1. +Qed. + +Lemma seqInd_ortho_cfuni : {in S, forall phi, '[phi, '1_K] = 0}. +Proof. +move=> phi /seqInd_ortho_Ind1/eqP; apply: contraTeq => not_o_phi_1K. +by rewrite cfInd_cfun1 // cfdotZr rmorph_nat mulf_neq0. +Qed. + +Lemma seqInd_ortho_1 : {in S, forall phi, '[phi, 1] = 0}. +Proof. +move=> _ /seqInd_sub/seqIndC1P[i nzi ->]. +by rewrite -cfdot_Res_r cfRes_cfun1 // -irr0 cfdot_irr (negbTE nzi). +Qed. + +Lemma sum_seqIndD_square : + H <| L -> M <| L -> M \subset H -> + \sum_(phi <- S) phi 1%g ^+ 2 / '[phi] = #|L : H|%:R * (#|H : M|%:R - 1). +Proof. +move=> nsHL nsML sMH; rewrite -(Lagrange_index sKL sHK) natrM -/e -mulrA. +rewrite -sum_Iirr_kerD_square ?(normalS _ sKL) ?(subset_trans sMH) //. +pose h i := @Ordinal (size S).+1 _ (index_size ('Ind 'chi[K]_i) S). +rewrite (partition_big h (ltn^~ (size S))) => /= [|i Xi]; last first. + by rewrite index_mem mem_seqInd. +rewrite big_distrr big_ord_narrow //= big_index_uniq ?seqInd_uniq //=. +apply: eq_big_seq => phi Sphi; rewrite /eq_op insubT ?index_mem //= => _. +have /seqIndP[i kHMi def_phi] := Sphi. +have/cfunP/(_ 1%g) := scaled_cfResInd_sum_cfclass i nsKL. +rewrite !cfunE sum_cfunE -def_phi cfResE // mulrAC => ->; congr (_ * _). +rewrite reindex_cfclass //=; apply/esym/eq_big => j; last by rewrite !cfunE. +rewrite (sameP (cfclass_Ind_irrP _ _ nsKL) eqP) -def_phi -mem_seqInd //. +by apply/andP/eqP=> [[/(nth_index 0){2}<- /eqP->] | -> //]; exact: nth_index. +Qed. + +Section Odd. + +Hypothesis oddL : odd #|L|. + +Lemma seqInd_conjC_ortho : {in S, forall phi, '[phi, phi^*] = 0}. +Proof. +by move=> _ /seqInd_sub/seqIndC1P[i nzi ->]; exact: odd_induced_orthogonal. +Qed. + +Lemma seqInd_conjC_neq : {in S, forall phi, phi^* != phi}%CF. +Proof. +move=> phi Sphi; apply: contraNneq (cfnorm_seqInd_neq0 Sphi) => {2}<-. +by rewrite seqInd_conjC_ortho. +Qed. + +Lemma seqInd_notReal : ~~ has cfReal S. +Proof. exact/hasPn/seqInd_conjC_neq. Qed. + +Variable chi : 'CF(L). +Hypotheses (irr_chi : chi \in irr L) (Schi : chi \in S). + +Lemma seqInd_conjC_ortho2 : orthonormal (chi :: chi^*)%CF. +Proof. +by rewrite /orthonormal/= cfnorm_conjC irrWnorm ?seqInd_conjC_ortho ?eqxx. +Qed. + +Lemma seqInd_nontrivial_irr : (#|[set i | 'chi_i \in S]| > 1)%N. +Proof. +have /irrP[i Dchi] := irr_chi; rewrite (cardsD1 i) (cardsD1 (conjC_Iirr i)). +rewrite !inE -(inj_eq irr_inj) conjC_IirrE -Dchi seqInd_conjC_neq //. +by rewrite cfAut_seqInd Schi. +Qed. + +Lemma seqInd_nontrivial : (size S > 1)%N. +Proof. +apply: (@leq_trans (size [seq 'chi_i | i in [pred i | 'chi_i \in S]])). + by rewrite size_map -cardE -cardsE seqInd_nontrivial_irr. +apply: uniq_leq_size => [| _ /imageP[i Schi_i ->] //]. +exact/dinjectiveP/(in2W irr_inj). +Qed. + +End Odd. + +End SeqIndD. + +Lemma sum_seqIndC1_square : + \sum_(phi <- seqIndD K 1) phi 1%g ^+ 2 / '[phi] = e * (#|K|%:R - 1). +Proof. by rewrite sum_seqIndD_square ?normal1 ?sub1G // indexg1. Qed. + +End InducedIrrs. + +Implicit Arguments seqIndP [gT K L calX phi]. +Implicit Arguments seqIndC1P [gT K L phi]. + +Section Five. + +Variable gT : finGroupType. + +Section Defs. + +Variables L G : {group gT}. + +(* This is Peterfalvi, Definition (5.1). *) +(* We depart from the text in Section 5 on three points: *) +(* - We drop non-triviality condition in Z[S, A], which is not used *) +(* consistently in the rest of the proof. In particular, it is *) +(* incompatible with the use of "not coherent" in (6.2), and it is only *) +(* really used in (7.8), where it is equivalent to the simpler condition *) +(* (size S > 1). For us the empty S is coherent; this avoids duplicate *) +(* work in some inductive proofs, e.g., subcoherent_norm - Lemma (5.4) - *) +(* belom. *) +(* - The preconditions for coherence (A < L, S < Z[irr L], and tau Z-linear *) +(* on some E < Z[irr L]) are not part of the definition of "coherent". *) +(* These will be captured as separate requirements; in particular in the *) +(* Odd Order proof tau will always be C-linear on all of 'CF(L). *) +(* - By contrast, our "coherent" only supplies an additive (Z-linear) *) +(* isometry, where the source text ambiguously specifies "linear" one. *) +(* When S consists of virtual characters this implies the existence of *) +(* a C-linear one: the linear extension of the restriction of the *) +(* isometry to a basis of the Z-module Z[S]; the latter being given by *) +(* the Smith normal form (see intdiv.v). The weaker requirement lets us *) +(* use the dual_iso construction when size S = 2. *) +(* Finally, note that although we have retained the A parameter, in the *) +(* sequel we shall always take A = L^#, as in the text it is always the case *) +(* that Z[S, A] = Z[S, L^#]. *) +Definition coherent_with S A tau (tau1 : {additive 'CF(L) -> 'CF(G)}) := + {in 'Z[S], isometry tau1, to 'Z[irr G]} /\ {in 'Z[S, A], tau1 =1 tau}. + +Definition coherent S A tau := exists tau1, coherent_with S A tau tau1. + +(* This is Peterfalvi, Hypothesis (5.2). *) +(* The Z-linearity constraint on tau will be expressed by an additive or *) +(* linear structure on tau. *) +Definition subcoherent S tau R := + [/\ (*a*) [/\ {subset S <= character}, ~~ has cfReal S & conjC_closed S], + (*b*) {in 'Z[S, L^#], isometry tau, to 'Z[@irr gT G, G^#]}, + (*c*) pairwise_orthogonal S, + (*d*) {in S, forall xi : 'CF(L : {set gT}), + [/\ {subset R xi <= 'Z[irr G]}, orthonormal (R xi) + & tau (xi - xi^*)%CF = \sum_(alpha <- R xi) alpha]} + & (*e*) {in S &, forall xi phi : 'CF(L), + orthogonal phi (xi :: xi^*%CF) -> orthogonal (R phi) (R xi)}]. + +Definition dual_iso (nu : {additive 'CF(L) -> 'CF(G)}) := + [additive of -%R \o nu \o cfAut conjC]. + +End Defs. + +Section SubsetCoherent. + +Variables L G : {group gT}. +Implicit Type tau : 'CF(L) -> 'CF(G). + +Lemma subgen_coherent S1 S2 A tau: + {subset S2 <= 'Z[S1]} -> coherent S1 A tau -> coherent S2 A tau. +Proof. +move/zchar_trans=> sS21 [tau1 [[Itau1 Ztau1] def_tau]]. +exists tau1; split; last exact: sub_in1 def_tau. +by split; [exact: sub_in2 Itau1 | exact: sub_in1 Ztau1]. +Qed. + +Lemma subset_coherent S1 S2 A tau: + {subset S2 <= S1} -> coherent S1 A tau -> coherent S2 A tau. +Proof. +by move=> sS21; apply: subgen_coherent => phi /sS21/mem_zchar->. +Qed. + +Lemma subset_coherent_with S1 S2 A tau (tau1 : {additive 'CF(L) -> 'CF(G)}) : + {subset S1 <= S2} -> coherent_with S2 A tau tau1 -> + coherent_with S1 A tau tau1. +Proof. +move=> /zchar_subset sS12 [Itau1 Dtau1]. +by split=> [|xi /sS12/Dtau1//]; exact: sub_iso_to Itau1. +Qed. + +Lemma perm_eq_coherent S1 S2 A tau: + perm_eq S1 S2 -> coherent S1 A tau -> coherent S2 A tau. +Proof. +by move=> eqS12; apply: subset_coherent => phi; rewrite (perm_eq_mem eqS12). +Qed. + +Lemma dual_coherence S tau R nu : + subcoherent S tau R -> coherent_with S L^# tau nu -> (size S <= 2)%N -> + coherent_with S L^# tau (dual_iso nu). +Proof. +move=> [[charS nrS ccS] _ oSS _ _] [[Inu Znu] Dnu] szS2. +split=> [|{Inu Znu oSS} phi ZSphi]. + have{oSS} ccZS := cfAut_zchar ccS. + have vcharS: {subset S <= 'Z[irr L]} by move=> phi /charS/char_vchar. + split=> [phi1 phi2 Sphi1 Sphi2 | phi Sphi]. + rewrite cfdotNl cfdotNr opprK Inu ?ccZS // cfdot_conjC aut_Cint //. + by rewrite Cint_cfdot_vchar ?(zchar_sub_irr vcharS). + by rewrite rpredN Znu ?ccZS. +rewrite -{}Dnu //; move: ZSphi; rewrite zcharD1E => /andP[]. +case/zchar_nth_expansion=> x Zx -> {phi} /=. +case: S charS nrS ccS szS2 x Zx => [_ _ _ _ x _| eta S1]. + by rewrite big_ord0 !raddf0. +case/allP/andP=> Neta _ /norP[eta'c _] /allP/andP[S1_etac _]. +rewrite inE [_ == _](negPf eta'c) /= in S1_etac. +case S1E: S1 S1_etac => [|u []] // /predU1P[] //= <- _ z Zz. +rewrite big_ord_recl big_ord1 !raddfD !raddfZ_Cint //=. +rewrite !cfunE (conj_Cnat (Cnat_char1 Neta)) -mulrDl mulf_eq0. +rewrite addr_eq0 char1_eq0 // !scalerN /= cfConjCK addrC. +by case/pred2P => ->; rewrite ?raddf0 //= !scaleNr opprK. +Qed. + +Lemma coherent_seqInd_conjCirr S tau R nu r : + subcoherent S tau R -> coherent_with S L^# tau nu -> + let chi := 'chi_r in let chi2 := (chi :: chi^*)%CF in + chi \in S -> + [/\ {subset map nu chi2 <= 'Z[irr G]}, orthonormal (map nu chi2), + chi - chi^*%CF \in 'Z[S, L^#] & (nu chi - nu chi^*)%CF 1%g == 0]. +Proof. +move=> [[charS nrS ccS] [_ Ztau] oSS _ _] [[Inu Znu] Dnu] chi chi2 Schi. +have sSZ: {subset S <= 'Z[S]} by apply: mem_zchar. +have vcharS: {subset S <= 'Z[irr L]} by move=> phi /charS/char_vchar. +have Schi2: {subset chi2 <= 'Z[S]} by apply/allP; rewrite /= !sSZ ?ccS. +have Schi_diff: chi - chi^*%CF \in 'Z[S, L^#]. + by rewrite sub_aut_zchar // zchar_onG sSZ ?ccS. +split=> // [_ /mapP[xi /Schi2/Znu ? -> //]||]. + apply: map_orthonormal; first by apply: sub_in2 Inu; exact: zchar_trans_on. + rewrite orthonormalE (conjC_pair_orthogonal ccS) //=. + by rewrite cfnorm_conjC !cfnorm_irr !eqxx. +by rewrite -raddfB -cfunD1E Dnu // irr_vchar_on ?Ztau. +Qed. + +End SubsetCoherent. + +(* This is Peterfalvi (5.3)(a). *) +Lemma irr_subcoherent (L G : {group gT}) S tau : + cfConjC_subset S (irr L) -> ~~ has cfReal S -> + {in 'Z[S, L^#], isometry tau, to 'Z[irr G, G^#]} -> + {R | subcoherent S tau R}. +Proof. +case=> U_S irrS ccS nrS [isoL Ztau]. +have N_S: {subset S <= character} by move=> _ /irrS/irrP[i ->]; exact: irr_char. +have vcS: {subset S <= 'Z[irr L]} by move=> chi /N_S/char_vchar. +have o1SS: orthonormal S by exact: sub_orthonormal (irr_orthonormal L). +have [[_ dotSS] oSS] := (orthonormalP o1SS, orthonormal_orthogonal o1SS). +have freeS := orthogonal_free oSS. +pose beta chi := tau (chi - chi^*)%CF; pose eqBP := _ =P beta _. +have Zbeta: {in S, forall chi, chi - (chi^*)%CF \in 'Z[S, L^#]}. + move=> chi Schi; rewrite /= zcharD1E rpredB ?mem_zchar ?ccS //= !cfunE. + by rewrite subr_eq0 conj_Cnat // Cnat_char1 ?N_S. +pose sum_beta chi R := \sum_(alpha <- R) alpha == beta chi. +pose Zortho R := all (mem 'Z[irr G]) R && orthonormal R. +have R chi: {R : 2.-tuple 'CF(G) | (chi \in S) ==> sum_beta chi R && Zortho R}. + apply: sigW; case Schi: (chi \in S) => /=; last by exists [tuple 0; 0]. + move/(_ _ Schi) in Zbeta; have /irrP[i def_chi] := irrS _ Schi. + have: '[beta chi] = 2%:R. + rewrite isoL // cfnormBd ?dotSS ?ccS ?eqxx // eq_sym -/(cfReal _). + by rewrite (negPf (hasPn nrS _ _)). + case/zchar_small_norm; rewrite ?(zcharW (Ztau _ _)) // => R [oR ZR sumR]. + by exists R; apply/and3P; split; [exact/eqP | exact/allP | ]. +exists (fun xi => val (val (R xi))); split=> // [chi Schi | chi phi Schi Sphi]. + by case: (R chi) => Rc /=; rewrite Schi => /and3P[/eqBP-> /allP]. +case/andP => /and3P[/= /eqP opx /eqP opx' _] _. +have{opx opx'} obpx: '[beta phi, beta chi] = 0. + rewrite isoL ?Zbeta // cfdotBl !cfdotBr -{3}[chi]cfConjCK. + by rewrite !cfdot_conjC opx opx' rmorph0 !subr0. +case: (R phi) => [[[|a [|b []]] //= _]]. +rewrite Sphi => /and3P[/eqBP sum_ab Zab o_ab]. +case: (R chi) => [[[|c [|d []]] //= _]]. +rewrite Schi => /and3P[/eqBP sum_cd Zcd o_cd]. +suffices: orthonormal [:: a; - b; c; d]. + rewrite (orthonormal_cat [:: a; _]) => /and3P[_ _]. + by rewrite /orthogonal /= !cfdotNl !oppr_eq0. +apply: vchar_pairs_orthonormal 1 (-1) _ _ _ _. +- by split; apply/allP; rewrite //= rpredN. +- by rewrite o_cd andbT /orthonormal/= cfnormN /orthogonal /= cfdotNr !oppr_eq0. +- by rewrite oppr_eq0 oner_eq0 rpredN rpred1. +rewrite !(big_seq1, big_cons) in sum_ab sum_cd. +rewrite scale1r scaleN1r !opprK sum_ab sum_cd obpx eqxx /=. +by rewrite !(cfun_on0 (zchar_on (Ztau _ _))) ?Zbeta ?inE ?eqxx. +Qed. + +(* This is Peterfalvi (5.3)(b). *) +Lemma prDade_subcoherent (G L K H W W1 W2 : {group gT}) A A0 S + (defW : W1 \x W2 = W) (ddA : prime_Dade_hypothesis G L K H A A0 defW) + (w_ := fun i j => cyclicTIirr defW i j) (sigma := cyclicTIiso ddA) + (mu := primeTIred ddA) (delta := fun j => primeTIsign ddA j) + (tau := Dade ddA) : + let dsw j k := [seq delta j *: sigma (w_ i k) | i : Iirr W1] in + let Rmu j := dsw j j ++ map -%R (dsw j (conjC_Iirr j)) in + cfConjC_subset S (seqIndD K L H 1) -> ~~ has cfReal S -> + {R | [/\ subcoherent S tau R, + {in [predI S & irr L] & irr W, + forall phi w, orthogonal (R phi) (sigma w)} + & forall j, R (mu j) = Rmu j ]}. +Proof. +pose mu2 i j := primeTIirr ddA i j. +set S0 := seqIndD K L H 1 => dsw Rmu [uS sSS0 ccS] nrS. +have nsKL: K <| L by have [[/sdprod_context[]]] := prDade_prTI ddA. +have /subsetD1P[sAK notA1]: A \subset K^# by have [_ []] := prDade_def ddA. +have [_ _ defA0] := prDade_def ddA. +have defSA: 'Z[S, L^#] =i 'Z[S, A]. + have sS0A1: {subset S0 <= 'CF(L, 1%g |: A)}. + move=> _ /seqIndP[i /setDP[_ kerH'i] ->]; rewrite inE in kerH'i. + exact: (prDade_Ind_irr_on ddA) kerH'i. + move=> phi; have:= zcharD1_seqInd_Dade nsKL notA1 sS0A1 phi. + rewrite !{1}(zchar_split _ A, zchar_split _ L^#) => eq_phiAL. + by apply: andb_id2l => /(zchar_subset sSS0) S0phi; rewrite S0phi in eq_phiAL. +have Itau: {in 'Z[S, L^#], isometry tau, to 'Z[irr G, G^#]}. + apply: sub_iso_to sub_refl (Dade_Zisometry _) => phi; rewrite defSA => SAphi. + rewrite defA0; apply: zchar_onS (subsetUl _ _) _ _. + by apply: zchar_sub_irr SAphi => ? /sSS0/seqInd_vcharW. +have orthoS: pairwise_orthogonal S. + exact: sub_pairwise_orthogonal sSS0 uS (seqInd_orthogonal nsKL _). +pose S1 := filter (mem (irr L)) S. +have sS1S: {subset S1 <= S} by apply: mem_subseq; exact: filter_subseq. +have sZS1S: {subset 'Z[S1, L^#] <= 'Z[S, L^#]}. + by apply: zchar_subset sS1S; exact: orthogonal_free. +have [||R1 cohR1] := irr_subcoherent _ _ (sub_iso_to sZS1S sub_refl Itau). +- split=> [|phi|phi]; rewrite ?mem_filter ?filter_uniq //; try case/andP=> //. + by case/irrP=> i {2}-> /=/ccS->; rewrite cfConjC_irr. +- by apply/hasPn=> phi /sS1S/(hasPn nrS). +have{cohR1} [[charS1 _ _] _ _ R1ok R1ortho] := cohR1. +pose R phi := oapp Rmu (R1 phi) [pick j | phi == mu j]. +have inS1 phi: [pred j | phi == mu j] =1 pred0 -> phi \in S -> phi \in S1. + move=> mu'phi Sphi; rewrite mem_filter Sphi andbT /=. + have{Sphi} /seqIndP[ell _ Dphi] := sSS0 _ Sphi; rewrite Dphi. + have [[j Dell] | [] //] := prTIres_irr_cases ddA ell. + by have /=/eqP[] := mu'phi j; rewrite Dphi Dell cfInd_prTIres. +have Smu_nz j: mu j \in S -> j != 0. + move/(hasPn nrS); apply: contraNneq => ->. + by rewrite /cfReal -(prTIred_aut ddA) aut_Iirr0. +have oS1sigma phi: phi \in S1 -> orthogonal (R1 phi) (map sigma (irr W)). + move=> S1phi; have [zR1 oR1] := R1ok _ S1phi; set psi := _ - _=> Dpsi. + suffices o_psi_sigma: orthogonal (tau psi) (map sigma (irr W)). + apply/orthogonalP=> aa sw R1aa Wsw; have:= orthoPl o_psi_sigma _ Wsw. + have{sw Wsw} /dirrP[bw [lw ->]]: sw \in dirr G. + have [_ /(cycTIirrP defW)[i [j ->]] ->] := mapP Wsw. + exact: cycTIiso_dirr. + have [|ba [la Daa]] := vchar_norm1P (zR1 _ R1aa). + by have [_ -> //] := orthonormalP oR1; rewrite eqxx. + rewrite Daa cfdotZl !cfdotZr cfdot_irr. + case: eqP => [<-{lw} | _ _]; last by rewrite !mulr0. + move/(congr1 ( *%R ((-1) ^+ (ba (+) bw))^*)); rewrite mulr0 => /eqP/idPn[]. + rewrite mulrA -rmorphM -signr_addb {bw}addbK -cfdotZr -{ba la}Daa. + rewrite Dpsi -(eq_bigr _ (fun _ _ => scale1r _)). + by rewrite cfproj_sum_orthonormal ?oner_eq0. + apply/orthoPl=> _ /mapP[_ /(cycTIirrP defW)[i [j ->]] ->]; rewrite -/w_. + pose w1 := #|W1|; pose w2 := #|W2|. + have minw_gt2: (2 < minn w1 w2)%N. + have [[_ ntW1 _ _] [ntW2 _ _] _] := prDade_prTI ddA. + rewrite -(dprod_card defW) odd_mul => /andP[oddW1 oddW2]. + by rewrite leq_min !odd_gt2 ?cardG_gt1. + apply: contraTeq (minw_gt2) => ntNC; rewrite -leqNgt. + pose NC := cyclicTI_NC ddA. + have /andP[/=/irrP[l Dphi] Sphi]: phi \in [predI irr L & S]. + by rewrite mem_filter in S1phi. + have Zpsi: psi \in 'Z[S, L^#]. + rewrite sub_aut_zchar ?mem_zchar_on ?orthogonal_free ?ccS ?cfun_onG //. + by move=> ? /sSS0/seqInd_vcharW. + have NCpsi_le2: (NC (tau psi) <= 2)%N. + have{Itau} [Itau Ztau] := Itau. + suff: '[tau psi] <= 2%:R by apply: cycTI_NC_norm; apply: zcharW (Ztau _ _). + rewrite Itau // cfnormBd; first by rewrite cfnorm_conjC Dphi cfnorm_irr. + have /pairwise_orthogonalP[_ -> //] := orthoS; first exact: ccS. + by rewrite eq_sym (hasPn nrS). + apply: leq_trans (NCpsi_le2). + have: (0 < NC (tau psi) < 2 * minn w1 w2)%N. + rewrite -(subnKC minw_gt2) (leq_ltn_trans NCpsi_le2) // andbT lt0n. + by apply/existsP; exists (i, j); rewrite /= topredE inE. + apply: cycTI_NC_minn (ddA) _ _ => x Vx. + rewrite Dade_id; last by rewrite defA0 inE orbC mem_class_support. + rewrite defSA in Zpsi; rewrite (cfun_on0 (zchar_on Zpsi)) // -in_setC. + by apply: subsetP (subsetP (prDade_supp_disjoint ddA) x Vx); rewrite setCS. +exists R; split=> [|phi w S1phi irr_w|j]; first 1 last. +- rewrite /R; case: pickP => [j /eqP Dphi | _ /=]. + by case/nandP: S1phi; right; rewrite /= Dphi (prTIred_not_irr ddA). + apply/orthoPr=> aa R1aa; rewrite (orthogonalP (oS1sigma phi _)) ?map_f //. + by rewrite mem_filter andbC. +- by rewrite /R; case: pickP => /= [k /eqP/(prTIred_inj ddA)-> | /(_ j)/eqP]. +have Zw i j: w_ i j \in 'Z[irr W] by exact: irr_vchar. +have{oS1sigma} oS1dsw psi j: psi \in S1 -> orthogonal (R1 psi) (dsw _ j). + move/oS1sigma/orthogonalP=> opsiW. + apply/orthogonalP=> aa _ R1aa /codomP[i ->]. + by rewrite cfdotZr opsiW ?map_f ?mem_irr ?mulr0. +have odsw j1 j2: j1 != j2 -> orthogonal (dsw _ j1) (dsw _ j2). + move/negPf=> j2'1; apply/orthogonalP=> _ _ /codomP[i1 ->] /codomP[i2 ->]. + by rewrite cfdotZl cfdotZr (cfdot_cycTIiso ddA) j2'1 andbF !mulr0. +split=> // [|phi Sphi|phi xi Sphi Sxi]. +- by split=> // phi /sSS0; exact: seqInd_char. +- rewrite /R; case: pickP => [j /eqP Dphi /= | /inS1/(_ Sphi)/R1ok//]. + have nz_j: j != 0 by rewrite Smu_nz -?Dphi. + have [Isig Zsig]: {in 'Z[irr W], isometry sigma, to 'Z[irr G]}. + exact: cycTI_Zisometry. + split=> [aa | |]. + - rewrite mem_cat -map_comp => /orP. + by case=> /codomP[i ->]; rewrite ?rpredN rpredZsign Zsig. + - rewrite orthonormal_cat orthogonal_oppr odsw ?andbT; last first. + rewrite -(inj_eq (prTIred_inj ddA)) (prTIred_aut ddA) -/mu -Dphi. + by rewrite eq_sym (hasPn nrS). + suffices oNdsw k: orthonormal (dsw j k). + by rewrite map_orthonormal ?oNdsw //; apply: in2W; exact: opp_isometry. + apply/orthonormalP; split=> [|_ _ /codomP[i1 ->] /codomP[i2 ->]]. + rewrite map_inj_uniq ?enum_uniq // => i1 i2 /(can_inj (signrZK _))/eqP. + by rewrite (cycTIiso_eqE ddA) eqxx andbT => /eqP. + rewrite cfdotZl cfdotZr rmorph_sign signrMK (cfdot_cycTIiso ddA). + by rewrite -(cycTIiso_eqE ddA) (inj_eq (can_inj (signrZK _))). + have [Tstruct [tau1 Dtau1 [_ Dtau]]] := uniform_prTIred_coherent ddA nz_j. + have{Tstruct} [/orthogonal_free freeT _ ccT _ _] := Tstruct. + have phi1c: (phi 1%g)^* = phi 1%g := conj_Cnat (Cnat_seqInd1 (sSS0 _ Sphi)). + rewrite -[tau _]Dtau; last first. + rewrite zcharD1E !cfunE phi1c subrr Dphi eqxx andbT. + by rewrite rpredB ?mem_zchar ?ccT ?image_f ?inE // nz_j eqxx. + rewrite linearB Dphi -(prTIred_aut ddA) !Dtau1 -/w_ -/sigma -/(delta j). + by rewrite big_cat /= !big_map !raddf_sum. +rewrite /R; case: pickP => [j1 /eqP Dxi | /inS1/(_ Sxi)S1xi]; last first. + case: pickP => [j2 _ _ | /inS1/(_ Sphi)S1phi]; last exact: R1ortho. + by rewrite orthogonal_catr orthogonal_oppr !oS1dsw. +case: pickP => [j2 /eqP Dphi | /inS1/(_ Sphi)S1phi _]; last first. + by rewrite orthogonal_sym orthogonal_catr orthogonal_oppr !oS1dsw. +case/andP=> /and3P[/= /eqP o_xi_phi /eqP o_xi_phi'] _ _. +have /eqP nz_xi: '[xi] != 0 := cfnorm_seqInd_neq0 nsKL (sSS0 _ Sxi). +have [Dj1 | j2'1] := eqVneq j1 j2. + by rewrite {2}Dxi Dj1 -Dphi o_xi_phi in nz_xi. +have [Dj1 | j2c'1] := eqVneq j1 (conjC_Iirr j2). + by rewrite {2}Dxi Dj1 /mu (prTIred_aut ddA) -/mu -Dphi o_xi_phi' in nz_xi. +rewrite orthogonal_catl andbC orthogonal_oppl. +rewrite !orthogonal_catr !orthogonal_oppr !odsw ?(inj_eq (aut_Iirr_inj _)) //. +by rewrite (inv_eq (@conjC_IirrK _ _)). +Qed. + +Section SubCoherentProperties. + +Variables (L G : {group gT}) (S : seq 'CF(L)) (R : 'CF(L) -> seq 'CF(G)). +Variable tau : {linear 'CF(L) -> 'CF(G)}. +Hypothesis cohS : subcoherent S tau R. + +Lemma nil_coherent A : coherent [::] A tau. +Proof. +exists [additive of 'Ind[G]]; split=> [|u /zchar_span]; last first. + by rewrite span_nil memv0 => /eqP-> /=; rewrite !raddf0. +split=> [u v | u] /zchar_span; rewrite span_nil memv0 => /eqP->. + by rewrite raddf0 !cfdot0l. +by rewrite raddf0 rpred0. +Qed. + +Lemma subset_subcoherent S1 : cfConjC_subset S1 S -> subcoherent S1 tau R. +Proof. +case=> uS1 sS1 ccS1; have [[N_S nrS _] Itau oS defR oR] := cohS. +split; last 1 [exact: sub_in1 defR | exact: sub_in2 oR]. +- split=> // [xi /sS1/N_S// | ]. + by apply/hasPn; exact: sub_in1 (hasPn nrS). +- by apply: sub_iso_to Itau => //; apply: zchar_subset. +exact: sub_pairwise_orthogonal oS. +Qed. + +Lemma subset_ortho_subcoherent S1 chi : + {subset S1 <= S} -> chi \in S -> chi \notin S1 -> orthogonal S1 chi. +Proof. +move=> sS1S Schi S1'chi; apply/orthoPr=> phi S1phi; have Sphi := sS1S _ S1phi. +have [_ _ /pairwise_orthogonalP[_ -> //]] := cohS. +by apply: contraNneq S1'chi => <-. +Qed. + +Lemma subcoherent_split chi beta : + chi \in S -> beta \in 'Z[irr G] -> + exists2 X, X \in 'Z[R chi] + & exists Y, [/\ beta = X - Y, '[X, Y] = 0 & orthogonal Y (R chi)]. +Proof. +move=> Schi Zbeta; have [_ _ _ /(_ _ Schi)[ZR oRR _] _] := cohS. +have [X RX [Y [defXY oXY oYR]]] := orthogonal_split (R chi) beta. +exists X; last first. + by exists (- Y); rewrite opprK (orthogonal_oppl Y) cfdotNr oXY oppr0. +have [_ -> ->] := orthonormal_span oRR RX; rewrite big_seq rpred_sum // => a Ra. +rewrite rpredZ_Cint ?mem_zchar // -(addrK Y X) -defXY. +by rewrite cfdotBl (orthoPl oYR) // subr0 Cint_cfdot_vchar // ZR. +Qed. + +(* This is Peterfalvi (5.4). *) +(* The assumption X \in 'Z[R chi] has been weakened to '[X, Y] = 0; this *) +(* stronger form of the lemma is needed to strengthen the proof of (5.6.3) so *) +(* that it can actually be reused in (9.11.8), as the text suggests. *) +Lemma subcoherent_norm chi psi (tau1 : {additive 'CF(L) -> 'CF(G)}) X Y : + [/\ chi \in S, psi \in 'Z[irr L] & orthogonal (chi :: chi^*)%CF psi] -> + let S0 := chi - psi :: chi - chi^*%CF in + {in 'Z[S0], isometry tau1, to 'Z[irr G]} -> + tau1 (chi - chi^*)%CF = tau (chi - chi^*)%CF -> + [/\ tau1 (chi - psi) = X - Y, '[X, Y] = 0 & orthogonal Y (R chi)] -> + [/\ (*a*) '[chi] <= '[X] + & (*b*) '[psi] <= '[Y] -> + [/\ '[X] = '[chi], '[Y] = '[psi] + & exists2 E, subseq E (R chi) & X = \sum_(xi <- E) xi]]. +Proof. +case=> Schi Zpsi /and3P[/andP[/eqP ochi_psi _] /andP[/eqP ochic_psi _] _] S0. +move=> [Itau1 Ztau1] tau1dchi [defXY oXY oYR]. +have [[ZS nrS ccS] [tS Zt] oS /(_ _ Schi)[ZR o1R tau_dchi] _] := cohS. +have [/=/andP[S'0 uS] oSS] := pairwise_orthogonalP oS. +have [nRchi Schic] := (hasPn nrS _ Schi, ccS _ Schi). +have ZtauS00: tau1 S0`_0 \in 'Z[irr G] by rewrite Ztau1 ?mem_zchar ?mem_head. +have{ZtauS00} [X1 R_X1 [Y1 [dXY1 oXY1 oY1R]]] := subcoherent_split Schi ZtauS00. +have [uR _] := orthonormalP o1R; have [a Za defX1] := zchar_expansion uR R_X1. +have dotS00R xi: xi \in R chi -> '[tau1 S0`_0, xi] = a xi. + move=> Rxi; rewrite dXY1 cfdotBl (orthoPl oY1R) // subr0. + by rewrite defX1 cfproj_sum_orthonormal. +have nchi: '[chi] = \sum_(xi <- R chi) a xi. + transitivity '[S0`_0, S0`_1]. + rewrite [rhs in _ = rhs]cfdotC cfdotBl !cfdotBr ochi_psi ochic_psi. + by rewrite (oSS _ _ Schic) // !subr0 -cfdotC. + rewrite -Itau1 ?mem_zchar ?mem_nth // tau1dchi tau_dchi cfdot_sumr. + exact: eq_big_seq. +have nX: '[X1] <= '[X] ?= iff (X == X1). + rewrite -subr_eq0 -{1 2}[X](subrK X1) cfnormDd. + rewrite -lerif_subLR subrr -cfnorm_eq0 eq_sym. + by apply: lerif_eq; apply: cfnorm_ge0. + rewrite defX1 cfdot_sumr big1_seq // => xi Rxi; rewrite cfdotZr cfdotBl. + rewrite cfproj_sum_orthonormal // -[X](subrK Y) cfdotDl -defXY dotS00R //. + by rewrite (orthoPl oYR) // addr0 subrr mulr0. +pose is01a xi := a xi == (a xi != 0)%:R. +have leXa xi: a xi <= `|a xi| ^+ 2 ?= iff is01a xi. + apply/lerifP; rewrite /is01a; have /CintP[b ->] := Za xi. + rewrite -intr_norm -rmorphX ltr_int intr_eq0 pmulrn !eqr_int. + by case: b => [[|[|n]]|] //=; rewrite ltr_eexpr. +have{nchi nX} part_a: '[chi] <= '[X] ?= iff all is01a (R chi) && (X == X1). + apply: lerif_trans nX; rewrite nchi defX1 cfnorm_sum_orthonormal //. + by rewrite -big_all !(big_tnth _ _ (R chi)) big_andE; apply: lerif_sum. +split=> [|/lerif_eq part_b]; first by case: part_a. +have [_ /esym] := lerif_add part_a part_b; rewrite -!cfnormBd // -defXY. +rewrite Itau1 ?mem_zchar ?mem_head // eqxx => /andP[a_eq /eqP->]. +split=> //; first by apply/esym/eqP; rewrite part_a. +have{a_eq} [/allP a01 /eqP->] := andP a_eq; rewrite defX1. +exists (filter [preim a of predC1 0] (R chi)); first exact: filter_subseq. +rewrite big_filter [rhs in _ = rhs]big_mkcond /=. +by apply: eq_big_seq => xi /a01/eqP{1}->; rewrite scaler_nat -mulrb. +Qed. + +(* This is Peterfalvi (5.5). *) +Lemma coherent_sum_subseq chi (tau1 : {additive 'CF(L) -> 'CF(G)}) : + chi \in S -> + {in 'Z[chi :: chi^*%CF], isometry tau1, to 'Z[irr G]} -> + tau1 (chi - chi^*%CF) = tau (chi - chi^*%CF) -> + exists2 E, subseq E (R chi) & tau1 chi = \sum_(a <- E) a. +Proof. +set S1 := (chi :: _) => Schi [iso_t1 Zt1] t1cc'. +have freeS1: free S1. + have [[_ nrS ccS] _ oS _ _] := cohS. + by rewrite orthogonal_free ?(conjC_pair_orthogonal ccS). +have subS01: {subset 'Z[chi - 0 :: chi - chi^*%CF] <= 'Z[S1]}. + apply: zchar_trans setT _; apply/allP; rewrite subr0 /= andbT. + by rewrite rpredB !mem_zchar ?inE ?eqxx ?orbT. +have Zt1c: tau1 (chi - 0) \in 'Z[irr G]. + by rewrite subr0 Zt1 ?mem_zchar ?mem_head. +have [X R_X [Y defXY]] := subcoherent_split Schi Zt1c. +case/subcoherent_norm: (defXY); last 2 [by []]. +- by rewrite /orthogonal /= !cfdot0r eqxx Schi cfun0_zchar. +- by split; [apply: sub_in2 iso_t1 | apply: sub_in1 Zt1]. +move=> _ [|_ /eqP]; rewrite cfdot0l ?cfnorm_ge0 // cfnorm_eq0 => /eqP Y0. +case=> E sER defX; exists E => //; rewrite -defX -[X]subr0 -Y0 -[chi]subr0. +by case: defXY. +Qed. + +(* A reformulation of (5.5) that is more convenient to use. *) +Corollary mem_coherent_sum_subseq S1 chi (tau1 : {additive 'CF(L) -> 'CF(G)}) : + cfConjC_subset S1 S -> coherent_with S1 L^# tau tau1 -> chi \in S1 -> + exists2 E, subseq E (R chi) & tau1 chi = \sum_(a <- E) a. +Proof. +move=> uccS1 [Itau1 Dtau1] S1chi; have [uS1 sS1S ccS1] := uccS1. +have S1chi_s: chi^*%CF \in S1 by exact: ccS1. +apply: coherent_sum_subseq; first exact: sS1S. + by apply: sub_iso_to Itau1 => //; apply: zchar_subset; apply/allP/and3P. +apply: Dtau1; rewrite sub_aut_zchar ?zchar_onG ?mem_zchar // => phi /sS1S. +by have [[charS _ _] _ _ _ _] := cohS => /charS/char_vchar. +Qed. + +(* A frequently used consequence of (5.5). *) +Corollary coherent_ortho_supp S1 chi (tau1 : {additive 'CF(L) -> 'CF(G)}) : + cfConjC_subset S1 S -> coherent_with S1 L^# tau tau1 -> + chi \in S -> chi \notin S1 -> + orthogonal (map tau1 S1) (R chi). +Proof. +move=> uccS1 cohS1 Schi S1'chi; have [uS1 sS1S ccS1] := uccS1. +apply/orthogonalP=> _ mu /mapP[phi S1phi ->] Rmu; have Sphi := sS1S _ S1phi. +have [e /mem_subseq Re ->] := mem_coherent_sum_subseq uccS1 cohS1 S1phi. +rewrite cfdot_suml big1_seq // => xi {e Re}/Re Rxi. +apply: orthogonalP xi mu Rxi Rmu; have [_ _ _ _ -> //] := cohS. +rewrite /orthogonal /= !andbT cfdot_conjCr fmorph_eq0. +by rewrite !(orthoPr (subset_ortho_subcoherent sS1S _ _)) ?ccS1 ?eqxx. +Qed. + +(* An even more frequently used corollary of the corollary above. *) +Corollary coherent_ortho S1 S2 (tau1 tau2 : {additive 'CF(L) -> 'CF(G)}) : + cfConjC_subset S1 S -> coherent_with S1 L^# tau tau1 -> + cfConjC_subset S2 S -> coherent_with S2 L^# tau tau2 -> + {subset S2 <= [predC S1]} -> + orthogonal (map tau1 S1) (map tau2 S2). +Proof. +move=> uccS1 cohS1 uccS2 cohS2 S1'2; have [_ sS2S _] := uccS2. +apply/orthogonalP=> mu _ S1mu /mapP[phi S2phi ->]. +have [e /mem_subseq Re ->] := mem_coherent_sum_subseq uccS2 cohS2 S2phi. +rewrite cfdot_sumr big1_seq // => xi {e Re}/Re; apply: orthogonalP mu xi S1mu. +by apply: coherent_ortho_supp; rewrite ?sS2S //; apply: S1'2. +Qed. + +(* A glueing lemma exploiting the corollary above. *) +Lemma bridge_coherent S1 S2 (tau1 tau2 : {additive 'CF(L) -> 'CF(G)}) chi phi : + cfConjC_subset S1 S -> coherent_with S1 L^# tau tau1 -> + cfConjC_subset S2 S -> coherent_with S2 L^# tau tau2 -> + {subset S2 <= [predC S1]} -> + [/\ chi \in S1, phi \in 'Z[S2] & chi - phi \in 'CF(L, L^#)] -> + tau (chi - phi) = tau1 chi - tau2 phi -> + coherent (S1 ++ S2) L^# tau. +Proof. +move=> uccS1 cohS1 uccS2 cohS2 S1'2 [S1chi S2phi chi1_phi] tau_chi_phi. +do [rewrite cfunD1E !cfunE subr_eq0 => /eqP] in chi1_phi. +have [[uS1 sS1S _] [uS2 sS2S _]] := (uccS1, uccS2). +have [[[Itau1 Ztau1] Dtau1] [[Itau2 Ztau2] Dtau2]] := (cohS1, cohS2). +have [[N_S1 _ _] _ oS11 _ _] := subset_subcoherent uccS1. +have [_ _ oS22 _ _] := subset_subcoherent uccS2. +have{N_S1} nz_chi1: chi 1%g != 0; last move/mem_zchar in S1chi. + by rewrite char1_eq0 ?N_S1 //; have [/memPn->] := andP oS11. +have oS12: orthogonal S1 S2. + apply/orthogonalP=> xi1 xi2 Sxi1 Sxi2; apply: orthoPr xi1 Sxi1. + by rewrite subset_ortho_subcoherent ?sS2S //; apply: S1'2. +pose S3 := S1 ++ S2; pose Y := map tau1 S1 ++ map tau2 S2. +have oS33: pairwise_orthogonal S3 by rewrite pairwise_orthogonal_cat oS11 oS22. +have oYY: pairwise_orthogonal Y. + by rewrite pairwise_orthogonal_cat !map_pairwise_orthogonal ?coherent_ortho. +have Z_Y: {subset Y <= 'Z[irr G]}. + move=> xi_tau; rewrite mem_cat => /orP[] /mapP[xi Sxi ->] {xi_tau}. + by rewrite Ztau1 ?mem_zchar. + by rewrite Ztau2 ?mem_zchar. +have nY: map cfnorm Y = map cfnorm (S1 ++ S2). + rewrite !map_cat -!map_comp; congr (_ ++ _). + by apply/eq_in_map => xi S1xi; rewrite /= Itau1 ?mem_zchar. + by apply/eq_in_map => xi S2xi; rewrite /= Itau2 ?mem_zchar. +have [tau3 /eqP defY ZItau3] := Zisometry_of_cfnorm oS33 oYY nY Z_Y. +exists tau3; split=> {ZItau3}// xi; rewrite zcharD1E /= => /andP[S3xi]. +have{defY} [defY1 defY2]: {in S1, tau3 =1 tau1} /\ {in S2, tau3 =1 tau2}. + have:= defY; rewrite map_cat eqseq_cat ?size_map // => /andP[]. + by split; apply/eq_in_map/eqP. +have{S3xi} [xi1 [xi2 [Sxi1 Sxi2 ->] {xi}]]: + exists xi1, exists xi2, [/\ xi1 \in 'Z[S1], xi2 \in 'Z[S2] & xi = xi1 + xi2]. +- have uS3 := free_uniq (orthogonal_free oS33). + have [z Zz ->] := zchar_expansion uS3 S3xi; rewrite big_cat. + pose Y_ S4 := \sum_(mu <- S4) z mu *: mu. + suffices ZS_Y S4: Y_ S4 \in 'Z[S4] by exists (Y_ S1), (Y_ S2). + by rewrite /Y_ big_seq rpred_sum // => psi /mem_zchar/rpredZ_Cint->. +rewrite cfunE addrC addr_eq0 linearD => /eqP xi2_1. +transitivity (tau1 xi1 + tau2 xi2). + have [z1 Zz1 ->] := zchar_nth_expansion Sxi1. + have [z2 Zz2 ->] := zchar_nth_expansion Sxi2. + rewrite !raddf_sum; congr(_ + _); apply: eq_bigr => i _; + by rewrite !raddfZ_Cint -?(defY1, defY2) ?mem_nth. +have Z_S1_1 zeta: zeta \in 'Z[S1] -> zeta 1%g \in Cint. + move=> Szeta; rewrite Cint_vchar1 // (zchar_sub_irr _ Szeta) {zeta Szeta}//. + by move=> zeta /sS1S Szeta; apply: char_vchar; have [[->]] := cohS. +have [Zchi1 Zxi1] := (Z_S1_1 _ S1chi, Z_S1_1 _ Sxi1). +apply: (scalerI nz_chi1); rewrite scalerDr -!raddfZ_Cint // scalerDr. +rewrite -[_ *: _](subrK (xi1 1%g *: chi)) raddfD -[_ + _]addrA. +rewrite -[rhs in _ = tau rhs]addrA linearD Dtau1; last first. + by rewrite zcharD1E rpredB ?rpredZ_Cint ?Z_S1_1 //= !cfunE mulrC subrr. +congr (_ + _); rewrite -[_ *: xi2](addKr (xi1 1%g *: phi)) (raddfD tau2). +rewrite [_ + _]addrA [rhs in tau rhs]addrA linearD; congr (_ + _); last first. + rewrite Dtau2 // zcharD1E rpredD ?rpredZ_Cint ?Z_S1_1 //=. + by rewrite !cfunE mulrC xi2_1 chi1_phi mulrN subrr. +rewrite raddfN (raddfZ_Cint tau1) // (raddfZ_Cint tau2) // -!scalerBr linearZ. +by congr (_ *: _). +Qed. + +(* This is essentially Peterfalvi (5.6.3), which gets reused in (9.11.8). *) +Lemma extend_coherent_with S1 (tau1 : {additive 'CF(L) -> 'CF(G)}) chi phi a X : + cfConjC_subset S1 S -> coherent_with S1 L^# tau tau1 -> + [/\ phi \in S1, chi \in S & chi \notin S1] -> + [/\ a \in Cint, chi 1%g = a * phi 1%g & '[X, a *: tau1 phi] = 0] -> + tau (chi - a *: phi) = X - a *: tau1 phi -> + coherent (chi :: chi^*%CF :: S1) L^# tau. +Proof. +set beta := _ - _ => sS10 cohS1 [S1phi Schi S1'chi] [Za chi1 oXaphi] tau_beta. +have [[uS1 sS1S ccS1] [[Itau1 Ztau1] _]] := (sS10, cohS1). +have [[N_S nrS ccS] ZItau _ R_P _] := cohS; have [Itau Ztau] := ZItau. +have [Sphi [ZR o1R sumR]] := (sS1S _ S1phi, R_P _ Schi). +have Zbeta: beta \in 'Z[S, L^#]. + by rewrite zcharD1E !cfunE -chi1 subrr rpredB ?scale_zchar ?mem_zchar /=. +have o_aphi_R: orthogonal (a *: tau1 phi) (R chi). + have /orthogonalP oS1R := coherent_ortho_supp sS10 cohS1 Schi S1'chi. + by apply/orthoPl=> xi Rxi; rewrite cfdotZl oS1R ?map_f ?mulr0. +have /orthoPl o_chi_S1: orthogonal chi S1. + by rewrite orthogonal_sym subset_ortho_subcoherent. +have Zdchi: chi - chi^*%CF \in 'Z[S, L^#]. + by rewrite sub_aut_zchar ?zchar_onG ?mem_zchar ?ccS // => xi /N_S/char_vchar. +have [||_] := subcoherent_norm _ _ (erefl _) (And3 tau_beta oXaphi o_aphi_R). +- rewrite Schi rpredZ_Cint ?char_vchar ?N_S /orthogonal //= !cfdotZr. + by rewrite cfdot_conjCl !o_chi_S1 ?ccS1 // conjC0 !mulr0 !eqxx. +- apply: sub_iso_to ZItau; [apply: zchar_trans_on; apply/allP | exact: zcharW]. + by rewrite /= Zbeta Zdchi. +case=> [|nX _ [e Re defX]]; first by rewrite !cfnormZ Itau1 ?mem_zchar. +have uR: uniq (R chi) by have [] := orthonormalP o1R. +have{uR} De: e = filter (mem e) (R chi) by apply/subseq_uniqP. +pose ec := filter [predC e] (R chi); pose Xc := - \sum_(xi <- ec) xi. +have defR: perm_eq (e ++ ec) (R chi) by rewrite De perm_filterC. +pose S2 := chi :: chi^*%CF; pose X2 := X :: Xc. +have{nrS} uS2: uniq S2 by rewrite /= andbT inE eq_sym (hasPn nrS). +have sS20: cfConjC_subset S2 S. + by split=> //; apply/allP; rewrite /= ?cfConjCK ?inE ?eqxx ?orbT // ccS Schi. +have oS2: pairwise_orthogonal S2 by have [] := subset_subcoherent sS20. +have nz_chi: chi != 0 by rewrite eq_sym; have [/norP[]] := andP oS2. +have o_chi_chic: '[chi, chi^*] = 0 by have [_ /andP[/andP[/eqP]]] := and3P oS2. +have def_XXc: X - Xc = tau (chi - chi^*%CF). + by rewrite opprK defX -big_cat sumR; apply: eq_big_perm. +have oXXc: '[X, Xc] = 0. + have /span_orthogonal o_e_ec: orthogonal e ec. + by move: o1R; rewrite -(eq_orthonormal defR) orthonormal_cat => /and3P[]. + by rewrite defX /Xc !big_seq o_e_ec ?rpredN ?rpred_sum // => xi /memv_span. +have{o_chi_chic} nXc: '[Xc] = '[chi^*]. + by apply: (addrI '[X]); rewrite -cfnormBd // nX def_XXc Itau // cfnormBd. +have{oXXc} oX2: pairwise_orthogonal X2. + rewrite /pairwise_orthogonal /= oXXc eqxx !inE !(eq_sym 0) -!cfnorm_eq0. + by rewrite nX nXc cfnorm_conjC cfnorm_eq0 orbb nz_chi. +have{nX nXc} nX2: map cfnorm X2 = map cfnorm S2 by congr [:: _; _]. +have [|tau2 [tau2X tau2Xc] Itau2] := Zisometry_of_cfnorm oS2 oX2 nX2. + apply/allP; rewrite /= defX De rpredN !big_seq. + by rewrite !rpred_sum // => xi; rewrite mem_filter => /andP[_ /ZR]. +have{Itau2} cohS2: coherent_with S2 L^# tau tau2. + split=> // psi; rewrite zcharD1E => /andP[/zchar_expansion[//|z Zz ->]]. + rewrite big_cons big_seq1 !cfunE conj_Cnat ?Cnat_char1 ?N_S // addrC addr_eq0. + rewrite -mulNr (inj_eq (mulIf _)) ?char1_eq0 ?N_S // => /eqP->. + by rewrite scaleNr -scalerBr !raddfZ_Cint // raddfB /= tau2X tau2Xc -def_XXc. +have: tau beta = tau2 chi - tau1 (a *: phi) by rewrite tau2X raddfZ_Cint. +apply: (bridge_coherent sS20 cohS2 sS10 cohS1) => //. + by apply/hasPn; rewrite has_sym !negb_or S1'chi (contra (ccS1 _)) ?cfConjCK. +by rewrite mem_head (zchar_on Zbeta) rpredZ_Cint ?mem_zchar. +Qed. + +(* This is Peterfalvi (5.6). *) +Lemma extend_coherent S1 xi1 chi : + cfConjC_subset S1 S -> [/\ xi1 \in S1, chi \in S & chi \notin S1] -> + [/\ (*a*) coherent S1 L^# tau, + (*b*) (xi1 1%g %| chi 1%g)%C + & (*c*) 2%:R * chi 1%g * xi1 1%g < \sum_(xi <- S1) xi 1%g ^+ 2 / '[xi]] -> + coherent (chi :: chi^*%CF :: S1) L^# tau. +Proof. +move=> ccsS1S [S1xi1 Schi notS1chi] [[tau1 cohS1] xi1_dv_chi1 ub_chi1]. +have [[uS1 sS1S ccS1] [[Itau1 Ztau1] Dtau1]] := (ccsS1S, cohS1). +have{xi1_dv_chi1} [a Za chi1] := dvdCP _ _ xi1_dv_chi1. +have [[N_S nrS ccS] ZItau oS R_P oR] := cohS; have [Itau Ztau] := ZItau. +have [Sxi1 [ZRchi o1Rchi sumRchi]] := (sS1S _ S1xi1, R_P _ Schi). +have ocS1 xi: xi \in S1 -> '[chi, xi] = 0. + by apply: orthoPl; rewrite orthogonal_sym subset_ortho_subcoherent. +have /andP[/memPn/=nzS _] := oS; have [Nchi nz_chi] := (N_S _ Schi, nzS _ Schi). +have oS1: pairwise_orthogonal S1 by exact: sub_pairwise_orthogonal oS. +have [freeS freeS1] := (orthogonal_free oS, orthogonal_free oS1). +have nz_nS1 xi: xi \in S1 -> '[xi] != 0 by rewrite cfnorm_eq0 => /sS1S/nzS. +have nz_xi11: xi1 1%g != 0 by rewrite char1_eq0 ?N_S ?nzS. +have inj_tau1: {in 'Z[S1] &, injective tau1} := Zisometry_inj Itau1. +have Z_S1: {subset S1 <= 'Z[S1]} by move=> xi /mem_zchar->. +have inj_tau1_S1: {in S1 &, injective tau1} := sub_in2 Z_S1 inj_tau1. +pose a_ t1xi := S1`_(index t1xi (map tau1 S1)) 1%g / xi1 1%g / '[t1xi]. +have a_E xi: xi \in S1 -> a_ (tau1 xi) = xi 1%g / xi1 1%g / '[xi]. + by move=> S1xi; rewrite /a_ nth_index_map // Itau1 ?Z_S1. +have a_xi1 : a_ (tau1 xi1) = '[xi1]^-1 by rewrite a_E // -mulrA mulVKf //. +have Zachi: chi - a *: xi1 \in 'Z[S, L^#]. + by rewrite zcharD1E !cfunE -chi1 subrr rpredB ?scale_zchar ?mem_zchar /=. +have Ztau_achi := zcharW (Ztau _ Zachi). +have [X R_X [Y defXY]] := subcoherent_split Schi Ztau_achi. +have [eqXY oXY oYRchi] := defXY; pose X1 := map tau1 (in_tuple S1). +have oX1: pairwise_orthogonal X1 by exact: map_pairwise_orthogonal. +have N_S1_1 xi: xi \in S1 -> xi 1%g \in Cnat by move/sS1S/N_S/Cnat_char1. +have oRchiX1 psi: psi \in 'Z[R chi] -> orthogonal psi X1. + move/zchar_span=> Rpsi; apply/orthoPl=> chi2 /memv_span. + by apply: span_orthogonal Rpsi; rewrite orthogonal_sym coherent_ortho_supp. +have [lam Zlam [Z oZS1 defY]]: + exists2 lam, lam \in Cint & exists2 Z : 'CF(G), orthogonal Z (map tau1 S1) & + Y = a *: tau1 xi1 - lam *: (\sum_(xi <- X1) a_ xi *: xi) + Z. +- pose lam := a * '[xi1] - '[Y, tau1 xi1]; exists lam. + rewrite rpredD ?mulr_natl ?rpredN //. + by rewrite rpredM // CintE Cnat_cfdot_char ?N_S. + rewrite Cint_cfdot_vchar ?Ztau1 ?Z_S1 // -(subrK X Y) -opprB -eqXY addrC. + by rewrite rpredB // (zchar_trans ZRchi). + set Z' := _ - _; exists (Y - Z'); last by rewrite addrC subrK. + have oXtau1 xi: xi \in S1 -> '[Y, tau1 xi] = - '[X - Y, tau1 xi]. + move=> S1xi; rewrite cfdotBl opprB. + by rewrite (orthogonalP (oRchiX1 X R_X) X) ?subr0 ?mem_head ?map_f. + apply/orthogonalP=> _ _ /predU1P[-> | //] /mapP[xi S1xi ->]. + rewrite !cfdotBl !cfdotZl Itau1 ?mem_zchar //. + rewrite cfproj_sum_orthogonal ?map_f // a_E // Itau1 ?Z_S1 //. + apply: (mulIf nz_xi11); rewrite divfK ?nz_nS1 // 2!mulrBl mulrA divfK //. + rewrite mul0r mulrBl opprB -addrA addrCA addrC !addrA !oXtau1 // !mulNr. + rewrite -(conj_Cnat (N_S1_1 _ S1xi)) -(conj_Cnat (N_S1_1 _ S1xi1)). + rewrite opprK [- _ + _]addrC -!(mulrC _^*) -!cfdotZr -cfdotBr. + rewrite -!raddfZ_Cnat ?N_S1_1 // -raddfB; set beta : 'CF(L) := _ - _. + have Zbeta: beta \in 'Z[S1, L^#]. + rewrite zcharD1E !cfunE mulrC subrr eqxx. + by rewrite rpredB ?rpredZ_Cint ?Z_S1 // CintE N_S1_1. + rewrite -eqXY Dtau1 // Itau // ?(zchar_subset sS1S) //. + rewrite cfdotBl !cfdotBr !cfdotZr !ocS1 // !mulr0 subrr add0r !cfdotZl. + by rewrite opprB addrAC subrK subrr. +have [|| leXchi _] := subcoherent_norm _ _ (erefl _) defXY. +- rewrite Schi scale_zchar ?char_vchar ?N_S /orthogonal //= !cfdotZr ocS1 //. + by rewrite -[xi1]cfConjCK cfdot_conjC ocS1 ?ccS1 // conjC0 mulr0 eqxx. +- apply: sub_iso_to ZItau; [apply: zchar_trans_on; apply/allP | exact: zcharW]. + rewrite /= Zachi sub_aut_zchar ?zchar_onG ?mem_zchar ?ccS //. + by move=> xi /N_S/char_vchar. +have{defY leXchi lam Z Zlam oZS1 ub_chi1} defY: Y = a *: tau1 xi1. + have nXY: '[X] + '[Y] = '[chi] + '[a *: xi1]. + by rewrite -!cfnormBd // ?cfdotZr ?ocS1 ?mulr0 // -eqXY Itau. + have{leXchi nXY}: '[Y] <= a ^+ 2 * '[xi1]. + by rewrite -(ler_add2l '[X]) nXY cfnormZ Cint_normK // ler_add2r. + rewrite defY cfnormDd; last first. + rewrite cfdotC (span_orthogonal oZS1) ?rmorph0 ?memv_span1 //. + rewrite big_seq memvB ?memvZ ?memv_suml ?memv_span ?map_f //. + by move=> theta S1theta; rewrite memvZ ?memv_span. + rewrite -subr_ge0 cfnormB cfnormZ Cint_normK // Itau1 ?Z_S1 //. + rewrite -2!addrA (opprD (_ * _)) addNKr cfnormZ Cint_normK // oppr_ge0. + rewrite cfnorm_sum_orthogonal //; set sum_a := \sum_(xi <- _) _. + rewrite -cfdotC cfdotC cfdotZl cfdotZr cfproj_sum_orthogonal ?map_f // a_xi1. + rewrite Itau1 ?Z_S1 // 3!rmorphM !(aut_Cint _ Za) fmorphV aut_Cint //. + rewrite -cfdotC -mulr2n 2!mulrA divfK ?nz_nS1 // -mulrnAr addrA => ub_lam. + have [lam0 | nz_lam] := eqVneq lam 0. + suffices /eqP->: Z == 0 by rewrite lam0 scale0r subr0 addr0. + rewrite -cfnorm_eq0 eqr_le cfnorm_ge0 andbT. + by rewrite lam0 -mulrA !mul0r subrr add0r in ub_lam. + set d := \sum_(xi <- _) _ in ub_chi1; pose b := 2%:R * chi 1%g * xi1 1%g / d. + have pos_S1_1 := Cnat_ge0 (Cnat_char1 (N_S _ (sS1S _ _))). + have xi11_gt0: 0 < xi1 1%g by rewrite ltr_def nz_xi11 pos_S1_1. + have d_gt0: 0 < d. + have a_xi_ge0 xi: xi \in S1 -> 0 <= xi 1%g ^+ 2 / '[xi]. + by move/pos_S1_1 => xi_1_pos; rewrite 2?mulr_ge0 // invr_ge0 cfnorm_ge0. + rewrite [d]big_seq; case defS1: {1 2}S1 S1xi1 => // [xi S1'] _. + have{defS1} S1xi: xi \in S1 by rewrite defS1 mem_head. + rewrite big_cons S1xi ltr_spaddl ?sumr_ge0 // ltr_def a_xi_ge0 //=. + by rewrite !mulf_neq0 ?invr_eq0 ?char1_eq0 -?cfnorm_eq0 ?nz_nS1 ?N_S ?sS1S. + have nz_d: d != 0 by rewrite eqr_le ltr_geF. + have b_gt0: 0 < b. + rewrite !pmulr_rgt0 ?ltr0n ?invr_gt0 // lt0r. + by rewrite Cnat_ge0 ?Cnat_char1 ?char1_eq0 ?N_S // nzS. + have{ub_chi1} b_lt1: b < 1 by rewrite ltr_pdivr_mulr ?mul1r. + have{ub_lam} ub_lam: lam ^+ 2 <= b * lam. + rewrite -(ler_pmul2r d_gt0) (mulrAC b) divfK //. + rewrite -[d](divfK (mulf_neq0 nz_xi11 nz_xi11)) chi1 mulr_natl -mulrnAl. + rewrite !mulrA 2!(mulrAC _ _ lam) 2?ler_pmul2r // -mulrA -expr2. + have ->: d / xi1 1%g ^+ 2 = sum_a. + rewrite big_distrl /sum_a big_map !big_seq; apply: eq_bigr => xi S1xi /=. + rewrite a_E // Itau1 ?Z_S1 //= (normr_idP _); last first. + by rewrite !(cfnorm_ge0, mulr_ge0, invr_ge0) ?pos_S1_1. + rewrite mulrAC 2!exprMn -!exprVn [p in p * '[xi]]mulrA. + by rewrite divfK ?nz_nS1. + rewrite -subr_ge0 -opprB oppr_ge0 (ler_trans _ ub_lam) //. + by rewrite (mulrC lam) -{1}[_ - _]addr0 ler_add2l cfnorm_ge0. + have lam_gt0: 0 < lam. + rewrite ltr_def nz_lam -(ler_pmul2l b_gt0) mulr0. + by apply: ler_trans ub_lam; rewrite -Cint_normK // mulr_ge0 ?normr_ge0. + rewrite ler_pmul2r // ltr_geF // in ub_lam. + rewrite (ltr_le_trans b_lt1) //; have:= lam_gt0. + have /CnatP[n ->]: lam \in Cnat by rewrite CnatEint Zlam ltrW. + by rewrite ltr0n ler1n. +by move: eqXY; rewrite defY; apply: extend_coherent_with => //; rewrite -defY. +Qed. + +(* This is Peterfalvi (5.7). *) +(* This is almost a superset of (1.4): we could use it to get a coherent *) +(* isometry, which would necessarily map irreducibles to signed irreducibles. *) +(* It would then only remain to show that the signs are chosen consistently, *) +(* by considering the degrees of the differences. *) +Lemma uniform_degree_coherence : + constant [seq chi 1%g | chi : 'CF(L) <- S] -> coherent S L^# tau. +Proof. +case defS: {1}S => /= [|chi S1] szS; first by rewrite defS; exact: nil_coherent. +have{szS} unifS xi: xi \in S -> xi 1%g = chi 1%g. + by rewrite defS => /predU1P[-> // | S'xi]; apply/eqP/(allP szS)/map_f. +have Schi: chi \in S by rewrite defS mem_head. +have [[N_S nrS ccS] IZtau oS R_P oR] := cohS; have [Itau Ztau] := IZtau. +have freeS := orthogonal_free oS. +have Zd: {in S &, forall xi1 xi2, xi1 - xi2 \in 'Z[S, L^#]}. + move=> xi1 xi2 Sxi1 Sxi2 /=. + by rewrite zcharD1E rpredB ?mem_zchar //= !cfunE !unifS ?subrr. +have [neq_chic Schic] := (hasPn nrS _ Schi, ccS _ Schi). +have [/andP[/memPn notS0 _] ooS] := pairwise_orthogonalP oS. +pose S' xi := [predD1 S & xi]; pose S'c xi := predD1 (S' xi) xi^*%CF. +have{oR} oR xi1 xi2: xi1 \in S -> xi2 \in S'c xi1 -> orthogonal (R xi1) (R xi2). + move=> Sxi1 /and3P[/= neq_xi21c neq_xi21 Sxi2]. + by rewrite orthogonal_sym oR // /orthogonal /= !ooS ?eqxx // ccS. +have oSc xi: xi \in S -> '[xi, xi^*] = 0. + by move=> Sxi; rewrite ooS ?ccS // -[_ == _]negbK eq_sym (hasPn nrS). +pose D xi := tau (chi - xi). +have Z_D xi: xi \in S -> D xi \in 'Z[irr G] by move/(Zd _ _ Schi)/Ztau/zcharW. +have /CnatP[N defN]: '[chi] \in Cnat by rewrite Cnat_cfdot_char ?N_S. +have dotD: {in S' chi &, forall xi1 xi2, '[D xi1, D xi2] = N%:R + '[xi1, xi2]}. +- move=> xi1 xi2 /andP[ne_xi1chi Sxi1] /andP[ne_xi2chi Sxi2]. + rewrite Itau ?Zd // cfdotBl !cfdotBr defN. + by rewrite 2?ooS // 1?eq_sym // opprB !subr0. +have /R_P[ZRchi o1Rchi defRchi] := Schi; have frRchi := orthonormal_free o1Rchi. +have szRchi: size (R chi) = (N + N)%N. + apply: (can_inj natCK); rewrite -cfnorm_orthonormal // -defRchi. + by rewrite dotD ?inE ?ccS ?(hasPn nrS) // cfnorm_conjC defN -natrD. +pose sub_Rchi X := exists2 E, subseq E (R chi) & X = \sum_(a <- E) a. +pose Xspec X := [/\ X \in 'Z[R chi], '[X]_G = N%:R & sub_Rchi X]. +pose Xi_spec X xi := X - D xi \in 'Z[R xi] /\ '[X, D xi] = N%:R. +have haveX xi: xi \in S'c chi -> exists2 X, Xspec X & Xi_spec X xi. + move=> S'xi; have /and3P[/= ne_xi_chi' ne_xi_chi Sxi] := S'xi. + have [neq_xi' Sxi'] := (hasPn nrS xi Sxi, ccS xi Sxi). + have [X RchiX [Y1 defXY1]] := subcoherent_split Schi (Z_D _ Sxi). + have [eqXY1 oXY1 oY1chi] := defXY1; have sRchiX := zchar_span RchiX. + have Z_Y1: Y1 \in 'Z[irr G]. + rewrite -[Y1](subrK X) -opprB -eqXY1 addrC rpredB ?Z_D //. + exact: (zchar_trans ZRchi). + have [X1 RxiX1 [Y defX1Y]] := subcoherent_split Sxi Z_Y1; pose Y2 := X + Y. + have [eqX1Y oX1Y oYxi] := defX1Y; pose D2 := tau (xi - chi). + have oY2Rxi: orthogonal Y2 (R xi). + apply/orthogonalP=> _ phi /predU1P[-> | //] Rxi_phi. + rewrite cfdotDl (orthoPl oYxi) // addr0. + by rewrite (span_orthogonal (oR _ _ _ S'xi)) // (memv_span Rxi_phi). + have{oY2Rxi} defX1Y2: [/\ D2 = X1 - Y2, '[X1, Y2] = 0 & orthogonal Y2 (R xi)]. + rewrite -opprB -addrA -opprB -eqX1Y -eqXY1 -linearN opprB cfdotC. + by rewrite (span_orthogonal oY2Rxi) ?conjC0 ?memv_span1 ?(zchar_span RxiX1). + have [||minX eqX1] := subcoherent_norm _ _ (erefl _) defXY1. + - by rewrite char_vchar ?N_S /orthogonal //= !ooS ?eqxx // eq_sym. + - apply: sub_iso_to IZtau; last exact: zcharW. + by apply: zchar_trans_on; apply/allP; rewrite /= !Zd. + have [||minX1 _]:= subcoherent_norm _ _ (erefl _) defX1Y2. + - rewrite char_vchar ?N_S /orthogonal //= !ooS ?eqxx //. + by rewrite (inv_eq (@cfConjCK _ _)). + - apply: sub_iso_to IZtau; last exact: zcharW. + by apply: zchar_trans_on; apply/allP; rewrite /= !Zd. + have span_head := memv_span (mem_head _ _); have sRxiX1 := zchar_span RxiX1. + have Y0: Y = 0. + apply/eqP; rewrite -cfnorm_eq0 eqr_le cfnorm_ge0 andbT. + rewrite -(ler_add2l ('[X] + '[X1])) -!addrA. + rewrite -2?cfnormBd -?eqX1Y -?eqXY1 ?addr0; first last. + - by rewrite cfdotC (span_orthogonal oYxi) ?rmorph0 ?span_head. + - by rewrite cfdotC (span_orthogonal oY1chi) ?rmorph0 ?span_head. + by rewrite dotD ?inE ?ne_xi_chi // -defN ler_add. + rewrite eqX1Y Y0 subr0 defN in eqX1. + have [nX _ defX] := eqX1 minX1; exists X => //; red. + rewrite eqXY1 eqX1Y Y0 subr0 opprD opprK addNKr cfdotBr nX. + by rewrite (span_orthogonal (oR _ _ _ S'xi)) ?subr0 ?(zchar_span RxiX1). +pose X_spec X := forall xi, X - D xi \in 'Z[irr G] /\ '[X, D xi] = N%:R. +have [X [RchiX nX defX] X_S'c]: exists2 X, Xspec X & {in S'c chi, X_spec X}. + have [S_chi | /allPn[xi1 Sxi1]] := altP (@allP _ (pred2 chi chi^*%CF) S). + pose E := take N (R chi); pose Ec := drop N (R chi). + have eqRchi: E ++ Ec = R chi by rewrite cat_take_drop. + have:= o1Rchi; rewrite -eqRchi orthonormal_cat => /and3P[onE onEc oEEc]. + exists (\sum_(a <- E) a) => [|xi /and3P[? ? /S_chi/norP[] //]]. + split; last by exists E; rewrite // -[E]cats0 -eqRchi cat_subseq ?sub0seq. + rewrite big_seq rpred_sum // => a Ea. + by rewrite mem_zchar // -eqRchi mem_cat Ea. + by rewrite cfnorm_orthonormal //= size_takel ?szRchi ?leq_addl. + case/norP=> ne_xi1chi ne_xi1chi'; have S'xi1: xi1 \in S'c chi by exact/and3P. + have [X [RchiX nX defX] [Rxi1X1 XD_N]] := haveX _ S'xi1. + exists X => // xi S'xi; have [ne_xi_chi' ne_xi_chi /= Sxi] := and3P S'xi. + have /R_P[ZRxi _ _] := Sxi; have /R_P[ZRxi1 _ defRxi1] := Sxi1. + have [-> | ne_xi_xi1] := eqVneq xi xi1; first by rewrite (zchar_trans ZRxi1). + have [sRchiX sRxi1X1] := (zchar_span RchiX, zchar_span Rxi1X1). + have [-> | ne_xi_xi1'] := eqVneq xi xi1^*%CF. + rewrite /D -[chi](subrK xi1) -addrA linearD cfdotDr XD_N opprD addrA. + rewrite defRxi1 big_seq (span_orthogonal (oR _ _ _ S'xi1)) ?addr0 //. + by rewrite rpredB ?rpred_sum // (zchar_trans ZRxi1). + by rewrite memv_suml // => a /memv_span. + have [X' [RchiX' nX' _] [RxiX' X'D_N]] := haveX _ S'xi. + have [ZXi sRxiX'] := (zchar_trans ZRxi RxiX', zchar_span RxiX'). + suffices: '[X - X'] == 0 by rewrite cfnorm_eq0 subr_eq0 => /eqP->. + rewrite cfnormB subr_eq0 nX nX' aut_Cint -?mulr2n; last first. + by rewrite Cint_cfdot_vchar ?(zchar_trans ZRchi). + apply/eqP; congr (_ *+ _); transitivity '[D xi1, D xi]. + by rewrite dotD ?inE ?ne_xi_chi ?ne_xi1chi ?ooS ?addr0 // eq_sym. + rewrite -[D xi](subrK X') -opprB addrC -[D _](subrK X) -opprB addrC. + rewrite cfdotBr cfdotBl -addrA addrC -addrA addrCA cfdotBl opprB. + rewrite (span_orthogonal (oR xi1 xi _ _)) //; last exact/and3P. + rewrite (span_orthogonal (oR chi xi _ _)) // subrr add0r. + rewrite cfdotC (span_orthogonal (oR chi xi1 _ _)) ?rmorph0 ?oppr0 ?add0r //. + exact: (zchar_span RchiX'). +have ZX: X \in 'Z[irr G] := zchar_trans ZRchi RchiX. +have{defX X_S'c} X_S': {in S' chi, X_spec X}. + move=> xi. + have [-> _| ne_xi_chi' S'xi] := eqVneq xi chi^*%CF; last exact/X_S'c/andP. + rewrite /D defRchi {1}big_seq rpredB ?rpred_sum //. + have{defX} [E sER defX] := defX; pose Ec := filter [predC E] (R chi). + have eqRchi: perm_eq (R chi) (E ++ Ec). + by rewrite -(perm_filterC (mem E)) -(subseq_uniqP _ _) ?free_uniq. + have:= o1Rchi; rewrite (eq_orthonormal eqRchi) orthonormal_cat. + case/and3P=> onE _ oEEc. + rewrite (eq_big_perm _ eqRchi) big_cat /= -defX cfdotDr nX defX !big_seq. + by rewrite (span_orthogonal oEEc) ?addr0 // memv_suml // => ? /memv_span. +pose X_ xi := X - D xi. +have X_chi: X_ chi = X by rewrite /X_ /D subrr linear0 subr0. +have{X_S'} ZI_X: {in S, isometry X_, to 'Z[irr G]}. + have dotXD_N xi: xi \in S' chi -> '[X, D xi] = N%:R by case/X_S'. + have S_S': {subset S <= [predU1 chi & [predD1 S' chi & chi]]}. + by move=> xi; rewrite !inE; case: eqP. + split=> [xi1 xi2 Sxi1 Sxi2 | xi]; last first. + by case/S_S'/predU1P=> [-> | /andP[_ /X_S'[]//]]; rewrite X_chi. + have /predU1P[-> | /andP[chi'xi1 S'xi1]] := S_S' _ Sxi1. + have /predU1P[->|/andP[chi'xi2 S'xi2]] := S_S' _ Sxi2; rewrite X_chi ?nX //. + by rewrite cfdotBr nX dotXD_N // subrr ooS // eq_sym. + have /predU1P[->|/andP[chi'xi2 S'xi2]] := S_S' _ Sxi2. + by rewrite X_chi cfdotBl nX cfdotC dotXD_N // rmorph_nat subrr ooS. + rewrite cfdotBl !cfdotBr nX (cfdotC _ X) !dotXD_N // conjC_nat. + by rewrite opprB subrr add0r dotD // addrC addKr. +have [tau1 Dtau1 Itau1] := Zisometry_of_iso oS ZI_X. +exists tau1; split=> // xi; rewrite zcharD1E. +case/andP=> /zchar_expansion[|z Zz ->{xi}]; first exact: free_uniq. +rewrite defS big_cons /= !cfunE addr_eq0 => eq_z. +have{eq_z} ->: z chi = - \sum_(xi <- S1) z xi. + have nz_chi1: chi 1%g != 0 by rewrite char1_eq0 ?N_S // notS0. + apply: (mulIf nz_chi1); rewrite (eqP eq_z) sum_cfunE mulNr mulr_suml. + congr (- _); apply: eq_big_seq => xi S1xi. + by rewrite cfunE unifS // defS !inE S1xi orbT. +rewrite scaleNr scaler_suml addrC -opprB -sumrB !linearN !linear_sum /=. +apply: eq_big_seq => xi S1xi; rewrite -scalerBr !linearZ /= -/(D _). +congr (_ *: - _); rewrite linearB !Dtau1 // ?defS 1?mem_behead //. +by rewrite X_chi opprD addNKr opprK. +Qed. + +End SubCoherentProperties. + +(* A corollary of Peterfalvi (5.7) used (sometimes implicitly!) in the proof *) +(* of lemmas (11.9), (12.4) and (12.5). *) +Lemma pair_degree_coherence L G S (tau : {linear _ -> 'CF(gval G)}) R : + subcoherent S tau R -> + {in S &, forall phi1 phi2 : 'CF(gval L), phi1 1%g == phi2 1%g -> + exists S1 : seq 'CF(L), + [/\ phi1 \in S1, phi2 \in S1, cfConjC_subset S1 S & coherent S1 L^# tau]}. +Proof. +move=> scohS phi1 phi2 Sphi1 Sphi2 /= eq_phi12_1. +have [[N_S _ ccS] _ _ _ _] := scohS. +pose S1 := undup (phi1 :: phi1^* :: phi2 :: phi2^*)%CF. +have sS1S: cfConjC_subset S1 S. + split=> [|chi|chi]; rewrite ?undup_uniq //= !mem_undup; move: chi; apply/allP. + by rewrite /= !ccS ?Sphi1 ?Sphi2. + by rewrite /= !inE !cfConjCK !eqxx !orbT. +exists S1; rewrite !mem_undup !inE !eqxx !orbT; split=> //. +apply: uniform_degree_coherence (subset_subcoherent scohS sS1S) _. +apply/(@all_pred1_constant _ (phi2 1%g))/allP=> _ /mapP[chi S1chi ->] /=. +rewrite mem_undup in S1chi; move: chi S1chi; apply/allP. +by rewrite /= !cfAut_char1 ?N_S // eqxx eq_phi12_1. +Qed. + +(* This is Peterfalvi (5.8). *) +Lemma coherent_prDade_TIred (G H L K W W1 W2 : {group gT}) S A A0 + k (tau1 : {additive 'CF(L) -> 'CF(G)}) + (defW : W1 \x W2 = W) (ddA : prime_Dade_hypothesis G L K H A A0 defW) + (sigma := cyclicTIiso ddA) + (eta_ := fun i j => sigma (cyclicTIirr defW i j)) + (mu := primeTIred ddA) (dk := primeTIsign ddA k) (tau := Dade ddA) : + cfConjC_subset S (seqIndD K L H 1) -> + [/\ ~~ has cfReal S, has (mem (irr L)) S & mu k \in S] -> + coherent_with S L^# tau tau1 -> + let j := conjC_Iirr k in + tau1 (mu k) = dk *: (\sum_i eta_ i k) + \/ tau1 (mu k) = - dk *: (\sum_i eta_ i j) + /\ (forall ell, mu ell \in S -> mu ell 1%g = mu k 1%g -> ell = k \/ ell = j). +Proof. +set phi := tau1 (mu k) => uccS [nrS /hasP[zeta Szeta irr_zeta] Sk] cohS j. +pose sum_eta a ell := \sum_i a i ell *: eta_ i ell. +have [R [subcohS oS1sig defR]] := prDade_subcoherent ddA uccS nrS. +have [[charS _ ccS] _ /orthogonal_free freeS Rok _] := subcohS. +have [[Itau1 _] Dtau1] := cohS. +have natS1 xi: xi \in S -> xi 1%g \in Cnat by move/charS/Cnat_char1. +have k'j: j != k by rewrite -(inj_eq (prTIred_inj ddA)) prTIred_aut (hasPn nrS). +have nzSmu l (Sl : mu l \in S): l != 0. + apply: contraNneq (hasPn nrS _ Sl) => ->. + by rewrite /cfReal -prTIred_aut aut_Iirr0. +have [nzk nzj]: k != 0 /\ j != 0 by rewrite !nzSmu // /mu (prTIred_aut ddA) ccS. +have sSS: cfConjC_subset S S by have:= free_uniq freeS; split. +have{sSS} Dtau1S:= mem_coherent_sum_subseq subcohS sSS cohS. +have o_sum_eta a j1 i j2: j1 != j2 -> '[sum_eta a j1, eta_ i j2] = 0. + move/negPf=> neq_j; rewrite cfdot_suml big1 // => i1 _. + by rewrite cfdotZl cfdot_cycTIiso neq_j andbF mulr0. +have proj_sum_eta a i j1: '[sum_eta a j1, eta_ i j1] = a i j1. + rewrite cfdot_suml (bigD1 i) //= cfdotZl cfdot_cycTIiso !eqxx mulr1. + rewrite big1 ?addr0 // => i1 /negPf i'i1. + by rewrite cfdotZl cfdot_cycTIiso i'i1 mulr0. +have [a Dphi Da0]: exists2 a, phi = sum_eta a k + sum_eta a j + & pred2 0 dk (a 0 k) /\ pred2 0 (- dk) (a 0 j). +- have uRk: uniq (R (mu k)) by have [_ /orthonormal_free/free_uniq] := Rok _ Sk. + have [E sER Dphi] := Dtau1S _ Sk; rewrite /phi Dphi (subseq_uniqP uRk sER). + pose a i ell (alpha := dk *: eta_ i ell) := + if alpha \in E then dk else if - alpha \in E then - dk else 0. + have sign_eq := inj_eq (can_inj (signrZK _)). + have E'Nsk i: (- (dk *: eta_ i k) \in E) = false. + apply/idP=> /(mem_subseq sER); rewrite defR -/dk -/sigma mem_cat -map_comp. + case/orP=> /codomP[i1 /esym/eqP/idPn[]]. + by rewrite -scalerN sign_eq cycTIiso_neqN. + by rewrite (inj_eq oppr_inj) sign_eq cycTIiso_eqE (negPf k'j) andbF. + have E'sj i: (dk *: eta_ i j \in E) = false. + apply/idP=> /(mem_subseq sER); rewrite defR -/dk -/sigma mem_cat -map_comp. + case/orP=> /codomP[i1 /eqP/idPn[]]. + by rewrite sign_eq cycTIiso_eqE (negPf k'j) andbF. + by rewrite /= -scalerN sign_eq cycTIiso_neqN. + exists a; last first. + by rewrite !(fun_if (pred2 _ _)) /= !eqxx !orbT E'Nsk !(if_same, E'sj). + rewrite big_filter big_mkcond defR big_cat !big_map -/dk -/sigma /=. + congr (_ + _); apply: eq_bigr => i _; rewrite /a -/(eta_ i _). + by rewrite E'Nsk; case: ifP => // _; rewrite scale0r. + by rewrite E'sj; case: ifP => _; rewrite (scaleNr, scale0r). +pose V := cyclicTIset defW; have zetaV0: {in V, tau1 zeta =1 \0}. + apply: (ortho_cycTIiso_vanish ddA); apply/orthoPl=> _ /mapP[ww Www ->]. + rewrite (span_orthogonal (oS1sig zeta ww _ _)) ?memv_span1 ?inE ?Szeta //. + have [E sER ->] := Dtau1S _ Szeta; rewrite big_seq rpred_sum // => aa Raa. + by rewrite memv_span ?(mem_subseq sER). +pose zeta1 := zeta 1%g *: mu k - mu k 1%g *: zeta. +have Zzeta1: zeta1 \in 'Z[S, L^#]. + rewrite zcharD1E !cfunE mulrC subrr eqxx andbT. + by rewrite rpredB ?scale_zchar ?mem_zchar // CintE ?natS1. +have /cfun_onP A1zeta1: zeta1 \in 'CF(L, 1%g |: A). + rewrite memvB ?memvZ ?prDade_TIred_on //; have [_ sSS0 _] := uccS. + have /seqIndP[kz /setIdP[kerH'kz _] Dzeta] := sSS0 _ Szeta. + by rewrite Dzeta (prDade_Ind_irr_on ddA) //; rewrite inE in kerH'kz. +have{A1zeta1} zeta1V0: {in V, zeta1 =1 \0}. + move=> x Vx; rewrite /= A1zeta1 // -in_setC. + apply: subsetP (subsetP (prDade_supp_disjoint ddA) x Vx); rewrite setCS. + by rewrite subUset sub1G; have [/= _ _ _ [_ [_ _ /subsetD1P[->]]]] := ddA. +have o_phi_0 i: '[phi, eta_ i 0] = 0 by rewrite Dphi cfdotDl !o_sum_eta ?addr0. +have{o_phi_0 zeta1V0} proj_phi0 i ell: '[phi, eta_ i ell] = '[phi, eta_ 0 ell]. + rewrite -[LHS]add0r -(o_phi_0 0) -[RHS]addr0 -(o_phi_0 i). + apply: (cycTIiso_cfdot_exchange ddA); rewrite -/V => x Vx. + have: tau zeta1 x == 0. + have [_ _ defA0] := prDade_def ddA; rewrite Dade_id ?zeta1V0 //. + by rewrite defA0 inE orbC mem_class_support. + rewrite -Dtau1 // raddfB !raddfZ_Cnat ?natS1 // !cfunE zetaV0 //. + rewrite oppr0 mulr0 addr0 mulf_eq0 => /orP[/idPn[] | /eqP->//]. + by have /irrP[iz ->] := irr_zeta; apply: irr1_neq0. +have Dphi_j i: '[phi, eta_ i j] = a i j. + by rewrite Dphi cfdotDl proj_sum_eta o_sum_eta 1?eq_sym ?add0r. +have Dphi_k i: '[phi, eta_ i k] = a i k. + by rewrite Dphi cfdotDl proj_sum_eta o_sum_eta ?addr0. +have Da_j i: a i j = a 0 j by rewrite -!Dphi_j. +have{proj_phi0} Da_k i: a i k = a 0 k by rewrite -!Dphi_k. +have oW1: #|W1| = #|Iirr W1|. + by rewrite card_Iirr_cyclic //; have [[]] := prDade_prTI ddA. +have{oW1}: `|a 0 j| ^+ 2 + `|a 0 k| ^+ 2 == 1. + apply/eqP/(mulfI (neq0CG W1)); rewrite mulr1 {}[in LHS]oW1. + transitivity '[phi]; last by rewrite Itau1 ?mem_zchar ?cfnorm_prTIred. + rewrite {2}Dphi cfdotDr !cfdot_sumr mulrDr addrC !mulr_natl -!sumr_const. + congr (_ + _); apply: eq_bigr => i _; rewrite cfdotZr mulrC normCK. + by rewrite Dphi_k (Da_k i). + by rewrite Dphi_j (Da_j i). +have{Da0}[/pred2P[]Da0k /pred2P[]Da0j] := Da0; rewrite Da0k Da0j; last 2 first. +- left; rewrite Dphi [sum_eta a j]big1 ?addr0 => [|i _]; last first. + by rewrite Da_j Da0j scale0r. + by rewrite scaler_sumr; apply: eq_bigr => i _; rewrite Da_k Da0k. +- by rewrite normrN normr_sign expr1n (eqr_nat _ 2 1). +- by rewrite normr0 expr0n add0r (eqr_nat _ 0 1). +have{Dphi} Dphi: phi = - dk *: (\sum_i eta_ i j). + rewrite Dphi [sum_eta a k]big1 ?add0r => [|i _]; last first. + by rewrite Da_k Da0k scale0r. + by rewrite raddf_sum; apply: eq_bigr => i _; rewrite Da_j Da0j. +clear 1; right; split=> // l Sl deg_l; apply/pred2P. +have [_ [tau2 Dtau2 [_ Dtau]]] := uniform_prTIred_coherent ddA nzk. +have nz_l: l != 0 := nzSmu l Sl. +have Tmukl: mu k - mu l \in 'Z[uniform_prTIred_seq ddA k, L^#]. + rewrite zcharD1E !cfunE deg_l subrr eqxx andbT. + by rewrite rpredB ?mem_zchar ?image_f // !inE ?nzk ?nz_l ?deg_l eqxx. +pose ak (_ : Iirr W1) (_ : Iirr W2) := dk. +have: phi - tau1 (mu l) = sum_eta ak k - sum_eta ak l. + rewrite -raddfB Dtau1; last first. + by rewrite zcharD1E rpredB ?mem_zchar //= !cfunE deg_l subrr. + by rewrite -[tau _]Dtau // raddfB /= !Dtau2 2!raddf_sum. +have [E /mem_subseq sER ->] := Dtau1S _ Sl. +move/esym/(congr1 (cfdotr (eta_ 0 k))); apply: contra_eqT => /norP[k'l j'l] /=. +rewrite !cfdotBl Dphi_k Da0k proj_sum_eta o_sum_eta // cfdot_suml. +rewrite big_seq big1 ?subr0 ?signr_eq0 // => aa /sER; rewrite defR -map_comp. +rewrite mem_cat => /orP[]/codomP[/= i ->]; rewrite -/(eta_ i _). + by rewrite cfdotZl cfdot_cycTIiso (negPf k'l) andbF mulr0. +rewrite cfdotNl cfdotZl cfdot_cycTIiso (inv_eq (@conjC_IirrK _ _)) -/j. +by rewrite (negPf j'l) andbF mulr0 oppr0. +Qed. + +Section DadeAut. + +Variables (L G : {group gT}) (A : {set gT}). +Implicit Types K H M : {group gT}. +Hypothesis ddA : Dade_hypothesis G L A. + +Local Notation tau := (Dade ddA). +Local Notation "alpha ^\tau" := (tau alpha). + +Section DadeAutIrr. +Variable u : {rmorphism algC -> algC}. +Local Notation "alpha ^u" := (cfAut u alpha). + +(* This is Peterfalvi (5.9)(a), slightly reformulated to allow calS to also *) +(* contain non-irreducible characters; for groups of odd order, the second *) +(* assumption holds uniformly for all calS of the form seqIndD. *) +(* We have stated the coherence assumption directly over L^#; this lets us *) +(* drop the Z{S, A] = Z{S, L^#] assumption, and is more consistent with the *) +(* rest of the proof. *) +Lemma cfAut_Dade_coherent calS tau1 chi : + coherent_with calS L^# tau tau1 -> + (1 < #|[set i | 'chi_i \in calS]|)%N /\ cfAut_closed u calS -> + chi \in irr L -> chi \in calS -> + (tau1 chi)^u = tau1 (chi^u). +Proof. +case=> [[Itau1 Ztau1] tau1_tau] [irrS_gt1 sSuS] /irrP[i {chi}->] Schi. +have sSZS: {subset calS <= 'Z[calS]} by move=> phi Sphi; apply: mem_zchar. +pose mu j := 'chi_j 1%g *: 'chi_i - 'chi_i 1%g *: 'chi_j. +have ZAmu j: 'chi_j \in calS -> mu j \in 'Z[calS, L^#]. + move=> Sxj; rewrite zcharD1E !cfunE mulrC subrr. + by rewrite rpredB //= scale_zchar ?sSZS // ?Cint_Cnat ?Cnat_irr1. +have Npsi j: 'chi_j \in calS -> '[tau1 'chi_j] = 1%:R. + by move=> Sxj; rewrite Itau1 ?sSZS ?cfnorm_irr. +have{Npsi} Dtau1 Sxj := vchar_norm1P (Ztau1 _ (sSZS _ Sxj)) (Npsi _ Sxj). +have [e [r tau1_chi]] := Dtau1 _ Schi; set eps := (-1) ^+ e in tau1_chi. +have{Dtau1} Dtau1 j: 'chi_j \in calS -> exists t, tau1 'chi_j = eps *: 'chi_t. + move=> Sxj; suffices: 0 <= (eps *: tau1 'chi_j) 1%g. + have [f [t ->]] := Dtau1 j Sxj. + have [-> | neq_f_eps] := eqVneq f e; first by exists t. + rewrite scalerA -signr_addb scaler_sign addbC -negb_eqb neq_f_eps. + by rewrite cfunE oppr_ge0 ltr_geF ?irr1_gt0. + rewrite -(pmulr_rge0 _ (irr1_gt0 i)) cfunE mulrCA. + have: tau1 (mu j) 1%g == 0 by rewrite tau1_tau ?ZAmu ?Dade1. + rewrite raddfB 2?raddfZ_Cnat ?Cnat_irr1 // !cfunE subr_eq0 => /eqP <-. + by rewrite tau1_chi cfunE mulrCA signrMK mulr_ge0 ?Cnat_ge0 ?Cnat_irr1. +have SuSirr j: 'chi_j \in calS -> 'chi_(aut_Iirr u j) \in calS. + by rewrite aut_IirrE => /sSuS. +have [j Sxj neq_ij]: exists2 j, 'chi_j \in calS & 'chi_i != 'chi_j. + move: irrS_gt1; rewrite (cardsD1 i) inE Schi ltnS card_gt0 => /set0Pn[j]. + by rewrite !inE -(inj_eq irr_inj) eq_sym => /andP[]; exists j. +have: (tau1 (mu j))^u == tau1 (mu j)^u. + by rewrite !tau1_tau ?cfAut_zchar ?ZAmu ?Dade_aut. +rewrite !raddfB [-%R]lock !raddfZ_Cnat ?Cnat_irr1 //= -lock -!aut_IirrE. +have [/Dtau1[ru ->] /Dtau1[tu ->]] := (SuSirr i Schi, SuSirr j Sxj). +have: (tau1 'chi_i)^u != (tau1 'chi_j)^u. + apply: contraNneq neq_ij => /cfAut_inj/(isometry_raddf_inj Itau1)/eqP. + by apply; rewrite ?sSZS //; apply: rpredB. +have /Dtau1[t ->] := Sxj; rewrite tau1_chi !cfAutZ_Cint ?rpred_sign //. +rewrite !scalerA -!(mulrC eps) -!scalerA -!scalerBr -!aut_IirrE. +rewrite !(inj_eq (scalerI _)) ?signr_eq0 // (inj_eq irr_inj) => /negPf neq_urt. +have [/CnatP[a ->] /CnatP[b xj1]] := (Cnat_irr1 i, Cnat_irr1 j). +rewrite xj1 eq_subZnat_irr neq_urt orbF andbC => /andP[_]. +by rewrite eqn0Ngt -ltC_nat -xj1 irr1_gt0 /= => /eqP->. +Qed. + +End DadeAutIrr. + +(* This covers all the uses of (5.9)(a) in the rest of Peterfalvi, except *) +(* one instance in (6.8.2.1). *) +Lemma cfConjC_Dade_coherent K H M (calS := seqIndD K L H M) tau1 chi : + coherent_with calS L^# (Dade ddA) tau1 -> + [/\ odd #|G|, K <| L & H \subset K] -> chi \in irr L -> chi \in calS -> + (tau1 chi)^*%CF = tau1 chi^*%CF. +Proof. +move=> cohS [oddG nsKL sHK] irr_chi Schi. +apply: (cfAut_Dade_coherent cohS) => //; split; last exact: cfAut_seqInd. +have oddL: odd #|L| by apply: oddSg oddG; have [_] := ddA. +exact: seqInd_nontrivial_irr Schi. +Qed. + +(* This is Peterfalvi (5.9)(b). *) +Lemma Dade_irr_sub_conjC chi (phi := chi - chi^*%CF) : + chi \in irr L -> chi \in 'CF(L, 1%g |: A) -> + exists t, phi^\tau = 'chi_t - ('chi_t)^*%CF. +Proof. +case/irrP=> i Dchi Achi; rewrite {chi}Dchi in phi Achi *. +have [Rchi | notRchi] := eqVneq (conjC_Iirr i) i. + by exists 0; rewrite irr0 cfConjC_cfun1 /phi -conjC_IirrE Rchi !subrr linear0. +have Zphi: phi \in 'Z[irr L, A]. + have notA1: 1%g \notin A by have [] := ddA. + by rewrite -(setU1K notA1) sub_conjC_vchar // zchar_split irr_vchar. +have Zphi_tau: phi^\tau \in 'Z[irr G, G^#]. + by rewrite zchar_split Dade_cfun Dade_vchar ?Zphi. +have norm_phi_tau : '[phi^\tau] = 2%:R. + rewrite Dade_isometry ?(zchar_on Zphi) // cfnormB -conjC_IirrE. + by rewrite !cfdot_irr !eqxx eq_sym (negPf notRchi) rmorph0 addr0 subr0. +have [j [k ne_kj phi_tau]] := vchar_norm2 Zphi_tau norm_phi_tau. +suffices def_k: conjC_Iirr j = k by exists j; rewrite -conjC_IirrE def_k. +have/esym:= eq_subZnat_irr 1 1 k j (conjC_Iirr j) (conjC_Iirr k). +rewrite (negPf ne_kj) orbF /= !scale1r !conjC_IirrE -rmorphB. +rewrite -opprB -phi_tau /= -Dade_conjC // rmorphB /= cfConjCK. +by rewrite -linearN opprB eqxx => /andP[/eqP->]. +Qed. + +End DadeAut. + +End Five. + +Implicit Arguments coherent_prDade_TIred + [gT G H L K W W1 W2 A0 A S0 k tau1 defW].
\ No newline at end of file diff --git a/mathcomp/odd_order/PFsection6.v b/mathcomp/odd_order/PFsection6.v new file mode 100644 index 0000000..d74185a --- /dev/null +++ b/mathcomp/odd_order/PFsection6.v @@ -0,0 +1,1649 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg poly finset center. +Require Import fingroup morphism perm automorphism quotient action zmodp. +Require Import gfunctor gproduct cyclic pgroup commutator gseries nilpotent. +Require Import sylow abelian maximal hall frobenius. +Require Import matrix mxalgebra mxrepresentation vector ssrnum algC algnum. +Require Import classfun character inertia vcharacter integral_char. +Require Import PFsection1 PFsection2 PFsection3 PFsection4 PFsection5. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 6: *) +(* Some Coherence Theorems *) +(* Defined here: *) +(* odd_Frobenius_quotient K L M <-> *) +(* L has odd order, M <| L, K with K / M is nilpotent, and L / H1 is a *) +(* Frobenius group with kernel K / H1, where H1 / M = (K / M)^(1). *) +(* This is the statement of Peterfalvi, Hypothesis (6.4), except for *) +(* the K <| L and subcoherence assumptions, to be required separately. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +(* The main section *) +Section Six. + +Variables (gT : finGroupType) (G : {group gT}). +Implicit Types H K L M : {group gT}. + +(* Grouping lemmas that assume Hypothesis (6.1). *) +Section GeneralCoherence. + +Variables K L : {group gT}. +Local Notation S M := (seqIndD K L K M). +Local Notation calS := (S 1). + +Variables (R : 'CF(L) -> seq 'CF(G)) (tau : {linear 'CF(L) -> 'CF(G)}). + +(* These may need to be grouped, in order to make the proofs of 6.8, 10.10, *) +(* and 12.6 more manageable. *) +Hypotheses (nsKL : K <| L) (solK : solvable K). +Hypothesis Itau : {in 'Z[calS, L^#] &, isometry tau}. +Hypothesis scohS : subcoherent calS tau R. + +Let sKL : K \subset L. Proof. exact: normal_sub. Qed. +Let nKL : L \subset 'N(K). Proof. exact: normal_norm. Qed. +Let orthS: pairwise_orthogonal calS. Proof. by case: scohS. Qed. +Let sSS M : {subset S M <= calS}. Proof. exact: seqInd_sub. Qed. +Let ccS M : conjC_closed (S M). Proof. exact: cfAut_seqInd. Qed. +Let uniqS M : uniq (S M). Proof. exact: seqInd_uniq. Qed. +Let nrS : ~~ has cfReal calS. Proof. by case: scohS => [[]]. Qed. + +Lemma exists_linInd M : + M \proper K -> M <| K -> exists2 phi, phi \in S M & phi 1%g = #|L : K|%:R. +Proof. +move=> ltMK nsMK; have [sMK nMK] := andP nsMK. +have ntKM: (K / M)%g != 1%g by rewrite -subG1 quotient_sub1 // proper_subn. +have [r /andP[_ r1] ntr] := solvable_has_lin_char ntKM (quotient_sol M solK). +exists ('Ind[L, K] ('chi_r %% M)%CF); last first. + by rewrite cfInd1 // cfModE // morph1 (eqP r1) mulr1. +apply/seqIndP; exists (mod_Iirr r); last by rewrite mod_IirrE. +rewrite !inE subGcfker mod_IirrE ?cfker_mod //= andbT. +apply: contraNneq ntr => /(canRL (mod_IirrK nsMK))->. +by rewrite quo_IirrE // irr0 ?cfker_cfun1 ?cfQuo_cfun1. +Qed. + +(* This is Peterfalvi (6.2). *) +Lemma coherent_seqIndD_bound (A B C D : {group gT}) : + [/\ A <| L, B <| L, C <| L & D <| L] -> + (*a*) [/\ A \proper K, B \subset D, D \subset C, C \subset K + & D / B \subset 'Z(C / B)]%g -> + (*b1*) coherent (S A) L^# tau -> + (*b2*) coherent (S B) L^# tau + \/ #|K : A|%:R - 1 <= 2%:R * #|L : C|%:R * sqrtC #|C : D|%:R. +Proof. +move=> [nsAL nsBL nsCL nsDL] [ltAK sBD sDC sCK sDbZC] cohA. +have [|not_ineq] := boolP (_ <= _); [by right | left]. +have sBC := subset_trans sBD sDC; have sBK := subset_trans sBC sCK. +have [sAK nsBK] := (proper_sub ltAK, normalS sBK sKL nsBL). +have{sBC} [nsAK nsBC] := (normalS sAK sKL nsAL, normalS sBC sCK nsBK). +pose wf S1 := [/\ uniq S1, {subset S1 <= calS} & conjC_closed S1]. +pose S1 := [::] ++ S A; set S2 := [::] in S1; rewrite -[S A]/S1 in cohA. +have wfS1: wf S1 by split; [apply: uniqS | apply: sSS | apply: ccS]. +move: {2}_.+1 (ltnSn (size calS - size S1)) => n. +elim: n => // n IHn in (S2) S1 wfS1 cohA *; rewrite ltnS => leSnS1. +have [uniqS1 sS1S ccS1] := wfS1. +have [sAB1 | /allPn[psi /= SBpsi notS1psi]] := altP (@allP _ (mem S1) (S B)). + by apply: subset_coherent cohA. +have [neq_psi_c SBpsic] := (hasPn nrS _ (sSS SBpsi), ccS SBpsi). +have wfS1': wf [:: psi, psi^* & S1]%CF. + split=> [|xi|xi]; rewrite /= !inE 1?andbC. + - rewrite negb_or eq_sym neq_psi_c notS1psi uniqS1 (contra (ccS1 _)) //. + by rewrite cfConjCK. + - by case/predU1P=> [|/predU1P[|/sS1S]] -> //; rewrite (@sSS B). + do 2![case/predU1P=> [-> |]; first by rewrite ?cfConjCK eqxx ?orbT // eq_sym]. + by move/ccS1=> ->; rewrite !orbT. +apply: (IHn [:: psi, psi^* & S2]%CF) => //; last first. + rewrite -subSn ?uniq_leq_size //; try by case: wfS1'. + by rewrite /= subSS (leq_trans _ leSnS1) // leq_sub2l ?leqW. +have [phi SAphi phi1] := exists_linInd ltAK nsAK. +have: [/\ phi \in S1, psi \in calS & psi \notin S1]. + by rewrite mem_cat SAphi orbT (@sSS B). +have /seqIndP[i /setDP[kBi _] def_psi] := SBpsi; rewrite inE in kBi. +move/(extend_coherent scohS); apply; rewrite // {phi SAphi}phi1; split=> //. + by rewrite def_psi cfInd1 // dvdC_mulr // CintE Cnat_irr1. +have Spos xi: xi \in calS -> 0 <= xi 1%g by move/Cnat_seqInd1/Cnat_ge0. +rewrite big_cat sum_seqIndD_square //= -subr_gt0 -addrA ltr_paddl //=. + rewrite big_seq sumr_ge0 // => xi S2xi. + by rewrite !mulr_ge0 ?invr_ge0 ?cfnorm_ge0 ?Spos ?sS1S // mem_cat S2xi. +rewrite mulrC -mulrBl pmulr_rgt0 ?gt0CiG // subr_gt0. +rewrite real_ltrNge ?rpredB ?rpredM ?rpred_nat ?rpred1 //; last first. + by rewrite realE Spos ?(sSS SBpsi). +apply: contra not_ineq => /ler_trans-> //. +rewrite -mulrA ler_pmul2l ?ltr0n // def_psi cfInd1 //. +rewrite -(Lagrange_index sKL sCK) natrM -mulrA ler_pmul2l ?gt0CiG //. +exact: irr1_bound_quo sDbZC. +Qed. + +(* This is Peterfalvi, Theorem (6.3). *) +Theorem bounded_seqIndD_coherent M H H1 : + [/\ M <| L, H <| L & H1 <| L] -> + [/\ M \subset H1, H1 \subset H & H \subset K] -> + (*a*) nilpotent (H / M)%g -> + (*b*) coherent (S H1) L^# tau -> + (*c*) (#|H : H1| > 4 * #|L : K| ^ 2 + 1)%N -> + coherent (S M) L^# tau. +Proof. +move: H1 => A [nsML nsHL nsAL] [sMA sAH sHK] nilHb cohA lbHA. +elim: {A}_.+1 {-2}A (ltnSn #|A|) => // m IHm A leAm in nsAL sMA sAH cohA lbHA *. +have [/group_inj-> // | ltMA] := eqVproper sMA; have [sAL nAL] := andP nsAL. +have{ltMA} [B maxB sMB]: {B : {group gT} | maxnormal B A L & M \subset B}. + by apply: maxgroup_exists; rewrite ltMA normal_norm. +have /andP[ltBA nBL] := maxgroupp maxB; have [sBA not_sAB] := andP ltBA. +have [sBH sBL] := (subset_trans sBA sAH, subset_trans sBA sAL). +have nsBL: B <| L by apply/andP. +suffices{m IHm leAm} cohB: coherent (S B) L^# tau. + apply: IHm cohB _ => //; first exact: leq_trans (proper_card ltBA) _. + by rewrite (leq_trans lbHA) // dvdn_leq // indexgS. +have /andP[sHL nHL] := nsHL. +have sAbZH: (A / B \subset 'Z(H / B))%g. + have nBA := subset_trans sAL nBL; have nsBA : B <| A by apply/andP. + have minBb: minnormal (A / B)%g (L / B)%g. + apply/mingroupP; split=> [|Db /andP[ntDb nDLb] sDAb]. + by rewrite -subG1 quotient_sub1 // not_sAB quotient_norms. + have: Db <| (L / B)%g by rewrite /normal (subset_trans sDAb) ?quotientS. + case/(inv_quotientN nsBL)=> D defDb sBD /andP[sDL nDL]. + apply: contraNeq ntDb => neqDAb; rewrite defDb quotientS1 //. + case/maxgroupP: maxB => /= _ /(_ D) {1}<- //. + rewrite -(quotient_proper (normalS sBD sDL nsBL)) // -defDb. + by rewrite properEneq sDAb neqDAb. + apply/setIidPl; case/mingroupP: minBb => /andP[ntAb nALb]. + apply; rewrite ?subsetIl //. + have nZHb := char_norm_trans (center_char (H / B)) (quotient_norms _ nHL). + rewrite andbC normsI //= meet_center_nil //=; last first. + by rewrite quotient_normal // (normalS sAH sHL). + suffices /homgP[f /= <-]: (H / B)%g \homg (H / M)%g by rewrite morphim_nil. + by apply: homg_quotientS; rewrite ?(subset_trans sHL) ?normal_norm. +have [defA | ltAH] := eqVproper sAH. + by rewrite addn1 defA indexgg in lbHA. +have [sAK ltAK] := (subset_trans sAH sHK, proper_sub_trans ltAH sHK). +case: (@coherent_seqIndD_bound A B H A) => // /idPn[]. +apply: contraL lbHA; rewrite -ltnNge -ltC_nat -(Lagrange_index sHK sAH) natrM. +set x := #|H : A|%:R => ub_x. +have nz_x2: sqrtC x != 0 by rewrite sqrtC_eq0 neq0CiG. +have x2_gt0: 0 < sqrtC x by rewrite ltr_def nz_x2 sqrtC_ge0 ler0n. +have{ub_x}: sqrtC x - (sqrtC x)^-1 <= (2 * #|L : K|)%:R. + rewrite -(ler_pmul2r (gt0CiG K H)) -natrM -mulnA Lagrange_index //. + rewrite natrM -(ler_pmul2r x2_gt0) mulrC mulrBl mulrBr. + rewrite !mulrA -expr2 sqrtCK divff // (ler_trans _ ub_x) // mulrC. + by rewrite ler_add2l ler_opp2 mul1r ler1n. +rewrite -(@ler_pexpn2r _ 2) -?topredE /= ?ler0n //; last first. + rewrite subr_ge0 -(ler_pmul2r x2_gt0) -expr2 mulVf // sqrtCK. + by rewrite ler1n. +rewrite -natrX expnMn -(ler_add2r 2%:R) -addnS natrD. +apply: ltr_le_trans; rewrite sqrrB // exprVn sqrtCK divff //. +by rewrite addrAC subrK addrC -subr_gt0 addrK invr_gt0 gt0CiG. +Qed. + +(* This is the statement of Peterfalvi, Hypothesis (6.4). *) +Definition odd_Frobenius_quotient M (H1 := K^`(1) <*> M) := + [/\ (*a*) odd #|L|, + (*b*) [/\ M <| L, M \subset K & nilpotent (K / M)] + & (*c*) [Frobenius L / H1 with kernel K / H1] ]%g. + +(* This is Peterfalvi (6.5). *) +Lemma non_coherent_chief M (H1 := (K^`(1) <*> M)%G) : + odd_Frobenius_quotient M -> + coherent (S M) L^# tau +\/ [/\ (*a*) chief_factor L H1 K /\ (#|K : H1| <= 4 * #|L : K| ^ 2 + 1)%N + & (*b*) exists2 p : nat, p.-group (K / M)%g /\ ~~ abelian (K / M)%g + & (*c*) ~~ (#|L : K| %| p - 1)]. +Proof. +case=> oddL [nsML sMK nilKb]; rewrite /= -(erefl (gval H1)) => frobLb. +set e := #|L : K|; have odd_e: odd e := dvdn_odd (dvdn_indexg L K) oddL. +have{odd_e} mod1e_lb m: (odd m -> m > 1 -> m == 1 %[mod e] -> 2 * e + 1 <= m)%N. + move=> odd_m m_gt1; rewrite eqn_mod_dvd ?(ltnW m_gt1) //. + rewrite -[m]odd_double_half odd_m subn1 /= -mul2n addn1 ltnS leq_pmul2l //. + rewrite Gauss_dvdr; last by rewrite coprime_sym prime_coprime // dvdn2 odd_e. + by apply: dvdn_leq; rewrite -(subnKC m_gt1). +have nsH1L: H1 <| L by rewrite normalY // (char_normal_trans (der_char 1 K)). +have sH1K: H1 \subset K by rewrite join_subG der_sub. +have cohH1: coherent (S H1) L^# tau. + apply: uniform_degree_coherence (subset_subcoherent scohS _) _ => //. + apply/(@all_pred1_constant _ #|L : K|%:R)/allP=> _ /mapP[chi Schi ->] /=. + have [i /setIdP[_]] := seqIndP Schi; rewrite inE join_subG -lin_irr_der1. + by do 2![case/andP]=> _ /eqP chi1 _ ->; rewrite cfInd1 // chi1 mulr1. +have sMH1: M \subset H1 by apply: joing_subr. +have [ubK | lbK] := leqP; last by left; apply: bounded_seqIndD_coherent lbK. +have{ubK} ubK: (#|K : H1| < (2 * e + 1) ^ 2)%N. + rewrite sqrnD expnMn (leq_ltn_trans ubK) // -subn_gt0 addKn. + by rewrite !muln_gt0 indexg_gt0. +have [-> | neqMH1] := eqVneq M H1; [by left | right]. +have{neqMH1} ltMH1: M \proper H1 by rewrite properEneq neqMH1. +have{frobLb} [[E1b frobLb] [sH1L nH1L]] := (existsP frobLb, andP nsH1L). +have [defLb ntKb _ _ /andP[sE1L _]] := Frobenius_context frobLb. +have nH1K: K \subset 'N(H1) := subset_trans sKL nH1L. +have chiefH1: chief_factor L H1 K. + have ltH1K: H1 \proper K by rewrite /proper sH1K -quotient_sub1 ?subG1. + rewrite /chief_factor nsKL andbT; apply/maxgroupP; rewrite ltH1K. + split=> // H2 /andP[ltH2K nH2L] sH12; have sH2K := proper_sub ltH2K. + have /eqVproper[// | ltH21] := sH12; case/idPn: ubK; rewrite -leqNgt. + have dv_e H3: H1 \subset H3 -> H3 \subset K -> L \subset 'N(H3) -> + #|H3 : H1| == 1 %[mod e]. + - move=> sH13 sH3K nH3L; rewrite eqn_mod_dvd // subn1. + rewrite /e -(index_quotient_eq _ sKL nH1L) ?subIset ?sH1K ?orbT //. + rewrite -[#|_ : _|]divgS ?quotientS // -(sdprod_card defLb) mulKn //. + rewrite -card_quotient ?(subset_trans (subset_trans sH3K sKL)) //. + rewrite regular_norm_dvd_pred ?(subset_trans sE1L) ?quotient_norms //. + apply: semiregular_sym; apply: sub_in1 (Frobenius_reg_compl frobLb). + by apply/subsetP; rewrite setSD ?quotientS. + have dv_H21 := dv_e H2 sH12 sH2K nH2L. + have dv_KH2: #|K : H2| == 1 %[mod e]. + have:= dv_e K sH1K (subxx K) nKL; rewrite -(Lagrange_index sH2K sH12). + by rewrite -modnMmr (eqP dv_H21) modnMmr muln1. + have odd_iK := dvdn_odd (dvdn_indexg _ _) (oddSg (subset_trans _ sKL) oddL). + have iK_gt1 H3 H4: H4 \proper H3 -> (#|H3 : H4| > 1)%N. + by rewrite indexg_gt1 => /andP[]. + by rewrite -(Lagrange_index sH2K sH12) leq_mul ?mod1e_lb ?odd_iK ?iK_gt1. +split=> //; have nMK := subset_trans sKL (normal_norm nsML). +have not_abKb: ~~ abelian (K / M). + apply: contra (proper_subn ltMH1) => /derG1P/trivgP. + rewrite /= join_subG subxx andbT -quotient_der ?quotient_sub1 //. + exact: subset_trans (der_sub 1 K) nMK. +have /is_abelemP[p p_pr /and3P[pKb _ _]]: is_abelem (K / H1)%g. + have: solvable (K / H1)%g by apply: quotient_sol solK. + by case/(minnormal_solvable (chief_factor_minnormal chiefH1)). +have [_ p_dv_Kb _] := pgroup_pdiv pKb ntKb. +have iso3M := third_isog sMH1 (normalS sMK sKL nsML) (normalS sH1K sKL nsH1L). +have pKM: p.-group (K / M)%g. + have /dprodP[_ defKM cKMpp' tiKMpp'] := nilpotent_pcoreC p nilKb. + rewrite -defKM (eqP (forall_inP (nilpotent_sol nilKb) 'O_p^'(_)%G _)). + by rewrite mulg1 pcore_pgroup. + have /isomP[inj_quo im_quo] := quotient_isom (cents_norm cKMpp') tiKMpp'. + rewrite subsetI pcore_sub /= -(injmSK inj_quo) // (morphim_der _ 1) //. + rewrite {inj_quo}im_quo /= -[Q in Q^`(1)%g]quotientMidl defKM. + rewrite -quotient_der ?gFnorm ?quotientS //. + rewrite -quotient_sub1 ?(subset_trans (pcore_sub _ _) (der_norm _ _)) //. + rewrite -[(_ / _)%g]setIid coprime_TIg //. + apply: pnat_coprime (quotient_pgroup _ (pcore_pgroup _ _)). + apply: pgroupS (quotientS _ (pcore_sub _ _)) _. + rewrite /= -quotient_der // -(quotientYidr (subset_trans (der_sub 1 K) nMK)). + by rewrite (isog_pgroup _ iso3M) ?(normalS sMK sKL nsML). +exists p => //; apply: contra not_abKb => e_dv_p1. +rewrite cyclic_abelian // Phi_quotient_cyclic //. +have /homgP[f <-]: (K / M / 'Phi(K / M) \homg K / H1)%g. + apply: homg_trans (isog_hom iso3M). + rewrite homg_quotientS ?gFnorm ?quotient_norms //=. + rewrite quotientYidr ?(subset_trans (der_sub 1 K)) // quotient_der //. + by rewrite (Phi_joing pKM) joing_subl. +rewrite {f}morphim_cyclic // abelian_rank1_cyclic; last first. + by rewrite sub_der1_abelian ?joing_subl. +rewrite (rank_pgroup pKb) (leq_trans (p_rank_le_logn _ _)) //. +rewrite -ltnS -(ltn_exp2l _ _ (prime_gt1 p_pr)) -p_part part_pnat_id //. +rewrite card_quotient // (leq_trans ubK) // leq_exp2r //. +have odd_p: odd p := dvdn_odd p_dv_Kb (quotient_odd _ (oddSg sKL oddL)). +by rewrite mod1e_lb // ?eqn_mod_dvd ?prime_gt0 ?prime_gt1. +Qed. + +(* This is Peterfalvi (6.6). *) +Lemma seqIndD_irr_coherence (Z : {group gT}) (calX := seqIndD K L Z 1) : + odd_Frobenius_quotient 1%G -> + [/\ Z <| L, Z :!=: 1 & Z \subset 'Z(K)]%g -> + {subset calX <= irr L} -> + calX =i [pred chi in irr L | ~~ (Z \subset cfker chi)] + /\ coherent calX L^#tau. +Proof. +move=> Frob_quo1 [nsZL ntZ sZ_ZK] irrX; have [sZL nZL] := andP nsZL. +have abZ: abelian Z by rewrite (abelianS sZ_ZK) ?center_abelian. +have /andP[sZK nZK]: Z <| K := sub_center_normal sZ_ZK. +split=> [chi|]. + apply/idP/andP=> [Xchi | [/irrP[r ->{chi}] nkerZr]]. + rewrite irrX //; case/seqIndP: Xchi => t /setIdP[nkerZt _] ->. + by rewrite inE in nkerZt; rewrite sub_cfker_Ind_irr. + have [t Res_r_t] := neq0_has_constt (Res_irr_neq0 K r). + pose chi := 'Ind[L] 'chi_t; have chi_r: '[chi, 'chi_r] != 0. + by rewrite -cfdot_Res_r cfdotC fmorph_eq0 -irr_consttE. + have Xchi: chi \in calX. + apply/seqIndP; exists t; rewrite // !inE sub1G andbT. + rewrite -(sub_cfker_Ind_irr t sKL nZL). + apply: contra nkerZr => /subset_trans-> //. + by rewrite cfker_constt // cfInd_char ?irr_char //. + case/irrX/irrP: Xchi chi_r (Xchi) => r' ->. + by rewrite cfdot_irr pnatr_eq0 -lt0n; case: eqP => // ->. +have [|[]] := non_coherent_chief Frob_quo1. + by apply: subset_coherent; apply: seqInd_sub. +have [oddL _] := Frob_quo1; rewrite /= joingG1 => frobLb _ [p []]. +set e := #|L : K|; have e_gt0: (e > 0)%N by apply: indexg_gt0. +have isoK1 := isog_symr (quotient1_isog K). +rewrite (isog_abelian isoK1) {isoK1}(isog_pgroup _ isoK1). +have [-> | ntK pK _ not_e_dv_p1] := eqsVneq K [1]; first by rewrite abelian1. +have{ntK} [p_pr p_dv_K _] := pgroup_pdiv pK ntK. +set Y := calX; pose d (xi : 'CF(L)) := logn p (truncC (xi 1%g) %/ e). +have: conjC_closed Y by apply: cfAut_seqInd. +have: perm_eq (Y ++ [::]) calX by rewrite cats0. +have: {in Y & [::], forall xi1 xi2, d xi1 <= d xi2}%N by []. +elim: {Y}_.+1 {-2}Y [::] (ltnSn (size Y)) => // m IHm Y X' leYm leYX' defX ccY. +have sYX: {subset Y <= calX}. + by move=> xi Yxi; rewrite -(perm_eq_mem defX) mem_cat Yxi. +have sX'X: {subset X' <= calX}. + by move=> xi X'xi; rewrite -(perm_eq_mem defX) mem_cat X'xi orbT. +have uniqY: uniq Y. + have: uniq calX := seqInd_uniq L _. + by rewrite -(perm_eq_uniq defX) cat_uniq => /and3P[]. +have sYS: {subset Y <= calS} by move=> xi /sYX/seqInd_sub->. +case homoY: (constant [seq xi 1%g | xi : 'CF(L) <- Y]). + exact: uniform_degree_coherence (subset_subcoherent scohS _) homoY. +have Ndg: {in calX, forall xi : 'CF(L), xi 1%g = (e * p ^ d xi)%:R}. + rewrite /d => _ /seqIndP[i _ ->]; rewrite cfInd1 // -/e. + have:= dvd_irr1_cardG i; have /CnatP[n ->] := Cnat_irr1 i. + rewrite -natrM natCK dvdC_nat mulKn // -p_part => dv_n_K. + by rewrite part_pnat_id // (pnat_dvd dv_n_K). +have [chi Ychi leYchi]: {chi | chi \in Y & {in Y, forall xi, d xi <= d chi}%N}. + have [/eqP/nilP Y0 | ntY] := posnP (size Y); first by rewrite Y0 in homoY. + pose i := [arg max_(i > Ordinal ntY) d Y`_i]. + exists Y`_i; [exact: mem_nth | rewrite {}/i; case: arg_maxP => //= i _ max_i]. + by move=> _ /(nthP 0)[j ltj <-]; apply: (max_i (Ordinal ltj)). +have{homoY} /hasP[xi1 Yxi1 lt_xi1_chi]: has (fun xi => d xi < d chi)%N Y. + apply: contraFT homoY => geYchi; apply: (@all_pred1_constant _ (chi 1%g)). + rewrite all_map; apply/allP=> xi Yxi; rewrite /= !Ndg ?sYX // eqr_nat. + rewrite eqn_pmul2l // eqn_exp2l ?prime_gt1 //. + by rewrite eqn_leq leYchi //= leqNgt (hasPn geYchi). +pose Y' := rem chi^*%CF (rem chi Y); pose X'' := [:: chi, chi^*%CF & X']. +have ccY': conjC_closed Y'. + move=> xi; rewrite !(inE, mem_rem_uniq) ?rem_uniq //. + by rewrite !(inv_eq (@cfConjCK _ _)) cfConjCK => /and3P[-> -> /ccY->]. +have Xchi := sYX _ Ychi; have defY: perm_eq [:: chi, chi^*%CF & Y'] Y. + rewrite (perm_eqrP (perm_to_rem Ychi)) perm_cons perm_eq_sym perm_to_rem //. + by rewrite mem_rem_uniq ?inE ?ccY // (seqInd_conjC_neq _ _ _ Xchi). +apply: perm_eq_coherent (defY) _. +have d_chic: d chi^*%CF = d chi. + by rewrite /d cfunE conj_Cnat // (Cnat_seqInd1 Xchi). +have /andP[uniqY' Y'x1]: uniq Y' && (xi1 \in Y'). + rewrite !(inE, mem_rem_uniq) ?rem_uniq // Yxi1 andbT -negb_or. + by apply: contraL lt_xi1_chi => /pred2P[] ->; rewrite ?d_chic ltnn. +have xi1P: [/\ xi1 \in Y', chi \in calS & chi \notin Y']. + by rewrite Y'x1 sYS ?(inE, mem_rem_uniq) ?rem_uniq // eqxx andbF. +have sY'Y: {subset Y' <= Y} by move=> xi /mem_rem/mem_rem. +apply: (extend_coherent scohS) xi1P _; first by split=> // xi /sY'Y/sYS. +have{defX} defX: perm_eq (Y' ++ X'') calX. + by rewrite (perm_catCA Y' [::_; _]) catA -(perm_eqrP defX) perm_cat2r. +have{d_chic} le_chi_X'': {in X'', forall xi, d chi <= d xi}%N. + by move=> xi /or3P[/eqP-> | /eqP-> | /leYX'->] //; rewrite d_chic. +rewrite !Ndg ?sYX // dvdC_nat dvdn_pmul2l // dvdn_exp2l 1?ltnW //; split=> //. + apply: IHm defX ccY' => [|xi xi' /sY'Y/leYchi le_xi_chi /le_chi_X'']. + by rewrite -ltnS // (leq_trans _ leYm) // -(perm_eq_size defY) ltnW. + exact: leq_trans. +have pos_p n: (0 < p ^ n)%N by rewrite expn_gt0 prime_gt0. +rewrite -!natrM; apply: (@ltr_le_trans _ (e ^ 2 * (p ^ d chi) ^ 2)%:R). + rewrite ltr_nat -expnMn -mulnn mulnAC !mulnA 2?ltn_pmul2r //. + rewrite -mulnA mulnCA ltn_pmul2l // -(subnK lt_xi1_chi) addnS expnS. + rewrite expnD mulnA ltn_pmul2r // -(muln1 3) leq_mul //. + rewrite ltn_neqAle prime_gt1 // eq_sym (sameP eqP (prime_oddPn p_pr)). + by rewrite (dvdn_odd p_dv_K) // (oddSg sKL). +have [r] := seqIndP (sYX _ Ychi); rewrite !inE => /andP[nkerZr _] def_chi. +have d_r: 'chi_r 1%g = (p ^ d chi)%:R. + by apply: (mulfI (neq0CiG L K)); rewrite -cfInd1 // -def_chi -natrM Ndg. +pose sum_p2d S := (\sum_(xi <- S) p ^ (d xi * 2))%N. +pose sum_xi1 (S : seq 'CF(L)) := \sum_(xi <- S) xi 1%g ^+ 2 / '[xi]. +have def_sum_xi1 S: {subset S <= calX} -> sum_xi1 S = (e ^ 2 * sum_p2d S)%:R. + move=> sSX; rewrite big_distrr natr_sum /=; apply: eq_big_seq => xi /sSX Xxi. + rewrite expnM -expnMn natrX -Ndg //. + by have /irrP[i ->] := irrX _ Xxi; rewrite cfnorm_irr divr1. +rewrite -/(sum_xi1 _) def_sum_xi1 ?leC_nat 1?dvdn_leq => [|||_ /sY'Y/sYX] //. + by rewrite muln_gt0 expn_gt0 e_gt0 [_ Y'](bigD1_seq xi1) //= addn_gt0 pos_p. +have coep: coprime e p. + have:= Frobenius_ker_coprime frobLb; rewrite coprime_sym. + have /andP[_ nK'L] := char_normal_trans (der_char 1 K) nsKL. + rewrite index_quotient_eq ?subIset ?der_sub ?orbT {nK'L}// -/e. + have ntKb: (K / K^`(1))%g != 1%g by case/Frobenius_kerP: frobLb. + have [_ _ [k ->]] := pgroup_pdiv (quotient_pgroup _ pK) ntKb. + by rewrite coprime_pexpr. +rewrite -expnM Gauss_dvd ?coprime_expl ?coprime_expr {coep}// dvdn_mulr //=. +have /dvdn_addl <-: p ^ (d chi * 2) %| e ^ 2 * sum_p2d X''. + rewrite big_distrr big_seq dvdn_sum //= => xi /le_chi_X'' le_chi_xi. + by rewrite dvdn_mull // dvdn_exp2l ?leq_pmul2r. +rewrite -mulnDr -big_cat (eq_big_perm _ defX) -(natCK (e ^ 2 * _)) /=. +rewrite -def_sum_xi1 // /sum_xi1 sum_seqIndD_square ?normal1 ?sub1G //. +rewrite indexg1 -(natrB _ (cardG_gt0 Z)) -natrM natCK. +rewrite -(Lagrange_index sKL sZK) mulnAC dvdn_mull //. +have /p_natP[k defKZ]: p.-nat #|K : Z| by rewrite (pnat_dvd (dvdn_indexg K Z)). +rewrite defKZ dvdn_exp2l // -(leq_exp2l _ _ (prime_gt1 p_pr)) -{k}defKZ. +rewrite -leC_nat expnM natrX -d_r ?(ler_trans (irr1_bound r).1) //. +rewrite ler_nat dvdn_leq ?indexgS ?(subset_trans sZ_ZK) //=. +by rewrite -cap_cfcenter_irr bigcap_inf. +Qed. + +End GeneralCoherence. + +(* This is Peterfalvi (6.7). *) +(* In (6.8) we only know initially the P is Sylow in L; perhaps the lemma *) +(* should be stated with this equivalent (but weaker) assumption. *) +Lemma constant_irr_mod_TI_Sylow (Z L P : {group gT}) p i : + p.-Sylow(G) P -> odd #|L| -> normedTI P^# G L -> + [/\ Z <| L, Z :!=: 1%g & Z \subset 'Z(P)] -> + {in Z^# &, forall x y, #|'C_L[x]| = #|'C_L[y]| } -> + let phi := 'chi[G]_i in + {in Z^# &, forall x y, phi x = phi y} -> + {in Z^#, forall x, phi x \in Cint /\ (#|P| %| phi x - phi 1%g)%C}. +Proof. +move=> sylP oddL tiP [/andP[sZL nZL] ntZ sZ_ZP] prZL; move: i. +pose a := @gring_classM_coef _ G; pose C (i : 'I_#|classes G|) := enum_val i. +have [[sPG pP p'PiG] [sZP cPZ]] := (and3P sylP, subsetIP sZ_ZP). +have [ntP sLG memJ_P1] := normedTI_memJ_P tiP; rewrite setD_eq0 subG1 in ntP. +have nsPL: P <| L. + by have [_ _ /eqP<-] := and3P tiP; rewrite normD1 normal_subnorm. +have [p_pr _ [e oP]] := pgroup_pdiv pP ntP. +have [sZG [sPL _]] := (subset_trans sZP sPG, andP nsPL). +pose dC i (A : {set gT}) := [disjoint C i & A]. +have actsGC i: {acts G, on C i | 'J}. + apply/actsP; rewrite astabsJ /C; have /imsetP[x _ ->] := enum_valP i. + by apply/normsP; apply: classGidr. +have{actsGC} PdvKa i j s: + ~~ dC i Z^# -> ~~ dC j Z^# -> dC s Z -> (#|P| %| a i j s * #|C s|)%N. +- pose Omega := [set uv in [predX C i & C j] | mulgm uv \in C s]%g. + pose to_fn uv x := prod_curry (fun u v : gT => (u ^ x, v ^ x)%g) uv. + have toAct: is_action setT to_fn. + by apply: is_total_action => [[u v]|[u v] x y] /=; rewrite ?conjg1 ?conjgM. + move=> Zi Zj Z's; pose to := Action toAct. + have actsPO: [acts P, on Omega | to]. + apply/(subset_trans sPG)/subsetP=> x Gx; rewrite !inE. + apply/subsetP=> [[u v] /setIdP[/andP/=[Ciu Cjv] Csuv]]. + by rewrite !inE /= -conjMg !actsGC // Ciu Cjv. + have <-: #|Omega| = (a i j s * #|C s|)%N. + have /repr_classesP[_ defCs] := enum_valP s; rewrite -/(C s) in defCs. + rewrite -sum1_card mulnC -sum_nat_const. + rewrite (partition_big mulgm (mem (C s))) => [|[u v] /setIdP[]//]. + apply: eq_bigr; rewrite /= defCs => _ /imsetP[z Gz ->]. + rewrite -[a i j s]sum1_card -!/(C _) (reindex_inj (act_inj to z)) /=. + apply: eq_bigl => [[u v]]; rewrite !inE /= -conjMg (inj_eq (conjg_inj _)). + by apply: andb_id2r => /eqP->; rewrite {2}defCs mem_imset ?andbT ?actsGC. + suffices regPO: {in Omega, forall uv, 'C_P[uv | to] = 1%g}. + rewrite -(acts_sum_card_orbit actsPO) dvdn_sum // => _ /imsetP[uv Ouv ->]. + by rewrite card_orbit regPO // indexg1. + case=> u v /setIdP[/andP[/= Ciu Cjv] Csuv]; apply: contraTeq Z's. + case/trivgPn=> x /setIP[Px /astab1P[/= cux cvx]] nt_x. + suffices inZ k y: y \in C k -> ~~ dC k Z^# -> y ^ x = y -> y \in Z. + apply/exists_inP; exists (u * v)%g => //=. + by rewrite groupM // (inZ i u, inZ j v). + rewrite /dC /C; have /imsetP[_ _ ->{k} /class_transr <-] := enum_valP k. + case/exists_inP=> _ /imsetP[g Gg ->] /setD1P[nt_yg Zyg] yx. + have xy: (x ^ y = x)%g by rewrite /conjg (conjgCV x) -{2}yx conjgK mulKg. + rewrite -(memJ_conjg _ g) (normsP nZL) //. + rewrite -(memJ_P1 y) ?inE //=; first by rewrite nt_yg (subsetP sZP). + rewrite -order_eq1 -(orderJ y g) order_eq1 nt_yg. + rewrite (mem_normal_Hall (pHall_subl sPL sLG sylP)) //. + by rewrite -(p_eltJ _ _ g) (mem_p_elt pP) ?(subsetP sZP). + rewrite -(memJ_P1 x) // ?xy ?inE ?nt_x // -[y](conjgK g) groupJ ?groupV //. + by rewrite (subsetP sZG). +pose a2 i j := (\sum_(s | ~~ dC s Z^#) a i j s)%N. +pose kerZ l := {in Z^# &, forall x y, 'chi[G]_l x = 'chi_l y}. +move=> l phi kerZl z Z1z; move: l @phi {kerZl}(kerZl : kerZ l). +have [ntz Zz] := setD1P Z1z. +have [[Pz Lz] Gz] := (subsetP sZP z Zz, subsetP sZL z Zz, subsetP sZG z Zz). +pose inC y Gy := enum_rank_in (@mem_classes _ y G Gy) (y ^: G). +have CE y Gy: C (inC y Gy) = y ^: G by rewrite /C enum_rankK_in ?mem_classes. +pose i0 := inC _ (group1 G); pose i1 := inC z Gz; pose i2 := inC _ (groupVr Gz). +suffices Ea2 l (phi := 'chi[G]_l): + kerZ l -> (phi z *+ a2 i1 i1 == phi 1%g + phi z *+ a2 i1 i2 %[mod #|P|])%A. +- move=> l phi kerZphi. + have Zphi1: phi 1%g \in Cint by rewrite irr1_degree rpred_nat. + have chi0 x: x \in Z -> 'chi[G]_0 x = 1. + by rewrite irr0 cfun1E => /(subsetP sZG) ->. + have: kerZ 0 by move=> x y /setD1P[_ Zx] /setD1P[_ Zy]; rewrite !chi0. + move/Ea2/(eqAmodMl (Aint_irr l z)); rewrite !{}chi0 // -/phi eqAmod_sym. + rewrite mulrDr mulr1 !mulr_natr => /eqAmod_trans/(_ (Ea2 l kerZphi)). + rewrite eqAmodDr -/phi eqAmod_rat ?rpred_nat ?(rpred_Cint _ Zphi1) //. + move=> PdvDphi; split; rewrite // -[phi z](subrK (phi 1%g)) rpredD //. + by have /dvdCP[b Zb ->] := PdvDphi; rewrite rpredM ?rpred_nat. + have nz_Z1: #|Z^#|%:R != 0 :> algC. + by rewrite pnatr_eq0 cards_eq0 setD_eq0 subG1. + rewrite -[phi z](mulfK nz_Z1) rpred_div ?rpred_nat // mulr_natr. + rewrite -(rpredDl _ (rpred_Cint _ Zphi1)) //. + rewrite -[_ + _](mulVKf (neq0CG Z)) rpredM ?rpred_nat //. + have: '['Res[Z] phi, 'chi_0] \in Crat. + by rewrite rpred_Cnat ?Cnat_cfdot_char ?cfRes_char ?irr_char. + rewrite irr0 cfdotE (big_setD1 _ (group1 Z)) cfun1E cfResE ?group1 //=. + rewrite rmorph1 mulr1; congr (_ * (_ + _) \in Crat). + rewrite -sumr_const; apply: eq_bigr => x Z1x; have [_ Zx] := setD1P Z1x. + by rewrite cfun1E cfResE ?Zx // rmorph1 mulr1; apply: kerZphi. +move=> kerZphi; pose alpha := 'omega_l['K_i1]; pose phi1 := phi 1%g. +have tiZG: {in Z^#, forall y, 'C_G[y] \subset L}. + move=> y /setD1P[nty /(subsetP sZP)Py]. + apply/subsetP=> u /setIP[Gu /cent1P cuy]. + by rewrite -(memJ_P1 y) // /conjg -?cuy ?mulKg !inE nty. +have Dalpha s: ~~ dC s Z^# -> alpha = 'omega_l['K_s]. + case/exists_inP=> x /= /gring_mode_class_sum_eq-> Z1x. + have Ci1z: z \in C i1 by rewrite CE class_refl. + rewrite [alpha](gring_mode_class_sum_eq _ Ci1z) -/phi (kerZphi z x) //. + have{tiZG} tiZG: {in Z^#, forall y, 'C_G[y] = 'C_L[y]}. + by move=> y /tiZG/setIidPr; rewrite setIA (setIidPl sLG). + by rewrite -!index_cent1 -!divgS ?subsetIl //= !tiZG ?(prZL z x). +have Ci01: 1%g \in C i0 by rewrite CE class_refl. +have rCi10: repr (C i0) = 1%g by rewrite CE class1G repr_set1. +have Dalpha2 i j: ~~ dC i Z^# -> ~~ dC j Z^# -> + (phi1 * alpha ^+ 2 == phi1 * ((a i j i0)%:R + alpha *+ a2 i j) %[mod #|P|])%A. +- move=> Z1i Z1j. + have ->: phi1 * alpha ^+ 2 = \sum_s (phi1 *+ a i j s) * 'omega_l['K_s]. + rewrite expr2 {1}(Dalpha i Z1i) (Dalpha j Z1j). + rewrite -gring_irr_modeM ?gring_class_sum_central //. + rewrite gring_classM_expansion raddf_sum mulr_sumr; apply: eq_bigr => s _. + by rewrite scaler_nat raddfMn mulrnAl mulrnAr. + rewrite (bigID (fun s => dC s Z^#)) (bigD1 i0) //=; last first. + by rewrite [dC _ _]disjoints_subset CE class1G sub1set !inE eqxx. + rewrite (gring_mode_class_sum_eq _ Ci01) mulfK ?irr1_neq0 //. + rewrite class1G cards1 mulr1 mulrDr mulr_natr -addrA eqAmodDl. + rewrite /eqAmod -addrA rpredD //; last first. + rewrite -mulr_natr natr_sum !mulr_sumr -sumrB rpred_sum // => s Z1s. + by rewrite -Dalpha // mulr_natr mulrnAl mulrnAr subrr rpred0. + apply: rpred_sum => // s /andP[Z1'Cs ntCs]; rewrite mulrnAl mulrC. + have /imsetP[x _ defCs] := enum_valP s. + have Cs_x: x \in C s by rewrite /C defCs class_refl. + rewrite (gring_mode_class_sum_eq _ Cs_x) divfK ?irr1_neq0 // -defCs -/(C s). + rewrite -mulrnAl -mulrnA mulnC -[_%:R]subr0 mulrBl. + apply: eqAmodMr; first exact: Aint_irr. + rewrite eqAmod0_rat ?rpred_nat // dvdC_nat PdvKa //. + rewrite -(setD1K (group1 Z)) [dC _ _]disjoint_sym disjoints_subset. + rewrite subUset sub1set inE -disjoints_subset disjoint_sym. + rewrite (contra _ ntCs) // [C s]defCs => /class_transr. + by rewrite -(inj_eq enum_val_inj) defCs -/(C _) CE => ->. +have zG'z1: (z^-1 \notin z ^: G)%g. + have genL2 y: y \in L -> <[y]> = <[y ^+ 2]>. + move=> Ly; apply/eqP; rewrite [_ == _]generator_coprime. + by rewrite coprime_sym prime_coprime // dvdn2 (oddSg _ oddL) ?cycle_subG. + apply: contra (ntz) => /imsetP[y Gy zy]. + have cz_y2: (y ^+ 2 \in 'C[z])%g. + by rewrite !inE conjg_set1 conjgM -zy conjVg -zy invgK. + rewrite -cycle_eq1 genL2 // cycle_eq1 -eq_invg_mul zy (sameP eqP conjg_fixP). + rewrite (sameP commgP cent1P) cent1C -cycle_subG genL2 ?cycle_subG //. + by rewrite -(memJ_P1 z) -?zy ?in_setD ?groupV ?inE ?ntz. +have a110: a i1 i1 i0 = 0%N. + apply: contraNeq zG'z1 => /existsP[[u v] /setIdP[/andP[/=]]]. + rewrite rCi10 -!/(C _) !CE -eq_invg_mul => /imsetP[x Gx ->] /class_transr <-. + by move/eqP <-; rewrite -conjVg classGidl ?class_refl. +have a120: a i1 i2 i0 = #|C i1|. + rewrite -(card_imset _ (@can_inj _ _ (fun y => (y, y^-1)%g) (@fst _ _) _)) //. + apply/eq_card=> [[u v]]; rewrite !inE rCi10 -eq_invg_mul -!/(C _) !CE -andbA. + apply/and3P/imsetP=> /= [[zGu _ /eqP<-] | [y zGy [-> ->]]]; first by exists u. + by rewrite classVg inE invgK. +have Z1i1: ~~ dC i1 Z^#. + by apply/exists_inP; exists z; rewrite //= CE class_refl. +have Z1i2: ~~ dC i2 Z^#. + apply/exists_inP; exists z^-1%g; first by rewrite /= CE class_refl. + by rewrite /= in_setD !groupV !inE ntz. +have{Dalpha2}: (phi1 * (alpha *+ a2 i1 i1) + == phi1 * (#|C i1|%:R + alpha *+ a2 i1 i2) %[mod #|P|])%A. +- rewrite -a120; apply: eqAmod_trans (Dalpha2 i1 i2 Z1i1 Z1i2). + by have:= Dalpha2 _ _ Z1i1 Z1i1; rewrite a110 add0r eqAmod_sym. +rewrite mulrDr !mulrnAr mulr1 -/phi1. +have ->: phi1 * alpha = phi z *+ #|C i1|. + have Ci1z: z \in C i1 by rewrite CE class_refl. + rewrite [alpha](gring_mode_class_sum_eq _ Ci1z) mulrC divfK ?irr1_neq0 //. + by rewrite mulr_natl CE. +rewrite -!mulrnA !(mulnC #|C _|) !mulrnA -mulrnDl. +have [|r _ /dvdnP[q Dqr]] := @Bezoutl #|C i1| #|P|. + by rewrite CE -index_cent1. +have Zq: q%:R \in Aint by apply: rpred_nat. +move/(eqAmodMr Zq); rewrite ![_ *+ #|C _| * _]mulrnAl -!mulrnAr -mulrnA -Dqr. +have /eqnP->: coprime #|C i1| #|P|. + rewrite (p'nat_coprime _ pP) // (pnat_dvd _ p'PiG) // CE -index_cent1. + by rewrite indexgS // subsetI sPG sub_cent1 (subsetP cPZ). +rewrite add1n !mulrS !mulrDr !mulr1 natrM !mulrA. +set u := _ * r%:R; set v := _ * r%:R; rewrite -[u](subrK v) mulrDl addrA. +rewrite eqAmodDr; apply: eqAmod_trans; rewrite eqAmod_sym addrC. +rewrite eqAmod_addl_mul // -mulrBl mulr_natr. +by rewrite !(rpredB, rpredD, rpredMn, Aint_irr). +Qed. + +(* This is Peterfalvi, Theorem (6.8). *) +(* We omit the semi-direct structure of L in assumption (a), since it is *) +(* implied by our statement of assumption (c). *) +Theorem Sibley_coherence (L H W1 : {group gT}) : + (*a*) [/\ odd #|L|, nilpotent H & normedTI H^# G L] -> + (*b*) let calS := seqIndD H L H 1 in let tau := 'Ind[G, L] in + (*c*) [\/ (*c1*) [Frobenius L = H ><| W1] + | (*c2*) exists2 W2 : {group gT}, prime #|W2| /\ W2 \subset H^`(1)%G + & exists A0, exists W : {group gT}, exists defW : W1 \x W2 = W, + prime_Dade_hypothesis G L H H H^# A0 defW] -> + coherent calS L^# tau. +Proof. +set A := H^# => [][oddL nilH tiA] S tau structL. +set case_c1 := [Frobenius L = H ><| W1] in structL. +have sLG: L \subset G by have [_ _ /eqP <-] := and3P tiA; apply: subsetIl. +have [defL ntH ntW1]: [/\ H ><| W1 = L, H :!=: 1 & W1 :!=: 1]%g. + have [/Frobenius_context[]// | [W2 _ [A0 [W [defW []]]]]] := structL. + by move=> _ [[-> -> _ _] [ntW2 /subG1_contra->]]. +have [nsHL _ /mulG_sub[sHL sW1L] _ _] := sdprod_context defL. +have [uccS nrS]: cfConjC_subset S S /\ ~~ has cfReal S. + by do 2?split; rewrite ?seqInd_uniq ?seqInd_notReal //; apply: cfAut_seqInd. +have defZS: 'Z[S, L^#] =i 'Z[S, A] by apply: zcharD1_seqInd. +have c1_irr: case_c1 -> {subset S <= irr L}. + move/FrobeniusWker=> frobL _ /seqIndC1P[i nz_i ->]. + exact: irr_induced_Frobenius_ker. +move defW2: 'C_H(W1)%G => W2; move defW: (W1 <*> W2)%G => W. +have{defW} defW: W1 \x W2 = W. + rewrite -defW dprodEY // -defW2 ?subsetIr // setICA setIA. + by have [_ _ _ ->] := sdprodP defL; rewrite setI1g. +pose V := cyclicTIset defW; pose A0 := A :|: class_support V L. +pose c2hyp := prime_Dade_hypothesis G L H H A A0 defW. +have c1W2: case_c1 -> W2 = 1%G by move/Frobenius_trivg_cent/group_inj <-. +have{structL} c2W2: ~~ case_c1 -> [/\ prime #|W2|, W2 \subset H^`(1)%G & c2hyp]. + case: structL => [-> // | [W20 [prW20 sW20H'] W20hyp] _]. + have{W20hyp} [A00 [W0 [defW0 W20hyp]]] := W20hyp. + suffices /group_inj defW20: W2 :=: W20. + have eqW0: W0 = W by apply: group_inj; rewrite -defW0 -defW20. + rewrite -defW20 eqW0 in prW20 sW20H' defW0 W20hyp; split=> //. + rewrite /c2hyp (eq_irrelevance defW defW0) /A0. + by have [_ _ <-] := prDade_def W20hyp. + have [[_ _ _ cycW1] [_ _ _ prW120] _] := prDade_prTI W20hyp. + have [x defW1] := cyclicP cycW1; rewrite -defW2 /= defW1 cent_cycle prW120 //. + by rewrite !inE defW1 cycle_id -cycle_eq1 -defW1 ntW1. +have{c2W2} [prW2 sW2H' c2W2] := all_and3 c2W2. +have{sW2H'} sW2H': W2 \subset H^`(1)%G. + by have [/c1W2-> | /sW2H'//] := boolP case_c1; apply: sub1G. +pose sigma := cyclicTIiso (c2W2 _). +have [R scohS oRW]: exists2 R, subcoherent S tau R & forall c2 : ~~ case_c1, + {in [predI S & irr L] & irr W, forall phi w, orthogonal (R phi) (sigma c2 w)}. +- have sAG: A \subset G^# by rewrite setSD // (subset_trans (normal_sub nsHL)). + have Itau: {in 'Z[S, L^#], isometry tau, to 'Z[irr G, G^#]}. + split=> [xi1 xi2 | xi]. + rewrite !defZS => /zchar_on Axi1 /zchar_on Axi2. + exact: normedTI_isometry Axi1 Axi2. + rewrite !zcharD1E => /andP[Zxi /eqP xi1_0]. + rewrite cfInd1 // xi1_0 mulr0 eqxx cfInd_vchar //. + by apply: zchar_trans_on Zxi; apply: seqInd_vcharW. + have [/= Hc1 | Hc2] := boolP (idfun case_c1). + suffices [R]: {R | subcoherent S tau R} by exists R => // /negP[]. + by apply: irr_subcoherent => //; first by case: uccS (c1_irr Hc1). + have ddA0 := c2W2 Hc2. + have [R [subcohR oRW _]] := prDade_subcoherent ddA0 uccS nrS. + exists R => [|not_c1 phi w irrSphi irr_w]; last first. + by rewrite /sigma -(cycTIiso_irrel ddA0) oRW. + set tau0 := Dade _ in subcohR. + have Dtau: {in 'CF(L, A), tau =1 tau0}. + have nAL: L \subset 'N(A) by rewrite normD1 normal_norm. + move=> phi Aphi; rewrite /tau0 -(restr_DadeE ddA0 (subsetUl _ _) nAL) //. + by rewrite /restr_Dade Dade_Ind. + have [Sok _ oSS Rok oRR] := subcohR; split=> // phi Sphi. + have [ZR oNR <-] := Rok _ Sphi; split=> //. + by rewrite Dtau ?irr_vchar_on ?sub_conjC_vchar ?(seqInd_vchar nsHL Sphi). +have [nsH'H nsH'L] := (der_normal 1 H, char_normal_trans (der_char 1 H) nsHL). +have [nH'L solH] := (normal_norm nsH'L, nilpotent_sol nilH). +have ltH'H: H^`(1)%g \proper H by rewrite ?(nil_comm_properl nilH) ?subsetIidl. +have coHW1: coprime #|H| #|W1|. + have [/Frobenius_coprime// | /c2W2[_ [[_ _]]]] := boolP case_c1. + by rewrite (coprime_sdprod_Hall_r defL). +have oW1: #|W1| = #|L : H| by rewrite -divgS // -(sdprod_card defL) mulKn. +have frobL1: [Frobenius L / H^`(1) = (H / H^`(1)) ><| (W1 / H^`(1))]%g. + apply: (Frobenius_coprime_quotient defL nsH'L) => //; split=> // x W1x. + have [/Frobenius_reg_ker-> //|] := boolP case_c1; first exact: sub1G. + by case/c2W2=> _ [_ [_ _ _ ->]]. +have odd_frobL1: odd_Frobenius_quotient H L 1. + have ? := FrobeniusWker frobL1. + by split=> //=; rewrite ?joingG1 // normal1 sub1G quotient_nil. +without loss [/p_groupP[p p_pr pH] not_cHH]: / p_group H /\ ~~ abelian H. + have [//| [_] [p []]] := non_coherent_chief nsHL solH scohS odd_frobL1. + rewrite (isog_abelian (quotient1_isog H)) -(isog_pgroup p (quotient1_isog H)). + by move=> /pgroup_p-> -> _; apply. +have sylH: p.-Sylow(G) H. (* required for (6.7) *) + have sylH: p.-Sylow(L) H. + apply/and3P; split=> //; rewrite -oW1 p'natE // -prime_coprime //. + by case/pgroup_pdiv: pH coHW1 => // _ _ [m ->]; rewrite coprime_pexpl. + have [P sylP sHP] := Sylow_superset (subset_trans sHL sLG) pH. + have [sPG pP _] := and3P sylP; have nilP := pgroup_nil pP. + rewrite -(nilpotent_sub_norm nilP sHP) // (sub_normal_Hall sylH) //. + exact: pgroupS (subsetIl P _) pP. + by have [_ _ /eqP <-] := and3P tiA; rewrite normD1 setSI. +pose caseA := 'Z(H) :&: W2 == [1]. +have caseB_P: ~~ caseA -> [/\ ~~ case_c1, W2 :!=: [1] & W2 \subset 'Z(H)]. + rewrite /caseA; have [-> |] := eqsVneq W2 [1]; first by rewrite setIg1 eqxx. + have [/c1W2->/eqP[]// | /prW2 pW2 ->] := boolP case_c1. + by rewrite setIC => /prime_meetG->. +pose Z := if caseA then ('Z(H) :&: H^`(1))%G else W2. +have /subsetIP[sZZ sZH']: Z \subset 'Z(H) :&: H^`(1)%g. + by rewrite /Z; case: ifPn => // /caseB_P[/c2W2[]] *; apply/subsetIP. +have caseB_cZL: ~~ caseA -> Z \subset 'Z(L). + move=> inB; have [_ _ /subsetIP[sW2H cW2H]] := caseB_P inB. + have [_ mulHW1 _ _] := sdprodP defL. + rewrite /Z (negPf inB) subsetI (subset_trans sW2H) //. + by rewrite -mulHW1 centM subsetI cW2H -defW2 subsetIr. +have nsZL: Z <| L. + have [inA | /caseB_cZL/sub_center_normal//] := boolP caseA. + by rewrite /Z inA (char_normal_trans _ nsHL) // charI ?gFchar. +have ntZ: Z :!=: 1%g. + rewrite /Z; case: ifPn => [_ | /caseB_P[]//]. + by rewrite /= setIC meet_center_nil // (sameP eqP derG1P). +have nsZH := sub_center_normal sZZ; have [sZH nZH] := andP nsZH. +have regZL: {in Z^# &, forall x y, #|'C_L[x]| = #|'C_L[y]| }. + have [inA | /caseB_cZL cZL] := boolP caseA; last first. + suffices defC x: x \in Z^# -> 'C_L[x] = L by move=> x y /defC-> /defC->. + by case/setD1P=> _ /(subsetP cZL)/setIP[_]; rewrite -sub_cent1 => /setIidPl. + suffices defC x: x \in Z^# -> 'C_L[x] = H by move=> x y /defC-> /defC->. + case/setD1P=> ntx Zx; have /setIP[Hx cHx] := subsetP sZZ x Zx. + have [_ <- _ _] := sdprodP defL; rewrite -group_modl ?sub_cent1 //=. + suffices ->: 'C_W1[x] = 1%g by rewrite mulg1. + have [/Frobenius_reg_compl-> // | in_c2] := boolP case_c1; first exact/setD1P. + have [_ [_ [_ _ _ regW1] _] _ _] := c2W2 in_c2. + apply: contraNeq ntx => /trivgPn[y /setIP[W1y cxy] nty]. + rewrite -in_set1 -set1gE -((_ =P [1]) inA) -(regW1 y) 2!inE ?nty //. + by rewrite inE cent1C cHx Hx. +have Zconst_modH := + constant_irr_mod_TI_Sylow sylH oddL tiA (And3 nsZL ntZ sZZ) regZL. +pose X := seqIndD H L Z 1; pose Y := seqIndD H L H H^`(1). +have ccsXS: cfConjC_subset X S by apply: seqInd_conjC_subset1. +have ccsYS: cfConjC_subset Y S by apply: seqInd_conjC_subset1. +have [[uX sXS ccX] [uY sYS ccY]] := (ccsXS, ccsYS). +have X'Y: {subset Y <= [predC X]}. + move=> _ /seqIndP[i /setIdP[_ kH'i] ->]; rewrite inE in kH'i. + by rewrite !inE mem_seqInd ?normal1 // !inE sub1G (subset_trans sZH'). +have irrY: {subset Y <= irr L}. + move=> _ /seqIndP[i /setIdP[not_kHi kH'i] ->]; rewrite !inE in not_kHi kH'i. + have kH'iInd: H^`(1)%g \subset cfker ('Ind[L] 'chi_i). + by rewrite sub_cfker_Ind_irr ?normal_norm. + rewrite -(cfQuoK nsH'L kH'iInd) -cfIndQuo // -quo_IirrE //. + set i1 := quo_Iirr _ i; have /irrP[k ->]: 'Ind 'chi_i1 \in irr (L / H^`(1)). + apply: irr_induced_Frobenius_ker; first exact: FrobeniusWker frobL1. + apply: contraNneq not_kHi; rewrite -(quo_IirrK nsH'H kH'i) -/i1 => ->. + by rewrite mod_IirrE // irr0 cfMod_cfun1 ?cfker_cfun1. + by rewrite -mod_IirrE ?mem_irr. +have uniY: {in Y, forall phi : 'CF(L), phi 1%g = #|W1|%:R}. + move=> _ /seqIndP[i /setIdP[_ kH'i] ->]; rewrite inE -lin_irr_der1 in kH'i. + rewrite cfInd1 // -divgS // -(sdprod_card defL) mulKn //. + by case/andP: kH'i => _ /eqP->; rewrite mulr1. +have scohY: subcoherent Y tau R by apply: (subset_subcoherent scohS). +have [tau1 cohY]: coherent Y L^# tau. + apply/(uniform_degree_coherence scohY)/(@all_pred1_constant _ #|W1|%:R). + by apply/allP=> _ /mapP[phi Yphi ->]; rewrite /= uniY. +have [[Itau1 Ztau1] Dtau1] := cohY. +have [eta1 Yeta1]: exists eta1, eta1 \in Y. + pose IY := Iirr_kerD H H H^`(1)%G. + have [IY0 | [j IYj]] := set_0Vmem IY; last first. + by exists ('Ind 'chi_j); apply/seqIndP; exists j. + have /idPn[]: \sum_(j in IY) ('chi_j 1%g) ^+ 2 == 0 by rewrite IY0 big_set0. + rewrite sum_Iirr_kerD_square ?der_sub // indexgg mul1r subr_eq0. + by rewrite pnatr_eq1 indexg_eq1 proper_subn. +have caseA_coh12: caseA -> coherent (X ++ Y) L^# tau. + move=> haveA. + have scohX: subcoherent X tau R by apply: subset_subcoherent ccsXS. + have irrX: {subset X <= irr L}. + have [/c1_irr irrS | in_c2] := boolP case_c1. + move=> phi Xphi; apply: irrS; apply: seqIndS phi Xphi. + by rewrite Iirr_kerDS // (subset_trans sZH') ?der_sub. + move/(_ in_c2) in prW2; have [_ ptiL _ _] := c2W2 in_c2. + have [[_ _ _ cycW1] [ntW2 sW2H cycW2 prW1H] oddW] := ptiL. + have nZL := normal_norm nsZL; have nZW1 := subset_trans sW1L nZL. + have isoW2: (W2 / Z)%g \isog W2. + apply/isog_symr/quotient_isog; first exact: subset_trans sW2H nZH. + by rewrite -(setIidPr sZZ) setIAC ((_ =P [1]) haveA) setI1g. + have [|defWb ptiLZ] := primeTIhyp_quotient ptiL _ sZH nsZL. + by rewrite (isog_eq1 isoW2). + pose Ichi := primeTI_Ires ptiL; pose IchiZ := primeTI_Ires ptiLZ. + have eq_Ichi: codom (fun j1 => mod_Iirr (IchiZ j1)) =i codom Ichi. + apply/subset_cardP. + rewrite !card_codom; last first; try exact: prTIres_inj. + apply: inj_comp (prTIres_inj ptiLZ). + exact: can_inj (mod_IirrK (sub_center_normal sZZ)). + by rewrite !card_ord !NirrE (nclasses_isog isoW2). + apply/subsetP=> _ /codomP[/= j1 ->]. + have [[j2 /irr_inj->] | ] := prTIres_irr_cases ptiL (mod_Iirr (IchiZ j1)). + exact: codom_f. + case=> /idPn[]; rewrite mod_IirrE // cfIndMod // cfInd_prTIres. + apply: contra (prTIred_not_irr ptiLZ j1) => /irrP[ell Dell]. + by rewrite -[_ j1]cfModK // Dell -quo_IirrE ?mem_irr // -Dell cfker_mod. + move=> _ /seqIndP[k /setDP[_ kZ'k] ->]. + have [[j /irr_inj Dk] | [] //] := prTIres_irr_cases ptiL k. + case/negP: kZ'k; have: k \in codom Ichi by rewrite Dk codom_f. + by rewrite -eq_Ichi => /codomP[j1 ->]; rewrite !inE mod_IirrE ?cfker_mod. + have [//|] := seqIndD_irr_coherence nsHL solH scohS odd_frobL1 _ irrX. + rewrite -/X => defX [tau2 cohX]; have [[Itau2 Ztau2] Dtau2] := cohX. + have [xi1 Xxi1 Nd]: + exists2 xi1, xi1 \in X & forall xi, xi \in X -> (xi1 1%g %| xi 1%g)%C. + - pose IX := Iirr_kerD H Z 1%G; have [i0 IXi0]: exists i0, i0 \in IX. + apply/set0Pn; apply: contraNneq ntZ => IX_0. + have: \sum_(i in IX) ('chi_i 1%g) ^+ 2 == 0 by rewrite IX_0 big_set0. + rewrite sum_Iirr_kerD_square ?normal1 ?sub1G // indexg1 mulf_eq0. + by rewrite (negPf (neq0CiG H Z)) subr_eq0 trivg_card1 -eqC_nat. + have:= erefl [arg min_(i < i0 in IX) truncC ('chi_i 1%g)]. + have [//|{i0 IXi0} i1 IXi1 min_i1 _] := arg_minP. + exists ('Ind 'chi_i1); first by apply/seqIndP; exists i1. + move=> _ /seqIndP[i /min_i1 le_i1_i] ->; rewrite !cfInd1 //. + have pHP := p_natP (pnat_dvd _ pH). + move: (dvd_irr1_cardG i1) (dvd_irr1_cardG i) le_i1_i. + rewrite !irr1_degree -!natrM !dvdC_nat => /pHP[m1 ->] /pHP[m ->]. + rewrite !natCK leq_exp2l ?prime_gt1 // => /subnKC <-. + by rewrite expnD mulnA dvdn_mulr. + pose d (xi : 'CF(L)) : algC := (truncC (xi 1%g / xi1 1%g))%:R. + have{Nd} def_d xi: xi \in X -> xi 1%g = d xi * xi1 1%g. + rewrite /d => Xxi; move: Xxi (Nd _ Xxi) => /irrX/irrP[i ->]. + have /irrX/irrP[i1 ->] := Xxi1. + rewrite !irr1_degree dvdC_nat => /dvdnP[q ->]. + by rewrite natrM -irr1_degree mulfK ?irr1_neq0 // natCK. + have d_xi1: d xi1 = 1. + by apply: (mulIf (seqInd1_neq0 nsHL Xxi1)); rewrite mul1r -def_d. + have oXY: orthogonal X Y. + apply/orthogonalP=> xi eta Xxi Yeta; apply: orthoPr xi Xxi. + exact: (subset_ortho_subcoherent scohS sXS (sYS _ Yeta) (X'Y _ Yeta)). + have [_ [Itau Ztau] /orthogonal_free freeS _ _] := scohS. + have o_tauXY: orthogonal (map tau2 X) (map tau1 Y). + exact: (coherent_ortho scohS). + have [a Na Dxi11]: exists2 a, a \in Cnat & xi1 1%g = a * #|W1|%:R. + have [i1 _ ->] := seqIndP Xxi1. + exists ('chi_i1 1%g); first exact: Cnat_irr1. + by rewrite cfInd1 // -divgS // -(sdprod_card defL) ?mulKn // mulrC. + pose psi1 := xi1 - a *: eta1; have Za: a \in Cint by rewrite CintE Na. + have Zpsi1: psi1 \in 'Z[S, L^#]. + rewrite zcharD1E !cfunE (uniY _ Yeta1) -Dxi11 subrr eqxx. + by rewrite rpredB ?scale_zchar ?mem_zchar ?(sXS _ Xxi1) // sYS. + have [Y1 dY1 [X1 [dX1 _ oX1tauY]]] := orthogonal_split (map tau1 Y)(tau psi1). + have oY: orthonormal Y by apply: sub_orthonormal (irr_orthonormal L). + have oYtau: orthonormal (map tau1 Y) by apply: map_orthonormal. + have{dX1 Y1 dY1} [b Zb Dpsi1]: {b | b \in Cint & + tau psi1 = X1 - a *: tau1 eta1 + b *: (\sum_(eta <- Y) tau1 eta)}. + - exists ('[tau psi1, tau1 eta1] + a). + rewrite rpredD // Cint_cfdot_vchar ?Ztau1 ?seqInd_zcharW //. + exact: zcharW (Ztau _ Zpsi1). + rewrite {1}dX1 addrC -addrA; congr (_ + _). + have [_ -> ->] := orthonormal_span oYtau dY1. + rewrite -[Y1](addrK X1) -dX1 big_map !(big_rem eta1 Yeta1) /=. + rewrite cfdotBl (orthoPl oX1tauY) ?map_f // subr0. + rewrite scalerDr addrA; congr (_ + _). + by rewrite addrC -scaleNr -scalerDl addrK. + rewrite raddf_sum; apply: eq_big_seq => eta. + rewrite mem_rem_uniq ?seqInd_uniq // => /andP[eta1'eta /= Yeta]. + congr (_ *: _); rewrite cfdotBl (orthoPl oX1tauY) ?map_f // subr0 addrC. + apply: canRL (subrK _) _; rewrite -2!raddfB /=. + have Zeta: (eta - eta1) \in 'Z[Y, L^#]. + by rewrite zcharD1E rpredB ?seqInd_zcharW //= !cfunE !uniY ?subrr. + rewrite Dtau1 // Itau // ?(zchar_subset sYS) //. + rewrite cfdotBl cfdotZl !cfdotBr 2?(orthogonalP oXY) // subr0 add0r. + have [_ oYY] := orthonormalP oY; rewrite !oYY // eqxx. + by rewrite eq_sym (negPf eta1'eta) add0r mulrN mulr1 opprK. + pose psi := 'Res[L] (tau1 eta1). + have [X2 dX2 [xi' [dxi' _ oxi'X]]] := orthogonal_split X psi. + have oX: orthonormal X by apply: sub_orthonormal (irr_orthonormal L). + have Zpsi: psi \in 'Z[irr L] by rewrite cfRes_vchar ?Ztau1 ?seqInd_zcharW. + pose sumXd := \sum_(xi <- X) d xi *: xi. + have Zxi1Xd xi: xi \in X -> xi - d xi *: xi1 \in 'Z[X, L^#]. + move=> Xxi; rewrite zcharD1E !cfunE -def_d // subrr eqxx. + by rewrite rpredB ?scale_zchar ?seqInd_zcharW ?rpred_nat. + have{dxi' X2 dX2} [c Zc Dpsi]: {c | c \in Cint & psi = c *: sumXd + xi'}. + exists '[psi, xi1]; first by rewrite Cint_cfdot_vchar ?(seqInd_vcharW Xxi1). + rewrite {1}dxi'; congr (_ + _); have [_ -> ->] := orthonormal_span oX dX2. + rewrite -[X2](addrK xi') -dxi' raddf_sum; apply: eq_big_seq => /= xi Xxi. + rewrite cfdotBl (orthoPl oxi'X) // subr0 scalerA; congr (_ *: _). + apply/eqP; rewrite -subr_eq0 mulrC -[d xi]conj_Cnat ?Cnat_nat //. + rewrite -cfdotZr -cfdotBr cfdot_Res_l -Dtau2 ?Zxi1Xd //. + rewrite cfdotC raddfB raddfZ_Cnat ?Cnat_nat // cfdotBl cfdotZl. + by rewrite !(orthogonalP o_tauXY) ?map_f // mulr0 subr0 conjC0. + have Exi' z: z \in Z -> xi' z = xi' 1%g. + move=> Zz; rewrite [xi']cfun_sum_cfdot !sum_cfunE; apply: eq_bigr => ell _. + have [Xell |] := boolP ('chi_ell \in X). + by rewrite !cfunE (orthoPl oxi'X) ?mul0r. + by rewrite !cfunE defX inE /= mem_irr negbK => /subsetP/(_ z Zz)/cfker1->. + have Eba: '[psi, psi1] = b - a. + rewrite cfdot_Res_l -/tau Dpsi1 -addrA !cfdotDr cfdotNr cfdotZr. + rewrite cfdotC (orthoPl oX1tauY) ?map_f // conjC0 add0r addrC. + rewrite cfdotC raddf_sum cfproj_sum_orthonormal // !aut_Cint //. + by have [_ ->] := orthonormalP oYtau; rewrite ?map_f // eqxx mulr1. + have nz_xi11: xi1 1%g != 0 by have /irrX/irrP[i ->] := Xxi1; apply: irr1_neq0. + have {Eba} Ebc: (a %| b - c)%C. + rewrite -[b](subrK a) -Eba cfdotBr {1}Dpsi cfdotDl cfdotZl. + rewrite cfproj_sum_orthonormal // (orthoPl oxi'X) // addr0 d_xi1 mulr1. + rewrite addrC -addrA addKr addrC rpredB ?dvdC_refl //= cfdotZr aut_Cint //. + by rewrite dvdC_mulr // Cint_cfdot_vchar ?(seqInd_vcharW Yeta1). + have DsumXd: sumXd = (xi1 1%g)^-1 *: (cfReg L - (cfReg (L / Z)%g %% Z)%CF). + apply: canRL (scalerK nz_xi11) _; rewrite !cfReg_sum 2!linear_sum /=. + pose F (xi : 'CF(L)) := xi 1%g *: xi; transitivity (\sum_(xi <- X) F xi). + by apply: eq_big_seq => xi Xxi; rewrite scalerA mulrC -def_d. + rewrite (bigID (mem (Iirr_ker L Z))) /=; apply: canRL (addrK _) _. + rewrite addrC; congr (_ + _). + rewrite (eq_bigl _ _ (in_set _)) (reindex _ (mod_Iirr_bij nsZL)) /=. + apply: eq_big => [i | i _]; first by rewrite mod_IirrE ?cfker_mod. + by rewrite linearZ mod_IirrE // cfMod1. + transitivity (\sum_(xi <- [seq 'chi_i | i in [predC Iirr_ker L Z]]) F xi). + apply: eq_big_perm; apply: uniq_perm_eq => // [|xi]. + by rewrite (map_inj_uniq irr_inj) ?enum_uniq. + rewrite defX inE /=; apply/andP/imageP=> [[/irrP[i ->] kerZ'i] | [i]]. + by exists i; rewrite ?inE. + by rewrite !inE => ? ->; rewrite mem_irr. + by rewrite big_map big_filter; apply: eq_bigl => i; rewrite !inE. + have eta1tauZ z: z \in Z^# -> tau1 eta1 z - tau1 eta1 1%g = - c * #|H|%:R / a. + case/setD1P=> ntz Zz; transitivity (psi z - psi 1%g). + by rewrite !cfResE ?(subsetP (normal_sub nsZL)). + rewrite Dpsi DsumXd !cfunE Exi' ?cfuniE ?normal1 // set11 inE (negPf ntz). + rewrite mulr0 mulr1 sub0r Dxi11 cfker1 ?cfker_reg_quo //. + set cc := c * _ + _; rewrite 2!mulrDr -[rhs in _ - rhs]addrA -/cc. + rewrite addrC opprD {cc}subrK -(sdprod_card defL) mulnC natrM. + by rewrite invfM !mulrA divfK ?neq0CG // mulrAC -2!mulNr. + have{eta1tauZ} dvHpsi: (#|H| %| - c * #|H|%:R / a)%C. + have /dirrP[e [i Deta1]]: tau1 eta1 \in dirr G. + rewrite dirrE Ztau1 ?Itau1 ?seqInd_zcharW //=. + by have [_ ->] := orthonormalP oY; rewrite ?eqxx. + have [z ntz Zz] := trivgPn _ ntZ; have Z1z: z \in Z^# by apply/setD1P. + have /(Zconst_modH i)[|_] := Z1z. + move=> z1 z2 Zz1 Zz2; rewrite -(canLR (signrZK e) Deta1) !cfunE. + by apply/eqP; do 2!rewrite eq_sym (canRL (subrK _) (eta1tauZ _ _)) //. + by rewrite -(canLR (signrZK e) Deta1) !cfunE -mulrBr eta1tauZ // rpredMsign. + have nz_a: a != 0 by rewrite Dxi11 mulf_eq0 negb_or neq0CG andbT in nz_xi11. + have{dvHpsi} dv_ac: (a %| c)%C. + move: dvHpsi; rewrite !mulNr mulrAC rpredN => /dvdCP[q Zq]. + by move/(mulIf (neq0CG H))/(canRL (divfK nz_a))->; apply: dvdC_mull. + have{Ebc dv_ac} /dvdCP[q Zq Db]: (a %| b)%C by rewrite -[b](subrK c) rpredD. + pose m : algC := (size Y)%:R. + have Da1: 1 + a ^+ 2 = '[X1] + a ^+ 2 * ((q - 1) ^+ 2 + (m - 1) * q ^+ 2). + transitivity '[psi1]. + rewrite cfnormBd; last by rewrite cfdotZr (orthogonalP oXY) ?mulr0. + rewrite cfnormZ Cint_normK //. + have [[_ -> //] [_ -> //]]:= (orthonormalP oX, orthonormalP oY). + by rewrite !eqxx mulr1. + rewrite -Itau // Dpsi1 -addrA cfnormDd; last first. + rewrite addrC cfdotBr !cfdotZr cfdot_sumr (orthoPl oX1tauY) ?map_f //. + rewrite big_seq big1 ?mulr0 ?subrr // => eta Yeta. + by rewrite (orthoPl oX1tauY) ?map_f //. + congr (_ + _); rewrite cfnormD cfnormN !cfnormZ. + have [_ ->] := orthonormalP oYtau; rewrite ?map_f // eqxx mulr1. + rewrite cfnorm_map_orthonormal // -/m !Cint_normK // cfdotNl cfdotZl. + rewrite linear_sum cfdotC cfproj_sum_orthonormal // rmorphN rmorphM. + rewrite conjCK !aut_Cint // -mulr2n mulNrn -[_ - _]addrAC. + rewrite mulrDr -{1}[m](addNKr 1) mulrDr mulr1 addrA -sqrrB. + congr (_ + _); last by rewrite mulrCA -exprMn (mulrC a) addrC -Db mulrC. + by rewrite -exprMn -sqrrN opprB mulrBr mulr1 (mulrC a) -Db. + have{Da1} maxq: ~~ (2%:R <= (q - 1) ^+ 2 + (m - 1) * q ^+ 2). + have a2_gt1: a ^+ 2 > 1. + have /seqIndP[i1 /setDP[_ not_kerH'i1] Dxi1] := Xxi1. + apply: contraR not_kerH'i1; rewrite inE expr_gt1 ?Cnat_ge0 //. + have [n Da] := CnatP a Na; rewrite Da ltr1n -leqNgt leq_eqVlt. + rewrite ltnNge lt0n -!eqC_nat -{n}Da nz_a orbF => /eqP a_eq1. + rewrite (subset_trans sZH') // -lin_irr_der1 qualifE irr_char. + rewrite -(inj_eq (mulfI (neq0CiG L H))) -cfInd1 // -Dxi1 Dxi11 a_eq1. + by rewrite mul1r mulr1 -divgS //= -(sdprod_card defL) mulKn. + rewrite -(ler_pmul2l (ltr_trans ltr01 a2_gt1)) ltr_geF // mulr_natr. + apply: ler_lt_trans (_ : 1 + a ^+ 2 < _); last by rewrite ltr_add2r. + by rewrite Da1 -subr_ge0 addrK cfnorm_ge0. + clear psi Dpsi Zpsi Zb c sumXd DsumXd Zc xi' Exi' oxi'X. + wlog{Dpsi1 Itau1 Ztau1 Dtau1 oYtau b q maxq Db Zq} Dpsi1: + tau1 cohY o_tauXY oX1tauY / tau psi1 = X1 - a *: tau1 eta1. + - move=> IH; have [q0 | nz_q] := eqVneq q 0. + by apply: (IH tau1) => //; rewrite Dpsi1 Db q0 mul0r scale0r addr0. + have m1_ge1: 1 <= m - 1. + rewrite -(@ler_add2r _ 1) subrK (ler_nat _ 2). + exact: seqInd_nontrivial (irrY _ Yeta1) (Yeta1). + have q1: q = 1. + apply: contraNeq maxq; rewrite -subr_eq0 => nz_q1. + rewrite ler_add // ?sqr_Cint_ge1 ?rpredB //. + rewrite (ler_trans m1_ge1) // -{1}[m - 1]mulr1. + by rewrite ler_pmul2l ?sqr_Cint_ge1 // (ltr_le_trans ltr01). + have szY2: (size Y <= 2)%N. + move: maxq; rewrite q1 subrr exprS mul0r add0r mulrA !mulr1. + by rewrite -(ler_add2r 1) subrK -mulrSr ler_nat -leqNgt. + have defY: perm_eq Y (eta1 :: eta1^*)%CF. + have uYeta: uniq (eta1 :: eta1^*)%CF. + by rewrite /= andbT inE eq_sym; have [[_ /hasPn/=->]] := scohY. + rewrite perm_eq_sym uniq_perm_eq //. + have [|//]:= leq_size_perm uYeta _ szY2. + by apply/allP; rewrite /= Yeta1 ccY. + have memYtau1c: {subset map (tau1 \o cfAut conjC) Y <= map tau1 Y}. + by move=> _ /mapP[eta Yeta ->]; rewrite /= map_f ?ccY. + apply: (IH _ (dual_coherence scohY cohY szY2)). + - rewrite (map_comp -%R) orthogonal_oppr. + by apply/orthogonalP=> phi psi ? /memYtau1c; apply: (orthogonalP o_tauXY). + - rewrite (map_comp -%R) orthogonal_oppr. + by apply/orthoPl=> psi /memYtau1c; apply: (orthoPl oX1tauY). + rewrite Dpsi1 (eq_big_perm _ defY) Db q1 /= mul1r big_cons big_seq1. + by rewrite scalerDr addrA subrK -scalerN opprK. + have [[[Itau1 Ztau1] Dtau1] [_ oXX]] := (cohY, orthonormalP oX). + have n1X1: '[X1] = 1. + apply: (addIr '[a *: tau1 eta1]); rewrite -cfnormBd; last first. + by rewrite cfdotZr (orthoPl oX1tauY) ?mulr0 ?map_f. + rewrite -Dpsi1 Itau // cfnormBd; last first. + by rewrite cfdotZr (orthogonalP oXY) ?mulr0. + by rewrite !cfnormZ Itau1 ?seqInd_zcharW // oXX ?eqxx. + without loss{Itau2 Ztau2 Dtau2} defX1: tau2 cohX o_tauXY / X1 = tau2 xi1. + move=> IH; have ZX: {subset X <= 'Z[X]} by apply: seqInd_zcharW. + have dirrXtau xi: xi \in X -> tau2 xi \in dirr G. + by move=> Xxi; rewrite dirrE Ztau2 1?Itau2 ?ZX //= oXX ?eqxx. + have dirrX1: X1 \in dirr G. + rewrite dirrE n1X1 eqxx -[X1](subrK (a *: tau1 eta1)) -Dpsi1. + rewrite rpredD ?scale_zchar ?(zcharW (Ztau _ _)) //. + by rewrite Ztau1 ?seqInd_zcharW. + have oX1_Xd xi: + xi \in X -> xi != xi1 -> '[d xi *: tau2 xi1 - tau2 xi, X1] = d xi. + - move=> Xxi xi1'xi; have ZXxi := Zxi1Xd xi Xxi. + rewrite -[X1](subrK (a *: tau1 eta1)) -Dpsi1 cfdotDr cfdotZr addrC. + rewrite cfdotBl cfdotZl 2?(orthogonalP o_tauXY) ?map_f //. + rewrite !(mulr0, subr0, conjC0) add0r -{1}raddfZ_Cnat ?Cnat_nat //. + rewrite -opprB cfdotNl -raddfB Dtau2 //. + rewrite Itau //; last exact: zchar_subset ZXxi. + rewrite cfdotBr cfdotZr addrC !cfdotBl !cfdotZl. + rewrite 2?(orthogonalP oXY) // !(mulr0, oppr0, add0r, conjC0). + by rewrite !oXX // eqxx (negPf xi1'xi) add0r opprK mulr1. + have Xxi2: xi1^*%CF \in X by apply: ccX. + have xi1'2: xi1^*%CF != xi1 by have [[_ /hasPn->]] := scohX. + have xi2tau_irr: - tau2 xi1^*%CF \in dirr G by rewrite dirr_opp dirrXtau. + have d_xi2: d xi1^*%CF = 1. + by rewrite /d cfunE conj_Cnat // (Cnat_seqInd1 Xxi1). + have [||def_X1]:= cfdot_add_dirr_eq1 (dirrXtau _ Xxi1) xi2tau_irr dirrX1. + - by rewrite -[tau2 xi1]scale1r -d_xi2 oX1_Xd. + - exact: IH. + have sX_xi12: {subset X <= xi1 :: xi1^*%CF}. + apply/allP/allPn=> [[xi3 Xxi3 /norP[xi1'3 /norP[xi2'3 _]]]]. + suffices d3_0: d xi3 = 0. + by have:= seqInd1_neq0 nsHL Xxi3; rewrite def_d // d3_0 mul0r eqxx. + rewrite -oX1_Xd // def_X1 cfdotNr cfdotBl cfdotZl !Itau2 ?ZX //. + by rewrite !oXX // (negPf xi2'3) eq_sym (negPf xi1'2) mulr0 add0r opprK. + have{sX_xi12 defX} defX: perm_eq X (xi1 :: xi1^*%CF). + have uXxi: uniq (xi1 :: xi1^*)%CF by rewrite /= andbT inE eq_sym. + rewrite perm_eq_sym uniq_perm_eq // => xi. + by apply/idP/idP; [rewrite !inE => /pred2P[]-> | apply: sX_xi12]. + have szX2: (size X <= 2)%N by rewrite (perm_eq_size defX). + apply: (IH _ (dual_coherence scohX cohX szX2)) def_X1. + rewrite (map_comp -%R) orthogonal_oppl. + apply/orthogonalP=> _ eta /mapP[xi Xxi ->]. + by apply: (orthogonalP o_tauXY); rewrite map_f ?ccX. + move: Dpsi1; rewrite -raddfZ_Cnat // defX1. + apply: (bridge_coherent scohS ccsXS cohX ccsYS cohY X'Y). + by rewrite (zchar_on Zpsi1) rpredZ_Cnat ?mem_zchar. +have{caseA_coh12} cohXY: coherent (X ++ Y) L^# tau. + have [/caseA_coh12// | caseB] := boolP caseA. + have defZ: Z = W2 by rewrite /Z (negPf caseB). + have{caseB} [case_c2 _ _] := caseB_P caseB. + move/(_ case_c2) in oRW; pose PtypeL := c2W2 case_c2. + have{prW2} pr_w2 := prW2 case_c2; set w2 := #|W2| in pr_w2. + have /cyclicP[z0 cycZ]: cyclic Z by rewrite defZ prime_cyclic. + have idYZ: {in Y & Z^#, forall (eta : 'CF(L)) x, tau1 eta x = tau1 eta z0}. + move=> eta x Yeta; rewrite !inE andbC cycZ => /andP[/cyclePmin[k]]. + rewrite orderE -cycZ defZ -/w2 => lt_k_w2 -> nt_z0k. + have k_gt0: (0 < k)%N by rewrite lt0n (contraNneq _ nt_z0k) // => ->. + have cokw2: coprime k w2 by rewrite coprime_sym prime_coprime // gtnNdvd. + have sW2G: W2 \subset G by rewrite -defW2 subIset // (subset_trans sHL). + have [u Du _]:= make_pi_cfAut G cokw2. + rewrite -Du ?Ztau1 ?seqInd_zcharW //; last by rewrite orderE -cycZ defZ. + have nAL: L \subset 'N(A) by rewrite normD1 normal_norm. + pose ddA := restr_Dade_hyp PtypeL (subsetUl _ _) nAL. + have cohY_Dade: coherent_with Y L^# (Dade ddA) tau1. + split=> // phi Yphi; rewrite Dtau1 ?Dade_Ind //. + by rewrite (@zchar_on _ _ Y) -?zcharD1_seqInd. + rewrite (cfAut_Dade_coherent cohY_Dade) ?irrY //; last first. + split; last exact: cfAut_seqInd. + exact: seqInd_nontrivial_irr (irrY _ Yeta) (Yeta). + rewrite -[cfAut u _](subrK eta) raddfD cfunE. + apply: canLR (subrK _) _; rewrite subrr. + have [_ ->] := cohY_Dade; last first. + by rewrite -opprB rpredN zcharD1_seqInd // seqInd_sub_aut_zchar. + rewrite Dade_id; last first. + by rewrite !inE -cycle_eq1 -cycle_subG -cycZ ntZ. + rewrite !cfunE cfker1 ?aut_Cnat ?subrr ?(Cnat_seqInd1 Yeta) //. + rewrite -cycle_subG -cycZ (subset_trans sZH') //. + have [j /setDP[kerH'j _] ->] := seqIndP Yeta. + by rewrite inE in kerH'j; rewrite sub_cfker_Ind_irr. + have [_ [Itau _] oSS _ _] := scohS. + have oY: orthonormal Y by apply: sub_orthonormal (irr_orthonormal L). + have oYtau: orthonormal (map tau1 Y) by apply: map_orthonormal. + have oXY: orthogonal X Y. + apply/orthogonalP=> xi eta Xxi Yeta; apply: orthoPr xi Xxi. + exact: (subset_ortho_subcoherent scohS sXS (sYS _ Yeta) (X'Y _ Yeta)). + have [Y1 Dpsi1 defY1]: exists2 Y1, + forall i : Iirr Z, i != 0 -> + exists2 X1 : 'CF(G), orthogonal X1 (map tau1 Y) + & tau ('Ind 'chi_i - #|H : Z|%:R *: eta1) = X1 - #|H : Z|%:R *: Y1 + & Y1 = tau1 eta1 \/ size Y = 2 /\ Y1 = dual_iso tau1 eta1. + - have [i0 nz_i0]: exists i0 : Iirr Z, i0 != 0. + by apply: (ex_intro _ (Sub 1%N _)) => //; rewrite NirrE classes_gt1. + pose psi1 := tau1 eta1; pose b := psi1 z0. + pose a := (psi1 1%g - b) / #|Z|%:R. + have sZL := normal_sub nsZL; have sZG := subset_trans sZL sLG. + have Dpsi1: 'Res psi1 = a *: cfReg Z + b%:A. + apply/cfun_inP=> z Zz. + rewrite cfResE // !cfunE cfun1E Zz mulr1 cfuniE ?normal1 // inE. + have [-> | ntz] := altP eqP; first by rewrite mulr1 divfK ?neq0CG ?subrK. + by rewrite !mulr0 add0r idYZ // !inE ntz. + have /dvdCP[x0 Zx0 Dx0]: (#|H : Z| %| a)%C. + have /dvdCP[x Zx Dx]: (#|H| %| b - psi1 1%g)%C. + have psi1Z z: z \in Z^# -> psi1 z = b. + case/setD1P=> ntz Zz; rewrite -(cfResE _ _ Zz) // Dpsi1 !cfunE cfun1E. + by rewrite cfuniE ?normal1 // Zz inE (negPf ntz) !mulr0 mulr1 add0r. + have /dirrP[e [i /(canLR (signrZK e)) Epsi1]]: psi1 \in dirr G. + have [_ oYt] := orthonormalP oYtau. + by rewrite dirrE oYt ?map_f // !eqxx Ztau1 ?seqInd_zcharW. + have Zz: z0 \in Z^# by rewrite !inE -cycle_eq1 -cycle_subG -cycZ ntZ /=. + have/(Zconst_modH i)[z1 Zz1 z2 Zz2 |_] := Zz. + by rewrite -Epsi1 !cfunE !psi1Z. + by rewrite -Epsi1 !cfunE -mulrBr rpredMsign psi1Z. + apply/dvdCP; exists (- x); first by rewrite rpredN. + rewrite /a -opprB Dx -(Lagrange sZH) mulnC [p in x * p]natrM -!mulNr. + by rewrite !mulrA !mulfK ?neq0CG. + pose x1 := '[eta1, 'Res psi1]; pose x := x0 + 1 - x1. + have Zx: x \in Cint. + rewrite rpredB ?rpredD // Cint_cfdot_vchar // ?(seqInd_vcharW Yeta1) //. + by rewrite cfRes_vchar // Ztau1 ?seqInd_zcharW. + pose Y1 := - \sum_(eta <- Y) (x - (eta == eta1)%:R) *: tau1 eta. + pose alpha i := 'Ind[L, Z] 'chi_i - #|H : Z|%:R *: eta1. + have IZfacts i: i != 0 -> + [/\ 'chi_i 1%g = 1, 'Ind[L, Z] 'chi_i \in 'Z[X] & alpha i \in 'Z[S, L^#]]. + - move=> nzi; have /andP[_ /eqP lin_i]: 'chi_i \is a linear_char. + by rewrite lin_irr_der1 (derG1P _) ?sub1G // cycZ cycle_abelian. + have Xchi: 'Ind 'chi_i \in 'Z[X]. + rewrite -(cfIndInd _ sHL) // ['Ind[H] _]cfun_sum_constt linear_sum. + apply: rpred_sum => k k_i; rewrite linearZ; apply: scale_zchar. + by rewrite Cint_cfdot_vchar_irr // cfInd_vchar ?irr_vchar. + rewrite seqInd_zcharW //; apply/seqIndP; exists k => //. + rewrite !inE sub1G andbT; apply: contra k_i => kerZk. + rewrite -Frobenius_reciprocity. + have ->: 'Res[Z] 'chi_k = ('chi_k 1%g)%:A. + apply: cfun_inP => z Zz; rewrite cfunE cfun1E Zz mulr1 cfResE //. + by rewrite cfker1 ?(subsetP kerZk). + by rewrite cfdotZr -irr0 cfdot_irr (negPf nzi) mulr0. + split=> //; rewrite zcharD1E !cfunE cfInd1 // uniY // lin_i mulr1. + rewrite -divgS // -(sdprod_card defL) -(Lagrange sZH) -mulnA mulKn //. + rewrite -natrM subrr rpredB //=; first by rewrite (zchar_subset sXS). + by rewrite scale_zchar ?rpred_nat // seqInd_zcharW ?sYS. + have Dalpha (i : Iirr Z) (nzi : i != 0) : + exists2 X1 : 'CF(G), orthogonal X1 (map tau1 Y) + & tau (alpha i) = X1 - #|H : Z|%:R *: Y1. + - have [lin_i Xchi Zalpha] := IZfacts i nzi. + have Da: '[tau (alpha i), psi1] = a - #|H : Z|%:R * x1. + rewrite !(=^~ Frobenius_reciprocity, cfdotBl) cfResRes // cfdotZl. + congr (_ - _); rewrite cfdotC Dpsi1 cfdotDl cfdotZl cfReg_sum. + rewrite cfdot_suml (bigD1 i) //= big1 => [|j i'j]; last first. + by rewrite cfdotZl cfdot_irr (negPf i'j) mulr0. + rewrite !cfdotZl cfnorm_irr lin_i addr0 !mulr1. + rewrite -irr0 cfdot_irr eq_sym (negPf nzi) mulr0 addr0. + by rewrite aut_Cint // Dx0 rpredM ?rpred_nat. + have [Y2 dY2 [X1 [dX1 _ oX1Yt]]] := + orthogonal_split (map tau1 Y) (tau (alpha i)). + exists X1 => //; rewrite dX1 addrC scalerN opprK scaler_sumr. + congr (_ + _); have [_ -> ->] := orthonormal_span oYtau dY2. + rewrite big_map; apply: eq_big_seq => eta Yeta. + rewrite scalerA -[Y2](addrK X1) -dX1 cfdotBl (orthoPl oX1Yt) ?map_f //. + congr (_ *: _); rewrite subr0 !mulrBr mulrDr mulrC -Dx0. + rewrite (addrAC a) -Da -addrA -mulrBr addrC; apply: canRL (subrK _) _. + have Zeta: eta - eta1 \in 'Z[Y, L^#]. + by rewrite zcharD1E rpredB ?seqInd_zcharW //= !cfunE !uniY ?subrr. + rewrite -cfdotBr -raddfB Dtau1 // Itau //; last first. + by rewrite (zchar_subset sYS) ?seqInd_free. + rewrite cfdotBl (span_orthogonal oXY) ?(zchar_span Xchi)//; last first. + by rewrite memvB ?memv_span. + have [_ oYY] := orthonormalP oY; rewrite cfdotZl cfdotBr !oYY //. + by rewrite eqxx sub0r -mulrN opprB eq_sym. + exists Y1 => //; have{Dalpha} [X1 oX1Y Dalpha] := Dalpha i0 nz_i0. + have [lin_i Xchi Zalpha] := IZfacts i0 nz_i0. + have norm_alpha: '[tau (alpha i0)] = (#|L : Z| + #|H : Z| ^ 2)%:R. + rewrite natrD Itau // cfnormBd; last first. + rewrite (span_orthogonal oXY) ?(zchar_span Xchi) //. + by rewrite memvZ ?memv_span. + rewrite cfnorm_Ind_irr //; congr (#|_ : _|%:R + _). + apply/setIidPl; apply: subset_trans (cent_sub_inertia _). + rewrite -(sdprodW defL) mulG_subG (centsS sZZ) centsC ?subsetIr //=. + by rewrite defZ -defW2 subsetIr. + have [_ oYY] := orthonormalP oY; rewrite cfnormZ oYY // eqxx mulr1. + by rewrite normCK rmorph_nat -natrM. + have{norm_alpha} ub_norm_alpha: '[tau (alpha i0)] < (#|H : Z| ^ 2).*2%:R. + rewrite norm_alpha -addnn ltr_nat ltn_add2r. + rewrite -divgS // -(sdprod_card defL) -(Lagrange sZH) -mulnA mulKn //. + rewrite ltn_pmul2l //. + have frobL2: [Frobenius L / Z = (H / Z) ><| (W1 / Z)]%g. + apply: (Frobenius_coprime_quotient defL nsZL) => //. + split=> [|y W1y]; first exact: sub_proper_trans ltH'H. + by rewrite defZ; have [/= ? [_ [_ _ _ ->]]] := PtypeL. + have nZW1 := subset_trans sW1L (normal_norm nsZL). + rewrite (card_isog (quotient_isog nZW1 _)); last first. + by rewrite coprime_TIg ?(coprimeSg sZH). + rewrite -(prednK (indexg_gt0 H Z)) ltnS -card_quotient //. + rewrite dvdn_leq ?(Frobenius_dvd_ker1 frobL2) // -subn1 subn_gt0. + by rewrite cardG_gt1; case/Frobenius_context: frobL2. + pose m : algC := (size Y)%:R. + have{ub_norm_alpha} ub_xm: ~~ (2%:R <= (x - 1) ^+ 2 + (m - 1) * x ^+ 2). + have: ~~ (2%:R <= '[Y1]). + rewrite -2!(ler_pmul2l (gt0CiG H Z)) -!natrM mulnA muln2. + rewrite ltr_geF //; apply: ler_lt_trans ub_norm_alpha. + rewrite Dalpha cfnormBd. + by rewrite cfnormZ normCK rmorph_nat mulrA -subr_ge0 addrK cfnorm_ge0. + rewrite scalerN -scaleNr cfdotZr cfdot_sumr big_seq. + rewrite big1 ?mulr0 // => eta Yeta. + by rewrite cfdotZr (orthoPl oX1Y) ?map_f ?mulr0. + rewrite cfnormN cfnorm_sum_orthonormal // (big_rem eta1) //= eqxx. + rewrite big_seq (eq_bigr (fun _ => (x ^+ 2))) => [|eta]; last first. + rewrite mem_rem_uniq // => /andP[/negPf-> _]. + by rewrite subr0 Cint_normK. + rewrite Cint_normK 1?rpredB //= -big_seq; congr (~~ (_ <= _ + _)). + rewrite big_const_seq count_predT // -Monoid.iteropE. + rewrite /m (perm_eq_size (perm_to_rem Yeta1)) /=. + by rewrite mulrSr addrK mulr_natl. + have [x_eq0 | nz_x] := eqVneq x 0. + left; rewrite /Y1 x_eq0 (big_rem eta1) //= eqxx sub0r scaleN1r. + rewrite big_seq big1 => [|eta]; last first. + by rewrite mem_rem_uniq // => /andP[/negPf-> _]; rewrite subrr scale0r. + by rewrite addr0 opprK. + have m1_ge1: 1 <= m - 1. + rewrite -(@ler_add2r _ 1) subrK (ler_nat _ 2). + exact: seqInd_nontrivial (irrY _ Yeta1) (Yeta1). + right; have x_eq1: x = 1. + apply: contraNeq ub_xm; rewrite -subr_eq0 => nz_x1; apply: ler_add. + by rewrite sqr_Cint_ge1 // rpredB. + rewrite (ler_trans m1_ge1) // -{1}[m - 1]mulr1 ler_pmul2l. + exact: sqr_Cint_ge1. + exact: ltr_le_trans ltr01 m1_ge1. + have szY2: size Y = 2. + apply: contraNeq ub_xm => Yneq2; rewrite x_eq1 /m subrr !exprS mul0r. + rewrite add0r !mul1r mulr1 -(ler_add2r 1) subrK -mulrSr ler_nat. + by rewrite ltn_neqAle eq_sym Yneq2 -leC_nat -/m -[m](subrK 1) ler_add2r. + have eta1'2: eta1^*%CF != eta1 by apply: seqInd_conjC_neq Yeta1. + have defY: perm_eq Y (eta1 :: eta1^*%CF). + have uY2: uniq (eta1 :: eta1^*%CF) by rewrite /= inE eq_sym eta1'2. + rewrite perm_eq_sym uniq_perm_eq //. + have sY2Y: {subset (eta1 :: eta1^*%CF) <= Y}. + by apply/allP; rewrite /= cfAut_seqInd ?Yeta1. + by have [|//]:= leq_size_perm uY2 sY2Y; rewrite szY2. + split=> //; congr (- _); rewrite (eq_big_perm _ defY) /= x_eq1. + rewrite big_cons big_seq1 eqxx (negPf eta1'2) subrr scale0r add0r. + by rewrite subr0 scale1r. + have [a Za Dxa]: exists2 a, forall xi, a xi \in Cint + & forall xi, xi \in X -> xi 1%g = a xi * #|W1|%:R + /\ (exists2 X1 : 'CF(G), orthogonal X1 (map tau1 Y) + & tau (xi - a xi *: eta1) = X1 - a xi *: Y1). + - pose aX (xi : 'CF(L)) : algC := (truncC (xi 1%g / #|W1|%:R))%:R. + exists aX => [xi | xi Xxi]; first exact: rpred_nat. + have [k kerZ'k def_xi] := seqIndP Xxi; rewrite !inE sub1G andbT in kerZ'k. + set a := aX xi; have Dxi1: xi 1%g = a * #|W1|%:R. + rewrite /a /aX def_xi cfInd1 // -divgS // -(sdprod_card defL) mulKn //. + by rewrite mulrC mulfK ?neq0CG // irr1_degree natCK. + split=> //; have Da: a = 'chi_k 1%g. + apply: (mulIf (neq0CG W1)); rewrite -Dxi1 def_xi cfInd1 //. + by rewrite mulrC -divgS // -(sdprod_card defL) mulKn. + have [i0 nzi0 Res_k]: exists2 i: Iirr Z, i != 0 & 'Res 'chi_k = a *: 'chi_i. + have [chi /andP[Nchi lin_chi] defRkZ] := cfcenter_Res 'chi_k. + have sZ_Zk: Z \subset 'Z('chi_k)%CF. + by rewrite (subset_trans sZZ) // -cap_cfcenter_irr bigcap_inf. + have{Nchi lin_chi} /irrP[i defRk] : 'Res chi \in irr Z. + by rewrite lin_char_irr // qualifE cfRes_char // cfRes1. + have{chi defRk defRkZ} defRk: 'Res 'chi_k = a *: 'chi_i. + by rewrite -defRk -linearZ -/a Da -defRkZ /= cfResRes ?cfcenter_sub. + exists i => //; apply: contra kerZ'k. + rewrite -subGcfker => /subsetP sZker. + apply/subsetP=> t Zt; rewrite cfkerEirr inE. + by rewrite -!(cfResE _ sZH) // defRk !cfunE cfker1 ?sZker. + set phi := 'chi_i0 in Res_k; pose a_ i := '['Ind[H] phi, 'chi_i]. + pose rp := irr_constt ('Ind[H] phi). + have defIphi: 'Ind phi = \sum_(i in rp) a_ i *: 'chi_i. + exact: cfun_sum_constt. + have a_k: a_ k = a. + by rewrite /a_ -cfdot_Res_r Res_k cfdotZr cfnorm_irr mulr1 rmorph_nat. + have rp_k: k \in rp by rewrite inE ['[_, _]]a_k Da irr1_neq0. + have resZr i: i \in rp -> 'Res[Z] 'chi_i = a_ i *: phi. + move=> r_i; rewrite ['Res _]cfun_sum_cfdot. + rewrite (bigD1 i0) // big1 => /= [|j i0'j]. + rewrite cfdot_Res_l addr0 -/phi cfdotC conj_Cnat //. + by rewrite Cnat_cfdot_char_irr ?cfInd_char ?irr_char. + apply/eqP; rewrite scaler_eq0 cfdot_Res_l. + rewrite -(inj_eq (mulfI r_i)) mulr0 -/(a_ i) -cfdotZl. + have: '['Ind[H] phi, 'Ind[H] 'chi_j] = 0. + apply: not_cfclass_Ind_ortho => //. + have defIj: 'I_H['chi_j] = H. + apply/setIidPl; apply: subset_trans (cent_sub_inertia _). + by rewrite centsC (subset_trans sZZ) ?subsetIr. + rewrite -(congr1 (cfclass _) defIj) cfclass_inertia inE. + by rewrite eq_sym (inj_eq irr_inj). + rewrite defIphi cfdot_suml => /psumr_eq0P-> //; first by rewrite eqxx. + move=> i1 _; rewrite cfdotZl. + by rewrite mulr_ge0 ?Cnat_ge0 ?Cnat_cfdot_char ?cfInd_char ?irr_char. + have lin_phi: phi 1%g = 1. + apply: (mulfI (irr1_neq0 k)); have /resZr/cfunP/(_ 1%g) := rp_k. + by rewrite cfRes1 // cfunE mulr1 a_k Da. + have Da_ i: i \in rp -> 'chi_i 1%g = a_ i. + move/resZr/cfunP/(_ 1%g); rewrite cfRes1 // cfunE => ->. + by rewrite lin_phi mulr1. + pose chi i := 'Ind[L, H] 'chi_i; pose alpha i := chi i - a_ i *: eta1. + have Aalpha i: i \in rp -> alpha i \in 'CF(L, A). + move=> r_i; rewrite cfun_onD1 !cfunE cfInd1 // (uniY _ Yeta1). + rewrite -divgS // -(sdprod_card defL) mulKn // Da_ // mulrC subrr eqxx. + by rewrite memvB ?cfInd_normal ?memvZ // (seqInd_on _ Yeta1). + have [sum_alpha sum_a2]: + 'Ind phi - #|H : Z|%:R *: eta1 = \sum_(i in rp) a_ i *: alpha i + /\ \sum_(i in rp) a_ i ^+ 2 = #|H : Z|%:R. + + set rhs2 := _%:R; set lhs1 := _ - _; set rhs1 := \sum_(i | _) _. + set lhs2 := \sum_(i | _) _. + have eq_diff: lhs1 - rhs1 = (lhs2 - rhs2) *: eta1. + rewrite scalerBl addrAC; congr (_ - _). + rewrite -(cfIndInd _ sHL sZH) defIphi linear_sum -sumrB scaler_suml. + apply: eq_bigr => i rp_i; rewrite linearZ scalerBr opprD addNKr. + by rewrite opprK scalerA. + have: (lhs1 - rhs1) 1%g == 0. + rewrite !cfunE -(cfIndInd _ sHL sZH) !cfInd1 // lin_phi mulr1. + rewrite -divgS // -(sdprod_card defL) mulKn // mulrC uniY // subrr. + rewrite sum_cfunE big1 ?subrr // => i rp_i. + by rewrite cfunE (cfun_on0 (Aalpha i rp_i)) ?mulr0 // !inE eqxx. + rewrite eq_diff cfunE mulf_eq0 subr_eq0 (negPf (seqInd1_neq0 _ Yeta1)) //. + rewrite orbF => /eqP sum_a2; split=> //; apply/eqP; rewrite -subr_eq0. + by rewrite eq_diff sum_a2 subrr scale0r. + have Xchi i: i \in rp -> chi i \in X. + move=> rp_i; apply/seqIndP; exists i => //; rewrite !inE sub1G andbT. + apply: contra rp_i => kerZi; rewrite -cfdot_Res_r. + suffices ->: 'Res[Z] 'chi_i = ('chi_i 1%g)%:A. + by rewrite cfdotZr -irr0 cfdot_irr (negPf nzi0) mulr0. + apply/cfun_inP=> t Zt; rewrite cfResE // cfunE cfun1E Zt mulr1. + by rewrite cfker1 ?(subsetP kerZi). + have oRY i: i \in rp -> orthogonal (R (chi i)) (map tau1 Y). + move/Xchi=> Xchi_i; rewrite orthogonal_sym. + by rewrite (coherent_ortho_supp scohS) // ?sXS // (contraL (X'Y _)). + have n1Y1: '[Y1] = 1. + have [_ oYYt] := orthonormalP oYtau. + have [-> | [_ ->]] := defY1; + by rewrite ?cfnormN oYYt ?eqxx ?map_f // cfAut_seqInd. + have YtauY1: Y1 \in 'Z[map tau1 Y]. + have [-> | [_ ->]] := defY1; + by rewrite ?rpredN mem_zchar ?map_f ?cfAut_seqInd. + have /fin_all_exists[XbZ defXbZ] i: exists XbZ, let: (X1, b, Z1) := XbZ in + [/\ tau (alpha i) = X1 - b *: Y1 + Z1, + i \in rp -> X1 \in 'Z[R (chi i)], i \in rp -> b \is Creal, + '[Z1, Y1] = 0 & i \in rp -> orthogonal Z1 (R (chi i))]. + - have [X1 dX1 [YZ1 [dXYZ _ oYZ1R]]] := + orthogonal_split (R (chi i)) (tau (alpha i)). + have [Y1b dY1b [Z1 [dYZ1 _ oZY1]]] := orthogonal_split Y1 YZ1. + have{dY1b} [|b Db dY1b] := orthogonal_span _ dY1b. + by rewrite /pairwise_orthogonal /= inE eq_sym -cfnorm_eq0 n1Y1 oner_eq0. + exists (X1, - b Y1, Z1); split. + + by rewrite dXYZ dYZ1 dY1b scaleNr big_seq1 opprK addrA. + + have [_ _ _ Rok _] := scohS => /Xchi/sXS/Rok[ZR oRR _]. + have [_ -> ->] := orthonormal_span oRR dX1. + rewrite big_seq rpred_sum // => aa Raa. + rewrite scale_zchar ?mem_zchar //. + rewrite -[X1](addrK YZ1) -dXYZ cfdotBl (orthoPl oYZ1R) // subr0. + rewrite Cint_cfdot_vchar ?(ZR aa) //. + rewrite !(rpredB, cfInd_vchar) ?irr_vchar //. + rewrite scale_zchar ?(seqInd_vcharW Yeta1) // Cint_cfdot_vchar_irr //. + by rewrite cfInd_vchar ?irr_vchar. + + move=> rp_i; rewrite Db -[Y1b](addrK Z1) -dYZ1 cfdotBl (orthoP oZY1). + rewrite subr0 n1Y1 divr1 -[YZ1](addKr X1) -dXYZ cfdotDl cfdotNl. + rewrite (span_orthogonal (oRY i rp_i)) ?(zchar_span YtauY1) //. + rewrite oppr0 add0r Creal_Cint // rpredN Cint_cfdot_vchar //. + rewrite !(cfInd_vchar, rpredB) ?irr_vchar //. + rewrite -Da_ // scale_zchar ?Cint_Cnat ?Cnat_irr1 //. + exact: (seqInd_vcharW Yeta1). + apply: zchar_trans_on YtauY1 => _ /mapP[eta Yeta ->]. + by rewrite Ztau1 ?seqInd_zcharW. + + exact: (orthoP oZY1). + move/oRY=> oRiY; apply/orthoPl=> aa Raa. + rewrite -[Z1](addKr Y1b) -dYZ1 cfdotDl cfdotNl cfdotC (orthoPl oYZ1R) //. + rewrite dY1b addr0 big_seq1 cfdotZr. + by have [-> | [_ ->]] := defY1; + rewrite ?cfdotNr (orthogonalP oRiY) ?map_f ?cfAut_seqInd //; + rewrite ?(oppr0, mulr0, rmorph0). + pose X1 i := (XbZ i).1.1; pose b i := (XbZ i).1.2; pose Z1 i := (XbZ i).2. + have R_X1 i: i \in rp -> X1 i \in 'Z[R (chi i)]. + by rewrite /X1; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. + have Rb i: i \in rp -> b i \is Creal. + by rewrite /b; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. + have oZY1 i: '[Z1 i, Y1] = 0. + by rewrite /Z1; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. + have oZ1R i: i \in rp -> orthogonal (Z1 i) (R (chi i)). + by rewrite /Z1; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. + have{defXbZ} defXbZ i: tau (alpha i) = X1 i - b i *: Y1 + Z1 i. + by rewrite /X1 /b /Z1; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. + have ub_alpha i: i \in rp -> + [/\ '[chi i] <= '[X1 i] + & '[a_ i *: eta1] <= '[b i *: Y1 - Z1 i] -> + [/\ '[X1 i] = '[chi i], '[b i *: Y1 - Z1 i] = '[a_ i *: eta1] + & exists2 E, subseq E (R (chi i)) & X1 i = \sum_(aa <- E) aa]]. + - move=> rp_i; apply: (subcoherent_norm scohS) (erefl _) _. + + rewrite sXS ?Xchi // scale_zchar ?(seqInd_vcharW Yeta1) //; last first. + by rewrite Cint_cfdot_vchar_irr // cfInd_vchar ?irr_vchar. + split=> //; apply/orthoPr=> xi2; rewrite !inE => Dxi2. + rewrite cfdotZr (orthogonalP oXY) ?mulr0 //. + by case/pred2P: Dxi2 => ->; rewrite ?cfAut_seqInd ?Xchi. + + have [_ IZtau _ _ _] := scohS. + apply: sub_iso_to IZtau; [apply: zchar_trans_on | exact: zcharW]. + apply/allP; rewrite /= zchar_split (cfun_onS (setSD _ sHL)) ?Aalpha //. + rewrite rpredB ?scale_zchar ?seqInd_zcharW ?(sYS eta1) ?sXS ?Xchi //. + by rewrite sub_aut_zchar ?zchar_onG ?seqInd_zcharW ?cfAut_seqInd; + rewrite ?sXS ?Xchi //; apply: seqInd_vcharW. + by rewrite -Da_ // irr1_degree rpred_nat. + suffices oYZ_R: orthogonal (b i *: Y1 - Z1 i) (R (chi i)). + rewrite opprD opprK addrA -defXbZ cfdotC. + rewrite (span_orthogonal oYZ_R) ?memv_span1 ?conjC0 //. + exact: (zchar_span (R_X1 i rp_i)). + apply/orthoPl=> aa Raa. + rewrite cfdotBl cfdotZl (orthoPl (oZ1R i _)) //. + by rewrite subr0 cfdotC; have [-> | [_ ->]] := defY1; + rewrite ?cfdotNr (orthogonalP (oRY i _)) ?map_f ?cfAut_seqInd //; + by rewrite ?(mulr0, oppr0, rmorph0). + have leba i: i \in rp -> b i <= a_ i. + move=> rp_i; have ai_gt0: a_ i > 0 by rewrite -Da_ ?irr1_gt0. + have /orP [b_le0|b_ge0] := Rb i rp_i; last first. + by rewrite (ler_trans b_ge0) ?ltrW. + rewrite -(@ler_pexpn2r _ 2) //; last exact: ltrW. + apply: ler_trans (_ : '[b i *: Y1 - Z1 i] <= _). + rewrite cfnormBd; last by rewrite cfdotZl cfdotC oZY1 ?conjC0 ?mulr0. + rewrite cfnormZ (normr_idP _) // n1Y1 mulr1 addrC -subr_ge0 addrK. + exact: cfnorm_ge0. + rewrite -(ler_add2l '[X1 i]) -cfnormBd; last first. + rewrite cfdotBr cfdotZr (span_orthogonal (oRY i _)) //; last first. + - exact: (zchar_span YtauY1). + - exact: (zchar_span (R_X1 i rp_i)). + rewrite mulr0 sub0r cfdotC (span_orthogonal (oZ1R i _)) ?raddf0 //. + exact: memv_span1. + exact: (zchar_span (R_X1 i rp_i)). + have Salpha: alpha i \in 'Z[S, L^#]. + rewrite zchar_split (cfun_onS (setSD _ sHL)) ?Aalpha //. + rewrite rpredB ?scale_zchar ?seqInd_zcharW + ?(sYS _ Yeta1) ?sXS ?Xchi //. + by rewrite -Da_ // irr1_degree rpred_nat. + rewrite opprD opprK addrA -defXbZ // Itau ?Salpha //. + rewrite cfnormBd; last first. + by rewrite cfdotZr (orthogonalP oXY) ?mulr0 ?Xchi. + rewrite cfnormZ (normr_idP _) ?(ltrW ai_gt0) //. + have [_ /(_ eta1)->//] := orthonormalP oY; rewrite eqxx mulr1 ler_add2r. + by have [] := ub_alpha i rp_i. + have{leba} eq_ab: {in rp, a_ =1 b}. + move=> i rp_i; apply/eqP; rewrite -subr_eq0; apply/eqP. + apply: (mulfI (irr1_neq0 i)); rewrite mulr0 Da_ // mulrBr. + move: i rp_i; apply: psumr_eq0P => [i rp_i | ]. + by rewrite subr_ge0 ler_pmul2l ?leba // -Da_ ?irr1_gt0. + have [X2 oX2Y /(congr1 (cfdotr Y1))] := Dpsi1 i0 nzi0. + rewrite sumrB sum_a2 sum_alpha /tau linear_sum /= cfdot_suml cfdotBl. + rewrite (span_orthogonal oX2Y) ?memv_span1 ?(zchar_span YtauY1) // add0r. + rewrite cfdotZl n1Y1 mulr1 => /(canLR (@opprK _)) <-. + rewrite -opprD -big_split big1 ?oppr0 //= => i rp_i. + rewrite linearZ cfdotZl /= -/tau defXbZ addrC cfdotDl oZY1 addr0. + rewrite cfdotBl cfdotZl n1Y1 mulr1. + rewrite (span_orthogonal (oRY i _)) ?(zchar_span YtauY1) //. + by rewrite add0r mulrN subrr. + exact: (zchar_span (R_X1 i rp_i)). + exists (X1 k). + apply/orthoPl=> psi /memv_span Ypsi. + by rewrite (span_orthogonal (oRY k _)) // (zchar_span (R_X1 k rp_k)). + apply/eqP; rewrite -/a def_xi -a_k defXbZ addrC -subr_eq0 eq_ab // addrK. + have n1eta1: '[eta1] = 1 by have [_ -> //] := orthonormalP oY; rewrite eqxx. + rewrite -cfnorm_eq0 -(inj_eq (addrI '[b k *: Y1])). + have [_ [|_]] := ub_alpha k rp_k. + rewrite cfnormBd; last by rewrite cfdotZl cfdotC oZY1 conjC0 mulr0. + by rewrite addrC !cfnormZ eq_ab // n1Y1 n1eta1 -subr_ge0 addrK cfnorm_ge0. + rewrite cfnormBd; last by rewrite cfdotZl cfdotC oZY1 conjC0 mulr0. + by move=> -> _; rewrite addr0 !cfnormZ eq_ab // n1Y1 n1eta1. + have oX: pairwise_orthogonal X by rewrite (sub_pairwise_orthogonal sXS). + have [_ oYY] := orthonormalP oY. + have [[N_S _ _] [_ /(_ _ _)/zcharW/=Ztau] _ _ _] := scohS. + have n1eta: '[eta1] = 1 by rewrite oYY ?eqxx. + have n1Y1: '[Y1] = 1. + have [_ oYYt] := orthonormalP oYtau. + have [-> | [_ ->]] := defY1; + by rewrite ?cfnormN oYYt ?eqxx ?map_f // cfAut_seqInd. + have YtauY1: Y1 \in <<map tau1 Y>>%VS. + by have [-> | [_ ->]] := defY1; + rewrite ?memvN memv_span ?map_f ?cfAut_seqInd. + have Z_Y1: Y1 \in 'Z[irr G]. + by case: defY1 => [|[_]] ->; rewrite ?rpredN Ztau1 ?mem_zchar ?ccY. + have Zalpha xi: xi \in X -> xi - a xi *: eta1 \in 'Z[S, L^#]. + move=> Xxi; rewrite zcharD1E rpredB ?scale_zchar; + rewrite ?seqInd_zcharW ?(sXS xi) ?sYS //. + by rewrite !cfunE (uniY eta1) //= subr_eq0; have [<-] := Dxa xi Xxi. + have Zbeta eta: eta \in Y -> eta - eta1 \in 'Z[S, L^#]. + move=> Yeta; rewrite zcharD1E rpredB ?seqInd_zcharW ?sYS //=. + by rewrite !cfunE !uniY ?subrr. + have nza xi: xi \in X -> a xi != 0. + move=> Xxi; have [/eqP Dxi1 _] := Dxa _ Xxi; apply: contraTneq Dxi1 => ->. + by rewrite mul0r (seqInd1_neq0 _ Xxi). + have alphaY xi: xi \in X -> '[tau (xi - a xi *: eta1), Y1] = - a xi. + case/Dxa=> _ [X1 oX1Y ->]; rewrite cfdotBl cfdotZl n1Y1 mulr1. + by rewrite (span_orthogonal oX1Y) ?memv_span1 ?add0r. + have betaY eta: eta \in Y -> '[tau (eta - eta1), Y1] = (eta == eta1)%:R - 1. + move=> Yeta; rewrite -Dtau1; last first. + by rewrite zchar_split (zchar_on (Zbeta eta _)) ?rpredB ?seqInd_zcharW. + rewrite raddfB cfdotBl. + have [-> | [szY2 ->]] := defY1. + by rewrite !{1}Itau1 ?seqInd_zcharW // !oYY // !eqxx. + rewrite !cfdotNr opprK !{1}Itau1 ?oYY ?seqInd_zcharW ?cfAut_seqInd //. + have defY: (eta1 :: eta1^*)%CF =i Y. + apply: proj1 (leq_size_perm _ _ _); last by rewrite szY2. + by rewrite /= inE eq_sym (seqInd_conjC_neq _ _ _ Yeta1). + by apply/allP; rewrite /= Yeta1 cfAut_seqInd. + rewrite -defY !inE in Yeta; case/pred2P: Yeta => ->. + rewrite eqxx eq_sym (negPf (seqInd_conjC_neq _ _ _ Yeta1)) //. + by rewrite addrC !subrr. + by rewrite eqxx eq_sym (negPf (seqInd_conjC_neq _ _ _ Yeta1)) ?add0r ?addr0. + pose tau2_X xi := tau (xi - a xi *: eta1) + a xi *: Y1. + pose tau3_Y eta := tau (eta - eta1) + Y1. + have Itau2_X: {in X, isometry tau2_X, to 'Z[irr G]}. + split=> [xi1 xi2 Xxi1 Xxi2 | xi Xxi]; last first. + by rewrite rpredD ?rpredZ_Cint ?Za ?Ztau ?Zalpha. + rewrite /= cfdotDl !cfdotDr Itau ?Zalpha // cfdotBl !cfdotBr opprB !cfdotZr. + rewrite !aut_Cint ?Za // !cfdotZl (cfdotC Y1) !alphaY // n1Y1 n1eta rmorphN. + rewrite aut_Cint // (cfdotC eta1) !(orthogonalP oXY _ eta1) // conjC0. + by rewrite !mulr0 !subr0 !mulr1 !mulrN mulrC !addrA subrK addrK. + have{Itau2_X} [tau2 Dtau2 Itau2] := Zisometry_of_iso oX Itau2_X. + have{Itau2} cohX: coherent_with X L^# tau tau2. + split=> // xi; rewrite zcharD1E => /andP[/zchar_expansion[]// z Zz ->{xi}]. + pose sum_za := \sum_(xi <- X) z xi * a xi => /eqP sum_xi_0. + have{sum_xi_0} sum_za_0: sum_za = 0. + apply: (mulIf (neq0CG W1)); rewrite mul0r -{}sum_xi_0 sum_cfunE mulr_suml. + by apply: eq_big_seq => xi /Dxa[xi_1 _]; rewrite !cfunE xi_1 mulrA. + rewrite -[rhs in tau rhs](subrK (sum_za *: eta1)) {1}scaler_suml -sumrB. + rewrite -[tau _](addrK (sum_za *: Y1)) {1 3}sum_za_0 !scale0r addr0 subr0. + rewrite scaler_suml !raddf_sum [tau _]raddf_sum -big_split /= -/tau. + apply: eq_big_seq => xi Xxi; rewrite raddfZ_Cint //= Dtau2 //. + by rewrite -!scalerA -scalerBr [tau _]linearZ -scalerDr. + have Itau3_Y: {in Y, isometry tau3_Y, to 'Z[irr G]}. + split=> [eta3 eta4 Yeta3 Yeta4 | eta Yeta]; last first. + by rewrite rpredD // Ztau ?Zbeta. + rewrite /= cfdotDl !cfdotDr n1Y1 (cfdotC Y1) !betaY // Itau ?Zbeta //. + rewrite cfdotBl !cfdotBr !oYY // eqxx rmorphB rmorph1 rmorph_nat subrK. + rewrite (eq_sym eta1) opprB !addrA 3!(addrAC _ (- _)) addrK. + by rewrite(addrAC _ 1) subrK addrK. + have{oY} oY := orthonormal_orthogonal oY. + have{Itau3_Y} [tau3 Dtau3 Itau3] := Zisometry_of_iso oY Itau3_Y. + have{Itau3 cohY} cohY: coherent_with Y L^# tau tau3. + split=> // eta; rewrite zcharD1E => /andP[/zchar_expansion[]//z Zz ->{eta}]. + pose sum_z := \sum_(eta <- Y) z eta => /eqP sum_eta_0. + have{sum_eta_0} sum_z_0: sum_z = 0. + apply: (mulIf (neq0CG W1)); rewrite mul0r -sum_eta_0 sum_cfunE mulr_suml. + by apply: eq_big_seq => xi /uniY eta_1; rewrite !cfunE eta_1. + rewrite -[rhs in tau rhs](subrK (sum_z *: eta1)) {1}scaler_suml -sumrB. + rewrite -[tau _](addrK (sum_z *: Y1)) {1 3}sum_z_0 !scale0r addr0 subr0. + rewrite scaler_suml !raddf_sum [tau _]raddf_sum -big_split /= -/tau. + apply: eq_big_seq => eta Yeta; rewrite raddfZ_Cint //= Dtau3 //. + by rewrite -scalerBr [tau _]linearZ -scalerDr. + have [-> | ] := altP (@nilP _ X); first by exists tau3. + rewrite -lt0n -has_predT => /hasP[xi1 Xxi1 _]. + have: tau (xi1 - a xi1 *: eta1) = tau2 xi1 - tau3 (a xi1 *: eta1). + rewrite [tau3 _]linearZ Dtau2 //= Dtau3 // /tau3_Y subrr [tau 0]linear0. + by rewrite add0r addrK. + apply: (bridge_coherent scohS ccsXS cohX ccsYS cohY X'Y). + by rewrite (zchar_on (Zalpha _ Xxi1)) // rpredZ_Cint ?mem_zchar. +pose wf S1 := cfConjC_subset S1 S /\ {subset X ++ Y <= S1}. +pose S1 := [::] ++ X ++ Y; set S2 := [::] in S1; rewrite -[X ++ Y]/S1 in cohXY. +have wfS1: wf S1. + do 2!split=> //; rewrite /S1 /= ?cat_uniq ?uX ?uY ?(introT hasPn) //. + by apply/allP; rewrite all_cat !(introT allP). + by move=> phi; rewrite !mem_cat => /orP[/ccX|/ccY]->; rewrite ?orbT. +move: {2}_.+1 (ltnSn (size S - size S1)) => n. +elim: n => // n IHn in (S2) S1 wfS1 cohXY *; rewrite ltnS => leSnS1. +have [ccsS1S sXYS1] := wfS1; have [uS1 sS1S ccS1] := ccsS1S. +have [sSS1 | /allPn[psi /= Spsi notS1psi]] := altP (@allP _ (mem S1) S). + exact: subset_coherent cohXY. +have [_ _ ccS] := uccS. +have [neq_psi_c Spsic] := (hasPn nrS _ Spsi, ccS _ Spsi). +have wfS1': wf [:: psi, psi^* & S1]%CF. + split; last by move=> xi XYxi; rewrite !inE sXYS1 ?orbT. + split=> [|xi|xi]; rewrite /= !inE. + - by rewrite negb_or eq_sym neq_psi_c notS1psi (contra (ccS1 _)) ?cfConjCK. + - by do 2?[case/predU1P=> [-> //|]] => /sS1S. + rewrite (inv_eq (@cfConjCK _ _)) (can_eq (@cfConjCK _ _)) orbCA !orbA. + by case: pred2P => // _ /ccS1. +apply: (IHn [:: psi, psi^* & S2]%CF) => //; last first. + rewrite -subSn ?uniq_leq_size //; try by have [[]] := wfS1'. + by rewrite /= subSS (leq_trans _ leSnS1) // leq_sub2l ?leqW. +have ltZH': Z \proper H^`(1)%g. + rewrite properEneq sZH' andbT; apply: contraNneq notS1psi => eqZH'. + have [i /setDP[_ nt_i] ->] := seqIndP Spsi; rewrite sXYS1 // mem_cat. + rewrite !mem_seqInd ?normal1 //= -eqZH' !inE in nt_i *. + by rewrite sub1G nt_i andbT orNb. +have: [/\ eta1 \in S1, psi \in S & psi \notin S1]. + by rewrite Spsi sXYS1 // mem_cat Yeta1 orbT. +have /seqIndC1P[i nzi Dpsi] := Spsi. +move/(extend_coherent scohS ccsS1S); apply; split=> //. + rewrite (uniY _ Yeta1) Dpsi cfInd1 // (index_sdprod defL) dvdC_mulr //. + by rewrite Cint_Cnat ?Cnat_irr1. +rewrite !big_cat //= (big_rem _ Yeta1) /= addrC -!addrA -big_cat //=. +rewrite sum_seqIndD_square ?normal1 ?sub1G // indexg1 addrC. +rewrite -subr_gt0 -!addrA ltr_spaddl //. + have /irrY/irrP[j ->] := Yeta1. + by rewrite cfnorm_irr divr1 exprn_gt0 ?irr1_gt0. +rewrite big_seq addr_ge0 ?sumr_ge0 // => [phi Sphi|]. + rewrite mulr_ge0 ?invr_ge0 ?cfnorm_ge0 ?exprn_ge0 // ?char1_pos //. + suffices /seqInd_char: phi \in S by apply: char1_ge0. + rewrite sS1S // !mem_cat; rewrite mem_cat in Sphi. + by case/orP: Sphi => [/mem_rem-> | ->]; rewrite ?orbT. +rewrite subr_ge0 -(Lagrange_index sHL sZH) -oW1 natrM mulrC -mulrA. +rewrite uniY ?ler_pmul2l ?gt0CG //. +rewrite -(prednK (cardG_gt0 Z)) [zz in zz - 1]mulrSr addrK -natrM. +rewrite Dpsi cfInd1 // -oW1. +rewrite -(@ler_pexpn2r _ 2) -?topredE /= ?mulr_ge0 ?ler0n //; last first. + by rewrite char1_ge0 ?irr_char. +rewrite !exprMn -!natrX mulrA -natrM. +apply: ler_trans (_ : (4 * #|W1| ^ 2)%:R * #|H : Z|%:R <= _). + rewrite ler_pmul2l; last by rewrite ltr0n muln_gt0 !expn_gt0 cardG_gt0. + rewrite (ler_trans (irr1_bound i)) // ler_nat dvdn_leq // indexgS //. + by rewrite (subset_trans sZZ) // -cap_cfcenter_irr bigcap_inf. +rewrite -natrM ler_nat expnMn mulnC -mulnA leq_pmul2l //. +have [in_caseA | in_caseB] := boolP caseA. + have regW1Z: semiregular Z W1. + have [in_c1 | in_c2] := boolP case_c1. + move=> x /(Frobenius_reg_ker in_c1) regHx; apply/trivgP. + by rewrite -regHx setSI. + have [/= _ [_ [_ _ _ prW1H] _] _ _] := c2W2 in_c2. + move=> x /prW1H prHx; apply/trivgP; rewrite -((_ =P [1]) in_caseA) -prHx. + by rewrite subsetI subIset ?sZZ // setSI. + rewrite -(mul1n (4 * _)%N) leq_mul // -(expnMn 2 _ 2) leq_exp2r //. + rewrite dvdn_leq //; first by rewrite -subn1 subn_gt0 cardG_gt1. + rewrite Gauss_dvd ?(@pnat_coprime 2) -?odd_2'nat ?(oddSg sW1L) //. + rewrite dvdn2 -{1}subn1 odd_sub // (oddSg (normal_sub nsZL)) //=. + by rewrite regular_norm_dvd_pred // (subset_trans sW1L) ?normal_norm. +rewrite -(muln1 (4 * _)%N) leq_mul //; last first. + by rewrite expn_gt0 -subn1 subn_gt0 orbF cardG_gt1. +rewrite -(expnMn 2 _ 2) -(Lagrange_index (der_sub 1 H) sZH') leq_mul //. + rewrite -(prednK (indexg_gt0 _ _)) leqW // dvdn_leq //. + by rewrite -subn1 subn_gt0 indexg_gt1 proper_subn. + rewrite Gauss_dvd ?(@pnat_coprime 2) -?odd_2'nat ?(oddSg sW1L) //. + rewrite dvdn2 -{1}subn1 odd_sub // -card_quotient ?der_norm //. + rewrite quotient_odd ?(oddSg sHL) //=. + rewrite (card_isog (quotient_isog (subset_trans sW1L nH'L) _)); last first. + by rewrite coprime_TIg ?(coprimeSg (der_sub 1 H)). + exact: Frobenius_dvd_ker1 frobL1. +rewrite -(prednK (indexg_gt0 _ _)) leqW // dvdn_leq //. + by rewrite -subn1 subn_gt0 indexg_gt1 proper_subn. +rewrite Gauss_dvd ?(@pnat_coprime 2) -?odd_2'nat ?(oddSg sW1L) //. +rewrite dvdn2 -{1}subn1 odd_sub //. +rewrite -card_quotient ?(subset_trans (der_sub 1 H)) //. +rewrite quotient_odd ?(oddSg (der_sub 1 H)) ?(oddSg sHL) //=. +have /andP[sZL nZL] := nsZL. +rewrite (card_isog (quotient_isog (subset_trans sW1L nZL) _)); last first. + by rewrite coprime_TIg ?(coprimeSg sZH). +suffices: [Frobenius (H^`(1) / Z) <*> (W1 / Z) = (H^`(1) / Z) ><| (W1 / Z)]%g. + exact: Frobenius_dvd_ker1. +suffices: [Frobenius (L / Z) = (H / Z) ><| (W1 / Z)]%g. + apply: Frobenius_subl (quotientS Z (der_sub 1 H)) _. + by rewrite quotient_neq1 // (normalS sZH' (der_sub 1 H)). + by rewrite quotient_norms ?(subset_trans sW1L). +apply: (Frobenius_coprime_quotient defL nsZL) => //; split=> [|x W1x]. + exact: sub_proper_trans sZH' ltH'H. +have /caseB_P[/c2W2[_ [_ [_ _ _ -> //] _] _ _] _ _] := in_caseB. +by rewrite /Z (negPf in_caseB). +Qed. + +End Six. + + diff --git a/mathcomp/odd_order/PFsection7.v b/mathcomp/odd_order/PFsection7.v new file mode 100644 index 0000000..eed77b7 --- /dev/null +++ b/mathcomp/odd_order/PFsection7.v @@ -0,0 +1,819 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg poly finset center. +Require Import fingroup morphism perm automorphism quotient action zmodp. +Require Import gfunctor gproduct cyclic pgroup commutator nilpotent frobenius. +Require Import matrix mxalgebra mxrepresentation BGsection3 vector. +Require Import ssrnum algC classfun character inertia vcharacter. +Require Import PFsection1 PFsection2 PFsection4 PFsection5 PFsection6. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 7: *) +(* Non-existence of a Certain Type of Group of Odd Order *) +(* Defined here: *) +(* inDade ddA == the right inverse to the Dade isometry with respect to G, *) +(* L, A, given ddA : Dade_hypothesis G L A. *) +(* phi^\rho == locally-bindable Notation for invDade ddA phi. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Reserved Notation "alpha ^\rho" (at level 2, format "alpha ^\rho"). + +Section Seven. + +Variables (gT : finGroupType) (G : {group gT}). +Implicit Types (L H P : {group gT}) (DH : gT -> {group gT}). + +(* Properties of the inverse to the Dade isometry (Peterfalvi (7.1) to (7.3). *) +Section InverseDade. + +Variables (A : {set gT}) (L : {group gT}). +Hypothesis ddA : Dade_hypothesis G L A. + +Local Notation "alpha ^\tau" := (Dade ddA alpha). +Local Notation Atau := (Dade_support ddA). +Local Notation H := (Dade_signalizer ddA). + +Let nsAL : A <| L. Proof. by have [] := ddA. Qed. +Let sAL : A \subset L. Proof. exact: normal_sub nsAL. Qed. +Let nAL : L \subset 'N(A). Proof. exact: normal_norm nsAL. Qed. +Let sLG : L \subset G. Proof. by have [] := ddA. Qed. + +(* This is the Definition embedded in Peterfalvi, Hypothesis (7.1). *) +Fact invDade_subproof (chi : 'CF(G)) : + is_class_fun <<L>> + [ffun a => #|H a|%:R^-1 * (\sum_(x in H a) chi (x * a)%g) *+ (a \in A)]. +Proof. +rewrite genGid; apply: intro_class_fun => [x y Lx Ly | x notLx]; last first. + by rewrite (contraNF (subsetP sAL x)). +rewrite memJ_norm ?(subsetP nAL) // !mulrb; case: ifP => // Ax. +rewrite (DadeJ ddA) // cardJg; congr (_ * _). +rewrite big_imset /= => [|z y0 _ _ /=]; last exact: conjg_inj. +by apply: eq_bigr => u Hu; rewrite -conjMg cfunJ // (subsetP sLG). +Qed. +Definition invDade alpha := Cfun 1 (invDade_subproof alpha). + +Local Notation "alpha ^\rho" := (invDade alpha). + +Fact invDade_is_linear : linear invDade. +Proof. +move=> mu alpha beta; apply/cfunP=> a; rewrite !cfunElock. +rewrite mulrnAr -mulrnDl mulrCA -mulrDr; congr (_ * _ *+ _). +by rewrite big_distrr -big_split; apply: eq_bigr => x _; rewrite !cfunE. +Qed. +Canonical invDade_linear := Linear invDade_is_linear. +Canonical invDade_additive := Additive invDade_is_linear. + +Lemma invDade_on chi : chi^\rho \in 'CF(L, A). +Proof. by apply/cfun_onP=> x notAx; rewrite cfunElock (negPf notAx). Qed. + +Lemma invDade_cfun1 : 1^\rho = '1_A. +Proof. +apply/cfunP=> x; rewrite cfuniE // cfunElock mulrb; case: ifP => //= Ax. +apply: canLR (mulKf (neq0CG _)) _; rewrite mulr1 -sumr_const. +apply: eq_bigr => u Hu; rewrite cfun1E (subsetP (subsetIl G 'C[x])) //. +have /sdprodP[_ <- _ _] := Dade_sdprod ddA Ax. +by rewrite mem_mulg // inE cent1id (subsetP sAL). +Qed. + +(* This is Peterfalvi (2.7), restated using invDade. *) +Lemma invDade_reciprocity chi alpha : + alpha \in 'CF(L, A) -> '[alpha^\tau, chi] = '[alpha, chi^\rho]. +Proof. +move=> Aalpha; apply: general_Dade_reciprocity => //= a Aa. +by rewrite cfunElock Aa. +Qed. + +(* This is Peterfalvi (7.2)(a). *) +Lemma DadeK alpha : alpha \in 'CF(L, A) -> (alpha^\tau)^\rho = alpha. +Proof. +move=> Aalpha; apply/cfunP=> a; rewrite cfunElock mulrb. +case: ifPn => [Aa | /cfun_on0-> //]; apply: canLR (mulKf (neq0CG _)) _. +rewrite mulr_natl -sumr_const; apply: eq_bigr => x Hx. +by rewrite (DadeE _ Aa) ?mem_class_support // mem_mulg ?set11. +Qed. + +(* This is Peterfalvi (7.2)(b); note that by (7.2)(a) chi is in the image of *) +(* tau iff chi = (chi^\rho)^\tau, and this condition is easier to write. *) +Lemma leC_norm_invDade chi : + '[chi^\rho] <= '[chi] ?= iff (chi == (chi^\rho)^\tau). +Proof. +have Achi_rho := invDade_on chi; rewrite -(Dade_isometry ddA) //. +set chi1 := _^\tau; rewrite -subr_eq0 -cfnorm_eq0; set mu := chi - chi1. +have owv: '[chi1, mu] = 0. + by rewrite invDade_reciprocity ?raddfB //= DadeK ?subrr. +rewrite -(subrK chi1 chi) -/mu addrC cfnormD owv conjC0 !addr0. +split; first by rewrite -subr_ge0 addrC addKr cfnorm_ge0. +by rewrite eq_sym addrC -subr_eq0 addrK. +Qed. + +(* This is Peterfalvi (7.3). *) +Lemma leC_cfnorm_invDade_support chi : + '[chi^\rho] <= #|G|%:R^-1 * (\sum_(g in Atau) `|chi g| ^+ 2) + ?= iff [forall a in A, forall u in H a, chi (u * a)%g == chi a]. +Proof. +have nsAtauG: Atau <| G := Dade_support_normal ddA. +pose chi1 := chi * '1_Atau; set RHS := _ * _. +have inA1 a x: a \in A -> x \in H a -> (x * a)%g \in Dade_support1 ddA a. + by move=> Aa Hx; rewrite mem_class_support ?mem_mulg ?set11. +have chi1E a x: a \in A -> x \in H a -> chi1 (x * a)%g = chi (x * a)%g. + move=> Aa Hx; rewrite cfunE cfuniE // mulr_natr mulrb. + by case: bigcupP => // [[]]; exists a; rewrite ?inA1. +have ->: chi^\rho = chi1^\rho. + apply/cfunP => a; rewrite !cfunElock !mulrb; case: ifP => // Aa. + by congr (_ * _); apply: eq_bigr => x /chi1E->. +have Achi1: chi1 \in 'CF(G, Atau). + by apply/cfun_onP=> x ?; rewrite cfunE (cfun_onP (cfuni_on _ _)) ?mulr0. +have{Achi1 RHS} <-: '[chi1] = RHS. + rewrite (cfnormE Achi1); congr (_ * _). + by apply: eq_bigr => x Atau_x; rewrite cfunE cfuniE // Atau_x mulr1. +congr (_ <= _ ?= iff _): (leC_norm_invDade chi1). +apply/eqP/forall_inP=> [chi1_id a Aa | chi_id]. + apply/forall_inP => x Ha_x; rewrite -{2}[a]mul1g -!chi1E // chi1_id mul1g. + by rewrite (DadeE _ Aa) ?inA1 ?Dade_id. +apply/cfunP => g; rewrite cfunE cfuniE // mulr_natr mulrb. +case: ifPn => [/bigcupP[a Aa] | /(cfun_onP (Dade_cfunS _ _))-> //]. +case/imset2P=> _ z /rcosetP[x Hx ->] Gz ->{g}; rewrite !cfunJ {z Gz}//. +have{chi_id} chi_id := eqP (forall_inP (chi_id a Aa) _ _). +rewrite chi_id // (DadeE _ Aa) ?inA1 {x Hx}// cfunElock mulrb Aa. +apply: canRL (mulKf (neq0CG _)) _; rewrite mulr_natl -sumr_const. +by apply: eq_bigr => x Hx; rewrite chi1E ?chi_id. +Qed. + +(* This is the norm expansion embedded in Peterfalvi (7.3). *) +Lemma cfnormE_invDade chi : + '[chi^\rho] = #|L|%:R^-1 * (\sum_(a in A) `|chi^\rho a| ^+ 2). +Proof. by apply: cfnormE; exact: invDade_on. Qed. + +End InverseDade. + +(* Hypothesis (7.4) and Lemma (7.5). *) +Section DadeCoverInequality. + +(* These declarations correspond to Peterfalvi, Hypothesis (7.4); as it is *) +(* only instantiated twice after this section we leave it unbundled. *) +Variables (I : finType) (L : I -> {group gT}) (A : I -> {set gT}). +Hypothesis ddA : forall i : I, Dade_hypothesis G (L i) (A i). + +Local Notation Atau i := (Dade_support (ddA i)). +Local Notation "alpha ^\rho" := (invDade (ddA _) alpha). +Hypothesis disjointA : forall i j, i != j -> [disjoint Atau i & Atau j]. + +(* This is Peterfalvi (7.5), generalised to all class functions of norm 1. *) +Lemma Dade_cover_inequality (chi : 'CF(G)) (G0 := G :\: \bigcup_i Atau i) : + '[chi] = 1 -> + #|G|%:R^-1 * (\sum_(g in G0) `|chi g| ^+ 2 - #|G0|%:R) + + \sum_i ('[chi^\rho]_(L i) - #|A i|%:R / #|L i|%:R) <= 0. +Proof. +move=> Nchi1; set vG := _^-1; rewrite sumrB /= addrCA mulrBr -addrA. +pose F (xi : 'CF(G)) (B : {set gT}) := vG * \sum_(g in B) `|xi g| ^+ 2. +have sumF xi: F xi G0 + \sum_i F xi (Atau i) = '[xi]. + rewrite (cfnormE (cfun_onG _)) -mulr_sumr -mulrDr; congr (_ * _). + rewrite -partition_disjoint_bigcup //=; set U_A := \bigcup_i _. + have sUG: U_A \subset G by apply/bigcupsP=> i _; apply: Dade_support_sub. + by rewrite -(setIidPr sUG) addrC -big_setID. +have ->: \sum_i #|A i|%:R / #|L i|%:R = \sum_i F 1 (Atau i). + apply: eq_bigr => i _; apply/eqP; rewrite /F. + have [[/andP[sAL nAL] _ _ _ _] sHG] := (ddA i, Dade_signalizer_sub (ddA i)). + rewrite -{1}[A i]setIid -cfdot_cfuni /normal ?sAL // -(invDade_cfun1 (ddA i)). + rewrite leC_cfnorm_invDade_support; apply/forall_inP=> a Aa. + by apply/forall_inP=> x Hx; rewrite !cfun1E groupMl // (subsetP (sHG a)). +have ->: vG * #|G0|%:R = F 1 G0. + congr (_ * _); rewrite -sumr_const; apply: eq_bigr => x /setDP[Gx _]. + by rewrite cfun1E Gx normr1 expr1n. +rewrite -opprD sumF cfnorm1 -Nchi1 -sumF opprD addNKr -oppr_ge0 opprB -sumrB. +by rewrite sumr_ge0 // => i _; rewrite subr_ge0 leC_cfnorm_invDade_support. +Qed. + +(* +set vG := _^-1; rewrite sumrB /= addrCA mulrBr -addrA. +pose F t (B : {set gT}) := vG * \sum_(g in B) `|'chi[G]_t g| ^+ 2. +have sumF t: F t G0 + \sum_i F t (Atau i) = 1. + rewrite -(cfnorm_irr t) (cfnormE (cfun_onG _)) -mulr_sumr -mulrDr. + congr (_ * _); rewrite -partition_disjoint_bigcup //=; set U_A := \bigcup_i _. + have sUG: U_A \subset G by apply/bigcupsP=> i _; apply: Dade_support_sub. + by rewrite -(setIidPr sUG) addrC -big_setID. +have ->: \sum_i #|A i|%:R / #|L i|%:R = \sum_i F 0 (Atau i). + apply: eq_bigr => i _; apply/eqP; rewrite /F irr0. + have [[/andP[sAL nAL] _ _ _ _] sHG] := (ddA i, Dade_signalizer_sub (ddA i)). + rewrite -{1}[A i]setIid -cfdot_cfuni /normal ?sAL // -(invDade_cfun1 (ddA i)). + rewrite leC_cfnorm_invDade_support; apply/forall_inP=> a Aa. + by apply/forall_inP=> x Hx; rewrite !cfun1E groupMl // (subsetP (sHG a)). +have ->: vG * #|G0|%:R = F 0 G0. + congr (_ * _); rewrite -sumr_const; apply: eq_bigr => x /setDP[Gx _]. + by rewrite irr0 cfun1E Gx normr1 expr1n. +rewrite -opprD sumF -(sumF r) opprD addNKr -oppr_ge0 opprB -sumrB. +by rewrite sumr_ge0 // => i _; rewrite subr_ge0 leC_cfnorm_invDade_support. +Qed. +*) + +End DadeCoverInequality. + +(* Hypothesis (7.6), and Lemmas (7.7) and (7.8) *) +Section Dade_seqIndC1. + +(* In this section, A = H^# with H <| L. *) +Variables L H : {group gT}. +Let A := H^#. +Hypothesis ddA : Dade_hypothesis G L A. + +Local Notation Atau := (Dade_support ddA). +Local Notation "alpha ^\tau" := (Dade ddA alpha). +Local Notation "alpha ^\rho" := (invDade ddA alpha). + +Let calT := seqIndT H L. +Local Notation calS := (seqIndD H L H 1). +Local Notation Ind1H := ('Ind[gval L, gval H] 1). +Let uniqS : uniq calS := seqInd_uniq _ _. + +Let h := #|H|%:R : algC. +Let e := #|L : H|%:R : algC. + +Let nsAL : A <| L. Proof. by have [] := ddA. Qed. +Let sLG : L \subset G. Proof. by have [] := ddA. Qed. +Let nsHL : H <| L. Proof. by rewrite -normalD1. Qed. +Let sHL := normal_sub nsHL. +Let nHL := normal_norm nsHL. + +Let nzh : h != 0 := neq0CG H. +Let nze : e != 0 := neq0CiG L H. +Let nzL : #|L|%:R != 0 := neq0CG L. + +Let eh : e * h = #|L|%:R. Proof. by rewrite -natrM mulnC Lagrange. Qed. + +Section InvDadeSeqInd. + +Variables (xi0 : 'CF(L)) (S : seq 'CF(L)) (chi : 'CF(G)). +Implicit Types xi mu : 'CF(L). + +Let d xi := xi 1%g / xi0 1%g. +Let psi xi := xi - d xi *: xi0. +Let c xi := '[(psi xi)^\tau, chi]. + +Let uc c xi mu := (c xi)^* * c mu / ('[xi] * '[mu]). +Let u c xi mu := uc c xi mu * ('[xi, mu] - xi 1%g * mu 1%g / (e * h)). + +(* This is Peterfalvi (7.7); it is stated using a bespoke concrete Prop so as *) +(* to encapsulate the coefficient definitions given above. *) +CoInductive is_invDade_seqInd_sum : Prop := + InvDadeSeqIndSum (c := c) (u := u c) of + (*a*) {in A, forall x, (chi^\rho) x = \sum_(xi <- S) (c xi)^* / '[xi] * xi x} + & (*b*) '[chi^\rho] = \sum_(xi <- S) \sum_(mu <- S) u xi mu. + +Lemma invDade_seqInd_sum : perm_eq calT (xi0 :: S) -> is_invDade_seqInd_sum. +Proof. +move=> defT; pose chi0 := \sum_(xi <- S) (c xi)^* / '[xi] *: xi. +have Txi0: xi0 \in calT by rewrite (perm_eq_mem defT) mem_head. +have sST : {subset S <= calT}. + by move=> xi Sxi; rewrite (perm_eq_mem defT) mem_behead. +have nz_xi01 : xi0 1%g != 0 by apply: seqInd1_neq0 Txi0. +have part_a: {in A, chi^\rho =1 chi0}. + pose phi := (chi^\rho - chi0) * '1_A. + have Aphi : phi \in 'CF(L, A) := mul_cfuni_on A _. + suffices: '[phi, chi^\rho - chi0] == 0; last clearbody phi. + rewrite -(eq_cfdotr Aphi (eq_mul_cfuni _ nsAL)) cfnorm_eq0 => /eqP phi0. + by move=> x Ax; rewrite -[chi0]add0r -phi0 cfunE eq_mul_cfuni ?cfunE ?subrK. + have{Aphi} [Hphi phi1]: phi \in 'CF(L, H) /\ phi 1%g = 0. + by move: Aphi; rewrite cfun_onD1 => /andP[-> /eqP]. + have{Hphi} def_phi: phi = e^-1 *: 'Ind ('Res[H] phi). + apply/cfunP=> y; have [Hy | notHy] := boolP (y \in H); last first. + by rewrite cfunE !(cfun_on0 _ notHy) ?mulr0 ?cfInd_normal. + rewrite cfunE cfIndE // mulrA -invfM eh. + apply: (canRL (mulKf nzL)); rewrite mulr_natl -sumr_const. + by apply: eq_bigr => z Lz; rewrite cfResE ?memJ_norm ?cfunJ ?(subsetP nHL). + have{def_phi} Tphi: phi \in <<calT>>%VS. + rewrite def_phi memvZ // ['Res _]cfun_sum_cfdot linear_sum. + apply: memv_suml => i _; rewrite linearZ memvZ ?memv_span //=. + by apply/seqIndP; exists i; rewrite ?inE. + have{Tphi} [z def_phi _] := free_span (seqInd_free nsHL _) Tphi. + have {phi def_phi phi1} ->: phi = \sum_(xi <- S) z xi *: psi xi. + rewrite def_phi (eq_big_perm _ defT) !big_cons /= 2!cfunE in phi1 *. + rewrite (canRL (mulfK nz_xi01) (canRL (addrK _) phi1)) add0r addrC mulNr. + rewrite sum_cfunE mulr_suml scaleNr scaler_suml -sumrB. + by apply: eq_bigr => xi _; rewrite cfunE -mulrA -scalerA -scalerBr. + rewrite cfdot_suml big1_seq //= => xi Sxi; have Txi := sST xi Sxi. + rewrite cfdotZl cfdotBr -invDade_reciprocity -/(c xi); last first. + rewrite cfun_onD1 !cfunE divfK // subrr eqxx andbT. + by rewrite memvB ?memvZ //= ((seqInd_on _) setT). + have [oSS /orthoPl o_xi0S]: pairwise_orthogonal S /\ orthogonal xi0 S. + have:= seqInd_orthogonal nsHL setT; rewrite (eq_pairwise_orthogonal defT). + by rewrite /= -cat1s pairwise_orthogonal_cat => /and3P[]. + rewrite cfdotBl cfdotC cfproj_sum_orthogonal {oSS}// cfdotZl cfdot_sumr. + rewrite big1_seq ?mulr0 => [|mu Smu]; last by rewrite cfdotZr o_xi0S ?mulr0. + by rewrite subr0 divfK ?(cfnorm_seqInd_neq0 _ Txi) // conjCK subrr mulr0. +split=> [x /part_a-> | ]. + by rewrite sum_cfunE; apply: eq_bigr => ?; rewrite cfunE. +rewrite (cfdotEl _ (invDade_on ddA _)) mulrC mulr_suml. +pose F xi mu x := uc c xi mu * (xi x * (mu x)^*) / #|L|%:R. +transitivity (\sum_(x in A) \sum_(xi <- S) \sum_(mu <- S) F xi mu x). + apply: eq_bigr => x Ax; rewrite part_a // sum_cfunE -mulrA mulr_suml. + apply: eq_bigr => xi _; rewrite mulrA -mulr_suml rmorph_sum; congr (_ * _). + rewrite mulr_sumr; apply: eq_bigr => mu _; rewrite !cfunE (cfdotC mu). + rewrite -{1}[mu x]conjCK -fmorph_div -rmorphM conjCK -4!mulrA 2!(mulrCA _^-1). + by rewrite (mulrA _^-1) -invfM 2!(mulrCA (xi x)) mulrA 2!(mulrA _^*). +rewrite exchange_big; apply: eq_bigr => xi _; rewrite exchange_big /=. +apply: eq_big_seq => mu Smu; have Tmu := sST mu Smu. +rewrite /u eh (cfdotEr _ (seqInd_on nsHL Tmu)) (mulrC _^-1) -mulrBl mulrA. +rewrite -mulr_suml -mulr_sumr (big_setD1 1%g (group1 H)) /=; congr (_ * _ * _). +by rewrite addrC conj_Cnat ?addKr // (Cnat_seqInd1 Tmu). +Qed. + +End InvDadeSeqInd. + +Notation "1" := (1 : 'CF(_)) : classfun_scope. + +(* This is Peterfalvi (7.8). *) +(* Our version is slightly stronger because we state the nontriviality of S *) +(* directly than through the coherence premise (see the discussion of (5.2)). *) +Lemma Dade_Ind1_sub_lin (nu : {additive 'CF(L) -> 'CF(G)}) zeta : + coherent_with calS L^# (Dade ddA) nu -> (1 < size calS)%N -> + zeta \in irr L -> zeta \in calS -> zeta 1%g = e -> + let beta := (Ind1H - zeta)^\tau in let calSnu := map nu calS in + let sumSnu := \sum_(xi <- calS) xi 1%g / e / '[xi] *: nu xi in + [/\ (*a1*) [/\ orthogonal calSnu 1%CF, '[beta, 1] = 1 & beta \in 'Z[irr G]], + exists2 Gamma : 'CF(G), + (*a2*) [/\ orthogonal calSnu Gamma, '[Gamma, 1] = 0 + & exists2 a, a \in Cint & beta = 1 - nu zeta + a *: sumSnu + Gamma] + & (*b*) e <= (h - 1) / 2%:R -> + '[(nu zeta)^\rho] >= 1 - e / h /\ '[Gamma] <= e - 1 + & (*c*) {in irr G, forall chi : 'CF(G), orthogonal calSnu chi -> + [/\ {in A, forall x, chi^\rho x = '[beta, chi]} + & '[chi^\rho] = #|A|%:R / #|L|%:R * '[beta, chi] ^+ 2]}]. +Proof. +move=> [[Inu Znu] nu_tau] nt_calS /irrWnorm Nzeta1 Szeta zeta1. +set mu := _ - _ => beta calSnu sumSnu; pose S1 := rem zeta calS. +have defS: perm_eq calS (zeta :: S1) := perm_to_rem Szeta. +have defZS: 'Z[calS, L^#] =i 'Z[calS, A] by apply: zcharD1_seqInd. +have S1P xi: xi \in S1 -> xi != zeta /\ xi \in calS. + by rewrite mem_rem_uniq // => /andP[]. +have defT: perm_eql calT [:: Ind1H, zeta & S1]. + apply/perm_eqlP; have Tind1: Ind1H \in calT := seqIndT_Ind1 H L. + by rewrite (perm_eqlP (perm_to_rem Tind1)) perm_cons -seqIndC1_rem. +have mu_vchar: mu \in 'Z[irr L, A] := cfInd1_sub_lin_vchar nsHL Szeta zeta1. +have beta_vchar: beta \in 'Z[irr G] by apply: Dade_vchar. +have [mu_on beta_on] := (zchar_on mu_vchar, zchar_on beta_vchar). +have{nt_calS} ntS1: (size S1 > 0)%N by rewrite size_rem // -subn1 subn_gt0. +case defS1: S1 ntS1 => // [phi S2] _. +have /S1P[neq_phi Sphi]: phi \in S1 by rewrite defS1 mem_head. +have nz_phi1: phi 1%g != 0 by rewrite (seqInd1_neq0 nsHL Sphi). +have NatS1e xi (Sxi : xi \in calS) := dvd_index_seqInd1 nsHL Sxi. +have oS1: {in calS, forall psi, '[psi, 1] = 0} by apply: seqInd_ortho_1. +have oS1H: {in calS, forall psi, '[psi, Ind1H] = 0} by apply: seqInd_ortho_Ind1. +have InuS: {in calS &, isometry nu} by apply: sub_in2 Inu; exact: seqInd_zcharW. +have ZnuS xi (Sxi : xi \in calS) := Znu xi (seqInd_zcharW Sxi). +have S_Se xi (Sxi : xi \in calS) := seqInd_sub_lin_vchar nsHL Szeta zeta1 Sxi. +have oSnu1: orthogonal calSnu 1%CF. + have dotSnu1 psi : psi \in calS -> '[nu psi, 1] = psi 1%g / e * '[nu zeta, 1]. + move=> Spsi; apply/eqP; rewrite -subr_eq0 -cfdotZl -cfdotBl. + rewrite -raddfZ_Cnat ?NatS1e // -raddfB; have Spi := S_Se _ Spsi. + rewrite nu_tau ?defZS // invDade_reciprocity ?(zchar_on Spi) //. + rewrite invDade_cfun1 (eq_cfdotr (zchar_on Spi) (eq_cfuni nsAL)). + by rewrite cfdotBl cfdotZl !oS1 // mulr0 subr0. + suffices oz1: '[nu zeta, 1] = 0. + by apply/orthoPr=> _ /mapP[psi Spsi ->]; rewrite dotSnu1 // oz1 mulr0. + have norm_nu_zeta : '[nu zeta] = 1 by rewrite InuS // irrWnorm. + have [et [t defz]] := vchar_norm1P (ZnuS _ Szeta) norm_nu_zeta. + rewrite defz cfdotZl -{1}irr0 cfdot_irr mulr_natr mulrb; case: eqP => // t0. + have /eqP/idPn[] := seqInd_ortho nsHL Sphi Szeta neq_phi. + rewrite -InuS // defz t0 cfdotZr irr0 dotSnu1 // mulrCA -irr0 -t0. + by rewrite -cfdotZr -defz norm_nu_zeta mulr1 mulf_neq0 ?invr_eq0. +have dot_beta_1: '[beta, 1] = 1. + rewrite invDade_reciprocity // invDade_cfun1 (eq_cfdotr _ (eq_cfuni nsAL)) //. + by rewrite cfdotBl -Frobenius_reciprocity cfRes_cfun1 ?cfnorm1 ?oS1 ?subr0. +have o_beta1: '[beta - 1, 1] = 0 by rewrite cfdotBl dot_beta_1 cfnorm1 subrr. +have [X SnuX [Gamma [def_beta1 _ oSnuG]]]:= orthogonal_split calSnu (beta - 1). +have oG1: '[Gamma, 1] = 0. + rewrite -(addKr X Gamma) -def_beta1 addrC cfdotBl o_beta1. + by rewrite (span_orthogonal oSnu1) ?subr0 // memv_span ?mem_head. +have oSS: pairwise_orthogonal calS by apply: seqInd_orthogonal. +have oSnuS: pairwise_orthogonal calSnu by apply: map_pairwise_orthogonal. +have [a_ def_a defX] := orthogonal_span oSnuS SnuX. +have{def_a} def_a: {in calS, forall xi, a_ (nu xi) = '[beta, nu xi] / '[xi]}. + move=> xi Sxi; rewrite (canRL (subrK 1) def_beta1) !cfdotDl def_a InuS //. + by rewrite (cfdotC 1) (orthoPl oSnuG) ?(orthoPr oSnu1) ?map_f ?conjC0 ?addr0. +pose a := '[beta, nu zeta] + 1; have Z1 := Cint1. +have{Z1} Za: a \in Cint by rewrite rpredD ?Cint_cfdot_vchar // ZnuS. +have {a_ def_a defX} defX: X = - nu zeta + a *: sumSnu. + rewrite linear_sum defX big_map !(eq_big_perm _ defS) !big_cons /= addrCA. + rewrite def_a // Nzeta1 !divr1 zeta1 divff // scalerDl !scale1r addrA. + rewrite addrK; congr (_ + _); apply: eq_big_seq => xi /S1P[neq_xi Sxi]. + rewrite def_a // scalerA mulrA mulrDl mul1r; congr (_ / _ *: _). + rewrite mulrC -(conj_Cnat (NatS1e _ Sxi)) -cfdotZr -raddfZ_Cnat ?NatS1e //. + rewrite addrC; apply: canRL (subrK _) _; rewrite -!raddfB /= -/e. + have Spi := S_Se xi Sxi; rewrite nu_tau ?defZS //. + rewrite Dade_isometry ?(zchar_on Spi) // cfdotC cfdotBl cfdotZl !cfdotBr. + by rewrite !oS1H ?(seqInd_ortho _ Sxi) // Nzeta1 subr0 !add0r mulrN1 opprK. +have Ind1H1: Ind1H 1%g = e by rewrite cfInd1 // cfun11 mulr1. +split=> // [ | chi /irrP[t def_chi] o_chiSnu]. + rewrite (canRL (subrK 1) def_beta1) defX addrC 2!addrA. + exists Gamma; first by rewrite orthogonal_sym; split; last exists a. + move=> lt_e_h2; pose v := h^-1; pose u := e^-1 * (1 - v); set w := 1 - e / h. + have hu: h * u = e^-1 * (h - 1) by rewrite mulrCA (mulrBr h) mulr1 divff. + have ->: '[(nu zeta)^\rho] = u * a ^+ 2 - v * a *+ 2 + w. + have defT1: perm_eq calT [:: phi, Ind1H, zeta & S2]. + by rewrite defT defS1 (perm_catCA [::_ ; _] phi). + have [c ua _ ->] := invDade_seqInd_sum (nu zeta) defT1. + have def_c xi: xi \in calS -> c xi = '[xi, zeta]. + move=> S2xi; rewrite /c mulrC -{1}[xi]scale1r -(mulVf nz_phi1) -!scalerA. + rewrite -scalerBr linearZ cfdotZl /=; set pi := _ - _. + have Spi: pi \in 'Z[calS, A] by apply: sub_seqInd_zchar. + rewrite -nu_tau ?defZS // Inu ?(zcharW Spi) ?seqInd_zcharW //. + by rewrite cfdotBl !cfdotZl (seqInd_ortho _ Sphi) // mulr0 subr0 mulKf. + have c2: c zeta = 1 by rewrite def_c. + have c1: c Ind1H = a. + by rewrite /a -c2 -cfdotDl -linearD !addrA subrK zeta1 -Ind1H1. + have{def_c} c3 xi: xi \in S2 -> c xi = 0. + move=> S2xi; have /S1P[neq_xi Sxi]: xi \in S1 by rewrite defS1 mem_behead. + by rewrite def_c // (seqInd_ortho _ Sxi). + rewrite !big_cons; have kill0 := (mul0r, mulr0, big1, conjC0). + rewrite !big1_seq /ua; try by move=> psi /c3->; do 2?rewrite ?kill0 => *. + rewrite !addr0 c1 c2 Nzeta1 cfnorm_Ind_cfun1 // -/e Ind1H1 zeta1 conjC1. + rewrite cfdotC (seqInd_ortho_Ind1 _ _ Szeta) // !kill0 sub0r !mulrN !mulr1. + rewrite divr1 !mul1r !invfM mulrBr !mulrA !mulfK ?divfK // -/w. + rewrite aut_Cint // -[_ / h]mulrA -{1}[e^-1]mulr1 -2!mulrBr -/u -/v. + by rewrite mulrC mulrA addrA (mulrC v) -[_ - _]addrA -opprD. + have ->: '[Gamma] = e - 1 - h * (u * a ^+ 2 - v * a *+ 2). + have /(canLR (addrK 1)) <-: '[beta] = e + 1. + rewrite Dade_isometry // cfnormBd ?cfnorm_Ind_cfun1 ?Nzeta1 //. + by rewrite cfdotC (seqInd_ortho_Ind1 _ _ Szeta) ?conjC0. + rewrite -[beta](subrK 1) cfnormDd; last first. + by rewrite cfdotBl dot_beta_1 cfnorm1 subrr. + rewrite cfnorm1 addrK def_beta1 (addrC X) cfnormDd; last first. + by rewrite (span_orthogonal oSnuG) // memv_span ?mem_head. + do 2!apply: canRL (addrK _) _; rewrite -addrA; congr (_ + _). + rewrite defX (addrC (- nu _)) cfnormB cfnormZ Cint_normK // InuS //. + rewrite cfdotZl cfproj_sum_orthogonal // Nzeta1 zeta1 divff // divr1. + rewrite !mulr1 aut_Cint // mulrBr mulrDr mulVKf // addrAC. + rewrite mulrA mulrC hu -[e^-1](divfK nze) -expr2; congr (_ * _ - _ + 1). + rewrite -mulrA -sum_seqIndC1_square // mulr_sumr cfnorm_sum_orthogonal //. + apply: eq_big_seq => xi Sxi. + have [nz_xi Nxi1] := (cfnorm_seqInd_neq0 nsHL Sxi, Cnat_seqInd1 Sxi). + rewrite (normr_idP _) ?mulr_ge0 ?invr_ge0 ?ler0n ?cfnorm_ge0 ?Cnat_ge0 //. + by rewrite mulrCA !exprMn ['[xi]]lock !mulrA divfK // -lock. + apply/andP; rewrite -subr_ge0 addrK andbC -subr_ge0 addrC opprB subrK. + rewrite pmulr_rge0 ?gt0CG // andbb -mulr_natr (mulrAC v). + have v_ge0: 0 <= v by [rewrite invr_ge0 ler0n]; have L_gt0 := gt0CG L. + have Lu: #|L|%:R * u = h - 1 by rewrite -eh -mulrA hu mulVKf. + have h1ge0: 0 <= h - 1 by rewrite subr_ge0 ler1n cardG_gt0. + have{h1ge0} u_ge0: 0 <= u by rewrite -Lu pmulr_rge0 in h1ge0. + have [a_le0 | ] := boolP (a <= 0). + by rewrite -mulrN -sqrrN addr_ge0 ?(u_ge0, mulr_ge0) ?oppr_ge0 ?ler0n. + rewrite -real_ltrNge ?Creal_Cint // ltr_def => /andP[]. + move/(norm_Cint_ge1 Za)=> a_ge1 a_ge0; rewrite mulrA -mulrBl. + rewrite (normr_idP _) // -(@mulVf _ 2%:R) ?pnatr_eq0 // in a_ge1. + rewrite mulr_ge0 // subr_ge0 (ler_trans _ (ler_wpmul2l u_ge0 a_ge1)) // mulrA. + by rewrite ler_wpmul2r ?ler0n // -(ler_pmul2l L_gt0) mulrA Lu -eh mulfK. +have Zchi: chi \in 'Z[irr G] by rewrite def_chi irr_vchar. +have def_chi0: {in A, chi^\rho =1 (fun _ => '[beta, chi])}. + have defT1: perm_eq calT [:: zeta, Ind1H & S1]. + by rewrite defT (perm_catCA Ind1H zeta). + move=> x Ax; have [_ Hx] := setD1P Ax. + have [c _ -> // _] := invDade_seqInd_sum chi defT1. + rewrite big_cons big1_seq ?addr0 /c => [|xi /S1P[neq_xi /= Sxi]]; last first. + rewrite zeta1 -nu_tau ?defZS ?S_Se // raddfB cfdotBl raddfZ_Cnat ?NatS1e //. + by rewrite cfdotZl !(orthoPr o_chiSnu) ?map_f // mulr0 subr0 conjC0 !mul0r. + rewrite Ind1H1 zeta1 divff // scale1r -/beta aut_Cint ?Cint_cfdot_vchar //. + by rewrite cfnorm_Ind_cfun1 ?cfInd_cfun1 // cfunE cfuniE // Hx mulr1 divfK. +split=> //; rewrite -mulrA mulrCA cfnormE_invDade; congr (_ * _). +rewrite mulr_natl -sumr_const; apply: eq_bigr => _ /def_chi0->. +by rewrite Cint_normK ?Cint_cfdot_vchar. +Qed. + +End Dade_seqIndC1. + +(* The other results of the section are specific to groups of odd order. *) +Hypothesis oddG : odd #|G|. + +(* We explicitly factor out several intermediate results from the proof of *) +(* (7.9) that are reused throughout the proof (including in (7.10) below). *) + +Import ssrint. +Lemma cfdot_real_vchar_even phi psi : + phi \in 'Z[irr G] /\ cfReal phi -> psi \in 'Z[irr G] /\ cfReal psi -> + (2 %| '[phi, psi])%C = (2 %| '[phi, 1])%C || (2 %| '[psi, 1])%C. +Proof. +move=> [Zphi Rphi] [Zpsi Rpsi]; rewrite cfdot_vchar_r // (bigD1 (0 : 'I__)) //=. +rewrite addrC -irr0 (bigID [pred i | conjC_Iirr i < i]%N) /=. +set a1 := \sum_(i | _) _; set a2 := \sum_(i | _) _; suffices ->: a1 = a2. + rewrite -mulr2n -mulr_natr (rpredDl _ (dvdC_mull _ _)) //; last first. + by rewrite rpred_sum // => i; rewrite rpredM ?Cint_cfdot_vchar_irr. + have /CintP[m1 ->] := Cint_cfdot_vchar_irr 0 Zphi. + have /CintP[m2 ->] := Cint_cfdot_vchar_irr 0 Zpsi. + rewrite [m1]intEsign [m2]intEsign !rmorphMsign mulrACA -!mulrA !rpredMsign. + by rewrite -natrM !(dvdC_nat 2) Euclid_dvdM. +rewrite /a2 (reindex_inj (inv_inj (@conjC_IirrK _ _))) /=. +apply: eq_big => [t | t _]; last first. + by rewrite !conjC_IirrE !cfdot_real_conjC ?aut_Cint ?Cint_cfdot_vchar_irr. +rewrite (inv_eq (@conjC_IirrK _ _)) conjC_IirrK -leqNgt ltn_neqAle val_eqE. +rewrite -!(inj_eq irr_inj) !conjC_IirrE irr0 cfConjC_cfun1 odd_eq_conj_irr1 //. +by rewrite andbA andbb. +Qed. + +Section DisjointDadeOrtho. + +Variables (L1 L2 H1 H2 : {group gT}). +Let A1 := H1^#. +Let A2 := H2^#. + +Hypothesis ddA1 : Dade_hypothesis G L1 A1. +Hypothesis ddA2 : Dade_hypothesis G L2 A2. +Let Atau1 := Dade_support ddA1. +Let tau1 := Dade ddA1. +Let Atau2 := Dade_support ddA2. +Let tau2 := Dade ddA2. + +Hypothesis disjointA : [disjoint Atau1 & Atau2]. + +Lemma disjoint_Dade_ortho phi psi : '[tau1 phi, tau2 psi] = 0. +Proof. +rewrite (cfdot_complement (Dade_cfunS _ _)) ?(cfun_onS _ (Dade_cfunS _ _)) //. +by rewrite subsetD disjoint_sym Dade_support_sub. +Qed. + +Let odd_Dade_context L H : Dade_hypothesis G L H^# -> H <| L /\ odd #|L|. +Proof. by case=> nsAL sLG _ _ _; rewrite -normalD1 (oddSg sLG). Qed. + +(* This lemma encapsulates uses of lemma (4.1) in sections 7 and 14. *) +Lemma disjoint_coherent_ortho nu1 nu2 chi1 chi2 : + let S1 := seqIndD H1 L1 H1 1 in coherent_with S1 L1^# tau1 nu1 -> + let S2 := seqIndD H2 L2 H2 1 in coherent_with S2 L2^# tau2 nu2 -> + chi1 \in irr L1 -> chi1 \in S1 -> chi2 \in irr L2 -> chi2 \in S2 -> + '[nu1 chi1, nu2 chi2] = 0. +Proof. +move=> S1 cohS1 S2 cohS2 /irrP[i1 ->] Schi1 /irrP[i2 ->] Schi2. +have [[nsHL1 oddL1] [[Inu1 Znu1] nu1tau]] := (odd_Dade_context ddA1, cohS1). +have [[nsHL2 oddL2] [[Inu2 Znu2] nu2tau]] := (odd_Dade_context ddA2, cohS2). +pose nu_chiC L (nu : 'CF(L) -> 'CF(G)) i := map nu ('chi_i :: ('chi_i)^*)%CF. +have: orthonormal (nu_chiC L1 nu1 i1) && orthonormal (nu_chiC L2 nu2 i2). + rewrite /orthonormal /= !andbT !Inu1 ?Inu2 ?seqInd_zcharW ?cfAut_seqInd //=. + rewrite !cfnorm_conjC !cfnorm_irr (seqInd_conjC_ortho _ _ _ Schi1) ?eqxx //=. + by rewrite (seqInd_conjC_ortho _ _ _ Schi2). +move/orthonormal_vchar_diff_ortho=> -> //. + by split; apply/allP; rewrite /= !(Znu1, Znu2) ?seqInd_zcharW ?cfAut_seqInd. +rewrite -!raddfB !(nu1tau, nu2tau) ?zcharD1_seqInd ?seqInd_sub_aut_zchar //. +by rewrite !Dade1 disjoint_Dade_ortho !eqxx. +Qed. + +(* This is Peterfalvi (7.9). *) +(* We have inlined Hypothesis (7.4) because although it is readily available *) +(* for the proof of (7.10), it would be inconvenient to establish in (14.4). *) +(* Note that our Delta corresponds to Delta - 1 in the Peterfalvi proof. *) +Let beta L H ddA zeta := @Dade _ G L H^# ddA ('Ind[L, H] 1 - zeta). +Lemma Dade_sub_lin_nonorthogonal nu1 nu2 zeta1 zeta2 : + let S1 := seqIndD H1 L1 H1 1 in coherent_with S1 L1^# tau1 nu1 -> + let S2 := seqIndD H2 L2 H2 1 in coherent_with S2 L2^# tau2 nu2 -> + zeta1 \in irr L1 -> zeta1 \in S1 -> zeta1 1%g = #|L1 : H1|%:R -> + zeta2 \in irr L2 -> zeta2 \in S2 -> zeta2 1%g = #|L2 : H2|%:R -> + '[beta ddA1 zeta1, nu2 zeta2] != 0 \/ '[beta ddA2 zeta2, nu1 zeta1] != 0. +Proof. +move=> S1 cohS1 S2 cohS2 irr_zeta1 Szeta1 zeta1_1 irr_zeta2 Szeta2 zeta2_1. +apply/nandP; pose Delta ddA nu zeta := beta ddA zeta + nu zeta. +have Delta_context L H (A := H^#) ddA (tau := Dade ddA) nu zeta : + let S := seqIndD H L H 1 in coherent_with S L^# tau nu -> + zeta \in irr L -> zeta \in S -> zeta 1%g = #|L : H|%:R -> + let D := Delta L H ddA nu zeta in '[D, 1] = 1 /\ D \in 'Z[irr G] /\ cfReal D. +- move=> S cohS irr_zeta Szeta zeta_1 D. + have [[nsHL oddL] [[_ Znu] nu_tau]] := (odd_Dade_context ddA, cohS). + have ntS: (size S > 1)%N by apply: seqInd_nontrivial Szeta. + have [[nuS1_0 beta1_1 Zbeta] _ _] := + Dade_Ind1_sub_lin cohS ntS irr_zeta Szeta zeta_1. + rewrite cfdotDl {}beta1_1 {nuS1_0}(orthoPr nuS1_0) ?map_f // addr0. + rewrite rpredD ?{}Znu ?seqInd_zcharW {Zbeta}// /cfReal; do !split=> //. + rewrite rmorphD /= -subr_eq0 opprD addrAC addrA -addrA addr_eq0 opprD. + rewrite (cfConjC_Dade_coherent cohS) // opprK -Dade_conjC -!raddfB /=. + rewrite nu_tau ?zcharD1_seqInd ?seqInd_sub_aut_zchar //=. + by rewrite rmorphB /= conj_cfInd cfConjC_cfun1 opprB addrC addrA subrK. +have: ~~ (2 %| '[Delta L1 H1 ddA1 nu1 zeta1, Delta L2 H2 ddA2 nu2 zeta2])%C. + have /Delta_context/(_ irr_zeta1 Szeta1 zeta1_1)[Delta1_1 ZR_Delta1] := cohS1. + have /Delta_context/(_ irr_zeta2 Szeta2 zeta2_1)[Delta2_1 ZR_Delta2] := cohS2. + by rewrite cfdot_real_vchar_even // Delta1_1 Delta2_1 (dvdC_nat 2 1). +rewrite cfdotDl !cfdotDr disjoint_Dade_ortho // add0r addrC cfdotC. +apply: contra => /andP[/eqP-> /eqP->]; rewrite conjC0 add0r addr0. +by rewrite (disjoint_coherent_ortho cohS1 cohS2) ?dvdC0. +Qed. + +End DisjointDadeOrtho. + +(* A numerical fact used in Sections 7 and 14. *) +Lemma odd_Frobenius_index_ler (R : numFieldType) (K L : {group gT}) : + odd #|L| -> [Frobenius L with kernel K] -> + #|L : K|%:R <= (#|K|%:R - 1) / 2%:R :> R. +Proof. +move=> oddL /existsP[H frobL]; rewrite ler_pdivl_mulr ?ltr0n // ler_subr_addl. +have ->: #|L : K| = #|H| by have [/index_sdprod] := Frobenius_context frobL. +by rewrite -natrM -mulrS ler_nat muln2 (ltn_odd_Frobenius_ker frobL). +Qed. + +(* This final section factors the assumptions common to (7.10) and (7.11). *) +(* We add solvability of the Frobenius groups, so as not to rely on the *) +(* theorem of Thompson asserting the nilpotence of Frobenius kernels. *) + +Section CoherentFrobeniusPartition. + +Variables (k : nat) (L H E : 'I_k -> {group gT}). + +Local Notation A i := (gval (H i))^#. +Let h_ i : algC := #|H i|%:R. +Let e_ i : algC := #|L i : H i|%:R. +Let G0 := G :\: \bigcup_(i < k) class_support (H i)^# G. + +Hypothesis k_ge2: (k >= 2)%N. + +(*a*) Hypothesis frobeniusL_G : + forall i, [/\ L i \subset G, solvable (L i) & [Frobenius L i = H i ><| E i]]. + +(*b*) Hypothesis normedTI_A : forall i, normedTI (A i) G (L i). + +(*c*) Hypothesis card_coprime : forall i j, i != j -> coprime #|H i| #|H j|. + +(* A numerical fact that is used in both (7.10) and (7.11) *) +Let e_bounds i : 1 < e_ i /\ e_ i <= (h_ i - 1) / 2%:R. +Proof. +have [/oddSg/(_ oddG)oddL _ frobL] := frobeniusL_G i. +rewrite ltr1n odd_Frobenius_index_ler ?(FrobeniusWker frobL) //. +by have [/index_sdprod <-] := Frobenius_context frobL; rewrite cardG_gt1. +Qed. + +(* This is Peterfalvi (7.10). *) +Lemma coherent_Frobenius_bound : exists i, let e := e_ i in let h := h_ i in + (e - 1) * ((h - 2%:R * e - 1) / (e * h) + 2%:R / (h * (h + 2%:R))) + <= (#|G0|%:R - 1) / #|G|%:R. +Proof. +have [sLG solL frobL] := all_and3 frobeniusL_G. +have oddL i := oddSg (sLG i) oddG. +have /all_and2[nsHL ntH] i: H i <| L i /\ H i :!=: 1%g. + by case/Frobenius_context: (frobL i) => /sdprod_context[]. +have sHL i: H i \subset L i by case/andP: (nsHL i). +pose DH i := @Dade_signalizer gT G (L i) (A i). +have /fin_all_exists[ddA DH1] i: exists dd, {in A i, forall a, DH i dd a = 1%G}. + have /Dade_normedTI_P[|ddAi _] := normedTI_A i; last by exists ddAi. + by apply: normedTI_Dade => //; rewrite setSD // (subset_trans (sHL i)). +pose tau i := Dade (ddA i); pose rho i := invDade (ddA i). +pose Atau i := Dade_support (ddA i). +have defAtau i: Atau i = class_support (A i) G. + rewrite class_supportEl; apply: eq_bigr => x Ax. + by rewrite /Dade_support1 -/(DH i) DH1 // mul1g class_support_set1l. +have disjoint_Atau i j : i != j -> [disjoint Atau i & Atau j]. + move=> neq_ij; rewrite !defAtau !class_supportEr -setI_eq0 big_distrlr /=. + rewrite pair_big big1 // => [[x y] _] /=; apply/eqP. + by rewrite !conjD1g -setDIl setD_eq0 coprime_TIg // !cardJg card_coprime. +have{defAtau} defG0: G0 = G :\: \bigcup_i Atau i. + by congr (_ :\: _); apply: eq_bigr => i; rewrite defAtau. +pose S i := seqIndD (H i) (L i) (H i) 1. +have irrS i: {subset S i <= irr (L i)}. + move=> _ /seqIndC1P[t nz_t ->]; rewrite irr_induced_Frobenius_ker //. + exact: FrobeniusWker (frobL i). +have /fin_all_exists[r lin_r] i: exists r, 'chi_r \in S i /\ 'chi_r 1%g = e_ i. + have lt1Hi: [1] \proper H i by rewrite proper1G. + have solHi := solvableS (sHL i) (solL i). + have [xi Sxi lin_xi] := exists_linInd (nsHL i) solHi lt1Hi (normal1 _). + by have /irrP[r def_xi] := irrS i xi Sxi; exists r; rewrite -def_xi. +have{lin_r} [Sr r1] := all_and2 lin_r. +have ntS i: (size (S i) > 1)%N by apply: seqInd_nontrivial (mem_irr _) (Sr i). +have /fin_all_exists[nu cohS] i: coherent (S i) (L i)^# 'Ind[G, L i]. + have [[[frobLi tiAiL] sLiG] oddLi] := (frobL i, normedTI_A i, sLG i, oddL i). + have [defLi ntHi ntEi _ _] := Frobenius_context frobLi. + have{ntEi} nilHi: nilpotent (H i) by apply: (Frobenius_sol_kernel_nil frobLi). + exact: Sibley_coherence (or_introl _ frobLi). +have{cohS} [/all_and2[Inu Znu] nu_Ind] := all_and2 cohS. +have{DH DH1 nu_Ind} cohS i: coherent_with (S i) (L i)^# (tau i) (nu i). + split=> // phi Sphi; rewrite /tau nu_Ind ?Dade_Ind //. + by rewrite (@zchar_on _ _ (S i)) -?zcharD1_seqInd. +have n1S i xi: xi \in S i -> '[xi] = 1. + by case/irrS/irrP=> t ->; rewrite cfnorm_irr. +have n1Snu i xi: xi \in S i -> '[nu i xi] = 1. + by move=> Sxi; rewrite Inu ?n1S ?seqInd_zcharW. +have o_nu i j: i != j -> {in S i & S j, forall xi xj, '[nu i xi, nu j xj] = 0}. + move/disjoint_Atau/disjoint_coherent_ortho=> o_ij xi xj Sxi Sxj. + by rewrite o_ij ?irrS //; apply: cohS. +have /all_and2[nze nzh] i: e_ i != 0 /\ h_ i != 0 by rewrite neq0CiG neq0CG. +have h_gt1 i: 1 < h_ i by rewrite ltr1n cardG_gt1. +have eh i: e_ i * h_ i = #|L i|%:R by rewrite -natrM mulnC Lagrange. +have def_h1 i: h_ i - 1 = #|A i|%:R. + by rewrite /h_ (cardsD1 1%g) group1 addnC natrD addrK. +have [i1 min_i1]: {i1 | forall i, i != i1 -> h_ i1 + 2%:R <= h_ i}. + exists [arg min_(i < Ordinal k_ge2) #|H i|]; case: arg_minP => // i1 _ min_i1. + have oddH i: #|H i| = #|H i|./2.*2.+1. + by rewrite -{1}[#|H i|]odd_double_half (oddSg (sHL i)). + move=> i neq_i; rewrite -natrD ler_nat (oddH i) oddH addn2 -doubleS ltnS. + rewrite leq_double ltn_neqAle andbC half_leq ?min_i1 //=. + apply: contraTneq (card_coprime neq_i) => eqHii1. + by rewrite oddH -eqHii1 -oddH /coprime gcdnn -trivg_card1. +exists i1 => e h; set lhs := (e - 1) * _. +have nzh2: h + 2%:R != 0 by rewrite -natrD addnC pnatr_eq0. +have{lhs} ->: lhs = 1 - e / h - (h - 1) / (e * h) - (e - 1) / (h + 2%:R). + rewrite {}/lhs -{2}(addrK h 2%:R) !invfM (mulrBl _ _ h) mulVKf ?nzh //. + rewrite addrCA (addrC _ h) mulrCA mulrA addrA mulrBr; congr (_ - _). + rewrite mulfK // mulrDr addrAC addrC mulrC mulrBl -mulrA mulVKf ?nze //. + rewrite mulrC mulrBr mulrBl mul1r addrAC addrC addrA; congr (_ - _). + rewrite mulrCA mulVKf ?nze // addrCA mulrCA mulr_natl opprD addNKr. + by rewrite !mulrBl opprB addrA subrK divff ?nzh. +pose beta i := tau i ('Ind[L i, H i] 1 - 'chi_(r i)). +have betaP i := Dade_Ind1_sub_lin (cohS i) (ntS i) (mem_irr _) (Sr i) (r1 i). +pose chi i := nu i 'chi_(r i); pose c i j := '[beta i, chi j]. +have:= betaP i1; rewrite -/(S _) -/(tau _) -/(beta _) -/(chi _) -/(e_ _) -/e. +move=> [[oSnu1 o_beta1 Zbeta1] [Gamma [oSnuGamma oGamma1] [a Za def_beta1]]]. +have [_ lt_e_h2] := e_bounds i1; rewrite -/(rho _) -/(h_ _) -/h. +case/(_ lt_e_h2)=> min_rho1 maxGamma _ {lt_e_h2}. +pose calB := [set i | (i != i1) && (c i i1 == 0)]. +pose sumB := \sum_(i in calB) (h_ i - 1) / (e_ i * h_ i). +suffices{min_rho1} sumB_max: sumB <= (e - 1) / (h + 2%:R). + rewrite -subr_ge0 opprB addrCA -opprB subr_ge0; apply: ler_trans sumB_max. + rewrite -subr_ge0 opprB addrCA -(opprB _ sumB) subr_ge0. + have Zchi1: chi i1 \in 'Z[irr G] by rewrite Znu ?seqInd_zcharW ?Sr. + have [eps [t def_chi1]] := vchar_norm1P Zchi1 (n1Snu i1 'chi_(r i1) (Sr i1)). + pose sumG0 := \sum_(g in G0) `|'chi_t g| ^+ 2. + apply: (@ler_trans _ ((#|G0|%:R - sumG0) / #|G|%:R)); last first. + rewrite ler_pmul2r ?invr_gt0 ?gt0CG // ler_add2l ler_opp2. + rewrite [sumG0](bigD1 1%g) /=; last first. + rewrite !inE group1 andbT; apply/bigcupP=> [[i _]]. + by rewrite class_supportEr => /bigcupP[x _]; rewrite conjD1g !inE eqxx. + rewrite -[1]addr0 ler_add ?sumr_ge0 // => [|x _]; last first. + by rewrite -normrX normr_ge0. + have Zchit1: 'chi_t 1%g \in Cint by rewrite CintE Cnat_irr1. + by rewrite expr_ge1 ?normr_ge0 // norm_Cint_ge1 ?irr1_neq0. + pose ea i : algC := #|(H i)^#|%:R / #|L i|%:R. + apply: (@ler_trans _ (\sum_i ('[rho i 'chi_t] - ea i))); last first. + rewrite -subr_ge0 -opprB oppr_ge0 -mulNr opprB addrC mulrC. + by rewrite /sumG0 defG0 Dade_cover_inequality ?cfnorm_irr. + rewrite (bigID (mem calB)) /= addrC ler_add //. + rewrite -subr_ge0 opprK -big_split sumr_ge0 //= => i _. + by rewrite def_h1 eh subrK cfnorm_ge0. + rewrite (bigD1 i1) ?inE ?eqxx ?andbF //= -ler_subl_addl (@ler_trans _ 0) //. + rewrite opprB /ea -def_h1 -eh -/h -/e addrA subrK subr_le0. + by rewrite -(cfnorm_sign eps) -linearZ -def_chi1. + rewrite sumr_ge0 // => i; rewrite inE /c andbC => /andP[neq_i]. + rewrite neq_i subr_ge0 def_chi1 cfdotZr mulf_eq0 => /norP[_ not_o_beta_chi]. + have [[_ _ Zbeta_i] _ /(_ _ (mem_irr t))[|_ ->]] := betaP i. + apply/orthoPr=> _ /mapP[xi Sxi ->]; rewrite -['chi_t](signrZK eps). + by rewrite -def_chi1 cfdotZr o_nu ?mulr0 ?Sr. + rewrite -[ea i]mulr1 /ea ler_wpmul2l ?mulr_ge0 ?invr_ge0 ?ler0n //. + by rewrite -/(tau i) -/(beta i) sqr_Cint_ge1 ?Cint_cfdot_vchar_irr. +rewrite -(mulfK nzh2 sumB) -{2 3}natrD ler_wpmul2r ?invr_ge0 ?ler0n //. +apply: ler_trans maxGamma; rewrite mulr_suml. +pose phi i : 'CF(G) := \sum_(xi <- S i) xi 1%g / e_ i / '[xi] *: nu i xi. +have o_phi_nu i j xi: i != j -> xi \in S j -> '[phi i, nu j xi] = 0. + move/o_nu=> o_ij Sxi; rewrite cfdot_suml big1_seq //= => pi Spi. + by rewrite cfdotZl o_ij ?mulr0. +have o_phi i j: i != j -> '[phi i, phi j] = 0. + move/o_phi_nu=> o_ij; rewrite cfdot_sumr big1_seq // => xi Sxi. + by rewrite cfdotZr o_ij ?mulr0. +pose X : 'CF(G) := \sum_(i in calB) c i1 i *: phi i; pose Gamma1 := Gamma - X. +have ->: Gamma = Gamma1 + X by rewrite subrK. +have{betaP def_beta1} /cfnormDd->: '[Gamma1, X] = 0. + rewrite cfdot_sumr big1 // => i Bi; have [neq_i _] := setIdP Bi. + rewrite cfdotZr cfdot_sumr big1_seq ?mulr0 //= => xi Sxi. + apply/eqP; rewrite cfdotZr cfdotBl mulf_eq0; apply/pred2P; right. + rewrite cfdot_suml (bigD1 i) ?big1 //= => [|j /andP[_ neq_j]]; last first. + by rewrite cfdotZl o_phi_nu ?mulr0. + rewrite cfdotZl cfproj_sum_orthogonal ?seqInd_orthogonal //; last exact: Inu. + rewrite n1S // divr1 mulr1 addr0 mulrC -(canLR (addKr _) def_beta1). + rewrite !(cfdotDl, cfdotNl) cfdotZl o_nu ?o_phi_nu ?Sr 1?eq_sym // mulr0. + have[[/orthoPr oSnui_1 _ _] _ _] := betaP i; rewrite -/(S i) in oSnui_1. + rewrite cfdotC oSnui_1 ?map_f // conjC0 !(add0r, oppr0). + have Nxie: xi 1%g / e_ i \in Cnat by apply: dvd_index_seqInd1 _ Sxi. + rewrite -(conj_Cnat Nxie) // -cfdotZr -raddfZ_Cnat // -!raddfB /=. + have [_ Dnu] := cohS i. + rewrite Dnu ?zcharD1_seqInd ?seqInd_sub_lin_vchar ?Sr ?r1 //. + by rewrite disjoint_Dade_ortho ?disjoint_Atau 1?eq_sym. +rewrite -subr_ge0 cfdot_sumr -addrA -sumrB addr_ge0 ?cfnorm_ge0 //. +rewrite sumr_ge0 // => i Bi; have [neq_i ci1_0] := setIdP Bi. +have n_phi: '[phi i] = (h_ i - 1) / e_ i. + rewrite cfnorm_sum_orthogonal ?seqInd_orthogonal //; last exact: Inu. + rewrite -[_ - 1](mulKf (nze i)) -sum_seqIndC1_square // -/(S i) mulrAC. + rewrite -invfM mulrC mulr_suml; apply: eq_big_seq => _ /irrS/irrP[t ->]. + rewrite cfnorm_irr !divr1 mulr1 -expr2 -exprVn -exprMn. + by rewrite (normr_idP _) // mulr_ge0 ?invr_ge0 ?ler0n // ltrW ?irr1_gt0. +rewrite subr_ge0 cfdotZr cfdot_suml (bigD1 i) //=. +rewrite big1 ?addr0 => [|j /andP[_ ne_j]]; last by rewrite cfdotZl o_phi ?mulr0. +rewrite cfdotZl invfM 2!mulrA -n_phi -[_ * _]mulrA mulrC. +rewrite ler_wpmul2r ?cfnorm_ge0 // (@ler_trans _ 1) //. + by rewrite -{2}(mulVf (nzh i)) ler_wpmul2l ?invr_ge0 ?ler0n ?min_i1. +rewrite mulrC -normCK expr_ge1 ?normr_ge0 // norm_Cint_ge1 //. + rewrite Cint_cfdot_vchar ?Znu ?seqInd_zcharW ?Sr //. +suffices /orP: c i i1 != 0 \/ c i1 i != 0 by rewrite ci1_0. +apply: Dade_sub_lin_nonorthogonal; rewrite ?mem_irr ?Sr ?r1 //; try exact: cohS. +exact: disjoint_Atau. +Qed. + +(* This is Peterfalvi (7.11). *) +Theorem no_coherent_Frobenius_partition : G0 != 1%G. +Proof. +have [i] := coherent_Frobenius_bound; apply: contraTneq => ->. +have [] := e_bounds i; set e := e_ i; set h := h_ i => e_gt1 le_e_h2. +rewrite cards1 subrr mul0r ltr_geF // pmulr_rgt0 ?subr_gt0 // ltr_paddl //. + rewrite ?(mulr_ge0, invr_ge0) ?ler0n // addrAC subr_ge0. + by rewrite -[_ - 1](@divfK _ 2%:R) ?pnatr_eq0 // mulrC ler_wpmul2r ?ler0n. +by rewrite -natrD addnC ?(pmulr_rgt0, invr_gt0) ?ltr0n. +Qed. + +End CoherentFrobeniusPartition. + +End Seven. + diff --git a/mathcomp/odd_order/PFsection8.v b/mathcomp/odd_order/PFsection8.v new file mode 100644 index 0000000..72a0d00 --- /dev/null +++ b/mathcomp/odd_order/PFsection8.v @@ -0,0 +1,1128 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg poly finset center. +Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +Require Import gfunctor gproduct cyclic commutator nilpotent pgroup. +Require Import sylow hall abelian maximal frobenius. +Require Import matrix mxalgebra mxrepresentation vector. +Require Import BGsection1 BGsection3 BGsection7 BGsection10. +Require Import BGsection14 BGsection15 BGsection16. +Require ssrnum. +Require Import algC classfun character inertia vcharacter. +Require Import PFsection1 PFsection2 PFsection3 PFsection4 PFsection5. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 8: Structure of a Minimal Simple *) +(* Group of Odd Order. Actually, most Section 8 definitions can be found in *) +(* BGsection16, which holds the conclusions of the Local Analysis part of the *) +(* proof, as the B & G text has been adapted to fit the usage in Section 8. *) +(* Most of the definitions of Peterfalvi Section 8 are covered in BGsection7, *) +(* BGsection15 and BGsection16; we only give here: *) +(* FT_Pstructure S T defW <-> the groups W, W1, W2, S, and T satisfy the *) +(* conclusion of Theorem (8.8)(b), in particular, S and T *) +(* are of type P, S = S^(1) ><| W1, and T = T^`(1) ><| W2. *) +(* The assumption defW : W1 \x W2 = W is a parameter. *) +(* 'R[x] == the "signalizer" group of x \in 'A1(M) for the Dade *) +(* hypothesis of M (note: this is only extensionally equal *) +(* to the 'R[x] defined in BGsection14). *) +(* 'R_M == the signalizer functor for the Dade hypothesis of M. *) +(* Note that this only maps x to 'R[x] for x in 'A1(M). *) +(* The casual use of the R(x) in Peterfalvi is improper, *) +(* as its meaning depends on which maximal group is *) +(* considered. *) +(* 'A~(M, A) == the support of the image of 'CF(M, A) under the Dade *) +(* isometry of a maximal group M. *) +(* 'A1~(M) := 'A~(M, 'A1(M)). *) +(* 'A~(M) := 'A~(M, 'A(M)). *) +(* 'A0~(M) := 'A~(M, 'A0(M)). *) +(* FT_Dade maxM, FT_Dade0 maxM, FT_Dade1 maxM, FT_DadeF maxM *) +(* FT_Dade_hyp maxM, FT_Dade0_hyp maxM, FT_Dade1_hyp maxM, FT_DadeF_hyp maxM *) +(* == for maxM : M \in 'M, the Dade isometry of M, with *) +(* domain 'A(M), 'A0(M), 'A1(M) and M`_\F^#, respectively, *) +(* and the proofs of the corresponding Dade hypotheses. *) +(* Note that we use an additional restriction (to M`_\F^#) *) +(* to fit better with the conventions of PFsection7. *) +(* FTsupports M L <-> L supports M in the sense of (8.14) and (8.18). This *) +(* definition is not used outside this file. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory. + +Local Open Scope ring_scope. + +(* Supercedes the notation in BGsection14. *) +Notation "''R' [ x ]" := 'C_((gval 'N[x])`_\F)[x] + (at level 8, format "''R' [ x ]") : group_scope. +Notation "''R' [ x ]" := 'C_('N[x]`_\F)[x]%G : Group_scope. + +Section Definitions. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types L M X : {set gT}. + +(* These cover Peterfalvi, Definition (8.14). *) +Definition FTsignalizer M x := if 'C[x] \subset M then 1%G else 'R[x]%G. + +Definition FTsupports M L := + [exists x in 'A(M), ~~ ('C[x] \subset M) && ('C[x] \subset L)]. + +Definition FT_Dade_support M X := + \bigcup_(x in X) class_support (FTsignalizer M x :* x) G. + +End Definitions. + +Notation "''R_' M" := (FTsignalizer M) + (at level 8, M at level 2, format "''R_' M") : group_scope. + +Notation "''A~' ( M , A )" := (FT_Dade_support M A) + (at level 8, format "''A~' ( M , A )"). + +Notation "''A1~' ( M )" := 'A~(M, 'A1(M)) (at level 8, format "''A1~' ( M )"). +Notation "''A~' ( M )" := 'A~(M, 'A(M)) (at level 8, format "''A~' ( M )"). +Notation "''A0~' ( M )" := 'A~(M, 'A0(M)) (at level 8, format "''A0~' ( M )"). + +Section Eight. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types (p q : nat) (x y z : gT) (A B : {set gT}). +Implicit Types H K L M N P Q R S T U V W : {group gT}. + +(* Peterfalvi, Definition (8.1) is covered by BGsection16.of_typeF. *) + +(* This is the remark following Definition (8.1). *) +Remark compl_of_typeF M U V (H := M`_\F) : + H ><| U = M -> of_typeF M V -> of_typeF M U. +Proof. +move=> defM_U [[]]; rewrite -/H => ntH ntV defM part_b part_c. +have oU: #|U| = #|V|. + apply/eqP; rewrite -(@eqn_pmul2l #|H|) ?cardG_gt0 //. + by rewrite (sdprod_card defM) (sdprod_card defM_U). +have [x Mx defU]: exists2 x, x \in M & U :=: V :^ x. + pose pi := \pi(V); have hallV: pi.-Hall(M) V. + by rewrite Hall_pi // -(sdprod_Hall defM) (pHall_Hall (Fcore_Hall M)). + apply: Hall_trans (hallV). + rewrite mFT_sol // (sub_proper_trans _ (mFT_norm_proper ntH _)) ?gFnorm //. + rewrite (proper_sub_trans _ (subsetT M)) // properEcard gFsub. + by rewrite -(sdprod_card defM) ltn_Pmulr ?cardG_gt0 ?cardG_gt1. + rewrite pHallE -(card_Hall hallV) oU eqxx andbT. + by case/sdprod_context: defM_U. +have nHx: x \in 'N(H) by apply: subsetP Mx; rewrite gFnorm. +split; first by rewrite {1}defU conjsg_eq1. + have [U1 [nsU1U abU1 prU1H]] := part_b. + rewrite defU; exists (U1 :^ x)%G; split; rewrite ?normalJ ?abelianJ //. + rewrite -/H -(normP nHx) -conjD1g => _ /imsetP[h Hh ->]. + by rewrite -conjg_set1 normJ -conjIg conjSg prU1H. +have [U0 [sU0V expU0 frobHU0]] := part_c. +have [defHU0 _ ntU0 _ _] := Frobenius_context frobHU0. +rewrite defU; exists (U0 :^ x)%G; split; rewrite ?conjSg ?exponentJ //. +by rewrite -/H -(normP nHx) -conjYg FrobeniusJ. +Qed. + +Lemma Frobenius_of_typeF M U (H := M`_\F) : + [Frobenius M = H ><| U] -> of_typeF M U. +Proof. +move=> frobM; have [defM ntH ntU _ _] := Frobenius_context frobM. +have [_ _ nHU tiHU] := sdprodP defM. +split=> //; last by exists U; split; rewrite // -sdprodEY ?defM. +exists 1%G; split; rewrite ?normal1 ?abelian1 //. +by move=> x /(Frobenius_reg_compl frobM)->. +Qed. + +(* This is Peterfalvi (8.2). *) +Lemma typeF_context M U (H := M`_\F) : + of_typeF M U -> + [/\ (*a*) forall U0, is_typeF_complement M U U0 -> #|U0| = exponent U, + (*b*) [Frobenius M = H ><| U] = Zgroup U + & (*c*) forall U1 (i : Iirr H), + is_typeF_inertia M U U1 -> i != 0 -> 'I_U['chi_i] \subset U1]. +Proof. +case; rewrite -/H => [[ntH ntM defM] _ exU0]; set part_a := forall U0, _. +have [nsHM sUG mulHU nHU _] := sdprod_context defM. +have oU0: part_a. + move=> U0 [sU0U <- /Frobenius_reg_ker regU0]; rewrite exponent_Zgroup //. + apply/forall_inP=> S /SylowP[p _ /and3P[sSU0 pS _]]. + apply: odd_regular_pgroup_cyclic pS (mFT_odd S) ntH _ _. + by rewrite (subset_trans (subset_trans sSU0 sU0U)). + by move=> x /setD1P[ntx /(subsetP sSU0) U0x]; rewrite regU0 // !inE ntx. +split=> // [|U1 i [nsU1U abU1 s_cUH_U1] nz_i]. + apply/idP/idP=> [frobU | ZgU]. + apply/forall_inP=> S /SylowP[p _ /and3P[sSU pS _]]. + apply: odd_regular_pgroup_cyclic pS (mFT_odd S) ntH _ _. + by rewrite (subset_trans sSU). + move=> x /setD1P[ntx /(subsetP sSU) Ux]. + by rewrite (Frobenius_reg_ker frobU) // !inE ntx. + have [U0 [sU0U expU0 frobU0]] := exU0; have regU0 := Frobenius_reg_ker frobU0. + suffices defU0: U0 :=: U by rewrite defU0 norm_joinEr ?mulHU // in frobU0. + by apply/eqP; rewrite eqEcard sU0U /= (oU0 U0) // exponent_Zgroup. +have itoP: is_action M (fun (j : Iirr H) x => conjg_Iirr j x). + split=> [x | j x y Mx My]. + apply: can_inj (fun j => conjg_Iirr j x^-1) _ => j. + by apply: irr_inj; rewrite !conjg_IirrE cfConjgK. + by apply: irr_inj; rewrite !conjg_IirrE (cfConjgM _ nsHM). +pose ito := Action itoP; pose cto := ('Js \ subsetT M)%act. +have actsMcH: [acts M, on classes H | cto]. + apply/subsetP=> x Mx; rewrite !inE Mx; apply/subsetP=> _ /imsetP[y Hy ->]. + have nHx: x \in 'N(H) by rewrite (subsetP (gFnorm _ _)). + rewrite !inE /= -class_rcoset norm_rlcoset // class_lcoset mem_classes //. + by rewrite memJ_norm. +apply/subsetP=> g /setIP[Ug /setIdP[nHg c_i_g]]; have Mg := subsetP sUG g Ug. +apply: contraR nz_i => notU1g; rewrite (sameP eqP set1P). +suffices <-: 'Fix_ito[g] = [set 0 : Iirr H]. + by rewrite !inE sub1set inE -(inj_eq (@irr_inj _ _)) conjg_IirrE. +apply/eqP; rewrite eq_sym eqEcard cards1 !(inE, sub1set) /=. +rewrite -(inj_eq (@irr_inj _ _)) conjg_IirrE irr0 cfConjg_cfun1 eqxx. +rewrite (card_afix_irr_classes Mg actsMcH) => [|j y z Hy /=]; last first. + case/imsetP=> _ /imsetP[t Ht ->] -> {z}. + by rewrite conjg_IirrE cfConjgE // conjgK cfunJ. +rewrite -(cards1 [1 gT]) subset_leq_card //= -/H. +apply/subsetP=> _ /setIP[/imsetP[a Ha ->] /afix1P caHg]; rewrite inE classG_eq1. +have{caHg} /imsetP[x Hgx cax]: a \in a ^: (H :* g). + by rewrite class_rcoset caHg class_refl. +have coHg: coprime #|H| #[g]. + apply: (coprime_dvdr (order_dvdG Ug)). + by rewrite (coprime_sdprod_Hall_l defM) (pHall_Hall (Fcore_Hall M)). +have /imset2P[z y cHgg_z Hy defx]: x \in class_support ('C_H[g] :* g) H. + have [/and3P[/eqP defUcHgg _ _] _] := partition_cent_rcoset nHg coHg. + by rewrite class_supportEr -cover_imset defUcHgg. +rewrite -(can_eq (conjgKV y)) conj1g; apply: contraR notU1g => nt_ay'. +have{nt_ay'} Hay': a ^ y^-1 \in H^# by rewrite !inE nt_ay' groupJ ?groupV. +rewrite (subsetP (s_cUH_U1 _ Hay')) // inE Ug. +have ->: g = z.`_(\pi(H)^'). + have [h /setIP[Hh /cent1P cgh] ->] := rcosetP cHgg_z. + rewrite consttM // (constt1P _) ?mul1g ?constt_p_elt //. + by rewrite /p_elt -coprime_pi' ?cardG_gt0. + by rewrite (mem_p_elt _ Hh) // pgroupNK pgroup_pi. +by rewrite groupX //= -conjg_set1 normJ mem_conjgV -defx !inE conjg_set1 -cax. +Qed. + +(* Peterfalvi, Definition (8.3) is covered by BGsection16.of_typeI. *) +(* Peterfalvi, Definition (8.4) is covered by BGsection16.of_typeP. *) + +Section TypeP_Remarks. +(* These correspond to the remarks following Definition (8.4). *) + +Variables (M U W W1 W2 : {group gT}) (defW : W1 \x W2 = W). +Let H := M`_\F. +Let M' := M^`(1)%g. + +Hypothesis MtypeP : of_typeP M U defW. + +Remark of_typeP_sol : solvable M. +Proof. +have [_ [nilU _ _ defM'] _ _ _] := MtypeP. +have [nsHM' _ mulHU _ _] := sdprod_context defM'. +rewrite (series_sol (der_normal 1 M)) (abelian_sol (der_abelian 0 M)) andbT. +rewrite (series_sol nsHM') (nilpotent_sol (Fcore_nil M)). +by rewrite -mulHU quotientMidl quotient_sol ?(nilpotent_sol nilU). +Qed. + +Remark typeP_cent_compl : 'C_M'(W1) = W2. +Proof. +have [[/cyclicP[x ->] _ ntW1 _] _ _ [_ _ _ _ prM'W1] _] := MtypeP. +by rewrite cent_cycle prM'W1 // !inE cycle_id -cycle_eq1 ntW1. +Qed. + +Remark typeP_cent_core_compl : 'C_H(W1) = W2. +Proof. +have [sW2H sHM']: W2 \subset H /\ H \subset M'. + by have [_ [_ _ _ /sdprodW/mulG_sub[-> _]] _ []] := MtypeP. +by apply/eqP; rewrite eqEsubset subsetI sW2H -typeP_cent_compl ?subsetIr ?setSI. +Qed. + +Lemma typePF_exclusion K : ~ of_typeF M K. +Proof. +move=> [[ntH ntU1 defM_K] _ [U0 [sU01 expU0] frobU0]]. +have [[cycW1 hallW1 ntW1 defM] [_ _ _ defM'] _ [_]] := MtypeP; case/negP. +pose p := pdiv #|W1|; rewrite -/M' -/H in defM defM' frobU0 *. +have piW1p: p \in \pi(W1) by rewrite pi_pdiv cardG_gt1. +have piU0p: p \in \pi(U0). + rewrite -pi_of_exponent expU0 pi_of_exponent (pi_of_dvd _ _ piW1p) //=. + rewrite -(@dvdn_pmul2l #|H|) ?cardG_gt0 // (sdprod_card defM_K). + rewrite -(sdprod_card defM) dvdn_pmul2r ?cardSg //. + by case/sdprodP: defM' => _ <- _ _; exact: mulG_subl. +have [|X EpX]:= @p_rank_geP _ p 1 U0 _; first by rewrite p_rank_gt0. +have [ntX [sXU0 abelX _]] := (nt_pnElem EpX isT, pnElemP EpX). +have piW1_X: \pi(W1).-group X by apply: pi_pgroup piW1p; case/andP: abelX. +have sXM: X \subset M. + by rewrite -(sdprodWY defM_K) joingC sub_gen ?subsetU // (subset_trans sXU0). +have nHM: M \subset 'N(H) by apply: gFnorm. +have [regU0 solM] := (Frobenius_reg_ker frobU0, of_typeP_sol). +have [a Ma sXaW1] := Hall_Jsub solM (Hall_pi hallW1) sXM piW1_X. +rewrite -subG1 -(conjs1g a) -(cent_semiregular regU0 sXU0 ntX) conjIg -centJ. +by rewrite (normsP nHM) // -typeP_cent_core_compl ?setIS ?centS. +Qed. + +Remark of_typeP_compl_conj W1x : M' ><| W1x = M -> W1x \in W1 :^: M. +Proof. +case/sdprodP=> [[{W1x}_ W1x _ ->] mulM'W1x _ tiM'W1x]. +have [[_ /Hall_pi hallW1 _ defM] _ _ _ _] := MtypeP. +apply/imsetP; apply: Hall_trans of_typeP_sol _ (hallW1). +rewrite pHallE -(card_Hall hallW1) -(@eqn_pmul2l #|M'|) ?cardG_gt0 //. +by rewrite (sdprod_card defM) -mulM'W1x mulG_subr /= TI_cardMg. +Qed. + +Remark conj_of_typeP x : + {defWx : W1 :^ x \x W2 :^ x = W :^ x | of_typeP (M :^ x) (U :^ x) defWx}. +Proof. +have defWx: W1 :^ x \x W2 :^ x = W :^ x by rewrite -dprodJ defW. +exists defWx; rewrite /of_typeP !derJ FcoreJ FittingJ centJ -conjIg normJ. +rewrite !cyclicJ !conjsg_eq1 /Hall !conjSg indexJg cardJg -[_ && _]/(Hall M W1). +rewrite -(isog_nil (conj_isog U x)) -!sdprodJ -conjsMg -conjD1g. +rewrite -(conjGid (in_setT x)) -conjUg -conjDg normedTI_J. +have [[-> -> -> ->] [-> -> -> ->] [-> -> -> ->] [-> -> -> -> prW1] ->]:= MtypeP. +by do 2![split]=> // _ /imsetP[y /prW1<- ->]; rewrite cent1J -conjIg. +Qed. + +(* This is Peterfalvi (8.5), with an extra clause in anticipation of (8.15). *) +Lemma typeP_context : + [/\ (*a*) H \x 'C_U(H) = 'F(M), + (*b*) U^`(1)%g \subset 'C(H) /\ (U :!=: 1%g -> ~~ (U \subset 'C(H))), + (*c*) normedTI (cyclicTIset defW) G W + & cyclicTI_hypothesis G defW]. +Proof. +have defW2 := typeP_cent_core_compl. +case: MtypeP; rewrite /= -/H => [] [cycW1 hallW1 ntW1 defM] [nilU _ _ defM']. +set V := W :\: _ => [] [_ sM''F defF sFM'] [cycW2 ntW2 sW2H _ _] TI_V. +have [/andP[sHM' nHM'] sUM' mulHU _ tiHU] := sdprod_context defM'. +have sM'M : M' \subset M by apply: der_sub. +have hallM': \pi(M').-Hall(M) M' by rewrite Hall_pi // (sdprod_Hall defM). +have hallH_M': \pi(H).-Hall(M') H := pHall_subl sHM' sM'M (Fcore_Hall M). +have{defF} defF: (H * 'C_U(H))%g = 'F(M). + rewrite -(setIidPl sFM') -defF -group_modl //= -/H. + rewrite setIAC (setIidPr (der_sub 1 M)). + rewrite -(coprime_mulG_setI_norm mulHU) ?norms_cent //; last first. + by rewrite (coprime_sdprod_Hall_l defM') (pHall_Hall hallH_M'). + by rewrite mulgA (mulGSid (subsetIl _ _)). +have coW12: coprime #|W1| #|W2|. + rewrite coprime_sym (coprimeSg (subset_trans sW2H sHM')) //. + by rewrite (coprime_sdprod_Hall_r defM). +have cycW: cyclic W by rewrite (cyclic_dprod defW). +have ctiW: cyclicTI_hypothesis G defW by split; rewrite ?mFT_odd. +split=> //; first by rewrite dprodE ?subsetIr //= setIA tiHU setI1g. +split. + apply: subset_trans (_ : U :&: 'F(M) \subset _). + by rewrite subsetI der_sub (subset_trans (dergS 1 sUM')). + by rewrite -defF -group_modr ?subsetIl // setIC tiHU mul1g subsetIr. +apply: contra => cHU; rewrite -subG1 -tiHU subsetIidr (subset_trans sUM') //. +by rewrite (Fcore_max hallM') ?der_normal // -mulHU mulg_nil ?Fcore_nil. +Qed. + +End TypeP_Remarks. + +Remark FTtypeP_witness M : + M \in 'M -> FTtype M != 1%N -> exists_typeP (of_typeP M). +Proof. +move=> maxM /negbTE typeMnot1. +have:= FTtype_range M; rewrite -mem_iota !inE typeMnot1 /=. +by case/or4P=> /FTtypeP[//|U W W1 W2 defW [[]]]; exists U W W1 W2 defW. +Qed. + +(* Peterfalvi, Definition (8.6) is covered by BGsection16.of_typeII_IV et al. *) +(* Peterfalvi, Definition (8.7) is covered by BGsection16.of_typeV. *) + +Section FTypeP_Remarks. +(* The remarks for Definition (8.4) also apply to (8.6) and (8.7). *) + +Variables (M U W W1 W2 : {group gT}) (defW : W1 \x W2 = W). +Let H := M`_\F. +Let M' := M^`(1)%g. + +Hypotheses (maxM : M \in 'M) (MtypeP : of_typeP M U defW). + +Remark of_typeP_conj (Ux W1x W2x Wx : {group gT}) (defWx : W1x \x W2x = Wx) : + of_typeP M Ux defWx -> + exists x, + [/\ x \in M, U :^ x = Ux, W1 :^ x = W1x, W2 :^ x = W2x & W :^ x = Wx]. +Proof. +move=> MtypePx; have [[_ _ _ defMx] [_ _ nUW1x defM'x] _ _ _] := MtypePx. +have [[_ hallW1 _ defM] [_ _ nUW1 defM'] _ _ _] := MtypeP. +have [/mulG_sub[/= sHM' sUM'] [_ _ nM'W1 _]] := (sdprodW defM', sdprodP defM). +rewrite -/M' -/H in defMx defM'x defM defM' sHM' sUM' nM'W1. +have /imsetP[x2 Mx2 defW1x2] := of_typeP_compl_conj MtypeP defMx. +have /andP[sM'M nM'M]: M' <| M by apply: der_normal. +have solM': solvable M' := solvableS sM'M (of_typeP_sol MtypeP). +have [hallU hallUx]: \pi(H)^'.-Hall(M') U /\ \pi(H)^'.-Hall(M') (Ux :^ x2^-1). + have hallH: \pi(H).-Hall(M') H by apply: pHall_subl (Fcore_Hall M). + rewrite pHallJnorm ?(subsetP nM'M) ?groupV // -!(compl_pHall _ hallH). + by rewrite (sdprod_compl defM') (sdprod_compl defM'x). +have coM'W1: coprime #|M'| #|W1| by rewrite (coprime_sdprod_Hall_r defM). +have nUxW1: W1 \subset 'N(Ux :^ x2^-1) by rewrite normJ -sub_conjg -defW1x2. +have [x1] := coprime_Hall_trans nM'W1 coM'W1 solM' hallUx nUxW1 hallU nUW1. +case/setIP=> /(subsetP sM'M) My /(normsP (cent_sub _)) nW1x1 defUx1. +pose x := (x1 * x2)%g; have Mx: x \in M by rewrite groupM. +have defW1x: W1 :^ x = W1x by rewrite conjsgM nW1x1. +have defW2x: W2 :^ x = W2x. + rewrite -(typeP_cent_compl MtypeP) -(typeP_cent_compl MtypePx). + by rewrite conjIg -centJ defW1x (normsP nM'M). +by exists x; rewrite -defW dprodJ defW1x defW2x conjsgM -defUx1 conjsgKV. +Qed. + +Lemma FTtypeP_neq1 : FTtype M != 1%N. +Proof. by apply/FTtypeP=> // [[V [/(typePF_exclusion MtypeP)]]]. Qed. + +Remark compl_of_typeII_IV : FTtype M != 5 -> of_typeII_IV M U defW. +Proof. +move=> Mtype'5. +have [Ux Wx W1x W2x defWx Mtype24]: exists_typeP (of_typeII_IV M). + have:= FTtype_range M; rewrite leq_eqVlt eq_sym (leq_eqVlt _ 5). + rewrite (negPf FTtypeP_neq1) (negPf Mtype'5) /= -mem_iota !inE. + by case/or3P=> /FTtypeP[]// Ux Wx W1x W2x dWx []; exists Ux Wx W1x W2x dWx. +have [MtypePx ntUx prW1x tiFM] := Mtype24. +have [x [Mx defUx defW1x _ _]] := of_typeP_conj MtypePx. +by rewrite -defUx -defW1x cardJg conjsg_eq1 in ntUx prW1x. +Qed. + +Remark compl_of_typeII : FTtype M == 2 -> of_typeII M U defW. +Proof. +move=> Mtype2. +have [Ux Wx W1x W2x defWx [[MtypePx _ _ _]]] := FTtypeP 2 maxM Mtype2. +have [x [Mx <- _ _ _]] := of_typeP_conj MtypePx; rewrite -/M' -/H. +rewrite abelianJ normJ -{1}(conjGid Mx) conjSg => cUU not_sNUM M'typeF defH. +split=> //; first by apply: compl_of_typeII_IV; rewrite // (eqP Mtype2). +by apply: compl_of_typeF M'typeF; rewrite defH; have [_ []] := MtypeP. +Qed. + +Remark compl_of_typeIII : FTtype M == 3 -> of_typeIII M U defW. +Proof. +move=> Mtype3. +have [Ux Wx W1x W2x defWx [[MtypePx _ _ _]]] := FTtypeP 3 maxM Mtype3. +have [x [Mx <- _ _ _]] := of_typeP_conj MtypePx; rewrite -/M' -/H. +rewrite abelianJ normJ -{1}(conjGid Mx) conjSg. +by split=> //; apply: compl_of_typeII_IV; rewrite // (eqP Mtype3). +Qed. + +Remark compl_of_typeIV : FTtype M == 4 -> of_typeIV M U defW. +Proof. +move=> Mtype4. +have [Ux Wx W1x W2x defWx [[MtypePx _ _ _]]] := FTtypeP 4 maxM Mtype4. +have [x [Mx <- _ _ _]] := of_typeP_conj MtypePx; rewrite -/M' -/H. +rewrite abelianJ normJ -{1}(conjGid Mx) conjSg. +by split=> //; apply: compl_of_typeII_IV; rewrite // (eqP Mtype4). +Qed. + +Remark compl_of_typeV : FTtype M == 5 -> of_typeV M U defW. +Proof. +move=> Mtype5. +have [Ux Wx W1x W2x defWx [[MtypePx /eqP]]] := FTtypeP 5 maxM Mtype5. +have [x [Mx <- <- _ _]] := of_typeP_conj MtypePx; rewrite -/M' -/H. +by rewrite cardJg conjsg_eq1 => /eqP. +Qed. + +End FTypeP_Remarks. + +(* This is the statement of Peterfalvi, Theorem (8.8)(a). *) +Definition all_FTtype1 := [forall M : {group gT} in 'M, FTtype M == 1%N]. + +(* This is the statement of Peterfalvi, Theorem (8.8)(b). *) +Definition typeP_pair S T (W W1 W2 : {set gT}) (defW : W1 \x W2 = W) := + [/\ [/\ cyclicTI_hypothesis G defW, S \in 'M & T \in 'M], + (*b1*) [/\ S^`(1) ><| W1 = S, T^`(1) ><| W2 = T & S :&: T = W]%g, + (*b2*) (FTtype S == 2) || (FTtype T == 2), + (*b3*) (1 < FTtype S <= 5 /\ 1 < FTtype T <= 5)%N + & (*b4*) {in 'M, forall M, FTtype M != 1%N -> gval M \in S :^: G :|: T :^: G}]. + +Lemma typeP_pair_sym S T W W1 W2 (defW : W1 \x W2 = W) (xdefW : W2 \x W1 = W) : + typeP_pair S T defW -> typeP_pair T S xdefW. +Proof. +by case=> [[/cyclicTIhyp_sym ? ? ?] [? ?]]; rewrite setIC setUC orbC => ? ? []. +Qed. + +(* This is Peterfalvi, Theorem (8.8). *) +Lemma FTtypeP_pair_cases : + (*a*) {in 'M, forall M, FTtype M == 1%N} + \/ (*b*) exists S, exists T, exists_typeP (fun _ => typeP_pair S T). +Proof. +have [_ [| [[S T] [[maxS maxT] [[W1 W2] /=]]]]] := BGsummaryI gT; first by left. +set W := W1 <*> W2; set V := W :\: (W1 :|: W2). +case=> [[cycW tiV _] [defS defT tiST]] b4 /orP b2 b3. +have [cWW /joing_sub[sW1W sW2W]] := (cyclic_abelian cycW, erefl W). +have ntV: V != set0 by have [] := andP tiV. +suffices{tiST tiV cWW sW1W sW2W b3 b4} tiW12: W1 :&: W2 = 1%g. + have defW: W1 \x W2 = W by rewrite dprodEY ?(centSS _ _ cWW). + right; exists S, T; exists S _ _ _ defW; split=> // [|M _ /b4[] // x]. + by do 2?split; rewrite ?mFT_odd // /normedTI tiV nVW setTI /=. + by case=> <-; rewrite inE mem_orbit ?orbT. +wlog {b2 T defT maxT} Stype2: S W1 W2 @W @V maxS defS cycW ntV / FTtype S == 2. + move=> IH; case/orP: b2 cycW ntV => /IH; first exact. + by rewrite setIC /V /W /= joingC setUC; apply. +have{maxS Stype2 defS} prW1: prime #|W1|. + have [U ? W1x ? ? [[StypeP _ prW1x _] _ _ _ _]] := FTtypeP 2 maxS Stype2. + by have /imsetP[x _ ->] := of_typeP_compl_conj StypeP defS; rewrite cardJg. +rewrite prime_TIg //; apply: contra ntV => sW12. +by rewrite setD_eq0 (setUidPr sW12) join_subG sW12 /=. +Qed. + +(* This is Peterfalvi (8.9). *) +(* We state the lemma using the of_typeP predicate, as it is the Skolemised *) +(* form of Peterfalvi, Definition (8.4). *) +Lemma typeP_pairW S T W W1 W2 (defW : W1 \x W2 = W) : + typeP_pair S T defW -> exists U : {group gT}, of_typeP S U defW. +Proof. +case=> [[[cycW _ /and3P[_ _ /eqP nVW]] maxS _] [defS _ defST] _ [Stype25 _] _]. +set S' := S^`(1)%g in defS; have [nsS'S _ _ _ tiS'W1] := sdprod_context defS. +have{Stype25} Stype'1: FTtype S != 1%N by apply: contraTneq Stype25 => ->. +have [/mulG_sub[sW1W sW2W] [_ mulW12 cW12 _]] := (dprodW defW, dprodP defW). +have [cycW1 cycW2] := (cyclicS sW1W cycW, cyclicS sW2W cycW). +have{cycW1 cycW2} coW12: coprime #|W1| #|W2| by rewrite -(cyclic_dprod defW). +have{maxS Stype'1} [Ux Wx W1x W2x defWx StypeP] := FTtypeP_witness maxS Stype'1. +have /imsetP[y Sy defW1] := of_typeP_compl_conj StypeP defS. +suffices defW2: W2 :=: W2x :^ y. + have [] := conj_of_typeP StypeP y; rewrite -defWx dprodJ -defW1 -defW2. + by rewrite (conjGid Sy) {-1}defW; exists (Ux :^ y)%G. +have [[_ hallW1x _ defSx] _ _ [/cyclic_abelian abW2x _ _ _ _] _] := StypeP. +have{Sy} nS'y: y \in 'N(S') by rewrite (subsetP (normal_norm nsS'S)). +have{nS'y} defW2xy: W2x :^ y = 'C_S'(W1). + by rewrite -(typeP_cent_compl StypeP) conjIg -centJ -defW1 (normP nS'y). +have{nsS'S} sW2S': W2 \subset S'. + have sW2S: W2 \subset S by rewrite (subset_trans sW2W) // -defST subsetIl. + have{hallW1x} hallW1: \pi(W1).-Hall(S) W1x by rewrite defW1 /= cardJg Hall_pi. + have hallS': \pi(W1)^'.-Hall(S) S' by apply/(sdprod_normal_pHallP _ hallW1). + by rewrite coprime_pi' // (sub_normal_Hall hallS') in coW12 *. +have sW2xy: W2 \subset W2x :^ y by rewrite defW2xy subsetI sW2S'. +have defW2: W2 :=: S' :&: W by rewrite -mulW12 -group_modr ?tiS'W1 ?mul1g. +apply/eqP; rewrite eqEsubset sW2xy defW2 subsetI {1}defW2xy subsetIl /=. +rewrite -nVW /= setTI cents_norm // (centsS (subsetDl _ _)) // -mulW12. +by rewrite centM subsetI {1}defW2xy subsetIr sub_abelian_cent // abelianJ. +Qed. + +Section OneMaximal. + +Variable M U W W1 W2 : {group gT}. (* W, W1 and W2 are only used later. *) +Hypothesis maxM : M \in 'M. + +(* Peterfalvi, Definition (8.10) is covered in BGsection16. *) + +(* This is Peterfalvi (8.11). *) +Lemma FTcore_facts : + [/\ Hall G M`_\F, Hall G M`_\s + & forall S, Sylow M`_\s S -> S :!=: 1%g -> 'N(S) \subset M]. +Proof. +have hallMs := Msigma_Hall_G maxM; have [_ sMs _] := and3P hallMs. +rewrite def_FTcore // (pHall_Hall hallMs). +split=> // [|S /SylowP[p _ sylS] ntS]. + have sMF_Ms:= Fcore_sub_Msigma maxM. + apply: (@pHall_Hall _ \pi(M`_\F)); apply: (subHall_Hall hallMs). + by move=> p /(piSg sMF_Ms)/(pnatPpi sMs). + exact: pHall_subl (pcore_sub _ M) (Fcore_Hall M). +have s_p: p \in \sigma(M). + by rewrite (pnatPpi sMs) // -p_rank_gt0 -(rank_Sylow sylS) rank_gt0. +by apply: (norm_sigma_Sylow s_p); exact: (subHall_Sylow (Msigma_Hall maxM)). +Qed. + +(* This is Peterfalvi (8.12). *) +(* (b) could be stated for subgroups of U wlog -- usage should be checked. *) +Lemma FTtypeI_II_facts n (H := M`_\F) : + FTtype M == n -> H ><| U = M ^`(n.-1)%g -> + if 0 < n <= 2 then + [/\ (*a*) forall p S, p.-Sylow(U) S -> abelian S /\ ('r(S) <= 2)%N, + (*b*) forall X, X != set0 -> X \subset U^# -> 'C_H(X) != 1%g -> + 'M('C(X)) = [set M] + & (*c*) let B := 'A(M) :\: 'A1(M) in B != set0 -> normedTI B G M + ] else True. +Proof. +move=> typeM defMn; have [n12 | //] := ifP; rewrite -mem_iota !inE in n12. +have defH: H = M`_\sigma. + by rewrite -def_FTcore -?(Fcore_eq_FTcore _ _) // (eqP typeM) !inE orbA n12. +have [K complU]: exists K : {group gT}, kappa_complement M U K. + have [[V K] /= complV] := kappa_witness maxM. + have [[hallV hallK gVK] [_ sUMn _ _ _]] := (complV, sdprod_context defMn). + have hallU: \sigma_kappa(M)^'.-Hall(M) U. + rewrite pHallE -(card_Hall hallV) (subset_trans sUMn) ?der_sub //=. + rewrite -(@eqn_pmul2l #|H|) ?cardG_gt0 // (sdprod_card defMn) defH. + rewrite (sdprod_card (sdprod_FTder maxM complV)) (eqP typeM). + by case/pred2P: n12 => ->. + have [x Mx defU] := Hall_trans (mmax_sol maxM) hallU hallV. + exists (K :^ x)%G; split; rewrite ?pHallJ // defU -conjsMg. + by rewrite -(gen_set_id gVK) groupP. +have [part_a _ _ [part_b part_c]] := BGsummaryB maxM complU. +rewrite eqEsubset FTsupp1_sub // andbT -setD_eq0 in part_c. +split=> // X notX0 /subsetD1P[sXU notX1]; rewrite -cent_gen defH. +apply: part_b; rewrite -?subG1 ?gen_subG //. +by rewrite -setD_eq0 setDE (setIidPl _) // subsetC sub1set inE. +Qed. + +(* This is Peterfalvi (8.13). *) +(* We have substituted the B & G notation for the unique maximal supergroup *) +(* of 'C[x], and specialized the lemma to X := 'A0(M). *) +Lemma FTsupport_facts (X := 'A0(M)) (D := [set x in X | ~~('C[x] \subset M)]) : + [/\ (*a*) {in X &, forall x, {subset x ^: G <= x ^: M}}, + (*b*) D \subset 'A1(M) /\ {in D, forall x, 'M('C[x]) = [set 'N[x]]} + & (*c*) {in D, forall x (L := 'N[x]) (H := L`_\F), + [/\ (*c1*) H ><| (M :&: L) = L /\ 'C_H[x] ><| 'C_M[x] = 'C[x], + (*c2*) {in X, forall y, coprime #|H| #|'C_M[y]| }, + (*c3*) x \in 'A(L) :\: 'A1(L) + & (*c4*) 1 <= FTtype L <= 2 + /\ (FTtype L == 2 -> [Frobenius M with kernel M`_\F])]}]. +Proof. +have defX: X \in pred2 'A(M) 'A0(M) by rewrite !inE eqxx orbT. +have [sDA1 part_a part_c] := BGsummaryII maxM defX. +have{part_a} part_a: {in X &, forall x, {subset x ^: G <= x ^: M}}. + move=> x y A0x A0y /= /imsetP[g Gg def_y]; rewrite def_y. + by apply/imsetP/part_a; rewrite -?def_y. +do [split=> //; first split=> //] => x /part_c[_ ] //. +rewrite /= -(mem_iota 1) !inE => -> [-> ? -> -> L2_frob]. +by do 2![split=> //] => /L2_frob[E /FrobeniusWker]. +Qed. + +(* A generic proof of the first assertion of Peterfalvi (8.15). *) +Let norm_FTsuppX A : + M \subset 'N(A) -> 'A1(M) \subset A -> A \subset 'A0(M) -> 'N(A) = M. +Proof. +move=> nAM sA1A sAA0; apply: mmax_max => //. +rewrite (sub_proper_trans (norm_gen _)) ?mFT_norm_proper //; last first. + rewrite (sub_proper_trans _ (mmax_proper maxM)) // gen_subG. + by rewrite (subset_trans sAA0) // (subset_trans (FTsupp0_sub M)) ?subsetDl. +rewrite (subG1_contra (genS sA1A)) //= genD1 ?group1 //. +by rewrite genGid /= def_FTcore ?Msigma_neq1. +Qed. + +Lemma norm_FTsupp1 : 'N('A1(M)) = M. +Proof. exact: norm_FTsuppX (FTsupp1_norm M) _ (FTsupp1_sub0 maxM). Qed. + +Lemma norm_FTsupp : 'N('A(M)) = M. +Proof. exact: norm_FTsuppX (FTsupp_norm M) (FTsupp1_sub _) (FTsupp_sub0 M). Qed. + +Lemma norm_FTsupp0 : 'N('A0(M)) = M. +Proof. exact: norm_FTsuppX (FTsupp0_norm M) (FTsupp1_sub0 _) _. Qed. + +Lemma FTsignalizerJ x y : 'R_(M :^ x) (y ^ x) :=: 'R_M y :^ x. +Proof. +rewrite /'R__ /= {1}cent1J conjSg; case: ifP => _ /=; first by rewrite conjs1g. +by rewrite cent1J FT_signalizer_baseJ FcoreJ -conjIg. +Qed. + +Let is_FTsignalizer : is_Dade_signalizer G M 'A0(M) 'R_M. +Proof. +rewrite /'R_M => x A0x /=; rewrite setTI. +case: ifPn => [sCxM | not_sCxM]; first by rewrite sdprod1g (setIidPr sCxM). +by have [_ _ /(_ x)[| [] //]] := FTsupport_facts; exact/setIdP. +Qed. + +(* This is Peterfalvi (8.15), second assertion. *) +Lemma FT_Dade0_hyp : Dade_hypothesis G M 'A0(M). +Proof. +have [part_a _ parts_bc] := FTsupport_facts. +have /subsetD1P[sA0M notA0_1] := FTsupp0_sub M. +split; rewrite // /normal ?sA0M ?norm_FTsupp0 //=. +exists 'R_M => [|x y A0x A0y]; first exact: is_FTsignalizer. +rewrite /'R_M; case: ifPn => [_ | not_sCxM]; first by rewrite cards1 coprime1n. +rewrite (coprimeSg (subsetIl _ _)) //=. +by have [| _ -> //] := parts_bc x; apply/setIdP. +Qed. + +Definition FT_Dade_hyp := + restr_Dade_hyp FT_Dade0_hyp (FTsupp_sub0 M) (FTsupp_norm M). + +Definition FT_Dade1_hyp := + restr_Dade_hyp FT_Dade0_hyp (FTsupp1_sub0 maxM) (FTsupp1_norm M). + +Definition FT_DadeF_hyp := + restr_Dade_hyp FT_Dade0_hyp (Fcore_sub_FTsupp0 maxM) (normsD1 (gFnorm _ _)). + +Lemma def_FTsignalizer0 : {in 'A0(M), Dade_signalizer FT_Dade0_hyp =1 'R_M}. +Proof. exact: def_Dade_signalizer. Qed. + +Lemma def_FTsignalizer : {in 'A(M), Dade_signalizer FT_Dade_hyp =1 'R_M}. +Proof. exact: restr_Dade_signalizer def_FTsignalizer0. Qed. + +Lemma def_FTsignalizer1 : {in 'A1(M), Dade_signalizer FT_Dade1_hyp =1 'R_M}. +Proof. exact: restr_Dade_signalizer def_FTsignalizer0. Qed. + +Lemma def_FTsignalizerF : {in M`_\F^#, Dade_signalizer FT_DadeF_hyp =1 'R_M}. +Proof. exact: restr_Dade_signalizer def_FTsignalizer0. Qed. + +Local Notation tau := (Dade FT_Dade0_hyp). +Local Notation FT_Dade := (Dade FT_Dade_hyp). +Local Notation FT_Dade1 := (Dade FT_Dade1_hyp). +Local Notation FT_DadeF := (Dade FT_DadeF_hyp). + +Lemma FT_DadeE : {in 'CF(M, 'A(M)), FT_Dade =1 tau}. +Proof. exact: restr_DadeE. Qed. + +Lemma FT_Dade1E : {in 'CF(M, 'A1(M)), FT_Dade1 =1 tau}. +Proof. exact: restr_DadeE. Qed. + +Lemma FT_DadeF_E : {in 'CF(M, M`_\F^#), FT_DadeF =1 tau}. +Proof. exact: restr_DadeE. Qed. + +Lemma FT_Dade_supportS A B : A \subset B -> 'A~(M, A) \subset 'A~(M, B). +Proof. +by move/subsetP=> sAB; apply/bigcupsP=> x Ax; rewrite (bigcup_max x) ?sAB. +Qed. + +Lemma FT_Dade0_supportE : Dade_support FT_Dade0_hyp = 'A0~(M). +Proof. by apply/eq_bigr=> x /def_FTsignalizer0 <-. Qed. + +Let defA A (sAA0 : A \subset 'A0(M)) (nAM : M \subset 'N(A)) : + Dade_support (restr_Dade_hyp FT_Dade0_hyp sAA0 nAM) = 'A~(M, A). +Proof. +by apply/eq_bigr=> x /(restr_Dade_signalizer sAA0 nAM def_FTsignalizer0) <-. +Qed. + +Lemma FT_Dade_supportE : Dade_support FT_Dade_hyp = 'A~(M). +Proof. exact: defA. Qed. + +Lemma FT_Dade1_supportE : Dade_support FT_Dade1_hyp = 'A1~(M). +Proof. exact: defA. Qed. + +Lemma FT_DadeF_supportE : Dade_support FT_DadeF_hyp = 'A~(M, M`_\F^#). +Proof. exact: defA. Qed. + +Lemma FT_Dade0_supportJ x : 'A0~(M :^ x) = 'A0~(M). +Proof. +rewrite /'A0~(_) FTsupp0J big_imset /=; last exact: in2W (conjg_inj x). +apply: eq_bigr => y _; rewrite FTsignalizerJ -conjg_set1 -conjsMg. +by rewrite class_supportGidl ?inE. +Qed. + +Lemma FT_Dade1_supportJ x : 'A1~(M :^ x) = 'A1~(M). +Proof. +rewrite /'A1~(_) FTsupp1J big_imset /=; last exact: in2W (conjg_inj x). +apply: eq_bigr => y _; rewrite FTsignalizerJ -conjg_set1 -conjsMg. +by rewrite class_supportGidl ?inE. +Qed. + +Lemma FT_Dade_supportJ x : 'A~(M :^ x) = 'A~(M). +Proof. +rewrite /'A~(_) FTsuppJ big_imset /=; last exact: in2W (conjg_inj x). +apply: eq_bigr => y _; rewrite FTsignalizerJ -conjg_set1 -conjsMg. +by rewrite class_supportGidl ?inE. +Qed. + +(* Subcoherence and cyclicTI properties of type II-V subgroups. *) +Hypotheses (defW : W1 \x W2 = W) (MtypeP : of_typeP M U defW). +Let H := M`_\F%G. +Let K := M^`(1)%G. + +Lemma FT_cyclicTI_hyp : cyclicTI_hypothesis G defW. +Proof. by case/typeP_context: MtypeP. Qed. +Let ctiW := FT_cyclicTI_hyp. + +(* This is a useful combination of Peterfalvi (8.8) and (8.9). *) +Lemma FTtypeP_pair_witness : + exists2 T, typeP_pair M T defW + & exists xdefW : W2 \x W1 = W, exists V : {group gT}, of_typeP T V xdefW. +Proof. +have Mtype'1 := FTtypeP_neq1 maxM MtypeP. +case: FTtypeP_pair_cases => [/(_ M maxM)/idPn[] // | [S [T]]]. +case=> _ Wx W1x W2x defWx pairST. +without loss /imsetP[y2 _ defSy]: S T W1x W2x defWx pairST / gval M \in S :^: G. + have [_ _ _ _ coverST] := pairST => IH. + have /setUP[] := coverST M maxM Mtype'1; first exact: IH pairST. + by apply: IH (typeP_pair_sym _ pairST); rewrite dprodC. +have [U_S StypeP] := typeP_pairW pairST. +have [[_ maxS maxT] [defS defT defST] b2 b3 b4] := pairST. +have [[[_ _ _ defM] _ _ _ _] defW2] := (MtypeP, typeP_cent_compl MtypeP). +have /imsetP[y1 Sy1 /(canRL (conjsgKV _)) defW1]: W1 :^ y2^-1 \in W1x :^: S. + apply: (of_typeP_compl_conj StypeP). + by rewrite -(conjsgK y2 S) -defSy derJ -sdprodJ defM. +pose y := (y1 * y2)%g; rewrite -conjsgM -/y in defW1. +have{defSy} defSy: S :^ y = M by rewrite conjsgM (conjGid Sy1). +have{defW2} defW2: W2 :=: W2x :^ y. + by rewrite -(typeP_cent_compl StypeP) conjIg -derJ -centJ defSy -defW1. +suffices pairMTy: typeP_pair M (T :^ y) defW. + exists (T :^ y)%G => //; have xdefW: W2 \x W1 = W by rewrite dprodC. + by exists xdefW; apply: typeP_pairW (typeP_pair_sym xdefW pairMTy). +do [split; rewrite ?defM -?defSy ?mmaxJ ?FTtypeJ //] => [|L maxL /(b4 L maxL)]. + by rewrite -defW defW1 defW2 derJ -sdprodJ -dprodJ -conjIg defT defST defWx. +by rewrite !conjugates_conj lcoset_id // inE. +Qed. + +(* A converse to the above. *) +Lemma of_typeP_pair (xdefW : W2 \x W1 = W) T V : + T \in 'M -> of_typeP T V xdefW -> typeP_pair M T defW. +Proof. +have [S pairMS [xdefW' [V1 StypeP]]] := FTtypeP_pair_witness => maxT TtypeP. +have [[cycW2 /andP[sW2T _] ntW2 _] _ _ [cycW1 _ _ sW1T'' _] _] := TtypeP. +have{sW1T'' sW2T} sWT: W \subset T. + by rewrite -(dprodW defW) mul_subG ?(subset_trans sW1T'') ?gFsub. +have [cycW _ /and3P[_ _ /eqP defNW]] := ctiW. +rewrite (@group_inj _ T S) //; have{pairMS} [_ _ _ _ defT] := pairMS. +have /defT/setUP[] := FTtypeP_neq1 maxT TtypeP => {defT}// /imsetP[x _ defT]. + have [defWx] := conj_of_typeP MtypeP x; rewrite -defT. + case/(of_typeP_conj TtypeP)=> y [_ _ _ defW1y _]. + have /idP[]:= negbF cycW; rewrite (cyclic_dprod defW) // /coprime. + by rewrite -(cardJg _ y) defW1y cardJg gcdnn -trivg_card1. +have [defWx] := conj_of_typeP StypeP x; rewrite -defT. +case/(of_typeP_conj TtypeP)=> y [Ty _ defW2y defW1y defWy]. +have Wyx: (y * x^-1)%g \in W. + by rewrite -defNW !inE /= conjDg conjUg !conjsgM defW2y defW1y defWy !conjsgK. +by rewrite -(conjGid (subsetP sWT _ Wyx)) conjsgM (conjGid Ty) defT conjsgK. +Qed. + +Lemma FT_primeTI_hyp : primeTI_hypothesis M K defW. +Proof. +have [[cycW1 ntW1 hallW1 defM] _ _ [cycW2 ntW2 _ sW2M'' prM'W1] _] := MtypeP. +by split; rewrite ?mFT_odd // (subset_trans sW2M'') ?der_subS. +Qed. +Let ptiWM := FT_primeTI_hyp. + +Lemma FTtypeP_supp0_def : + 'A0(M) = 'A(M) :|: class_support (cyclicTIset defW) M. +Proof. +rewrite -(setID 'A0(M) 'A(M)) (FTsupp0_typeP maxM MtypeP) (setIidPr _) //. +exact: FTsupp_sub0. +Qed. + +Fact FT_Fcore_prime_Dade_def : prime_Dade_definition M K H 'A(M) 'A0(M) defW. +Proof. +have [_ [_ _ _ /sdprodW/mulG_sub[sHK _]] _ [_ _ sW2H _ _] _] := MtypeP. +split; rewrite ?gFnormal //; last exact: FTtypeP_supp0_def. +rewrite /normal FTsupp_norm andbT /'A(M) (FTtypeP_neq1 maxM MtypeP) /=. +do ?split=> //; apply/bigcupsP=> x A1x; last by rewrite setSD ?subsetIl. + by rewrite setDE -setIA subIset // gFsub. +by rewrite (bigcup_max x) // (subsetP _ x A1x) // setSD ?Fcore_sub_FTcore. +Qed. + +Definition FT_prDade_hypF : prime_Dade_hypothesis _ M K H 'A(M) 'A0(M) defW := + PrimeDadeHypothesis ctiW ptiWM FT_Dade0_hyp FT_Fcore_prime_Dade_def. + +Fact FT_core_prime_Dade_def : prime_Dade_definition M K M`_\s 'A(M) 'A0(M) defW. +Proof. +have [[_ sW2H sHK] [nsAM sCA sAK] defA0] := FT_Fcore_prime_Dade_def. +have [_ [_ sW2K _ _] _] := ptiWM. +split=> //=; first by rewrite FTcore_normal /M`_\s; case: ifP. +rewrite nsAM /= /'A(M) /M`_\s (FTtypeP_neq1 maxM MtypeP); split=> //=. +by apply/bigcupsP=> x _; rewrite setSD ?subsetIl. +Qed. + +Definition FT_prDade_hyp : prime_Dade_hypothesis _ M K M`_\s 'A(M) 'A0(M) defW + := PrimeDadeHypothesis ctiW ptiWM FT_Dade0_hyp FT_core_prime_Dade_def. + +Let calS := seqIndD K M M`_\s 1. + +Fact FTtypeP_cohererence_base_subproof : cfConjC_subset calS calS. +Proof. exact: seqInd_conjC_subset1. Qed. + +Fact FTtypeP_cohererence_nonreal_subproof : ~~ has cfReal calS. +Proof. by rewrite seqInd_notReal ?mFT_odd ?FTcore_sub_der1 ?der_normal. Qed. + +Definition FTtypeP_coh_base_sig := + prDade_subcoherent FT_prDade_hyp + FTtypeP_cohererence_base_subproof FTtypeP_cohererence_nonreal_subproof. + +Definition FTtypeP_coh_base := sval FTtypeP_coh_base_sig. + +Local Notation R := FTtypeP_coh_base. + +Lemma FTtypeP_subcoherent : subcoherent calS tau R. +Proof. by rewrite /R; case: FTtypeP_coh_base_sig => R1 []. Qed. +Let scohS := FTtypeP_subcoherent. + +Let w_ i j := cyclicTIirr defW i j. +Let sigma := cyclicTIiso ctiW. +Let eta_ i j := sigma (w_ i j). +Let mu_ := primeTIred ptiWM. +Let delta_ := fun j => primeTIsign ptiWM j. + +Lemma FTtypeP_base_ortho : + {in [predI calS & irr M] & irr W, forall phi w, orthogonal (R phi) (sigma w)}. +Proof. by rewrite /R; case: FTtypeP_coh_base_sig => R1 []. Qed. + +Lemma FTtypeP_base_TIred : + let dsw j k := [seq delta_ j *: eta_ i k | i : Iirr W1] in + let Rmu j := dsw j j ++ map -%R (dsw j (conjC_Iirr j)) in + forall j, R (mu_ j) = Rmu j. +Proof. by rewrite /R; case: FTtypeP_coh_base_sig => R1 []. Qed. + +Lemma coherent_ortho_cycTIiso calS1 (tau1 : {additive 'CF(M) -> 'CF(G)}) : + cfConjC_subset calS1 calS -> coherent_with calS1 M^# tau tau1 -> + forall chi i j, chi \in calS1 -> chi \in irr M -> '[tau1 chi, eta_ i j] = 0. +Proof. +move=> ccsS1S cohS1 chi i j S1chi chi_irr; have [_ sS1S _] := ccsS1S. +have [e /mem_subseq Re ->] := mem_coherent_sum_subseq scohS ccsS1S cohS1 S1chi. +rewrite cfdot_suml big1_seq // => xi /Re; apply: orthoPr. +by apply: FTtypeP_base_ortho (mem_irr _); rewrite !inE sS1S. +Qed. + +Import ssrnum Num.Theory. + +(* A reformuation of Peterfalvi (5.8) for the Odd Order proof context. *) +Lemma FTtypeP_coherent_TIred calS1 tau1 zeta j : + cfConjC_subset calS1 calS -> coherent_with calS1 M^# tau tau1 -> + zeta \in irr M -> zeta \in calS1 -> mu_ j \in calS1 -> + let d := primeTI_Isign ptiWM j in let k := conjC_Iirr j in + {dk : bool * Iirr W2 | tau1 (mu_ j) = (-1) ^+ dk.1 *: (\sum_i eta_ i dk.2) + & dk.1 = d /\ dk.2 = j + \/ [/\ dk.1 = ~~ d, dk.2 = k + & forall l, mu_ l \in calS1 -> mu_ l 1%g = mu_ j 1%g -> pred2 j k l]}. +Proof. +move=> ccsS1S cohS1 irr_zeta S1zeta S1mu_j d k. +have irrS1: [/\ ~~ has cfReal calS1, has (mem (irr M)) calS1 & mu_ j \in calS1]. + have [[_ -> _] _ _ _ _] := subset_subcoherent scohS ccsS1S. + by split=> //; apply/hasP; exists zeta. +have Dmu := coherent_prDade_TIred FT_prDade_hyp ccsS1S irrS1 cohS1. +rewrite -/mu_ -/d in Dmu; pose mu_sum d1 k1 := (-1) ^+ d1 *: (\sum_i eta_ i k1). +have mu_sumK (d1 d2 : bool) k1 k2: + ('[mu_sum d1 k1, (-1) ^+ d2 *: eta_ 0 k2] > 0) = (d1 == d2) && (k1 == k2). +- rewrite cfdotZl cfdotZr rmorph_sign mulrA -signr_addb cfdot_suml. + rewrite (bigD1 0) //= cfdot_cycTIiso !eqxx big1 => [|i nz_i]; last first. + by rewrite cfdot_cycTIiso (negPf nz_i). + rewrite addr0 /= andbC; case: (k1 == k2); rewrite ?mulr0 ?ltrr //=. + by rewrite mulr1 signr_gt0 negb_add. +have [dk tau1mu_j]: {dk : bool * Iirr W2 | tau1 (mu_ j) = mu_sum dk.1 dk.2}. + apply: sig_eqW; case: Dmu => [-> | [-> _]]; first by exists (d, j). + by exists (~~ d, k); rewrite -signrN. +exists dk => //; have:= mu_sumK dk.1 dk.1 dk.2 dk.2; rewrite !eqxx -tau1mu_j. +case: Dmu => [-> | [-> all_jk]]; + rewrite -?signrN mu_sumK => /andP[/eqP <- /eqP <-]; [by left | right]. +by split=> // j1 S1j1 /(all_jk j1 S1j1)/pred2P. +Qed. + +Lemma size_red_subseq_seqInd_typeP (calX : {set Iirr K}) calS1 : + uniq calS1 -> {subset calS1 <= seqInd M calX} -> + {subset calS1 <= [predC irr M]} -> + size calS1 = #|[set i : Iirr K | 'Ind 'chi_i \in calS1]|. +Proof. +move=> uS1 sS1S redS1; pose h s := 'Ind[M, K] 'chi_s. +apply/eqP; rewrite cardE -(size_map h) -uniq_size_uniq // => [|xi]; last first. + apply/imageP/idP=> [[i] | S1xi]; first by rewrite inE => ? ->. + by have /seqIndP[s _ Dxi] := sS1S _ S1xi; exists s; rewrite ?inE -?Dxi. +apply/dinjectiveP; pose h1 xi := cfIirr (#|W1|%:R^-1 *: 'Res[K, M] xi). +apply: can_in_inj (h1) _ => s; rewrite inE => /redS1 red_s. +have cycW1: cyclic W1 by have [[]] := MtypeP. +have [[j /irr_inj->] | [/idPn[]//]] := prTIres_irr_cases ptiWM s. +by rewrite /h cfInd_prTIres /h1 cfRes_prTIred scalerK ?neq0CG ?irrK. +Qed. + +End OneMaximal. + +(* This is Peterfalvi (8.16). *) +Lemma FTtypeII_ker_TI M : + M \in 'M -> FTtype M == 2 -> + [/\ normedTI 'A0(M) G M, normedTI 'A(M) G M & normedTI 'A1(M) G M]. +Proof. +move=> maxM typeM; have [sA1A sAA0] := (FTsupp1_sub maxM, FTsupp_sub0 M). +have [sA10 sA0M] := (subset_trans sA1A sAA0, FTsupp0_sub M). +have nzA1: 'A1(M) != set0 by rewrite setD_eq0 def_FTcore ?subG1 ?Msigma_neq1. +have [nzA nzA0] := (subset_neq0 sA1A nzA1, subset_neq0 sA10 nzA1). +suffices nTI_A0: normedTI 'A0(M) G M. + by rewrite nTI_A0 !(normedTI_S _ _ _ nTI_A0) // ?FTsupp_norm ?FTsupp1_norm. +have [U W W1 W2 defW [[MtypeP _ _ tiFM] _ _ _ _]] := FTtypeP 2 maxM typeM. +apply/(Dade_normedTI_P (FT_Dade0_hyp maxM)); split=> // x A0x. +rewrite /= def_FTsignalizer0 /'R_M //=; have [// | not_sCxM] := ifPn. +have [y cxy /negP[]] := subsetPn not_sCxM. +apply: subsetP cxy; rewrite -['C[x]]setTI (cent1_normedTI tiFM) //. +have /setD1P[ntx Ms_x]: x \in 'A1(M). + by have [_ [/subsetP-> // ]] := FTsupport_facts maxM; apply/setIdP. +rewrite !inE ntx (subsetP (Fcore_sub_Fitting M)) //. +by rewrite (Fcore_eq_FTcore _ _) ?(eqP typeM). +Qed. + +(* This is Peterfalvi, Theorem (8.17). *) +Theorem FT_Dade_support_partition : + [/\ (*a1*) + \pi(G) =i [pred p | [exists M : {group gT} in 'M, p \in \pi(M`_\s)]], + (*a2*) {in 'M &, forall M L, + gval L \notin M :^: G -> coprime #|M`_\s| #|L`_\s| }, + (*b*) {in 'M, forall M, #|'A1~(M)| = (#|M`_\s|.-1 * #|G : M|)%N} + & (*c*) let PG := [set 'A1~(Mi) | Mi : {group gT} in 'M^G] in + [/\ {in 'M^G &, injective (fun M => 'A1~(M))}, + all_FTtype1 -> partition PG G^# + & forall S T W W1 W2 (defW : W1 \x W2 = W), + let VG := class_support (cyclicTIset defW) G in + typeP_pair S T defW -> partition (VG |: PG) G^# /\ VG \notin PG]]. +Proof. +have defDsup M: M \in 'M -> class_support M^~~ G = 'A1~(M). + move=> maxM; rewrite class_supportEr /'A1~(M) /'A1(M) def_FTcore //. + rewrite -(eq_bigr _ (fun _ _ => bigcupJ _ _ _ _)) exchange_big /=. + apply: eq_bigr => x Ms_x; rewrite -class_supportEr. + rewrite -norm_rlcoset ?(subsetP (cent_sub _)) ?cent_FT_signalizer //=. + congr (class_support (_ :* x) G); rewrite /'R_M. + have [_ _ /(_ x Ms_x)[_ defCx _] /(_ x Ms_x)defNF]:= BGsummaryD maxM. + have [sCxM | /defNF[[_ <-]] //] := ifPn. + apply/eqP; rewrite trivg_card1 -(eqn_pmul2r (cardG_gt0 'C_M[x])). + by rewrite (sdprod_card defCx) mul1n /= (setIidPr _). +have [b [a1 a2] [/and3P[_ _ not_PG_set0] _ _]] := BGsummaryE gT. +split=> [p | M L maxM maxL /a2 | M maxM | {b a1 a2}PG]. +- apply/idP/exists_inP=> [/a1[M maxM sMp] | [M _]]. + by exists M => //; rewrite def_FTcore // pi_Msigma. + exact: piSg (subsetT _) p. +- move/(_ maxM maxL)=> coML; rewrite coprime_pi' // !def_FTcore //. + apply: sub_pgroup (pcore_pgroup _ L) => p; apply/implyP. + by rewrite implybN /= pi_Msigma // implybE -negb_and [_ && _]coML. +- by rewrite -defDsup // def_FTcore // b. +have [/subsetP sMG_M _ injMG sM_MG] := mmax_transversalP gT. +have{PG} ->: PG = [set class_support M^~~ G | M : {group gT} in 'M]. + apply/setP=> AG; apply/imsetP/imsetP=> [] [M maxM ->]. + by move/sMG_M in maxM; exists M; rewrite ?defDsup //. + have [x MG_Mx] := sM_MG M maxM. + by exists (M :^ x)%G; rewrite // defDsup ?mmaxJ ?FT_Dade1_supportJ. +have [c1 c2] := mFT_partition gT. +split=> [M H maxM maxH eq_MH | Gtype1 | S T W W1 W2 defW VG pairST]. +- apply: injMG => //; move/sMG_M in maxM; move/sMG_M in maxH. + apply/orbit_transl/idPn => not_HG_M. + have /negP[]: ~~ [disjoint 'A1~(M) & 'A1~(H)]. + rewrite eq_MH -setI_eq0 setIid -defDsup //. + by apply: contraNneq not_PG_set0 => <-; exact: mem_imset. + rewrite -!defDsup // -setI_eq0 class_supportEr big_distrl -subset0. + apply/bigcupsP=> x /class_supportGidr <- /=; rewrite -conjIg sub_conjg conj0g. + rewrite class_supportEr big_distrr /=; apply/bigcupsP=> {x}x _. + rewrite subset0 setI_eq0 -sigma_supportJ sigma_support_disjoint ?mmaxJ //. + by rewrite (orbit_transr _ (mem_orbit _ _ _)) ?in_setT // orbit_sym. +- rewrite c1 // setD_eq0; apply/subsetP=> M maxM. + by rewrite FTtype_Fmax ?(forall_inP Gtype1). +have [[[cycW maxS _] _ _ _ _] [U_S StypeP]] := (pairST, typeP_pairW pairST). +have Stype'1 := FTtypeP_neq1 maxS StypeP. +have maxP_S: S \in TypeP_maxgroups _ by rewrite FTtype_Pmax. +have hallW1: \kappa(S).-Hall(S) W1. + have [[U1 K] /= complU1] := kappa_witness maxS. + have ntK: K :!=: 1%g by rewrite -(trivgPmax maxS complU1). + have [[defS_K _ _] [//|defS' _] _ _ _] := kappa_structure maxS complU1. + rewrite {}defS' in defS_K. + have /imsetP[x Sx defK] := of_typeP_compl_conj StypeP defS_K. + by have [_ hallK _] := complU1; rewrite defK pHallJ in hallK. +have{cycW} [[ntW1 ntW2] [cycW _ _]] := (cycTI_nontrivial cycW, cycW). +suffices defW2: 'C_(S`_\sigma)(W1) = W2. + by have [] := c2 _ _ maxP_S hallW1; rewrite defW2 /= (dprodWY defW). +have [U1 complU1] := ex_kappa_compl maxS hallW1. +have [[_ [_ _ sW2'F] _] _ _ _] := BGsummaryC maxS complU1 ntW1. +rewrite -(setIidPr sW2'F) setIA (setIidPl (Fcore_sub_Msigma maxS)). +exact: typeP_cent_core_compl StypeP. +Qed. + +(* This is Peterfalvi (8.18). Note that part (a) is not actually used later. *) +Lemma FT_Dade_support_disjoint S T : + S \in 'M -> T \in 'M -> gval T \notin S :^: G -> + [/\ (*a*) FTsupports S T = ~~ [disjoint 'A1(S) & 'A(T)] + /\ {in 'A1(S) :&: 'A(T), forall x, + ~~ ('C[x] \subset S) /\ 'C[x] \subset T}, + (*b*) [exists x, FTsupports S (T :^ x)] = ~~ [disjoint 'A1~(S) & 'A~(T)] + & (*c*) [disjoint 'A1~(S) & 'A~(T)] \/ [disjoint 'A1~(T) & 'A~(S)]]. +Proof. +move: S T; pose NC S T := gval T \notin S :^: G. +have part_a2 S T (maxS : S \in 'M) (maxT : T \in 'M) (ncST : NC S T) : + {in 'A1(S) :&: 'A(T), forall x, ~~ ('C[x] \subset S) /\ 'C[x] \subset T}. +- move=> x /setIP[/setD1P[ntx Ss_x] ATx]. + have coxTs: coprime #[x] #|T`_\s|. + apply: (coprime_dvdl (order_dvdG Ss_x)). + by have [_ ->] := FT_Dade_support_partition. + have [z /setD1P[ntz Ts_z] /setD1P[_ /setIP[Tn_x czx]]] := bigcupP ATx. + set n := FTtype T != 1%N in Tn_x. + have typeT: FTtype T == n.+1. + have notTs_x: x \notin T`_\s. + apply: contra ntx => Ts_x. + by rewrite -order_eq1 -dvdn1 -(eqnP coxTs) dvdn_gcd dvdnn order_dvdG. + apply: contraLR ATx => typeT; rewrite FTsupp_eq1 // ?inE ?ntx //. + move: (FTtype_range T) typeT; rewrite -mem_iota /n. + by do 5!case/predU1P=> [-> // | ]. + have defTs: T`_\s = T`_\F. + by apply/esym/Fcore_eq_FTcore; rewrite // (eqP typeT); case n. + have [U Ux defTn]: exists2 U : {group gT}, x \in U & T`_\F ><| U = T^`(n)%g. + have [[U K] /= complU] := kappa_witness maxT. + have defTn: T`_\s ><| U = T^`(n)%g. + by rewrite def_FTcore // (sdprod_FTder maxT complU). + have nsTsTn: T`_\s <| T^`(n)%g by case/sdprod_context: defTn. + have [sTsTn nTsTn] := andP nsTsTn. + have hallTs: \pi(T`_\s).-Hall(T^`(n)%g) T`_\s. + by rewrite defTs (pHall_subl _ (der_sub n T) (Fcore_Hall T)) //= -defTs. + have hallU: \pi(T`_\s)^'.-Hall(T^`(n)%g) U. + by apply/sdprod_Hall_pcoreP; rewrite /= (normal_Hall_pcore hallTs). + have solTn: solvable T^`(n)%g := solvableS (der_sub n T) (mmax_sol maxT). + rewrite coprime_sym coprime_pi' // in coxTs. + have [|y Tn_y] := Hall_subJ solTn hallU _ coxTs; rewrite cycle_subG //. + exists (U :^ y)%G; rewrite // -defTs. + by rewrite -(normsP nTsTn y Tn_y) -sdprodJ defTn conjGid. + have uniqCx: 'M('C[x]) = [set T]. + have:= FTtypeI_II_facts maxT typeT defTn; rewrite !ltnS leq_b1 -cent_set1. + case=> _ -> //; first by rewrite -cards_eq0 cards1. + by rewrite sub1set !inE ntx. + by apply/trivgPn; exists z; rewrite //= -defTs inE Ts_z cent_set1 cent1C. + split; last by case/mem_uniq_mmax: uniqCx. + by apply: contra ncST => /(eq_uniq_mmax uniqCx maxS)->; exact: orbit_refl. +have part_a1 S T (maxS : S \in 'M) (maxT : T \in 'M) (ncST : NC S T) : + FTsupports S T = ~~ [disjoint 'A1(S) & 'A(T)]. +- apply/existsP/pred0Pn=> [[x /and3P[ASx not_sCxS sCxT]] | [x /andP[A1Sx Atx]]]. + have [_ [/subsetP]] := FTsupport_facts maxS; set D := finset _. + have Dx: x \in D by rewrite !inE ASx. + move=> /(_ x Dx) A1x /(_ x Dx)uniqCx /(_ x Dx)[_ _ /setDP[ATx _] _]. + by rewrite (eq_uniq_mmax uniqCx maxT sCxT); exists x; exact/andP. + exists x; rewrite (subsetP (FTsupp1_sub maxS)) //=. + by apply/andP/part_a2=> //; exact/setIP. +have part_b S T (maxS : S \in 'M) (maxT : T \in 'M) (ncST : NC S T) : + [exists x, FTsupports S (T :^ x)] = ~~ [disjoint 'A1~(S) & 'A~(T)]. +- apply/existsP/pred0Pn=> [[x] | [y /andP[/= A1GSy AGTy]]]. + rewrite part_a1 ?mmaxJ // => [/pred0Pn[y /andP/=[A1Sy ATyx]]|]; last first. + by rewrite /NC -(rcoset_id (in_setT x)) orbit_rcoset. + rewrite FTsuppJ mem_conjg in ATyx; exists (y ^ x^-1); apply/andP; split. + by apply/bigcupP; exists y => //; rewrite mem_imset2 ?rcoset_refl ?inE. + apply/bigcupP; exists (y ^ x^-1) => //. + by rewrite mem_class_support ?rcoset_refl. + have{AGTy} [x2 ATx2 x2R_yG] := bigcupP AGTy. + have [sCx2T | not_sCx2T] := boolP ('C[x2] \subset T); last first. + have [_ _ _ [injA1G pGI pGP]] := FT_Dade_support_partition. + have{pGI pGP} tiA1g: trivIset [set 'A1~(M) | M : {group gT} in 'M^G]. + case: FTtypeP_pair_cases => [/forall_inP/pGI/and3P[] // | [M [L]]]. + by case=> _ W W1 W2 defW1 /pGP[]/and3P[_ /(trivIsetS (subsetUr _ _))]. + have [_ _ injMG sM_MG] := mmax_transversalP gT. + have [_ [sDA1T _] _] := FTsupport_facts maxT. + have [[z1 maxSz] [z2 maxTz]] := (sM_MG S maxS, sM_MG T maxT). + case/imsetP: ncST; exists (z1 * z2^-1)%g; first by rewrite inE. + rewrite conjsgM; apply/(canRL (conjsgK _))/congr_group/injA1G=> //. + apply/eqP/idPn=> /(trivIsetP tiA1g)/pred0Pn[]; try exact: mem_imset. + exists y; rewrite !FT_Dade1_supportJ /= A1GSy andbT. + by apply/bigcupP; exists x2; rewrite // (subsetP sDA1T) ?inE ?ATx2. + have{x2R_yG} /imsetP[z _ def_y]: y \in x2 ^: G. + by rewrite /'R_T {}sCx2T mul1g class_support_set1l in x2R_yG. + have{A1GSy} [x1 A1Sx1] := bigcupP A1GSy; rewrite {y}def_y -mem_conjgV. + rewrite class_supportGidr ?inE {z}//. + case/imset2P=> _ z /rcosetP[y Hy ->] _ def_x2. + exists z^-1%g; rewrite part_a1 ?mmaxJ //; last first. + by rewrite /NC (orbit_transr _ (mem_orbit _ _ _)) ?inE. + apply/pred0Pn; exists x1; rewrite /= A1Sx1 FTsuppJ mem_conjgV; apply/bigcupP. + pose ddS := FT_Dade1_hyp maxS; have [/andP[sA1S _] _ notA1_1 _ _] := ddS. + have [ntx1 Sx1] := (memPn notA1_1 _ A1Sx1, subsetP sA1S _ A1Sx1). + have [coHS defCx1] := (Dade_coprime ddS A1Sx1 A1Sx1, Dade_sdprod ddS A1Sx1). + rewrite def_FTsignalizer1 // in coHS defCx1. + have[u Ts_u /setD1P[_ cT'ux2]] := bigcupP ATx2. + exists u => {Ts_u}//; rewrite 2!inE -(conj1g z) (can_eq (conjgK z)) ntx1. + suffices{u cT'ux2} ->: x1 = (y * x1).`_(\pi('R_S x1)^'). + by rewrite -consttJ -def_x2 groupX. + have /setIP[_ /cent1P cx1y]: y \in 'C_G[x1]. + by case/sdprod_context: defCx1 => /andP[/subsetP->]. + rewrite consttM // (constt1P _) ?p_eltNK ?(mem_p_elt (pgroup_pi _)) // mul1g. + have piR'_Cx1: \pi('R_S x1)^'.-group 'C_S[x1] by rewrite coprime_pi' in coHS. + by rewrite constt_p_elt ?(mem_p_elt piR'_Cx1) // inE Sx1 cent1id. +move=> S T maxS maxT ncST; split; first split; auto. +apply/orP/idPn; rewrite negb_or -part_b // => /andP[suppST /negP[]]. +without loss{suppST} suppST: T maxT ncST / FTsupports S T. + move=> IH; case/existsP: suppST => x /IH {IH}. + rewrite FT_Dade1_supportJ (orbit_transr _ (mem_orbit _ _ _)) ?in_setT //. + by rewrite mmaxJ => ->. +have{suppST} [y /and3P[ASy not_sCyS sCyT]] := existsP suppST. +have Dy: y \in [set z in 'A0(S) | ~~ ('C[z] \subset S)] by rewrite !inE ASy. +have [_ [_ /(_ y Dy) uCy] /(_ y Dy)[_ coTcS _ typeT]] := FTsupport_facts maxS. +rewrite -mem_iota -(eq_uniq_mmax uCy maxT sCyT) !inE in coTcS typeT. +apply/negbNE; rewrite -part_b /NC 1?orbit_sym // negb_exists. +apply/forallP=> x; rewrite part_a1 ?mmaxJ ?negbK //; last first. + by rewrite /NC (orbit_transr _ (mem_orbit _ _ _)) ?in_setT // orbit_sym. +rewrite -setI_eq0 -subset0 FTsuppJ -bigcupJ big_distrr; apply/bigcupsP=> z Sxz. +rewrite conjD1g /= -setDIl coprime_TIg ?setDv //= cardJg. +rewrite -(Fcore_eq_FTcore maxT _) ?inE ?orbA; last by have [->] := typeT. +by rewrite (coprimegS _ (coTcS z _)) ?(subsetP (FTsupp1_sub0 _)) ?setSI ?gFsub. +Qed. + +(* A corollary to the above, which Peterfalvi derives from (8.17a) (i.e., *) +(* FT_Dade_support_partition) in the proof of (12.16). *) +Lemma FT_Dade1_support_disjoint S T : + S \in 'M -> T \in 'M -> gval T \notin S :^: G -> [disjoint 'A1~(S) & 'A1~(T)]. +Proof. +move=> maxS maxT /FT_Dade_support_disjoint[] // _ _ tiA1A. +without loss{tiA1A maxT}: S T maxS / [disjoint 'A1~(T) & 'A~(S)]. + by move=> IH_ST; case: tiA1A => /IH_ST; first rewrite disjoint_sym; apply. +by rewrite disjoint_sym; apply/disjoint_trans/FT_Dade_supportS/FTsupp1_sub. +Qed. + +End Eight. + +Notation FT_Dade0 maxM := (Dade (FT_Dade0_hyp maxM)). +Notation FT_Dade maxM := (Dade (FT_Dade_hyp maxM)). +Notation FT_Dade1 maxM := (Dade (FT_Dade1_hyp maxM)). +Notation FT_DadeF maxM := (Dade (FT_DadeF_hyp maxM)). + diff --git a/mathcomp/odd_order/PFsection9.v b/mathcomp/odd_order/PFsection9.v new file mode 100644 index 0000000..361d5fe --- /dev/null +++ b/mathcomp/odd_order/PFsection9.v @@ -0,0 +1,2205 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime binomial ssralg poly finset. +Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +Require Import gfunctor gproduct cyclic commutator center gseries nilpotent. +Require Import pgroup sylow hall abelian maximal frobenius. +Require Import matrix mxalgebra mxrepresentation mxabelem vector. +Require Import BGsection1 BGsection3 BGsection7 BGsection15 BGsection16. +Require Import algC classfun character inertia vcharacter. +Require Import PFsection1 PFsection2 PFsection3 PFsection4. +Require Import PFsection5 PFsection6 PFsection8. + +(******************************************************************************) +(* This file covers Peterfalvi, Section 9: On the maximal subgroups of Types *) +(* II, III and IV. For defW : W1 \x W2 = W, MtypeP : of_typeP M U defW, and *) +(* H := M`_\F we define : *) +(* Ptype_Fcore_kernel MtypeP == a maximal normal subgroup of M contained *) +(* (locally) H0 in H and containing 'C_H(U), provided M is *) +(* not a maximal subgroup of type V. *) +(* Ptype_Fcore_kernel MtypeP == the stabiliser of Hbar := H / H0 in U; this *) +(* (locally to this file) C is locked for performance reasons. *) +(* typeP_Galois MtypeP <=> U acts irreducibly on Hbar; this implies *) +(* that M / H0C is isomorphic to a Galois group *) +(* acting on the semidirect product of the *) +(* additive group of a finite field with a *) +(* a subgroup of its multiplicative group. *) +(* --> This predicate reflects alternative (b) in Peterfalvi (9.7). *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory FinRing.Theory. + +Section Nine. + +Variable gT : minSimpleOddGroupType. +Local Notation G := (TheMinSimpleOddGroup gT). +Implicit Types (p q : nat) (x y z : gT). +Implicit Types H K L N P Q R S T U V W : {group gT}. + +(* Peterfalvi (9.1) is covered by BGsection3.Frobenius_Wielandt_fixpoint. *) + +(* These assumptions correspond to Peterfalvi, Hypothesis (9.2). *) + +Variables M U W W1 W2 : {group gT}. +Hypotheses (maxM : M \in 'M) (defW : W1 \x W2 = W) (MtypeP: of_typeP M U defW). +Hypothesis notMtype5 : FTtype M != 5. + +Local Notation "` 'M'" := (gval M) (at level 0, only parsing) : group_scope. +Local Notation "` 'U'" := (gval U) (at level 0, only parsing) : group_scope. +Local Notation "` 'W1'" := (gval W1) (at level 0, only parsing) : group_scope. +Local Notation H := `M`_\F%G. +Local Notation "` 'H'" := `M`_\F (at level 0) : group_scope. +Local Notation "` 'W2'" := (gval W2) (at level 0, only parsing) : group_scope. +Local Notation HU := M^`(1)%G. +Local Notation "` 'HU'" := `M^`(1) (at level 0) : group_scope. +Local Notation U' := U^`(1)%G. +Local Notation "` 'U''" := `U^`(1) (at level 0) : group_scope. + +Let q := #|W1|. + +Let defM : HU ><| W1 = M. Proof. by have [[]] := MtypeP. Qed. +Let defHU : H ><| U = HU. Proof. by have [_ []] := MtypeP. Qed. +Let nUW1 : W1 \subset 'N(U). Proof. by have [_ []] := MtypeP. Qed. +Let cHU' : U' \subset 'C(H). Proof. by have [_ []] := typeP_context MtypeP. Qed. + +Let notMtype1 : FTtype M != 1%N. Proof. exact: FTtypeP_neq1 MtypeP. Qed. + +Local Notation Mtype24 := (compl_of_typeII_IV maxM MtypeP notMtype5). +Let ntU : U :!=: 1. Proof. by have [] := Mtype24. Qed. +Let pr_q : prime q. Proof. by have [] := Mtype24. Qed. +Let ntW2 : W2 :!=: 1. Proof. by have [_ _ _ []] := MtypeP. Qed. +Let sW2H : W2 \subset H. Proof. by have [_ _ _ []] := MtypeP. Qed. +Let defW2 : 'C_H(W1) = W2. Proof. exact: typeP_cent_core_compl MtypeP. Qed. + +Lemma Ptype_Fcore_sdprod : H ><| (U <*> W1) = M. +Proof. +have [_ /= sW1M mulHUW1 _ tiHUW1] := sdprod_context defM. +have [/= /andP[sHHU _] sUHU mulHU nHU tiHU] := sdprod_context defHU. +rewrite sdprodE /= norm_joinEr // ?mulgA ?mulHU //. + by rewrite mulG_subG nHU (subset_trans sW1M) ?gFnorm. +rewrite setIC -(setIidPr sHHU) setIA -group_modl //. +by rewrite (setIC W1) tiHUW1 mulg1 setIC tiHU. +Qed. +Local Notation defHUW1 := Ptype_Fcore_sdprod. + +Lemma Ptype_Fcore_coprime : coprime #|H| #|U <*> W1|. +Proof. +by rewrite (coprime_sdprod_Hall_l defHUW1) ?(pHall_Hall (Fcore_Hall M)). +Qed. +Let coH_UW1 := Ptype_Fcore_coprime. +Let coHU : coprime #|H| #|U|. +Proof. exact: coprimegS (joing_subl U W1) coH_UW1. Qed. + +Let not_cHU : ~~ (U \subset 'C(H)). +Proof. by have [_ [_ ->]] := typeP_context MtypeP. Qed. + +Lemma Ptype_compl_Frobenius : [Frobenius U <*> W1 = U ><| W1]. +Proof. +have [[_ _ ntW1 _] _ _ [_ _ _ _ prHU_W1] _] := MtypeP. +have [[_ _ _ tiHUW1] [_ sUHU _ _ tiHU]] := (sdprodP defM, sdprod_context defHU). +apply/Frobenius_semiregularP=> // [|x /prHU_W1 defCx]. + by rewrite sdprodEY //; apply/trivgP; rewrite -tiHUW1 setSI. +by apply/trivgP; rewrite -tiHU /= -{1}(setIidPr sUHU) setIAC defCx setSI. +Qed. +Local Notation frobUW1 := Ptype_compl_Frobenius. + +Let nilH : nilpotent H. Proof. exact: Fcore_nil. Qed. +Let solH : solvable H. Proof. exact: nilpotent_sol. Qed. + +(* This is Peterfalvi (9.3). *) +Lemma typeII_IV_core (p := #|W2|) : + if FTtype M == 2 then 'C_H(U) = 1 /\ #|H| = (#|W2| ^ q)%N + else [/\ prime p, 'C_H(U <*> W1) = 1 & #|H| = (p ^ q * #|'C_H(U)|)%N]. +Proof. +have [_ _ nHUW1 _] := sdprodP defHUW1. +have /= [oH _ oH1] := Frobenius_Wielandt_fixpoint frobUW1 nHUW1 coH_UW1 solH. +have [Mtype2 {oH}| notMtype2 {oH1}] := boolP (FTtype M == 2). + suffices regHU: 'C_H(U) = 1 by rewrite -defW2 oH1. + have [_ _ _ HUtypeF defHUF] := compl_of_typeII maxM MtypeP Mtype2. + have [_ _ [U0 [sU0U _]]] := HUtypeF; rewrite {}defHUF => frobHU0. + have /set0Pn[x U0x]: U0^# != set0. + by rewrite setD_eq0 subG1; case/Frobenius_context: frobHU0. + apply/trivgP; rewrite -(Frobenius_reg_ker frobHU0 U0x) setIS // -cent_cycle. + by rewrite centS // cycle_subG (subsetP sU0U) //; case/setD1P: U0x. +have p_pr: prime p. + have [S pairMS [xdefW [U_S StypeP]]] := FTtypeP_pair_witness maxM MtypeP. + have [[_ _ maxS] _] := pairMS; rewrite {1}(negPf notMtype2) /= => Stype2 _ _. + by have [[]] := compl_of_typeII maxS StypeP Stype2. +rewrite -/q -/p centY setICA defW2 setIC in oH *. +suffices regW2U: 'C_W2(U) = 1 by rewrite -oH regW2U cards1 exp1n mul1n. +apply: prime_TIg => //=; apply: contra not_cHU => /setIidPl cUW2. +rewrite centsC (sameP setIidPl eqP) eqEcard subsetIl. +by rewrite -(@leq_pmul2l (p ^ q)) -?oH ?cUW2 //= expn_gt0 cardG_gt0. +Qed. + +(* Existential witnesses for Peterfalvi (9.4). *) +Definition Ptype_Fcore_kernel of of_typeP M U defW := + odflt 1%G [pick H0 : {group gT} | chief_factor M H0 H & 'C_H(U) \subset H0]. +Let H0 := (Ptype_Fcore_kernel MtypeP). +Local Notation "` 'H0'" := (gval H0) (at level 0, only parsing) : group_scope. +Local Notation Hbar := (H / `H0)%G. +Local Notation "` 'Hbar'" := (`H / `H0)%g (at level 0) : group_scope. +Let p := pdiv #|Hbar|. + +(* This corresponds to Peterfalvi (9.4). *) +Lemma Ptype_Fcore_kernel_exists : chief_factor M H0 H /\ 'C_H(U) \subset H0. +Proof. +pose S := <<class_support 'C_H(U) H>> . +suffices [H1 maxH sCH1]: {H1 : {group gT} | maxnormal H1 H M & S \subset H1}. + apply/andP; rewrite /H0 /Ptype_Fcore_kernel; case: pickP => // /(_ H1)/idP[]. + rewrite /chief_factor maxH Fcore_normal (subset_trans _ sCH1) ?sub_gen //. + exact: sub_class_support. +apply/maxgroup_exists/andP; split. + have snCH: 'C_H(U) <|<| H by rewrite nilpotent_subnormal ?subsetIl. + by have [/setIidPl/idPn[] | // ] := subnormalEsupport snCH; rewrite centsC. +have [_ {3}<- nHUW1 _] := (sdprodP defHUW1). +rewrite norms_gen // mulG_subG class_support_norm norms_class_support //. +by rewrite normsI ?norms_cent // join_subG normG. +Qed. + +Let chiefH0 : chief_factor M H0 H. +Proof. by have [] := Ptype_Fcore_kernel_exists. Qed. +Let ltH0H : H0 \proper H. +Proof. by case/andP: chiefH0 => /maxgroupp/andP[]. Qed. +Let nH0M : M \subset 'N(H0). +Proof. by case/andP: chiefH0 => /maxgroupp/andP[]. Qed. +Let sH0H : H0 \subset H. Proof. exact: proper_sub ltH0H. Qed. +Let nsH0M : H0 <| M. Proof. by rewrite /normal (subset_trans sH0H) ?gFsub. Qed. +Let nsH0H : H0 <| H. Proof. by rewrite (normalS _ (Fcore_sub _)). Qed. +Let minHbar : minnormal Hbar (M / H0). +Proof. exact: chief_factor_minnormal. Qed. +Let ntHbar : Hbar :!=: 1. Proof. by case/mingroupp/andP: minHbar. Qed. +Let solHbar: solvable Hbar. Proof. by rewrite quotient_sol. Qed. +Let abelHbar : p.-abelem Hbar. +Proof. by have [] := minnormal_solvable minHbar _ solHbar. Qed. +Let p_pr : prime p. Proof. by have [/pgroup_pdiv[]] := and3P abelHbar. Qed. +Let abHbar : abelian Hbar. Proof. exact: abelem_abelian abelHbar. Qed. + +(* This is Peterfalvi, Hypothesis (9.5). *) +Fact Ptype_Fcompl_kernel_key : unit. Proof. by []. Qed. +Definition Ptype_Fcompl_kernel := + locked_with Ptype_Fcompl_kernel_key 'C_U(Hbar | 'Q)%G. +Local Notation C := Ptype_Fcompl_kernel. +Local Notation "` 'C'" := (gval C) (at level 0, only parsing) : group_scope. +Local Notation Ubar := (U / `C)%G. +Local Notation "` 'Ubar'" := (`U / `C)%g (at level 0) : group_scope. +Local Notation W1bar := (W1 / `H0)%G. +Local Notation "` 'W1bar'" := (`W1 / `H0)%g (at level 0) : group_scope. +Local Notation W2bar := 'C_Hbar(`W1bar)%G. +Local Notation "` 'W2bar'" := 'C_`Hbar(`W1bar) (at level 0) : group_scope. +Let c := #|C|. +Let u := #|Ubar|. +Local Notation tau := (FT_Dade0 maxM). +Local Notation "chi ^\tau" := (tau chi). +Let calX := Iirr_kerD M^`(1) H 1. +Let calS := seqIndD M^`(1) M M`_\F 1. +Let X_ Y := Iirr_kerD M^`(1) H Y. +Let S_ Y := seqIndD M^`(1) M M`_\F Y. + +Local Notation inMb := (coset (gval H0)). + +Local Notation H0C := (`H0 <*> `C)%G. +Local Notation "` 'H0C'" := (`H0 <*> `C) (at level 0) : group_scope. +Local Notation HC := (`H <*> `C)%G. +Local Notation "` 'HC'" := (`H <*> `C) (at level 0) : group_scope. +Local Notation H0U' := (`H0 <*> `U')%G. +Local Notation "` 'H0U''" := (gval H0 <*> `U')%G (at level 0) : group_scope. +Local Notation H0C' := (`H0 <*> `C^`(1)%g)%G. +Local Notation "` 'H0C''" := (`H0 <*> `C^`(1)) (at level 0) : group_scope. + +Let defW2bar : W2bar :=: W2 / H0. +Proof. +rewrite -defW2 coprime_quotient_cent ?(subset_trans _ nH0M) //. + by have [_ /mulG_sub[]] := sdprodP defM. +exact: coprimegS (joing_subr _ _) coH_UW1. +Qed. + +Let sCU : C \subset U. Proof. by rewrite [C]unlock subsetIl. Qed. + +Let nsCUW1 : C <| U <*> W1. +Proof. +have [_ sUW1M _ nHUW1 _] := sdprod_context defHUW1. +rewrite /normal [C]unlock subIset ?joing_subl // normsI //. + by rewrite join_subG normG. +rewrite /= astabQ norm_quotient_pre ?norms_cent ?quotient_norms //. +exact: subset_trans sUW1M nH0M. +Qed. + +Lemma Ptype_Fcore_extensions_normal : + [/\ H0C <| M, HC <| M, H0U' <| M & H0C' <| M]. +Proof. +have [nsHUM sW1M /mulG_sub[sHUM _] nHUW1 tiHUW1] := sdprod_context defM. +have [nsHHU sUHU /mulG_sub[sHHU _] nHU tiHU] := sdprod_context defHU. +have [sHM sUM] := (subset_trans sHHU sHUM, subset_trans sUHU sHUM). +have sCM: C \subset M := subset_trans sCU sUM. +have sH0C_M: H0C \subset M by rewrite /normal join_subG (subset_trans sH0H). +have [nH0C nH0_H0C] := (subset_trans sCM nH0M, subset_trans sH0C_M nH0M). +have nsH0C: H0C <| M. + rewrite /normal sH0C_M -{1}defM sdprodEY //= -defHU sdprodEY //= -joingA. + rewrite join_subG andbC normsY ?(normal_norm nsCUW1) //=; last first. + by rewrite (subset_trans _ nH0M) // join_subG sUM. + rewrite -quotientYK // -{1}(quotientGK nsH0H) morphpre_norms //= [C]unlock. + by rewrite cents_norm // centsC -quotient_astabQ quotientS ?subsetIr. +split=> //; first by rewrite /= -{1}(joing_idPl sH0H) -joingA normalY ?gFnormal. + rewrite normalY // /normal (subset_trans (der_sub 1 U)) //=. + rewrite -{1}defM sdprodEY //= -defHU sdprodEY //=. + rewrite !join_subG gFnorm cents_norm 1?centsC //=. + by rewrite (char_norm_trans (der_char _ _)). +suffices ->: H0C' :=: H0 <*> H0C^`(1). + by rewrite normalY ?(char_normal_trans (der_char _ _)). +rewrite /= -?quotientYK ?(subset_trans (der_sub _ _)) ?subsetIl //=. +by rewrite !quotient_der ?cosetpreK ?subsetIl. +Qed. +Local Notation nsH0xx_M := Ptype_Fcore_extensions_normal. + +Let Du : u = #|HU : HC|. +Proof. +have nCU := subset_trans (joing_subl U W1) (normal_norm nsCUW1). +by rewrite -(index_sdprodr defHU) -?card_quotient. +Qed. + +(* This is Peterfalvi (9.6). *) +Lemma Ptype_Fcore_factor_facts : + [/\ C :!=: U, #|W2bar| = p & #|Hbar| = p ^ q]%N. +Proof. +have [defUW1 _ ntW1 _ _] := Frobenius_context Ptype_compl_Frobenius. +have coHW1: coprime #|H| #|W1| := coprimegS (joing_subr U W1) coH_UW1. +have [_ sUW1M _ nHUW1 _] := sdprod_context defHUW1. +have nH0UW1 := subset_trans sUW1M nH0M; have [nH0U nH0W1] := joing_subP nH0UW1. +have regUHb: 'C_Hbar(U / H0) = 1. + have [_ sCH0] := Ptype_Fcore_kernel_exists. + by rewrite -coprime_quotient_cent ?(nilpotent_sol nilH) ?quotientS1. +have ->: C != U. + apply: contraNneq ntHbar => defU; rewrite -subG1 -regUHb subsetIidl centsC. + by rewrite -defU [C]unlock -quotient_astabQ quotientS ?subsetIr. +have frobUW1b: [Frobenius U <*> W1 / H0 = (U / H0) ><| W1bar]. + have tiH0UW1 := coprime_TIg (coprimeSg sH0H coH_UW1). + have /isomP[inj_f im_f] := quotient_isom nH0UW1 tiH0UW1. + have:= injm_Frobenius (subxx _) inj_f frobUW1. + by rewrite im_f !morphim_restrm !(setIidPr _) ?joing_subl ?joing_subr. +have{frobUW1b} oHbar: #|Hbar| = (#|W2bar| ^ q)%N. + have nHbUW1 : U <*> W1 / H0 \subset 'N(Hbar) := quotient_norms H0 nHUW1. + have coHbUW1 : coprime #|Hbar| #|U <*> W1 / H0| by apply: coprime_morph. + have [//|_ _ -> //] := Frobenius_Wielandt_fixpoint frobUW1b nHbUW1 coHbUW1 _. + by rewrite -(card_isog (quotient_isog _ _)) // coprime_TIg ?(coprimeSg sH0H). +have abelW2bar: p.-abelem W2bar := abelemS (subsetIl _ _) abelHbar. +rewrite -(part_pnat_id (abelem_pgroup abelW2bar)) p_part in oHbar *. +suffices /eqP cycW2bar: logn p #|W2bar| == 1%N by rewrite oHbar cycW2bar. +have cycW2: cyclic W2 by have [_ _ _ []] := MtypeP. +rewrite eqn_leq -abelem_cyclic //= -/W2bar {1}defW2bar quotient_cyclic //=. +rewrite lt0n; apply: contraNneq ntHbar => W2bar1. +by rewrite trivg_card1 oHbar W2bar1 exp1n. +Qed. + +Lemma def_Ptype_factor_prime : prime #|W2| -> p = #|W2|. +Proof. +move=> prW2; suffices: p \in \pi(W2) by rewrite !(primes_prime, inE) // => /eqP. +rewrite mem_primes p_pr cardG_gt0; have [_ <- _] := Ptype_Fcore_factor_facts. +by rewrite defW2bar dvdn_quotient. +Qed. + +(* The first assertion of (9.4)(b) (the rest is subsumed by (9.6)). *) +Lemma typeIII_IV_core_prime : FTtype M != 2 -> p = #|W2|. +Proof. +by have:= typeII_IV_core => /=; case: ifP => // _ [/def_Ptype_factor_prime]. +Qed. + +Let frobUW1c : [Frobenius U <*> W1 / C = Ubar ><| W1 / C]. +Proof. +apply: Frobenius_quotient frobUW1 _ nsCUW1 _. + by apply: nilpotent_sol; have [_ []] := MtypeP. +by have [] := Ptype_Fcore_factor_facts; rewrite eqEsubset sCU. +Qed. + +Definition typeP_Galois := acts_irreducibly U Hbar 'Q. + +(* This is Peterfalvi (9.7)(a). *) +Lemma typeP_Galois_Pn : + ~~ typeP_Galois -> + {H1 : {group coset_of H0} | + [/\ #|H1| = p, U / H0 \subset 'N(H1), [acts U, on H1 | 'Q], + \big[dprod/1]_(w in W1bar) H1 :^ w = Hbar + & let a := #|U : 'C_U(H1 | 'Q)| in + [/\ a > 1, a %| p.-1, cyclic (U / 'C_U(H1 | 'Q)) + & exists V : {group 'rV['Z_a]_q.-1}, Ubar \isog V]]}. +Proof. +have [_ sUW1M defHUW1 nHUW1 _] := sdprod_context defHUW1. +have [nHU nHW1] := joing_subP nHUW1. +have nH0UW1 := subset_trans sUW1M nH0M; have [nH0U nH0W1] := joing_subP nH0UW1. +rewrite /typeP_Galois acts_irrQ //= => not_minHbarU. +have [H1 minH1 sH1Hb]: {H1 | minnormal (gval H1) (U / H0) & H1 \subset Hbar}. + by apply: mingroup_exists; rewrite ntHbar quotient_norms. +exists H1; have [defH1 | ltH1H] := eqVproper sH1Hb. + by rewrite -defH1 minH1 in not_minHbarU. +have [/andP[ntH1 nH1U] _] := mingroupP minH1. +have actsUH1: [acts U, on H1 | 'Q]. + by rewrite -(cosetpreK H1) actsQ ?norm_quotient_pre. +have [nH0H [neqCU _ oHbar]] := (normal_norm nsH0H, Ptype_Fcore_factor_facts). +have nUW1b: W1bar \subset 'N(U / H0) by apply: quotient_norms. +have oW1b: #|W1bar| = q. + rewrite -(card_isog (quotient_isog _ _)) // coprime_TIg //. + by rewrite (coprimeSg sH0H) // (coprimegS (joing_subr U W1)). +have [oH1 defHbar]: #|H1| = p /\ \big[dprod/1]_(w in W1bar) H1 :^ w = Hbar. + have nHbUW1: U <*> W1 / H0 \subset 'N(Hbar) by apply: quotient_norms. + pose rUW1 := abelem_repr abelHbar ntHbar nHbUW1. + have irrUW1: mx_irreducible rUW1. + apply/abelem_mx_irrP/mingroupP; split=> [|H2]; first by rewrite ntHbar. + case/andP=> ntH2 nH2UW1 sH2H; case/mingroupP: minHbar => _; apply=> //. + by rewrite ntH2 -defHUW1 quotientMl // mulG_subG sub_abelian_norm. + have nsUUW1: U / H0 <| U <*> W1 / H0 by rewrite quotient_normal // normalYl. + pose rU := subg_repr rUW1 (normal_sub nsUUW1). + pose V1 := rowg_mx (abelem_rV abelHbar ntHbar @* H1). + have simV1: mxsimple rU V1 by apply/mxsimple_abelem_subg/mxsimple_abelemGP. + have [W0 /subsetP sW01 [sumW0 dxW0]] := Clifford_basis irrUW1 simV1. + have def_q: q = (#|W0| * \rank V1)%N. + transitivity (\rank (\sum_(w in W0) V1 *m rUW1 w))%R. + by rewrite sumW0 mxrank1 /= (dim_abelemE abelHbar) // oHbar pfactorK. + rewrite (mxdirectP dxW0) -sum_nat_const; apply: eq_bigr => x /sW01/= Wx. + by rewrite mxrankMfree ?row_free_unit ?repr_mx_unit. + have oH1: #|H1| = (p ^ \rank V1)%N. + by rewrite -{1}(card_Fp p_pr) -card_rowg rowg_mxK card_injm ?abelem_rV_injm. + have oW0: #|W0| = q. + apply/prime_nt_dvdP=> //; last by rewrite def_q dvdn_mulr. + apply: contraTneq (proper_card ltH1H) => trivW0. + by rewrite oHbar def_q trivW0 mul1n -oH1 ltnn. + have q_gt0 := prime_gt0 pr_q. + rewrite oH1 -(mulKn (\rank V1) q_gt0) -{1}oW0 -def_q divnn q_gt0. + have defHbar: \big[dprod/1]_(w in W0) H1 :^ w = Hbar. + have inj_rV_Hbar := rVabelem_injm abelHbar ntHbar. + have/(injm_bigdprod _ inj_rV_Hbar)/= := bigdprod_rowg sumW0 dxW0. + rewrite sub_im_abelem_rV rowg1 im_rVabelem => <- //=; apply: eq_bigr => w. + by move/sW01=> Ww; rewrite abelem_rowgJ ?rowg_mxK ?abelem_rV_mK. + have injW0: {in W0 &, injective (fun w => H1 :^ w)}. + move=> x y Wx Wy /= eq_Hxy; apply: contraNeq ntH1 => neq_xy. + rewrite -(conjsg_eq1 _ x) -[H1 :^ x]setIid {1}eq_Hxy; apply/eqP. + rewrite (bigD1 y) // (bigD1 x) /= ?Wx // dprodA in defHbar. + by case/dprodP: defHbar => [[_ _ /dprodP[_ _ _ ->] _]]. + have defH1W0: [set H1 :^ w | w in W0] = [set H1 :^ w | w in W1 / H0]. + apply/eqP; rewrite eqEcard (card_in_imset injW0) oW0 -oW1b leq_imset_card. + rewrite andbT; apply/subsetP=> _ /imsetP[w /sW01/= Ww ->]. + move: Ww; rewrite norm_joinEr ?quotientMl // => /mulsgP[x w1 Ux Ww1 ->]. + by rewrite conjsgM (normsP nH1U) // mem_imset. + have injW1: {in W1 / H0 &, injective (fun w => H1 :^ w)}. + by apply/imset_injP; rewrite -defH1W0 (card_in_imset injW0) oW0 oW1b. + by rewrite -(big_imset id injW1) -defH1W0 big_imset. +split=> //; set a := #|_ : _|; pose q1 := #|(W1 / H0)^#|. +have a_gt1: a > 1. + rewrite indexg_gt1 subsetIidl /= astabQ -sub_quotient_pre //. + apply: contra neqCU => cH1U; rewrite [C]unlock (sameP eqP setIidPl) /= astabQ. + rewrite -sub_quotient_pre // -(bigdprodWY defHbar) cent_gen centsC. + by apply/bigcupsP=> w Ww; rewrite centsC centJ -(normsP nUW1b w) ?conjSg. +have Wb1: 1 \in W1bar := group1 _. +have ->: q.-1 = q1 by rewrite -oW1b (cardsD1 1) Wb1. +have /cyclicP[h defH1]: cyclic H1 by rewrite prime_cyclic ?oH1. +have o_h: #[h] = p by rewrite defH1 in oH1. +have inj_Zp_h w := injm_Zp_unitm (h ^ w). +pose phi w := invm (inj_Zp_h w) \o restr_perm <[h ^ w]> \o actperm 'Q. +have dU w: w \in W1bar -> {subset U <= 'dom (phi w)}. + move=> Ww x Ux; have Qx := subsetP (acts_dom actsUH1) x Ux. + rewrite inE Qx /= im_Zp_unitm inE mem_morphpre //=; last first. + by apply: Aut_restr_perm (actperm_Aut 'Q _); rewrite //= quotientT. + rewrite cycleJ -defH1 !inE /=; apply/subsetP=> z H1w_z; rewrite inE actpermK. + rewrite qactJ (subsetP nH0U) ?memJ_norm // normJ mem_conjg. + by rewrite (subsetP nH1U) // -mem_conjg (normsP nUW1b) ?mem_quotient. +have sUD := introT subsetP (dU _ _). +have Kphi w: 'ker (phi w) = 'C(H1 :^ w | 'Q). + rewrite !ker_comp ker_invm -kerE ker_restr_perm defH1 -cycleJ. + apply/setP=> x; rewrite !inE; congr (_ && _) => /=. + by apply: eq_subset_r => h1; rewrite !inE actpermK. +have o_phiU w: w \in W1bar -> #|phi w @* U| = a. + move=> Ww; have [w1 Nw1 Ww1 def_w] := morphimP Ww. + rewrite card_morphim Kphi (setIidPr _) ?sUD // /a indexgI /= !astabQ. + by rewrite centJ def_w morphpreJ // -{1}(normsP nUW1 w1 Ww1) indexJg. +have a_dv_p1: a %| p.-1. + rewrite -(o_phiU 1) // (dvdn_trans (cardSg (subsetT _))) // card_units_Zp //. + by rewrite conjg1 o_h (@totient_pfactor p 1) ?muln1. +have cycZhw w: cyclic (units_Zp #[h ^ w]). + rewrite -(injm_cyclic (inj_Zp_h w)) // im_Zp_unitm Aut_prime_cyclic //=. + by rewrite -orderE orderJ o_h. +have cyc_phi1U: cyclic (phi 1 @* U) := cyclicS (subsetT _) (cycZhw 1). +split=> //; last have{cyc_phi1U a_dv_p1} [z def_z] := cyclicP cyc_phi1U. + by rewrite -(conjsg1 H1) -Kphi (isog_cyclic (first_isog_loc _ _)) ?sUD. +have o_hw w: #[h ^ w] = #[h ^ 1] by rewrite !orderJ. +pose phi1 w x := eq_rect _ (fun m => {unit 'Z_m}) (phi w x) _ (o_hw w). +have val_phi1 w x: val (phi1 w x) = val (phi w x) :> nat. + by rewrite /phi1; case: _ / (o_hw _). +have mem_phi1 w x: w \in W1bar -> x \in U -> phi1 w x \in <[z]>%G. + move=> Ww Ux; have: #|<[z]>%G| = a by rewrite /= -def_z o_phiU. + rewrite /phi1; case: _ / (o_hw w) <[z]>%G => A oA /=. + suffices <-: phi w @* U = A by rewrite mem_morphim // dU. + by apply/eqP; rewrite (eq_subG_cyclic (cycZhw w)) ?subsetT // oA o_phiU. +have o_z: #[z] = a by rewrite orderE -def_z o_phiU. +pose phi0 w x := ecast m 'Z_m o_z (invm (injm_Zpm z) (phi1 w x)). +pose psi x := (\row_(i < q1) (phi0 (enum_val i) x * (phi0 1 x)^-1)%g)%R. +have psiM: {in U &, {morph psi: x y / x * y}}. + have phi0M w: w \in W1bar -> {in U &, {morph phi0 w: x y / x * y}}. + move=> Ww x y Ux Uy; rewrite /phi0; case: (a) / (o_z) => /=. + rewrite -morphM; first 1 [congr (invm _ _)] || by rewrite im_Zpm mem_phi1. + by rewrite /phi1; case: _ / (o_hw w); rewrite /= -morphM ?dU. + move=> x y Ux Uy; apply/rowP=> i; have /setD1P[_ Ww] := enum_valP i. + by rewrite !{1}mxE !{1}phi0M // addrCA -addrA -opprD addrCA addrA. +suffices Kpsi: 'ker (Morphism psiM) = C. + by exists [group of Morphism psiM @* U]; rewrite /Ubar -Kpsi first_isog. +apply/esym/eqP; rewrite eqEsubset; apply/andP; split. + rewrite [C]unlock -(bigdprodWY defHbar); apply/subsetP=> x /setIP[Ux cHx]. + suffices phi0x1 w: w \in W1bar -> phi0 w x = 1. + rewrite !inE Ux; apply/eqP/rowP=> i; have /setD1P[_ Ww] := enum_valP i. + by rewrite !mxE !phi0x1 ?mulgV. + move=> Ww; apply: val_inj; rewrite /phi0; case: (a) / (o_z); congr (val _). + suffices /eqP->: phi1 w x == 1 by rewrite morph1. + rewrite -2!val_eqE [val _]val_phi1 -(o_hw w) [phi _ _]mker // Kphi. + by apply: subsetP (astabS _ _) _ cHx; rewrite sub_gen // (bigcup_sup w). +have sKU: 'ker (Morphism psiM) \subset U by apply: subsetIl. +rewrite -quotient_sub1 -?(Frobenius_trivg_cent frobUW1c); last first. + by apply: subset_trans (normal_norm nsCUW1); rewrite subIset ?joing_subl. +rewrite subsetI quotientS //= quotient_cents2r // [C]unlock subsetI. +rewrite (subset_trans (commSg W1 sKU)) ?commg_subl //= astabQ gen_subG /=. +apply/subsetP=> _ /imset2P[x w1 Kx Ww1 ->]. +have:= Kx; rewrite -groupV 2!inE groupV => /andP[Ux /set1P/rowP psi_x'0]. +have [nH0x Ux'] := (subsetP nH0U x Ux, groupVr Ux); pose x'b := (inMb x)^-1. +rewrite mem_morphpre ?groupR ?morphR //= ?(subsetP nH0W1) //. +have conj_x'b w: w \in W1bar -> (h ^ w) ^ x'b = (h ^ w) ^+ val (phi 1 x^-1). + move=> Ww; transitivity (Zp_unitm (phi w x^-1) (h ^ w)). + have /morphpreP[_ /morphpreP[Px' Rx']] := dU w Ww x^-1 Ux'. + rewrite invmK ?restr_permE ?cycle_id //. + by rewrite actpermE qactJ groupV nH0x morphV. + have:= Ww; rewrite -(setD1K Wb1) autE ?cycle_id // => /setU1P[-> // | W'w]. + have /eqP := psi_x'0 (enum_rank_in W'w w); rewrite 2!mxE enum_rankK_in //. + rewrite -eq_mulgV1 -val_eqE /phi0; case: (a) / (o_z); rewrite /= val_eqE. + rewrite (inj_in_eq (injmP (injm_invm _))) /= ?im_Zpm ?mem_phi1 //. + by rewrite -2!val_eqE /= !val_phi1 // => /eqP->. +rewrite -sub_cent1 -(bigdprodWY defHbar) gen_subG; apply/bigcupsP=> w2 Ww2. +rewrite defH1 -cycleJ cycle_subG cent1C inE conjg_set1 !conjgM // conj_x'b //. +rewrite conjXg -!conjgM -conj_x'b ?groupM ?groupV ?mem_quotient //. +by rewrite !conjgM !conjgKV. +Qed. + +(* This is Peterfalvi (9.7)(b). *) +(* Note that part of this statement feeds directly into the final chapter of *) +(* the proof (PFsection14 and BGappendixC) and is not used before; we have *) +(* thus chosen to formulate the statement of (9.7)(b) accordingly. *) +(* For example, we supply separately the three component of the semi-direct *) +(* product isomorphism, because no use is made of the global isomorphism. We *) +(* also state explicitly that the image of W2bar is Fp because this is the *) +(* fact used in B & G, Appendix C, it is readily available during the proof, *) +(* whereas it can only be derived from the original statement of (9.7)(b) by *) +(* using Galois theory. Indeed the Galois part of the isomorphism is only *) +(* needed for this -- so with the formulation below it will not be used. *) +(* In order to avoid the use of the Wedderburn theorem on finite division *) +(* rings we build the field F from the enveloping algebra of the *) +(* representation of U rather than its endomorphism ring: then the fact that *) +(* Ubar is abelian yields commutativity directly. *) +Lemma typeP_Galois_P : + typeP_Galois -> + {F : finFieldType & {phi : {morphism Hbar >-> F} + & {psi : {morphism U >-> {unit F}} & {eta : {morphism W1 >-> {perm F}} + & forall alpha : {perm F}, reflect (rmorphism alpha) (alpha \in eta @* W1) + & [/\ 'injm eta, {in Hbar & W1, morph_act 'Q 'P phi eta} + & {in U & W1, forall x w, val (psi (x ^ w)) = eta w (val (psi x))}]} + & 'ker psi = C /\ {in Hbar & U, morph_act 'Q 'U phi psi}} + & [/\ #|F| = (p ^ q)%N, isom Hbar [set: F] phi & phi @* W2bar = <[1%R : F]>]} + & [/\ cyclic Ubar, coprime u p.-1 & u %| (p ^ q).-1 %/ p.-1]}. +Proof. +move=> irrU; have [_ sUW1M _ /joing_subP[nHU nHW1] _] := sdprod_context defHUW1. +have [nHbU nHbW1] := (quotient_norms H0 nHU, quotient_norms H0 nHW1). +have{sUW1M} /joing_subP[nH0U nH0W1] := subset_trans sUW1M nH0M. +have [ltCU oW2b oHb] := Ptype_Fcore_factor_facts. +pose rU := abelem_repr abelHbar ntHbar nHbU. +pose inHb := rVabelem abelHbar ntHbar; pose outHb := abelem_rV abelHbar ntHbar. +have{irrU} irrU: mx_irreducible rU by apply/abelem_mx_irrP; rewrite -acts_irrQ. +pose E_U := [pred A | (A \in enveloping_algebra_mx rU)%MS]. +have cEE A: A \in E_U -> centgmx rU A. + case/envelop_mxP=> z_ ->{A}; rewrite -memmx_cent_envelop linear_sum. + rewrite summx_sub // => x Ux; rewrite linearZ scalemx_sub {z_}//=. + rewrite memmx_cent_envelop; apply/centgmxP=> y Uy. + rewrite -repr_mxM // commgC 2?repr_mxM ?(groupR, groupM) // -/rU. + apply/row_matrixP=> i; rewrite row_mul; move: (row i _) => h. + have cHbH': (U / H0)^`(1) \subset 'C(Hbar). + by rewrite -quotient_der ?quotient_cents. + apply: rVabelem_inj; rewrite rVabelemJ ?groupR //. + by apply: (canLR (mulKg _)); rewrite -(centsP cHbH') ?mem_commg ?mem_rVabelem. +have{cEE} [F [outF [inF outFK inFK] E_F]]: + {F : finFieldType & {outF : {rmorphism F -> 'M(Hbar)%Mg} + & {inF : {additive _} | cancel outF inF & {in E_U, cancel inF outF}} + & forall a, outF a \in E_U}}%R. +- pose B := row_base (enveloping_algebra_mx rU). + have freeB: row_free B by apply: row_base_free. + pose outF := [additive of vec_mx \o mulmxr B]. + pose inF := [additive of mulmxr (pinvmx B) \o mxvec]. + have E_F a: outF a \in E_U by rewrite !inE vec_mxK mulmx_sub ?eq_row_base. + have inK: {in E_U, cancel inF outF}. + by move=> A E_A; rewrite /= mulmxKpV ?mxvecK ?eq_row_base. + have outI: injective outF := inj_comp (can_inj vec_mxK) (row_free_inj freeB). + have outK: cancel outF inF by move=> a; apply: outI; rewrite inK ?E_F. + pose one := inF 1%R; pose mul a b := inF (outF a * outF b)%R. + have outM: {morph outF: a b / mul a b >-> a * b}%R. + by move=> a b; rewrite inK //; apply: envelop_mxM; exact: E_F. + have out0: outF 0%R = 0%R by apply: raddf0. + have out1: outF one = 1%R by rewrite inK //; exact: envelop_mx1. + have nzFone: one != 0%R by rewrite -(inj_eq outI) out1 out0 oner_eq0. + have mulA: associative mul by move=> *; apply: outI; rewrite !{1}outM mulrA. + have mulC: commutative mul. + move=> a b; apply: outI; rewrite !{1}outM. + by apply: cent_mxP (E_F a); rewrite memmx_cent_envelop cEE ?E_F. + have mul1F: left_id one mul by move=> a; apply: outI; rewrite outM out1 mul1r. + have mulD: left_distributive mul +%R%R. + by move=> a1 a2 b; apply: canLR outK _; rewrite !raddfD mulrDl -!{1}outM. + pose Fring_NC := RingType 'rV__ (ComRingMixin mulA mulC mul1F mulD nzFone). + pose Fring := ComRingType Fring_NC mulC. + have outRM: multiplicative (outF : Fring -> _) by []. + have mulI (nza : {a | a != 0%R :> Fring}): GRing.rreg (val nza). + case: nza => a /=; rewrite -(inj_eq outI) out0 => nzA b1 b2 /(congr1 outF). + rewrite !{1}outM => /row_free_inj eqB12; apply/outI/eqB12. + by rewrite row_free_unit (mx_Schur irrU) ?cEE ?E_F. + pose inv (a : Fring) := oapp (fun nza => invF (mulI nza) one) a (insub a). + have inv0: (inv 0 = 0)%R by rewrite /inv insubF ?eqxx. + have mulV: GRing.Field.axiom inv. + by move=> a nz_a; rewrite /inv insubT /= (f_invF (mulI (exist _ _ _))). + pose Funit := FieldUnitMixin mulV inv0. + pose FringUcl := @GRing.ComUnitRing.Class _ (GRing.ComRing.class Fring) Funit. + have Ffield := @FieldMixin (GRing.ComUnitRing.Pack FringUcl nat) _ mulV inv0. + pose F := FieldType (IdomainType _ (FieldIdomainMixin Ffield)) Ffield. + by exists [finFieldType of F], (AddRMorphism outRM); first exists inF. +pose in_uF (a : F) : {unit F} := insubd (1 : {unit F}) a. +have in_uF_E a: a != 1 -> val (in_uF a) = a. + by move=> nt_a; rewrite insubdK /= ?unitfE. +have [psi psiK]: {psi : {morphism U >-> {unit F}} + | {in U, forall x, outF (val (psi x)) = rU (inMb x)}}. +- pose psi x := in_uF (inF (rU (inMb x))). + have psiK x: x \in U -> outF (val (psi x)) = rU (inMb x). + move/(mem_quotient H0)=> Ux; have EUx := envelop_mx_id rU Ux. + rewrite in_uF_E ?inFK //; apply: contraTneq (repr_mx_unitr rU Ux). + by move/(canRL_in inFK EUx)->; rewrite rmorph0 unitr0. + suffices psiM: {in U &, {morph psi: x y / x * y}} by exists (Morphism psiM). + move=> x y Ux Uy /=; apply/val_inj/(can_inj outFK); rewrite rmorphM //. + by rewrite !{1}psiK ?groupM // morphM ?(subsetP nH0U) ?repr_mxM ?mem_quotient. +have /trivgPn/sig2W[s W2s nts]: W2bar != 1%G. + by rewrite -cardG_gt1 oW2b prime_gt1. +pose sb := outHb s; have [Hs cW1s] := setIP W2s. +have nz_sb: sb != 0%R by rewrite morph_injm_eq1 ?abelem_rV_injm. +pose phi' a : coset_of H0 := inHb (sb *m outF a)%R. +have Hphi' a: phi' a \in Hbar by apply: mem_rVabelem. +have phi'D: {in setT &, {morph phi' : a b / a * b}}. + by move=> a b _ _; rewrite /phi' !raddfD [inHb _]morphM ?mem_im_abelem_rV. +have inj_phi': injective phi'. + move=> a b /rVabelem_inj eq_sab; apply: contraNeq nz_sb. + rewrite -[sb]mulmx1 idmxE -(rmorph1 outF) -subr_eq0 => /divff <-. + by rewrite rmorphM mulmxA !raddfB /= eq_sab subrr mul0mx. +have injm_phi': 'injm (Morphism phi'D) by apply/injmP; exact: in2W. +have Dphi: 'dom (invm injm_phi') = Hbar. + apply/setP=> h; apply/morphimP/idP=> [[a _ _ ->] // | Hh]. + have /cyclic_mxP[A E_A def_h]: (outHb h <= cyclic_mx rU sb)%MS. + by rewrite -(mxsimple_cyclic irrU) ?submx1. + by exists (inF A); rewrite ?inE //= /phi' inFK // -def_h [inHb _]abelem_rV_K. +have [phi [def_phi Kphi _ im_phi]] := domP _ Dphi. +have{Kphi} inj_phi: 'injm phi by rewrite Kphi injm_invm. +have{im_phi} im_phi: phi @* Hbar = setT by rewrite im_phi -Dphi im_invm. +have phiK: {in Hbar, cancel phi phi'} by rewrite def_phi -Dphi; exact: invmK. +have{def_phi Dphi injm_phi'} phi'K: cancel phi' phi. + by move=> a; rewrite def_phi /= invmE ?inE. +have phi'1: phi' 1%R = s by rewrite /phi' rmorph1 mulmx1 [inHb _]abelem_rV_K. +have phi_s: phi s = 1%R by rewrite -phi'1 phi'K. +have phiJ: {in Hbar & U, forall h x, phi (h ^ inMb x) = phi h * val (psi x)}%R. + move=> h x Hh Ux; have Uxb := mem_quotient H0 Ux. + apply: inj_phi'; rewrite phiK ?memJ_norm ?(subsetP nHbU) // /phi' rmorphM. + by rewrite psiK // mulmxA [inHb _]rVabelemJ // -/inHb [inHb _]phiK. +have Kpsi: 'ker psi = C. + apply/setP=> x; rewrite [C]unlock 2!in_setI /= astabQ; apply: andb_id2l => Ux. + have Ubx := mem_quotient H0 Ux; rewrite 3!inE (subsetP nH0U) //= inE. + apply/eqP/centP=> [psi_x1 h Hh | cHx]; last first. + by apply/val_inj; rewrite -[val _]mul1r -phi_s -phiJ // conjgE -cHx ?mulKg. + red; rewrite (conjgC h) -[h ^ _]phiK ?memJ_norm ?(subsetP nHbU) ?phiJ //. + by rewrite psi_x1 mulr1 phiK. +have etaP (w : subg_of W1): injective (fun a => phi (phi' a ^ inMb (val w))). + case: w => w /=/(mem_quotient H0)/(subsetP nHbW1) => nHw a b eq_ab. + apply/inj_phi'/(conjg_inj (inMb w)). + by apply: (injmP inj_phi) eq_ab; rewrite memJ_norm ?mem_rVabelem. +pose eta w : {perm F} := perm (etaP (subg W1 w)). +have etaK: {in Hbar & W1, forall h w, eta w (phi h) = phi (h ^ inMb w)}. + by move=> h w Hh Ww; rewrite /= permE subgK ?phiK. +have eta1 w: w \in W1 -> eta w 1%R = 1%R. + move=> Ww; rewrite -phi_s etaK //. + by rewrite conjgE (centP cW1s) ?mulKg ?mem_quotient. +have etaM: {in W1 &, {morph eta: w1 w2 / w1 * w2}}. + move=> w1 w2 Ww1 Ww2; apply/permP=> a; rewrite -[a]phi'K permM. + rewrite !etaK ?memJ_norm ?groupM ?(subsetP nHbW1) ?mem_quotient //. + by rewrite -conjgM -morphM ?(subsetP nH0W1). +have etaMpsi a: {in U & W1, forall x w, + eta w (a * val (psi x)) = eta w a * val (psi (x ^ w)%g)}%R. +- move=> x w Ux Ww; rewrite -[a]phi'K (etaK _ w (Hphi' a) Ww). + rewrite -!phiJ // ?memJ_norm ?(subsetP nHbW1, subsetP nUW1) ?mem_quotient //. + rewrite etaK ?memJ_norm ?(subsetP nHbU) ?mem_quotient // -!conjgM. + by rewrite conjgC -morphJ ?(subsetP nH0U x Ux, subsetP nH0W1 w Ww). +have psiJ: {in U & W1, forall x w, val (psi (x ^ w)) = eta w (val (psi x))}. + by move=> x w Ux Ww /=; rewrite -[val _]mul1r -(eta1 w Ww) -etaMpsi ?mul1r. +have etaRM w: w \in W1 -> rmorphism (eta w). + move=> Ww; have nUw := subsetP nHbW1 _ (mem_quotient _ Ww). + have etaD: additive (eta w). + move=> a b; rewrite -[a]phi'K -[b]phi'K -!zmodMgE -!zmodVgE. + rewrite -morphV // -morphM ?{1}etaK ?groupM ?groupV // conjMg conjVg. + by rewrite morphM 1?morphV ?groupV // memJ_norm. + do 2![split=> //] => [a b|]; last exact: eta1. + rewrite -[a]outFK; have /envelop_mxP[d ->] := E_F a. + rewrite raddf_sum mulr_suml ) mulr_suml. + apply: eq_bigr => _ /morphimP[x Nx Ux ->]; move: {d}(d _) => dx. + rewrite -[dx]natr_Zp scaler_nat !(mulrnAl, raddfMn); congr (_ *+ dx)%R. + by rewrite -psiK //= outFK mulrC etaMpsi // mulrC psiJ. +have oF: #|F| = (p ^ q)%N by rewrite -cardsT -im_phi card_injm. +pose nF := <[1%R : F]>; have o_nF: #|nF| = p. + by rewrite -orderE -phi_s (order_injm inj_phi) // (abelem_order_p abelHbar). +have cyc_uF := @field_unit_group_cyclic F. +exists F. + exists phi; last first. + split=> //; first exact/isomP; apply/esym/eqP; rewrite eqEcard o_nF -phi_s. + by rewrite (@cycle_subG F) mem_morphim //= card_injm ?subsetIl ?oW2b. + exists psi => //; last first. + by split=> // h x Hh Ux; rewrite qactJ (subsetP nH0U) ?phiJ. + have inj_eta: 'injm (Morphism etaM). + have /properP[_ [h Hh notW2h]]: W2bar \proper Hbar. + by rewrite properEcard subsetIl oW2b oHb (ltn_exp2l 1) prime_gt1. + apply/subsetP=> w /morphpreP[Ww /set1P/permP/(_ (phi h))]. + rewrite etaK // permE => /(injmP inj_phi) => chw. + rewrite -(@prime_TIg _ W1 <[w]>) //; first by rewrite inE Ww cycle_id. + rewrite proper_subn // properEneq cycle_subG Ww andbT. + apply: contraNneq notW2h => defW1; rewrite inE Hh /= -defW1. + rewrite quotient_cycle ?(subsetP nH0W1) // cent_cycle cent1C inE. + by rewrite conjg_set1 chw ?memJ_norm // (subsetP nHbW1) ?mem_quotient. + exists (Morphism etaM) => [alpha |]; last first. + by split=> // h w Hh Ww /=; rewrite qactJ (subsetP nH0W1) -?etaK. + pose autF (f : {perm F}) := rmorphism f. (* Bits of Galois theory... *) + have [r prim_r]: {r : F | forall f g, autF f -> autF g -> f r = g r -> f = g}. + have /cyclicP/sig_eqW[r def_uF] := cyc_uF [set: {unit F}]%G. + exists (val r) => f g fRM gRM eq_fgr; apply/permP=> a. + rewrite (_ : f =1 RMorphism fRM) // (_ : g =1 RMorphism gRM) //. + have [-> | /in_uF_E <-] := eqVneq a 0%R; first by rewrite !rmorph0. + have /cycleP[m ->]: in_uF a \in <[r]> by rewrite -def_uF inE. + by rewrite val_unitX !rmorphX /= eq_fgr. + have /sigW[P /and3P[Pr0 nP lePq]]: + exists P: {poly F}, [&& root P r, all (mem nF) P & #|root P| <= q]. + - pose Mr := (\matrix_(i < q.+1) (sb *m outF (r ^+ i)))%R. + have /rowV0Pn[v /sub_kermxP vMr0 nz_v]: kermx Mr != 0%R. + rewrite kermx_eq0 neq_ltn ltnS (leq_trans (rank_leq_col Mr)) //. + by rewrite (dim_abelemE abelHbar) // oHb pfactorK. + pose P : {poly F} := (\poly_(i < q.+1) (v 0 (inord i))%:R)%R. + have szP: size P <= q.+1 by apply: size_poly. + exists P; apply/and3P; split. + + apply/eqP/inj_phi'; congr (inHb _); rewrite rmorph0 mulmx0 -vMr0. + rewrite horner_poly !raddf_sum mulmx_sum_row; apply: eq_bigr => i _. + rewrite rowK inord_val //= mulr_natl rmorphMn -scaler_nat scalemxAr. + by rewrite natr_Zp. + + apply/(all_nthP 0%R)=> i /leq_trans/(_ szP) le_i_q. + by rewrite coef_poly /= le_i_q mem_cycle. + rewrite cardE -ltnS (leq_trans _ szP) //. + rewrite max_poly_roots ?enum_uniq //; last first. + by apply/allP=> r'; rewrite mem_enum. + apply: contraNneq nz_v => /polyP P0; apply/eqP/rowP=> i; apply/eqP. + have /eqP := P0 i; rewrite mxE coef0 coef_poly ltn_ord inord_val. + have charF: p \in [char F]%R by rewrite !inE p_pr -order_dvdn -o_nF /=. + by rewrite -(dvdn_charf charF) (dvdn_charf (char_Fp p_pr)) natr_Zp. + have{Pr0 nP} fPr0 f: autF f -> root P (f r). + move=> fRM; suff <-: map_poly (RMorphism fRM) P = P by apply: rmorph_root. + apply/polyP=> i; rewrite coef_map. + have [/(nth_default _)-> | lt_i_P] := leqP (size P) i; first exact: rmorph0. + by have /cycleP[n ->] := all_nthP 0%R nP i lt_i_P; exact: rmorph_nat. + apply: (iffP morphimP) => [[w _ Ww ->] | alphaRM]; first exact: etaRM. + suffices /setP/(_ (alpha r)): [set (eta w) r | w in W1] = [set t | root P t]. + rewrite inE fPr0 // => /imsetP[w Ww def_wr]; exists w => //. + by apply: prim_r => //; exact: etaRM. + apply/eqP; rewrite eqEcard; apply/andP; split. + by apply/subsetP=> _ /imsetP[w Ww ->]; rewrite inE fPr0 //; exact: etaRM. + rewrite (@cardsE F) card_in_imset // => w1 w2 Ww1 Ww2 /= /prim_r eq_w12. + by apply: (injmP inj_eta) => //; apply: eq_w12; exact: etaRM. +have isoUb: isog Ubar (psi @* U) by rewrite /Ubar -Kpsi first_isog. +pose unF := [set in_uF a | a in nF^#]. +have unF_E: {in nF^#, cancel in_uF val} by move=> a /setD1P[/in_uF_E]. +have unFg: group_set unF. + apply/group_setP; split=> [|_ _ /imsetP[a nFa ->] /imsetP[b nFb ->]]. + have nF1: 1%R \in nF^# by rewrite !inE cycle_id oner_eq0. + by apply/imsetP; exists 1%R => //; apply: val_inj; rewrite unF_E. + have nFab: (a * b)%R \in nF^#. + rewrite !inE mulf_eq0 negb_or. + have [[-> /cycleP[m ->]] [-> /cycleP[n ->]]] := (setD1P nFa, setD1P nFb). + by rewrite -natrM mem_cycle. + by apply/imsetP; exists (a * b)%R => //; apply: val_inj; rewrite /= !unF_E. +have <-: #|Group unFg| = p.-1. + by rewrite -o_nF (cardsD1 1 nF) group1 (card_in_imset (can_in_inj unF_E)). +have <-: #|[set: {unit F}]| = (p ^ q).-1. + rewrite -oF -(cardC1 1) cardsT card_sub; apply: eq_card => a /=. + by rewrite !inE unitfE. +rewrite /u (isog_cyclic isoUb) (card_isog isoUb) cyc_uF. +suffices co_u_p1: coprime #|psi @* U| #|Group unFg|. + by rewrite -(Gauss_dvdr _ co_u_p1) mulnC divnK ?cardSg ?subsetT. +rewrite -(cyclic_dprod (dprodEY _ _)) ?cyc_uF //. + by rewrite (sub_abelian_cent2 (cyclic_abelian (cyc_uF [set:_]%G))) ?subsetT. +apply/trivgP/subsetP=> _ /setIP[/morphimP[x Nx Ux ->] /imsetP[a nFa /eqP]]. +have nCx: x \in 'N(C) by rewrite -Kpsi (subsetP (ker_norm _)). +rewrite -val_eqE (unF_E a) //; case/setD1P: nFa => _ /cycleP[n {a}->]. +rewrite zmodXgE => /eqP def_psi_x; rewrite mker ?set11 // Kpsi coset_idr //. +apply/set1P; rewrite -set1gE -(Frobenius_trivg_cent frobUW1c) /= -/C. +rewrite inE mem_quotient //= -sub1set -quotient_set1 ?quotient_cents2r //. +rewrite gen_subG /= -/C -Kpsi; apply/subsetP=> _ /imset2P[_ w /set1P-> Ww ->]. +have Uxw: x ^ w \in U by rewrite memJ_norm ?(subsetP nUW1). +apply/kerP; rewrite (morphM, groupM) ?morphV ?groupV //. +apply/(canLR (mulKg _))/val_inj; rewrite psiJ // mulg1 def_psi_x. +exact: (rmorph_nat (RMorphism (etaRM w Ww))). +Qed. + +Local Open Scope ring_scope. + +Let redM := [predC irr M]. +Let mu_ := filter redM (S_ H0). + +(* This subproof is shared between (9.8)(b) and (9.9)(b). *) +Let nb_redM_H0 : size mu_ = p.-1 /\ {subset mu_ <= S_ H0C}. +Proof. +have pddM := FT_prDade_hypF maxM MtypeP; pose ptiWM := prDade_prTI pddM. +have [nsHUM sW1M /mulG_sub[sHUM _] nHUW1 tiHUW1] := sdprod_context defM. +have [nsHHU sUHU /mulG_sub[sHHU _] nHU tiHU] := sdprod_context defHU. +have nb_redM K: + K <| M -> K \subset HU -> K :&: H = H0 -> count redM (S_ K) = p.-1. +- move=> nsKM sKHU tiKHbar; have [sKM nKM] := andP nsKM; pose b L := (L / K)%G. + have [nsKHU [_ [_ sW2HU cycW2 _] _]] := (normalS sKHU sHUM nsKM, ptiWM). + have nKW2 := subset_trans sW2HU (normal_norm nsKHU). + have oW2b: #|b W2| = p. + have [_ <- _] := Ptype_Fcore_factor_facts; rewrite defW2bar. + rewrite !card_quotient ?(subset_trans (subset_trans sW2HU sHUM)) //. + by rewrite -indexgI -{2}(setIidPl sW2H) setIAC -setIA tiKHbar indexgI. + have{cycW2} cycW2b: cyclic (b W2) by apply: quotient_cyclic. + have ntW2b: (W2 / K != 1)%g by rewrite -cardG_gt1 oW2b prime_gt1. + have{ntW2b} [defWb ptiWMb]:= primeTIhyp_quotient ptiWM ntW2b sKHU nsKM. + pose muK j := (primeTIred ptiWMb j %% K)%CF. + apply/eqP; have <-: size (image muK (predC1 0)) = p.-1. + by rewrite size_map -cardE cardC1 card_Iirr_cyclic ?oW2b. + rewrite -size_filter -uniq_size_uniq ?filter_uniq ?seqInd_uniq // => [|phi]. + by apply/dinjectiveP=> j1 j2 _ _ /(can_inj (cfModK nsKM))/prTIred_inj. + rewrite mem_filter; apply/imageP/andP=> [[j nz_j ->] | [red_phi]]; last first. + case/seqIndP=> s /setDP[kerK ker'H] Dphi; rewrite !inE in kerK ker'H. + pose s1 := quo_Iirr K s; have Ds: s = mod_Iirr s1 by rewrite quo_IirrK. + rewrite {phi}Dphi Ds mod_IirrE ?cfIndMod // in kerK ker'H red_phi *. + have [[j Ds1] | [/idPn[]]] := prTIres_irr_cases ptiWMb s1. + rewrite Ds1 cfInd_prTIres -/(muK j) in ker'H *; exists j => //. + by apply: contraNneq ker'H => ->; rewrite prTIres0 rmorph1 cfker_cfun1. + by apply: contra red_phi => /cfMod_irr/= ->. + have red_j: redM (muK j). + apply: contra (prTIred_not_irr ptiWMb j) => /(cfQuo_irr nsKM). + by rewrite cfker_mod ?cfModK // => ->. + have [s DmuKj]: exists s, muK j = 'Ind[M, HU] 'chi_s. + exists (mod_Iirr (primeTI_Ires ptiWMb j)). + by rewrite mod_IirrE // cfIndMod // cfInd_prTIres. + split=> //; apply/seqIndP; exists s; rewrite // !inE andbC. + rewrite -(@sub_cfker_Ind_irr _ M) ?gFnorm // -DmuKj cfker_mod //=. + have [[j1 Ds] | [/idPn]] := prTIres_irr_cases ptiWM s; last by rewrite -DmuKj. + rewrite Ds cfker_prTIres //; apply: contraNneq nz_j => j1_0. + apply/eqP/(prTIred_inj ptiWMb)/(can_inj (cfModK nsKM)); rewrite -{1}/(muK j). + by rewrite DmuKj Ds j1_0 -cfInd_prTIres !prTIres0 -cfIndMod ?rmorph1. +have [sH0HU sH0M] := (subset_trans sH0H sHHU, subset_trans sH0H (gFsub _ _)). +have sz_mu: size mu_ = p.-1. + by rewrite size_filter nb_redM ?(setIidPl sH0H) // /normal sH0M. +have s_muC_mu: {subset filter redM (S_ H0C) <= mu_}. + move=> phi; rewrite /= !mem_filter => /andP[->]; apply: seqIndS. + by rewrite setSD // Iirr_kerS ?joing_subl. +have UmuC: uniq (filter redM (S_ H0C)) by rewrite filter_uniq ?seqInd_uniq. +have [|Dmu _] := leq_size_perm UmuC s_muC_mu; last first. + by split=> // phi; rewrite -Dmu mem_filter => /andP[]. +have [nsH0C_M _ _ _] := nsH0xx_M. +have sCHU := subset_trans sCU sUHU; have sCM := subset_trans sCHU sHUM. +have sHOC_HU: H0C \subset HU by apply/joing_subP. +rewrite sz_mu size_filter nb_redM //= norm_joinEr ?(subset_trans sCM) //. +by rewrite -group_modl //= setIC [C]unlock setIA tiHU setI1g mulg1. +Qed. + +Let isIndHC (zeta : 'CF(M)) := + [/\ zeta 1%g = (q * u)%:R, zeta \in S_ H0C + & exists2 xi : 'CF(HC), xi \is a linear_char & zeta = 'Ind xi]. + +(* This is Peterfalvi (9.8). *) +Lemma typeP_nonGalois_characters (not_Galois : ~~ typeP_Galois) : + let a := #|U : 'C_U(sval (typeP_Galois_Pn not_Galois) | 'Q)| in + [/\ (*a*) {in X_ H0, forall s, (a %| 'chi_s 1%g)%C}, + (*b*) size mu_ = p.-1 /\ {in mu_, forall mu_j, isIndHC mu_j}, + (*c*) exists t, isIndHC 'chi_t + & (*d*) let irr_qa := [pred zeta in irr M | zeta 1%g == (q * a)%:R] in + let lb_n := (p.-1 * #|U|)%N in let lb_d := (a ^ 2 * #|U'|)%N in + (lb_d %| lb_n /\ lb_n %/ lb_d <= count irr_qa (S_ H0U'))%N]. +Proof. +case: (typeP_Galois_Pn _) => H1 [oH1 nH1U nH1Uq defHbar aP]; rewrite [sval _]/=. +move => a; case: aP; rewrite -/a => a_gt1 a_dv_p1 cycUb1 isoUb. +set part_a := ({in _, _}); pose HCbar := (HC / H0)%G. +have [_ /mulG_sub[sHUM sW1M] nHUW1 tiHUW1] := sdprodP defM. +have [nsHHU _ /mulG_sub[sHHU sUHU] nHU tiHU] := sdprod_context defHU. +have [nH0H nHHU] := (normal_norm nsH0H, normal_norm nsHHU). +have sHHC: H \subset HC by rewrite joing_subl. +have [nH0HU sCHU] := (subset_trans sHUM nH0M, subset_trans sCU sUHU). +have nsH0_HU: H0 <| HU by rewrite /normal (subset_trans sH0H). +have nH0C := subset_trans sCHU nH0HU. +have [nsH0C_M nsHC_M nsH0U'_M _] := nsH0xx_M; have [sHC_M _] := andP nsHC_M. +have nsH0HC: H0 <| HC := normalS (subset_trans sH0H sHHC) sHC_M nsH0M. +have defHCbar: Hbar \x (C / H0) = HCbar. + rewrite /= quotientY // [C]unlock /= astabQ quotient_setIpre. + by rewrite dprodEY ?subsetIr // setIA -quotientGI // tiHU quotient1 setI1g. +have sHC_HU: HC \subset HU by rewrite join_subG sHHU. +have nsHC_HU: HC <| HU := normalS sHC_HU sHUM nsHC_M. +have defHb1 := defHbar; rewrite (big_setD1 1%g) ?group1 ?conjsg1 //= in defHb1. +have [[_ H1c _ defH1c] _ _ _] := dprodP defHb1; rewrite defH1c in defHb1. +have [nsH1H _] := dprod_normal2 defHb1; have [sH1H nH1H] := andP nsH1H. +have nHW1: W1 \subset 'N(H) := subset_trans sW1M (gFnorm _ _). +have nHbW1: W1bar \subset 'N(Hbar) by rewrite quotient_norms. +have sH1wH w: w \in W1bar -> H1 :^ w \subset Hbar. + by move/(normsP nHbW1) <-; rewrite conjSg. +have nsH1wHUb w: w \in W1bar -> H1 :^ w <| HU / H0. + move=> W1w; rewrite -(normsP (quotient_norms _ nHUW1) w W1w) normalJ. + rewrite /normal (subset_trans sH1H) ?quotientS //. + by rewrite -defHU sdprodE // quotientMl // mulG_subG nH1H. +have nH1wHUb := normal_norm (nsH1wHUb _ _). +have Part_a: part_a. + move=> s; rewrite !inE => /andP[kers'H kersH0]. + have [t sHt] := constt_cfRes_irr H s; pose theta := ('chi_t / H0)%CF. + have{kers'H} t_neq0: t != 0. + by rewrite -subGcfker (sub_cfker_constt_Res_irr sHt). + have{kersH0} kertH0: H0 \subset cfker 'chi_t. + by rewrite (sub_cfker_constt_Res_irr sHt). + have Ltheta: theta \is a linear_char. + by rewrite /theta -quo_IirrE // (char_abelianP _ _). + have Dtheta : _ = theta := cfBigdprod_Res_lin defHbar Ltheta. + set T := 'I_HU['chi_t]; have sHT: H \subset T by rewrite sub_Inertia. + have sTHU: T \subset HU by rewrite Inertia_sub. + suffices{s sHt} a_dv_iTHU: a %| #|HU : T|. + have [_ defInd_t _ imInd_t _] := cfInd_sum_Inertia t nsHHU. + have /imsetP[r tTr ->]: s \in Ind_Iirr HU @: irr_constt ('Ind[T] 'chi_t). + by rewrite imInd_t constt_Ind_Res. + by rewrite defInd_t ?cfInd1 // dvdC_mulr ?dvdC_nat // Cint_Cnat ?Cnat_irr1. + have /exists_inP[w W1w nt_t_w]: [exists w in W1bar, 'Res[H1 :^ w] theta != 1]. + rewrite -negb_forall_in; apply: contra t_neq0 => /forall_inP=> tH1w1. + rewrite -irr_eq1 -(cfQuoK nsH0H kertH0) -/theta -Dtheta. + rewrite [cfBigdprod _ _]big1 ?rmorph1 // => w /tH1w1/eqP->. + by rewrite /cfBigdprodi rmorph1. + have defT: H ><| (U :&: T) = T. + by rewrite (sdprod_modl defHU) // (setIidPr sTHU). + have /irrP[k Dk]: 'Res theta \in irr (H1 :^ w). + by rewrite lin_char_irr ?cfRes_lin_char. + rewrite -divgS // -(sdprod_card defHU) -(sdprod_card defT) divnMl // divgI. + rewrite -indexgI; have ->: a = #|U : 'C_U(H1 :^ w | 'Q)|. + have [w1 nH0w1 W1w1 ->] := morphimP W1w; rewrite astabQ centJ morphpreJ //. + by rewrite -astabQ indexgI -(normsP nUW1 _ W1w1) indexJg -indexgI. + rewrite indexgS ?setIS // sub_astabQ ?(subset_trans sTHU) //= -inertia_quo //. + apply: subset_trans (sub_inertia_Res _ (nH1wHUb w W1w)) _. + by rewrite Dk (inertia_irr_prime _ p_pr) ?subsetIr ?cardJg // -irr_eq1 -Dk. +pose isoJ := conj_isom H1; pose cfJ w i := 'chi_(isom_Iirr (isoJ w) i). +pose thetaH (f : {ffun _}) := cfBigdprod defHbar (fun w => cfJ w (f w)). +pose theta f := cfDprodl defHCbar (thetaH f). +have abH1: abelian H1 by rewrite cyclic_abelian ?prime_cyclic ?oH1. +have linH1 i: 'chi[H1]_i \is a linear_char by apply/char_abelianP. +have lin_thetaH f: thetaH f \is a linear_char. + by apply: cfBigdprod_lin_char => w _; rewrite /cfJ isom_IirrE cfIsom_lin_char. +have nz_thetaH f: thetaH f 1%g != 0 by rewrite lin_char_neq0. +have Dtheta f: {in W1bar & H1, forall w xb, theta f (xb ^ w) = 'chi_(f w) xb}. + move=> w xb W1w H1xb /=; have sHHCb := quotientS H0 sHHC. + transitivity ('Res[H1 :^ w] ('Res[Hbar] (theta f)) (xb ^ w)); last first. + by rewrite cfDprodlK cfBigdprodKabelian // isom_IirrE cfIsomE. + by rewrite cfResRes ?sH1wH // cfResE ?memJ_conjg ?(subset_trans (sH1wH w _)). +have lin_theta f: theta f \is a linear_char by apply: cfDprodl_lin_char. +pose Ftheta := pffun_on (0 : Iirr H1) W1bar (predC1 0). +have inj_theta: {in Ftheta &, injective theta}. + move=> f1 f2 /pffun_onP[/supportP W1f1 _] /pffun_onP[/supportP W1f2 _] eq_f12. + apply/ffunP=> w. + have [W1w | W1'w] := boolP (w \in W1bar); last by rewrite W1f1 ?W1f2. + by apply/irr_inj/cfun_inP=> x H1x; rewrite -!Dtheta ?eq_f12. +have irr_thetaH0 f: (theta f %% H0)%CF \in irr HC. + by rewrite cfMod_irr ?lin_char_irr. +have def_Itheta f: f \in Ftheta -> 'I_HU[theta f %% H0]%CF = HC. + case/pffun_onP=> _ nz_fW1; apply/eqP; rewrite eqEsubset sub_Inertia //. + rewrite inertia_mod_pre //= -{1}(sdprodW defHU) -group_modl; last first. + rewrite (subset_trans sHHC) // -sub_quotient_pre ?normal_norm //. + by rewrite sub_Inertia ?quotientS. + rewrite -gen_subG genM_join genS ?setUS //= {2}[C]unlock setIS //= astabQ. + rewrite morphpreS // centsC -{1}(bigdprodWY defHbar) gen_subG. + apply/bigcupsP=> w W1w; rewrite centsC. + apply: subset_trans (sub_inertia_Res _ (quotient_norms _ nHHU)) _. + rewrite cfDprodlK inertia_bigdprod_irr // subIset // orbC (bigcap_min w) //. + rewrite (inertia_irr_prime _ p_pr) ?cardJg ?subsetIr // isom_Iirr_eq0. + by apply: nz_fW1; apply: image_f. +have irrXtheta f: f \in Ftheta -> 'Ind (theta f %% H0)%CF \in irr HU. + move/def_Itheta; rewrite -(cfIirrE (irr_thetaH0 f)) => I_f_HC. + by rewrite inertia_Ind_irr ?I_f_HC //. +pose Mtheta := [set cfIirr (theta f %% H0)%CF | f in Ftheta]. +pose Xtheta := [set cfIirr ('Ind[HU] 'chi_t) | t in Mtheta]. +have oXtheta: (u * #|Xtheta| = p.-1 ^ q)%N. + transitivity #|Ftheta|; last first. + rewrite card_pffun_on cardC1 card_Iirr_abelian // oH1. + rewrite -(card_isog (quotient_isog _ _)) ?oW1 ?(subset_trans sW1M) //. + by apply/trivgP; rewrite -tiHUW1 setSI ?(subset_trans sH0H). + rewrite Du -card_imset_Ind_irr ?card_in_imset //. + - move=> f1 f2 Df1 Df2 /(congr1 (tnth (irr HC))); rewrite !{1}cfIirrE //. + by move/(can_inj (cfModK nsH0HC)); apply: inj_theta. + - by move=> _ /imsetP[f Df ->]; rewrite cfIirrE ?irrXtheta. + move=> _ y /imsetP[f /familyP Ff ->] HUy; apply/imsetP. + pose yb := inMb y; have HUyb: yb \in (HU / H0)%g by apply: mem_quotient. + have nHb_y: inMb y \in 'N(Hbar) by rewrite (subsetP (quotient_norms _ nHHU)). + have nH1b_y := subsetP (nH1wHUb _ _) yb HUyb. + exists [ffun w => conjg_Iirr (f w) (inMb y ^ w^-1)]. + apply/familyP=> w; rewrite ffunE. + by case: ifP (Ff w) => _; rewrite !inE conjg_Iirr_eq0. + apply: irr_inj; rewrite !(cfIirrE, conjg_IirrE) // (cfConjgMod _ nsHC_HU) //. + rewrite cfConjgDprodl //; first congr (cfDprodl _ _ %% H0)%CF; last first. + rewrite /= -quotientYidl // (subsetP _ _ HUyb) ?quotient_norms //. + by rewrite (subset_trans sHUM) ?normal_norm. + rewrite cfConjgBigdprod //; apply: eq_bigr => w W1w; congr (cfBigdprodi _ _). + rewrite ffunE /cfJ !isom_IirrE conjg_IirrE. + apply/cfun_inP=> _ /imsetP[x Hx ->]; rewrite cfIsomE // cfConjgE ?nH1b_y //. + rewrite -conjgM conjgCV conjVg conjgM cfIsomE //; last first. + by rewrite -mem_conjg (normP _) // -mem_conjg -normJ ?nH1b_y. + by rewrite cfConjgE // -mem_conjg -normJ ?nH1b_y. +have sXthXH0C: Xtheta \subset X_ H0C. + apply/subsetP=> _ /imsetP[t Mt ->]; have{Mt} [f Ff Dt] := imsetP Mt. + rewrite !inE cfIirrE; last by rewrite Dt cfIirrE ?irrXtheta. + rewrite !sub_cfker_Ind_irr ?(subset_trans sHUM) ?normal_norm ?gFnormal //. + rewrite {t}Dt cfIirrE // join_subG andbCA {1}cfker_mod //. + rewrite !{1}sub_cfker_mod //= andbC {1}cfker_sdprod /=. + apply: contraL (familyP Ff 1%g) => kerHb; rewrite group1 negbK. + have:= sub_cfker_Res (subxx _) kerHb; rewrite cfDprodlK. + move/(subset_trans (sH1wH _ (group1 _)))/(sub_cfker_Res (subxx _)). + rewrite cfBigdprodKabelian // isom_IirrE cfker_isom morphim_conj /=. + by rewrite !conjsg1 subsetIidl subGcfker. +pose mu_f (i : Iirr H1) := [ffun w => if w \in W1bar then i else 0]. +have Fmu_f (i : Iirr H1): i != 0 -> mu_f i \in Ftheta. + by move=> nz_i; apply/familyP=> w; rewrite ffunE; case: ifP; rewrite !inE. +pose mk_mu i := 'Ind[HU] (theta (mu_f i) %% H0)%CF. +have sW1_Imu i: W1 \subset 'I[theta (mu_f i) %% H0]%CF. + apply/subsetP=> w W1w; have Mw := subsetP sW1M w W1w. + have nHC_W1 := subset_trans sW1M (normal_norm nsHC_M). + rewrite inE (subsetP nHC_W1) ?(cfConjgMod _ nsHC_M) //; apply/eqP. + have W1wb: inMb w \in W1bar by rewrite mem_quotient. + rewrite cfConjgDprodl ?(subsetP _ _ W1wb) ?quotient_norms //; last first. + by rewrite (subset_trans (joing_subr U W1)) ?normal_norm. + congr (cfDprodl _ _ %% H0)%CF. + apply/cfun_inP=> _ /(mem_bigdprod defHbar)[x [H1x -> _]]. + have Hx w1: w1 \in W1bar -> x w1 \in Hbar. + by move=> W1w1; rewrite (subsetP (sH1wH w1 _)) ?H1x. + rewrite !lin_char_prod ?cfConjg_lin_char //; apply/eq_bigr=> w1 W1w1. + rewrite cfConjgE ?(subsetP nHbW1) //. + have W1w1w: (w1 * (inMb w)^-1)%g \in W1bar by rewrite !in_group. + rewrite -(cfResE _ (sH1wH _ W1w1w)) -?mem_conjg -?conjsgM ?mulgKV ?H1x //. + rewrite -(cfResE _ (sH1wH _ W1w1)) ?H1x ?cfBigdprodKabelian //. + rewrite !ffunE W1w1 W1w1w -[x w1](conjgKV w1) -conjgM !isom_IirrE. + by rewrite !cfIsomE -?mem_conjg ?H1x. +have inj_mu: {in predC1 0 &, injective (fun i => cfIirr (mk_mu i))}. + move=> i1 i2 nz_i1 nz_i2 /(congr1 (tnth (irr HU))). + rewrite !{1}cfIirrE ?irrXtheta ?Fmu_f // /mk_mu. + do 2![move/esym; rewrite -{1}(cfIirrE (irr_thetaH0 _))]. + move/(cfclass_Ind_irrP _ _ nsHC_HU); rewrite !{1}cfIirrE //. + case/cfclassP=> _ /(mem_sdprod defHU)[x [y [Hx Uy -> _]]]. + rewrite (cfConjgM _ nsHC_HU) ?(subsetP sHHU x) ?(subsetP sUHU) //. + rewrite {x Hx}(cfConjg_id _ (subsetP sHHC x Hx)) => Dth1. + suffices /setIP[_ /inertiaJ]: y \in 'I_HU[theta (mu_f i2) %% H0]%CF. + rewrite -Dth1 => /(can_inj (cfModK nsH0HC))/inj_theta/ffunP/(_ 1%g). + by rewrite !ffunE group1 => -> //; apply: Fmu_f. + rewrite def_Itheta ?Fmu_f //= (subsetP (joing_subr _ _)) //. + have nCy: y \in 'N(C). + by rewrite (subsetP (normal_norm nsCUW1)) ?mem_gen ?inE ?Uy. + have [_ _ /trivgPn[wb W1wb ntwb] _ _] := Frobenius_context frobUW1c. + have /morphimP[w nCw W1w Dw] := W1wb; have Mw := subsetP sW1M w W1w. + rewrite coset_idr //; apply/set1P; rewrite -set1gE; apply: wlog_neg => nty. + rewrite -((Frobenius_reg_ker frobUW1c) wb); last by rewrite !inE ntwb. + rewrite inE mem_quotient //=; apply/cent1P/commgP. + rewrite Dw -morphR //= coset_id //. + suffices: [~ y, w] \in U :&: HC. + rewrite /= norm_joinEr ?(subset_trans sCU) // -group_modr ?subsetIl //=. + by rewrite setIC tiHU mul1g. + have Uyw: [~ y, w] \in U; last rewrite inE Uyw. + by rewrite {1}commgEl groupMl ?groupV // memJ_norm ?(subsetP nUW1) // Uy. + rewrite -(def_Itheta _ (Fmu_f _ nz_i1)) 2!inE /= andbA -in_setI. + rewrite (setIidPl (normal_norm nsHC_HU)) (subsetP sUHU) //=. + rewrite Dth1 -(cfConjgM _ nsHC_HU) ?(subsetP sUHU) //. + have My: y \in M := subsetP (subset_trans sUHU sHUM) y Uy. + rewrite mulKVg (cfConjgM _ nsHC_M) ?in_group //. + have /inertiaJ-> := subsetP (sW1_Imu i2) _ (groupVr W1w). + rewrite (cfConjgM _ nsHC_M) // -Dth1. + by have /inertiaJ-> := subsetP (sW1_Imu i1) w W1w. +pose Xmu := [set cfIirr (mk_mu i) | i in predC1 0]. +have def_IXmu: {in Xmu, forall s, 'I_M['chi_s] = M}. + move=> _ /imsetP[i nz_i ->]; apply/setIidPl. + rewrite -subsetIidl -{1}(sdprodW defM) mulG_subG sub_Inertia //. + rewrite !cfIirrE ?irrXtheta ?Fmu_f //=. + apply: subset_trans (sub_inertia_Ind _ (der_norm 1 M)). + by rewrite subsetI sW1M /=. +pose Smu := [seq 'Ind[M] 'chi_s | s in Xmu]. +have sSmu_mu: {subset Smu <= mu_}. + move=> _ /imageP[s Xmu_s ->]; rewrite mem_filter /=. + rewrite irrEchar cfnorm_Ind_irr ?gFnormal // def_IXmu //. + rewrite -(index_sdprod defM) (eqC_nat _ 1) gtn_eqF ?prime_gt1 // andbF. + rewrite mem_seqInd ?gFnormal /normal ?(subset_trans sH0H) ?gFsub //=. + suffices /(subsetP sXthXH0C): s \in Xtheta. + by apply: subsetP; rewrite setSD // Iirr_kerS ?joing_subl. + have /imsetP[i nz_i ->] := Xmu_s; rewrite /Xtheta -imset_comp. + by apply/imsetP; exists (mu_f i); rewrite /= ?cfIirrE ?Fmu_f. +have ResIndXmu: {in Xmu, forall s, 'Res ('Ind[M] 'chi_s) = q%:R *: 'chi_s}. + move=> s /def_IXmu Imu_s; rewrite cfResInd_sum_cfclass ?gFnormal ?Imu_s //. + by rewrite -(index_sdprod defM) -Imu_s cfclass_inertia big_seq1. +have uSmu: uniq Smu. + apply/dinjectiveP=> s1 s2 Xs1 Xs2 /(congr1 'Res[HU]); rewrite !ResIndXmu //. + by move/(can_inj (scalerK (neq0CG W1)))/irr_inj. +have sz_Smu: size Smu = p.-1. + by rewrite size_map -cardE card_in_imset // cardC1 card_Iirr_abelian ?oH1. +have [sz_mu s_mu_H0C] := nb_redM_H0. +have Dmu: Smu =i mu_. + by have [|//] := leq_size_perm uSmu sSmu_mu; rewrite sz_mu sz_Smu. +split=> {Part_a part_a}//. +- split=> // phi mu_phi; have S_HOC_phi := s_mu_H0C _ mu_phi. + move: mu_phi; rewrite -Dmu => /imageP[_ /imsetP[i nz_i ->]]. + rewrite cfIirrE ?irrXtheta ?Fmu_f // => Dphi. + split=> //; rewrite Dphi ?cfInd1 ?cfIndInd //. + rewrite -(index_sdprod defM) -/q -Du mulrA -natrM. + by rewrite lin_char1 1?cfMod_lin_char ?mulr1. + by exists (theta (mu_f i) %% H0)%CF; rewrite 1?cfMod_lin_char. +- have /eqVproper: Xmu \subset Xtheta. + apply/subsetP=> _ /imsetP[i nz_i ->]; rewrite -[Xtheta]imset_comp /=. + by apply/imsetP; exists (mu_f i); rewrite /= ?cfIirrE ?Fmu_f. + case=> [defXmu | /andP[_ /subsetPn[s theta_s Xmu'_s]]]; last first. + have [_ /imsetP[f Dth_f ->] Ds] := imsetP theta_s; rewrite cfIirrE // in Ds. + have /irrP[t Dt]: 'Ind 'chi_s \in irr M; last 1 [exists t; rewrite -{t}Dt]. + apply: contraR Xmu'_s => red_Ind_s. + have: 'Ind 'chi_s \in mu_. + rewrite mem_filter /= red_Ind_s mem_seqInd ?gFnormal //=. + apply: subsetP theta_s; rewrite (subset_trans sXthXH0C) ?setSD //. + by rewrite Iirr_kerS ?joing_subl. + rewrite -Dmu => /imageP[s1 Xmu_s1] /(congr1 (cfdot ('Ind 'chi_s1)))/eqP. + rewrite cfnorm_Ind_irr ?gFnormal // eq_sym -cfdot_Res_l. + rewrite ResIndXmu // cfdotZl cfdot_irr -natrM mulnC. + by case: (s1 =P s) => [<- // | _] /idPn[]; apply: neq0CiG. + split; first 2 [by rewrite mem_seqInd ?gFnormal ?(subsetP sXthXH0C)]. + rewrite Ds cfIirrE ?irrXtheta ?cfInd1 // -Du -(index_sdprod defM) -/q. + by rewrite mulrA -natrM lin_char1 ?cfMod_lin_char ?mulr1. + exists (theta f %% H0)%CF; first by rewrite cfMod_lin_char. + by rewrite Ds cfIirrE ?irrXtheta //= cfIndInd. + suffices /(congr1 odd): u = (p.-1 ^ q.-1)%N. + rewrite odd_exp -(subnKC (prime_gt1 pr_q)) /= -subn1 odd_sub ?prime_gt0 //. + by rewrite -oH1 (oddSg sH1H) ?quotient_odd // mFT_odd. + have p1_gt0: (0 < p.-1)%N by rewrite -(subnKC (prime_gt1 p_pr)). + apply/eqP; rewrite -(eqn_pmul2r p1_gt0) -expnSr prednK ?prime_gt0 //. + by rewrite -oXtheta -defXmu card_in_imset // cardC1 card_Iirr_abelian ?oH1. +clear Xmu def_IXmu Smu sSmu_mu ResIndXmu uSmu sz_Smu sz_mu s_mu_H0C Dmu. +clear Mtheta Xtheta irrXtheta oXtheta sXthXH0C mu_f Fmu_f mk_mu sW1_Imu inj_mu. +clear nz_thetaH lin_thetaH lin_theta Ftheta inj_theta irr_thetaH0 def_Itheta. +clear theta Dtheta => irr_qa lb_n lb_d. +have sU'U: U' \subset U := der_sub 1 U. +have nH0U := subset_trans sUHU nH0HU; have nH0U' := subset_trans sU'U nH0U. +have sU'CH1: U' \subset 'C_U(H1 | 'Q). + by rewrite subsetI sU'U sub_astabQ nH0U' (centsS sH1H) ?quotient_cents. +have sCH1_U: 'C_U(H1 | 'Q) \subset U := subsetIl U _. +have dvd_lb: lb_d %| lb_n. + rewrite -[lb_d]mulnA dvdn_mul // -(Lagrange sCH1_U). + by rewrite mulnC dvdn_pmul2r ?cardSg ?indexg_gt0. +split; rewrite ?leq_divLR // /lb_n -(Lagrange sCH1_U) -/a -(Lagrange sU'CH1). +rewrite mulnCA -mulnA mulnC !mulnA !leq_pmul2r ?cardG_gt0 ?indexg_gt0 // mulnC. +pose H1CH1 := (H1 <*> 'C_(U / H0)(H1))%G; pose HCH1 := (H <*> 'C_U(H1 | 'Q))%G. +have defH1CH1: H1 \x 'C_(U / H0)(H1) = H1CH1. + rewrite dprodEY ?subsetIr ?coprime_TIg ?(coprimeSg sH1H) //. + by rewrite (coprimegS (subsetIl _ _)) ?coprime_morph. +have sHCH1_HU: HCH1 \subset HU by rewrite join_subG sHHU (subset_trans sCH1_U). +have nsHCH1_HU: HCH1 <| HU. + rewrite /normal sHCH1_HU -(sdprodW defHU) mulG_subG normsG ?joing_subl //=. + by rewrite normsY // sub_der1_norm. +have nsH0_HCH1: H0 <| HCH1. + by rewrite (normalS _ sHCH1_HU) // (subset_trans sH0H) ?joing_subl. +have nsH1cHU: H1c <| HU / H0. + rewrite -(bigdprodWY defH1c) /normal gen_subG norms_gen ?andbT //. + by apply/bigcupsP=> w /setD1P[_ /nsH1wHUb/andP[]]. + by apply/norms_bigcup/bigcapsP=> w /setD1P[_ /nH1wHUb]. +have defHCH1: H1c ><| H1CH1 = (HCH1 / H0)%G. + have /sdprodP[_ /mulG_sub[sH1cH _] nH1cH1 tiH1cH1] := dprodWsdC defHb1. + rewrite sdprodE /= -(dprodW defH1CH1). + - rewrite mulgA (dprodWC defHb1) -morphim_setIpre -astabQ -quotientMl //. + by rewrite norm_joinEr // (subset_trans sCH1_U). + - rewrite mul_subG ?subIset // (subset_trans (quotientS _ sUHU)) //. + exact: normal_norm nsH1cHU. + rewrite -(setIidPl sH1cH) setIAC -setIA -group_modl // coprime_TIg ?mulg1 //. + by rewrite coprime_sym (coprimegS (subsetIl _ _)) ?coprime_morph. +have [nsH1cHCH1 sH1CH1_HCH1 _ nH1cH1C _] := sdprod_context defHCH1. +pose Clam := ('C_(U / H0)(H1) / (U' / H0))%G. +pose lam (j : Iirr Clam) := 'chi_(mod_Iirr j). +pose theta i j := cfSdprod defHCH1 (cfDprod defH1CH1 'chi_i (lam j)). +have nsU'CH1: U' <| 'C_U(H1 | 'Q) by rewrite (normalS _ sCH1_U) ?gFnormal. +have nsU'CH1b: U' / H0 <| 'C_(U / H0)(H1). + by rewrite -morphim_setIpre -astabQ quotient_normal. +have abClam: abelian Clam. + by rewrite sub_der1_abelian //= quotient_der ?dergS ?subsetIl. +have lam_lin j: lam j \is a linear_char. + by rewrite /lam mod_IirrE ?cfMod_lin_char //; apply/char_abelianP. +have theta_lin i j: theta i j \is a linear_char. + by rewrite cfSdprod_lin_char ?cfDprod_lin_char. +have <-: #|Clam| = #|'C_U(H1 | 'Q) : U'|. + rewrite -card_quotient ?normal_norm //= /= -morphim_setIpre -astabQ. + have nsU'U : U' <| U by apply: der_normal. + rewrite -(restrmEsub _ _ sCH1_U) -(restrm_quotientE _ sU'U) -morphim_quotm. + rewrite card_injm ?quotientS ?injm_quotm ?(isom_inj (quotient_isom _ _)) //. + by rewrite coprime_TIg ?(coprimeSg sH0H). +pose Mtheta := [set mod_Iirr (cfIirr (theta i j)) | i in [set~ 0], j in setT]. +have ->: (p.-1 * #|Clam|)%N = #|Mtheta|. + rewrite [Mtheta]curry_imset2X card_imset ?cardsX => [|[i1 j1] [i2 j2] /=/eqP]. + by rewrite cardsC1 cardsT !card_Iirr_abelian ?(abelianS sH1H) ?oH1. + rewrite (can_eq (mod_IirrK _)) // -(inj_eq irr_inj) !cfIirrE ?lin_char_irr //. + rewrite (can_eq (cfSdprodK _)) -!dprod_IirrE (inj_eq irr_inj). + by rewrite (can_eq (dprod_IirrK _)) => /eqP[->] /(can_inj (mod_IirrK _))->. +have{lam_lin} thetaH1 i j: 'Res[H1] (theta i j) = 'chi_i. + rewrite -(cfResRes _ _ sH1CH1_HCH1) ?joing_subl // cfSdprodK cfDprodKl //. + exact: lin_char1. +have Itheta r: r \in Mtheta -> 'I_HU['chi_r]%CF = HCH1. + case/imset2P=> i j; rewrite /= in_setC1 => nz_i _ Dr; apply/eqP. + rewrite eqEsubset sub_Inertia //= Dr mod_IirrE // cfIirrE ?lin_char_irr //. + rewrite andbT -(quotientSGK _ (normal_sub nsH0_HCH1)) ?subIset ?nH0HU //. + rewrite inertia_mod_quo //. + apply: subset_trans (sub_inertia_Res _ (nH1wHUb _ (group1 _))) _. + rewrite /= conjsg1 thetaH1 (inertia_irr_prime _ p_pr) //. + rewrite -quotient_setIpre -astabQ quotientS // -{1}(sdprodW defHU). + by rewrite -genM_join sub_gen // group_modl // sub_astabQ nH0H (centsS sH1H). +have irr_Xtheta: {in Mtheta, forall r, 'Ind[HU] 'chi_r \in irr HU}. + by move=> r Mr; rewrite /= inertia_Ind_irr ?Itheta. +pose Xtheta := [set cfIirr ('Ind[HU] 'chi_r) | r in Mtheta]. +have Da: a = #|HU : HCH1| by rewrite -(index_sdprodr defHU). +have Xtheta_1: {in Xtheta, forall s, 'chi_s 1%g = a%:R}. + move=> _ /imsetP[r Mr ->]; have /imset2P[i j _ _ Dr] := Mr. + rewrite cfIirrE ?irr_Xtheta ?cfInd1 //= -Da lin_char1 ?mulr1 //. + by rewrite Dr mod_IirrE ?cfMod_lin_char // cfIirrE ?lin_char_irr. +have nsH0U'HU: H0U' <| HU. + by apply: normalS nsH0U'_M; rewrite // -(sdprodWY defHU) genS ?setUSS. +have sXthetaXH0U': Xtheta \subset X_ H0U'. + apply/subsetP=> _ /imsetP[r Mr ->]; have [i j nz_i _ Dr] := imset2P Mr. + rewrite !inE cfIirrE ?irr_Xtheta ?sub_cfker_Ind_irr //= ?normal_norm //. + rewrite Dr mod_IirrE // cfIirrE ?lin_char_irr // join_subG andbCA. + rewrite {1}cfker_mod //= !{1}sub_cfker_mod //; apply/andP; split; last first. + rewrite -(sdprodWY (sdprod_cfker _ _)) sub_gen ?subsetU // orbC. + rewrite (subset_trans _ (cfker_dprod _ _ _)) // sub_gen ?subsetU // orbC. + by rewrite /lam mod_IirrE ?cfker_mod. + apply: contraL nz_i => /(subset_trans sH1H); rewrite !inE negbK. + by move/(sub_cfker_Res (subxx _)); rewrite thetaH1 subGcfker. +have nsCH1_U: 'C_U(H1 | 'Q) <| U by rewrite sub_der1_normal. +have nH1cU: (U / H0)%g \subset 'N(H1c). + rewrite -(bigdprodWY defH1c) norms_gen ?norms_bigcup //. + apply/bigcapsP=> w /setD1P[_ W1w]. + by rewrite normJ -sub_conjgV (normsP (quotient_norms H0 nUW1)) ?groupV. +have ->: #|Mtheta| = (#|Xtheta| * a)%N. + rewrite Da mulnC -card_imset_Ind_irr // => _ xy /imset2P[i j nz_i _ ->]. + case/(mem_sdprod defHU)=> x [y [Hx Uy -> _]]; have HUy := subsetP sUHU y Uy. + pose yb := inMb y; have Uyb: yb \in (U / H0)%g by rewrite mem_quotient. + pose iy := conjg_Iirr i yb; pose jy := conjg_Iirr j (coset (U' / H0)%g yb). + apply/imset2P; exists iy jy; rewrite !inE ?conjg_Iirr_eq0 // in nz_i *. + apply: irr_inj; have HCH1x: x \in HCH1 by rewrite mem_gen ?inE ?Hx. + rewrite conjg_IirrE (cfConjgM _ nsHCH1_HU) ?(subsetP sHHU x) {Hx}//. + rewrite {x HCH1x}(cfConjg_id _ HCH1x) !{1}mod_IirrE //. + rewrite !{1}cfIirrE ?lin_char_irr //. + rewrite cfConjgMod_norm ?(subsetP nH0U) ?(subsetP (normal_norm nsHCH1_HU)) //. + have nCH1_Ub: (U / H0)%g \subset 'N('C_(U / H0)(H1)). + by rewrite normsI ?normG ?norms_cent. + rewrite cfConjgSdprod ?cfConjgDprod ?(subsetP _ _ Uyb) ?normsY //. + rewrite /theta /lam !{1}mod_IirrE // !{1}conjg_IirrE. + by rewrite cfConjgMod_norm ?(subsetP _ _ Uyb) // quotient_norms ?gFnorm. +rewrite leq_pmul2r ?indexg_gt0 // cardE -(size_map (fun s => 'Ind[M] 'chi_s)). +have kerH1c s: s \in Xtheta -> H1c \subset (cfker 'chi_s / H0)%g. + case/imsetP=> r Mr ->; have [i j _ _ Dr] := imset2P Mr. + rewrite -(setIidPr (normal_sub nsH1cHCH1)) -morphim_setIpre quotientS //. + rewrite cfIirrE ?irr_Xtheta ?sub_cfker_Ind_irr //; last first. + by rewrite normsI ?normal_norm // -(quotientGK nsH0_HU) cosetpre_normal. + rewrite Dr mod_IirrE // cfker_morph ?normal_norm // cfIirrE ?lin_char_irr //. + by rewrite setIS ?joing_subl ?morphpreS // cfker_sdprod. +have injXtheta: + {in M & Xtheta &, forall w s1 s2, 'chi_s1 = 'chi_s2 ^ w -> w \in HU}%CF. +- move=> _ s1 s2 /(mem_sdprod defM)[y [w [HUy W1w -> _]]] Xs1 Xs2. + rewrite groupMl // cfConjgMnorm ?(subsetP (normG _) y) ?(subsetP nHUW1) //. + rewrite {y HUy}(cfConjg_id _ HUy) => Ds1. + have nH0w: w \in 'N(H0) by rewrite ?(subsetP nH0M) ?(subsetP sW1M). + rewrite (subsetP (normal_sub nsH0_HU)) // coset_idr //. + have /setDP[]:= subsetP sXthetaXH0U' s1 Xs1; rewrite !inE join_subG /=. + case/andP=> kerH0s1 _; apply: contraNeq; rewrite -eq_invg1 => ntw. + rewrite -(quotientSGK nH0H) // -(dprodW defHb1) mul_subG ?kerH1c //=. + rewrite Ds1 cfker_conjg ?(subsetP nHUW1) // quotientJ // -sub_conjgV. + rewrite (subset_trans _ (kerH1c s2 Xs2)) // -(bigdprodWY defH1c) sub_gen //. + by rewrite (bigcup_max (inMb w)^-1%g) // !inE ntw groupV mem_quotient. +rewrite -size_filter uniq_leq_size //. + apply/dinjectiveP=> s1 s2 Xs1 Xs2. + case/(cfclass_Ind_irrP _ _ (der_normal 1 M))/cfclassP=> y My Ds2. + by apply: irr_inj; rewrite Ds2 cfConjg_id ?(injXtheta y s1 s2). +move=> _ /imageP[s Xs ->]; rewrite mem_filter /= cfInd1 // -(index_sdprod defM). +rewrite Xtheta_1 // -natrM eqxx mem_seqInd ?gFnormal //. +rewrite (subsetP sXthetaXH0U') // !andbT inertia_Ind_irr ?gFnormal //. +by apply/subsetP=> y /setIP[My /inertiaJ/esym/injXtheta->]. +Qed. + +Import ssrnum Num.Theory. + +(* This is Peterfalvi (9.9); we have exported the fact that HU / H0 is a *) +(* Frobenius group in case (c), as this is directly used in (9.10). *) +Lemma typeP_Galois_characters (is_Galois : typeP_Galois) : + [/\ (*a*) {in X_ H0, forall s, (u %| 'chi_s 1%g)%Cx}, + {in X_ H0C', forall s, 'chi_s 1%g = u%:R /\ + (exists2 xi : 'CF(HC), xi \is a linear_char & 'chi_s = 'Ind xi)}, + (*b*) size mu_ = p.-1 /\ {in mu_, forall mu_j, isIndHC mu_j} + & (*c*) all redM (S_ H0C') -> + [/\ C :=: 1%g, u = (p ^ q).-1 %/ p.-1 + & [Frobenius HU / H0 = Hbar ><| (U / H0)]]]. +Proof. +have [F [phi [psi _ [Kpsi phiJ]]]] := typeP_Galois_P is_Galois. +case=> [oF /isomP[inj_phi im_phi] phiW2] [cycUbar co_u_p1 u_dv_pq1]. +have [nsHUM sW1M /mulG_sub[sHUM _] nHUW1 tiHUW1] := sdprod_context defM. +have [nsHHU sUHU /mulG_sub[sHHU _] nHU tiHU] := sdprod_context defHU. +have [nsH0C_M nsHC_M _ nsH0C'_M] := nsH0xx_M; have nH0H := normal_norm nsH0H. +have nsH0HU: H0 <| HU := normalS (subset_trans sH0H sHHU) sHUM nsH0M. +have nH0U: U \subset 'N(H0) := subset_trans sUHU (normal_norm nsH0HU). +have nH0C := subset_trans sCU nH0U. +have sH0C_HU: H0C \subset HU by rewrite -(sdprodWY defHU) genS ?setUSS. +have nsH0C_HU: H0C <| HU := normalS sH0C_HU sHUM nsH0C_M. +have nH0C_HU := normal_norm nsH0C_HU. +have [coH0U coHC] := (coprimeSg sH0H coHU, coprimegS sCU coHU). +have [nH0C_H nH0C_U] := (subset_trans sHHU nH0C_HU, subset_trans sUHU nH0C_HU). +have tiHOC_H: H0C :&: H = H0. + by rewrite /= norm_joinEr // -group_modl // setIC coprime_TIg ?mulg1. +have{coH0U} tiHOC_U: H0C :&: U = C. + by rewrite /= norm_joinEr // setIC -group_modr // setIC coprime_TIg ?mul1g. +have isoHbar: Hbar \isog H / H0C. + by have:= second_isog nH0C_H; rewrite tiHOC_H. +have isoUbar: Ubar \isog U / H0C. + by have:= second_isog nH0C_U; rewrite tiHOC_U. +have frobHU: [Frobenius HU / H0C = (H / H0C) ><| (U / H0C)]. + have defHUbar: (H / H0C) ><| (U / H0C) = (HU / H0C)%g. + exact: quotient_coprime_sdprod. + apply/Frobenius_semiregularP=> //; first by rewrite -(isog_eq1 isoHbar). + by rewrite -(isog_eq1 isoUbar); have [] := Frobenius_context frobUW1c. + move=> yb /setD1P[ntyb /morphimP[y nH0Cy Uy] Dyb] /=; rewrite Dyb. + apply/trivgP/subsetP=> _ /setIP[/morphimP[/= x nHOCx Hx ->] /cent1P/commgP]. + rewrite -morphR //; set xy := [~ x, y] => /eqP/coset_idr/=H0Cxy. + have [nH0x nH0y] := (subsetP nH0H x Hx, subsetP nH0U y Uy). + rewrite inE coset_id ?mem_gen // inE coset_idr //; apply: contraNeq ntyb. + rewrite -(morph_injm_eq1 inj_phi) ?mem_quotient // => nz_x. + rewrite {yb}Dyb /= coset_id ?mem_gen // -Kpsi !inE Uy orbC /= -val_eqE. + rewrite -(inj_eq (mulfI nz_x)) mulr1 -[_ * _]phiJ ?mem_quotient // qactJ nH0y. + rewrite -morphJ // conjg_mulR -/xy mkerr ?eqxx // ker_coset -tiHOC_H inE. + by rewrite andbC groupM ?groupV ?memJ_norm ?(subsetP nHU) //= H0Cxy ?groupR. +have{coHC} tiHbC: (Hbar :&: C / H0 = 1)%g by rewrite coprime_TIg ?coprime_morph. +have{tiHbC} defHCbar: Hbar \x (C / H0) = (HC / H0)%g. + by rewrite dprodEY ?quotientY // [C]unlock/= astabQ quotient_setIpre subsetIr. +have sHCHU: HC \subset HU by rewrite -(sdprodWY defHU) genS ?setUS. +have nsHCHU: HC <| HU := normalS sHCHU sHUM nsHC_M. +have sH0HC: H0 \subset HC := subset_trans sH0H (joing_subl H C). +have nsH0HC: H0 <| HC := normalS sH0HC sHCHU nsH0HU. +have nsHHUb: Hbar <| HU / H0 by rewrite quotient_normal. +have I_XH0_C i: i != 0 -> 'I_HU['chi[Hbar]_i %% H0]%CF = HC. + move=> /= nz_i; apply/esym/eqP. + have nsCHUb: C / H0 <| HU / H0 by rewrite -quotientYidl ?quotient_normal. + have sH0C_HC: H0C \subset HC by rewrite genS ?setSU. + have nsH0C_HC: H0C <| HC := normalS sH0C_HC sHCHU nsH0C_HU. + have [j Dj]: exists j, 'chi_j = (cfDprodl defHCbar 'chi_i %% H0)%CF. + by rewrite -dprodl_IirrE -mod_IirrE //; set j := mod_Iirr _; exists j. + have kerH0Cj: H0C \subset cfker 'chi_j. + by rewrite Dj sub_cfker_mod ?join_subG ?normG // quotientYidl ?cfker_sdprod. + rewrite inertia_mod_pre // -(inertia_dprodl defHCbar) ?normal_norm //. + rewrite -inertia_mod_pre // -Dj eqEsubset sub_Inertia //=. + rewrite -(quotientSGK _ sH0C_HC) ?subIset ?nH0C_HU -?inertia_quo //. + rewrite -(quotientYidr nH0C_H) joingA (joing_idPl sH0H) in frobHU. + rewrite -?quo_IirrE ?(inertia_Frobenius_ker (FrobeniusWker frobHU)) //. + by rewrite quo_Iirr_eq0 // -irr_eq1 Dj cfMod_eq1 // cfDprodl_eq1 irr_eq1. +have{I_XH0_C} irr_IndHC r: r \in Iirr_kerD HC H H0 -> 'Ind 'chi_r \in irr HU. + rewrite !inE => /andP[ker'H kerH0]; apply: inertia_Ind_irr => //. + apply: subset_trans (sub_inertia_Res _ (normal_norm nsHHU)) _. + rewrite -{kerH0}(quo_IirrK _ kerH0) // mod_IirrE // in ker'H *. + have /codomP[[i j] Dr] := dprod_Iirr_onto defHCbar (quo_Iirr H0 r). + rewrite {r}Dr dprod_IirrE cfResMod ?joing_subl ?sub_cfker_mod //= in ker'H *. + rewrite cfDprod_Resl linearZ inertia_scale_nz ?irr1_neq0 ?I_XH0_C //. + by apply: contraNneq ker'H => ->; rewrite irr0 cfDprod_cfun1l cfker_sdprod. +have [nb_mu H0C_mu] := nb_redM_H0; set part_a' := ({in X_ H0C', _}). +have Part_a s: s \in X_ H0 -> exists r, 'chi_s = 'Ind[HU, HC] 'chi_r. + rewrite !inE => /andP[Ks'H KsH0]; have [r sHCr] := constt_cfRes_irr HC s. + have{KsH0} KrH0: H0 \subset cfker 'chi_r. + by rewrite (sub_cfker_constt_Res_irr sHCr) // ?normal_norm. + have{Ks'H} Kr'H: ~~ (H \subset cfker 'chi_r). + by rewrite (sub_cfker_constt_Res_irr sHCr) ?joing_subl // ?normal_norm. + have [|s1 Ds1] := irrP _ (irr_IndHC r _); first by rewrite !inE Kr'H. + rewrite -constt_Ind_Res Ds1 constt_irr inE in sHCr. + by rewrite (eqP sHCr) -Ds1; exists r. +have [nH0HC nH0C'] := (normal_norm nsH0HC, subset_trans (der_sub 1 _) nH0C). +have Part_a': part_a'. + move=> s /setDP[KsH0C' Ks'H]; have [|r Ds] := Part_a s. + by rewrite inE Ks'H (subsetP (Iirr_kerS _ _) _ KsH0C') ?joing_subl. + suffices lin_r: 'chi_r \is a linear_char. + by split; [rewrite Du Ds cfInd1 ?lin_char1 ?mulr1 | exists 'chi_r]. + have KrH0C': H0C' \subset cfker 'chi_r. + rewrite inE Ds sub_cfker_Ind_irr // in KsH0C'. + by rewrite (subset_trans sHUM) ?normal_norm. + rewrite lin_irr_der1 (subset_trans _ KrH0C') //= (norm_joinEr nH0C'). + rewrite -quotientSK ?(subset_trans (der_sub 1 _)) ?quotient_der //= -/C. + by rewrite -(der_dprod 1 defHCbar) (derG1P abHbar) dprod1g. +split=> // [s /Part_a[r ->] | | {Part_a' part_a'}red_H0C']. +- by rewrite Du cfInd1 // dvdC_mulr // Cint_Cnat ?Cnat_irr1. +- split=> // mu_j /H0C_mu H0C_mu_j; have [s XH0Cs Dmuj] := seqIndP H0C_mu_j. + have [|s1u [xi lin_xi Ds]] := Part_a' s. + by rewrite (subsetP _ _ XH0Cs) ?Iirr_kerDS // genS ?setUS ?der_sub. + split=> //; first by rewrite Dmuj cfInd1 // s1u -natrM -(index_sdprod defM). + by rewrite Dmuj Ds cfIndInd //; exists xi. +have C1: C :=: 1%g. + apply: contraTeq red_H0C' => ntC; apply/allPn. + have sCM: C \subset M := subset_trans sCU (subset_trans sUHU sHUM). + have{sCM} solCbar: solvable (C / H0). + by rewrite quotient_sol ?(solvableS sCM) ?mmax_sol. + have [|{ntC solCbar} j lin_j nz_j] := solvable_has_lin_char _ solCbar. + rewrite -(isog_eq1 (quotient_isog _ _)) ?(subset_trans sCU) //. + by rewrite coprime_TIg ?(coprimegS sCU) ?(coprimeSg sH0H). + have [i lin_i nz_i] := solvable_has_lin_char ntHbar solHbar. + pose r := mod_Iirr (dprod_Iirr defHCbar (i, j)). + have KrH0: H0 \subset cfker 'chi_r by rewrite mod_IirrE ?cfker_mod. + have Kr'H: ~~ (H \subset cfker 'chi_r). + rewrite -subsetIidl -cfker_Res ?joing_subl ?irr_char // mod_IirrE //. + rewrite cfResMod ?joing_subl // sub_cfker_mod // dprod_IirrE. + by rewrite cfDprodKl ?lin_char1 // subGcfker -irr_eq1. + have [|s Ds] := irrP _ (irr_IndHC r _); first by rewrite !inE Kr'H. + have Ks'H: s \notin Iirr_ker HU H. + by rewrite inE -Ds sub_cfker_Ind_irr ?normal_norm. + exists ('Ind 'chi_s). + rewrite mem_seqInd ?gFnormal // inE Ks'H inE -Ds. + rewrite sub_cfker_Ind_irr // ?(subset_trans sHUM) ?normal_norm //=. + rewrite mod_IirrE // join_subG cfker_mod // sub_cfker_mod ?quotient_der //. + apply: subset_trans (dergS 1 (quotientS H0 (joing_subr H C))) _. + by rewrite -lin_irr_der1 dprod_IirrE cfDprod_lin_char. + apply: contra nz_j => red_j; have /implyP := H0C_mu ('Ind 'chi_s). + rewrite mem_filter red_j !mem_seqInd ?gFnormal // !in_setD Ks'H !inE -Ds. + rewrite irr_eq1 !sub_cfker_Ind_irr ?(normal_norm nsH0HU) //. + rewrite mod_IirrE // join_subG cfker_mod //= sub_cfker_mod // dprod_IirrE. + by move/(sub_cfker_Res (subxx _)); rewrite cfDprodKr ?lin_char1 ?subGcfker. +rewrite /= -/C C1 joingG1 in frobHU; split=> //; move/FrobeniusWker in frobHU. +have nsHbHU: Hbar <| (HU / H0) by rewrite quotient_normal. +have ->: (p ^ q).-1 = (#|X_ H0| * u)%N. + rewrite -oF -cardsT -im_phi card_injm // -(card_Iirr_abelian abHbar). + rewrite -(cardsC1 0) (card_imset_Ind_irr nsHbHU) => [|i|i y]; last first. + - by rewrite !inE conjg_Iirr_eq0. + - by rewrite !inE => nz_i; rewrite inertia_Ind_irr ?inertia_Frobenius_ker. + rewrite index_quotient_eq ?(subset_trans sHUM) ?subIset ?sH0H ?orbT //. + apply/eqP; rewrite Du /= C1 joingG1 mulnC eqn_pmul2r //. + rewrite -(card_imset _ (can_inj (mod_IirrK _))) // -imset_comp. + apply/eqP/eq_card=> s; apply/imsetP/idP=> [[i nz_i -> /=] | Xs]. + rewrite !inE mod_IirrE 1?{1}cfker_mod // andbT in nz_i *. + rewrite cfIirrE ?inertia_Ind_irr ?inertia_Frobenius_ker // sub_cfker_mod //. + by rewrite sub_cfker_Ind_irr ?quotientS ?normal_norm // subGcfker. + have [[]] := (Part_a s Xs, setDP Xs). + rewrite /= C1 joingG1 !inE => r Ds [kerH0s]. + have:= kerH0s; rewrite Ds !sub_cfker_Ind_irr ?normal_norm // => kerH0 ker'H. + exists (quo_Iirr H0 r). + by rewrite !inE -subGcfker quo_IirrE // cfker_quo ?quotientSGK. + by rewrite quo_IirrE // cfIndQuo // -Ds -quo_IirrE // irrK quo_IirrK. +suffices ->: #|X_ H0| = p.-1 by rewrite -(subnKC (prime_gt1 p_pr)) mulKn. +rewrite -nb_mu (size_red_subseq_seqInd_typeP MtypeP _ H0C_mu) //; last first. +- exact/allP/filter_all. +- by rewrite filter_uniq ?seqInd_uniq. +apply/esym/eq_card => i; rewrite inE mem_filter mem_seqInd ?gFnormal //. +rewrite andb_idl // => Xi; rewrite (allP red_H0C') //. +by rewrite mem_seqInd ?gFnormal //= C1 (trivgP (der_sub 1 _)) joingG1. +Qed. + +(* This combination of (9.8)(b) and (9.9)(b) covers most uses of these lemmas *) +(* in sections 10-14. *) +Lemma typeP_reducible_core_Ind (ptiWM := FT_primeTI_hyp MtypeP) : + [/\ size mu_ = p.-1, has predT mu_, + {subset mu_ <= [seq primeTIred ptiWM j | j in predC1 0]} + & {in mu_, forall mu_j, isIndHC mu_j}]. +Proof. +have [[sz_mu _] /mulG_sub[sHHU _]] := (nb_redM_H0, sdprodW defHU). +rewrite has_predT sz_mu -subn1 subn_gt0 prime_gt1 //; split=> // [mu_j|]. + rewrite mem_filter => /andP[red_chi /seqIndP[s /setDP[_ kerH's] Dchi]]. + have [[j Ds] | [/idPn[]]] := prTIres_irr_cases ptiWM s; last by rewrite -Dchi. + rewrite Dchi Ds cfInd_prTIres image_f ?inE //=. + by apply: contraNneq kerH's => j0; rewrite inE Ds j0 prTIres0 cfker_cfun1. +have[/typeP_Galois_characters[_ _ []] // | Gal'M] := boolP typeP_Galois. +by have [_ []] := typeP_nonGalois_characters Gal'M. +Qed. + +(* This is Peterfalvi (9.10), formulated as a constructive alternative. *) +Lemma typeP_reducible_core_cases : + {t : Iirr M & 'chi_t \in S_ H0C' /\ 'chi_t 1%g = (q * u)%:R + & {xi | xi \is a linear_char & 'chi_t = 'Ind[M, HC] xi}} + + [/\ typeP_Galois, [Frobenius HU / H0 = Hbar ><| (U / H0)], + cyclic U, #|U| = (p ^ q).-1 %/ p.-1 + & FTtype M == 2 -> [Frobenius HU = H ><| U]]. +Proof. +have [GalM | Gal'M] := boolP typeP_Galois; last first. + pose eqInHCb nu r := ('chi_r \is a linear_char) && (nu == 'Ind[M, HC] 'chi_r). + pose isIndHCb (nu : 'CF(M)) := + (nu 1%g == (q * u)%:R) && [exists r, eqInHCb nu r]. + suffices /sig2W[t H0C't]: exists2 t, 'chi_t \in S_ H0C' & isIndHCb 'chi_t. + case/andP=> /eqP t1qu /exists_inP/sig2W[r lin_r /eqP def_t]. + by left; exists t => //; exists 'chi_r. + have [_ _ [t [t1qu H0Ct IndHCt]] _] := typeP_nonGalois_characters Gal'M. + exists t; first by rewrite (seqIndS _ H0Ct) ?Iirr_kerDS ?genS ?setUS ?der_sub. + rewrite /isIndHCb t1qu eqxx; have [xi lin_xi ->] := IndHCt. + by apply/exists_inP; exists (cfIirr xi); rewrite cfIirrE ?lin_char_irr. +have [_ IndHC_SH0C' _] := typeP_Galois_characters GalM; rewrite all_predC. +case: hasP => [/sig2W[eta H0C'eta /irrP/sig_eqW[t Dt]] _ | _ [//|C1 <- frobHU]]. + have /sig2_eqW[s /IndHC_SH0C'[s1u IndHCs] Deta] := seqIndP H0C'eta. + have [joinHU [xi lin_xi Ds]] := (sdprodWY defHU, sig2_eqW IndHCs). + left; exists t; first split; rewrite -Dt // Deta. + by rewrite cfInd1 ?der_sub // -(index_sdprod defM) s1u -natrM. + by exists xi; rewrite ?Ds ?cfIndInd ?der_sub // -joinHU genS ?setUS ?subsetIl. +have cycU: cyclic U. + rewrite (isog_cyclic (quotient1_isog _)) -C1. + by have [_ _ []] := typeP_Galois_P GalM. +right; split=> //; first by rewrite /u /Ubar C1 -(card_isog (quotient1_isog _)). +case/(compl_of_typeII maxM MtypeP) => /= _ _ _ UtypeF <-. +have [_ -> _] := typeF_context UtypeF. +by apply/forall_inP=> S /and3P[_ /cyclicS->]. +Qed. + +Import ssrint. + +(* This is Peterfalvi (9.11) *) +(* We had to cover a small gap in step (9.11.4) of the proof, which starts by *) +(* proving that U1 \subset {1} u A(M), then asserts this obviously implies *) +(* HU1 \subset {1} u A(M). It is not, as while {1} u A(M) does contain H, it *) +(* is not (necessarily) a subgroup. We had to use the solvability of HU1 in a *) +(* significant way (using Philip Hall's theorems) to bridge the gap; it's *) +(* also possible to exploit lemma (2.1) (partition_cent_rcoset in PFsection1) *) +(* in a slightly different argument, but the inference is nontrivial in *) +(* either case. *) +Lemma Ptype_core_coherence : coherent (S_ H0C') M^# tau. +Proof. +have [nsHUM sW1M /mulG_sub[sHUM _] nHUW1 tiHUW1] := sdprod_context defM. +have [nsHHU sUHU /mulG_sub[sHHU _] nHU tiHU] := sdprod_context defHU. +have nsCU: C <| U := normalS sCU (joing_subl _ _) nsCUW1. +have [_ nCU] := andP nsCU; have sC'C: C^`(1)%g \subset C := der_sub 1 _. +have coHC := coprimegS sCU coHU; have coH0C := coprimeSg sH0H coHC. +have [nsH0C_M nsHC_M nsH0U'_M nsH0C'_M] := nsH0xx_M; have [_ nH0H]:= andP nsH0H. +have nH0HU := subset_trans sHUM nH0M; have nH0U := subset_trans sUHU nH0HU. +have nH0C := subset_trans sCU nH0U; have nH0C' := subset_trans sC'C nH0C. +have sHCHU: HC \subset HU by rewrite join_subG sHHU (subset_trans sCU). +have [nsHCHU nHC] := (normalS sHCHU sHUM nsHC_M, subset_trans sCU nHU). +have tiHCbar: Hbar :&: (C / H0)%g = 1%g by rewrite coprime_TIg ?coprime_morph. +have defHCbar: Hbar \x (C / H0) = (HC / H0)%g. + by rewrite dprodEY ?quotientY // [C]unlock/= astabQ quotient_setIpre subsetIr. +have{tiHCbar} defHC'bar: (HC / H0)^`(1)%g = (C^`(1) / H0)%g. + by rewrite -(der_dprod 1 defHCbar) (derG1P abHbar) dprod1g quotient_der. +have sU'U := der_sub 1 U; have nH0U' := subset_trans sU'U nH0U. +have sU'C: U' \subset C. + by rewrite [C]unlock subsetI sub_astabQ sU'U nH0U' quotient_cents. +have uS0: uniq (S_ H0C') by apply: seqInd_uniq. +have [rmR scohS0]: exists R : 'CF(M) -> seq 'CF(G), subcoherent (S_ H0C') tau R. + move: (FTtypeP_coh_base _ _) (FTtypeP_subcoherent maxM MtypeP) => R scohR. + exists R; apply: (subset_subcoherent scohR). + split=> //; last exact: cfAut_seqInd. + by apply: seqIndS; rewrite Iirr_kerDS ?sub1G ?Fcore_sub_FTcore. +have [GalM | Gal'M] := boolP typeP_Galois. + have [_ XOC'u _ _] := typeP_Galois_characters GalM. + apply: uniform_degree_coherence scohS0 _. + apply: all_pred1_constant (#|M : HU| * u)%:R _ _; rewrite all_map. + by apply/allP=> _ /seqIndP[s /XOC'u[s1u _] ->] /=; rewrite natrM cfInd1 ?s1u. +have:= typeP_nonGalois_characters Gal'M. +set U1 := 'C_U(_ | _); set a := #|_ : _|. +case: (_ Gal'M) => /= H1 [oH1 nH1U _ defHbar aP] in U1 a *. +rewrite -/U1 -/a in aP; case: aP => a_gt1 a_dv_p1 cycU1 _. +case=> [a_dv_XH0 [nb_mu IndHCmu] has_irrHC lb_Sqa]; rewrite -[S_ _]/(S_ H0C'). +have defHb1 := defHbar; rewrite (big_setD1 1%g) ?group1 ?conjsg1 //= in defHb1. +have [[_ H1c _ defH1c] _ _ _] := dprodP defHb1; rewrite defH1c in defHb1. +have [nsH1H _] := dprod_normal2 defHb1; have [sH1H _] := andP nsH1H. +have nsU1U: U1 <| U; last have [sU1U nU1U] := andP nsU1U. + by rewrite norm_normalI // astabQ norm_quotient_pre ?norms_cent. +have Da: a = #|HU : H <*> U1|. + rewrite /a -!divgS /= ?join_subG ?sHHU ?norm_joinEr ?(subset_trans sU1U) //=. + by rewrite -(sdprod_card defHU) coprime_cardMg ?(coprimegS sU1U) ?divnMl. +have sCU1: C \subset U1 by rewrite [C]unlock setIS ?astabS. +have a_dv_u: a %| u by rewrite Da Du indexgS ?genS ?setUS. +have [a_gt0 q_gt0 u_gt0 p1_gt0]: [/\ 0 < a, 0 < q, 0 < u & 0 < p.-1]%N. + rewrite !cardG_gt0 ltnW // -subn1 subn_gt0 prime_gt1 //. +have [odd_p odd_q odd_a]: [/\ odd p, odd q & odd a]. + by rewrite mFT_odd -oH1 (oddSg sH1H) ?(dvdn_odd a_dv_u) ?mFT_quo_odd. +have Dp: p = (2 * p.-1./2 + 1)%N. + by rewrite mul2n -[p]odd_double_half odd_p half_double addn1. +(* Start of main proof. *) +pose S1 := filter [pred zeta : 'CF(M) | zeta 1%g == (q * a)%:R] (S_ H0C'). +have ntS1: (0 < size S1)%N. + have [lb_dv lbS1] := lb_Sqa; apply: leq_trans (leq_trans lbS1 _). + by rewrite ltn_divRL // mul0n muln_gt0 p1_gt0 cardG_gt0. + rewrite -size_filter uniq_leq_size ?filter_uniq ?seqInd_uniq // => chi. + rewrite !mem_filter -andbA /= => /and3P[_ ->]. + by apply: seqIndS; rewrite Iirr_kerDS // genS ?setUS ?dergS ?subsetIl. +have sS10: cfConjC_subset S1 (S_ H0C'). + split=> [||chi]; first by rewrite filter_uniq. + by apply: mem_subseq; apply: filter_subseq. + rewrite !mem_filter !inE cfunE => /andP[/eqP <- S0chi]. + by rewrite cfAut_seqInd // andbT conj_Cnat ?(Cnat_seqInd1 S0chi). +have cohS1: coherent S1 M^# tau. + apply: uniform_degree_coherence (subset_subcoherent scohS0 sS10) _. + by apply: all_pred1_constant (q * a)%:R _ _; rewrite all_map filter_all. +pose S3 := filter [predC S1] (S_ H0C'); move: {2}_.+1 (ltnSn (size S3)) => nS. +move: @S3 (sS10) (cohS1); have: {subset S1 <= S1} by []. +elim: nS {-1}S1 => // nS IHnS S2 => sS12 S3 sS20 cohS2; rewrite ltnS => leS3nS. +have [ntS3|] := boolP (size S3 > 0)%N; last first. + rewrite size_filter -has_count has_predC negbK => /allP sS02. + exact: subset_coherent sS02 cohS2. +(* Ultimateley we'll contradict the maximality of S2 in (9.11.1) & (9.11.8). *) +suff [chi]: exists2 chi, chi \in S3 & coherent (chi :: chi^* :: S2)%CF M^# tau. + rewrite mem_filter => /andP[/= S2'chi S0chi]; have [uS2 sS2S0 ccS2] := sS20. + move/IHnS; apply=> [psi /sS12 S1psi||]; first by rewrite 2?mem_behead. + split. + - rewrite /= !inE negb_or S2'chi (contra (ccS2 _)) ?cfConjCK // eq_sym. + by rewrite (seqInd_conjC_neq _ _ _ S0chi) ?mFT_odd. + - by apply/allP; rewrite /= S0chi cfAut_seqInd //=; apply/allP. + apply/allP; rewrite /= !inE cfConjCK !eqxx orbT /=. + by apply/allP=> psi /ccS2; rewrite !inE orbA orbC => ->. + apply: leq_trans leS3nS; rewrite ltnNge; apply: contra S2'chi. + case/leq_size_perm=> [|psi|/(_ chi)]; first by rewrite filter_uniq. + by rewrite !mem_filter !inE orbA negb_or -andbA => /andP[]. + by rewrite !mem_filter !inE eqxx S0chi !andbT => /esym/negbFE. +(* This is step (9.11.1). *) clear nS IHnS leS3nS. +without loss [[eqS12 irrS1 H0C_S1] [Da_p defC] [S3qu ne_qa_qu] [oS1 oS1ua]]: + / [/\ [/\ S1 =i S2, {subset S1 <= irr M} & {subset S1 <= S_ H0C}], + a = p.-1./2 /\ C :=: U', + (forall chi, chi \in S3 -> chi 1%g == (q * u)%:R) /\ (q * u != q * a)%N + & size S1 = (p.-1 * u %/ a ^ 2)%N /\ size S1 = (2 * u %/ a)%N]. +- move=> IHwlog; have{sS20} [[uS2 sS20 ccS2] [uS1 _ _]] := (sS20, sS10). + pose is_qu := [pred chi : 'CF(M) | chi 1%g == (q * u)%:R]. + pose isn't_qu := [pred chi | is_qu chi ==> all is_qu S3]. + have /hasP[chi S3chi qu'chi]: has isn't_qu S3. + rewrite /isn't_qu; have [_|] := boolP (all _ _); last by rewrite has_predC. + by rewrite (eq_has (fun _ => implybT _)) has_predT. + have [S2'chi S0chi]: chi \notin S2 /\ chi \in S_ H0C'. + by apply/andP; rewrite mem_filter in S3chi. + have [s X0C's Dchi] := seqIndP S0chi. + have Dchi1: chi 1%g = q%:R * 'chi_s 1%g. + by rewrite Dchi cfInd1 // -(index_sdprod defM). + (* We'll show lb0 <= lb1 <= lb <= lb3 <= sumnS S1' <= sumnS S2 <= lb0, *) + (* with equality under conditions that imply the conclusion of (9.11.1). *) + pose lb0 := (2 * q * a)%:R * chi 1%g. + pose lb1 : algC := (2 * a * q ^ 2 * u)%:R. + pose lb2 : algC := (p.-1 * q ^ 2 * u)%:R. + pose lb3 : algC := (p.-1 * q ^ 2 * #|U : U'|)%:R. + pose S1' := filter [predI irr M & S_ H0U'] S1. + pose szS1' := ((p.-1 * #|U : U'|) %/ a ^ 2)%N; set lbS1' := _ %/ _ in lb_Sqa. + pose Snorm (psi : 'CF(M)) := psi 1%g ^+ 2 / '[psi]. + pose sumnS Si := \sum_(psi <- Si) Snorm psi. + have lb01: lb0 <= lb1 ?= iff (chi 1%g == (q * u)%:R). + rewrite /lb1 mulnA -mulnA natrM /lb0 mulnAC mono_lerif; last first. + by apply: ler_pmul2l; rewrite ltr0n !muln_gt0 a_gt0. + apply: lerif_eq; rewrite Dchi1 natrM ler_pmul2l ?gt0CG //. + have [KsH0C' _] := setDP X0C's; rewrite inE in KsH0C'. + have [t sHCt] := constt_cfRes_irr HC s. + have KtH0C': H0C' \subset cfker 'chi_t. + apply: subset_trans (cfker_constt (cfRes_char _ (irr_char s)) sHCt). + by rewrite cfker_Res ?irr_char // subsetI genS ?setUSS. + rewrite -constt_Ind_Res in sHCt. + apply: ler_trans (char1_ge_constt (cfInd_char _ (irr_char t)) sHCt) _. + rewrite cfInd1 // -Du lin_char1 ?mulr1 // lin_irr_der1. + apply: subset_trans KtH0C'; rewrite /= (norm_joinEr nH0C') /= -/C. + rewrite -quotientSK ?(subset_trans (der_sub _ _)) ?(subset_trans sHCHU) //. + by rewrite -defHC'bar quotient_der ?(subset_trans sHCHU). + have lb12: lb1 <= lb2 ?= iff (a == p.-1./2). + rewrite -(@eqn_pmul2l 2) // -(canLR (addnK 1) Dp) subn1 lerif_nat. + rewrite !(mono_leqif (fun _ _ => leq_pmul2r _)) ?expn_gt0 ?q_gt0 //. + apply: leqif_eq; rewrite dvdn_leq // Gauss_dvd //. + by rewrite {1}Dp addn1 dvdn_mulr. + by rewrite prime_coprime ?dvdn2 ?negbK. + have lb23: lb2 <= lb3 ?= iff (C :==: U') :> algC. + rewrite lerif_nat [u]card_quotient //. + rewrite (mono_leqif (fun _ _ => leq_pmul2l _)) ?muln_gt0 ?p1_gt0 ?q_gt0 //. + rewrite -(mono_leqif (fun _ _ => leq_pmul2l (cardG_gt0 C))) Lagrange //. + rewrite -(Lagrange sU'U) (mono_leqif (fun _ _ => leq_pmul2r _)) //. + by rewrite eq_sym; apply: subset_leqif_cards. + have lb3S1': lb3 <= sumnS S1' ?= iff (size S1' == szS1'). + rewrite /szS1' -(divnMr (cardG_gt0 U')) mulnAC -mulnA Lagrange // -/lbS1'. + have{lb_Sqa} [dv_lb lbSqa] := lb_Sqa; rewrite [sumnS S1']big_seq. + rewrite (eq_bigr (fun _ => ((q * a) ^ 2)%:R)) => [|psi]; last first. + rewrite !mem_filter -!andbA => /and4P[/irrP[r ->] _ /=/eqP r1qa _]. + by rewrite /Snorm cfnorm_irr divr1 r1qa natrX. + rewrite -big_seq (big_nth 0) -natr_sum sum_nat_const_nat subn0. + rewrite mulnC natrM [*%R]lock /lb3 natrM natf_indexg ?der_sub // mulrA. + rewrite -natrM mulnAC -(divnK dv_lb) mulnAC mulnA natrM mulfK ?neq0CG //. + rewrite -/lbS1' -mulnA -expnMn natrM mulrC -lock mono_lerif; last first. + by apply: ler_pmul2l; rewrite ltr0n !muln_gt0 a_gt0 q_gt0. + rewrite eq_sym lerif_nat; apply: leqif_eq; rewrite (leq_trans lbSqa) //. + rewrite -size_filter uniq_leq_size ?filter_uniq ?seqInd_uniq // => psi. + rewrite !mem_filter -!andbA /= => /and3P[-> -> S0psi]; rewrite S0psi. + by apply: seqIndS S0psi; rewrite Iirr_kerDS //= genS ?setUS ?dergS. + have lbS1'2: sumnS S1' <= sumnS S2 ?= iff ~~ has [predC S1'] S2. + have Ds2: perm_eq S2 (S1' ++ filter [predC S1'] S2). + rewrite -(perm_filterC (mem S1')) perm_cat2r. + rewrite uniq_perm_eq ?filter_uniq // => psi. + by rewrite mem_filter andb_idr //= mem_filter => /andP[_ /sS12]. + rewrite [sumnS S2](eq_big_perm _ Ds2) big_cat /= -/(sumnS S1') big_filter. + rewrite -all_predC -big_all_cond !(big_tnth _ _ S2) big_andE. + rewrite -{1}[_ S1']addr0 mono_lerif; last exact: ler_add2l. + set sumS2' := \sum_(i | _) _; rewrite -[0]/(sumS2' *+ 0) -sumrMnl. + apply: lerif_sum => i _; apply/lerifP; rewrite lt0r !mulf_eq0 invr_eq0. + set psi := tnth _ i; have Spsi := sS20 psi (mem_tnth _ _). + rewrite !negb_or (seqInd1_neq0 _ Spsi) //= (cfnorm_seqInd_neq0 _ Spsi) //=. + by rewrite divr_ge0 ?exprn_ge0 ?cfnorm_ge0 ?Cnat_ge0 ?(Cnat_seqInd1 Spsi). + have [lb0S2 | ] := boolP (lb0 < sumnS S2). + exists chi => //; have /hasP[xi S1xi _]: has predT S1 by rewrite has_predT. + have xi1: xi 1%g = (q * a)%:R. + by rewrite mem_filter in S1xi; have [/eqP] := andP S1xi. + apply: ((extend_coherent scohS0) _ xi) => //; first by rewrite S0chi sS12. + split=> //; last by rewrite mulrAC xi1 -natrM mulnA. + rewrite xi1 Dchi1 irr1_degree -natrM dvdC_nat dvdn_pmul2l ?cardG_gt0 //. + rewrite -dvdC_nat /= !nCdivE -irr1_degree a_dv_XH0 //. + by rewrite (subsetP (Iirr_kerDS _ _ _) _ X0C's) ?joing_subl. + have lb1S2 := lerif_trans lb12 (lerif_trans lb23 (lerif_trans lb3S1' lbS1'2)). + rewrite ltr_neqAle !(lerif_trans lb01 lb1S2) andbT has_predC !negbK. + case/and5P=> /eqP chi1qu /eqP Da_p /eqP defC /eqP sz_S1' /allP sS21'. + have defS1': S1' = S1. + apply/eqP; rewrite -(geq_leqif (size_subseq_leqif (filter_subseq _ S1))). + by rewrite uniq_leq_size // => psi /sS12/sS21'. + apply: IHwlog; split=> //. + + split=> psi; do 1?[rewrite -defS1' mem_filter andbC => /and3P[_ _] //=]. + by apply/idP/idP=> [/sS12 | /sS21']; rewrite ?defS1'. + by congr (_ \in S_ _); apply/group_inj; rewrite /= defC. + + split; first by apply/allP; apply: contraLR qu'chi; rewrite /= chi1qu eqxx. + rewrite -eqC_nat -chi1qu; apply: contra S2'chi => chi1qa. + by rewrite sS12 // mem_filter /= chi1qa. + rewrite -defS1' sz_S1' /szS1' -defC -card_quotient // -/u. + by split=> //; rewrite -mulnn {1}Dp addn1 -Da_p mulnAC divnMr. +have nCW1: W1 \subset 'N(C). + by rewrite (subset_trans (joing_subr U W1)) ?normal_norm. +(* This is step (9.11.2). *) clear sS20 cohS2 sS12 has_irrHC lb_Sqa sU'C. +have [tiU1 le_u_a2]: {in W1^#, forall w, U1 :&: U1 :^ w = C} /\ (u <= a ^ 2)%N. + have tiU1 w: w \in W1^# -> U1 :&: U1 :^ w = C; last split => //. + case/setD1P=> ntw W1w; have nH0w := subsetP (subset_trans sW1M nH0M) w W1w. + pose wb := coset H0 w; have W1wb: wb \in W1bar^#. + rewrite !inE mem_quotient ?(contraNneq _ ntw) // => /coset_idr H0w. + rewrite -in_set1 -set1gE -tiHUW1 inE (subsetP sHHU) // (subsetP sH0H) //. + by rewrite H0w. + have ntH1 w1: H1 :^ w1 :!=: 1%g by rewrite -cardG_gt1 cardJg oH1 prime_gt1. + pose t_ w1 := + if pred2 1%g wb w1 then sval (has_nonprincipal_irr (ntH1 w1)) else 0. + pose theta := + cfDprodl defHCbar (cfBigdprod defHbar (fun w1 => 'chi_(t_ w1))). + have lin_theta : theta \is a linear_char. + rewrite cfDprodl_lin_char ?cfBigdprod_lin_char // => w1 _. + by rewrite irr_prime_lin ?cardJg ?oH1. + have nsH0HC: H0 <| HC. + by rewrite /normal join_subG nH0H sub_gen ?subsetU ?sH0H. + move defK: H0C => K. (* to avoid divergence in Coq kernel *) + have kerK: K \subset cfker (theta %% H0). + rewrite -defK sub_cfker_mod ?join_subG ?normG ?quotientYidl //. + exact: cfker_sdprod. + have sKHC: K \subset HC by rewrite -defK genS ?setSU. + have nsKHC: K <| HC by rewrite (normalS sKHC (normal_sub nsHC_M)) -?defK. + have sH0K: H0 \subset K by rewrite -defK joing_subl. + have nsKHU: K <| HU. + by rewrite (normalS (subset_trans sKHC sHCHU) sHUM) -?defK. + have [t2 Dt2]: {t2 : Iirr (HC / K) | 'chi_t2 %% K = theta %% H0}%CF. + exists (cfIirr ((theta %% H0) / K)). + by rewrite cfIirrE ?cfQuoK ?cfQuo_irr ?cfMod_irr ?lin_char_irr. + have nsHChatHU: HC / K <| HU / K by rewrite quotient_normal. + have sHChatHU := normal_sub nsHChatHU. + pose That := 'I_(HU / K)['chi_t2]%G. + have sThatHU: That \subset (HU / K)%G := Inertia_sub _ _. + have abThatHC: abelian (That / (HC / K)). + rewrite (abelianS (quotientS _ sThatHU)) //. + rewrite (isog_abelian (third_isog _ _ _)) // defC. + rewrite -(isog_abelian (quotient_sdprodr_isog defHU _)) ?gFnormal //. + by rewrite sub_der1_abelian. + have hallHChat: Hall That (HC / K). + rewrite /Hall -(card_isog (third_isog sH0K nsH0HC nsKHC)) /=. + rewrite sub_Inertia // -[in #|_|]defK /= quotientYidl //. + rewrite -(card_isog (sdprod_isog (dprodWsdC defHCbar))). + apply: coprime_dvdr (indexSg (sub_Inertia _ sHChatHU) sThatHU) _. + apply: coprime_dvdr (index_quotient _) _. + by rewrite subIset // normal_norm. + by rewrite -Du coprime_morphl // coprime_morphr. + have [s t2HUs] := constt_cfInd_irr t2 sHChatHU. + have s_1: ('chi_s %% K)%CF 1%g = #|U : U1 :&: U1 :^ w|%:R. + rewrite cfMod1. + have [||_ _ ->] // := cfInd_Hall_central_Inertia _ abThatHC. + rewrite -cfMod1 Dt2 cfMod1 lin_char1 //= mulr1 -inertia_mod_quo // Dt2. + rewrite index_quotient_eq ?normal_norm ?Inertia_sub ?setIS //; last first. + by rewrite (subset_trans sKHC) ?sub_inertia. + rewrite /= inertia_morph_pre //= -quotientE inertia_dprodl; first 1 last. + - by rewrite quotient_norms ?normal_norm. + - rewrite /= -(quotientYidl nH0C) quotient_norms ?normal_norm //. + by rewrite -defK in nsKHU. + have nH1wHU w1: w1 \in (W1 / H0)%g -> (HU / H0)%g \subset 'N(H1 :^ w1). + move=> W1w1; rewrite -(normsP (quotient_norms H0 nHUW1) _ W1w1) normJ. + rewrite conjSg /= -(sdprodW defHU) quotientMl ?mul_subG //. + exact: normal_norm. + rewrite indexgI /= inertia_bigdprod_irr // (big_setD1 1%g) ?group1 //=. + rewrite 2!{1}setIA setIid (bigD1 wb) //= {1 2}/t_ /= !eqxx ?orbT /=. + rewrite !(inertia_irr_prime _ p_pr) ?cardJg //=; + try by case: (has_nonprincipal_irr _). + rewrite conjsg1 centJ setIA -setIIr /=. + elim/big_rec: _ => [|w1 Uk /andP[/setD1P[ntw1 Ww1] w'w1] IHk]; last first. + rewrite /t_ -if_neg negb_or ntw1 w'w1 irr0 Inertia1 -?setIIr 1?setIA //. + rewrite /normal nH1wHU //= -(normsP (quotient_norms H0 nHUW1) _ Ww1). + by rewrite conjSg (subset_trans sH1H) ?quotientS. + rewrite setIT !morphpreI morphpreJ ?(subsetP nH0W1) //= -astabQ. + rewrite quotientGK //; last by rewrite /normal (subset_trans sH0H). + rewrite -(sdprodWY (sdprod_modl defHU _)); last first. + rewrite subsetI -sub_conjgV. + rewrite (normsP (gFnorm _ _)) ?groupV ?(subsetP sW1M) //= andbb. + by rewrite sub_astabQ nH0H sub_abelian_cent. + rewrite -(index_sdprodr defHU) ?subsetIl // conjIg (normsP nUW1) //. + by rewrite -setIIr. + apply/esym/eqP; rewrite eqEcard subsetI -sub_conjgV. + rewrite (normsP _ _ (groupVr W1w)) ?sCU1 /=; last first. + by rewrite (subset_trans (joing_subr U W1)) ?normal_norm. + have{s_1} : pred2 u a #|U : U1 :&: U1 :^ w|. + rewrite /= -!eqC_nat -{}s_1 -mod_IirrE //. + pose phi := 'Ind[M] 'chi_(mod_Iirr s). + have Sphi: phi \in S_ H0C'. + rewrite mem_seqInd ?gFnormal // !inE mod_IirrE //. + rewrite andbC (subset_trans _ (cfker_mod _ _)) //=; last first. + by rewrite -defK genS ?setUS ?der_sub. + rewrite sub_cfker_mod ?(subset_trans sHHU) ?normal_norm //. + have sHHC: H \subset HC by rewrite joing_subl. + rewrite -(sub_cfker_constt_Ind_irr t2HUs) ?quotientS //; last first. + by rewrite quotient_norms ?normal_norm. + rewrite -sub_cfker_mod ?(subset_trans sHHU (normal_norm nsKHU)) //. + rewrite Dt2 sub_cfker_mod //. + apply: contra (valP (has_nonprincipal_irr (ntH1 1%g))). + move/eq_cfker_Res; rewrite cfDprodlK => kerHbar. + have:= sH1H; rewrite -{1}(conjsg1 H1) -kerHbar => /eq_cfker_Res. + by rewrite cfBigdprodKabelian ?group1 // /t_ /= eqxx -subGcfker => ->. + have [/S3qu | ] := boolP (phi \in S3). + rewrite cfInd1 // natrM -(index_sdprod defM). + by rewrite (inj_eq (mulfI (neq0CG _))) => ->. + rewrite mem_filter Sphi andbT negbK /= -eqS12 mem_filter Sphi andbT /=. + rewrite cfInd1 // natrM -(index_sdprod defM) (inj_eq (mulfI (neq0CG _))). + by rewrite orbC => ->. + case/pred2P=> [iUCu | iUCa]. + rewrite -(leq_pmul2r u_gt0) -{1}iUCu /u card_quotient ?Lagrange //. + by rewrite /= -setIA subsetIl. + rewrite subset_leq_card // subIset // [C]unlock subsetI subsetIl sub_astabQ. + rewrite subIset ?nH0U //= centsC -(bigdprodWY defHbar) gen_subG. + apply/orP; left; apply/bigcupsP=> w1 Ww1; rewrite centsC centJ -sub_conjgV. + rewrite (normsP _ _ (groupVr Ww1)) ?quotient_norms //. + by rewrite /U1 astabQ quotient_setIpre subsetIr. + rewrite prime_meetG //; apply/trivgPn; exists w; rewrite // !inE W1w. + rewrite (sameP setIidPr eqP) eqEcard subsetIr /= cardJg. + by rewrite -(leq_pmul2r a_gt0) -{2}iUCa !Lagrange //= -setIA subsetIl. + have /trivgPn[w W1w ntw]: W1 :!=: 1%g by rewrite -cardG_gt1 prime_gt1. + rewrite -(leq_pmul2l u_gt0) mulnn. + have nCU1 := subset_trans sU1U nCU. + have {1}->: u = (#|(U1 / C)%g| * a)%N. + by rewrite mulnC /u !card_quotient // Lagrange_index. + rewrite expnMn leq_pmul2r ?expn_gt0 ?orbF // -mulnn. + rewrite -{2}[U1](conjsgK w) quotientJ ?groupV ?(subsetP nCW1) //. + rewrite cardJg -TI_cardMg /= -/U1 ?subset_leq_card //. + rewrite mul_subG ?quotientS ?subsetIl //. + by rewrite -(normsP nUW1 w W1w) conjSg subsetIl. + by rewrite -quotientGI // tiU1 ?trivg_quotient // !inE ntw. +pose S4 := filter [predD S_ H0C & redM] S3. +have sS43: {subset S4 <= S3} by apply: mem_subseq; apply: filter_subseq. +(* This is step (9.11.3). *) +have nsHM: H <| M by apply: gFnormal. +have oS4: (q * u * size S4 + p.-1 * (q + u))%N = (p ^ q).-1. + rewrite mulnAC {1}[q](index_sdprod defM) -[S4]filter_predI. + rewrite (size_irr_subseq_seqInd _ (filter_subseq _ _)) //; last first. + by move=> xi; rewrite mem_filter -!andbA negbK => /andP[]. + set Xn := finset _; pose sumH0C := \sum_(s in X_ H0C) 'chi_s 1%g ^+ 2. + have /eqP: sumH0C = (q * size S1 * a ^ 2 + (#|Xn| + p.-1) * u ^ 2)%:R. + rewrite [q](index_sdprod defM) natrD natrM natrX. + rewrite (size_irr_subseq_seqInd _ (filter_subseq _ _)) //= -/S1. + have sH0CC': {subset S_ H0C <= S_ H0C'}. + by apply: seqIndS; rewrite Iirr_kerDS // genS ?setUS ?der_sub. + rewrite [sumH0C](big_setID [set s | 'Ind 'chi_s \in S1]) /=; congr (_ + _). + rewrite mulr_natl -[rhs in _ = rhs]sumr_const; apply: eq_big => s. + by rewrite in_setI inE andb_idl // => /H0C_S1; rewrite mem_seqInd. + rewrite 2!inE mem_filter andbCA /= cfInd1 // -(index_sdprod defM) natrM. + by case/andP=> /eqP/(mulfI (neq0CG _))->. + rewrite (eq_bigr (fun _ => u%:R ^+ 2)) => [|s]; last first. + rewrite 2!inE eqS12 => /andP[S2's H0Cs]; congr (_ ^+ 2). + have /S3qu/eqP: 'Ind 'chi_s \in S3. + by rewrite mem_filter /= S2's sH0CC' ?mem_seqInd. + by rewrite natrM cfInd1 // -(index_sdprod defM) => /(mulfI (neq0CG _)). + rewrite sumr_const -mulr_natl natrM natrX -nb_mu; congr (_%:R * _). + have [_ s_mu_H0C] := nb_redM_H0. + rewrite (size_red_subseq_seqInd_typeP MtypeP _ s_mu_H0C); last first. + - by apply/allP; apply: filter_all. + - by rewrite filter_uniq ?seqInd_uniq. + rewrite -/mu_ -[#|_|](cardsID Xn) (setIidPr _); last first. + apply/subsetP=> s; rewrite inE in_setD mem_filter /= -!andbA -eqS12. + rewrite mem_seqInd ?gFnormal // => /and4P[_ -> S1'xi _]. + by rewrite inE S1'xi. + congr (_ + _)%N; apply: eq_card => i; rewrite inE -/mu_ 2!inE. + rewrite -[seqInd M _]/(S_ H0C') mem_filter /= andbC 2!inE eqS12 -!andbA. + rewrite -(mem_seqInd nsHUM) // -/(S_ H0C); set xi := 'Ind _. + apply/idP/idP=> [/and3P[-> H0Cxi] | mu_xi]. + rewrite H0Cxi sH0CC' //= andbT negbK mem_filter unfold_in => ->. + by rewrite (seqIndS _ H0Cxi) // Iirr_kerDS ?joing_subl. + have [xi1u H0Cxi _] := IndHCmu _ mu_xi. + rewrite H0Cxi -eqS12 mem_filter sH0CC' //= !andbT xi1u eqC_nat ne_qa_qu. + by rewrite andbT negbK mem_filter in mu_xi *; case/andP: mu_xi. + rewrite oS1 -mulnA divnK ?dvdn_mul // !mulnA -mulnDl mulnC natrM {}/sumH0C. + rewrite /X_ -Iirr_kerDY joingC joingA (joing_idPl sH0H) /=. + rewrite sum_Iirr_kerD_square ?genS ?setSU //; last first. + by apply: normalS nsH0C_M; rewrite // -(sdprodWY defHU) genS ?setUSS. + rewrite -Du (inj_eq (mulfI (neq0CG _))) -(prednK (indexg_gt0 _ _)). + rewrite mulrSr addrK eqC_nat addnC mulnDl addnAC (mulnC q) -addnA -mulnDr. + move/eqP <-; congr _.-1. + have nH0HC: HC \subset 'N(H0) by rewrite join_subG nH0H. + rewrite -(index_quotient_eq _ _ nH0HC) ?genS ?setSU //; last first. + by rewrite setIC subIset ?joing_subl. + rewrite quotientYidl // -(index_sdprod (dprodWsdC defHCbar)). + by case: Ptype_Fcore_factor_facts. +have /hasP[psi1 S1psi1 _]: has predT S1 by rewrite has_predT. +pose gamma := 'Ind[M, H <*> U1] 1; pose alpha := gamma - psi1. +(* This is step (9.11.4). *) +pose nm_alpha : algC := a%:R + 1 + (q.-1 * a ^ 2)%:R / u%:R. +have [Aalpha Nalpha]: alpha \in 'CF(M, 'A(M)) /\ '[alpha] = nm_alpha. + have sHU1_HU: H <*> U1 \subset HU by rewrite -(sdprodWY defHU) genS ?setUS. + have sHU1_M := subset_trans sHU1_HU sHUM. + have sHU1_A1: H <*> U1 \subset 1%g |: 'A(M). + pose pi := \pi(H); rewrite -subDset; apply/subsetP=> y /setD1P[nty HU1y]. + apply/bigcupP; rewrite notMtype1 /=; have sHMs := Fcore_sub_FTcore maxM. + have defHU1: H ><| U1 = H <*> U1 := sdprod_subr defHU sU1U. + have nsH_HU1: H <| H <*> U1 by have [] := sdprod_context defHU1. + have [HUy [_ nH_HU1]] := (subsetP sHU1_HU y HU1y, normalP nsH_HU1). + have hallH: pi.-Hall(H <*> U1) H. + by rewrite Hall_pi // -(coprime_sdprod_Hall_l defHU1) (coprimegS sU1U). + have hallU1: pi^'.-Hall(H <*> U1) U1. + by rewrite -(compl_pHall _ hallH) sdprod_compl. + have [pi'y | pi_y] := boolP (pi^'.-elt y); last first. + exists y.`_pi; last by rewrite 3!inE nty HUy cent1C groupX ?cent1id. + rewrite !inE (sameP eqP constt1P) pi_y (subsetP sHMs) //. + by rewrite (mem_normal_Hall hallH) ?groupX ?p_elt_constt. + have solHU1: solvable (H <*> U1) by rewrite (solvableS sHU1_M) ?mmax_sol. + have [||z HU1z U1yz] := Hall_Jsub _ hallU1 _ pi'y; rewrite ?cycle_subG //. + have /trivgPn[x /setIP[Hx cxyz] ntx]: 'C_H[y ^ z] != 1%g. + apply: contraTneq (prime_gt1 p_pr) => regHy; rewrite -oH1 cardG_gt1 negbK. + move: U1yz; rewrite -cycleJ subsetI sub_astabQ => /and3P[sYU nHY cH1Y]. + rewrite centsC in cH1Y; rewrite -(setIidPl cH1Y) -(setIidPl sH1H) -setIA. + rewrite -coprime_quotient_cent ?(coprimegS sYU) // cent_cycle regHy. + by rewrite quotient1 setIg1. + exists (x ^ z^-1)%g; last by rewrite 3!inE nty HUy cent1J mem_conjgV cent1C. + by rewrite 2!inE conjg_eq1 ntx (subsetP sHMs) // -mem_conjg nH_HU1. + have{sHU1_A1} Aalpha: alpha \in 'CF(M, 'A(M)). + have A'1: 1%g \notin 'A(M) by have /subsetD1P[] := FTsupp_sub M. + rewrite -['A(M)](setU1K A'1) cfun_onD1 !cfunE subr_eq0 cfInd1 // cfun11. + rewrite mulr1 -(Lagrange_index sHUM) // -(index_sdprod defM) -/q. + rewrite -(index_sdprodr defHU) ?subsetIl // -/a eq_sym andbC. + have:= S1psi1; rewrite mem_filter /= => /andP[-> _] /=. + rewrite rpredB //. + apply: cfun_onS (cfInd_on sHU1_M (cfun_onG _)). + rewrite class_supportEr; apply/bigcupsP=> w Mw. + by rewrite sub_conjg conjUg conjs1g (normsP (FTsupp_norm M)) ?groupV. + have /seqIndP[s /setDP[_ ker'H ] ->] := H0C_S1 _ S1psi1. + rewrite (prDade_Ind_irr_on (FT_prDade_hypF maxM MtypeP)) //. + by rewrite inE in ker'H. + have ->: '[alpha] = '[gamma] + 1. + have /irrP[t Dt] := irrS1 _ S1psi1. + rewrite cfnormBd; first by rewrite Dt cfnorm_irr. + have /seqIndP[s /setDP[_ ker'H ] Dpsi1] := H0C_S1 _ S1psi1. + apply: contraNeq ker'H; rewrite Dt /gamma -irr0 -irr_consttE => tHU1_0. + rewrite inE -(sub_cfker_Ind_irr _ sHUM) ?gFnorm // -Dpsi1 Dt. + rewrite -(sub_cfker_constt_Ind_irr tHU1_0) ?gFnorm ?joing_subl //. + by rewrite irr0 cfker_cfun1 joing_subl. + split=> //; rewrite /nm_alpha addrAC natrM mulrAC mulrC; congr (_ + 1). + rewrite -{1}(mulnK a a_gt0) natf_div ?dvdn_mull // -mulrDr mulnn natrX. + have /sdprod_isom[nH_UW1 isomMH]: H ><| (U <*> W1) = M. + rewrite sdprodEY ?join_subG ?nHU ?(subset_trans sW1M) ?gFnorm //. + by rewrite joingA (sdprodWY defHU) (sdprodWY defM). + rewrite /= -(setIidPl sHHU) norm_joinEr // setIAC -setIA -group_modl //. + by rewrite (setIC W1) tiHUW1 mulg1. + have sU1_UW1: U1 \subset U <*> W1 by rewrite subIset ?joing_subl. + rewrite /gamma -(cfMod_cfun1 _ H) cfIndMod ?joing_subl //. + rewrite cfMod_iso //= quotientYidl ?(subset_trans sU1_UW1) //. + rewrite -(restrm_quotientE _ sU1_UW1) -(cfIsom_cfun1 (restr_isom _ isomMH)). + rewrite (cfIndIsom isomMH) // {nH_UW1 isomMH}cfIsom_iso. + rewrite -(cfIndInd _ (joing_subl U W1)) // cfInd_cfun1 //= -/U1 -/a. + rewrite linearZ cfnormZ normr_nat /=; congr (_ * _). + have defUW1: U ><| W1 = U <*> W1. + by rewrite sdprodEY // -(setIidPl sUHU) -setIA tiHUW1 setIg1. + apply: canLR (mulKf (neq0CG _)) _; rewrite -(sdprod_card defUW1) natrM -/q. + rewrite mulrAC mulrDr mulrCA -{1}(Lagrange sU1U) /= -/U1 -/a -(Lagrange sCU). + rewrite -card_quotient // !natrM !mulfK ?neq0CiG ?neq0CG //. + transitivity (\sum_(x in U <*> W1) \sum_(w1 in W1) \sum_(w2 in W1) + (x ^ w1 \in U1 :&: U1 :^ w2)%g%:R : algC). + - apply: eq_bigr => x _; rewrite (cfIndEsdprod _ _ defUW1) mulr_suml. + apply: eq_bigr => w1 W1w1; rewrite rmorph_sum mulr_sumr. + rewrite (reindex_inj invg_inj) (eq_bigl _ _ (groupV W1)) /=. + rewrite (reindex_acts 'R _ (groupVr W1w1)) ?astabsR //=. + apply: eq_bigr => w2 _; rewrite inE !cfuniE // rmorph_nat -natrM mulnb. + by congr (_ && _)%:R; rewrite invMg invgK conjgM -mem_conjg. + rewrite exchange_big /= mulr_natr -sumr_const; apply: eq_bigr => w1 W1w1. + transitivity (\sum_(w in W1) #|U1 :&: U1 :^ w|%:R : algC). + rewrite exchange_big /=; apply: eq_bigr => w W1w. + rewrite (reindex_acts 'J _ (groupVr W1w1)) ?astabsJ ?normsG ?joing_subr //=. + symmetry; rewrite big_mkcond -sumr_const /= big_mkcond /=. + apply: eq_bigr => x _; rewrite conjgKV. + by case: setIP => [[/(subsetP sU1_UW1)-> //] | _]; rewrite if_same. + rewrite (big_setD1 1%g) //= conjsg1 setIid; congr (_ + _). + rewrite [q](cardsD1 1%g) group1 /= mulr_natl -sumr_const. + by apply: eq_bigr => w W1w; rewrite tiU1. +(* This is step (9.11.5). *) +have [gtS4alpha s4gt0]: (size S4)%:R > '[alpha] /\ (size S4 > 0)%N. + suffices gtS4alpha: (size S4)%:R > '[alpha]. + by split=> //; rewrite -ltC_nat (ler_lt_trans (cfnorm_ge0 alpha)). + rewrite Nalpha -(@ltr_pmul2r _ u%:R) ?ltr0n // mulrDl divfK ?neq0CG //. + rewrite -(ltr_pmul2l (gt0CG W1)) -/q -mulrSr -!(natrM, natrD) ltC_nat. + rewrite mulnA mulnAC -(ltn_add2r (p.-1 * (q + u))) oS4 {1}Dp addn1 -Da_p /=. + apply: leq_ltn_trans (_ : q.+2 * a ^ 3 + q ^ 2 * a ^ 2 + 2 * q * a < _)%N. + rewrite (addnC q) 2!mulnDr addnA (mulnAC _ a q) leq_add2r. + rewrite mulnA addnAC -mulnDl mulnS -addnA -mulnDl addn2 mulnCA -mulnA. + rewrite -[q in (_ <= _ + q * _)%N](prednK q_gt0) (mulSn q.-1) addnA. + by rewrite leq_add2r mulnA -mulnDl addnC leq_mul. + have q_gt2: (2 < q)%N. + by rewrite ltn_neqAle prime_gt1 ?(contraTneq _ odd_q) => // <-. + apply: leq_trans (_ : a.*2 ^ q + 'C(q, 2) * a.*2 ^ 2 + q * a.*2 <= _)%N. + rewrite -mul2n (mulnCA q) (mulnA 2) ltn_add2r !expnMn -addSn leq_add //. + apply: leq_ltn_trans (_ : q.-1.*2.+1 * a ^ q < _)%N. + rewrite leq_mul ?leq_pexp2l //. + by rewrite -(subnKC q_gt2) -addnn !addnS !ltnS leq_addl. + rewrite ltn_pmul2r ?expn_gt0 ?a_gt0 // -doubleS. + by rewrite -(prednK q_gt0) expnS mul2n leq_double ltn_expl. + rewrite mulnA leq_pmul2r ?expn_gt0 ?a_gt0 // -(subnKC q_gt2). + rewrite mulnCA mulnA addSn -mul_Sm_binm bin1 -mulnA leq_pmul2l //. + by rewrite mulnS -addSnnS leq_addr. + rewrite Dp -Da_p mul2n (addnC a.*2) expnDn -(subnKC q_gt2) !addSn add0n. + rewrite 3!big_ord_recl big_ord_recr /= !exp1n /= bin1 binn !mul1n /bump /=. + by do 2!rewrite addnC leq_add2l; apply: leq_addl. +have{cohS1} [tau1 cohS1] := cohS1; have [[Itau1 Ztau1] Dtau1] := cohS1. +have sS30: cfConjC_subset S3 (S_ H0C'). + split=> [|chi|chi]; first by rewrite filter_uniq ?seqInd_uniq. + by rewrite mem_filter => /andP[]. + rewrite !mem_filter /= -!eqS12 => /andP[S1'chi S_chi]. + rewrite cfAut_seqInd // (contra _ S1'chi) //. + by have [_ _ ccS1] := sS10; move/ccS1; rewrite cfConjCK. +have scohS3: subcoherent S3 tau rmR := subset_subcoherent scohS0 sS30. +have [tau3 cohS3]: coherent S3 M^# tau. + apply: uniform_degree_coherence scohS3 _. + apply: all_pred1_constant (q * u)%:R _ _. + by rewrite all_map; apply/allP=> chi /S3qu. +have [IZtau3 Dtau3] := cohS3; have [Itau3 Ztau3] := IZtau3. +have notA1: 1%g \notin 'A(M) by have /subsetD1P[] := FTsupp_sub M. +have sS0_1A: {subset S_ H0C' <= 'CF(M, 1%g |: 'A(M))}. + move=> _ /seqIndP[s /setDP[_ ker'H] ->]; rewrite inE in ker'H. + by rewrite (prDade_Ind_irr_on (FT_prDade_hypF maxM MtypeP)). +have sS0A: {subset 'Z[S_ H0C', M^#] <= 'Z[irr M, 'A(M)]}. + move=> chi; rewrite (zcharD1_seqInd_Dade _ notA1) //. + by apply: zchar_sub_irr; apply: seqInd_vcharW. +have Zalpha: alpha \in 'Z[irr M]. + rewrite rpredB ?char_vchar ?cfInd_char ?rpred1 //. + exact: seqInd_char (H0C_S1 _ S1psi1). +have ZAalpha: alpha \in 'Z[irr M, 'A(M)] by rewrite zchar_split Zalpha. +have [Itau Ztau]: {in 'Z[irr M, 'A(M)], isometry tau, to 'Z[irr G]}. + apply: sub_iso_to (Dade_Zisometry _); last exact: zcharW. + by apply: zchar_onS; apply: FTsupp_sub0. +have oSgamma: {in S_ H0C', forall lam, '[gamma, lam] = 0}. + move=> _ /seqIndP[s /setDP[_ ker'H ] ->]. + rewrite ['Ind _]cfun_sum_constt cfdot_sumr big1 // => t sMt. + rewrite cfdotZr [gamma]cfun_sum_constt cfdot_suml big1 ?mulr0 // => t0 gMt0. + rewrite cfdotZl cfdot_irr (negPf (contraNneq _ ker'H)) ?mulr0 // => Dt0. + rewrite inE (sub_cfker_constt_Ind_irr sMt) ?gFnorm // -Dt0. + rewrite /gamma -irr0 in gMt0. + rewrite -(sub_cfker_constt_Ind_irr gMt0) ?gFnorm ?joing_subl //. + by rewrite irr0 cfker_cfun1 joing_subl. + by rewrite (subset_trans _ sHUM) // join_subG sHHU subIset ?sUHU. +(* This is step (9.11.6). *) +have [/eqP psi1qa Spsi1]: psi1 1%g == (q * a)%:R /\ psi1 \in S_ H0C'. + by move: S1psi1; rewrite mem_filter => /andP[]. +have o_alpha_S3: orthogonal alpha^\tau (map tau3 S3). + rewrite /orthogonal /= andbT all_map. + apply: contraFT (ltr_geF gtS4alpha) => /allPn[lam0 S3lam0 /= alpha_lam0]. + set ca := '[_, _] in alpha_lam0; pose al0 := (-1) ^+ (ca < 0)%R *: alpha. + have{alpha_lam0} al0_lam0: '[al0^\tau, tau3 lam0] > 0. + have Zca: ca \in Cint by rewrite Cint_cfdot_vchar ?Ztau // Ztau3 ?mem_zchar. + by rewrite linearZ cfdotZl (canLR (signrMK _) (CintEsign Zca)) normr_gt0. + rewrite -Itau // -(cfnorm_sign (ca < 0)%R) -linearZ /= -/al0. + have S4_dIirrK: {in map tau3 S4, cancel (dirr_dIirr id) (@dchi _ _)}. + apply: dirr_dIirrPE => _ /mapP[lam S4lam ->]. + rewrite mem_filter -andbA negbK in S4lam. + have [/irrP[i Dlam] _ S3lam] := and3P S4lam. + by rewrite dirrE Itau3 ?Ztau3 ?mem_zchar //= Dlam cfnorm_irr. + rewrite -(size_map tau3) -(size_map (dirr_dIirr id)). + rewrite -(card_uniqP _); last first. + rewrite !map_inj_in_uniq ?filter_uniq ?seqInd_uniq //. + apply: sub_in2 (Zisometry_inj Itau3) => lam. + by rewrite mem_filter => /andP[_ /mem_zchar->]. + exact: can_in_inj S4_dIirrK. + apply: ler_trans (_ : #|dirr_constt al0^\tau|%:R <= _); last first. + have Zal0: al0^\tau \in 'Z[irr G] by rewrite Ztau ?rpredZsign. + rewrite cnorm_dconstt // -sumr_const ler_sum // => i al0_i. + by rewrite sqr_Cint_ge1 ?gtr_eqF -?dirr_consttE // Cint_Cnat ?Cnat_dirr. + rewrite leC_nat subset_leq_card //; apply/subsetP=> _ /mapP[nu S4nu ->]. + rewrite dirr_consttE S4_dIirrK //; congr (_ > 0): al0_lam0. + rewrite {al0}linearZ !cfdotZl /=; congr (_ * _) => {ca}; apply/eqP. + have{nu S4nu} [lam S4lam ->] := mapP S4nu. + rewrite mem_filter in S4lam; have{S4lam} [_ S3lam] := andP S4lam. + have Zdlam: lam0 - lam \in 'Z[S3, M^#]. + rewrite zcharD1E rpredB ?mem_zchar //= !cfunE subr_eq0. + by have [/eqP->] := (S3qu _ S3lam, S3qu _ S3lam0). + rewrite -subr_eq0 -cfdotBr -raddfB Dtau3 //. + rewrite Itau // ?sS0A //; last exact: zchar_filter Zdlam. + suffices{lam S3lam Zdlam} oS3a: {in S3, forall lam, '[alpha, lam] = 0}. + by rewrite cfdotBr subr_eq0 !oS3a. + move=> lam; rewrite mem_filter /= -eqS12 => /andP[S1'lam H0C'lam]. + by rewrite cfdotBl oSgamma // (seqInd_ortho _ Spsi1) ?(memPn S1'lam) // subr0. +have{s4gt0 gtS4alpha} /hasP[lam1 S4lam1 _]: has predT S4 by rewrite has_predT. +have [/irrP[l1 Dl1] S3lam1]: lam1 \in irr M /\ lam1 \in S3. + by move: S4lam1; rewrite mem_filter -andbA negbK => /and3P[]. +have [S1'lam1 Slam1]: lam1 \notin S1 /\ lam1 \in S_ H0C'. + by move: S3lam1; rewrite mem_filter eqS12 => /andP[]. +have S3lam1s: lam1^*%CF \in S3 by have [[_ _ ->]] := scohS3. +have ZS3dlam1: lam1 - lam1^*%CF \in 'Z[S3, M^#]. + rewrite zcharD1E rpredB ?mem_zchar //. + by have:= seqInd_sub_aut_zchar nsHUM conjC Slam1; rewrite zcharD1 => /andP[]. +have ZAdlam1: lam1 - lam1^*%CF \in 'Z[irr M, 'A(M)]. + rewrite sS0A // zchar_split rpredB ?mem_zchar ?cfAut_seqInd //. + by rewrite (zchar_on ZS3dlam1). +pose beta := lam1 - (u %/ a)%:R *: psi1. +have ZAbeta: beta \in 'Z[irr M, 'A(M)]. + apply: sS0A; rewrite zcharD1E rpredB ?scaler_nat ?rpredMn ?mem_zchar //=. + by rewrite !cfunE subr_eq0 psi1qa -natrM mulnCA divnK // S3qu. +have [_ _ poSS _ _] := scohS0; have [_ oSS] := pairwise_orthogonalP poSS. +have o1S1: orthonormal S1. + rewrite orthonormalE filter_pairwise_orthogonal // andbT. + by apply/allP=> _ /irrS1/irrP[t ->]; rewrite /= cfnorm_irr. +have o1S4: orthonormal S4. + rewrite orthonormalE !filter_pairwise_orthogonal // andbT. + apply/allP=> nu; rewrite mem_filter /= -andbA negbK. + by case/andP=> /irrP[t ->]; rewrite cfnorm_irr. +have n1psi1: '[psi1] = 1 by have [_ -> //] := orthonormalP o1S1; rewrite eqxx. +have n1lam1: '[lam1] = 1 by have [_ -> //] := orthonormalP o1S4; rewrite eqxx. +have oS14tau: orthogonal (map tau1 S1) (map tau3 S4). + apply/orthogonalP=> psi _ S1psi /mapP[lam /sS43 S3lam ->]. + apply: {psi lam S3lam}orthogonalP S1psi (map_f tau3 S3lam). + apply: (coherent_ortho scohS0 sS10 cohS1 sS30 cohS3) => psi /=. + by rewrite mem_filter !inE eqS12 => /andP[-> _]. +have [Gamma [S4_Gamma normGamma [b Dbeta]]]: + exists Gamma, [/\ Gamma \in 'Z[map tau3 S4], '[Gamma] = 1 + & exists b : bool, beta^\tau + = Gamma - (u %/ a)%:R *: tau1 psi1 + b%:R *: \sum_(psi <- S1) tau1 psi]. +- have [G S4G [G' [Dbeta _ oG'4]]] := orthogonal_split (map tau3 S4) beta^\tau. + have [B S1B [Delta [dG' _ oD1]]] := orthogonal_split (map tau1 S1) G'. + have sZS43: {subset 'Z[S4] <= 'Z[S3]} := zchar_subset sS43. + have [Itau34 Ztau34] := sub_iso_to sZS43 sub_refl IZtau3. + have Z_G: G \in 'Z[map tau3 S4]. + have [_ -> ->] := orthonormal_span (map_orthonormal Itau34 o1S4) S4G. + rewrite big_seq rpred_sum // => xi S4xi; rewrite rpredZ_Cint ?mem_zchar //. + rewrite -(addrK G' G) -Dbeta cfdotBl (orthoPl oG'4) // subr0. + rewrite Cint_cfdot_vchar ?Ztau //. + by have{xi S4xi} [xi S4xi ->] := mapP S4xi; rewrite Ztau34 ?mem_zchar. + have oD4: orthogonal Delta (map tau3 S4). + apply/orthoPl=> xi S4xi; rewrite -(addKr B Delta) addrC -dG' cfdotBl. + by rewrite (orthoPl oG'4) // (span_orthogonal oS14tau) ?subrr // memv_span. + have [_ -> dB] := orthonormal_span (map_orthonormal Itau1 o1S1) S1B. + pose b := (u %/ a)%:R + '[B, tau1 psi1]. + have betaS1_B: {in S1, forall psi, '[beta^\tau, tau1 psi] = '[B, tau1 psi]}. + move=> psi S1psi; rewrite Dbeta dG' !cfdotDl (orthoPl oD1) ?map_f // addr0. + rewrite cfdotC (span_orthogonal oS14tau) ?rmorph0 ?add0r //. + by rewrite memv_span ?map_f. + have Zb: b \in Cint. + rewrite rpredD ?rpred_nat // -betaS1_B // Cint_cfdot_vchar ?Ztau //. + by rewrite Ztau1 ?mem_zchar. + have{dB} dB: B = - (u %/ a)%:R *: tau1 psi1 + b *: \sum_(psi <- S1) tau1 psi. + rewrite dB big_map !(big_rem _ S1psi1) /= scalerDr addrA -scalerDl addKr. + rewrite scaler_sumr; congr (_ + _); apply: eq_big_seq => psi. + rewrite mem_rem_uniq ?filter_uniq ?seqInd_uniq // => /andP[/= psi_1' S1psi]. + apply/esym/eqP; rewrite -subr_eq0 -scalerBl -addrA -!betaS1_B // -cfdotBr. + have [/eqP psi_qa Spsi]: psi 1%g == (q * a)%:R /\ psi \in S_ H0C'. + by move: S1psi; rewrite mem_filter => /andP[]. + have Z1dpsi: psi1 - psi \in 'Z[S1, M^#]. + by rewrite zcharD1E rpredB ?mem_zchar //= !cfunE psi1qa psi_qa subrr. + rewrite -raddfB Dtau1 // Itau //; last first. + by rewrite sS0A // zchar_split rpredB ?mem_zchar ?(zchar_on Z1dpsi). + rewrite cfdotC cfdotBr cfdotZr !cfdotBl 2?oSS ?(memPn S1'lam1) // subrr. + by rewrite add0r n1psi1 oSS // subr0 mulr1 rmorphN conjCK subrr scale0r. + have Gge1: 1 <= '[G] ?= iff ('[G] == 1). + rewrite eq_sym; apply: lerif_eq. + have N_G: '[G] \in Cnat. + apply: Cnat_cfnorm_vchar; apply: zchar_sub_irr Z_G => _ /mapP[nu S4nu ->]. + by rewrite Ztau34 ?mem_zchar. + rewrite -(truncCK N_G) ler1n lt0n -eqC_nat truncCK {N_G}// cfnorm_eq0. + have: '[beta^\tau, (lam1 - lam1^*%CF)^\tau] != 0. + rewrite Itau // cfdotBl cfdotZl !cfdotBr n1lam1. + rewrite (seqInd_conjC_ortho _ _ _ Slam1) ?mFT_odd // subr0. + rewrite !oSS ?cfAut_seqInd -?(inv_eq (@cfConjCK _ _)) ?(memPn S1'lam1) //. + by rewrite !(subr0, mulr0) oner_eq0. + by have [_ _ ->] := sS10. + rewrite Dbeta -Dtau3 //; apply: contraNneq => ->. + rewrite add0r raddfB cfdotBr !(orthoPl oG'4) ?map_f ?subr0 //. + rewrite mem_filter /= negbK /= S3lam1s irr_aut. + move: S4lam1; rewrite mem_filter /= negbK /= -andbA => /and3P[-> H0Clam1 _]. + by rewrite cfAut_seqInd. + have ubG: '[G] + (b ^+ 2 - b) * (u %/ a).*2%:R + '[Delta] = 1. + apply: (addrI ((u %/ a) ^ 2)%:R); transitivity '[beta^\tau]. + rewrite -!addrA addrCA Dbeta cfnormDd; last first. + by rewrite cfdotC (span_orthogonal oG'4) ?rmorph0 // memv_span ?inE. + congr (_ + _); rewrite !addrA dG' cfnormDd; last first. + by rewrite cfdotC (span_orthogonal oD1) ?rmorph0 // memv_span ?inE. + congr (_ + _); rewrite dB scaleNr [- _ + _]addrC cfnormB !cfnormZ. + rewrite normr_nat Cint_normK // scaler_sumr cfdotZr rmorph_nat. + rewrite cfnorm_map_orthonormal // cfproj_sum_orthonormal //. + rewrite Itau1 ?mem_zchar // n1psi1 mulr1 rmorphM rmorph_nat conj_Cint //. + rewrite -mulr2n oS1ua -muln_divA // mul2n -addrA addrCA -natrX mulrBl. + by congr (_ + (_ - _)); rewrite -mulrnAl -mulrnA muln2 mulrC. + rewrite Itau // cfnormBd; last first. + by rewrite cfdotZr oSS ?mulr0 // (memPnC S1'lam1). + by rewrite cfnormZ normr_nat n1psi1 n1lam1 mulr1 addrC -natrX. + have ubDelta: '[G] <= '[G] + '[Delta] ?= iff (Delta == 0). + rewrite addrC -lerif_subLR subrr -cfnorm_eq0 eq_sym. + by apply: lerif_eq; apply: cfnorm_ge0. + have{ubG} ubDeltaG: '[G] + '[Delta] <= 1 ?= iff pred2 0 1 b. + rewrite -{1}ubG addrAC [_ + _ + _] addrC -lerif_subLR subrr /=. + set n := _%:R; rewrite mulrC -{1}(mulr0 n) mono_lerif; last first. + by apply: ler_pmul2l; rewrite ltr0n double_gt0 divn_gt0 // dvdn_leq. + rewrite /= -(subr_eq0 b 1) -mulf_eq0 mulrBr mulr1 eq_sym. + apply: lerif_eq; rewrite subr_ge0. + have [-> | nz_b] := eqVneq b 0; first by rewrite expr2 mul0r. + rewrite (ler_trans (real_ler_norm _)) ?Creal_Cint // -[`|b|]mulr1. + by rewrite -Cint_normK // ler_pmul2l ?normr_gt0 // norm_Cint_ge1. + have [_ /esym] := lerif_trans Gge1 (lerif_trans ubDelta ubDeltaG). + rewrite eqxx => /and3P[/eqP normG1 /eqP Delta0 /pred2P b01]. + exists G; split=> //; exists (b != 0). + rewrite Dbeta dG' Delta0 addr0 dB scaleNr addrA; congr (_ + _ *: _). + by case: b01 => ->; rewrite ?eqxx ?oner_eq0. +(* Final step (9.11.8). *) +have alpha_beta: '[alpha^\tau, beta^\tau] = (u %/ a)%:R. + rewrite Itau // cfdotBr cfdotZr rmorph_nat !cfdotBl !oSgamma // !sub0r. + by rewrite n1psi1 mulrN opprK mulr1 addrC oSS ?subr0 // (memPn S1'lam1). +have [X S1X [Delta [Dalpha _ oD1]]]:= orthogonal_split (map tau1 S1) alpha^\tau. +pose x := 1 + '[X, tau1 psi1]. +have alphaS1_X: {in S1, forall psi, '[alpha^\tau, tau1 psi] = '[X, tau1 psi]}. + by move=> psi S1psi; rewrite Dalpha cfdotDl (orthoPl oD1) ?map_f // addr0. +have Zx: x \in Cint. + rewrite rpredD ?rpred1 // -alphaS1_X // Cint_cfdot_vchar ?Ztau //. + by rewrite Ztau1 ?mem_zchar. +have{alphaS1_X S1X} defX: X = x *: (\sum_(psi <- S1) tau1 psi) - tau1 psi1. + have [_ -> ->] := orthonormal_span (map_orthonormal Itau1 o1S1) S1X. + rewrite addrC -scaleN1r big_map !(big_rem _ S1psi1) /= scalerDr addrA. + rewrite -scalerDl addKr scaler_sumr; congr (_ + _); apply: eq_big_seq => psi. + rewrite mem_rem_uniq ?filter_uniq ?seqInd_uniq // => /andP[/= psi_1' S1psi]. + apply/esym/eqP; rewrite -subr_eq0 -scalerBl -addrA -!alphaS1_X // -cfdotBr. + have [/eqP psi_qa Spsi]: psi 1%g == (q * a)%:R /\ psi \in S_ H0C'. + by move: S1psi; rewrite mem_filter => /andP[]. + have Z1dpsi: psi1 - psi \in 'Z[S1, M^#]. + by rewrite zcharD1E rpredB ?mem_zchar //= !cfunE psi1qa psi_qa subrr. + rewrite -raddfB Dtau1 // Itau //; last first. + by rewrite sS0A // zchar_split rpredB ?mem_zchar ?(zchar_on Z1dpsi). + rewrite cfdotBr !cfdotBl !oSgamma // n1psi1 cfdotC oSS // rmorph0. + by rewrite !subr0 add0r subrr scale0r. +have{x Zx X defX Delta Dalpha oD1} b_mod_ua: (b == 0 %[mod u %/ a])%C. + rewrite -oppr0 -eqCmodN (eqCmod_trans _ (eqCmodm0 _)) // {2}nCdivE. + rewrite -alpha_beta Dbeta -addrA cfdotDr. + rewrite (span_orthogonal o_alpha_S3) ?add0r; first 1 last. + - by rewrite memv_span ?inE. + - apply: subvP (zchar_span S4_Gamma); apply: sub_span; apply: mem_subseq. + by rewrite map_subseq ?filter_subseq. + rewrite Dalpha addrC cfdotDl (span_orthogonal oD1); first 1 last. + - by rewrite memv_span ?inE. + - rewrite addrC rpredB ?rpredZ //; last by rewrite memv_span ?map_f. + by rewrite big_seq rpred_sum // => psi S1psi; rewrite memv_span ?map_f. + rewrite add0r addrC defX cfdotBr cfdotBl cfdotZl cfdotZr !scaler_sumr. + rewrite cfdotZr !rmorph_nat cfdotBl Itau1 ?mem_zchar // n1psi1. + rewrite cfnorm_map_orthonormal // cfdotC !cfproj_sum_orthonormal //. + rewrite rmorph_nat oS1ua -muln_divA // natrM !mulrA addrC mulrC addrA. + rewrite -mulNr -mulrDl eqCmod_sym eqCmod_addl_mul // addrC !rpredB ?rpred1 //. + by rewrite !rpredM ?rpred_nat. +have{b_mod_ua alpha_beta} b0: b = 0%N :> nat. + have:= b_mod_ua; rewrite /eqCmod subr0 dvdC_nat => /eqnP. + rewrite modn_small // (leq_ltn_trans (leq_b1 b)) // ltn_divRL // mul1n. + by rewrite ltn_neqAle -(eqn_pmul2l q_gt0) eq_sym ne_qa_qu dvdn_leq. +exists lam1 => //; suffices: coherent (lam1 :: lam1^* :: S1)%CF M^# tau. + by apply: subset_coherent => phi; rewrite !inE eqS12. +move: Dbeta; rewrite b0 scale0r addr0. +apply: (extend_coherent_with scohS0 sS10 cohS1); first by []. +rewrite rpred_nat psi1qa -natrM mulnCA (eqP (S3qu _ S3lam1)) divnK //. +rewrite cfdotC (span_orthogonal oS14tau) ?(zchar_span S4_Gamma) ?conjC0 //. +by rewrite rpredZ ?memv_span ?map_f. +Qed. + +End Nine. diff --git a/mathcomp/odd_order/all.v b/mathcomp/odd_order/all.v new file mode 100644 index 0000000..613acb0 --- /dev/null +++ b/mathcomp/odd_order/all.v @@ -0,0 +1,33 @@ +Require Export BGappendixAB. +Require Export BGappendixC. +Require Export BGsection10. +Require Export BGsection11. +Require Export BGsection12. +Require Export BGsection13. +Require Export BGsection14. +Require Export BGsection15. +Require Export BGsection16. +Require Export BGsection1. +Require Export BGsection2. +Require Export BGsection3. +Require Export BGsection4. +Require Export BGsection5. +Require Export BGsection6. +Require Export BGsection7. +Require Export BGsection8. +Require Export BGsection9. +Require Export PFsection10. +Require Export PFsection11. +Require Export PFsection12. +Require Export PFsection13. +Require Export PFsection14. +Require Export PFsection1. +Require Export PFsection2. +Require Export PFsection3. +Require Export PFsection4. +Require Export PFsection5. +Require Export PFsection6. +Require Export PFsection7. +Require Export PFsection8. +Require Export PFsection9. +Require Export stripped_odd_order_theorem. diff --git a/mathcomp/odd_order/stripped_odd_order_theorem.v b/mathcomp/odd_order/stripped_odd_order_theorem.v new file mode 100644 index 0000000..8123683 --- /dev/null +++ b/mathcomp/odd_order/stripped_odd_order_theorem.v @@ -0,0 +1,204 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Prelude ssreflect ssrbool ssrfun eqtype ssrnat fintype finset fingroup. +Require morphism quotient action gfunctor gproduct commutator gseries nilpotent. +Require PFsection14. + +(******************************************************************************) +(* This file contains a minimal, self-contained reformulation of the Odd *) +(* Order theorem, using only the bare Coq logic, and avoiding any use of *) +(* extra-logical features such as notation, coercion or implicit arguments. *) +(* This stripped theorem would hardly be usable, however; it just provides *) +(* evidence for the sceptics. *) +(******************************************************************************) + +(* Equivalence and equality *) + +Inductive equivalent P Q := Equivalent (P_to_Q : P -> Q) (Q_to_P : Q -> P). + +Inductive equal T (x : T) : T -> Type := Equal : equal T x x. + +(* Arithmetic *) + +Inductive natural := Zero | Add_1_to (n : natural). + +Fixpoint add (m n : natural) : natural := + match m with Zero => n | Add_1_to m_minus_1 => add m_minus_1 (Add_1_to n) end. + +Definition double (n : natural) : natural := add n n. + +Inductive odd (n : natural) := + Odd (half : natural) + (n_odd : equal natural n (Add_1_to (double half))). + +Inductive less_than (m n : natural) := + LessThan (diff : natural) + (m_lt_n : equal natural n (Add_1_to (add m diff))). + +(* Finite subsets *) + +Definition injective_in T R (D : T -> Type) (f : T -> R) := + forall x y, D x -> D y -> equal R (f x) (f y) -> equal T x y. + +Inductive in_image T R (D : T -> Type) (f : T -> R) (a : R) := + InImage (x : T) (x_in_D : D x) (a_is_fx : equal R a (f x)). + +Inductive finite_of_order T (D : T -> Type) (n : natural) := + FiniteOfOrder (rank : T -> natural) + (rank_injective : injective_in T natural D rank) + (rank_onto : + forall i, equivalent (less_than i n) (in_image T natural D rank i)). + +(* Elementary group theory *) + +Inductive group_axioms T (mul : T -> T -> T) (one : T) (inv : T -> T) := + GroupAxioms + (associativity : forall x y z, equal T (mul x (mul y z)) (mul (mul x y) z)) + (left_identity : forall x, equal T (mul one x) x) + (left_inverse : forall x, equal T (mul (inv x) x) one). + +Inductive group T mul one inv (G : T -> Type) := + Group + (G_closed_under_mul : forall x y, G x -> G y -> G (mul x y)) + (one_in_G : G one) + (G_closed_under_inv : forall x, G x -> G (inv x)). + +Inductive subgroup T mul one inv (H G : T -> Type) := + Subgroup + (H_group : group T mul one inv H) + (H_subset_G : forall x, H x -> G x). + +Inductive normal_subgroup T mul one inv (H G : T -> Type) := + NormalSubgroup + (H_subgroup_G : subgroup T mul one inv H G) + (H_is_G_invariant : forall x y, H x -> G y -> H (mul (inv y) (mul x y))). + +Inductive commute_mod T mul (x y : T) (H : T -> Type) := + CommuteMod (z : T) + (z_in_H : H z) + (xy_eq_zyx : equal T (mul x y) (mul z (mul y x))). + +Inductive abelian_factor T mul one inv (G H : T -> Type) := + AbelianFactor + (G_group : group T mul one inv G) + (H_normal_in_G : normal_subgroup T mul one inv H G) + (G_on_H_abelian : forall x y, G x -> G y -> commute_mod T mul x y H). + +Inductive solvable_group T mul one inv (G : T -> Type) := +| TrivialGroupSolvable + (G_trivial : forall x, equivalent (G x) (equal T x one)) +| AbelianExtensionSolvable (H : T -> Type) + (H_solvable : solvable_group T mul one inv H) + (G_on_H_abelian : abelian_factor T mul one inv G H). + +(* begin hide *) +Module InternalSkepticOddOrderProof. + +Local Notation Aeq := (equal _). +Local Notation Aadd := add. +Local Notation Adouble := double. +Local Notation Aodd := odd. +Local Notation Alt := less_than. +Local Notation Agroup := group. +Local Notation Asubgroup := subgroup. +Local Notation Anormal := normal_subgroup. +Local Notation Aabel_quo := abelian_factor. +Local Notation Asol := solvable_group. + +Import Prelude ssreflect ssrbool ssrfun eqtype ssrnat fintype finset fingroup. +Import morphism quotient action gfunctor gproduct commutator gseries nilpotent. +Import GroupScope. + +Lemma main T mul one inv G nn : + group_axioms T mul one inv -> Agroup T mul one inv G -> + finite_of_order T G nn -> Aodd nn -> + Asol T mul one inv G. +Proof. +pose fix natN n := if n is n1.+1 then Add_1_to (natN n1) else Zero. +pose fix Nnat mm := if mm is Add_1_to mm1 then (Nnat mm1).+1 else 0. +have natN_K: cancel natN Nnat by elim=> //= n ->. +have NnatK: cancel Nnat natN by elim=> //= mm ->. +have AaddE nn1 nn2: Nnat (Aadd nn1 nn2) = Nnat nn1 + Nnat nn2. + by elim: nn1 nn2 => //= nn1 IHnn nn2; rewrite IHnn addnS. +have AltE m n: Alt (natN m) (natN n) -> m < n. + by rewrite -{2}[n]natN_K => [[dd ->]]; rewrite /= ltnS AaddE natN_K leq_addr. +have AltI m n: m < n -> Alt (natN m) (natN n). + move/subnKC <-; exists (natN (n - m.+1)). + by rewrite -[Add_1_to _]NnatK /= AaddE !natN_K. +have AoddE n: Aodd (natN n) -> odd n. + by rewrite -{2}[n]natN_K => [[hh ->]]; rewrite /= AaddE addnn odd_double. +case=> mulA mul1T mulVT [mulG oneG invG] [rG inj_rG im_rG] odd_nn. +pose n := Nnat nn; have{odd_nn} odd_n: odd n by rewrite AoddE ?NnatK. +have{rG inj_rG im_rG} [gT o_gT [f [g Gf [fK gK]] [fM f1 fV]]]: + {gT : finGroupType & #|gT| = n & {f : gT -> T + & {g : _ & forall a, G (f a) & cancel f g /\ forall x, G x -> f (g x) = x} + & [/\ {morph f : a b / a * b >-> mul a b}, f 1 = one + & {morph f : a / a^-1 >-> inv a}]}}. +- pose gT := 'I_n.-1.+1; pose g x : gT := inord (Nnat (rG x)). + have ub_rG x: G x -> Nnat (rG x) < n. + move=> Gx; rewrite AltE ?NnatK //. + by have [_] := im_rG (rG x); apply; exists x. + have Dn: n.-1.+1 = n := ltn_predK (ub_rG one oneG). + have fP a: {x : T & G x * (g x = a)}%type. + have a_lt_n: Alt (natN a) nn by rewrite -(canLR NnatK Dn); apply: AltI. + have [/(_ a_lt_n)[x Gx rGx] _] := im_rG (natN a). + by exists x; split; rewrite // /g -rGx natN_K inord_val. + pose f a := tag (fP a); have Gf a: G (f a) by rewrite /f; case: (fP) => x []. + have fK: cancel f g by rewrite /f => a; case: (fP a) => x []. + have Ng x & G x: natN (g x) = rG x by rewrite inordK ?Dn ?ub_rG ?NnatK. + have{Ng} gK x: G x -> f (g x) = x. + by move=> Gx; rewrite (inj_rG (f (g x)) x) // -!Ng ?fK. + pose m a b := g (mul (f a) (f b)). + pose o := g one; pose v a := g (inv (f a)). + have fM: {morph f: a b / m a b >-> mul a b} by move=> a b; apply/gK/mulG. + have f1: f o = one by apply: gK. + have fV: {morph f: a / v a >-> inv a} by move=> a; apply/gK/invG. + have mA: associative m by move=> a b c; apply: canLR fK _; rewrite !fM mulA. + have m1: left_id o m by move=> a; apply: canLR fK _; rewrite f1 mul1T. + have mV: left_inverse o v m. + by move=> a; apply: canLR fK _; rewrite fV f1 mulVT. + pose bT := BaseFinGroupType _ (FinGroup.Mixin mA m1 mV). + exists (@FinGroupType bT mV); first by rewrite card_ord Dn. + by exists f; first exists g. +pose im (H : {group gT}) x := (G x * (g x \in H))%type. +have imG H : Agroup T mul one inv (im H). + split=> [x y [Gx Hx] [Gy Hy] | | x [Gx Hx]]; first 1 last. + - by split; rewrite // -(canRL fK f1). + - by split; [auto | rewrite -(gK x Gx) -fV fK groupV]. + by split; [auto | rewrite -(gK x Gx) -(gK y Gy) -fM fK groupM]. +pose G0 := [set: gT]%G. +have sGG0 x: G x -> im G0 x by split; rewrite ?inE. +have mulVV1 x: mul (inv (inv x)) one = x by rewrite -(mulVT x) mulA mulVT mul1T. +have{mulVV1} mulT1 x: mul x one = x by rewrite -[x]mulVV1 -mulA mul1T. +pose comm x y := mul (mul x y) (inv (mul y x)). +suffices solH: Asol T mul one inv (im G0). + right with (im G0) => //; split=> // [|x y Gx Gy]. + by split=> // [|x y [Gx _] Gy]; [split=> // x [] | apply: sGG0; auto]. + by exists (comm x y); [rewrite /comm; auto | rewrite -mulA mulVT -mulA mulT1]. +have solG0: solvable G0 by rewrite PFsection14.Feit_Thompson ?cardsT ?o_gT. +elim: _.+1 {-2}G0 (ltnSn #|G0|) => // m IHm H; rewrite ltnS => leHm. +have [-> | ntH] := eqVneq H 1%G. + left=> // x; split=> [[Gx /set1P] | ->]. + by rewrite -f1 => <-; rewrite gK. + by split; rewrite // -f1 fK. +have ltH'H: H^`(1) \proper H := sol_der1_proper solG0 (subsetT H) ntH. +right with (im H^`(1)%G); first exact: IHm (leq_trans (proper_card _) leHm). +have /andP[/subsetP sH'H /subsetP nH'H]: H^`(1) <| H := der_normal 1 H. +split=> // [|x y [Gx Hx] [Gy Hy]]. + split=> // [|x y [Gx H'x] [Gy Hy]]; first by split=> // x [Gx /sH'H]. + split; first by [auto]; rewrite -(gK x Gx) -(gK y Gy) -!(fM, fV) !fK. + by rewrite memJ_norm ?nH'H. +exists (comm x y); last by rewrite -mulA mulVT -mulA mulT1. +rewrite /comm; split; first by [auto]; rewrite -(gK x Gx) -(gK y Gy). +by rewrite -!(fM, fV) fK -[g x * g y]invgK !invMg -mulgA mem_commg ?groupV. +Qed. + +End InternalSkepticOddOrderProof. +(* end hide *) + +(* The Odd Order theorem *) + +Theorem stripped_Odd_Order T mul one inv (G : T -> Type) (n : natural) : + group_axioms T mul one inv -> group T mul one inv G -> + finite_of_order T G n -> odd n -> + solvable_group T mul one inv G. +Proof. exact (InternalSkepticOddOrderProof.main T mul one inv G n). Qed. |
