aboutsummaryrefslogtreecommitdiff
path: root/mathcomp/character
diff options
context:
space:
mode:
authorEnrico Tassi2015-03-09 11:07:53 +0100
committerEnrico Tassi2015-03-09 11:24:38 +0100
commitfc84c27eac260dffd8f2fb1cb56d599f1e3486d9 (patch)
treec16205f1637c80833a4c4598993c29fa0fd8c373 /mathcomp/character
Initial commit
Diffstat (limited to 'mathcomp/character')
-rw-r--r--mathcomp/character/all.v7
-rw-r--r--mathcomp/character/character.v2976
-rw-r--r--mathcomp/character/classfun.v2463
-rw-r--r--mathcomp/character/inertia.v1607
-rw-r--r--mathcomp/character/integral_char.v708
-rw-r--r--mathcomp/character/mxabelem.v1057
-rw-r--r--mathcomp/character/mxrepresentation.v5853
-rw-r--r--mathcomp/character/vcharacter.v987
8 files changed, 15658 insertions, 0 deletions
diff --git a/mathcomp/character/all.v b/mathcomp/character/all.v
new file mode 100644
index 0000000..936fa6c
--- /dev/null
+++ b/mathcomp/character/all.v
@@ -0,0 +1,7 @@
+Require Export character.
+Require Export classfun.
+Require Export inertia.
+Require Export integral_char.
+Require Export mxabelem.
+Require Export mxrepresentation.
+Require Export vcharacter.
diff --git a/mathcomp/character/character.v b/mathcomp/character/character.v
new file mode 100644
index 0000000..ac2f491
--- /dev/null
+++ b/mathcomp/character/character.v
@@ -0,0 +1,2976 @@
+(* (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 gproduct.
+Require Import fingroup morphism perm automorphism quotient finalg action.
+Require Import zmodp commutator cyclic center pgroup nilpotent sylow abelian.
+Require Import matrix mxalgebra mxpoly mxrepresentation vector ssrnum algC.
+Require Import classfun.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope GRing.Theory Num.Theory.
+Local Open Scope ring_scope.
+
+(******************************************************************************)
+(* This file contains the basic notions of character theory, based on Isaacs. *)
+(* irr G == tuple of the elements of 'CF(G) that are irreducible *)
+(* characters of G. *)
+(* Nirr G == number of irreducible characters of G. *)
+(* Iirr G == index type for the irreducible characters of G. *)
+(* := 'I_(Nirr G). *)
+(* 'chi_i == the i-th element of irr G, for i : Iirr G. *)
+(* 'chi[G]_i Note that 'chi_0 = 1, the principal character of G. *)
+(* 'Chi_i == an irreducible representation that affords 'chi_i. *)
+(* socle_of_Iirr i == the Wedderburn component of the regular representation *)
+(* of G, corresponding to 'Chi_i. *)
+(* Iirr_of_socle == the inverse of socle_of_Iirr (which is one-to-one). *)
+(* phi.[A]%CF == the image of A \in group_ring G under phi : 'CF(G). *)
+(* cfRepr rG == the character afforded by the representation rG of G. *)
+(* cfReg G == the regular character, afforded by the regular *)
+(* representation of G. *)
+(* detRepr rG == the linear character afforded by the determinant of rG. *)
+(* cfDet phi == the linear character afforded by the determinant of a *)
+(* representation affording phi. *)
+(* 'o(phi) == the "determinential order" of phi (the multiplicative *)
+(* order of cfDet phi. *)
+(* phi \is a character <=> phi : 'CF(G) is a character of G or 0. *)
+(* i \in irr_constt phi <=> 'chi_i is an irreducible constituent of phi: phi *)
+(* has a non-zero coordinate on 'chi_i over the basis irr G. *)
+(* xi \is a linear_char xi <=> xi : 'CF(G) is a linear character of G. *)
+(* 'Z(chi)%CF == the center of chi when chi is a character of G, i.e., *)
+(* rcenter rG where rG is a representation that affords phi. *)
+(* If phi is not a character then 'Z(chi)%CF = cfker phi. *)
+(* aut_Iirr u i == the index of cfAut u 'chi_i in irr G. *)
+(* conjC_Iirr i == the index of 'chi_i^*%CF in irr G. *)
+(* morph_Iirr i == the index of cfMorph 'chi[f @* G]_i in irr G. *)
+(* isom_Iirr isoG i == the index of cfIsom isoG 'chi[G]_i in irr R. *)
+(* mod_Iirr i == the index of ('chi[G / H]_i %% H)%CF in irr G. *)
+(* quo_Iirr i == the index of ('chi[G]_i / H)%CF in irr (G / H). *)
+(* Ind_Iirr G i == the index of 'Ind[G, H] 'chi_i, provided it is an *)
+(* irreducible character (such as when if H is the inertia *)
+(* group of 'chi_i). *)
+(* Res_Iirr H i == the index of 'Res[H, G] 'chi_i, provided it is an *)
+(* irreducible character (such as when 'chi_i is linear). *)
+(* sdprod_Iirr defG i == the index of cfSdprod defG 'chi_i in irr G, given *)
+(* defG : K ><| H = G. *)
+(* And, for KxK : K \x H = G. *)
+(* dprodl_Iirr KxH i == the index of cfDprodl KxH 'chi[K]_i in irr G. *)
+(* dprodr_Iirr KxH j == the index of cfDprodr KxH 'chi[H]_j in irr G. *)
+(* dprod_Iirr KxH (i, j) == the index of cfDprod KxH 'chi[K]_i 'chi[H]_j. *)
+(* inv_dprod_Iirr KxH == the inverse of dprod_Iirr KxH. *)
+(* The following are used to define and exploit the character table: *)
+(* character_table G == the character table of G, whose i-th row lists the *)
+(* values taken by 'chi_i on the conjugacy classes *)
+(* of G; this is a square Nirr G x NirrG matrix. *)
+(* irr_class i == the conjugacy class of G with index i : Iirr G. *)
+(* class_Iirr xG == the index of xG \in classes G, in Iirr G. *)
+(******************************************************************************)
+
+Local Notation algCF := [fieldType of algC].
+
+Section AlgC.
+
+Variable (gT : finGroupType).
+
+Lemma groupC : group_closure_field algCF gT.
+Proof. exact: group_closure_closed_field. Qed.
+
+End AlgC.
+
+Section Tensor.
+
+Variable (F : fieldType).
+
+Fixpoint trow (n1 : nat) :
+ forall (A : 'rV[F]_n1) m2 n2 (B : 'M[F]_(m2,n2)), 'M[F]_(m2,n1 * n2) :=
+ if n1 is n'1.+1
+ then
+ fun (A : 'M[F]_(1,(1 + n'1))) m2 n2 (B : 'M[F]_(m2,n2)) =>
+ (row_mx (lsubmx A 0 0 *: B) (trow (rsubmx A) B))
+ else (fun _ _ _ _ => 0).
+
+Lemma trow0 n1 m2 n2 B : @trow n1 0 m2 n2 B = 0.
+Proof.
+elim: n1=> //= n1 IH.
+rewrite !mxE scale0r linear0.
+rewrite IH //; apply/matrixP=> i j; rewrite !mxE.
+by case: split=> *; rewrite mxE.
+Qed.
+
+Definition trowb n1 m2 n2 B A := @trow n1 A m2 n2 B.
+
+Lemma trowbE n1 m2 n2 A B : trowb B A = @trow n1 A m2 n2 B.
+Proof. by []. Qed.
+
+Lemma trowb_is_linear n1 m2 n2 (B : 'M_(m2,n2)) : linear (@trowb n1 m2 n2 B).
+Proof.
+elim: n1=> [|n1 IH] //= k A1 A2 /=; first by rewrite scaler0 add0r.
+rewrite linearD /= linearZ.
+apply/matrixP=> i j.
+rewrite !mxE.
+case: split=> a.
+ by rewrite !mxE mulrDl mulrA.
+by rewrite linearD /= linearZ IH !mxE.
+Qed.
+
+Canonical Structure trowb_linear n1 m2 n2 B :=
+ Linear (@trowb_is_linear n1 m2 n2 B).
+
+Lemma trow_is_linear n1 m2 n2 (A : 'rV_n1) : linear (@trow n1 A m2 n2).
+Proof.
+elim: n1 A => [|n1 IH] //= A k A1 A2 /=; first by rewrite scaler0 add0r.
+rewrite linearD /= linearZ /=.
+apply/matrixP=> i j; rewrite !mxE.
+by case: split=> a; rewrite ?IH !mxE.
+Qed.
+
+Canonical Structure trow_linear n1 m2 n2 A :=
+ Linear (@trow_is_linear n1 m2 n2 A).
+
+Fixpoint tprod (m1 : nat) :
+ forall n1 (A : 'M[F]_(m1,n1)) m2 n2 (B : 'M[F]_(m2,n2)),
+ 'M[F]_(m1 * m2,n1 * n2) :=
+ if m1 is m'1.+1
+ return forall n1 (A : 'M[F]_(m1,n1)) m2 n2 (B : 'M[F]_(m2,n2)),
+ 'M[F]_(m1 * m2,n1 * n2)
+ then
+ fun n1 (A : 'M[F]_(1 + m'1,n1)) m2 n2 B =>
+ (col_mx (trow (usubmx A) B) (tprod (dsubmx A) B))
+ else (fun _ _ _ _ _ => 0).
+
+Lemma dsumx_mul m1 m2 n p A B :
+ dsubmx ((A *m B) : 'M[F]_(m1 + m2, n)) = dsubmx (A : 'M_(m1 + m2, p)) *m B.
+Proof.
+apply/matrixP=> i j; rewrite !mxE; apply: eq_bigr=> k _.
+by rewrite !mxE.
+Qed.
+
+Lemma usumx_mul m1 m2 n p A B :
+ usubmx ((A *m B) : 'M[F]_(m1 + m2, n)) = usubmx (A : 'M_(m1 + m2, p)) *m B.
+Proof.
+by apply/matrixP=> i j; rewrite !mxE; apply: eq_bigr=> k _; rewrite !mxE.
+Qed.
+
+Let trow_mul (m1 m2 n2 p2 : nat)
+ (A : 'rV_m1) (B1: 'M[F]_(m2,n2)) (B2 :'M[F]_(n2,p2)) :
+ trow A (B1 *m B2) = B1 *m trow A B2.
+Proof.
+elim: m1 A => [|m1 IH] A /=; first by rewrite mulmx0.
+by rewrite IH mul_mx_row -scalemxAr.
+Qed.
+
+Lemma tprodE m1 n1 p1 (A1 :'M[F]_(m1,n1)) (A2 :'M[F]_(n1,p1))
+ m2 n2 p2 (B1 :'M[F]_(m2,n2)) (B2 :'M[F]_(n2,p2)) :
+ tprod (A1 *m A2) (B1 *m B2) = (tprod A1 B1) *m (tprod A2 B2).
+Proof.
+elim: m1 n1 p1 A1 A2 m2 n2 p2 B1 B2 => /= [|m1 IH].
+ by move=> *; rewrite mul0mx.
+move=> n1 p1 A1 A2 m2 n2 p2 B1 B2.
+rewrite mul_col_mx -IH.
+congr col_mx; last by rewrite dsumx_mul.
+rewrite usumx_mul.
+elim: n1 {A1}(usubmx (A1: 'M_(1 + m1, n1))) p1 A2=> //= [u p1 A2|].
+ by rewrite [A2](flatmx0) !mulmx0 -trowbE linear0.
+move=> n1 IH1 A p1 A2 //=.
+set Al := lsubmx _; set Ar := rsubmx _.
+set Su := usubmx _; set Sd := dsubmx _.
+rewrite mul_row_col -IH1.
+rewrite -{1}(@hsubmxK F 1 1 n1 A).
+rewrite -{1}(@vsubmxK F 1 n1 p1 A2).
+rewrite (@mul_row_col F 1 1 n1 p1).
+rewrite -trowbE linearD /= trowbE -/Al.
+congr (_ + _).
+rewrite {1}[Al]mx11_scalar mul_scalar_mx.
+by rewrite -trowbE linearZ /= trowbE -/Su trow_mul scalemxAl.
+Qed.
+
+Let tprod_tr m1 n1 (A :'M[F]_(m1, 1 + n1)) m2 n2 (B :'M[F]_(m2, n2)) :
+ tprod A B = row_mx (trow (lsubmx A)^T B^T)^T (tprod (rsubmx A) B).
+Proof.
+elim: m1 n1 A m2 n2 B=> [|m1 IH] n1 A m2 n2 B //=.
+ by rewrite trmx0 row_mx0.
+rewrite !IH.
+pose A1 := A : 'M_(1 + m1, 1 + n1).
+have F1: dsubmx (rsubmx A1) = rsubmx (dsubmx A1).
+ by apply/matrixP=> i j; rewrite !mxE.
+have F2: rsubmx (usubmx A1) = usubmx (rsubmx A1).
+ by apply/matrixP=> i j; rewrite !mxE.
+have F3: lsubmx (dsubmx A1) = dsubmx (lsubmx A1).
+ by apply/matrixP=> i j; rewrite !mxE.
+rewrite tr_row_mx -block_mxEv -block_mxEh !(F1,F2,F3); congr block_mx.
+- by rewrite !mxE linearZ /= trmxK.
+by rewrite -trmx_dsub.
+Qed.
+
+Lemma tprod1 m n : tprod (1%:M : 'M[F]_(m,m)) (1%:M : 'M[F]_(n,n)) = 1%:M.
+Proof.
+elim: m n => [|m IH] n //=; first by rewrite [1%:M]flatmx0.
+rewrite tprod_tr.
+set u := rsubmx _; have->: u = 0.
+ apply/matrixP=> i j; rewrite !mxE.
+ by case: i; case: j=> /= j Hj; case.
+set v := lsubmx (dsubmx _); have->: v = 0.
+ apply/matrixP=> i j; rewrite !mxE.
+ by case: i; case: j; case.
+set w := rsubmx _; have->: w = 1%:M.
+ apply/matrixP=> i j; rewrite !mxE.
+ by case: i; case: j; case.
+rewrite IH -!trowbE !linear0.
+rewrite -block_mxEv.
+set z := (lsubmx _) 0 0; have->: z = 1.
+ by rewrite /z !mxE eqxx.
+by rewrite scale1r scalar_mx_block.
+Qed.
+
+Lemma mxtrace_prod m n (A :'M[F]_(m)) (B :'M[F]_(n)) :
+ \tr (tprod A B) = \tr A * \tr B.
+Proof.
+elim: m n A B => [|m IH] n A B //=.
+ by rewrite [A]flatmx0 mxtrace0 mul0r.
+rewrite tprod_tr -block_mxEv mxtrace_block IH.
+rewrite linearZ /= -mulrDl; congr (_ * _).
+rewrite -trace_mx11 .
+pose A1 := A : 'M_(1 + m).
+rewrite -{3}[A](@submxK _ 1 m 1 m A1).
+by rewrite (@mxtrace_block _ _ _ (ulsubmx A1)).
+Qed.
+
+End Tensor.
+
+(* Representation sigma type and standard representations. *)
+Section StandardRepresentation.
+
+Variables (R : fieldType) (gT : finGroupType) (G : {group gT}).
+Local Notation reprG := (mx_representation R G).
+
+Record representation :=
+ Representation {rdegree; mx_repr_of_repr :> reprG rdegree}.
+
+Lemma mx_repr0 : mx_repr G (fun _ : gT => 1%:M : 'M[R]_0).
+Proof. by split=> // g h Hg Hx; rewrite mulmx1. Qed.
+
+Definition grepr0 := Representation (MxRepresentation mx_repr0).
+
+Lemma add_mx_repr (rG1 rG2 : representation) :
+ mx_repr G (fun g => block_mx (rG1 g) 0 0 (rG2 g)).
+Proof.
+split=> [|x y Hx Hy]; first by rewrite !repr_mx1 -scalar_mx_block.
+by rewrite mulmx_block !(mulmx0, mul0mx, addr0, add0r, repr_mxM).
+Qed.
+
+Definition dadd_grepr rG1 rG2 :=
+ Representation (MxRepresentation (add_mx_repr rG1 rG2)).
+
+Section DsumRepr.
+
+Variables (n : nat) (rG : reprG n).
+
+Lemma mx_rsim_dadd (U V W : 'M_n) (rU rV : representation)
+ (modU : mxmodule rG U) (modV : mxmodule rG V) (modW : mxmodule rG W) :
+ (U + V :=: W)%MS -> mxdirect (U + V) ->
+ mx_rsim (submod_repr modU) rU -> mx_rsim (submod_repr modV) rV ->
+ mx_rsim (submod_repr modW) (dadd_grepr rU rV).
+Proof.
+case: rU; case: rV=> nV rV nU rU defW dxUV /=.
+have tiUV := mxdirect_addsP dxUV.
+move=> [fU def_nU]; rewrite -{nU}def_nU in rU fU * => inv_fU hom_fU.
+move=> [fV def_nV]; rewrite -{nV}def_nV in rV fV * => inv_fV hom_fV.
+pose pU := in_submod U (proj_mx U V) *m fU.
+pose pV := in_submod V (proj_mx V U) *m fV.
+exists (val_submod 1%:M *m row_mx pU pV) => [||g Gg].
+- by rewrite -defW (mxdirectP dxUV).
+- apply/row_freeP.
+ pose pU' := invmx fU *m val_submod 1%:M.
+ pose pV' := invmx fV *m val_submod 1%:M.
+ exists (in_submod _ (col_mx pU' pV')).
+ rewrite in_submodE mulmxA -in_submodE -mulmxA mul_row_col mulmxDr.
+ rewrite -[pU *m _]mulmxA -[pV *m _]mulmxA !mulKVmx -?row_free_unit //.
+ rewrite addrC (in_submodE V) 2![val_submod 1%:M *m _]mulmxA -in_submodE.
+ rewrite addrC (in_submodE U) 2![val_submod 1%:M *m _]mulmxA -in_submodE.
+ rewrite -!val_submodE !in_submodK ?proj_mx_sub //.
+ by rewrite add_proj_mx ?val_submodK // val_submod1 defW.
+rewrite mulmxA -val_submodE -[submod_repr _ g]mul1mx val_submodJ //.
+rewrite -(mulmxA _ (rG g)) mul_mx_row -mulmxA mul_row_block !mulmx0 addr0 add0r.
+rewrite !mul_mx_row; set W' := val_submod 1%:M; congr (row_mx _ _).
+ rewrite 3!mulmxA in_submodE mulmxA.
+ have hom_pU: (W' <= dom_hom_mx rG (proj_mx U V))%MS.
+ by rewrite val_submod1 -defW proj_mx_hom.
+ rewrite (hom_mxP hom_pU) // -in_submodE (in_submodJ modU) ?proj_mx_sub //.
+ rewrite -(mulmxA _ _ fU) hom_fU // in_submodE -2!(mulmxA W') -in_submodE.
+ by rewrite -mulmxA (mulmxA _ fU).
+rewrite 3!mulmxA in_submodE mulmxA.
+have hom_pV: (W' <= dom_hom_mx rG (proj_mx V U))%MS.
+ by rewrite val_submod1 -defW addsmxC proj_mx_hom // capmxC.
+rewrite (hom_mxP hom_pV) // -in_submodE (in_submodJ modV) ?proj_mx_sub //.
+rewrite -(mulmxA _ _ fV) hom_fV // in_submodE -2!(mulmxA W') -in_submodE.
+by rewrite -mulmxA (mulmxA _ fV).
+Qed.
+
+Lemma mx_rsim_dsum (I : finType) (P : pred I) U rU (W : 'M_n)
+ (modU : forall i, mxmodule rG (U i)) (modW : mxmodule rG W) :
+ let S := (\sum_(i | P i) U i)%MS in (S :=: W)%MS -> mxdirect S ->
+ (forall i, mx_rsim (submod_repr (modU i)) (rU i : representation)) ->
+ mx_rsim (submod_repr modW) (\big[dadd_grepr/grepr0]_(i | P i) rU i).
+Proof.
+move=> /= defW dxW rsimU.
+rewrite mxdirectE /= -!(big_filter _ P) in dxW defW *.
+elim: {P}(filter P _) => [|i e IHe] in W modW dxW defW *.
+ rewrite !big_nil /= in defW *.
+ by exists 0 => [||? _]; rewrite ?mul0mx ?mulmx0 // /row_free -defW !mxrank0.
+rewrite !big_cons /= in dxW defW *.
+rewrite 2!(big_nth i) !big_mkord /= in IHe dxW defW.
+set Wi := (\sum_i _)%MS in defW dxW IHe.
+rewrite -mxdirectE mxdirect_addsE !mxdirectE eqxx /= -/Wi in dxW.
+have modWi: mxmodule rG Wi by exact: sumsmx_module.
+case/andP: dxW; move/(IHe Wi modWi) {IHe}; move/(_ (eqmx_refl _))=> rsimWi.
+by move/eqP; move/mxdirect_addsP=> dxUiWi; exact: mx_rsim_dadd (rsimU i) rsimWi.
+Qed.
+
+Definition muln_grepr rW k := \big[dadd_grepr/grepr0]_(i < k) rW.
+
+Lemma mx_rsim_socle (sG : socleType rG) (W : sG) (rW : representation) :
+ let modW : mxmodule rG W := component_mx_module rG (socle_base W) in
+ mx_rsim (socle_repr W) rW ->
+ mx_rsim (submod_repr modW) (muln_grepr rW (socle_mult W)).
+Proof.
+set M := socle_base W => modW rsimM.
+have simM: mxsimple rG M := socle_simple W.
+have rankM_gt0: (\rank M > 0)%N by rewrite lt0n mxrank_eq0; case: simM.
+have [I /= U_I simU]: mxsemisimple rG W by exact: component_mx_semisimple.
+pose U (i : 'I_#|I|) := U_I (enum_val i).
+have reindexI := reindex _ (onW_bij I (enum_val_bij I)).
+rewrite mxdirectE /= !reindexI -mxdirectE /= => defW dxW.
+have isoU: forall i, mx_iso rG M (U i).
+ move=> i; have sUiW: (U i <= W)%MS by rewrite -defW (sumsmx_sup i).
+ exact: component_mx_iso (simU _) sUiW.
+have ->: socle_mult W = #|I|.
+ rewrite -(mulnK #|I| rankM_gt0); congr (_ %/ _)%N.
+ rewrite -defW (mxdirectP dxW) /= -sum_nat_const reindexI /=.
+ by apply: eq_bigr => i _; rewrite -(mxrank_iso (isoU i)).
+have modU: mxmodule rG (U _) := mxsimple_module (simU _).
+suff: mx_rsim (submod_repr (modU _)) rW by exact: mx_rsim_dsum defW dxW.
+by move=> i; apply: mx_rsim_trans (mx_rsim_sym _) rsimM; exact/mx_rsim_iso.
+Qed.
+
+End DsumRepr.
+
+Section ProdRepr.
+
+Variables (n1 n2 : nat) (rG1 : reprG n1) (rG2 : reprG n2).
+
+Lemma prod_mx_repr : mx_repr G (fun g => tprod (rG1 g) (rG2 g)).
+Proof.
+split=>[|i j InG JnG]; first by rewrite !repr_mx1 tprod1.
+by rewrite !repr_mxM // tprodE.
+Qed.
+
+Definition prod_repr := MxRepresentation prod_mx_repr.
+
+End ProdRepr.
+
+Lemma prod_repr_lin n2 (rG1 : reprG 1) (rG2 : reprG n2) :
+ {in G, forall x, let cast_n2 := esym (mul1n n2) in
+ prod_repr rG1 rG2 x = castmx (cast_n2, cast_n2) (rG1 x 0 0 *: rG2 x)}.
+Proof.
+move=> x Gx /=; set cast_n2 := esym _; rewrite /prod_repr /= !mxE !lshift0.
+apply/matrixP=> i j; rewrite castmxE /=.
+do 2![rewrite mxE; case: splitP => [? ? | []//]].
+by congr ((_ *: rG2 x) _ _); apply: val_inj.
+Qed.
+
+End StandardRepresentation.
+
+Implicit Arguments grepr0 [R gT G].
+Prenex Implicits grepr0 dadd_grepr.
+
+Section Char.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Fact cfRepr_subproof n (rG : mx_representation algCF G n) :
+ is_class_fun <<G>> [ffun x => \tr (rG x) *+ (x \in G)].
+Proof.
+rewrite genGid; apply: intro_class_fun => [x y Gx Gy | _ /negbTE-> //].
+by rewrite groupJr // !repr_mxM ?groupM ?groupV // mxtrace_mulC repr_mxK.
+Qed.
+Definition cfRepr n rG := Cfun 0 (@cfRepr_subproof n rG).
+
+Lemma cfRepr1 n rG : @cfRepr n rG 1%g = n%:R.
+Proof. by rewrite cfunE group1 repr_mx1 mxtrace1. Qed.
+
+Lemma cfRepr_sim n1 n2 rG1 rG2 :
+ mx_rsim rG1 rG2 -> @cfRepr n1 rG1 = @cfRepr n2 rG2.
+Proof.
+case/mx_rsim_def=> f12 [f21] fK def_rG1; apply/cfun_inP=> x Gx.
+by rewrite !cfunE def_rG1 // mxtrace_mulC mulmxA fK mul1mx.
+Qed.
+
+Lemma cfRepr0 : cfRepr grepr0 = 0.
+Proof. by apply/cfun_inP=> x Gx; rewrite !cfunE Gx mxtrace1. Qed.
+
+Lemma cfRepr_dadd rG1 rG2 :
+ cfRepr (dadd_grepr rG1 rG2) = cfRepr rG1 + cfRepr rG2.
+Proof. by apply/cfun_inP=> x Gx; rewrite !cfunE Gx mxtrace_block. Qed.
+
+Lemma cfRepr_dsum I r (P : pred I) rG :
+ cfRepr (\big[dadd_grepr/grepr0]_(i <- r | P i) rG i)
+ = \sum_(i <- r | P i) cfRepr (rG i).
+Proof. exact: (big_morph _ cfRepr_dadd cfRepr0). Qed.
+
+Lemma cfRepr_muln rG k : cfRepr (muln_grepr rG k) = cfRepr rG *+ k.
+Proof. by rewrite cfRepr_dsum /= sumr_const card_ord. Qed.
+
+Section StandardRepr.
+
+Variables (n : nat) (rG : mx_representation algCF G n).
+Let sG := DecSocleType rG.
+Let iG : irrType algCF G := DecSocleType _.
+
+Definition standard_irr (W : sG) := irr_comp iG (socle_repr W).
+
+Definition standard_socle i := pick [pred W | standard_irr W == i].
+Local Notation soc := standard_socle.
+
+Definition standard_irr_coef i := oapp (fun W => socle_mult W) 0%N (soc i).
+
+Definition standard_grepr :=
+ \big[dadd_grepr/grepr0]_i
+ muln_grepr (Representation (socle_repr i)) (standard_irr_coef i).
+
+Lemma mx_rsim_standard : mx_rsim rG standard_grepr.
+Proof.
+pose W i := oapp val 0 (soc i); pose S := (\sum_i W i)%MS.
+have C'G: [char algC]^'.-group G := algC'G G.
+have [defS dxS]: (S :=: 1%:M)%MS /\ mxdirect S.
+ rewrite /S mxdirectE /= !(bigID soc xpredT) /=.
+ rewrite addsmxC big1 => [|i]; last by rewrite /W; case (soc i).
+ rewrite adds0mx_id addnC (@big1 nat) ?add0n => [|i]; last first.
+ by rewrite /W; case: (soc i); rewrite ?mxrank0.
+ have <-: Socle sG = 1%:M := reducible_Socle1 sG (mx_Maschke rG C'G).
+ have [W0 _ | noW] := pickP sG; last first.
+ suff no_i: (soc : pred iG) =1 xpred0 by rewrite /Socle !big_pred0 ?mxrank0.
+ by move=> i; rewrite /soc; case: pickP => // W0; have:= noW W0.
+ have irrK Wi: soc (standard_irr Wi) = Some Wi.
+ rewrite /soc; case: pickP => [W' | /(_ Wi)] /= /eqP // eqWi.
+ apply/eqP/socle_rsimP.
+ apply: mx_rsim_trans (rsim_irr_comp iG C'G (socle_irr _)) (mx_rsim_sym _).
+ by rewrite [irr_comp _ _]eqWi; exact: rsim_irr_comp (socle_irr _).
+ have bij_irr: {on [pred i | soc i], bijective standard_irr}.
+ exists (odflt W0 \o soc) => [Wi _ | i]; first by rewrite /= irrK.
+ by rewrite inE /soc /=; case: pickP => //= Wi; move/eqP.
+ rewrite !(reindex standard_irr) {bij_irr}//=.
+ have all_soc Wi: soc (standard_irr Wi) by rewrite irrK.
+ rewrite (eq_bigr val) => [|Wi _]; last by rewrite /W irrK.
+ rewrite !(eq_bigl _ _ all_soc); split=> //.
+ rewrite (eq_bigr (mxrank \o val)) => [|Wi _]; last by rewrite /W irrK.
+ by rewrite -mxdirectE /= Socle_direct.
+pose modW i : mxmodule rG (W i) :=
+ if soc i is Some Wi as oWi return mxmodule rG (oapp val 0 oWi) then
+ component_mx_module rG (socle_base Wi)
+ else mxmodule0 rG n.
+apply: mx_rsim_trans (mx_rsim_sym (rsim_submod1 (mxmodule1 rG) _)) _ => //.
+apply: mx_rsim_dsum (modW) _ defS dxS _ => i.
+rewrite /W /standard_irr_coef /modW /soc; case: pickP => [Wi|_] /=; last first.
+ rewrite /muln_grepr big_ord0.
+ by exists 0 => [||x _]; rewrite ?mxrank0 ?mulmx0 ?mul0mx.
+by move/eqP=> <-; apply: mx_rsim_socle; exact: rsim_irr_comp (socle_irr Wi).
+Qed.
+
+End StandardRepr.
+
+Definition cfReg (B : {set gT}) : 'CF(B) := #|B|%:R *: '1_[1].
+
+Lemma cfRegE x : @cfReg G x = #|G|%:R *+ (x == 1%g).
+Proof. by rewrite cfunE cfuniE ?normal1 // inE mulr_natr. Qed.
+
+(* This is Isaacs, Lemma (2.10). *)
+Lemma cfReprReg : cfRepr (regular_repr algCF G) = cfReg G.
+Proof.
+apply/cfun_inP=> x Gx; rewrite cfRegE.
+have [-> | ntx] := altP (x =P 1%g); first by rewrite cfRepr1.
+rewrite cfunE Gx [\tr _]big1 // => i _; rewrite 2!mxE /=.
+rewrite -(inj_eq enum_val_inj) gring_indexK ?groupM ?enum_valP //.
+by rewrite eq_mulVg1 mulKg (negbTE ntx).
+Qed.
+
+Definition xcfun (chi : 'CF(G)) A :=
+ (gring_row A *m (\col_(i < #|G|) chi (enum_val i))) 0 0.
+
+Lemma xcfun_is_additive phi : additive (xcfun phi).
+Proof. by move=> A B; rewrite /xcfun linearB mulmxBl !mxE. Qed.
+Canonical xcfun_additive phi := Additive (xcfun_is_additive phi).
+
+Lemma xcfunZr a phi A : xcfun phi (a *: A) = a * xcfun phi A.
+Proof. by rewrite /xcfun linearZ -scalemxAl mxE. Qed.
+
+(* In order to add a second canonical structure on xcfun *)
+Definition xcfun_r_head k A phi := let: tt := k in xcfun phi A.
+Local Notation xcfun_r A := (xcfun_r_head tt A).
+
+Lemma xcfun_rE A chi : xcfun_r A chi = xcfun chi A. Proof. by []. Qed.
+
+Fact xcfun_r_is_additive A : additive (xcfun_r A).
+Proof.
+move=> phi psi; rewrite /= /xcfun !mxE -sumrB; apply: eq_bigr => i _.
+by rewrite !mxE !cfunE mulrBr.
+Qed.
+Canonical xcfun_r_additive A := Additive (xcfun_r_is_additive A).
+
+Lemma xcfunZl a phi A : xcfun (a *: phi) A = a * xcfun phi A.
+Proof.
+rewrite /xcfun !mxE big_distrr; apply: eq_bigr => i _ /=.
+by rewrite !mxE cfunE mulrCA.
+Qed.
+
+Lemma xcfun_repr n rG A : xcfun (@cfRepr n rG) A = \tr (gring_op rG A).
+Proof.
+rewrite gring_opE [gring_row A]row_sum_delta !linear_sum /xcfun !mxE.
+apply: eq_bigr => i _; rewrite !mxE /= !linearZ cfunE enum_valP /=.
+by congr (_ * \tr _) => {A} /=; rewrite /gring_mx /= -rowE rowK mxvecK.
+Qed.
+
+End Char.
+Notation xcfun_r A := (xcfun_r_head tt A).
+Notation "phi .[ A ]" := (xcfun phi A) : cfun_scope.
+
+Definition pred_Nirr gT B := #|@classes gT B|.-1.
+Arguments Scope pred_Nirr [_ group_scope].
+Notation Nirr G := (pred_Nirr G).+1.
+Notation Iirr G := 'I_(Nirr G).
+
+Section IrrClassDef.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Let sG := DecSocleType (regular_repr algCF G).
+
+Lemma NirrE : Nirr G = #|classes G|.
+Proof. by rewrite /pred_Nirr (cardD1 [1]) classes1. Qed.
+
+Fact Iirr_cast : Nirr G = #|sG|.
+Proof. by rewrite NirrE ?card_irr ?algC'G //; exact: groupC. Qed.
+
+Let offset := cast_ord (esym Iirr_cast) (enum_rank [1 sG]%irr).
+
+Definition socle_of_Iirr (i : Iirr G) : sG :=
+ enum_val (cast_ord Iirr_cast (i + offset)).
+Definition irr_of_socle (Wi : sG) : Iirr G :=
+ cast_ord (esym Iirr_cast) (enum_rank Wi) - offset.
+Local Notation W := socle_of_Iirr.
+
+Lemma socle_Iirr0 : W 0 = [1 sG]%irr.
+Proof. by rewrite /W add0r cast_ordKV enum_rankK. Qed.
+
+Lemma socle_of_IirrK : cancel W irr_of_socle.
+Proof. by move=> i; rewrite /irr_of_socle enum_valK cast_ordK addrK. Qed.
+
+Lemma irr_of_socleK : cancel irr_of_socle W.
+Proof. by move=> Wi; rewrite /W subrK cast_ordKV enum_rankK. Qed.
+Hint Resolve socle_of_IirrK irr_of_socleK.
+
+Lemma irr_of_socle_bij (A : pred (Iirr G)) : {on A, bijective irr_of_socle}.
+Proof. by apply: onW_bij; exists W. Qed.
+
+Lemma socle_of_Iirr_bij (A : pred sG) : {on A, bijective W}.
+Proof. by apply: onW_bij; exists irr_of_socle. Qed.
+
+End IrrClassDef.
+
+Prenex Implicits socle_of_IirrK irr_of_socleK.
+Arguments Scope socle_of_Iirr [_ ring_scope].
+
+Notation "''Chi_' i" := (irr_repr (socle_of_Iirr i))
+ (at level 8, i at level 2, format "''Chi_' i").
+
+Fact irr_key : unit. Proof. by []. Qed.
+Definition irr_def gT B : (Nirr B).-tuple 'CF(B) :=
+ let irr_of i := 'Res[B, <<B>>] (@cfRepr gT _ _ 'Chi_(inord i)) in
+ [tuple of mkseq irr_of (Nirr B)].
+Definition irr := locked_with irr_key irr_def.
+
+Arguments Scope irr [_ group_scope].
+
+Notation "''chi_' i" := (tnth (irr _) i%R)
+ (at level 8, i at level 2, format "''chi_' i") : ring_scope.
+Notation "''chi[' G ]_ i" := (tnth (irr G) i%R)
+ (at level 8, i at level 2, only parsing) : ring_scope.
+
+Section IrrClass.
+
+Variable (gT : finGroupType) (G : {group gT}).
+Implicit Types (i : Iirr G) (B : {set gT}).
+Open Scope group_ring_scope.
+
+Lemma congr_irr i1 i2 : i1 = i2 -> 'chi_i1 = 'chi_i2. Proof. by move->. Qed.
+
+Lemma Iirr1_neq0 : G :!=: 1%g -> inord 1 != 0 :> Iirr G.
+Proof. by rewrite -classes_gt1 -NirrE -val_eqE /= => /inordK->. Qed.
+
+Lemma has_nonprincipal_irr : G :!=: 1%g -> {i : Iirr G | i != 0}.
+Proof. by move/Iirr1_neq0; exists (inord 1). Qed.
+
+Lemma irrRepr i : cfRepr 'Chi_i = 'chi_i.
+Proof.
+rewrite [irr]unlock (tnth_nth 0) nth_mkseq // -[<<G>>]/(gval _) genGidG.
+by rewrite cfRes_id inord_val.
+Qed.
+
+Lemma irr0 : 'chi[G]_0 = 1.
+Proof.
+apply/cfun_inP=> x Gx; rewrite -irrRepr cfun1E cfunE Gx.
+by rewrite socle_Iirr0 irr1_repr // mxtrace1 degree_irr1.
+Qed.
+
+Lemma cfun1_irr : 1 \in irr G.
+Proof. by rewrite -irr0 mem_tnth. Qed.
+
+Lemma mem_irr i : 'chi_i \in irr G.
+Proof. exact: mem_tnth. Qed.
+
+Lemma irrP xi : reflect (exists i, xi = 'chi_i) (xi \in irr G).
+Proof.
+apply: (iffP idP) => [/(nthP 0)[i] | [i ->]]; last exact: mem_irr.
+rewrite size_tuple => lt_i_G <-.
+by exists (Ordinal lt_i_G); rewrite (tnth_nth 0).
+Qed.
+
+Let sG := DecSocleType (regular_repr algCF G).
+Let C'G := algC'G G.
+Let closG := @groupC _ G.
+Local Notation W i := (@socle_of_Iirr _ G i).
+Local Notation "''n_' i" := 'n_(W i).
+Local Notation "''R_' i" := 'R_(W i).
+Local Notation "''e_' i" := 'e_(W i).
+
+Lemma irr1_degree i : 'chi_i 1%g = ('n_i)%:R.
+Proof. by rewrite -irrRepr cfRepr1. Qed.
+
+Lemma Cnat_irr1 i : 'chi_i 1%g \in Cnat.
+Proof. by rewrite irr1_degree rpred_nat. Qed.
+
+Lemma irr1_gt0 i : 0 < 'chi_i 1%g.
+Proof. by rewrite irr1_degree ltr0n irr_degree_gt0. Qed.
+
+Lemma irr1_neq0 i : 'chi_i 1%g != 0.
+Proof. by rewrite eqr_le ltr_geF ?irr1_gt0. Qed.
+
+Lemma irr_neq0 i : 'chi_i != 0.
+Proof. by apply: contraNneq (irr1_neq0 i) => ->; rewrite cfunE. Qed.
+
+Definition cfIirr B (chi : 'CF(B)) : Iirr B := inord (index chi (irr B)).
+
+Lemma cfIirrE chi : chi \in irr G -> 'chi_(cfIirr chi) = chi.
+Proof.
+move=> chi_irr; rewrite (tnth_nth 0) inordK ?nth_index //.
+by rewrite -index_mem size_tuple in chi_irr.
+Qed.
+
+Lemma cfIirrPE J (f : J -> 'CF(G)) (P : pred J) :
+ (forall j, P j -> f j \in irr G) ->
+ forall j, P j -> 'chi_(cfIirr (f j)) = f j.
+Proof. by move=> irr_f j /irr_f; apply: cfIirrE. Qed.
+
+(* This is Isaacs, Corollary (2.7). *)
+Corollary irr_sum_square : \sum_i ('chi[G]_i 1%g) ^+ 2 = #|G|%:R.
+Proof.
+rewrite -(sum_irr_degree sG) // natr_sum (reindex _ (socle_of_Iirr_bij _)) /=.
+by apply: eq_bigr => i _; rewrite irr1_degree natrX.
+Qed.
+
+(* This is Isaacs, Lemma (2.11). *)
+Lemma cfReg_sum : cfReg G = \sum_i 'chi_i 1%g *: 'chi_i.
+Proof.
+apply/cfun_inP=> x Gx; rewrite -cfReprReg cfunE Gx (mxtrace_regular sG) //=.
+rewrite sum_cfunE (reindex _ (socle_of_Iirr_bij _)); apply: eq_bigr => i _.
+by rewrite -irrRepr cfRepr1 !cfunE Gx mulr_natl.
+Qed.
+
+Let aG := regular_repr algCF G.
+Let R_G := group_ring algCF G.
+
+Lemma xcfun_annihilate i j A : i != j -> (A \in 'R_j)%MS -> ('chi_i).[A]%CF = 0.
+Proof.
+move=> neq_ij RjA; rewrite -irrRepr xcfun_repr.
+by rewrite (irr_repr'_op0 _ _ RjA) ?raddf0 // eq_sym (can_eq socle_of_IirrK).
+Qed.
+
+Lemma xcfunG phi x : x \in G -> phi.[aG x]%CF = phi x.
+Proof.
+by move=> Gx; rewrite /xcfun /gring_row rowK -rowE !mxE !(gring_indexK, mul1g).
+Qed.
+
+Lemma xcfun_mul_id i A :
+ (A \in R_G)%MS -> ('chi_i).['e_i *m A]%CF = ('chi_i).[A]%CF.
+Proof.
+move=> RG_A; rewrite -irrRepr !xcfun_repr gring_opM //.
+by rewrite op_Wedderburn_id ?mul1mx.
+Qed.
+
+Lemma xcfun_id i j : ('chi_i).['e_j]%CF = 'chi_i 1%g *+ (i == j).
+Proof.
+have [<-{j} | /xcfun_annihilate->//] := altP eqP; last exact: Wedderburn_id_mem.
+by rewrite -xcfunG // repr_mx1 -(xcfun_mul_id _ (envelop_mx1 _)) mulmx1.
+Qed.
+
+Lemma irr_free : free (irr G).
+Proof.
+apply/freeP=> s s0 i; apply: (mulIf (irr1_neq0 i)).
+rewrite mul0r -(raddf0 (xcfun_r_additive 'e_i)) -{}s0 raddf_sum /=.
+rewrite (bigD1 i) //= -tnth_nth xcfunZl xcfun_id eqxx big1 ?addr0 // => j ne_ji.
+by rewrite -tnth_nth xcfunZl xcfun_id (negbTE ne_ji) mulr0.
+Qed.
+
+Lemma irr_inj : injective (tnth (irr G)).
+Proof. by apply/injectiveP/free_uniq; rewrite map_tnth_enum irr_free. Qed.
+
+Lemma irrK : cancel (tnth (irr G)) (@cfIirr G).
+Proof. by move=> i; apply: irr_inj; rewrite cfIirrE ?mem_irr. Qed.
+
+Lemma irr_eq1 i : ('chi_i == 1) = (i == 0).
+Proof. by rewrite -irr0 (inj_eq irr_inj). Qed.
+
+Lemma cforder_irr_eq1 i : (#['chi_i]%CF == 1%N) = (i == 0).
+Proof. by rewrite -dvdn1 dvdn_cforder irr_eq1. Qed.
+
+Lemma irr_basis : basis_of 'CF(G)%VS (irr G).
+Proof.
+rewrite /basis_of irr_free andbT -dimv_leqif_eq ?subvf //.
+by rewrite dim_cfun (eqnP irr_free) size_tuple NirrE.
+Qed.
+
+Lemma eq_sum_nth_irr a : \sum_i a i *: 'chi[G]_i = \sum_i a i *: (irr G)`_i.
+Proof. by apply: eq_bigr => i; rewrite -tnth_nth. Qed.
+
+(* This is Isaacs, Theorem (2.8). *)
+Theorem cfun_irr_sum phi : {a | phi = \sum_i a i *: 'chi[G]_i}.
+Proof.
+rewrite (coord_basis irr_basis (memvf phi)) -eq_sum_nth_irr.
+by exists ((coord (irr G))^~ phi).
+Qed.
+
+Lemma cfRepr_standard n (rG : mx_representation algCF G n) :
+ cfRepr (standard_grepr rG)
+ = \sum_i (standard_irr_coef rG (W i))%:R *: 'chi_i.
+Proof.
+rewrite cfRepr_dsum (reindex _ (socle_of_Iirr_bij _)).
+by apply: eq_bigr => i _; rewrite scaler_nat cfRepr_muln irrRepr.
+Qed.
+
+Lemma cfRepr_inj n1 n2 rG1 rG2 :
+ @cfRepr _ G n1 rG1 = @cfRepr _ G n2 rG2 -> mx_rsim rG1 rG2.
+Proof.
+move=> eq_repr12; pose c i : algC := (standard_irr_coef _ (W i))%:R.
+have [rsim1 rsim2] := (mx_rsim_standard rG1, mx_rsim_standard rG2).
+apply: mx_rsim_trans (rsim1) (mx_rsim_sym _).
+suffices ->: standard_grepr rG1 = standard_grepr rG2 by [].
+apply: eq_bigr => Wi _; congr (muln_grepr _ _); apply/eqP; rewrite -eqC_nat.
+rewrite -[Wi]irr_of_socleK -!/(c _ _ _) -!(coord_sum_free (c _ _) _ irr_free).
+rewrite -!eq_sum_nth_irr -!cfRepr_standard.
+by rewrite -(cfRepr_sim rsim1) -(cfRepr_sim rsim2) eq_repr12.
+Qed.
+
+Lemma cfRepr_rsimP n1 n2 rG1 rG2 :
+ reflect (mx_rsim rG1 rG2) (@cfRepr _ G n1 rG1 == @cfRepr _ G n2 rG2).
+Proof. by apply: (iffP eqP) => [/cfRepr_inj | /cfRepr_sim]. Qed.
+
+Lemma irr_reprP xi :
+ reflect (exists2 rG : representation _ G, mx_irreducible rG & xi = cfRepr rG)
+ (xi \in irr G).
+Proof.
+apply: (iffP (irrP xi)) => [[i ->] | [[n rG] irr_rG ->]].
+ by exists (Representation 'Chi_i); [exact: socle_irr | rewrite irrRepr].
+exists (irr_of_socle (irr_comp sG rG)); rewrite -irrRepr irr_of_socleK /=.
+exact/cfRepr_sim/rsim_irr_comp.
+Qed.
+
+(* This is Isaacs, Theorem (2.12). *)
+Lemma Wedderburn_id_expansion i :
+ 'e_i = #|G|%:R^-1 *: \sum_(x in G) 'chi_i 1%g * 'chi_i x^-1%g *: aG x.
+Proof.
+have Rei: ('e_i \in 'R_i)%MS by exact: Wedderburn_id_mem.
+have /envelop_mxP[a def_e]: ('e_i \in R_G)%MS; last rewrite -/aG in def_e.
+ by move: Rei; rewrite genmxE mem_sub_gring => /andP[].
+apply: canRL (scalerK (neq0CG _)) _; rewrite def_e linear_sum /=.
+apply: eq_bigr => x Gx; have Gx' := groupVr Gx; rewrite scalerA; congr (_ *: _).
+transitivity (cfReg G).['e_i *m aG x^-1%g]%CF.
+ rewrite def_e mulmx_suml raddf_sum (bigD1 x) //= -scalemxAl xcfunZr.
+ rewrite -repr_mxM // mulgV xcfunG // cfRegE eqxx mulrC big1 ?addr0 //.
+ move=> y /andP[Gy /negbTE neq_xy]; rewrite -scalemxAl xcfunZr -repr_mxM //.
+ by rewrite xcfunG ?groupM // cfRegE -eq_mulgV1 neq_xy mulr0.
+rewrite cfReg_sum -xcfun_rE raddf_sum /= (bigD1 i) //= xcfunZl.
+rewrite xcfun_mul_id ?envelop_mx_id ?xcfunG ?groupV ?big1 ?addr0 // => j ne_ji.
+rewrite xcfunZl (xcfun_annihilate ne_ji) ?mulr0 //.
+have /andP[_ /(submx_trans _)-> //] := Wedderburn_ideal (W i).
+by rewrite mem_mulsmx // envelop_mx_id ?groupV.
+Qed.
+
+End IrrClass.
+
+Arguments Scope cfReg [_ group_scope].
+Prenex Implicits cfIirr.
+Implicit Arguments irr_inj [[gT] [G] x1 x2].
+
+Section IsChar.
+
+Variable gT : finGroupType.
+
+Definition character {G : {set gT}} :=
+ [qualify a phi : 'CF(G) | [forall i, coord (irr G) i phi \in Cnat]].
+Fact character_key G : pred_key (@character G). Proof. by []. Qed.
+Canonical character_keyed G := KeyedQualifier (character_key G).
+
+Variable G : {group gT}.
+Implicit Types (phi chi xi : 'CF(G)) (i : Iirr G).
+
+Lemma irr_char i : 'chi_i \is a character.
+Proof.
+by apply/forallP=> j; rewrite (tnth_nth 0) coord_free ?irr_free ?isNatC_nat.
+Qed.
+
+Lemma cfun1_char : (1 : 'CF(G)) \is a character.
+Proof. by rewrite -irr0 irr_char. Qed.
+
+Lemma cfun0_char : (0 : 'CF(G)) \is a character.
+Proof. by apply/forallP=> i; rewrite linear0 rpred0. Qed.
+
+Fact add_char : addr_closed (@character G).
+Proof.
+split=> [|chi xi /forallP-Nchi /forallP-Nxi]; first exact: cfun0_char.
+by apply/forallP=> i; rewrite linearD rpredD /=.
+Qed.
+Canonical character_addrPred := AddrPred add_char.
+
+Lemma char_sum_irrP {phi} :
+ reflect (exists n, phi = \sum_i (n i)%:R *: 'chi_i) (phi \is a character).
+Proof.
+apply: (iffP idP)=> [/forallP-Nphi | [n ->]]; last first.
+ by apply: rpred_sum => i _; rewrite scaler_nat rpredMn // irr_char.
+do [have [a ->] := cfun_irr_sum phi] in Nphi *; exists (truncC \o a).
+apply: eq_bigr => i _; congr (_ *: _); have:= eqP (Nphi i).
+by rewrite eq_sum_nth_irr coord_sum_free ?irr_free.
+Qed.
+
+Lemma char_sum_irr chi :
+ chi \is a character -> {r | chi = \sum_(i <- r) 'chi_i}.
+Proof.
+move=> Nchi; apply: sig_eqW; case/char_sum_irrP: Nchi => n {chi}->.
+elim/big_rec: _ => [|i _ _ [r ->]]; first by exists nil; rewrite big_nil.
+exists (ncons (n i) i r); rewrite scaler_nat.
+by elim: {n}(n i) => [|n IHn]; rewrite ?add0r //= big_cons mulrS -addrA IHn.
+Qed.
+
+Lemma Cnat_char1 chi : chi \is a character -> chi 1%g \in Cnat.
+Proof.
+case/char_sum_irr=> r ->{chi}.
+by elim/big_rec: _ => [|i chi _ Nchi1]; rewrite cfunE ?rpredD // Cnat_irr1.
+Qed.
+
+Lemma char1_ge0 chi : chi \is a character -> 0 <= chi 1%g.
+Proof. by move/Cnat_char1/Cnat_ge0. Qed.
+
+Lemma char1_eq0 chi : chi \is a character -> (chi 1%g == 0) = (chi == 0).
+Proof.
+case/char_sum_irr=> r ->; apply/idP/idP=> [|/eqP->]; last by rewrite cfunE.
+case: r => [|i r]; rewrite ?big_nil // sum_cfunE big_cons.
+rewrite paddr_eq0 ?sumr_ge0 => // [||j _]; rewrite 1?ltrW ?irr1_gt0 //.
+by rewrite (negbTE (irr1_neq0 i)).
+Qed.
+
+Lemma char1_gt0 chi : chi \is a character -> (0 < chi 1%g) = (chi != 0).
+Proof. by move=> Nchi; rewrite -char1_eq0 // Cnat_gt0 ?Cnat_char1. Qed.
+
+Lemma char_reprP phi :
+ reflect (exists rG : representation algCF G, phi = cfRepr rG)
+ (phi \is a character).
+Proof.
+apply: (iffP char_sum_irrP) => [[n ->] | [[n rG] ->]]; last first.
+ exists (fun i => standard_irr_coef rG (socle_of_Iirr i)).
+ by rewrite -cfRepr_standard (cfRepr_sim (mx_rsim_standard rG)).
+exists (\big[dadd_grepr/grepr0]_i muln_grepr (Representation 'Chi_i) (n i)).
+rewrite cfRepr_dsum; apply: eq_bigr => i _.
+by rewrite cfRepr_muln irrRepr scaler_nat.
+Qed.
+
+Local Notation reprG := (mx_representation algCF G).
+
+Lemma cfRepr_char n (rG : reprG n) : cfRepr rG \is a character.
+Proof. by apply/char_reprP; exists (Representation rG). Qed.
+
+Lemma cfReg_char : cfReg G \is a character.
+Proof. by rewrite -cfReprReg cfRepr_char. Qed.
+
+Lemma cfRepr_prod n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) :
+ cfRepr rG1 * cfRepr rG2 = cfRepr (prod_repr rG1 rG2).
+Proof. by apply/cfun_inP=> x Gx; rewrite !cfunE /= Gx mxtrace_prod. Qed.
+
+Lemma mul_char : mulr_closed (@character G).
+Proof.
+split=> [|_ _ /char_reprP[rG1 ->] /char_reprP[rG2 ->]]; first exact: cfun1_char.
+apply/char_reprP; exists (Representation (prod_repr rG1 rG2)).
+by rewrite cfRepr_prod.
+Qed.
+Canonical char_mulrPred := MulrPred mul_char.
+Canonical char_semiringPred := SemiringPred mul_char.
+
+End IsChar.
+Prenex Implicits character.
+Implicit Arguments char_reprP [gT G phi].
+
+Section AutChar.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Implicit Type u : {rmorphism algC -> algC}.
+
+Lemma cfRepr_map u n (rG : mx_representation algCF G n) :
+ cfRepr (map_repr u rG) = cfAut u (cfRepr rG).
+Proof. by apply/cfun_inP=> x Gx; rewrite !cfunE Gx map_reprE trace_map_mx. Qed.
+
+Lemma cfAut_char u (chi : 'CF(G)) :
+ chi \is a character -> cfAut u chi \is a character.
+Proof.
+case/char_reprP=> rG ->; apply/char_reprP.
+by exists (Representation (map_repr u rG)); rewrite cfRepr_map.
+Qed.
+
+Lemma cfConjC_char (chi : 'CF(G)) :
+ chi \is a character -> chi^*%CF \is a character.
+Proof. exact: cfAut_char. Qed.
+
+Lemma cfAut_char1 u (chi : 'CF(G)) :
+ chi \is a character -> cfAut u chi 1%g = chi 1%g.
+Proof. by move/Cnat_char1=> Nchi1; rewrite cfunE aut_Cnat. Qed.
+
+Lemma cfAut_irr1 u i : (cfAut u 'chi[G]_i) 1%g = 'chi_i 1%g.
+Proof. exact: cfAut_char1 (irr_char i). Qed.
+
+Lemma cfConjC_char1 (chi : 'CF(G)) :
+ chi \is a character -> chi^*%CF 1%g = chi 1%g.
+Proof. exact: cfAut_char1. Qed.
+
+Lemma cfConjC_irr1 u i : ('chi[G]_i)^*%CF 1%g = 'chi_i 1%g.
+Proof. exact: cfAut_irr1. Qed.
+
+End AutChar.
+
+Section Linear.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Definition linear_char {B : {set gT}} :=
+ [qualify a phi : 'CF(B) | (phi \is a character) && (phi 1%g == 1)].
+
+Section OneChar.
+
+Variable xi : 'CF(G).
+Hypothesis CFxi : xi \is a linear_char.
+
+Lemma lin_char1: xi 1%g = 1.
+Proof. by case/andP: CFxi => _ /eqP. Qed.
+
+Lemma lin_charW : xi \is a character.
+Proof. by case/andP: CFxi. Qed.
+
+Lemma cfun1_lin_char : (1 : 'CF(G)) \is a linear_char.
+Proof. by rewrite qualifE cfun1_char /= cfun11. Qed.
+
+Lemma lin_charM : {in G &, {morph xi : x y / (x * y)%g >-> x * y}}.
+Proof.
+move=> x y Gx Gy; case/andP: CFxi => /char_reprP[[n rG] -> /=].
+rewrite cfRepr1 pnatr_eq1 => /eqP n1; rewrite {n}n1 in rG *.
+rewrite !cfunE Gx Gy groupM //= !mulr1n repr_mxM //.
+by rewrite [rG x]mx11_scalar [rG y]mx11_scalar -scalar_mxM !mxtrace_scalar.
+Qed.
+
+Lemma lin_char_prod I r (P : pred I) (x : I -> gT) :
+ (forall i, P i -> x i \in G) ->
+ xi (\prod_(i <- r | P i) x i)%g = \prod_(i <- r | P i) xi (x i).
+Proof.
+move=> Gx; elim/(big_load (fun y => y \in G)): _.
+elim/big_rec2: _ => [|i a y Pi [Gy <-]]; first by rewrite lin_char1.
+by rewrite groupM ?lin_charM ?Gx.
+Qed.
+
+Let xiMV x : x \in G -> xi x * xi (x^-1)%g = 1.
+Proof. by move=> Gx; rewrite -lin_charM ?groupV // mulgV lin_char1. Qed.
+
+Lemma lin_char_neq0 x : x \in G -> xi x != 0.
+Proof.
+by move/xiMV/(congr1 (predC1 0)); rewrite /= oner_eq0 mulf_eq0 => /norP[].
+Qed.
+
+Lemma lin_charV x : x \in G -> xi x^-1%g = (xi x)^-1.
+Proof. by move=> Gx; rewrite -[_^-1]mulr1 -(xiMV Gx) mulKf ?lin_char_neq0. Qed.
+
+Lemma lin_charX x n : x \in G -> xi (x ^+ n)%g = xi x ^+ n.
+Proof.
+move=> Gx; elim: n => [|n IHn]; first exact: lin_char1.
+by rewrite expgS exprS lin_charM ?groupX ?IHn.
+Qed.
+
+Lemma lin_char_unity_root x : x \in G -> xi x ^+ #[x] = 1.
+Proof. by move=> Gx; rewrite -lin_charX // expg_order lin_char1. Qed.
+
+Lemma normC_lin_char x : x \in G -> `|xi x| = 1.
+Proof.
+move=> Gx; apply/eqP; rewrite -(@pexpr_eq1 _ _ #[x]) ?normr_ge0 //.
+by rewrite -normrX // lin_char_unity_root ?normr1.
+Qed.
+
+Lemma lin_charV_conj x : x \in G -> xi x^-1%g = (xi x)^*.
+Proof.
+move=> Gx; rewrite lin_charV // invC_norm mulrC normC_lin_char //.
+by rewrite expr1n divr1.
+Qed.
+
+Lemma lin_char_irr : xi \in irr G.
+Proof.
+case/andP: CFxi => /char_reprP[rG ->]; rewrite cfRepr1 pnatr_eq1 => /eqP n1.
+by apply/irr_reprP; exists rG => //; exact/mx_abs_irrW/linear_mx_abs_irr.
+Qed.
+
+Lemma mul_conjC_lin_char : xi * xi^*%CF = 1.
+Proof.
+apply/cfun_inP=> x Gx.
+by rewrite !cfunE cfun1E Gx -normCK normC_lin_char ?expr1n.
+Qed.
+
+Lemma lin_char_unitr : xi \in GRing.unit.
+Proof. by apply/unitrPr; exists xi^*%CF; apply: mul_conjC_lin_char. Qed.
+
+Lemma invr_lin_char : xi^-1 = xi^*%CF.
+Proof. by rewrite -[_^-1]mulr1 -mul_conjC_lin_char mulKr ?lin_char_unitr. Qed.
+
+Lemma cfAut_lin_char u : cfAut u xi \is a linear_char.
+Proof. by rewrite qualifE cfunE lin_char1 rmorph1 cfAut_char ?lin_charW /=. Qed.
+
+Lemma cfConjC_lin_char : xi^*%CF \is a linear_char.
+Proof. exact: cfAut_lin_char. Qed.
+
+Lemma fful_lin_char_inj : cfaithful xi -> {in G &, injective xi}.
+Proof.
+move=> fful_phi x y Gx Gy xi_xy; apply/eqP; rewrite eq_mulgV1 -in_set1.
+rewrite (subsetP fful_phi) // inE groupM ?groupV //=; apply/forallP=> z.
+have [Gz | G'z] := boolP (z \in G); last by rewrite !cfun0 ?groupMl ?groupV.
+by rewrite -mulgA lin_charM ?xi_xy -?lin_charM ?groupM ?groupV // mulKVg.
+Qed.
+
+End OneChar.
+
+Lemma card_Iirr_abelian : abelian G -> #|Iirr G| = #|G|.
+Proof. by rewrite card_ord NirrE card_classes_abelian => /eqP. Qed.
+
+Lemma card_Iirr_cyclic : cyclic G -> #|Iirr G| = #|G|.
+Proof. by move/cyclic_abelian/card_Iirr_abelian. Qed.
+
+Lemma char_abelianP :
+ reflect (forall i : Iirr G, 'chi_i \is a linear_char) (abelian G).
+Proof.
+apply: (iffP idP) => [cGG i | CF_G].
+ rewrite qualifE irr_char /= irr1_degree.
+ by rewrite irr_degree_abelian //; last exact: groupC.
+rewrite card_classes_abelian -NirrE -eqC_nat -irr_sum_square //.
+rewrite -{1}[Nirr G]card_ord -sumr_const; apply/eqP/eq_bigr=> i _.
+by rewrite lin_char1 ?expr1n ?CF_G.
+Qed.
+
+Lemma irr_repr_lin_char (i : Iirr G) x :
+ x \in G -> 'chi_i \is a linear_char ->
+ irr_repr (socle_of_Iirr i) x = ('chi_i x)%:M.
+Proof.
+move=> Gx CFi; rewrite -irrRepr cfunE Gx.
+move: (_ x); rewrite -[irr_degree _]natCK -irr1_degree lin_char1 //.
+by rewrite (natCK 1) => A; rewrite trace_mx11 -mx11_scalar.
+Qed.
+
+Fact linear_char_key B : pred_key (@linear_char B). Proof. by []. Qed.
+Canonical linear_char_keted B := KeyedQualifier (linear_char_key B).
+Fact linear_char_divr : divr_closed (@linear_char G).
+Proof.
+split=> [|chi xi Lchi Lxi]; first exact: cfun1_lin_char.
+rewrite invr_lin_char // qualifE cfunE.
+by rewrite rpredM ?lin_char1 ?mulr1 ?lin_charW //= cfConjC_lin_char.
+Qed.
+Canonical lin_char_mulrPred := MulrPred linear_char_divr.
+Canonical lin_char_divrPred := DivrPred linear_char_divr.
+
+Lemma irr_cyclic_lin i : cyclic G -> 'chi[G]_i \is a linear_char.
+Proof. by move/cyclic_abelian/char_abelianP. Qed.
+
+Lemma irr_prime_lin i : prime #|G| -> 'chi[G]_i \is a linear_char.
+Proof. by move/prime_cyclic/irr_cyclic_lin. Qed.
+
+End Linear.
+
+Prenex Implicits linear_char.
+
+Section Restrict.
+
+Variable (gT : finGroupType) (G H : {group gT}).
+
+Lemma cfRepr_sub n (rG : mx_representation algCF G n) (sHG : H \subset G) :
+ cfRepr (subg_repr rG sHG) = 'Res[H] (cfRepr rG).
+Proof.
+by apply/cfun_inP => x Hx; rewrite cfResE // !cfunE Hx (subsetP sHG).
+Qed.
+
+Lemma cfRes_char chi : chi \is a character -> 'Res[H, G] chi \is a character.
+Proof.
+have [sHG | not_sHG] := boolP (H \subset G).
+ by case/char_reprP=> rG ->; rewrite -(cfRepr_sub rG sHG) cfRepr_char.
+by move/Cnat_char1=> Nchi1; rewrite cfResEout // rpredZ_Cnat ?rpred1.
+Qed.
+
+Lemma cfRes_eq0 phi : phi \is a character -> ('Res[H, G] phi == 0) = (phi == 0).
+Proof. by move=> Nchi; rewrite -!char1_eq0 ?cfRes_char // cfRes1. Qed.
+
+Lemma cfRes_lin_char chi :
+ chi \is a linear_char -> 'Res[H, G] chi \is a linear_char.
+Proof. by case/andP=> Nchi; rewrite qualifE cfRes_char ?cfRes1. Qed.
+
+Lemma Res_irr_neq0 i : 'Res[H, G] 'chi_i != 0.
+Proof. by rewrite cfRes_eq0 ?irr_neq0 ?irr_char. Qed.
+
+Lemma cfRes_lin_lin (chi : 'CF(G)) :
+ chi \is a character -> 'Res[H] chi \is a linear_char -> chi \is a linear_char.
+Proof. by rewrite !qualifE cfRes1 => -> /andP[]. Qed.
+
+Lemma cfRes_irr_irr chi :
+ chi \is a character -> 'Res[H] chi \in irr H -> chi \in irr G.
+Proof.
+have [sHG /char_reprP[rG ->] | not_sHG Nchi] := boolP (H \subset G).
+ rewrite -(cfRepr_sub _ sHG) => /irr_reprP[rH irrH def_rH]; apply/irr_reprP.
+ suffices /subg_mx_irr: mx_irreducible (subg_repr rG sHG) by exists rG.
+ by apply: mx_rsim_irr irrH; exact/cfRepr_rsimP/eqP.
+rewrite cfResEout // => /irrP[j Dchi_j]; apply/lin_char_irr/cfRes_lin_lin=> //.
+suffices j0: j = 0 by rewrite cfResEout // Dchi_j j0 irr0 rpred1.
+apply: contraNeq (irr1_neq0 j) => nz_j.
+have:= xcfun_id j 0; rewrite -Dchi_j cfunE xcfunZl -irr0 xcfun_id eqxx => ->.
+by rewrite (negPf nz_j).
+Qed.
+
+Definition Res_Iirr (A B : {set gT}) i := cfIirr ('Res[B, A] 'chi_i).
+
+Lemma Res_Iirr0 : Res_Iirr H (0 : Iirr G) = 0.
+Proof. by rewrite /Res_Iirr irr0 rmorph1 -irr0 irrK. Qed.
+
+Lemma lin_Res_IirrE i : 'chi[G]_i 1%g = 1 -> 'chi_(Res_Iirr H i) = 'Res 'chi_i.
+Proof.
+move=> chi1; rewrite cfIirrE ?lin_char_irr ?cfRes_lin_char //.
+by rewrite qualifE irr_char /= chi1.
+Qed.
+
+End Restrict.
+
+Arguments Scope Res_Iirr [_ group_scope group_scope ring_scope].
+
+Section Morphim.
+
+Variables (aT rT : finGroupType) (G D : {group aT}) (f : {morphism D >-> rT}).
+Implicit Type chi : 'CF(f @* G).
+
+Lemma cfRepr_morphim n (rfG : mx_representation algCF (f @* G) n) sGD :
+ cfRepr (morphim_repr rfG sGD) = cfMorph (cfRepr rfG).
+Proof.
+apply/cfun_inP=> x Gx; have Dx: x \in D := subsetP sGD x Gx.
+by rewrite cfMorphE // !cfunE ?mem_morphim ?Gx.
+Qed.
+
+Lemma cfMorph_char chi : chi \is a character -> cfMorph chi \is a character.
+Proof.
+have [sGD /char_reprP[rG ->] | outGD Nchi] := boolP (G \subset D); last first.
+ by rewrite cfMorphEout // rpredZ_Cnat ?rpred1 ?Cnat_char1.
+apply/char_reprP; exists (Representation (morphim_repr rG sGD)).
+by rewrite cfRepr_morphim.
+Qed.
+
+Lemma cfMorph_lin_char chi :
+ chi \is a linear_char -> cfMorph chi \is a linear_char.
+Proof. by case/andP=> Nchi; rewrite qualifE cfMorph_char ?cfMorph1. Qed.
+
+Lemma cfMorph_irr chi :
+ G \subset D -> chi \in irr (f @* G) -> cfMorph chi \in irr G.
+Proof.
+move=> sGD /irr_reprP[rG irrG ->]; apply/irr_reprP.
+exists (Representation (morphim_repr rG sGD)); first exact/morphim_mx_irr.
+apply/cfun_inP=> x Gx; rewrite !cfunElock /= sGD Gx.
+by rewrite mem_morphim ?(subsetP sGD).
+Qed.
+
+Definition morph_Iirr i := cfIirr (cfMorph 'chi[f @* G]_i).
+
+Lemma morph_Iirr0 : morph_Iirr 0 = 0.
+Proof. by rewrite /morph_Iirr irr0 rmorph1 -irr0 irrK. Qed.
+
+Hypothesis sGD : G \subset D.
+
+Lemma morph_IirrE i : 'chi_(morph_Iirr i) = cfMorph 'chi_i.
+Proof. by rewrite cfIirrE ?cfMorph_irr ?mem_irr. Qed.
+
+Lemma morph_Iirr_inj : injective morph_Iirr.
+Proof.
+by move=> i j eq_ij; apply/irr_inj/cfMorph_inj; rewrite // -!morph_IirrE eq_ij.
+Qed.
+
+Lemma morph_Iirr_eq0 i : (morph_Iirr i == 0) = (i == 0).
+Proof. by rewrite -!irr_eq1 morph_IirrE cfMorph_eq1. Qed.
+
+End Morphim.
+
+Section Isom.
+
+Variables (aT rT : finGroupType) (G : {group aT}) (f : {morphism G >-> rT}).
+Variables (R : {group rT}) (isoGR : isom G R f).
+Implicit Type chi : 'CF(G).
+
+Lemma cfIsom_char chi : chi \is a character -> cfIsom isoGR chi \is a character.
+Proof.
+by move=> Nchi; rewrite [cfIsom _]locked_withE cfMorph_char ?cfRes_char.
+Qed.
+
+Lemma cfIsom_lin_char chi :
+ chi \is a linear_char -> cfIsom isoGR chi \is a linear_char.
+Proof. by case/andP=> Nchi; rewrite qualifE cfIsom_char ?cfIsom1. Qed.
+
+Lemma cfIsom_irr chi : chi \in irr G -> cfIsom isoGR chi \in irr R.
+Proof.
+move=> irr_chi; rewrite [cfIsom _]locked_withE cfMorph_irr //.
+by rewrite (isom_im (isom_sym isoGR)) cfRes_id.
+Qed.
+
+Definition isom_Iirr i := cfIirr (cfIsom isoGR 'chi_i).
+
+Lemma isom_IirrE i : 'chi_(isom_Iirr i) = cfIsom isoGR 'chi_i.
+Proof. by rewrite cfIirrE ?cfIsom_irr ?mem_irr. Qed.
+
+Lemma isom_Iirr_inj : injective isom_Iirr.
+Proof.
+by move=> i j eqij; apply/irr_inj/(cfIsom_inj isoGR); rewrite -!isom_IirrE eqij.
+Qed.
+
+Lemma isom_Iirr_eq0 i : (isom_Iirr i == 0) = (i == 0).
+Proof. by rewrite -!irr_eq1 isom_IirrE cfIsom_eq1. Qed.
+
+Lemma isom_Iirr0 : isom_Iirr 0 = 0.
+Proof. by apply/eqP; rewrite isom_Iirr_eq0. Qed.
+
+End Isom.
+
+Implicit Arguments isom_Iirr_inj [aT rT G f R x1 x2].
+
+Section IsomInv.
+
+Variables (aT rT : finGroupType) (G : {group aT}) (f : {morphism G >-> rT}).
+Variables (R : {group rT}) (isoGR : isom G R f).
+
+Lemma isom_IirrK : cancel (isom_Iirr isoGR) (isom_Iirr (isom_sym isoGR)).
+Proof. by move=> i; apply: irr_inj; rewrite !isom_IirrE cfIsomK. Qed.
+
+Lemma isom_IirrKV : cancel (isom_Iirr (isom_sym isoGR)) (isom_Iirr isoGR).
+Proof. by move=> i; apply: irr_inj; rewrite !isom_IirrE cfIsomKV. Qed.
+
+End IsomInv.
+
+Section OrthogonalityRelations.
+
+Variables aT gT : finGroupType.
+
+(* This is Isaacs, Lemma (2.15) *)
+Lemma repr_rsim_diag (G : {group gT}) f (rG : mx_representation algCF G f) x :
+ x \in G -> let chi := cfRepr rG in
+ exists e,
+ [/\ (*a*) exists2 B, B \in unitmx & rG x = invmx B *m diag_mx e *m B,
+ (*b*) (forall i, e 0 i ^+ #[x] = 1) /\ (forall i, `|e 0 i| = 1),
+ (*c*) chi x = \sum_i e 0 i /\ `|chi x| <= chi 1%g
+ & (*d*) chi x^-1%g = (chi x)^*].
+Proof.
+move=> Gx; without loss cGG: G rG Gx / abelian G.
+ have sXG: <[x]> \subset G by rewrite cycle_subG.
+ move/(_ _ (subg_repr rG sXG) (cycle_id x) (cycle_abelian x)).
+ by rewrite /= !cfunE !groupV Gx (cycle_id x) !group1.
+have [I U W simU W1 dxW]: mxsemisimple rG 1%:M.
+ rewrite -(reducible_Socle1 (DecSocleType rG) (mx_Maschke _ (algC'G G))).
+ exact: Socle_semisimple.
+have linU i: \rank (U i) = 1%N.
+ by apply: mxsimple_abelian_linear cGG (simU i); exact: groupC.
+have castI: f = #|I|.
+ by rewrite -(mxrank1 algCF f) -W1 (eqnP dxW) /= -sum1_card; exact/eq_bigr.
+pose B := \matrix_j nz_row (U (enum_val (cast_ord castI j))).
+have rowU i: (nz_row (U i) :=: U i)%MS.
+ apply/eqmxP; rewrite -(geq_leqif (mxrank_leqif_eq (nz_row_sub _))) linU.
+ by rewrite lt0n mxrank_eq0 (nz_row_mxsimple (simU i)).
+have unitB: B \in unitmx.
+ rewrite -row_full_unit -sub1mx -W1; apply/sumsmx_subP=> i _.
+ pose j := cast_ord (esym castI) (enum_rank i).
+ by rewrite (submx_trans _ (row_sub j B)) // rowK cast_ordKV enum_rankK rowU.
+pose e := \row_j row j (B *m rG x *m invmx B) 0 j.
+have rGx: rG x = invmx B *m diag_mx e *m B.
+ rewrite -mulmxA; apply: canRL (mulKmx unitB) _.
+ apply/row_matrixP=> j; rewrite 2!row_mul; set u := row j B.
+ have /sub_rVP[a def_ux]: (u *m rG x <= u)%MS.
+ rewrite /u rowK rowU (eqmxMr _ (rowU _)).
+ exact: (mxmoduleP (mxsimple_module (simU _))).
+ rewrite def_ux [u]rowE scalemxAl; congr (_ *m _).
+ apply/rowP=> k; rewrite 5!mxE !row_mul def_ux [u]rowE scalemxAl mulmxK //.
+ by rewrite !mxE !eqxx !mulr_natr eq_sym.
+have exp_e j: e 0 j ^+ #[x] = 1.
+ suffices: (diag_mx e j j) ^+ #[x] = (B *m rG (x ^+ #[x])%g *m invmx B) j j.
+ by rewrite expg_order repr_mx1 mulmx1 mulmxV // [e]lock !mxE eqxx.
+ elim: #[x] => [|n IHn]; first by rewrite repr_mx1 mulmx1 mulmxV // !mxE eqxx.
+ rewrite expgS repr_mxM ?groupX // {1}rGx -!mulmxA mulKVmx //.
+ by rewrite mul_diag_mx mulmxA [M in _ = M]mxE -IHn exprS {1}mxE eqxx.
+have norm1_e j: `|e 0 j| = 1.
+ apply/eqP; rewrite -(@pexpr_eq1 _ _ #[x]) ?normr_ge0 //.
+ by rewrite -normrX exp_e normr1.
+exists e; split=> //; first by exists B.
+ rewrite cfRepr1 !cfunE Gx rGx mxtrace_mulC mulKVmx // mxtrace_diag.
+ split=> //=; apply: (ler_trans (ler_norm_sum _ _ _)).
+ by rewrite (eq_bigr _ (in1W norm1_e)) sumr_const card_ord lerr.
+rewrite !cfunE groupV !mulrb Gx rGx mxtrace_mulC mulKVmx //.
+rewrite -trace_map_mx map_diag_mx; set d' := diag_mx _.
+rewrite -[d'](mulKVmx unitB) mxtrace_mulC -[_ *m _](repr_mxK rG Gx) rGx.
+rewrite -!mulmxA mulKVmx // (mulmxA d').
+suffices->: d' *m diag_mx e = 1%:M by rewrite mul1mx mulKmx.
+rewrite mulmx_diag -diag_const_mx; congr diag_mx; apply/rowP=> j.
+by rewrite [e]lock !mxE mulrC -normCK -lock norm1_e expr1n.
+Qed.
+
+Variables (A : {group aT}) (G : {group gT}).
+
+(* This is Isaacs, Lemma (2.15) (d). *)
+Lemma char_inv (chi : 'CF(G)) x : chi \is a character -> chi x^-1%g = (chi x)^*.
+Proof.
+case Gx: (x \in G); last by rewrite !cfun0 ?rmorph0 ?groupV ?Gx.
+by case/char_reprP=> rG ->; have [e [_ _ _]] := repr_rsim_diag rG Gx.
+Qed.
+
+Lemma irr_inv i x : 'chi[G]_i x^-1%g = ('chi_i x)^*.
+Proof. exact/char_inv/irr_char. Qed.
+
+(* This is Isaacs, Theorem (2.13). *)
+Theorem generalized_orthogonality_relation y (i j : Iirr G) :
+ #|G|%:R^-1 * (\sum_(x in G) 'chi_i (x * y)%g * 'chi_j x^-1%g)
+ = (i == j)%:R * ('chi_i y / 'chi_i 1%g).
+Proof.
+pose W := @socle_of_Iirr _ G; pose e k := Wedderburn_id (W k).
+pose aG := regular_repr algCF G.
+have [Gy | notGy] := boolP (y \in G); last first.
+ rewrite cfun0 // mul0r big1 ?mulr0 // => x Gx.
+ by rewrite cfun0 ?groupMl ?mul0r.
+transitivity (('chi_i).[e j *m aG y]%CF / 'chi_j 1%g).
+ rewrite [e j]Wedderburn_id_expansion -scalemxAl xcfunZr -mulrA; congr (_ * _).
+ rewrite mulmx_suml raddf_sum big_distrl; apply: eq_bigr => x Gx /=.
+ rewrite -scalemxAl xcfunZr -repr_mxM // xcfunG ?groupM // mulrAC mulrC.
+ by congr (_ * _); rewrite mulrC mulKf ?irr1_neq0.
+rewrite mulr_natl mulrb; have [<-{j} | neq_ij] := altP eqP.
+ by congr (_ / _); rewrite xcfun_mul_id ?envelop_mx_id ?xcfunG.
+rewrite (xcfun_annihilate neq_ij) ?mul0r //.
+case/andP: (Wedderburn_ideal (W j)) => _; apply: submx_trans.
+by rewrite mem_mulsmx ?Wedderburn_id_mem ?envelop_mx_id.
+Qed.
+
+(* This is Isaacs, Corollary (2.14). *)
+Corollary first_orthogonality_relation (i j : Iirr G) :
+ #|G|%:R^-1 * (\sum_(x in G) 'chi_i x * 'chi_j x^-1%g) = (i == j)%:R.
+Proof.
+have:= generalized_orthogonality_relation 1 i j.
+rewrite mulrA mulfK ?irr1_neq0 // => <-; congr (_ * _).
+by apply: eq_bigr => x; rewrite mulg1.
+Qed.
+
+(* The character table. *)
+
+Definition irr_class i := enum_val (cast_ord (NirrE G) i).
+Definition class_Iirr xG :=
+ cast_ord (esym (NirrE G)) (enum_rank_in (classes1 G) xG).
+
+Local Notation c := irr_class.
+Local Notation g i := (repr (c i)).
+Local Notation iC := class_Iirr.
+
+Definition character_table := \matrix_(i, j) 'chi[G]_i (g j).
+Local Notation X := character_table.
+
+Lemma irr_classP i : c i \in classes G.
+Proof. exact: enum_valP. Qed.
+
+Lemma repr_irr_classK i : g i ^: G = c i.
+Proof. by case/repr_classesP: (irr_classP i). Qed.
+
+Lemma irr_classK : cancel c iC.
+Proof. by move=> i; rewrite /iC enum_valK_in cast_ordK. Qed.
+
+Lemma class_IirrK : {in classes G, cancel iC c}.
+Proof. by move=> xG GxG; rewrite /c cast_ordKV enum_rankK_in. Qed.
+
+Lemma reindex_irr_class R idx (op : @Monoid.com_law R idx) F :
+ \big[op/idx]_(xG in classes G) F xG = \big[op/idx]_i F (c i).
+Proof.
+rewrite (reindex c); first by apply: eq_bigl => i; exact: enum_valP.
+by exists iC; [apply: in1W; exact: irr_classK | exact: class_IirrK].
+Qed.
+
+(* The explicit value of the inverse is needed for the proof of the second *)
+(* orthogonality relation. *)
+Let X' := \matrix_(i, j) (#|'C_G[g i]|%:R^-1 * ('chi[G]_j (g i))^*).
+Let XX'_1: X *m X' = 1%:M.
+Proof.
+apply/matrixP=> i j; rewrite !mxE -first_orthogonality_relation mulr_sumr.
+rewrite sum_by_classes => [|u v Gu Gv]; last by rewrite -conjVg !cfunJ.
+rewrite reindex_irr_class /=; apply/esym/eq_bigr=> k _.
+rewrite !mxE irr_inv // -/(g k) -divg_index -indexgI /=.
+rewrite (char0_natf_div Cchar) ?dvdn_indexg // index_cent1 invfM invrK.
+by rewrite repr_irr_classK mulrCA mulrA mulrCA.
+Qed.
+
+Lemma character_table_unit : X \in unitmx.
+Proof. by case/mulmx1_unit: XX'_1. Qed.
+Let uX := character_table_unit.
+
+(* This is Isaacs, Theorem (2.18). *)
+Theorem second_orthogonality_relation x y :
+ y \in G ->
+ \sum_i 'chi[G]_i x * ('chi_i y)^* = #|'C_G[x]|%:R *+ (x \in y ^: G).
+Proof.
+move=> Gy; pose i_x := iC (x ^: G); pose i_y := iC (y ^: G).
+have [Gx | notGx] := boolP (x \in G); last first.
+ rewrite (contraNF (subsetP _ x) notGx) ?class_subG ?big1 // => i _.
+ by rewrite cfun0 ?mul0r.
+transitivity ((#|'C_G[repr (y ^: G)]|%:R *: (X' *m X)) i_y i_x).
+ rewrite scalemxAl !mxE; apply: eq_bigr => k _; rewrite !mxE mulrC -!mulrA.
+ by rewrite !class_IirrK ?mem_classes // !cfun_repr mulVKf ?neq0CG.
+rewrite mulmx1C // !mxE -!divg_index !(index_cent1, =^~ indexgI).
+rewrite (class_transr (mem_repr y _)) ?class_refl // mulr_natr.
+rewrite (can_in_eq class_IirrK) ?mem_classes //.
+have [-> | not_yGx] := altP eqP; first by rewrite class_refl.
+by rewrite [x \in _](contraNF _ not_yGx) // => /class_transr->.
+Qed.
+
+Lemma eq_irr_mem_classP x y :
+ y \in G -> reflect (forall i, 'chi[G]_i x = 'chi_i y) (x \in y ^: G).
+Proof.
+move=> Gy; apply: (iffP idP) => [/imsetP[z Gz ->] i | xGy]; first exact: cfunJ.
+have Gx: x \in G.
+ congr is_true: Gy; apply/eqP; rewrite -(can_eq oddb) -eqC_nat -!cfun1E.
+ by rewrite -irr0 xGy.
+congr is_true: (class_refl G x); apply/eqP; rewrite -(can_eq oddb).
+rewrite -(eqn_pmul2l (cardG_gt0 'C_G[x])) -eqC_nat !mulrnA; apply/eqP.
+by rewrite -!second_orthogonality_relation //; apply/eq_bigr=> i _; rewrite xGy.
+Qed.
+
+(* This is Isaacs, Theorem (6.32) (due to Brauer). *)
+Lemma card_afix_irr_classes (ito : action A (Iirr G)) (cto : action A _) a :
+ a \in A -> [acts A, on classes G | cto] ->
+ (forall i x y, x \in G -> y \in cto (x ^: G) a ->
+ 'chi_i x = 'chi_(ito i a) y) ->
+ #|'Fix_ito[a]| = #|'Fix_(classes G | cto)[a]|.
+Proof.
+move=> Aa actsAG stabAchi; apply/eqP; rewrite -eqC_nat; apply/eqP.
+have [[cP cK] iCK] := (irr_classP, irr_classK, class_IirrK).
+pose icto b i := iC (cto (c i) b).
+have Gca i: cto (c i) a \in classes G by rewrite (acts_act actsAG).
+have inj_qa: injective (icto a).
+ by apply: can_inj (icto a^-1%g) _ => i; rewrite /icto iCK ?actKin ?cK.
+pose Pa : 'M[algC]_(Nirr G) := perm_mx (actperm ito a).
+pose qa := perm inj_qa; pose Qa : 'M[algC]_(Nirr G) := perm_mx qa^-1^-1%g.
+transitivity (\tr Pa).
+ rewrite -sumr_const big_mkcond; apply: eq_bigr => i _.
+ by rewrite !mxE permE inE sub1set inE; case: ifP.
+symmetry; transitivity (\tr Qa).
+ rewrite cardsE -sumr_const -big_filter_cond big_mkcond big_filter /=.
+ rewrite reindex_irr_class; apply: eq_bigr => i _; rewrite !mxE invgK permE.
+ by rewrite inE sub1set inE -(can_eq cK) iCK //; case: ifP.
+rewrite -[Pa](mulmxK uX) -[Qa](mulKmx uX) mxtrace_mulC; congr (\tr(_ *m _)).
+rewrite -row_permE -col_permE; apply/matrixP=> i j; rewrite !mxE.
+rewrite -{2}[j](permKV qa); move: {j}(_ j) => j; rewrite !permE iCK //.
+apply: stabAchi; first by case/repr_classesP: (cP j).
+by rewrite repr_irr_classK (mem_repr_classes (Gca _)).
+Qed.
+
+End OrthogonalityRelations.
+
+Arguments Scope character_table [_ group_scope].
+
+Section InnerProduct.
+
+Variable (gT : finGroupType) (G : {group gT}).
+
+Lemma cfdot_irr i j : '['chi_i, 'chi_j]_G = (i == j)%:R.
+Proof.
+rewrite -first_orthogonality_relation; congr (_ * _).
+by apply: eq_bigr => x Gx; rewrite irr_inv.
+Qed.
+
+Lemma cfnorm_irr i : '['chi[G]_i] = 1.
+Proof. by rewrite cfdot_irr eqxx. Qed.
+
+Lemma irr_orthonormal : orthonormal (irr G).
+Proof.
+apply/orthonormalP; split; first exact: free_uniq (irr_free G).
+move=> _ _ /irrP[i ->] /irrP[j ->].
+by rewrite cfdot_irr (inj_eq (@irr_inj _ G)).
+Qed.
+
+Lemma coord_cfdot phi i : coord (irr G) i phi = '[phi, 'chi_i].
+Proof.
+rewrite {2}(coord_basis (irr_basis G) (memvf phi)).
+rewrite cfdot_suml (bigD1 i) // cfdotZl /= -tnth_nth cfdot_irr eqxx mulr1.
+rewrite big1 ?addr0 // => j neq_ji; rewrite cfdotZl /= -tnth_nth cfdot_irr.
+by rewrite (negbTE neq_ji) mulr0.
+Qed.
+
+Lemma cfun_sum_cfdot phi : phi = \sum_i '[phi, 'chi_i]_G *: 'chi_i.
+Proof.
+rewrite {1}(coord_basis (irr_basis G) (memvf phi)).
+by apply: eq_bigr => i _; rewrite coord_cfdot -tnth_nth.
+Qed.
+
+Lemma cfdot_sum_irr phi psi :
+ '[phi, psi]_G = \sum_i '[phi, 'chi_i] * '[psi, 'chi_i]^*.
+Proof.
+rewrite {1}[phi]cfun_sum_cfdot cfdot_suml; apply: eq_bigr => i _.
+by rewrite cfdotZl -cfdotC.
+Qed.
+
+Lemma Cnat_cfdot_char_irr i phi :
+ phi \is a character -> '[phi, 'chi_i]_G \in Cnat.
+Proof. by move/forallP/(_ i); rewrite coord_cfdot. Qed.
+
+Lemma cfdot_char_r phi chi :
+ chi \is a character -> '[phi, chi]_G = \sum_i '[phi, 'chi_i] * '[chi, 'chi_i].
+Proof.
+move=> Nchi; rewrite cfdot_sum_irr; apply: eq_bigr => i _; congr (_ * _).
+by rewrite conj_Cnat ?Cnat_cfdot_char_irr.
+Qed.
+
+Lemma Cnat_cfdot_char chi xi :
+ chi \is a character -> xi \is a character -> '[chi, xi]_G \in Cnat.
+Proof.
+move=> Nchi Nxi; rewrite cfdot_char_r ?rpred_sum // => i _.
+by rewrite rpredM ?Cnat_cfdot_char_irr.
+Qed.
+
+Lemma cfdotC_char chi xi :
+ chi \is a character-> xi \is a character -> '[chi, xi]_G = '[xi, chi].
+Proof. by move=> Nchi Nxi; rewrite cfdotC conj_Cnat ?Cnat_cfdot_char. Qed.
+
+Lemma irrEchar chi : (chi \in irr G) = (chi \is a character) && ('[chi] == 1).
+Proof.
+apply/irrP/andP=> [[i ->] | [Nchi]]; first by rewrite irr_char cfnorm_irr.
+rewrite cfdot_sum_irr => /eqP/Cnat_sum_eq1[i _| i [_ ci1 cj0]].
+ by rewrite rpredM // ?conj_Cnat ?Cnat_cfdot_char_irr.
+exists i; rewrite [chi]cfun_sum_cfdot (bigD1 i) //=.
+rewrite -(@normr_idP _ _ (@Cnat_ge0 _ (Cnat_cfdot_char_irr i Nchi))).
+rewrite normC_def {}ci1 sqrtC1 scale1r big1 ?addr0 // => j neq_ji.
+by rewrite (('[_] =P 0) _) ?scale0r // -normr_eq0 normC_def cj0 ?sqrtC0.
+Qed.
+
+Lemma irrWchar chi : chi \in irr G -> chi \is a character.
+Proof. by rewrite irrEchar => /andP[]. Qed.
+
+Lemma irrWnorm chi : chi \in irr G -> '[chi] = 1.
+Proof. by rewrite irrEchar => /andP[_ /eqP]. Qed.
+
+Lemma mul_lin_irr xi chi :
+ xi \is a linear_char -> chi \in irr G -> xi * chi \in irr G.
+Proof.
+move=> Lxi; rewrite !irrEchar => /andP[Nphi /eqP <-].
+rewrite rpredM // ?lin_charW //=; apply/eqP; congr (_ * _).
+apply: eq_bigr => x Gx; rewrite !cfunE rmorphM mulrACA -(lin_charV_conj Lxi) //.
+by rewrite -lin_charM ?groupV // mulgV lin_char1 ?mul1r.
+Qed.
+
+Lemma eq_scaled_irr a b i j :
+ (a *: 'chi[G]_i == b *: 'chi_j) = (a == b) && ((a == 0) || (i == j)).
+Proof.
+apply/eqP/andP=> [|[/eqP-> /pred2P[]-> //]]; last by rewrite !scale0r.
+move/(congr1 (cfdotr 'chi__)) => /= eq_ai_bj.
+move: {eq_ai_bj}(eq_ai_bj i) (esym (eq_ai_bj j)); rewrite !cfdotZl !cfdot_irr.
+by rewrite !mulr_natr !mulrb !eqxx eq_sym orbC; case: ifP => _ -> //= ->.
+Qed.
+
+Lemma eq_signed_irr (s t : bool) i j :
+ ((-1) ^+ s *: 'chi[G]_i == (-1) ^+ t *: 'chi_j) = (s == t) && (i == j).
+Proof. by rewrite eq_scaled_irr signr_eq0 (inj_eq (@signr_inj _)). Qed.
+
+Lemma eq_scale_irr a (i j : Iirr G) :
+ (a *: 'chi_i == a *: 'chi_j) = (a == 0) || (i == j).
+Proof. by rewrite eq_scaled_irr eqxx. Qed.
+
+Lemma eq_addZ_irr a b (i j r t : Iirr G) :
+ (a *: 'chi_i + b *: 'chi_j == a *: 'chi_r + b *: 'chi_t)
+ = [|| [&& (a == 0) || (i == r) & (b == 0) || (j == t)],
+ [&& i == t, j == r & a == b] | [&& i == j, r == t & a == - b]].
+Proof.
+rewrite -!eq_scale_irr; apply/eqP/idP; last first.
+ case/orP; first by case/andP=> /eqP-> /eqP->.
+ case/orP=> /and3P[/eqP-> /eqP-> /eqP->]; first by rewrite addrC.
+ by rewrite !scaleNr !addNr.
+have [-> /addrI/eqP-> // | /= ] := altP eqP.
+rewrite eq_scale_irr => /norP[/negP nz_a /negPf neq_ir].
+move/(congr1 (cfdotr 'chi__))/esym/eqP => /= eq_cfdot.
+move: {eq_cfdot}(eq_cfdot i) (eq_cfdot r); rewrite eq_sym !cfdotDl !cfdotZl.
+rewrite !cfdot_irr !mulr_natr !mulrb !eqxx -!(eq_sym i) neq_ir !add0r.
+have [<- _ | _] := i =P t; first by rewrite neq_ir addr0; case: ifP => // _ ->.
+rewrite 2!fun_if if_arg addr0 addr_eq0; case: eqP => //= <- ->.
+by rewrite neq_ir 2!fun_if if_arg eq_sym addr0; case: ifP.
+Qed.
+
+Lemma eq_subZnat_irr (a b : nat) (i j r t : Iirr G) :
+ (a%:R *: 'chi_i - b%:R *: 'chi_j == a%:R *: 'chi_r - b%:R *: 'chi_t)
+ = [|| a == 0%N | i == r] && [|| b == 0%N | j == t]
+ || [&& i == j, r == t & a == b].
+Proof.
+rewrite -!scaleNr eq_addZ_irr oppr_eq0 opprK -addr_eq0 -natrD eqr_nat.
+by rewrite !pnatr_eq0 addn_eq0; case: a b => [|a] [|b]; rewrite ?andbF.
+Qed.
+
+End InnerProduct.
+
+Section Sdprod.
+
+Variables (gT : finGroupType) (K H G : {group gT}).
+Hypothesis defG : K ><| H = G.
+
+Lemma cfSdprod_char chi :
+ chi \is a character -> cfSdprod defG chi \is a character.
+Proof. by move=> Nchi; rewrite unlock cfMorph_char ?cfIsom_char. Qed.
+
+Lemma cfSdprod_lin_char chi :
+ chi \is a linear_char -> cfSdprod defG chi \is a linear_char.
+Proof. by move=> Nphi; rewrite unlock cfMorph_lin_char ?cfIsom_lin_char. Qed.
+
+Lemma cfSdprod_irr chi : chi \in irr H -> cfSdprod defG chi \in irr G.
+Proof.
+have [/andP[_ nKG] _ _ _ _] := sdprod_context defG.
+by move=> Nphi; rewrite unlock cfMorph_irr ?cfIsom_irr.
+Qed.
+
+Definition sdprod_Iirr j := cfIirr (cfSdprod defG 'chi_j).
+
+Lemma sdprod_IirrE j : 'chi_(sdprod_Iirr j) = cfSdprod defG 'chi_j.
+Proof. by rewrite cfIirrE ?cfSdprod_irr ?mem_irr. Qed.
+
+Lemma sdprod_IirrK : cancel sdprod_Iirr (Res_Iirr H).
+Proof. by move=> j; rewrite /Res_Iirr sdprod_IirrE cfSdprodK irrK. Qed.
+
+Lemma sdprod_Iirr_inj : injective sdprod_Iirr.
+Proof. exact: can_inj sdprod_IirrK. Qed.
+
+Lemma sdprod_Iirr_eq0 i : (sdprod_Iirr i == 0) = (i == 0).
+Proof. by rewrite -!irr_eq1 sdprod_IirrE cfSdprod_eq1. Qed.
+
+Lemma sdprod_Iirr0 : sdprod_Iirr 0 = 0.
+Proof. by apply/eqP; rewrite sdprod_Iirr_eq0. Qed.
+
+Lemma Res_sdprod_irr phi :
+ K \subset cfker phi -> phi \in irr G -> 'Res phi \in irr H.
+Proof.
+move=> kerK /irrP[i Dphi]; rewrite irrEchar -(cfSdprod_iso defG).
+by rewrite cfRes_sdprodK // Dphi cfnorm_irr cfRes_char ?irr_char /=.
+Qed.
+
+Lemma sdprod_Res_IirrE i :
+ K \subset cfker 'chi[G]_i -> 'chi_(Res_Iirr H i) = 'Res 'chi_i.
+Proof. by move=> kerK; rewrite cfIirrE ?Res_sdprod_irr ?mem_irr. Qed.
+
+Lemma sdprod_Res_IirrK i :
+ K \subset cfker 'chi_i -> sdprod_Iirr (Res_Iirr H i) = i.
+Proof.
+by move=> kerK; rewrite /sdprod_Iirr sdprod_Res_IirrE ?cfRes_sdprodK ?irrK.
+Qed.
+
+End Sdprod.
+
+Implicit Arguments sdprod_Iirr_inj [gT K H G x1 x2].
+
+Section DProd.
+
+Variables (gT : finGroupType) (G K H : {group gT}).
+Hypothesis KxH : K \x H = G.
+
+Lemma cfDprodKl_abelian j : abelian H -> cancel ((cfDprod KxH)^~ 'chi_j) 'Res.
+Proof. by move=> cHH; apply: cfDprodKl; apply/lin_char1/char_abelianP. Qed.
+
+Lemma cfDprodKr_abelian i : abelian K -> cancel (cfDprod KxH 'chi_i) 'Res.
+Proof. by move=> cKK; apply: cfDprodKr; apply/lin_char1/char_abelianP. Qed.
+
+Lemma cfDprodl_char phi :
+ phi \is a character -> cfDprodl KxH phi \is a character.
+Proof. exact: cfSdprod_char. Qed.
+
+Lemma cfDprodr_char psi :
+ psi \is a character -> cfDprodr KxH psi \is a character.
+Proof. exact: cfSdprod_char. Qed.
+
+Lemma cfDprod_char phi psi :
+ phi \is a character -> psi \is a character ->
+ cfDprod KxH phi psi \is a character.
+Proof. by move=> /cfDprodl_char Nphi /cfDprodr_char; apply: rpredM. Qed.
+
+Lemma cfDprod_eq1 phi psi :
+ phi \is a character -> psi \is a character ->
+ (cfDprod KxH phi psi == 1) = (phi == 1) && (psi == 1).
+Proof.
+move=> /Cnat_char1 Nphi /Cnat_char1 Npsi.
+apply/eqP/andP=> [phi_psi_1 | [/eqP-> /eqP->]]; last by rewrite cfDprod_cfun1.
+have /andP[/eqP phi1 /eqP psi1]: (phi 1%g == 1) && (psi 1%g == 1).
+ by rewrite -Cnat_mul_eq1 // -(cfDprod1 KxH) phi_psi_1 cfun11.
+rewrite -[phi](cfDprodKl KxH psi1) -{2}[psi](cfDprodKr KxH phi1) phi_psi_1.
+by rewrite !rmorph1.
+Qed.
+
+Lemma cfDprodl_lin_char phi :
+ phi \is a linear_char -> cfDprodl KxH phi \is a linear_char.
+Proof. exact: cfSdprod_lin_char. Qed.
+
+Lemma cfDprodr_lin_char psi :
+ psi \is a linear_char -> cfDprodr KxH psi \is a linear_char.
+Proof. exact: cfSdprod_lin_char. Qed.
+
+Lemma cfDprod_lin_char phi psi :
+ phi \is a linear_char -> psi \is a linear_char ->
+ cfDprod KxH phi psi \is a linear_char.
+Proof. by move=> /cfDprodl_lin_char Lphi /cfDprodr_lin_char; apply: rpredM. Qed.
+
+Lemma cfDprodl_irr chi : chi \in irr K -> cfDprodl KxH chi \in irr G.
+Proof. exact: cfSdprod_irr. Qed.
+
+Lemma cfDprodr_irr chi : chi \in irr H -> cfDprodr KxH chi \in irr G.
+Proof. exact: cfSdprod_irr. Qed.
+
+Definition dprodl_Iirr i := cfIirr (cfDprodl KxH 'chi_i).
+
+Lemma dprodl_IirrE i : 'chi_(dprodl_Iirr i) = cfDprodl KxH 'chi_i.
+Proof. exact: sdprod_IirrE. Qed.
+Lemma dprodl_IirrK : cancel dprodl_Iirr (Res_Iirr K).
+Proof. exact: sdprod_IirrK. Qed.
+Lemma dprodl_Iirr_eq0 i : (dprodl_Iirr i == 0) = (i == 0).
+Proof. exact: sdprod_Iirr_eq0. Qed.
+Lemma dprodl_Iirr0 : dprodl_Iirr 0 = 0.
+Proof. exact: sdprod_Iirr0. Qed.
+
+Definition dprodr_Iirr j := cfIirr (cfDprodr KxH 'chi_j).
+
+Lemma dprodr_IirrE j : 'chi_(dprodr_Iirr j) = cfDprodr KxH 'chi_j.
+Proof. exact: sdprod_IirrE. Qed.
+Lemma dprodr_IirrK : cancel dprodr_Iirr (Res_Iirr H).
+Proof. exact: sdprod_IirrK. Qed.
+Lemma dprodr_Iirr_eq0 j : (dprodr_Iirr j == 0) = (j == 0).
+Proof. exact: sdprod_Iirr_eq0. Qed.
+Lemma dprodr_Iirr0 : dprodr_Iirr 0 = 0.
+Proof. exact: sdprod_Iirr0. Qed.
+
+Lemma cfDprod_irr i j : cfDprod KxH 'chi_i 'chi_j \in irr G.
+Proof.
+rewrite irrEchar cfDprod_char ?irr_char //=.
+by rewrite cfdot_dprod !cfdot_irr !eqxx mul1r.
+Qed.
+
+Definition dprod_Iirr ij := cfIirr (cfDprod KxH 'chi_ij.1 'chi_ij.2).
+
+Lemma dprod_IirrE i j : 'chi_(dprod_Iirr (i, j)) = cfDprod KxH 'chi_i 'chi_j.
+Proof. by rewrite cfIirrE ?cfDprod_irr. Qed.
+
+Lemma dprod_IirrEl i : 'chi_(dprod_Iirr (i, 0)) = cfDprodl KxH 'chi_i.
+Proof. by rewrite dprod_IirrE /cfDprod irr0 rmorph1 mulr1. Qed.
+
+Lemma dprod_IirrEr j : 'chi_(dprod_Iirr (0, j)) = cfDprodr KxH 'chi_j.
+Proof. by rewrite dprod_IirrE /cfDprod irr0 rmorph1 mul1r. Qed.
+
+Lemma dprod_Iirr_inj : injective dprod_Iirr.
+Proof.
+move=> [i1 j1] [i2 j2] /eqP; rewrite -[_ == _]oddb -(natCK (_ == _)).
+rewrite -cfdot_irr !dprod_IirrE cfdot_dprod !cfdot_irr -natrM mulnb.
+by rewrite natCK oddb -xpair_eqE => /eqP.
+Qed.
+
+Lemma dprod_Iirr0 : dprod_Iirr (0, 0) = 0.
+Proof. by apply/irr_inj; rewrite dprod_IirrE !irr0 cfDprod_cfun1. Qed.
+
+Lemma dprod_Iirr0l j : dprod_Iirr (0, j) = dprodr_Iirr j.
+Proof.
+by apply/irr_inj; rewrite dprod_IirrE irr0 dprodr_IirrE cfDprod_cfun1l.
+Qed.
+
+Lemma dprod_Iirr0r i : dprod_Iirr (i, 0) = dprodl_Iirr i.
+Proof.
+by apply/irr_inj; rewrite dprod_IirrE irr0 dprodl_IirrE cfDprod_cfun1r.
+Qed.
+
+Lemma dprod_Iirr_eq0 i j : (dprod_Iirr (i, j) == 0) = (i == 0) && (j == 0).
+Proof. by rewrite -xpair_eqE -(inj_eq dprod_Iirr_inj) dprod_Iirr0. Qed.
+
+Lemma cfdot_dprod_irr i1 i2 j1 j2 :
+ '['chi_(dprod_Iirr (i1, j1)), 'chi_(dprod_Iirr (i2, j2))]
+ = ((i1 == i2) && (j1 == j2))%:R.
+Proof. by rewrite cfdot_irr (inj_eq dprod_Iirr_inj). Qed.
+
+Lemma dprod_Iirr_onto k : k \in codom dprod_Iirr.
+Proof.
+set D := codom _; have Df: dprod_Iirr _ \in D := codom_f dprod_Iirr _.
+have: 'chi_k 1%g ^+ 2 != 0 by rewrite mulf_neq0 ?irr1_neq0.
+apply: contraR => notDk; move/eqP: (irr_sum_square G).
+rewrite (bigID (mem D)) (reindex _ (bij_on_codom dprod_Iirr_inj (0, 0))) /=.
+have ->: #|G|%:R = \sum_i \sum_j 'chi_(dprod_Iirr (i, j)) 1%g ^+ 2.
+ rewrite -(dprod_card KxH) natrM.
+ do 2![rewrite -irr_sum_square (mulr_suml, mulr_sumr); apply: eq_bigr => ? _].
+ by rewrite dprod_IirrE -exprMn -{3}(mulg1 1%g) cfDprodE.
+rewrite (eq_bigl _ _ Df) pair_bigA addrC -subr_eq0 addrK.
+by move/eqP/psumr_eq0P=> -> //= i _; rewrite irr1_degree -natrX ler0n.
+Qed.
+
+Definition inv_dprod_Iirr i := iinv (dprod_Iirr_onto i).
+
+Lemma dprod_IirrK : cancel dprod_Iirr inv_dprod_Iirr.
+Proof. by move=> p; exact: (iinv_f dprod_Iirr_inj). Qed.
+
+Lemma inv_dprod_IirrK : cancel inv_dprod_Iirr dprod_Iirr.
+Proof. by move=> i; exact: f_iinv. Qed.
+
+Lemma inv_dprod_Iirr0 : inv_dprod_Iirr 0 = (0, 0).
+Proof. by apply/(canLR dprod_IirrK); rewrite dprod_Iirr0. Qed.
+
+End DProd.
+
+Implicit Arguments dprod_Iirr_inj [gT G K H x1 x2].
+
+Lemma dprod_IirrC (gT : finGroupType) (G K H : {group gT})
+ (KxH : K \x H = G) (HxK : H \x K = G) i j :
+ dprod_Iirr KxH (i, j) = dprod_Iirr HxK (j, i).
+Proof. by apply: irr_inj; rewrite !dprod_IirrE; apply: cfDprodC. Qed.
+
+Section BigDprod.
+
+Variables (gT : finGroupType) (I : finType) (P : pred I).
+Variables (A : I -> {group gT}) (G : {group gT}).
+Hypothesis defG : \big[dprod/1%g]_(i | P i) A i = G.
+
+Let sAG i : P i -> A i \subset G.
+Proof. by move=> Pi; rewrite -(bigdprodWY defG) (bigD1 i) ?joing_subl. Qed.
+
+Lemma cfBigdprodi_char i (phi : 'CF(A i)) :
+ phi \is a character -> cfBigdprodi defG phi \is a character.
+Proof. by move=> Nphi; rewrite cfDprodl_char ?cfRes_char. Qed.
+
+Lemma cfBigdprod_char phi :
+ (forall i, P i -> phi i \is a character) ->
+ cfBigdprod defG phi \is a character.
+Proof.
+by move=> Nphi; apply: rpred_prod => i /Nphi; apply: cfBigdprodi_char.
+Qed.
+
+Lemma cfBigdprodi_lin_char i (phi : 'CF(A i)) :
+ phi \is a linear_char -> cfBigdprodi defG phi \is a linear_char.
+Proof. by move=> Lphi; rewrite cfDprodl_lin_char ?cfRes_lin_char. Qed.
+
+Lemma cfBigdprod_lin_char phi :
+ (forall i, P i -> phi i \is a linear_char) ->
+ cfBigdprod defG phi \is a linear_char.
+Proof.
+by move=> Lphi; apply/rpred_prod=> i /Lphi; apply: cfBigdprodi_lin_char.
+Qed.
+
+Lemma cfBigdprodi_irr i chi :
+ P i -> chi \in irr (A i) -> cfBigdprodi defG chi \in irr G.
+Proof. by move=> Pi Nchi; rewrite cfDprodl_irr // Pi cfRes_id. Qed.
+
+Lemma cfBigdprod_irr chi :
+ (forall i, P i -> chi i \in irr (A i)) -> cfBigdprod defG chi \in irr G.
+Proof.
+move=> Nchi; rewrite irrEchar cfBigdprod_char => [|i /Nchi/irrWchar] //=.
+by rewrite cfdot_bigdprod big1 // => i /Nchi/irrWnorm.
+Qed.
+
+Lemma cfBigdprod_eq1 phi :
+ (forall i, P i -> phi i \is a character) ->
+ (cfBigdprod defG phi == 1) = [forall (i | P i), phi i == 1].
+Proof.
+move=> Nphi; set Phi := cfBigdprod defG phi.
+apply/eqP/eqfun_inP=> [Phi1 i Pi | phi1]; last first.
+ by apply: big1 => i /phi1->; rewrite rmorph1.
+have Phi1_1: Phi 1%g = 1 by rewrite Phi1 cfun1E group1.
+have nz_Phi1: Phi 1%g != 0 by rewrite Phi1_1 oner_eq0.
+have [_ <-] := cfBigdprodK nz_Phi1 Pi.
+rewrite Phi1_1 divr1 -/Phi Phi1 rmorph1.
+rewrite prod_cfunE // in Phi1_1; have := Cnat_prod_eq1 _ Phi1_1 Pi.
+rewrite -(cfRes1 (A i)) cfBigdprodiK // => ->; first by rewrite scale1r.
+by move=> {i Pi} j /Nphi Nphi_j; rewrite Cnat_char1 ?cfBigdprodi_char.
+Qed.
+
+Lemma cfBigdprod_Res_lin chi :
+ chi \is a linear_char -> cfBigdprod defG (fun i => 'Res[A i] chi) = chi.
+Proof.
+move=> Lchi; apply/cfun_inP=> _ /(mem_bigdprod defG)[x [Ax -> _]].
+rewrite (lin_char_prod Lchi) ?cfBigdprodE // => [|i Pi]; last first.
+ by rewrite (subsetP (sAG Pi)) ?Ax.
+by apply/eq_bigr=> i Pi; rewrite cfResE ?sAG ?Ax.
+Qed.
+
+Lemma cfBigdprodKlin phi :
+ (forall i, P i -> phi i \is a linear_char) ->
+ forall i, P i -> 'Res (cfBigdprod defG phi) = phi i.
+Proof.
+move=> Lphi i Pi; have Lpsi := cfBigdprod_lin_char Lphi.
+have [_ <-] := cfBigdprodK (lin_char_neq0 Lpsi (group1 G)) Pi.
+by rewrite !lin_char1 ?Lphi // divr1 scale1r.
+Qed.
+
+Lemma cfBigdprodKabelian Iphi (phi := fun i => 'chi_(Iphi i)) :
+ abelian G -> forall i, P i -> 'Res (cfBigdprod defG phi) = 'chi_(Iphi i).
+Proof.
+move=> /(abelianS _) cGG.
+by apply: cfBigdprodKlin => i /sAG/cGG/char_abelianP->.
+Qed.
+
+End BigDprod.
+
+Section Aut.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Implicit Type u : {rmorphism algC -> algC}.
+
+Lemma conjC_charAut u (chi : 'CF(G)) x :
+ chi \is a character -> (u (chi x))^* = u (chi x)^*.
+Proof.
+have [Gx | /cfun0->] := boolP (x \in G); last by rewrite !rmorph0.
+case/char_reprP=> rG ->; have [e [_ [en1 _] [-> _] _]] := repr_rsim_diag rG Gx.
+by rewrite !rmorph_sum; apply: eq_bigr => i _; exact: aut_unity_rootC (en1 i).
+Qed.
+
+Lemma conjC_irrAut u i x : (u ('chi[G]_i x))^* = u ('chi_i x)^*.
+Proof. exact: conjC_charAut (irr_char i). Qed.
+
+Lemma cfdot_aut_char u (phi chi : 'CF(G)) :
+ chi \is a character -> '[cfAut u phi, cfAut u chi] = u '[phi, chi].
+Proof. by move/conjC_charAut=> Nchi; apply: cfdot_cfAut => _ /mapP[x _ ->]. Qed.
+
+Lemma cfdot_aut_irr u phi i :
+ '[cfAut u phi, cfAut u 'chi[G]_i] = u '[phi, 'chi_i].
+Proof. exact: cfdot_aut_char (irr_char i). Qed.
+
+Lemma cfAut_irr u chi : chi \in irr G -> cfAut u chi \in irr G.
+Proof.
+case/irrP=> i ->; rewrite irrEchar cfAut_char ?irr_char //=.
+by rewrite cfdot_aut_irr // cfdot_irr eqxx rmorph1.
+Qed.
+
+Lemma cfConjC_irr i : (('chi_i)^*)%CF \in irr G.
+Proof. by rewrite cfAut_irr ?mem_irr. Qed.
+
+Lemma irr_aut_closed u : cfAut_closed u (irr G).
+Proof. exact: cfAut_irr. Qed.
+
+Definition aut_Iirr u i := cfIirr (cfAut u 'chi[G]_i).
+
+Lemma aut_IirrE u i : 'chi_(aut_Iirr u i) = cfAut u 'chi_i.
+Proof. by rewrite cfIirrE ?cfAut_irr ?mem_irr. Qed.
+
+Definition conjC_Iirr := aut_Iirr conjC.
+
+Lemma conjC_IirrE i : 'chi[G]_(conjC_Iirr i) = ('chi_i)^*%CF.
+Proof. exact: aut_IirrE. Qed.
+
+Lemma conjC_IirrK : involutive conjC_Iirr.
+Proof. by move=> i; apply: irr_inj; rewrite !conjC_IirrE cfConjCK. Qed.
+
+Lemma aut_Iirr0 u : aut_Iirr u 0 = 0 :> Iirr G.
+Proof. by apply/irr_inj; rewrite aut_IirrE irr0 cfAut_cfun1. Qed.
+
+Lemma conjC_Iirr0 : conjC_Iirr 0 = 0 :> Iirr G.
+Proof. exact: aut_Iirr0. Qed.
+
+Lemma aut_Iirr_eq0 u i : (aut_Iirr u i == 0) = (i == 0).
+Proof. by rewrite -!irr_eq1 aut_IirrE cfAut_eq1. Qed.
+
+Lemma conjC_Iirr_eq0 i : (conjC_Iirr i == 0 :> Iirr G) = (i == 0).
+Proof. exact: aut_Iirr_eq0. Qed.
+
+Lemma aut_Iirr_inj u : injective (aut_Iirr u).
+Proof.
+by move=> i j eq_ij; apply/irr_inj/(cfAut_inj u); rewrite -!aut_IirrE eq_ij.
+Qed.
+
+Lemma char_aut u (chi : 'CF(G)) :
+ (cfAut u chi \is a character) = (chi \is a character).
+Proof.
+apply/idP/idP=> [Nuchi|]; last exact: cfAut_char.
+rewrite [chi]cfun_sum_cfdot rpred_sum // => i _.
+rewrite rpredZ_Cnat ?irr_char // -(Cnat_aut u) -cfdot_aut_irr.
+by rewrite -aut_IirrE Cnat_cfdot_char_irr.
+Qed.
+
+Lemma irr_aut u chi : (cfAut u chi \in irr G) = (chi \in irr G).
+Proof.
+rewrite !irrEchar char_aut; apply/andb_id2l=> /cfdot_aut_char->.
+by rewrite fmorph_eq1.
+Qed.
+
+End Aut.
+
+Section IrrConstt.
+
+Variable (gT : finGroupType) (G H : {group gT}).
+
+Lemma char1_ge_norm (chi : 'CF(G)) x :
+ chi \is a character -> `|chi x| <= chi 1%g.
+Proof.
+case/char_reprP=> rG ->; case Gx: (x \in G); last first.
+ by rewrite cfunE cfRepr1 Gx normr0 ler0n.
+by have [e [_ _ []]] := repr_rsim_diag rG Gx.
+Qed.
+
+Lemma max_cfRepr_norm_scalar n (rG : mx_representation algCF G n) x :
+ x \in G -> `|cfRepr rG x| = cfRepr rG 1%g ->
+ exists2 c, `|c| = 1 & rG x = c%:M.
+Proof.
+move=> Gx; have [e [[B uB def_x] [_ e1] [-> _] _]] := repr_rsim_diag rG Gx.
+rewrite cfRepr1 -[n in n%:R]card_ord -sumr_const -(eq_bigr _ (in1W e1)).
+case/normC_sum_eq1=> [i _ | c /eqP norm_c_1 def_e]; first by rewrite e1.
+have{def_e} def_e: e = const_mx c by apply/rowP=> i; rewrite mxE def_e ?andbT.
+by exists c => //; rewrite def_x def_e diag_const_mx scalar_mxC mulmxKV.
+Qed.
+
+Lemma max_cfRepr_mx1 n (rG : mx_representation algCF G n) x :
+ x \in G -> cfRepr rG x = cfRepr rG 1%g -> rG x = 1%:M.
+Proof.
+move=> Gx kerGx; have [|c _ def_x] := @max_cfRepr_norm_scalar n rG x Gx.
+ by rewrite kerGx cfRepr1 normr_nat.
+move/eqP: kerGx; rewrite cfRepr1 cfunE Gx {rG}def_x mxtrace_scalar.
+case: n => [_|n]; first by rewrite ![_%:M]flatmx0.
+rewrite mulrb -subr_eq0 -mulrnBl -mulr_natl mulf_eq0 pnatr_eq0 /=.
+by rewrite subr_eq0 => /eqP->.
+Qed.
+
+Definition irr_constt (B : {set gT}) phi := [pred i | '[phi, 'chi_i]_B != 0].
+
+Lemma irr_consttE i phi : (i \in irr_constt phi) = ('[phi, 'chi_i]_G != 0).
+Proof. by []. Qed.
+
+Lemma constt_charP (i : Iirr G) chi :
+ chi \is a character ->
+ reflect (exists2 chi', chi' \is a character & chi = 'chi_i + chi')
+ (i \in irr_constt chi).
+Proof.
+move=> Nchi; apply: (iffP idP) => [i_in_chi| [chi' Nchi' ->]]; last first.
+ rewrite inE /= cfdotDl cfdot_irr eqxx -(eqP (Cnat_cfdot_char_irr i Nchi')).
+ by rewrite -natrD pnatr_eq0.
+exists (chi - 'chi_i); last by rewrite addrC subrK.
+apply/forallP=> j; rewrite coord_cfdot cfdotBl cfdot_irr.
+have [<- | _] := eqP; last by rewrite subr0 Cnat_cfdot_char_irr.
+have := i_in_chi; rewrite inE /= -(eqP (Cnat_cfdot_char_irr i Nchi)) pnatr_eq0.
+by case: (truncC _) => // n _; rewrite mulrSr addrK ?isNatC_nat.
+Qed.
+
+Lemma cfun_sum_constt (phi : 'CF(G)) :
+ phi = \sum_(i in irr_constt phi) '[phi, 'chi_i] *: 'chi_i.
+Proof.
+rewrite {1}[phi]cfun_sum_cfdot (bigID [pred i | '[phi, 'chi_i] == 0]) /=.
+by rewrite big1 ?add0r // => i /eqP->; rewrite scale0r.
+Qed.
+
+Lemma neq0_has_constt (phi : 'CF(G)) :
+ phi != 0 -> exists i, i \in irr_constt phi.
+Proof.
+move=> nz_phi; apply/existsP; apply: contra nz_phi => /pred0P phi0.
+by rewrite [phi]cfun_sum_constt big_pred0.
+Qed.
+
+Lemma constt_irr i : irr_constt 'chi[G]_i =i pred1 i.
+Proof.
+by move=> j; rewrite !inE cfdot_irr pnatr_eq0 (eq_sym j); case: (i == j).
+Qed.
+
+Lemma char1_ge_constt (i : Iirr G) chi :
+ chi \is a character -> i \in irr_constt chi -> 'chi_i 1%g <= chi 1%g.
+Proof.
+move=> {chi} _ /constt_charP[// | chi Nchi ->].
+by rewrite cfunE addrC -subr_ge0 addrK char1_ge0.
+Qed.
+
+Lemma constt_ortho_char (phi psi : 'CF(G)) i j :
+ phi \is a character -> psi \is a character ->
+ i \in irr_constt phi -> j \in irr_constt psi ->
+ '[phi, psi] = 0 -> '['chi_i, 'chi_j] = 0.
+Proof.
+move=> _ _ /constt_charP[//|phi1 Nphi1 ->] /constt_charP[//|psi1 Npsi1 ->].
+rewrite cfdot_irr; case: eqP => // -> /eqP/idPn[].
+rewrite cfdotDl !cfdotDr cfnorm_irr -addrA gtr_eqF ?ltr_paddr ?ltr01 //.
+by rewrite Cnat_ge0 ?rpredD ?Cnat_cfdot_char ?irr_char.
+Qed.
+
+End IrrConstt.
+
+Arguments Scope irr_constt [_ group_scope cfun_scope].
+Implicit Arguments aut_Iirr_inj [gT G x1 x2].
+
+Section MoreConstt.
+
+Variables (gT : finGroupType) (G H : {group gT}).
+
+Lemma constt_Ind_Res i j :
+ i \in irr_constt ('Ind[G] 'chi_j) = (j \in irr_constt ('Res[H] 'chi_i)).
+Proof. by rewrite !irr_consttE cfdotC conjC_eq0 -cfdot_Res_l. Qed.
+
+Lemma cfdot_Res_ge_constt i j psi :
+ psi \is a character -> j \in irr_constt psi ->
+ '['Res[H, G] 'chi_j, 'chi_i] <= '['Res[H] psi, 'chi_i].
+Proof.
+move=> {psi} _ /constt_charP[// | psi Npsi ->].
+rewrite linearD cfdotDl addrC -subr_ge0 addrK Cnat_ge0 //=.
+by rewrite Cnat_cfdot_char_irr // cfRes_char.
+Qed.
+
+Lemma constt_Res_trans j psi :
+ psi \is a character -> j \in irr_constt psi ->
+ {subset irr_constt ('Res[H, G] 'chi_j) <= irr_constt ('Res[H] psi)}.
+Proof.
+move=> Npsi Cj i; apply: contraNneq; rewrite eqr_le => {1}<-.
+rewrite cfdot_Res_ge_constt ?Cnat_ge0 ?Cnat_cfdot_char_irr //.
+by rewrite cfRes_char ?irr_char.
+Qed.
+
+End MoreConstt.
+
+Section Kernel.
+
+Variable (gT : finGroupType) (G : {group gT}).
+Implicit Types (phi chi xi : 'CF(G)) (H : {group gT}).
+
+Lemma cfker_repr n (rG : mx_representation algCF G n) :
+ cfker (cfRepr rG) = rker rG.
+Proof.
+apply/esym/setP=> x; rewrite inE mul1mx /=.
+case Gx: (x \in G); last by rewrite inE Gx.
+apply/eqP/idP=> Kx; last by rewrite max_cfRepr_mx1 // cfker1.
+rewrite inE Gx; apply/forallP=> y; rewrite !cfunE !mulrb groupMl //.
+by case: ifP => // Gy; rewrite repr_mxM // Kx mul1mx.
+Qed.
+
+Lemma cfkerEchar chi :
+ chi \is a character -> cfker chi = [set x in G | chi x == chi 1%g].
+Proof.
+move=> Nchi; apply/setP=> x; apply/idP/setIdP=> [Kx | [Gx /eqP chi_x]].
+ by rewrite (subsetP (cfker_sub chi)) // cfker1.
+case/char_reprP: Nchi => rG -> in chi_x *; rewrite inE Gx; apply/forallP=> y.
+rewrite !cfunE groupMl // !mulrb; case: ifP => // Gy.
+by rewrite repr_mxM // max_cfRepr_mx1 ?mul1mx.
+Qed.
+
+Lemma cfker_nzcharE chi :
+ chi \is a character -> chi != 0 -> cfker chi = [set x | chi x == chi 1%g].
+Proof.
+move=> Nchi nzchi; apply/setP=> x; rewrite cfkerEchar // !inE andb_idl //.
+by apply: contraLR => /cfun0-> //; rewrite eq_sym char1_eq0.
+Qed.
+
+Lemma cfkerEirr i : cfker 'chi[G]_i = [set x | 'chi_i x == 'chi_i 1%g].
+Proof. by rewrite cfker_nzcharE ?irr_char ?irr_neq0. Qed.
+
+Lemma cfker_irr0 : cfker 'chi[G]_0 = G.
+Proof. by rewrite irr0 cfker_cfun1. Qed.
+
+Lemma cfaithful_reg : cfaithful (cfReg G).
+Proof.
+apply/subsetP=> x; rewrite cfkerEchar ?cfReg_char // !inE !cfRegE eqxx.
+by case/andP=> _; apply: contraLR => /negbTE->; rewrite eq_sym neq0CG.
+Qed.
+
+Lemma cfkerE chi :
+ chi \is a character ->
+ cfker chi = G :&: \bigcap_(i in irr_constt chi) cfker 'chi_i.
+Proof.
+move=> Nchi; rewrite cfkerEchar //; apply/setP=> x; rewrite !inE.
+apply: andb_id2l => Gx; rewrite {1 2}[chi]cfun_sum_constt !sum_cfunE.
+apply/eqP/bigcapP=> [Kx i Ci | Kx]; last first.
+ by apply: eq_bigr => i /Kx Kx_i; rewrite !cfunE cfker1.
+rewrite cfkerEirr inE /= -(inj_eq (mulfI Ci)).
+have:= (normC_sum_upper _ Kx) i; rewrite !cfunE => -> // {i Ci} i _.
+have chi_i_ge0: 0 <= '[chi, 'chi_i].
+ by rewrite Cnat_ge0 ?Cnat_cfdot_char_irr.
+by rewrite !cfunE normrM (normr_idP _) ?ler_wpmul2l ?char1_ge_norm ?irr_char.
+Qed.
+
+Lemma TI_cfker_irr : \bigcap_i cfker 'chi[G]_i = [1].
+Proof.
+apply/trivgP; apply: subset_trans cfaithful_reg; rewrite cfkerE ?cfReg_char //.
+rewrite subsetI (bigcap_min 0) //=; last by rewrite cfker_irr0.
+by apply/bigcapsP=> i _; rewrite bigcap_inf.
+Qed.
+
+Lemma cfker_constt i chi :
+ chi \is a character -> i \in irr_constt chi ->
+ cfker chi \subset cfker 'chi[G]_i.
+Proof. by move=> Nchi Ci; rewrite cfkerE ?subIset ?(bigcap_min i) ?orbT. Qed.
+
+Section KerLin.
+
+Variable xi : 'CF(G).
+Hypothesis lin_xi : xi \is a linear_char.
+Let Nxi: xi \is a character. Proof. by have [] := andP lin_xi. Qed.
+
+Lemma lin_char_der1 : G^`(1)%g \subset cfker xi.
+Proof.
+rewrite gen_subG /=; apply/subsetP=> _ /imset2P[x y Gx Gy ->].
+rewrite cfkerEchar // inE groupR //= !lin_charM ?lin_charV ?in_group //.
+by rewrite mulrCA mulKf ?mulVf ?lin_char_neq0 // lin_char1.
+Qed.
+
+Lemma cforder_lin_char : #[xi]%CF = exponent (G / cfker xi)%g.
+Proof.
+apply/eqP; rewrite eqn_dvd; apply/andP; split.
+ apply/dvdn_cforderP=> x Gx; rewrite -lin_charX // -cfQuoEker ?groupX //.
+ rewrite morphX ?(subsetP (cfker_norm xi)) //= expg_exponent ?mem_quotient //.
+ by rewrite cfQuo1 ?cfker_normal ?lin_char1.
+have abGbar: abelian (G / cfker xi) := sub_der1_abelian lin_char_der1.
+have [_ /morphimP[x Nx Gx ->] ->] := exponent_witness (abelian_nil abGbar).
+rewrite order_dvdn -morphX //= coset_id cfkerEchar // !inE groupX //=.
+by rewrite lin_charX ?lin_char1 // (dvdn_cforderP _ _ _).
+Qed.
+
+Lemma cforder_lin_char_dvdG : #[xi]%CF %| #|G|.
+Proof.
+by rewrite cforder_lin_char (dvdn_trans (exponent_dvdn _)) ?dvdn_morphim.
+Qed.
+
+Lemma cforder_lin_char_gt0 : (0 < #[xi]%CF)%N.
+Proof. by rewrite cforder_lin_char exponent_gt0. Qed.
+
+End KerLin.
+
+End Kernel.
+
+Section Coset.
+
+Variable (gT : finGroupType).
+
+Implicit Types G H : {group gT}.
+
+Lemma cfQuo_char G H (chi : 'CF(G)) :
+ chi \is a character -> (chi / H)%CF \is a character.
+Proof.
+move=> Nchi; case KchiH: (H \subset cfker chi); last first.
+ suffices ->: (chi / H)%CF = (chi 1%g)%:A.
+ by rewrite rpredZ_Cnat ?Cnat_char1 ?rpred1.
+ by apply/cfunP=> x; rewrite cfunE cfun1E mulr_natr cfunElock KchiH.
+have sHG := subset_trans KchiH (cfker_sub _).
+pose N := 'N_G(H); pose phi := 'Res[N] chi.
+have nsHN: H <| N by [rewrite normal_subnorm]; have [sHN nHN] := andP nsHN.
+have{Nchi} Nphi: phi \is a character by apply: cfRes_char.
+have KphiH: H \subset cfker phi.
+ apply/subsetP=> x Hx; have [Kx Nx] := (subsetP KchiH x Hx, subsetP sHN x Hx).
+ by rewrite cfkerEchar // inE Nx cfRes1 cfResE ?subsetIl //= cfker1.
+pose psi := 'Res[(N / H)%g] (chi / H)%CF.
+have ->: (chi / H)%CF = 'Res psi by rewrite /psi quotientInorm !cfRes_id.
+have{KchiH} ->: psi = (phi / H)%CF.
+ apply/cfun_inP => _ /morphimP[x nHx Nx ->]; have [Gx _] := setIP Nx.
+ rewrite cfResE ?mem_quotient ?quotientS ?subsetIl // cfQuoEnorm //.
+ by rewrite cfQuoE ?cfResE ?subsetIl.
+have [rG Dphi] := char_reprP Nphi; rewrite {phi Nphi}Dphi cfker_repr in KphiH *.
+apply/cfRes_char/char_reprP; exists (Representation (quo_repr KphiH nHN)).
+apply/cfun_inP=> _ /morphimP[x nHx Nx ->]; rewrite cfQuoE ?cfker_repr //=.
+by rewrite !cfunE Nx quo_repr_coset ?mem_quotient.
+Qed.
+
+Lemma cfQuo_lin_char G H (chi : 'CF(G)) :
+ chi \is a linear_char -> (chi / H)%CF \is a linear_char.
+Proof. by case/andP=> Nchi; rewrite qualifE cfQuo_char ?cfQuo1. Qed.
+
+Lemma cfMod_char G H (chi : 'CF(G / H)) :
+ chi \is a character -> (chi %% H)%CF \is a character.
+Proof. exact: cfMorph_char. Qed.
+
+Lemma cfMod_lin_char G H (chi : 'CF(G / H)) :
+ chi \is a linear_char -> (chi %% H)%CF \is a linear_char.
+Proof. exact: cfMorph_lin_char. Qed.
+
+Lemma cfMod_irr G H chi :
+ H <| G -> chi \in irr (G / H) -> (chi %% H)%CF \in irr G.
+Proof. by case/andP=> _; apply: cfMorph_irr. Qed.
+
+Definition mod_Iirr G H i := cfIirr ('chi[G / H]_i %% H)%CF.
+
+Lemma mod_Iirr0 G H : mod_Iirr (0 : Iirr (G / H)) = 0.
+Proof. exact: morph_Iirr0. Qed.
+
+Lemma mod_IirrE G H i : H <| G -> 'chi_(mod_Iirr i) = ('chi[G / H]_i %% H)%CF.
+Proof. by move=> nsHG; rewrite cfIirrE ?cfMod_irr ?mem_irr. Qed.
+
+Lemma mod_Iirr_eq0 G H i :
+ H <| G -> (mod_Iirr i == 0) = (i == 0 :> Iirr (G / H)).
+Proof. by case/andP=> _ /morph_Iirr_eq0->. Qed.
+
+Lemma cfQuo_irr G H chi :
+ H <| G -> H \subset cfker chi -> chi \in irr G ->
+ (chi / H)%CF \in irr (G / H).
+Proof.
+move=> nsHG sHK /irr_reprP[rG irrG Dchi]; have [sHG nHG] := andP nsHG.
+have sHKr: H \subset rker rG by rewrite -cfker_repr -Dchi.
+apply/irr_reprP; exists (Representation (quo_repr sHKr nHG)).
+ exact/quo_mx_irr.
+apply/cfun_inP=> _ /morphimP[x Nx Gx ->].
+by rewrite cfQuoE //= Dchi !cfunE Gx quo_repr_coset ?mem_quotient.
+Qed.
+
+Definition quo_Iirr G H i := cfIirr ('chi[G]_i / H)%CF.
+
+Lemma quo_Iirr0 G H : quo_Iirr H (0 : Iirr G) = 0.
+Proof. by rewrite /quo_Iirr irr0 cfQuo_cfun1 -irr0 irrK. Qed.
+
+Lemma quo_IirrE G H i :
+ H <| G -> H \subset cfker 'chi[G]_i -> 'chi_(quo_Iirr H i) = ('chi_i / H)%CF.
+Proof. by move=> nsHG kerH; rewrite cfIirrE ?cfQuo_irr ?mem_irr. Qed.
+
+Lemma quo_Iirr_eq0 G H i :
+ H <| G -> H \subset cfker 'chi[G]_i -> (quo_Iirr H i == 0) = (i == 0).
+Proof. by move=> nsHG kerH; rewrite -!irr_eq1 quo_IirrE ?cfQuo_eq1. Qed.
+
+Lemma mod_IirrK G H : H <| G -> cancel (@mod_Iirr G H) (@quo_Iirr G H).
+Proof.
+move=> nsHG i; apply: irr_inj.
+by rewrite quo_IirrE ?mod_IirrE ?cfker_mod // cfModK.
+Qed.
+
+Lemma quo_IirrK G H i :
+ H <| G -> H \subset cfker 'chi[G]_i -> mod_Iirr (quo_Iirr H i) = i.
+Proof.
+by move=> nsHG kerH; apply: irr_inj; rewrite mod_IirrE ?quo_IirrE ?cfQuoK.
+Qed.
+
+Lemma quo_IirrKeq G H :
+ H <| G ->
+ forall i, (mod_Iirr (quo_Iirr H i) == i) = (H \subset cfker 'chi[G]_i).
+Proof.
+move=> nsHG i; apply/eqP/idP=> [<- | ]; last exact: quo_IirrK.
+by rewrite mod_IirrE ?cfker_mod.
+Qed.
+
+Lemma mod_Iirr_bij H G :
+ H <| G -> {on [pred i | H \subset cfker 'chi_i], bijective (@mod_Iirr G H)}.
+Proof.
+by exists (quo_Iirr H) => [i _ | i]; [exact: mod_IirrK | exact: quo_IirrK].
+Qed.
+
+Lemma sum_norm_irr_quo H G x :
+ x \in G -> H <| G ->
+ \sum_i `|'chi[G / H]_i (coset H x)| ^+ 2
+ = \sum_(i | H \subset cfker 'chi_i) `|'chi[G]_i x| ^+ 2.
+Proof.
+move=> Gx nsHG; rewrite (reindex _ (mod_Iirr_bij nsHG)) /=.
+by apply/esym/eq_big=> [i | i _]; rewrite mod_IirrE ?cfker_mod ?cfModE.
+Qed.
+
+Lemma cap_cfker_normal G H :
+ H <| G -> \bigcap_(i | H \subset cfker 'chi[G]_i) (cfker 'chi_i) = H.
+Proof.
+move=> nsHG; have [sHG nHG] := andP nsHG; set lhs := \bigcap_(i | _) _.
+have nHlhs: lhs \subset 'N(H) by rewrite (bigcap_min 0) ?cfker_irr0.
+apply/esym/eqP; rewrite eqEsubset (introT bigcapsP) //= -quotient_sub1 //.
+rewrite -(TI_cfker_irr (G / H)); apply/bigcapsP=> i _.
+rewrite sub_quotient_pre // (bigcap_min (mod_Iirr i)) ?mod_IirrE ?cfker_mod //.
+by rewrite cfker_morph ?subsetIr.
+Qed.
+
+Lemma cfker_reg_quo G H : H <| G -> cfker (cfReg (G / H)%g %% H) = H.
+Proof.
+move=> nsHG; have [sHG nHG] := andP nsHG.
+apply/setP=> x; rewrite cfkerEchar ?cfMod_char ?cfReg_char //.
+rewrite -[in RHS in _ = RHS](setIidPr sHG) !inE; apply: andb_id2l => Gx.
+rewrite !cfModE // !cfRegE // morph1 eqxx.
+rewrite (sameP eqP (kerP _ (subsetP nHG x Gx))) ker_coset.
+by rewrite -!mulrnA eqr_nat eqn_pmul2l ?cardG_gt0 // (can_eq oddb) eqb_id.
+Qed.
+
+End Coset.
+
+Section Derive.
+
+Variable gT : finGroupType.
+Implicit Types G H : {group gT}.
+
+Lemma lin_irr_der1 G i :
+ ('chi_i \is a linear_char) = (G^`(1)%g \subset cfker 'chi[G]_i).
+Proof.
+apply/idP/idP=> [|sG'K]; first by apply: lin_char_der1.
+have nsG'G: G^`(1) <| G := der_normal 1 G.
+rewrite qualifE irr_char -[i](quo_IirrK nsG'G) // mod_IirrE //=.
+by rewrite cfModE // morph1 lin_char1 //; exact/char_abelianP/der_abelian.
+Qed.
+
+Lemma subGcfker G i : (G \subset cfker 'chi[G]_i) = (i == 0).
+Proof.
+rewrite -irr_eq1; apply/idP/eqP=> [chiG1 | ->]; last by rewrite cfker_cfun1.
+apply/cfun_inP=> x Gx; rewrite cfun1E Gx cfker1 ?(subsetP chiG1) ?lin_char1 //.
+by rewrite lin_irr_der1 (subset_trans (der_sub 1 G)).
+Qed.
+
+Lemma irr_prime_injP G i :
+ prime #|G| -> reflect {in G &, injective 'chi[G]_i} (i != 0).
+Proof.
+move=> pr_G; apply: (iffP idP) => [nz_i | inj_chi].
+ apply: fful_lin_char_inj (irr_prime_lin i pr_G) _.
+ by rewrite cfaithfulE -(setIidPr (cfker_sub _)) prime_TIg // subGcfker.
+have /trivgPn[x Gx ntx]: G :!=: 1%g by rewrite -cardG_gt1 prime_gt1.
+apply: contraNneq ntx => i0; apply/eqP/inj_chi=> //.
+by rewrite i0 irr0 !cfun1E Gx group1.
+Qed.
+
+(* This is Isaacs (2.23)(a). *)
+Lemma cap_cfker_lin_irr G :
+ \bigcap_(i | 'chi[G]_i \is a linear_char) (cfker 'chi_i) = G^`(1)%g.
+Proof.
+rewrite -(cap_cfker_normal (der_normal 1 G)).
+by apply: eq_bigl => i; rewrite lin_irr_der1.
+Qed.
+
+(* This is Isaacs (2.23)(b) *)
+Lemma card_lin_irr G :
+ #|[pred i | 'chi[G]_i \is a linear_char]| = #|G : G^`(1)%g|.
+Proof.
+have nsG'G := der_normal 1 G; rewrite (eq_card (@lin_irr_der1 G)).
+rewrite -(on_card_preimset (mod_Iirr_bij nsG'G)).
+rewrite -card_quotient ?normal_norm //.
+move: (der_abelian 0 G); rewrite card_classes_abelian; move/eqP<-.
+rewrite -NirrE -[X in _ = X]card_ord.
+by apply: eq_card => i; rewrite !inE mod_IirrE ?cfker_mod.
+(* Alternative: use the equivalent result in modular representation theory
+transitivity #|@socle_of_Iirr _ G @^-1: linear_irr _|; last first.
+ rewrite (on_card_preimset (socle_of_Iirr_bij _)).
+ by rewrite card_linear_irr ?algC'G; last exact: groupC.
+by apply: eq_card => i; rewrite !inE /lin_char irr_char irr1_degree -eqC_nat.
+*)
+Qed.
+
+(* A non-trivial solvable group has a nonprincipal linear character. *)
+Lemma solvable_has_lin_char G :
+ G :!=: 1%g -> solvable G ->
+ exists2 i, 'chi[G]_i \is a linear_char & 'chi_i != 1.
+Proof.
+move=> ntG solG.
+suff /subsetPn[i]: ~~ ([pred i | 'chi[G]_i \is a linear_char] \subset pred1 0).
+ by rewrite !inE -(inj_eq irr_inj) irr0; exists i.
+rewrite (contra (@subset_leq_card _ _ _)) // -ltnNge card1 card_lin_irr.
+by rewrite indexg_gt1 proper_subn // (sol_der1_proper solG).
+Qed.
+
+(* A combinatorial group isommorphic to the linear characters. *)
+Lemma lin_char_group G :
+ {linG : finGroupType & {cF : linG -> 'CF(G) |
+ [/\ injective cF, #|linG| = #|G : G^`(1)|,
+ forall u, cF u \is a linear_char
+ & forall phi, phi \is a linear_char -> exists u, phi = cF u]
+ & [/\ cF 1%g = 1%R,
+ {morph cF : u v / (u * v)%g >-> (u * v)%R},
+ forall k, {morph cF : u / (u^+ k)%g >-> u ^+ k},
+ {morph cF: u / u^-1%g >-> u^-1%CF}
+ & {mono cF: u / #[u]%g >-> #[u]%CF} ]}}.
+Proof.
+pose linT := {i : Iirr G | 'chi_i \is a linear_char}.
+pose cF (u : linT) := 'chi_(sval u).
+have cFlin u: cF u \is a linear_char := svalP u.
+have cFinj: injective cF := inj_comp irr_inj val_inj.
+have inT xi : xi \is a linear_char -> {u | cF u = xi}.
+ move=> lin_xi; have /irrP/sig_eqW[i Dxi] := lin_char_irr lin_xi.
+ by apply: (exist _ (Sub i _)) => //; rewrite -Dxi.
+have [one cFone] := inT 1 (rpred1 _).
+pose inv u := sval (inT _ (rpredVr (cFlin u))).
+pose mul u v := sval (inT _ (rpredM (cFlin u) (cFlin v))).
+have cFmul u v: cF (mul u v) = cF u * cF v := svalP (inT _ _).
+have cFinv u: cF (inv u) = (cF u)^-1 := svalP (inT _ _).
+have mulA: associative mul by move=> u v w; apply: cFinj; rewrite !cFmul mulrA.
+have mul1: left_id one mul by move=> u; apply: cFinj; rewrite cFmul cFone mul1r.
+have mulV: left_inverse one inv mul.
+ by move=> u; apply: cFinj; rewrite cFmul cFinv cFone mulVr ?lin_char_unitr.
+pose linGm := FinGroup.Mixin mulA mul1 mulV.
+pose linG := @FinGroupType (BaseFinGroupType linT linGm) mulV.
+have cFexp k: {morph cF : u / ((u : linG) ^+ k)%g >-> u ^+ k}.
+ by move=> u; elim: k => // k IHk; rewrite expgS exprS cFmul IHk.
+do [exists linG, cF; split=> //] => [|xi /inT[u <-]|u]; first 2 [by exists u].
+ have inj_cFI: injective (cfIirr \o cF).
+ apply: can_inj (insubd one) _ => u; apply: val_inj.
+ by rewrite insubdK /= ?irrK //; apply: cFlin.
+ rewrite -(card_image inj_cFI) -card_lin_irr.
+ apply/eq_card=> i; rewrite inE; apply/codomP/idP=> [[u ->] | /inT[u Du]].
+ by rewrite /= irrK; apply: cFlin.
+ by exists u; apply: irr_inj; rewrite /= irrK.
+apply/eqP; rewrite eqn_dvd; apply/andP; split.
+ by rewrite dvdn_cforder; rewrite -cFexp expg_order cFone.
+by rewrite order_dvdn -(inj_eq cFinj) cFone cFexp exp_cforder.
+Qed.
+
+Lemma cfExp_prime_transitive G (i j : Iirr G) :
+ prime #|G| -> i != 0 -> j != 0 ->
+ exists2 k, coprime k #['chi_i]%CF & 'chi_j = 'chi_i ^+ k.
+Proof.
+set p := #|G| => pr_p nz_i nz_j; have cycG := prime_cyclic pr_p.
+have [L [h [injh oL Lh h_ontoL]] [h1 hM hX _ o_h]] := lin_char_group G.
+rewrite (derG1P (cyclic_abelian cycG)) indexg1 -/p in oL.
+have /fin_all_exists[h' h'K] := h_ontoL _ (irr_cyclic_lin _ cycG).
+have o_h' k: k != 0 -> #[h' k] = p.
+ rewrite -cforder_irr_eq1 h'K -o_h => nt_h'k.
+ by apply/prime_nt_dvdP=> //; rewrite cforder_lin_char_dvdG.
+have{oL} genL k: k != 0 -> generator [set: L] (h' k).
+ move=> /o_h' o_h'k; rewrite /generator eq_sym eqEcard subsetT /=.
+ by rewrite cardsT oL -o_h'k.
+have [/(_ =P <[_]>)-> gen_j] := (genL i nz_i, genL j nz_j).
+have /cycleP[k Dj] := cycle_generator gen_j.
+by rewrite !h'K Dj o_h hX generator_coprime coprime_sym in gen_j *; exists k.
+Qed.
+
+(* This is Isaacs (2.24). *)
+Lemma card_subcent1_coset G H x :
+ x \in G -> H <| G -> (#|'C_(G / H)[coset H x]| <= #|'C_G[x]|)%N.
+Proof.
+move=> Gx nsHG; rewrite -leC_nat.
+move: (second_orthogonality_relation x Gx); rewrite mulrb class_refl => <-.
+have GHx: coset H x \in (G / H)%g by apply: mem_quotient.
+move: (second_orthogonality_relation (coset H x) GHx).
+rewrite mulrb class_refl => <-.
+rewrite -2!(eq_bigr _ (fun _ _ => normCK _)) sum_norm_irr_quo // -subr_ge0.
+rewrite (bigID (fun i => H \subset cfker 'chi[G]_i)) //= addrC addKr.
+by apply: sumr_ge0 => i _; rewrite normCK mul_conjC_ge0.
+Qed.
+
+End Derive.
+
+Implicit Arguments irr_prime_injP [gT G i].
+
+(* Determinant characters and determinential order. *)
+Section DetOrder.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Section DetRepr.
+
+Variables (n : nat) (rG : mx_representation [fieldType of algC] G n).
+
+Definition det_repr_mx x : 'M_1 := (\det (rG x))%:M.
+
+Fact det_is_repr : mx_repr G det_repr_mx.
+Proof.
+split=> [|g h Gg Gh]; first by rewrite /det_repr_mx repr_mx1 det1.
+by rewrite /det_repr_mx repr_mxM // det_mulmx !mulmxE scalar_mxM.
+Qed.
+
+Canonical det_repr := MxRepresentation det_is_repr.
+Definition detRepr := cfRepr det_repr.
+
+Lemma detRepr_lin_char : detRepr \is a linear_char.
+Proof.
+by rewrite qualifE cfRepr_char cfunE group1 repr_mx1 mxtrace1 mulr1n /=.
+Qed.
+
+End DetRepr.
+
+Definition cfDet phi := \prod_i detRepr 'Chi_i ^+ truncC '[phi, 'chi[G]_i].
+
+Lemma cfDet_lin_char phi : cfDet phi \is a linear_char.
+Proof. by apply: rpred_prod => i _; apply: rpredX; apply: detRepr_lin_char. Qed.
+
+Lemma cfDetD :
+ {in character &, {morph cfDet : phi psi / phi + psi >-> phi * psi}}.
+Proof.
+move=> phi psi Nphi Npsi; rewrite /= -big_split; apply: eq_bigr => i _ /=.
+by rewrite -exprD cfdotDl truncCD ?nnegrE ?Cnat_ge0 // Cnat_cfdot_char_irr.
+Qed.
+
+Lemma cfDet0 : cfDet 0 = 1.
+Proof. by rewrite /cfDet big1 // => i _; rewrite cfdot0l truncC0. Qed.
+
+Lemma cfDetMn k :
+ {in character, {morph cfDet : phi / phi *+ k >-> phi ^+ k}}.
+Proof.
+move=> phi Nphi; elim: k => [|k IHk]; rewrite ?cfDet0 // mulrS exprS -{}IHk.
+by rewrite cfDetD ?rpredMn.
+Qed.
+
+Lemma cfDetRepr n rG : cfDet (cfRepr rG) = @detRepr n rG.
+Proof.
+transitivity (\prod_W detRepr (socle_repr W) ^+ standard_irr_coef rG W).
+ rewrite (reindex _ (socle_of_Iirr_bij _)) /cfDet /=.
+ apply: eq_bigr => i _; congr (_ ^+ _).
+ rewrite (cfRepr_sim (mx_rsim_standard rG)) cfRepr_standard.
+ rewrite cfdot_suml (bigD1 i) ?big1 //= => [|j i'j]; last first.
+ by rewrite cfdotZl cfdot_irr (negPf i'j) mulr0.
+ by rewrite cfdotZl cfnorm_irr mulr1 addr0 natCK.
+apply/cfun_inP=> x Gx; rewrite prod_cfunE //.
+transitivity (detRepr (standard_grepr rG) x); last first.
+ rewrite !cfunE Gx !trace_mx11 !mxE eqxx !mulrb.
+ case: (standard_grepr rG) (mx_rsim_standard rG) => /= n1 rG1 [B Dn1].
+ rewrite -{n1}Dn1 in rG1 B *; rewrite row_free_unit => uB rG_B.
+ by rewrite -[rG x](mulmxK uB) rG_B // !det_mulmx mulrC -!det_mulmx mulKmx.
+rewrite /standard_grepr; elim/big_rec2: _ => [|W y _ _ ->].
+ by rewrite cfunE trace_mx11 mxE Gx det1.
+rewrite !cfunE Gx /= !{1}trace_mx11 !{1}mxE det_ublock; congr (_ * _).
+rewrite exp_cfunE //; elim: (standard_irr_coef rG W) => /= [|k IHk].
+ by rewrite /muln_grepr big_ord0 det1.
+rewrite exprS /muln_grepr big_ord_recl det_ublock -IHk; congr (_ * _).
+by rewrite cfunE trace_mx11 mxE Gx.
+Qed.
+
+Lemma cfDet_id xi : xi \is a linear_char -> cfDet xi = xi.
+Proof.
+move=> lin_xi; have /irrP[i Dxi] := lin_char_irr lin_xi.
+apply/cfun_inP=> x Gx; rewrite Dxi -irrRepr cfDetRepr !cfunE trace_mx11 mxE.
+move: lin_xi (_ x) => /andP[_]; rewrite Dxi irr1_degree pnatr_eq1 => /eqP-> X.
+by rewrite {1}[X]mx11_scalar det_scalar1 trace_mx11.
+Qed.
+
+Definition cfDet_order phi := #[cfDet phi]%CF.
+
+Definition cfDet_order_lin xi :
+ xi \is a linear_char -> cfDet_order xi = #[xi]%CF.
+Proof. by rewrite /cfDet_order => /cfDet_id->. Qed.
+
+Definition cfDet_order_dvdG phi : cfDet_order phi %| #|G|.
+Proof. by rewrite cforder_lin_char_dvdG ?cfDet_lin_char. Qed.
+
+End DetOrder.
+
+Notation "''o' ( phi )" := (cfDet_order phi)
+ (at level 8, format "''o' ( phi )") : cfun_scope.
+
+Section CfDetOps.
+
+Implicit Types gT aT rT : finGroupType.
+
+Lemma cfDetRes gT (G H : {group gT}) phi :
+ phi \is a character -> cfDet ('Res[H, G] phi) = 'Res (cfDet phi).
+Proof.
+move=> Nphi; have [sGH | not_sHG] := boolP (H \subset G); last first.
+ have /CnatP[n Dphi1] := Cnat_char1 Nphi.
+ rewrite !cfResEout // Dphi1 lin_char1 ?cfDet_lin_char // scale1r.
+ by rewrite scaler_nat cfDetMn ?cfDet_id ?rpred1 // expr1n.
+have [rG ->] := char_reprP Nphi; rewrite !(=^~ cfRepr_sub, cfDetRepr) //.
+apply: cfRepr_sim; exists 1%:M; rewrite ?row_free_unit ?unitmx1 // => x Hx.
+by rewrite mulmx1 mul1mx.
+Qed.
+
+Lemma cfDetMorph aT rT (D G : {group aT}) (f : {morphism D >-> rT})
+ (phi : 'CF(f @* G)) :
+ phi \is a character -> cfDet (cfMorph phi) = cfMorph (cfDet phi).
+Proof.
+move=> Nphi; have [sGD | not_sGD] := boolP (G \subset D); last first.
+ have /CnatP[n Dphi1] := Cnat_char1 Nphi.
+ rewrite !cfMorphEout // Dphi1 lin_char1 ?cfDet_lin_char // scale1r.
+ by rewrite scaler_nat cfDetMn ?cfDet_id ?rpred1 // expr1n.
+have [rG ->] := char_reprP Nphi; rewrite !(=^~ cfRepr_morphim, cfDetRepr) //.
+apply: cfRepr_sim; exists 1%:M; rewrite ?row_free_unit ?unitmx1 // => x Hx.
+by rewrite mulmx1 mul1mx.
+Qed.
+
+Lemma cfDetIsom aT rT (G : {group aT}) (R : {group rT})
+ (f : {morphism G >-> rT}) (isoGR : isom G R f) phi :
+ cfDet (cfIsom isoGR phi) = cfIsom isoGR (cfDet phi).
+Proof.
+rewrite rmorph_prod /cfDet (reindex (isom_Iirr isoGR)); last first.
+ by exists (isom_Iirr (isom_sym isoGR)) => i; rewrite ?isom_IirrK ?isom_IirrKV.
+apply: eq_bigr => i; rewrite -!cfDetRepr !irrRepr isom_IirrE rmorphX cfIsom_iso.
+by rewrite /= ![in cfIsom _]unlock cfDetMorph ?cfRes_char ?cfDetRes ?irr_char.
+Qed.
+
+Lemma cfDet_mul_lin gT (G : {group gT}) (lambda phi : 'CF(G)) :
+ lambda \is a linear_char -> phi \is a character ->
+ cfDet (lambda * phi) = lambda ^+ truncC (phi 1%g) * cfDet phi.
+Proof.
+case/andP=> /char_reprP[[n1 rG1] ->] /= n1_1 /char_reprP[[n2 rG2] ->] /=.
+do [rewrite !cfRepr1 pnatr_eq1 natCK; move/eqP] in n1_1 *.
+rewrite {n1}n1_1 in rG1 *; rewrite cfRepr_prod cfDetRepr.
+apply/cfun_inP=> x Gx; rewrite !cfunE cfDetRepr cfunE Gx !mulrb !trace_mx11.
+rewrite !mxE prod_repr_lin ?mulrb //=; case: _ / (esym _); rewrite detZ.
+congr (_ * _); case: {rG2}n2 => [|n2]; first by rewrite cfun1E Gx.
+by rewrite expS_cfunE //= cfunE Gx trace_mx11.
+Qed.
+
+End CfDetOps.
+
+Definition cfcenter (gT : finGroupType) (G : {set gT}) (phi : 'CF(G)) :=
+ if phi \is a character then [set g in G | `|phi g| == phi 1%g] else cfker phi.
+
+Notation "''Z' ( phi )" := (cfcenter phi) : cfun_scope.
+
+Section Center.
+
+Variable (gT : finGroupType) (G : {group gT}).
+Implicit Types (phi chi : 'CF(G)) (H : {group gT}).
+
+(* This is Isaacs (2.27)(a). *)
+Lemma cfcenter_repr n (rG : mx_representation algCF G n) :
+ 'Z(cfRepr rG)%CF = rcenter rG.
+Proof.
+rewrite /cfcenter /rcenter cfRepr_char /=.
+apply/setP=> x; rewrite !inE; apply/andb_id2l=> Gx.
+apply/eqP/is_scalar_mxP=> [|[c rG_c]].
+ by case/max_cfRepr_norm_scalar=> // c; exists c.
+rewrite -(sqrCK (char1_ge0 (cfRepr_char rG))) normC_def; congr (sqrtC _).
+rewrite expr2 -{2}(mulgV x) -char_inv ?cfRepr_char ?cfunE ?groupM ?groupV //.
+rewrite Gx group1 repr_mx1 repr_mxM ?repr_mxV ?groupV // !mulrb rG_c.
+by rewrite invmx_scalar -scalar_mxM !mxtrace_scalar mulrnAr mulrnAl mulr_natl.
+Qed.
+
+(* This is part of Isaacs (2.27)(b). *)
+Fact cfcenter_group_set phi : group_set ('Z(phi))%CF.
+Proof.
+have [[rG ->] | /negbTE notNphi] := altP (@char_reprP _ G phi).
+ by rewrite cfcenter_repr groupP.
+by rewrite /cfcenter notNphi groupP.
+Qed.
+Canonical cfcenter_group f := Group (cfcenter_group_set f).
+
+Lemma char_cfcenterE chi x :
+ chi \is a character -> x \in G ->
+ (x \in ('Z(chi))%CF) = (`|chi x| == chi 1%g).
+Proof. by move=> Nchi Gx; rewrite /cfcenter Nchi inE Gx. Qed.
+
+Lemma irr_cfcenterE i x :
+ x \in G -> (x \in 'Z('chi[G]_i)%CF) = (`|'chi_i x| == 'chi_i 1%g).
+Proof. by move/char_cfcenterE->; rewrite ?irr_char. Qed.
+
+(* This is also Isaacs (2.27)(b). *)
+Lemma cfcenter_sub phi : ('Z(phi))%CF \subset G.
+Proof. by rewrite /cfcenter /cfker !setIdE -fun_if subsetIl. Qed.
+
+Lemma cfker_center_normal phi : cfker phi <| 'Z(phi)%CF.
+Proof.
+apply: normalS (cfcenter_sub phi) (cfker_normal phi).
+rewrite /= /cfcenter; case: ifP => // Hphi; rewrite cfkerEchar //.
+apply/subsetP=> x; rewrite !inE => /andP[-> /eqP->] /=.
+by rewrite ger0_norm ?char1_ge0.
+Qed.
+
+Lemma cfcenter_normal phi : 'Z(phi)%CF <| G.
+Proof.
+have [[rG ->] | /negbTE notNphi] := altP (@char_reprP _ _ phi).
+ by rewrite cfcenter_repr rcenter_normal.
+by rewrite /cfcenter notNphi cfker_normal.
+Qed.
+
+(* This is Isaacs (2.27)(c). *)
+Lemma cfcenter_Res chi :
+ exists2 chi1, chi1 \is a linear_char & 'Res['Z(chi)%CF] chi = chi 1%g *: chi1.
+Proof.
+have [[rG ->] | /negbTE notNphi] := altP (@char_reprP _ _ chi); last first.
+ exists 1; first exact: cfun1_lin_char.
+ rewrite /cfcenter notNphi; apply/cfun_inP=> x Kx.
+ by rewrite cfunE cfun1E Kx mulr1 cfResE ?cfker_sub // cfker1.
+rewrite cfcenter_repr -(cfRepr_sub _ (normal_sub (rcenter_normal _))).
+case: rG => [[|n] rG] /=; rewrite cfRepr1.
+ exists 1; first exact: cfun1_lin_char.
+ by apply/cfun_inP=> x Zx; rewrite scale0r !cfunE flatmx0 raddf0 Zx.
+pose rZmx x := ((rG x 0 0)%:M : 'M_(1,1)).
+have rZmxP: mx_repr [group of rcenter rG] rZmx.
+ split=> [|x y]; first by rewrite /rZmx repr_mx1 mxE eqxx.
+ move=> /setIdP[Gx /is_scalar_mxP[a rGx]] /setIdP[Gy /is_scalar_mxP[b rGy]].
+ by rewrite /rZmx repr_mxM // rGx rGy -!scalar_mxM !mxE.
+exists (cfRepr (MxRepresentation rZmxP)).
+ by rewrite qualifE cfRepr_char cfRepr1 eqxx.
+apply/cfun_inP=> x Zx; rewrite !cfunE Zx /= /rZmx mulr_natl.
+by case/setIdP: Zx => Gx /is_scalar_mxP[a ->]; rewrite mxE !mxtrace_scalar.
+Qed.
+
+(* This is Isaacs (2.27)(d). *)
+Lemma cfcenter_cyclic chi : cyclic ('Z(chi)%CF / cfker chi)%g.
+Proof.
+case Nchi: (chi \is a character); last first.
+ by rewrite /cfcenter Nchi trivg_quotient cyclic1.
+have [-> | nz_chi] := eqVneq chi 0.
+ rewrite quotientS1 ?cyclic1 //= /cfcenter cfkerEchar ?cfun0_char //.
+ by apply/subsetP=> x /setIdP[Gx _]; rewrite inE Gx /= !cfunE.
+have [xi Lxi def_chi] := cfcenter_Res chi.
+set Z := ('Z(_))%CF in xi Lxi def_chi *.
+have sZG: Z \subset G by exact: cfcenter_sub.
+have ->: cfker chi = cfker xi.
+ rewrite -(setIidPr (normal_sub (cfker_center_normal _))) -/Z.
+ rewrite !cfkerEchar // ?lin_charW //= -/Z.
+ apply/setP=> x; rewrite !inE; apply: andb_id2l => Zx.
+ rewrite (subsetP sZG) //= -!(cfResE chi sZG) ?group1 // def_chi !cfunE.
+ by rewrite (inj_eq (mulfI _)) ?char1_eq0.
+have: abelian (Z / cfker xi) by rewrite sub_der1_abelian ?lin_char_der1.
+have [rG irrG ->] := irr_reprP _ (lin_char_irr Lxi); rewrite cfker_repr.
+apply: mx_faithful_irr_abelian_cyclic (kquo_mx_faithful rG) _.
+exact/quo_mx_irr.
+Qed.
+
+(* This is Isaacs (2.27)(e). *)
+Lemma cfcenter_subset_center chi :
+ ('Z(chi)%CF / cfker chi)%g \subset 'Z(G / cfker chi)%g.
+Proof.
+case Nchi: (chi \is a character); last first.
+ by rewrite /cfcenter Nchi trivg_quotient sub1G.
+rewrite subsetI quotientS ?cfcenter_sub // quotient_cents2r //=.
+case/char_reprP: Nchi => rG ->{chi}; rewrite cfker_repr cfcenter_repr gen_subG.
+apply/subsetP=> _ /imset2P[x y /setIdP[Gx /is_scalar_mxP[c rGx]] Gy ->].
+rewrite inE groupR //= !repr_mxM ?groupM ?groupV // rGx -(scalar_mxC c) -rGx.
+by rewrite !mulmxA !repr_mxKV.
+Qed.
+
+(* This is Isaacs (2.27)(f). *)
+Lemma cfcenter_eq_center (i : Iirr G) :
+ ('Z('chi_i)%CF / cfker 'chi_i)%g = 'Z(G / cfker 'chi_i)%g.
+Proof.
+apply/eqP; rewrite eqEsubset; rewrite cfcenter_subset_center ?irr_char //.
+apply/subsetP=> _ /setIP[/morphimP[x /= _ Gx ->] cGx]; rewrite mem_quotient //=.
+rewrite -irrRepr cfker_repr cfcenter_repr inE Gx in cGx *.
+apply: mx_abs_irr_cent_scalar 'Chi_i _ _ _.
+ by apply: groupC; apply: socle_irr.
+have nKG: G \subset 'N(rker 'Chi_i) by exact: rker_norm.
+(* GG -- locking here is critical to prevent Coq kernel divergence. *)
+apply/centgmxP=> y Gy; rewrite [eq]lock -2?(quo_repr_coset (subxx _) nKG) //.
+move: (quo_repr _ _) => rG; rewrite -2?repr_mxM ?mem_quotient // -lock.
+by rewrite (centP cGx) // mem_quotient.
+Qed.
+
+(* This is Isaacs (2.28). *)
+Lemma cap_cfcenter_irr : \bigcap_i 'Z('chi[G]_i)%CF = 'Z(G).
+Proof.
+apply/esym/eqP; rewrite eqEsubset (introT bigcapsP) /= => [|i _]; last first.
+ rewrite -(quotientSGK _ (normal_sub (cfker_center_normal _))).
+ by rewrite cfcenter_eq_center morphim_center.
+ by rewrite subIset // normal_norm // cfker_normal.
+set Z := \bigcap_i _.
+have sZG: Z \subset G by rewrite (bigcap_min 0) ?cfcenter_sub.
+rewrite subsetI sZG (sameP commG1P trivgP) -(TI_cfker_irr G).
+apply/bigcapsP=> i _; have nKiG := normal_norm (cfker_normal 'chi_i).
+rewrite -quotient_cents2 ?(subset_trans sZG) //.
+rewrite (subset_trans (quotientS _ (bigcap_inf i _))) //.
+by rewrite cfcenter_eq_center subsetIr.
+Qed.
+
+(* This is Isaacs (2.29). *)
+Lemma cfnorm_Res_lerif H phi :
+ H \subset G ->
+ '['Res[H] phi] <= #|G : H|%:R * '[phi] ?= iff (phi \in 'CF(G, H)).
+Proof.
+move=> sHG; rewrite cfun_onE mulrCA natf_indexg // -mulrA mulKf ?neq0CG //.
+rewrite (big_setID H) (setIidPr sHG) /= addrC.
+rewrite (mono_lerif (ler_pmul2l _)) ?invr_gt0 ?gt0CG // -lerif_subLR -sumrB.
+rewrite big1 => [|x Hx]; last by rewrite !cfResE ?subrr.
+have ->: (support phi \subset H) = (G :\: H \subset [set x | phi x == 0]).
+ rewrite subDset setUC -subDset; apply: eq_subset => x.
+ by rewrite !inE (andb_idr (contraR _)) // => /cfun0->.
+rewrite (sameP subsetP forall_inP); apply: lerif_0_sum => x _.
+by rewrite !inE /<?=%R mul_conjC_ge0 eq_sym mul_conjC_eq0.
+Qed.
+
+(* This is Isaacs (2.30). *)
+Lemma irr1_bound (i : Iirr G) :
+ ('chi_i 1%g) ^+ 2 <= #|G : 'Z('chi_i)%CF|%:R
+ ?= iff ('chi_i \in 'CF(G, 'Z('chi_i)%CF)).
+Proof.
+congr (_ <= _ ?= iff _): (cfnorm_Res_lerif 'chi_i (cfcenter_sub 'chi_i)).
+ have [xi Lxi ->] := cfcenter_Res 'chi_i.
+ have /irrP[j ->] := lin_char_irr Lxi; rewrite cfdotZl cfdotZr cfdot_irr eqxx.
+ by rewrite mulr1 irr1_degree conjC_nat.
+by rewrite cfdot_irr eqxx mulr1.
+Qed.
+
+(* This is Isaacs (2.31). *)
+Lemma irr1_abelian_bound (i : Iirr G) :
+ abelian (G / 'Z('chi_i)%CF) -> ('chi_i 1%g) ^+ 2 = #|G : 'Z('chi_i)%CF|%:R.
+Proof.
+move=> AbGc; apply/eqP; rewrite irr1_bound cfun_onE; apply/subsetP=> x nz_chi_x.
+have Gx: x \in G by apply: contraR nz_chi_x => /cfun0->.
+have nKx := subsetP (normal_norm (cfker_normal 'chi_i)) _ Gx.
+rewrite -(quotientGK (cfker_center_normal _)) inE nKx inE /=.
+rewrite cfcenter_eq_center inE mem_quotient //=.
+apply/centP=> _ /morphimP[y nKy Gy ->]; apply/commgP; rewrite -morphR //=.
+set z := [~ x, y]; rewrite coset_id //.
+have: z \in 'Z('chi_i)%CF.
+ apply: subsetP (mem_commg Gx Gy).
+ by rewrite der1_min // normal_norm ?cfcenter_normal.
+rewrite -irrRepr cfker_repr cfcenter_repr !inE in nz_chi_x *.
+case/andP=> Gz /is_scalar_mxP[c Chi_z]; rewrite Gz Chi_z mul1mx /=.
+apply/eqP; congr _%:M; apply: (mulIf nz_chi_x); rewrite mul1r.
+rewrite -{2}(cfunJ _ x Gy) conjg_mulR -/z !cfunE Gx groupM // !{1}mulrb.
+by rewrite repr_mxM // Chi_z mul_mx_scalar mxtraceZ.
+Qed.
+
+(* This is Isaacs (2.32)(a). *)
+Lemma irr_faithful_center i : cfaithful 'chi[G]_i -> cyclic 'Z(G).
+Proof.
+rewrite (isog_cyclic (isog_center (quotient1_isog G))) /=.
+by move/trivgP <-; rewrite -cfcenter_eq_center cfcenter_cyclic.
+Qed.
+
+Lemma cfcenter_fful_irr i : cfaithful 'chi[G]_i -> 'Z('chi_i)%CF = 'Z(G).
+Proof.
+move/trivgP=> Ki1; have:= cfcenter_eq_center i; rewrite {}Ki1.
+have inj1: 'injm (@coset gT 1%g) by rewrite ker_coset.
+by rewrite -injm_center; first apply: injm_morphim_inj; rewrite ?norms1.
+Qed.
+
+(* This is Isaacs (2.32)(b). *)
+Lemma pgroup_cyclic_faithful (p : nat) :
+ p.-group G -> cyclic 'Z(G) -> exists i, cfaithful 'chi[G]_i.
+Proof.
+pose Z := 'Ohm_1('Z(G)) => pG cycZG; have nilG := pgroup_nil pG.
+have [-> | ntG] := eqsVneq G [1]; first by exists 0; exact: cfker_sub.
+have{pG} [[p_pr _ _] pZ] := (pgroup_pdiv pG ntG, pgroupS (center_sub G) pG).
+have ntZ: 'Z(G) != [1] by rewrite center_nil_eq1.
+have{pZ} oZ: #|Z| = p by exact: Ohm1_cyclic_pgroup_prime.
+apply/existsP; apply: contraR ntZ; rewrite negb_exists => /forallP-not_ffulG.
+rewrite -Ohm1_eq1 -subG1 /= -/Z -(TI_cfker_irr G); apply/bigcapsP=> i _.
+rewrite prime_meetG ?oZ // setIC meet_Ohm1 // meet_center_nil ?cfker_normal //.
+by rewrite -subG1 not_ffulG.
+Qed.
+
+End Center.
+
+Section Induced.
+
+Variables (gT : finGroupType) (G H : {group gT}).
+Implicit Types (phi : 'CF(G)) (chi : 'CF(H)).
+
+Lemma cfInd_char chi : chi \is a character -> 'Ind[G] chi \is a character.
+Proof.
+move=> Nchi; apply/forallP=> i; rewrite coord_cfdot -Frobenius_reciprocity //.
+by rewrite Cnat_cfdot_char ?cfRes_char ?irr_char.
+Qed.
+
+Lemma cfInd_eq0 chi :
+ H \subset G -> chi \is a character -> ('Ind[G] chi == 0) = (chi == 0).
+Proof.
+move=> sHG Nchi; rewrite -!(char1_eq0) ?cfInd_char // cfInd1 //.
+by rewrite (mulrI_eq0 _ (mulfI _)) ?neq0CiG.
+Qed.
+
+Lemma Ind_irr_neq0 i : H \subset G -> 'Ind[G, H] 'chi_i != 0.
+Proof. by move/cfInd_eq0->; rewrite ?irr_neq0 ?irr_char. Qed.
+
+Definition Ind_Iirr (A B : {set gT}) i := cfIirr ('Ind[B, A] 'chi_i).
+
+Lemma constt_cfRes_irr i : {j | j \in irr_constt ('Res[H, G] 'chi_i)}.
+Proof. apply/sigW/neq0_has_constt/Res_irr_neq0. Qed.
+
+Lemma constt_cfInd_irr i :
+ H \subset G -> {j | j \in irr_constt ('Ind[G, H] 'chi_i)}.
+Proof. by move=> sHG; apply/sigW/neq0_has_constt/Ind_irr_neq0. Qed.
+
+Lemma cfker_Res phi :
+ H \subset G -> phi \is a character -> cfker ('Res[H] phi) = H :&: cfker phi.
+Proof.
+move=> sHG Nphi; apply/setP=> x; rewrite !cfkerEchar ?cfRes_char // !inE.
+by apply/andb_id2l=> Hx; rewrite (subsetP sHG) ?cfResE.
+Qed.
+
+(* This is Isaacs Lemma (5.11). *)
+Lemma cfker_Ind chi :
+ H \subset G -> chi \is a character -> chi != 0 ->
+ cfker ('Ind[G, H] chi) = gcore (cfker chi) G.
+Proof.
+move=> sHG Nchi nzchi; rewrite !cfker_nzcharE ?cfInd_char ?cfInd_eq0 //.
+apply/setP=> x; rewrite inE cfIndE // (can2_eq (mulVKf _) (mulKf _)) ?neq0CG //.
+rewrite cfInd1 // mulrA -natrM Lagrange // mulr_natl -sumr_const.
+apply/eqP/bigcapP=> [/normC_sum_upper ker_chiG_x y Gy | ker_chiG_x].
+ by rewrite mem_conjg inE ker_chiG_x ?groupV // => z _; exact: char1_ge_norm.
+by apply: eq_bigr => y /groupVr/ker_chiG_x; rewrite mem_conjgV inE => /eqP.
+Qed.
+
+Lemma cfker_Ind_irr i :
+ H \subset G -> cfker ('Ind[G, H] 'chi_i) = gcore (cfker 'chi_i) G.
+Proof. by move/cfker_Ind->; rewrite ?irr_neq0 ?irr_char. Qed.
+
+End Induced.
+
+Arguments Scope Ind_Iirr [_ group_scope group_scope ring_scope]. \ No newline at end of file
diff --git a/mathcomp/character/classfun.v b/mathcomp/character/classfun.v
new file mode 100644
index 0000000..16d044b
--- /dev/null
+++ b/mathcomp/character/classfun.v
@@ -0,0 +1,2463 @@
+(* (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.
+Require Import fingroup morphism perm automorphism quotient finalg action.
+Require Import gproduct zmodp commutator cyclic center pgroup sylow.
+Require Import matrix vector falgebra ssrnum algC algnum.
+
+(******************************************************************************)
+(* This file contains the basic theory of class functions: *)
+(* 'CF(G) == the type of class functions on G : {group gT}, i.e., *)
+(* which map gT to the type algC of complex algebraics, *)
+(* have support in G, and are constant on each conjugacy *)
+(* class of G. 'CF(G) implements the algebraType interface *)
+(* of finite-dimensional F-algebras. *)
+(* The identity 1 : 'CF(G) is the indicator function of G, *)
+(* and (later) the principal character. *)
+(* --> The %CF scope (cfun_scope) is bound to the 'CF(_) types. *)
+(* 'CF(G)%VS == the (total) vector space of 'CF(G). *)
+(* 'CF(G, A) == the subspace of functions in 'CF(G) with support in A. *)
+(* phi x == the image of x : gT under phi : 'CF(G). *)
+(* #[phi]%CF == the multiplicative order of phi : 'CF(G). *)
+(* cfker phi == the kernel of phi : 'CF(G); note that cfker phi <| G. *)
+(* cfaithful phi <=> phi : 'CF(G) is faithful (has a trivial kernel). *)
+(* '1_A == the indicator function of A as a function of 'CF(G). *)
+(* (Provided A <| G; G is determined by the context.) *)
+(* phi^*%CF == the function conjugate to phi : 'CF(G). *)
+(* cfAut u phi == the function conjugate to phi by an algC-automorphism u *)
+(* phi^u The notation "_ ^u" is only reserved; it is up to *)
+(* clients to set Notation "phi ^u" := (cfAut u phi). *)
+(* '[phi, psi] == the convolution of phi, psi : 'CF(G) over G, normalised *)
+(* '[phi, psi]_G by #|G| so that '[1, 1]_G = 1 (G is usually inferred). *)
+(* cfdotr psi phi == '[phi, psi] (self-expanding). *)
+(* '[phi], '[phi]_G == the squared norm '[phi, phi] of phi : 'CF(G). *)
+(* orthogonal R S <=> each phi in R : seq 'CF(G) is orthogonal to each psi in *)
+(* S, i.e., '[phi, psi] = 0. As 'CF(G) coerces to seq, one *)
+(* can write orthogonal phi S and orthogonal phi psi. *)
+(* pairwise_orthogonal S <=> the class functions in S are pairwise orthogonal *)
+(* AND non-zero. *)
+(* orthonormal S <=> S is pairwise orthogonal and all class functions in S *)
+(* have norm 1. *)
+(* isometry tau <-> tau : 'CF(D) -> 'CF(R) is an isometry, mapping *)
+(* '[_, _]_D to '[_, _]_R. *)
+(* {in CD, isometry tau, to CR} <-> in the domain CD, tau is an isometry *)
+(* whose range is contained in CR. *)
+(* cfReal phi <=> phi is real, i.e., phi^* == phi. *)
+(* cfAut_closed u S <-> S : seq 'CF(G) is closed under conjugation by u. *)
+(* conjC_closed S <-> S : seq 'CF(G) is closed under complex conjugation. *)
+(* conjC_subset S1 S2 <-> S1 : seq 'CF(G) represents a subset of S2 closed *)
+(* under complex conjugation. *)
+(* := [/\ uniq S1, {subset S1 <= S2} & conjC_closed S1]. *)
+(* 'Res[H] phi == the restriction of phi : 'CF(G) to a function of 'CF(H) *)
+(* 'Res[H, G] phi 'Res[H] phi x = phi x if x \in H (when H \subset G), *)
+(* 'Res phi 'Res[H] phi x = 0 if x \notin H. The syntax variants *)
+(* allow H and G to be inferred; the default is to specify *)
+(* H explicitly, and infer G from the type of phi. *)
+(* 'Ind[G] phi == the class function of 'CF(G) induced by phi : 'CF(H), *)
+(* 'Ind[G, H] phi when H \subset G. As with 'Res phi, both G and H can *)
+(* 'Ind phi be inferred, though usually G isn't. *)
+(* cfMorph phi == the class function in 'CF(G) that maps x to phi (f x), *)
+(* where phi : 'CF(f @* G), provided G \subset 'dom f. *)
+(* cfIsom isoGR phi == the class function in 'CF(R) that maps f x to phi x, *)
+(* given isoGR : isom G R f, f : {morphism G >-> rT} and *)
+(* phi : 'CF(G). *)
+(* (phi %% H)%CF == special case of cfMorph phi, when phi : 'CF(G / H). *)
+(* (phi / H)%CF == the class function in 'CF(G / H) that coincides with *)
+(* phi : 'CF(G) on cosets of H \subset cfker phi. *)
+(* For a group G that is a semidirect product (defG : K ><| H = G), we have *)
+(* cfSdprod KxH phi == for phi : 'CF(H), the class function of 'CF(G) that *)
+(* maps k * h to psi h when k \in K and h \in H. *)
+(* For a group G that is a direct product (with KxH : K \x H = G), we have *)
+(* cfDprodl KxH phi == for phi : 'CF(K), the class function of 'CF(G) that *)
+(* maps k * h to phi k when k \in K and h \in H. *)
+(* cfDprodr KxH psi == for psi : 'CF(H), the class function of 'CF(G) that *)
+(* maps k * h to psi h when k \in K and h \in H. *)
+(* cfDprod KxH phi psi == for phi : 'CF(K), psi : 'CF(H), the class function *)
+(* of 'CF(G) that maps k * h to phi k * psi h (this is *)
+(* the product of the two functions above). *)
+(* Finally, given defG : \big[dprod/1]_(i | P i) A i = G, with G and A i *)
+(* groups and i ranges over a finType, we have *)
+(* cfBigdprodi defG phi == for phi : 'CF(A i) s.t. P i, the class function *)
+(* of 'CF(G) that maps x to phi x_i, where x_i is the *)
+(* (A i)-component of x : G. *)
+(* cfBigdprodi defG phi == for phi : forall i, 'CF(A i), the class function *)
+(* of 'CF(G) that maps x to \prod_(i | P i) phi i x_i, *)
+(* where x_i is the (A i)-component of x : G. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope GRing.Theory Num.Theory.
+Local Open Scope ring_scope.
+Delimit Scope cfun_scope with CF.
+
+Reserved Notation "''CF' ( G , A )" (at level 8, format "''CF' ( G , A )").
+Reserved Notation "''CF' ( G )" (at level 8, format "''CF' ( G )").
+Reserved Notation "''1_' G" (at level 8, G at level 2, format "''1_' G").
+Reserved Notation "''Res[' H , G ]" (at level 8, only parsing).
+Reserved Notation "''Res[' H ]" (at level 8, format "''Res[' H ]").
+Reserved Notation "''Res'" (at level 8, only parsing).
+Reserved Notation "''Ind[' G , H ]" (at level 8, only parsing).
+Reserved Notation "''Ind[' G ]" (at level 8, format "''Ind[' G ]").
+Reserved Notation "''Ind'" (at level 8, only parsing).
+Reserved Notation "'[ phi , psi ]_ G" (at level 2, only parsing).
+Reserved Notation "'[ phi , psi ]"
+ (at level 2, format "'[hv' ''[' phi , '/ ' psi ] ']'").
+Reserved Notation "'[ phi ]_ G" (at level 2, only parsing).
+Reserved Notation "'[ phi ]" (at level 2, format "''[' phi ]").
+Reserved Notation "phi ^u" (at level 3, format "phi ^u").
+
+Section AlgC.
+(* Arithmetic properties of group orders in the characteristic 0 field algC. *)
+
+Variable (gT : finGroupType).
+Implicit Types (G : {group gT}) (B : {set gT}).
+
+Lemma neq0CG G : (#|G|)%:R != 0 :> algC. Proof. exact: natrG_neq0. Qed.
+Lemma neq0CiG G B : (#|G : B|)%:R != 0 :> algC.
+Proof. exact: natr_indexg_neq0. Qed.
+Lemma gt0CG G : 0 < #|G|%:R :> algC. Proof. exact: natrG_gt0. Qed.
+Lemma gt0CiG G B : 0 < #|G : B|%:R :> algC. Proof. exact: natr_indexg_gt0. Qed.
+
+Lemma algC'G G : [char algC]^'.-group G.
+Proof. by apply/pgroupP=> p _; rewrite inE /= char_num. Qed.
+
+End AlgC.
+
+Section Defs.
+
+Variable gT : finGroupType.
+
+Definition is_class_fun (B : {set gT}) (f : {ffun gT -> algC}) :=
+ [forall x, forall y in B, f (x ^ y) == f x] && (support f \subset B).
+
+Lemma intro_class_fun (G : {group gT}) f :
+ {in G &, forall x y, f (x ^ y) = f x} ->
+ (forall x, x \notin G -> f x = 0) ->
+ is_class_fun G (finfun f).
+Proof.
+move=> fJ Gf; apply/andP; split; last first.
+ by apply/supportP=> x notAf; rewrite ffunE Gf.
+apply/'forall_eqfun_inP=> x y Gy; rewrite !ffunE.
+by have [/fJ-> // | notGx] := boolP (x \in G); rewrite !Gf ?groupJr.
+Qed.
+
+Variable B : {set gT}.
+Local Notation G := <<B>>.
+
+Record classfun : predArgType :=
+ Classfun {cfun_val; _ : is_class_fun G cfun_val}.
+Implicit Types phi psi xi : classfun.
+(* The default expansion lemma cfunE requires key = 0. *)
+Fact classfun_key : unit. Proof. by []. Qed.
+Definition Cfun := locked_with classfun_key (fun flag : nat => Classfun).
+
+Canonical cfun_subType := Eval hnf in [subType for cfun_val].
+Definition cfun_eqMixin := Eval hnf in [eqMixin of classfun by <:].
+Canonical cfun_eqType := Eval hnf in EqType classfun cfun_eqMixin.
+Definition cfun_choiceMixin := Eval hnf in [choiceMixin of classfun by <:].
+Canonical cfun_choiceType := Eval hnf in ChoiceType classfun cfun_choiceMixin.
+
+Definition fun_of_cfun phi := cfun_val phi : gT -> algC.
+Coercion fun_of_cfun : classfun >-> Funclass.
+
+Lemma cfunElock k f fP : @Cfun k (finfun f) fP =1 f.
+Proof. by rewrite locked_withE; apply: ffunE. Qed.
+
+Lemma cfunE f fP : @Cfun 0 (finfun f) fP =1 f.
+Proof. exact: cfunElock. Qed.
+
+Lemma cfunP phi psi : phi =1 psi <-> phi = psi.
+Proof. by split=> [/ffunP/val_inj | ->]. Qed.
+
+Lemma cfun0gen phi x : x \notin G -> phi x = 0.
+Proof. by case: phi => f fP; case: (andP fP) => _ /supportP; exact. Qed.
+
+Lemma cfun_in_genP phi psi : {in G, phi =1 psi} -> phi = psi.
+Proof.
+move=> eq_phi; apply/cfunP=> x.
+by have [/eq_phi-> // | notAx] := boolP (x \in G); rewrite !cfun0gen.
+Qed.
+
+Lemma cfunJgen phi x y : y \in G -> phi (x ^ y) = phi x.
+Proof.
+case: phi => f fP Gy; apply/eqP.
+by case: (andP fP) => /'forall_forall_inP->.
+Qed.
+
+Fact cfun_zero_subproof : is_class_fun G (0 : {ffun _}).
+Proof. exact: intro_class_fun. Qed.
+Definition cfun_zero := Cfun 0 cfun_zero_subproof.
+
+Fact cfun_comp_subproof f phi :
+ f 0 = 0 -> is_class_fun G [ffun x => f (phi x)].
+Proof.
+by move=> f0; apply: intro_class_fun => [x y _ /cfunJgen | x /cfun0gen] ->.
+Qed.
+Definition cfun_comp f f0 phi := Cfun 0 (@cfun_comp_subproof f phi f0).
+Definition cfun_opp := cfun_comp (oppr0 _).
+
+Fact cfun_add_subproof phi psi : is_class_fun G [ffun x => phi x + psi x].
+Proof.
+apply: intro_class_fun => [x y Gx Gy | x notGx]; rewrite ?cfunJgen //.
+by rewrite !cfun0gen ?add0r.
+Qed.
+Definition cfun_add phi psi := Cfun 0 (cfun_add_subproof phi psi).
+
+Fact cfun_indicator_subproof (A : {set gT}) :
+ is_class_fun G [ffun x => ((x \in G) && (x ^: G \subset A))%:R].
+Proof.
+apply: intro_class_fun => [x y Gx Gy | x /negbTE/= -> //].
+by rewrite groupJr ?classGidl.
+Qed.
+Definition cfun_indicator A := Cfun 1 (cfun_indicator_subproof A).
+Local Notation "''1_' A" := (cfun_indicator A) : ring_scope.
+
+Lemma cfun1Egen x : '1_G x = (x \in G)%:R.
+Proof. by rewrite cfunElock andb_idr // => /class_subG->. Qed.
+
+Fact cfun_mul_subproof phi psi : is_class_fun G [ffun x => phi x * psi x].
+Proof.
+apply: intro_class_fun => [x y Gx Gy | x notGx]; rewrite ?cfunJgen //.
+by rewrite cfun0gen ?mul0r.
+Qed.
+Definition cfun_mul phi psi := Cfun 0 (cfun_mul_subproof phi psi).
+
+Definition cfun_unit := [pred phi : classfun | [forall x in G, phi x != 0]].
+Definition cfun_inv phi :=
+ if phi \in cfun_unit then cfun_comp (invr0 _) phi else phi.
+
+Definition cfun_scale a := cfun_comp (mulr0 a).
+
+Fact cfun_addA : associative cfun_add.
+Proof. by move=> phi psi xi; apply/cfunP=> x; rewrite !cfunE addrA. Qed.
+Fact cfun_addC : commutative cfun_add.
+Proof. by move=> phi psi; apply/cfunP=> x; rewrite !cfunE addrC. Qed.
+Fact cfun_add0 : left_id cfun_zero cfun_add.
+Proof. by move=> phi; apply/cfunP=> x; rewrite !cfunE add0r. Qed.
+Fact cfun_addN : left_inverse cfun_zero cfun_opp cfun_add.
+Proof. by move=> phi; apply/cfunP=> x; rewrite !cfunE addNr. Qed.
+
+Definition cfun_zmodMixin := ZmodMixin cfun_addA cfun_addC cfun_add0 cfun_addN.
+Canonical cfun_zmodType := ZmodType classfun cfun_zmodMixin.
+
+Lemma muln_cfunE phi n x : (phi *+ n) x = phi x *+ n.
+Proof. by elim: n => [|n IHn]; rewrite ?mulrS !cfunE ?IHn. Qed.
+
+Lemma sum_cfunE I r (P : pred I) (phi : I -> classfun) x :
+ (\sum_(i <- r | P i) phi i) x = \sum_(i <- r | P i) (phi i) x.
+Proof. by elim/big_rec2: _ => [|i _ psi _ <-]; rewrite cfunE. Qed.
+
+Fact cfun_mulA : associative cfun_mul.
+Proof. by move=> phi psi xi; apply/cfunP=> x; rewrite !cfunE mulrA. Qed.
+Fact cfun_mulC : commutative cfun_mul.
+Proof. by move=> phi psi; apply/cfunP=> x; rewrite !cfunE mulrC. Qed.
+Fact cfun_mul1 : left_id '1_G cfun_mul.
+Proof.
+by move=> phi; apply: cfun_in_genP => x Gx; rewrite !cfunE cfun1Egen Gx mul1r.
+Qed.
+Fact cfun_mulD : left_distributive cfun_mul cfun_add.
+Proof. by move=> phi psi xi; apply/cfunP=> x; rewrite !cfunE mulrDl. Qed.
+Fact cfun_nz1 : '1_G != 0.
+Proof.
+by apply/eqP=> /cfunP/(_ 1%g)/eqP; rewrite cfun1Egen cfunE group1 oner_eq0.
+Qed.
+
+Definition cfun_ringMixin :=
+ ComRingMixin cfun_mulA cfun_mulC cfun_mul1 cfun_mulD cfun_nz1.
+Canonical cfun_ringType := RingType classfun cfun_ringMixin.
+Canonical cfun_comRingType := ComRingType classfun cfun_mulC.
+
+Lemma expS_cfunE phi n x : (phi ^+ n.+1) x = phi x ^+ n.+1.
+Proof. by elim: n => //= n IHn; rewrite !cfunE IHn. Qed.
+
+Fact cfun_mulV : {in cfun_unit, left_inverse 1 cfun_inv *%R}.
+Proof.
+move=> phi Uphi; rewrite /cfun_inv Uphi; apply/cfun_in_genP=> x Gx.
+by rewrite !cfunE cfun1Egen Gx mulVf ?(forall_inP Uphi).
+Qed.
+Fact cfun_unitP phi psi : psi * phi = 1 -> phi \in cfun_unit.
+Proof.
+move/cfunP=> phiK; apply/forall_inP=> x Gx; rewrite -unitfE; apply/unitrP.
+by exists (psi x); have:= phiK x; rewrite !cfunE cfun1Egen Gx mulrC.
+Qed.
+Fact cfun_inv0id : {in [predC cfun_unit], cfun_inv =1 id}.
+Proof. by rewrite /cfun_inv => phi /negbTE/= ->. Qed.
+
+Definition cfun_unitMixin := ComUnitRingMixin cfun_mulV cfun_unitP cfun_inv0id.
+Canonical cfun_unitRingType := UnitRingType classfun cfun_unitMixin.
+Canonical cfun_comUnitRingType := [comUnitRingType of classfun].
+
+Fact cfun_scaleA a b phi :
+ cfun_scale a (cfun_scale b phi) = cfun_scale (a * b) phi.
+Proof. by apply/cfunP=> x; rewrite !cfunE mulrA. Qed.
+Fact cfun_scale1 : left_id 1 cfun_scale.
+Proof. by move=> phi; apply/cfunP=> x; rewrite !cfunE mul1r. Qed.
+Fact cfun_scaleDr : right_distributive cfun_scale +%R.
+Proof. by move=> a phi psi; apply/cfunP=> x; rewrite !cfunE mulrDr. Qed.
+Fact cfun_scaleDl phi : {morph cfun_scale^~ phi : a b / a + b}.
+Proof. by move=> a b; apply/cfunP=> x; rewrite !cfunE mulrDl. Qed.
+
+Definition cfun_lmodMixin :=
+ LmodMixin cfun_scaleA cfun_scale1 cfun_scaleDr cfun_scaleDl.
+Canonical cfun_lmodType := LmodType algC classfun cfun_lmodMixin.
+
+Fact cfun_scaleAl a phi psi : a *: (phi * psi) = (a *: phi) * psi.
+Proof. by apply/cfunP=> x; rewrite !cfunE mulrA. Qed.
+Fact cfun_scaleAr a phi psi : a *: (phi * psi) = phi * (a *: psi).
+Proof. by rewrite !(mulrC phi) cfun_scaleAl. Qed.
+
+Canonical cfun_lalgType := LalgType algC classfun cfun_scaleAl.
+Canonical cfun_algType := AlgType algC classfun cfun_scaleAr.
+Canonical cfun_unitAlgType := [unitAlgType algC of classfun].
+
+Section Automorphism.
+
+Variable u : {rmorphism algC -> algC}.
+
+Definition cfAut := cfun_comp (rmorph0 u).
+
+Lemma cfAut_cfun1i A : cfAut '1_A = '1_A.
+Proof. by apply/cfunP=> x; rewrite !cfunElock rmorph_nat. Qed.
+
+Lemma cfAutZ a phi : cfAut (a *: phi) = u a *: cfAut phi.
+Proof. by apply/cfunP=> x; rewrite !cfunE rmorphM. Qed.
+
+Lemma cfAut_is_rmorphism : rmorphism cfAut.
+Proof.
+by do 2?split=> [phi psi|]; last exact: cfAut_cfun1i;
+ apply/cfunP=> x; rewrite !cfunE (rmorphB, rmorphM).
+Qed.
+Canonical cfAut_additive := Additive cfAut_is_rmorphism.
+Canonical cfAut_rmorphism := RMorphism cfAut_is_rmorphism.
+
+Lemma cfAut_cfun1 : cfAut 1 = 1. Proof. exact: rmorph1. Qed.
+
+Lemma cfAut_scalable : scalable_for (u \; *:%R) cfAut.
+Proof. by move=> a phi; apply/cfunP=> x; rewrite !cfunE rmorphM. Qed.
+Canonical cfAut_linear := AddLinear cfAut_scalable.
+Canonical cfAut_lrmorphism := [lrmorphism of cfAut].
+
+Definition cfAut_closed (S : seq classfun) :=
+ {in S, forall phi, cfAut phi \in S}.
+
+End Automorphism.
+
+Definition cfReal phi := cfAut conjC phi == phi.
+
+Definition cfConjC_subset (S1 S2 : seq classfun) :=
+ [/\ uniq S1, {subset S1 <= S2} & cfAut_closed conjC S1].
+
+Fact cfun_vect_iso : Vector.axiom #|classes G| classfun.
+Proof.
+exists (fun phi => \row_i phi (repr (enum_val i))) => [a phi psi|].
+ by apply/rowP=> i; rewrite !(mxE, cfunE).
+set n := #|_|; pose eK x : 'I_n := enum_rank_in (classes1 _) (x ^: G).
+have rV2vP v : is_class_fun G [ffun x => v (eK x) *+ (x \in G)].
+ apply: intro_class_fun => [x y Gx Gy | x /negbTE/=-> //].
+ by rewrite groupJr // /eK classGidl.
+exists (fun v : 'rV_n => Cfun 0 (rV2vP (v 0))) => [phi | v].
+ apply/cfun_in_genP=> x Gx; rewrite cfunE Gx mxE enum_rankK_in ?mem_classes //.
+ by have [y Gy ->] := repr_class <<B>> x; rewrite cfunJgen.
+apply/rowP=> i; rewrite mxE cfunE; have /imsetP[x Gx def_i] := enum_valP i.
+rewrite def_i; have [y Gy ->] := repr_class <<B>> x.
+by rewrite groupJ // /eK classGidl // -def_i enum_valK_in.
+Qed.
+Definition cfun_vectMixin := VectMixin cfun_vect_iso.
+Canonical cfun_vectType := VectType algC classfun cfun_vectMixin.
+Canonical cfun_FalgType := [FalgType algC of classfun].
+
+Definition cfun_base A : #|classes B ::&: A|.-tuple classfun :=
+ [tuple of [seq '1_xB | xB in classes B ::&: A]].
+Definition classfun_on A := <<cfun_base A>>%VS.
+
+Definition cfdot phi psi := #|B|%:R^-1 * \sum_(x in B) phi x * (psi x)^*.
+Definition cfdotr_head k psi phi := let: tt := k in cfdot phi psi.
+Definition cfnorm_head k phi := let: tt := k in cfdot phi phi.
+
+Coercion seq_of_cfun phi := [:: phi].
+
+Definition cforder phi := \big[lcmn/1%N]_(x in <<B>>) #[phi x]%C.
+
+End Defs.
+
+Bind Scope cfun_scope with classfun.
+
+Arguments Scope classfun [_ group_scope].
+Arguments Scope classfun_on [_ group_scope group_scope].
+Arguments Scope cfun_indicator [_ group_scope].
+Arguments Scope cfAut [_ group_scope _ cfun_scope].
+Arguments Scope cfReal [_ group_scope cfun_scope].
+Arguments Scope cfdot [_ group_scope cfun_scope cfun_scope].
+Arguments Scope cfdotr_head [_ group_scope _ cfun_scope cfun_scope].
+Arguments Scope cfdotr_head [_ group_scope _ cfun_scope].
+
+Notation "''CF' ( G )" := (classfun G) : type_scope.
+Notation "''CF' ( G )" := (@fullv _ (cfun_vectType G)) : vspace_scope.
+Notation "''1_' A" := (cfun_indicator _ A) : ring_scope.
+Notation "''CF' ( G , A )" := (classfun_on G A) : ring_scope.
+Notation "1" := (@GRing.one (cfun_ringType _)) (only parsing) : cfun_scope.
+
+Notation "phi ^*" := (cfAut conjC phi) : cfun_scope.
+Notation conjC_closed := (cfAut_closed conjC).
+Prenex Implicits cfReal.
+(* Workaround for overeager projection reduction. *)
+Notation eqcfP := (@eqP (cfun_eqType _) _ _) (only parsing).
+
+Notation "#[ phi ]" := (cforder phi) : cfun_scope.
+Notation "''[' u , v ]_ G":= (@cfdot _ G u v) (only parsing) : ring_scope.
+Notation "''[' u , v ]" := (cfdot u v) : ring_scope.
+Notation "''[' u ]_ G" := '[u, u]_G (only parsing) : ring_scope.
+Notation "''[' u ]" := '[u, u] : ring_scope.
+Notation cfdotr := (cfdotr_head tt).
+Notation cfnorm := (cfnorm_head tt).
+
+Section Predicates.
+
+Variables (gT rT : finGroupType) (D : {set gT}) (R : {set rT}).
+Implicit Types (phi psi : 'CF(D)) (S : seq 'CF(D)) (tau : 'CF(D) -> 'CF(R)).
+
+Definition cfker phi := [set x in D | [forall y, phi (x * y)%g == phi y]].
+
+Definition cfaithful phi := cfker phi \subset [1].
+
+Definition ortho_rec S1 S2 :=
+ all [pred phi | all [pred psi | '[phi, psi] == 0] S2] S1.
+
+Fixpoint pair_ortho_rec S :=
+ if S is psi :: S' then ortho_rec psi S' && pair_ortho_rec S' else true.
+
+(* We exclude 0 from pairwise orthogonal sets. *)
+Definition pairwise_orthogonal S := (0 \notin S) && pair_ortho_rec S.
+
+Definition orthonormal S := all [pred psi | '[psi] == 1] S && pair_ortho_rec S.
+
+Definition isometry tau := forall phi psi, '[tau phi, tau psi] = '[phi, psi].
+
+Definition isometry_from_to mCFD tau mCFR :=
+ prop_in2 mCFD (inPhantom (isometry tau))
+ /\ prop_in1 mCFD (inPhantom (forall phi, in_mem (tau phi) mCFR)).
+
+End Predicates.
+
+(* Outside section so the nosimpl does not get "cooked" out. *)
+Definition orthogonal gT D S1 S2 := nosimpl (@ortho_rec gT D S1 S2).
+
+Arguments Scope cfker [_ group_scope cfun_scope].
+Arguments Scope cfaithful [_ group_scope cfun_scope].
+Arguments Scope orthogonal [_ group_scope cfun_scope cfun_scope].
+Arguments Scope pairwise_orthogonal [_ group_scope cfun_scope].
+Arguments Scope orthonormal [_ group_scope cfun_scope].
+Arguments Scope isometry [_ _ group_scope group_scope cfun_scope].
+
+Notation "{ 'in' CFD , 'isometry' tau , 'to' CFR }" :=
+ (isometry_from_to (mem CFD) tau (mem CFR))
+ (at level 0, format "{ 'in' CFD , 'isometry' tau , 'to' CFR }")
+ : type_scope.
+
+Section ClassFun.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Implicit Types (A B : {set gT}) (H K : {group gT}) (phi psi xi : 'CF(G)).
+
+Local Notation "''1_' A" := (cfun_indicator G A).
+
+Lemma cfun0 phi x : x \notin G -> phi x = 0.
+Proof. by rewrite -{1}(genGid G) => /(cfun0gen phi). Qed.
+
+Lemma support_cfun phi : support phi \subset G.
+Proof. by apply/subsetP=> g; apply: contraR => /cfun0->. Qed.
+
+Lemma cfunJ phi x y : y \in G -> phi (x ^ y) = phi x.
+Proof. by rewrite -{1}(genGid G) => /(cfunJgen phi)->. Qed.
+
+Lemma cfun_repr phi x : phi (repr (x ^: G)) = phi x.
+Proof. by have [y Gy ->] := repr_class G x; exact: cfunJ. Qed.
+
+Lemma cfun_inP phi psi : {in G, phi =1 psi} -> phi = psi.
+Proof. by rewrite -{1}genGid => /cfun_in_genP. Qed.
+
+Lemma cfuniE A x : A <| G -> '1_A x = (x \in A)%:R.
+Proof.
+case/andP=> sAG nAG; rewrite cfunElock genGid.
+by rewrite class_sub_norm // andb_idl // => /(subsetP sAG).
+Qed.
+
+Lemma support_cfuni A : A <| G -> support '1_A =i A.
+Proof. by move=> nsAG x; rewrite !inE cfuniE // pnatr_eq0 -lt0n lt0b. Qed.
+
+Lemma eq_mul_cfuni A phi : A <| G -> {in A, phi * '1_A =1 phi}.
+Proof. by move=> nsAG x Ax; rewrite cfunE cfuniE // Ax mulr1. Qed.
+
+Lemma eq_cfuni A : A <| G -> {in A, '1_A =1 (1 : 'CF(G))}.
+Proof. by rewrite -['1_A]mul1r; exact: eq_mul_cfuni. Qed.
+
+Lemma cfuniG : '1_G = 1.
+Proof. by rewrite -[G in '1_G]genGid. Qed.
+
+Lemma cfun1E g : (1 : 'CF(G)) g = (g \in G)%:R.
+Proof. by rewrite -cfuniG cfuniE. Qed.
+
+Lemma cfun11 : (1 : 'CF(G)) 1%g = 1.
+Proof. by rewrite cfun1E group1. Qed.
+
+Lemma prod_cfunE I r (P : pred I) (phi : I -> 'CF(G)) x :
+ x \in G -> (\prod_(i <- r | P i) phi i) x = \prod_(i <- r | P i) (phi i) x.
+Proof.
+by move=> Gx; elim/big_rec2: _ => [|i _ psi _ <-]; rewrite ?cfunE ?cfun1E ?Gx.
+Qed.
+
+Lemma exp_cfunE phi n x : x \in G -> (phi ^+ n) x = phi x ^+ n.
+Proof. by rewrite -[n]card_ord -!prodr_const; apply: prod_cfunE. Qed.
+
+Lemma mul_cfuni A B : '1_A * '1_B = '1_(A :&: B) :> 'CF(G).
+Proof.
+apply/cfunP=> g; rewrite !cfunElock -natrM mulnb subsetI.
+by rewrite andbCA !andbA andbb.
+Qed.
+
+Lemma cfun_classE x y : '1_(x ^: G) y = ((x \in G) && (y \in x ^: G))%:R.
+Proof.
+rewrite cfunElock genGid class_sub_norm ?class_norm //; congr (_ : bool)%:R.
+by apply: andb_id2r => /imsetP[z Gz ->]; rewrite groupJr.
+Qed.
+
+Lemma cfun_on_sum A :
+ 'CF(G, A) = (\sum_(xG in classes G | xG \subset A) <['1_xG]>)%VS.
+Proof.
+rewrite ['CF(G, A)]span_def big_map big_filter.
+by apply: eq_bigl => xG; rewrite !inE.
+Qed.
+
+Lemma cfun_onP A phi :
+ reflect (forall x, x \notin A -> phi x = 0) (phi \in 'CF(G, A)).
+Proof.
+apply: (iffP idP) => [/coord_span-> x notAx | Aphi].
+ set b := cfun_base G A; rewrite sum_cfunE big1 // => i _; rewrite cfunE.
+ have /mapP[xG]: b`_i \in b by rewrite -tnth_nth mem_tnth.
+ rewrite mem_enum => /setIdP[/imsetP[y Gy ->] Ay] ->.
+ by rewrite cfun_classE Gy (contraNF (subsetP Ay x)) ?mulr0.
+suffices <-: \sum_(xG in classes G) phi (repr xG) *: '1_xG = phi.
+ apply: memv_suml => _ /imsetP[x Gx ->]; rewrite rpredZeq cfun_repr.
+ have [s_xG_A | /subsetPn[_ /imsetP[y Gy ->]]] := boolP (x ^: G \subset A).
+ by rewrite cfun_on_sum [_ \in _](sumv_sup (x ^: G)) ?mem_classes ?orbT.
+ by move/Aphi; rewrite cfunJ // => ->; rewrite eqxx.
+apply/cfun_inP=> x Gx; rewrite sum_cfunE (bigD1 (x ^: G)) ?mem_classes //=.
+rewrite cfunE cfun_repr cfun_classE Gx class_refl mulr1.
+rewrite big1 ?addr0 // => _ /andP[/imsetP[y Gy ->]]; apply: contraNeq.
+rewrite cfunE cfun_repr cfun_classE Gy mulf_eq0 => /norP[_].
+by rewrite pnatr_eq0 -lt0n lt0b => /class_transr->.
+Qed.
+Implicit Arguments cfun_onP [A phi].
+
+Lemma cfun_on0 A phi x : phi \in 'CF(G, A) -> x \notin A -> phi x = 0.
+Proof. by move/cfun_onP; exact. Qed.
+
+Lemma sum_by_classes (R : ringType) (F : gT -> R) :
+ {in G &, forall g h, F (g ^ h) = F g} ->
+ \sum_(g in G) F g = \sum_(xG in classes G) #|xG|%:R * F (repr xG).
+Proof.
+move=> FJ; rewrite {1}(partition_big _ _ ((@mem_classes gT)^~ G)) /=.
+apply: eq_bigr => _ /imsetP[x Gx ->]; have [y Gy ->] := repr_class G x.
+rewrite mulr_natl -sumr_const FJ {y Gy}//; apply/esym/eq_big=> y /=.
+ apply/idP/andP=> [xGy | [Gy /eqP<-]]; last exact: class_refl.
+ by rewrite (class_transr xGy) (subsetP (class_subG Gx (subxx _))).
+by case/imsetP=> z Gz ->; rewrite FJ.
+Qed.
+
+Lemma cfun_base_free A : free (cfun_base G A).
+Proof.
+have b_i (i : 'I_#|classes G ::&: A|) : (cfun_base G A)`_i = '1_(enum_val i).
+ by rewrite /enum_val -!tnth_nth tnth_map.
+apply/freeP => s S0 i; move/cfunP/(_ (repr (enum_val i))): S0.
+rewrite sum_cfunE (bigD1 i) //= big1 ?addr0 => [|j].
+ rewrite b_i !cfunE; have /setIdP[/imsetP[x Gx ->] _] := enum_valP i.
+ by rewrite cfun_repr cfun_classE Gx class_refl mulr1.
+apply: contraNeq; rewrite b_i !cfunE mulf_eq0 => /norP[_].
+rewrite -(inj_eq enum_val_inj).
+have /setIdP[/imsetP[x _ ->] _] := enum_valP i; rewrite cfun_repr.
+have /setIdP[/imsetP[y Gy ->] _] := enum_valP j; rewrite cfun_classE Gy.
+by rewrite pnatr_eq0 -lt0n lt0b => /class_transr->.
+Qed.
+
+Lemma dim_cfun : \dim 'CF(G) = #|classes G|.
+Proof. by rewrite dimvf /Vector.dim /= genGid. Qed.
+
+Lemma dim_cfun_on A : \dim 'CF(G, A) = #|classes G ::&: A|.
+Proof. by rewrite (eqnP (cfun_base_free A)) size_tuple. Qed.
+
+Lemma dim_cfun_on_abelian A : abelian G -> A \subset G -> \dim 'CF(G, A) = #|A|.
+Proof.
+move/abelian_classP=> cGG sAG; rewrite -(card_imset _ set1_inj) dim_cfun_on.
+apply/eq_card=> xG; rewrite !inE.
+apply/andP/imsetP=> [[/imsetP[x Gx ->] Ax] | [x Ax ->]] {xG}.
+ by rewrite cGG ?sub1set // in Ax *; exists x.
+by rewrite -{1}(cGG x) ?mem_classes ?(subsetP sAG) ?sub1set.
+Qed.
+
+Lemma cfuni_on A : '1_A \in 'CF(G, A).
+Proof.
+apply/cfun_onP=> x notAx; rewrite cfunElock genGid.
+by case: andP => // [[_ s_xG_A]]; rewrite (subsetP s_xG_A) ?class_refl in notAx.
+Qed.
+
+Lemma mul_cfuni_on A phi : phi * '1_A \in 'CF(G, A).
+Proof.
+by apply/cfun_onP=> x /(cfun_onP (cfuni_on A)) Ax0; rewrite cfunE Ax0 mulr0.
+Qed.
+
+Lemma cfun_onE phi A : (phi \in 'CF(G, A)) = (support phi \subset A).
+Proof. exact: (sameP cfun_onP supportP). Qed.
+
+Lemma cfun_onT phi : phi \in 'CF(G, [set: gT]).
+Proof. by rewrite cfun_onE. Qed.
+
+Lemma cfun_onD1 phi A :
+ (phi \in 'CF(G, A^#)) = (phi \in 'CF(G, A)) && (phi 1%g == 0).
+Proof.
+by rewrite !cfun_onE -!(eq_subset (in_set (support _))) subsetD1 !inE negbK.
+Qed.
+
+Lemma cfun_onG phi : phi \in 'CF(G, G).
+Proof. by rewrite cfun_onE support_cfun. Qed.
+
+Lemma cfunD1E phi : (phi \in 'CF(G, G^#)) = (phi 1%g == 0).
+Proof. by rewrite cfun_onD1 cfun_onG. Qed.
+
+Lemma cfunGid : 'CF(G, G) = 'CF(G)%VS.
+Proof. by apply/vspaceP=> phi; rewrite cfun_onG memvf. Qed.
+
+Lemma cfun_onS A B phi : B \subset A -> phi \in 'CF(G, B) -> phi \in 'CF(G, A).
+Proof. by rewrite !cfun_onE => sBA /subset_trans->. Qed.
+
+Lemma cfun_complement A :
+ A <| G -> ('CF(G, A) + 'CF(G, G :\: A)%SET = 'CF(G))%VS.
+Proof.
+case/andP=> sAG nAG; rewrite -cfunGid [rhs in _ = rhs]cfun_on_sum.
+rewrite (bigID (fun B => B \subset A)) /=.
+congr (_ + _)%VS; rewrite cfun_on_sum; apply: eq_bigl => /= xG.
+ rewrite andbAC; apply/esym/andb_idr=> /andP[/imsetP[x Gx ->] _].
+ by rewrite class_subG.
+rewrite -andbA; apply: andb_id2l => /imsetP[x Gx ->].
+by rewrite !class_sub_norm ?normsD ?normG // inE andbC.
+Qed.
+
+Lemma cfConjCE phi x : (phi^*)%CF x = (phi x)^*.
+Proof. by rewrite cfunE. Qed.
+
+Lemma cfConjCK : involutive (fun phi => phi^*)%CF.
+Proof. by move=> phi; apply/cfunP=> x; rewrite !cfunE conjCK. Qed.
+
+Lemma cfConjC_cfun1 : (1^*)%CF = 1 :> 'CF(G).
+Proof. exact: rmorph1. Qed.
+
+(* Class function kernel and faithful class functions *)
+
+Fact cfker_is_group phi : group_set (cfker phi).
+Proof.
+apply/group_setP; split=> [|x y]; rewrite !inE ?group1.
+ by apply/forallP=> y; rewrite mul1g.
+case/andP=> Gx /forallP-Kx /andP[Gy /forallP-Ky]; rewrite groupM //.
+by apply/forallP=> z; rewrite -mulgA (eqP (Kx _)) Ky.
+Qed.
+Canonical cfker_group phi := Group (cfker_is_group phi).
+
+Lemma cfker_sub phi : cfker phi \subset G.
+Proof. by rewrite /cfker setIdE subsetIl. Qed.
+
+Lemma cfker_norm phi : G \subset 'N(cfker phi).
+Proof.
+apply/subsetP=> z Gz; have phiJz := cfunJ phi _ (groupVr Gz).
+rewrite inE; apply/subsetP=> _ /imsetP[x /setIdP[Gx /forallP-Kx] ->].
+rewrite inE groupJ //; apply/forallP=> y.
+by rewrite -(phiJz y) -phiJz conjMg conjgK Kx.
+Qed.
+
+Lemma cfker_normal phi : cfker phi <| G.
+Proof. by rewrite /normal cfker_sub cfker_norm. Qed.
+
+Lemma cfkerMl phi x y : x \in cfker phi -> phi (x * y)%g = phi y.
+Proof. by case/setIdP=> _ /eqfunP->. Qed.
+
+Lemma cfkerMr phi x y : x \in cfker phi -> phi (y * x)%g = phi y.
+Proof.
+by move=> Kx; rewrite conjgC cfkerMl ?cfunJ ?(subsetP (cfker_sub phi)).
+Qed.
+
+Lemma cfker1 phi x : x \in cfker phi -> phi x = phi 1%g.
+Proof. by move=> Kx; rewrite -[x]mulg1 cfkerMl. Qed.
+
+Lemma cfker_cfun0 : @cfker _ G 0 = G.
+Proof.
+apply/setP=> x; rewrite !inE andb_idr // => Gx; apply/forallP=> y.
+by rewrite !cfunE.
+Qed.
+
+Lemma cfker_add phi psi : cfker phi :&: cfker psi \subset cfker (phi + psi).
+Proof.
+apply/subsetP=> x /setIP[Kphi_x Kpsi_x]; have [Gx _] := setIdP Kphi_x.
+by rewrite inE Gx; apply/forallP=> y; rewrite !cfunE !cfkerMl.
+Qed.
+
+Lemma cfker_sum I r (P : pred I) (Phi : I -> 'CF(G)) :
+ G :&: \bigcap_(i <- r | P i) cfker (Phi i)
+ \subset cfker (\sum_(i <- r | P i) Phi i).
+Proof.
+elim/big_rec2: _ => [|i K psi Pi sK_psi]; first by rewrite setIT cfker_cfun0.
+by rewrite setICA; apply: subset_trans (setIS _ sK_psi) (cfker_add _ _).
+Qed.
+
+Lemma cfker_scale a phi : cfker phi \subset cfker (a *: phi).
+Proof.
+apply/subsetP=> x Kphi_x; have [Gx _] := setIdP Kphi_x.
+by rewrite inE Gx; apply/forallP=> y; rewrite !cfunE cfkerMl.
+Qed.
+
+Lemma cfker_scale_nz a phi : a != 0 -> cfker (a *: phi) = cfker phi.
+Proof.
+move=> nz_a; apply/eqP.
+by rewrite eqEsubset -{2}(scalerK nz_a phi) !cfker_scale.
+Qed.
+
+Lemma cfker_opp phi : cfker (- phi) = cfker phi.
+Proof. by rewrite -scaleN1r cfker_scale_nz // oppr_eq0 oner_eq0. Qed.
+
+Lemma cfker_cfun1 : @cfker _ G 1 = G.
+Proof.
+apply/setP=> x; rewrite !inE andb_idr // => Gx; apply/forallP=> y.
+by rewrite !cfun1E groupMl.
+Qed.
+
+Lemma cfker_mul phi psi : cfker phi :&: cfker psi \subset cfker (phi * psi).
+Proof.
+apply/subsetP=> x /setIP[Kphi_x Kpsi_x]; have [Gx _] := setIdP Kphi_x.
+by rewrite inE Gx; apply/forallP=> y; rewrite !cfunE !cfkerMl.
+Qed.
+
+Lemma cfker_prod I r (P : pred I) (Phi : I -> 'CF(G)) :
+ G :&: \bigcap_(i <- r | P i) cfker (Phi i)
+ \subset cfker (\prod_(i <- r | P i) Phi i).
+Proof.
+elim/big_rec2: _ => [|i K psi Pi sK_psi]; first by rewrite setIT cfker_cfun1.
+by rewrite setICA; apply: subset_trans (setIS _ sK_psi) (cfker_mul _ _).
+Qed.
+
+Lemma cfaithfulE phi : cfaithful phi = (cfker phi \subset [1]).
+Proof. by []. Qed.
+
+End ClassFun.
+
+Arguments Scope classfun_on [_ group_scope group_scope].
+Notation "''CF' ( G , A )" := (classfun_on G A) : ring_scope.
+
+Implicit Arguments cfun_onP [gT G A phi].
+Hint Resolve cfun_onT.
+
+Section DotProduct.
+
+Variable (gT : finGroupType) (G : {group gT}).
+Implicit Types (M : {group gT}) (phi psi xi : 'CF(G)) (R S : seq 'CF(G)).
+
+Lemma cfdotE phi psi :
+ '[phi, psi] = #|G|%:R^-1 * \sum_(x in G) phi x * (psi x)^*.
+Proof. by []. Qed.
+
+Lemma cfdotElr A B phi psi :
+ phi \in 'CF(G, A) -> psi \in 'CF(G, B) ->
+ '[phi, psi] = #|G|%:R^-1 * \sum_(x in A :&: B) phi x * (psi x)^*.
+Proof.
+move=> Aphi Bpsi; rewrite (big_setID G) cfdotE (big_setID (A :&: B)) setIC /=.
+congr (_ * (_ + _)); rewrite !big1 // => x /setDP[_].
+ by move/cfun0->; rewrite mul0r.
+rewrite inE; case/nandP=> notABx; first by rewrite (cfun_on0 Aphi) ?mul0r.
+by rewrite (cfun_on0 Bpsi) // rmorph0 mulr0.
+Qed.
+
+Lemma cfdotEl A phi psi :
+ phi \in 'CF(G, A) ->
+ '[phi, psi] = #|G|%:R^-1 * \sum_(x in A) phi x * (psi x)^*.
+Proof. by move=> Aphi; rewrite (cfdotElr Aphi (cfun_onT psi)) setIT. Qed.
+
+Lemma cfdotEr A phi psi :
+ psi \in 'CF(G, A) ->
+ '[phi, psi] = #|G|%:R^-1 * \sum_(x in A) phi x * (psi x)^*.
+Proof. by move=> Apsi; rewrite (cfdotElr (cfun_onT phi) Apsi) setTI. Qed.
+
+Lemma cfdot_complement A phi psi :
+ phi \in 'CF(G, A) -> psi \in 'CF(G, G :\: A) -> '[phi, psi] = 0.
+Proof.
+move=> Aphi A'psi; rewrite (cfdotElr Aphi A'psi).
+by rewrite setDE setICA setICr setI0 big_set0 mulr0.
+Qed.
+
+Lemma cfnormE A phi :
+ phi \in 'CF(G, A) -> '[phi] = #|G|%:R^-1 * (\sum_(x in A) `|phi x| ^+ 2).
+Proof. by move/cfdotEl->; rewrite (eq_bigr _ (fun _ _ => normCK _)). Qed.
+
+Lemma eq_cfdotl A phi1 phi2 psi :
+ psi \in 'CF(G, A) -> {in A, phi1 =1 phi2} -> '[phi1, psi] = '[phi2, psi].
+Proof.
+move/cfdotEr=> eq_dot eq_phi; rewrite !eq_dot; congr (_ * _).
+by apply: eq_bigr => x Ax; rewrite eq_phi.
+Qed.
+
+Lemma cfdot_cfuni A B :
+ A <| G -> B <| G -> '['1_A, '1_B]_G = #|A :&: B|%:R / #|G|%:R.
+Proof.
+move=> nsAG nsBG; rewrite (cfdotElr (cfuni_on G A) (cfuni_on G B)) mulrC.
+congr (_ / _); rewrite -sumr_const; apply: eq_bigr => x /setIP[Ax Bx].
+by rewrite !cfuniE // Ax Bx rmorph1 mulr1.
+Qed.
+
+Lemma cfnorm1 : '[1]_G = 1.
+Proof. by rewrite cfdot_cfuni ?genGid // setIid divff ?neq0CG. Qed.
+
+Lemma cfdotrE psi phi : cfdotr psi phi = '[phi, psi]. Proof. by []. Qed.
+
+Lemma cfdotr_is_linear xi : linear (cfdotr xi : 'CF(G) -> algC^o).
+Proof.
+move=> a phi psi; rewrite scalerAr -mulrDr; congr (_ * _).
+rewrite linear_sum -big_split; apply: eq_bigr => x _ /=.
+by rewrite !cfunE mulrDl -mulrA.
+Qed.
+Canonical cfdotr_additive xi := Additive (cfdotr_is_linear xi).
+Canonical cfdotr_linear xi := Linear (cfdotr_is_linear xi).
+
+Lemma cfdot0l xi : '[0, xi] = 0.
+Proof. by rewrite -cfdotrE linear0. Qed.
+Lemma cfdotNl xi phi : '[- phi, xi] = - '[phi, xi].
+Proof. by rewrite -!cfdotrE linearN. Qed.
+Lemma cfdotDl xi phi psi : '[phi + psi, xi] = '[phi, xi] + '[psi, xi].
+Proof. by rewrite -!cfdotrE linearD. Qed.
+Lemma cfdotBl xi phi psi : '[phi - psi, xi] = '[phi, xi] - '[psi, xi].
+Proof. by rewrite -!cfdotrE linearB. Qed.
+Lemma cfdotMnl xi phi n : '[phi *+ n, xi] = '[phi, xi] *+ n.
+Proof. by rewrite -!cfdotrE linearMn. Qed.
+Lemma cfdot_suml xi I r (P : pred I) (phi : I -> 'CF(G)) :
+ '[\sum_(i <- r | P i) phi i, xi] = \sum_(i <- r | P i) '[phi i, xi].
+Proof. by rewrite -!cfdotrE linear_sum. Qed.
+Lemma cfdotZl xi a phi : '[a *: phi, xi] = a * '[phi, xi].
+Proof. by rewrite -!cfdotrE linearZ. Qed.
+
+Lemma cfdotC phi psi : '[phi, psi] = ('[psi, phi])^*.
+Proof.
+rewrite /cfdot rmorphM fmorphV rmorph_nat rmorph_sum; congr (_ * _).
+by apply: eq_bigr=> x _; rewrite rmorphM conjCK mulrC.
+Qed.
+
+Lemma eq_cfdotr A phi psi1 psi2 :
+ phi \in 'CF(G, A) -> {in A, psi1 =1 psi2} -> '[phi, psi1] = '[phi, psi2].
+Proof. by move=> Aphi /eq_cfdotl eq_dot; rewrite cfdotC eq_dot // -cfdotC. Qed.
+
+Lemma cfdotBr xi phi psi : '[xi, phi - psi] = '[xi, phi] - '[xi, psi].
+Proof. by rewrite !(cfdotC xi) -rmorphB cfdotBl. Qed.
+Canonical cfun_dot_additive xi := Additive (cfdotBr xi).
+
+Lemma cfdot0r xi : '[xi, 0] = 0. Proof. exact: raddf0. Qed.
+Lemma cfdotNr xi phi : '[xi, - phi] = - '[xi, phi].
+Proof. exact: raddfN. Qed.
+Lemma cfdotDr xi phi psi : '[xi, phi + psi] = '[xi, phi] + '[xi, psi].
+Proof. exact: raddfD. Qed.
+Lemma cfdotMnr xi phi n : '[xi, phi *+ n] = '[xi, phi] *+ n.
+Proof. exact: raddfMn. Qed.
+Lemma cfdot_sumr xi I r (P : pred I) (phi : I -> 'CF(G)) :
+ '[xi, \sum_(i <- r | P i) phi i] = \sum_(i <- r | P i) '[xi, phi i].
+Proof. exact: raddf_sum. Qed.
+Lemma cfdotZr a xi phi : '[xi, a *: phi] = a^* * '[xi, phi].
+Proof. by rewrite !(cfdotC xi) cfdotZl rmorphM. Qed.
+
+Lemma cfdot_cfAut (u : {rmorphism algC -> algC}) phi psi :
+ {in image psi G, {morph u : x / x^*}} ->
+ '[cfAut u phi, cfAut u psi] = u '[phi, psi].
+Proof.
+move=> uC; rewrite rmorphM fmorphV rmorph_nat rmorph_sum; congr (_ * _).
+by apply: eq_bigr => x Gx; rewrite !cfunE rmorphM uC ?map_f ?mem_enum.
+Qed.
+
+Lemma cfdot_conjC phi psi : '[phi^*, psi^*] = '[phi, psi]^*.
+Proof. by rewrite cfdot_cfAut. Qed.
+
+Lemma cfdot_conjCl phi psi : '[phi^*, psi] = '[phi, psi^*]^*.
+Proof. by rewrite -cfdot_conjC cfConjCK. Qed.
+
+Lemma cfdot_conjCr phi psi : '[phi, psi^*] = '[phi^*, psi]^*.
+Proof. by rewrite -cfdot_conjC cfConjCK. Qed.
+
+Lemma cfnorm_ge0 phi : 0 <= '[phi].
+Proof.
+by rewrite mulr_ge0 ?invr_ge0 ?ler0n ?sumr_ge0 // => x _; exact: mul_conjC_ge0.
+Qed.
+
+Lemma cfnorm_eq0 phi : ('[phi] == 0) = (phi == 0).
+Proof.
+apply/idP/eqP=> [|->]; last by rewrite cfdot0r.
+rewrite mulf_eq0 invr_eq0 (negbTE (neq0CG G)) /= => /eqP/psumr_eq0P phi0.
+apply/cfun_inP=> x Gx; apply/eqP; rewrite cfunE -mul_conjC_eq0.
+by rewrite phi0 // => y _; exact: mul_conjC_ge0.
+Qed.
+
+Lemma cfnorm_gt0 phi : ('[phi] > 0) = (phi != 0).
+Proof. by rewrite ltr_def cfnorm_ge0 cfnorm_eq0 andbT. Qed.
+
+Lemma sqrt_cfnorm_ge0 phi : 0 <= sqrtC '[phi].
+Proof. by rewrite sqrtC_ge0 cfnorm_ge0. Qed.
+
+Lemma sqrt_cfnorm_eq0 phi : (sqrtC '[phi] == 0) = (phi == 0).
+Proof. by rewrite sqrtC_eq0 cfnorm_eq0. Qed.
+
+Lemma sqrt_cfnorm_gt0 phi : (sqrtC '[phi] > 0) = (phi != 0).
+Proof. by rewrite sqrtC_gt0 cfnorm_gt0. Qed.
+
+Lemma cfnormZ a phi : '[a *: phi]= `|a| ^+ 2 * '[phi]_G.
+Proof. by rewrite cfdotZl cfdotZr mulrA normCK. Qed.
+
+Lemma cfnormN phi : '[- phi] = '[phi].
+Proof. by rewrite cfdotNl raddfN opprK. Qed.
+
+Lemma cfnorm_sign n phi : '[(-1) ^+ n *: phi] = '[phi].
+Proof. by rewrite -signr_odd scaler_sign; case: (odd n); rewrite ?cfnormN. Qed.
+
+Lemma cfnormD phi psi :
+ let d := '[phi, psi] in '[phi + psi] = '[phi] + '[psi] + (d + d^*).
+Proof. by rewrite /= addrAC -cfdotC cfdotDl !cfdotDr !addrA. Qed.
+
+Lemma cfnormB phi psi :
+ let d := '[phi, psi] in '[phi - psi] = '[phi] + '[psi] - (d + d^*).
+Proof. by rewrite /= cfnormD cfnormN cfdotNr rmorphN -opprD. Qed.
+
+Lemma cfnormDd phi psi : '[phi, psi] = 0 -> '[phi + psi] = '[phi] + '[psi].
+Proof. by move=> ophipsi; rewrite cfnormD ophipsi rmorph0 !addr0. Qed.
+
+Lemma cfnormBd phi psi : '[phi, psi] = 0 -> '[phi - psi] = '[phi] + '[psi].
+Proof.
+by move=> ophipsi; rewrite cfnormDd ?cfnormN // cfdotNr ophipsi oppr0.
+Qed.
+
+Lemma cfnorm_conjC phi : '[phi^*] = '[phi].
+Proof. by rewrite cfdot_conjC geC0_conj // cfnorm_ge0. Qed.
+
+Lemma cfCauchySchwarz phi psi :
+ `|'[phi, psi]| ^+ 2 <= '[phi] * '[psi] ?= iff ~~ free (phi :: psi).
+Proof.
+rewrite free_cons span_seq1 seq1_free -negb_or negbK orbC.
+have [-> | nz_psi] /= := altP (psi =P 0).
+ by apply/lerifP; rewrite !cfdot0r normCK mul0r mulr0.
+without loss ophi: phi / '[phi, psi] = 0.
+ move=> IHo; pose a := '[phi, psi] / '[psi]; pose phi1 := phi - a *: psi.
+ have ophi: '[phi1, psi] = 0.
+ by rewrite cfdotBl cfdotZl divfK ?cfnorm_eq0 ?subrr.
+ rewrite (canRL (subrK _) (erefl phi1)) rpredDr ?rpredZ ?memv_line //.
+ rewrite cfdotDl ophi add0r cfdotZl normrM (ger0_norm (cfnorm_ge0 _)).
+ rewrite exprMn mulrA -cfnormZ cfnormDd; last by rewrite cfdotZr ophi mulr0.
+ by have:= IHo _ ophi; rewrite mulrDl -lerif_subLR subrr ophi normCK mul0r.
+rewrite ophi normCK mul0r; split; first by rewrite mulr_ge0 ?cfnorm_ge0.
+rewrite eq_sym mulf_eq0 orbC cfnorm_eq0 (negPf nz_psi) /=.
+apply/idP/idP=> [|/vlineP[a {2}->]]; last by rewrite cfdotZr ophi mulr0.
+by rewrite cfnorm_eq0 => /eqP->; apply: rpred0.
+Qed.
+
+Lemma cfCauchySchwarz_sqrt phi psi :
+ `|'[phi, psi]| <= sqrtC '[phi] * sqrtC '[psi] ?= iff ~~ free (phi :: psi).
+Proof.
+rewrite -(sqrCK (normr_ge0 _)) -sqrtCM ?qualifE ?cfnorm_ge0 //.
+rewrite (mono_in_lerif ler_sqrtC) 1?rpredM ?qualifE ?normr_ge0 ?cfnorm_ge0 //.
+exact: cfCauchySchwarz.
+Qed.
+
+Lemma cf_triangle_lerif phi psi :
+ sqrtC '[phi + psi] <= sqrtC '[phi] + sqrtC '[psi]
+ ?= iff ~~ free (phi :: psi) && (0 <= coord [tuple psi] 0 phi).
+Proof.
+rewrite -(mono_in_lerif ler_sqr) ?rpredD ?qualifE ?sqrtC_ge0 ?cfnorm_ge0 //.
+rewrite andbC sqrrD !sqrtCK addrAC cfnormD (mono_lerif (ler_add2l _)).
+rewrite -mulr_natr -[_ + _](divfK (negbT (eqC_nat 2 0))) -/('Re _).
+rewrite (mono_lerif (ler_pmul2r _)) ?ltr0n //.
+have:= lerif_trans (lerif_Re_Creal '[phi, psi]) (cfCauchySchwarz_sqrt phi psi).
+congr (_ <= _ ?= iff _); apply: andb_id2r.
+rewrite free_cons span_seq1 seq1_free -negb_or negbK orbC.
+have [-> | nz_psi] := altP (psi =P 0); first by rewrite cfdot0r coord0.
+case/vlineP=> [x ->]; rewrite cfdotZl linearZ pmulr_lge0 ?cfnorm_gt0 //=.
+by rewrite (coord_free 0) ?seq1_free // eqxx mulr1.
+Qed.
+
+Lemma orthogonal_cons phi R S :
+ orthogonal (phi :: R) S = orthogonal phi S && orthogonal R S.
+Proof. by rewrite /orthogonal /= andbT. Qed.
+
+Lemma orthoP phi psi : reflect ('[phi, psi] = 0) (orthogonal phi psi).
+Proof. by rewrite /orthogonal /= !andbT; exact: eqP. Qed.
+
+Lemma orthogonalP S R :
+ reflect {in S & R, forall phi psi, '[phi, psi] = 0} (orthogonal S R).
+Proof.
+apply: (iffP allP) => oSR phi => [psi /oSR/allP opS /opS/eqP // | /oSR opS].
+by apply/allP=> psi /= /opS->.
+Qed.
+
+Lemma orthoPl phi S :
+ reflect {in S, forall psi, '[phi, psi] = 0} (orthogonal phi S).
+Proof.
+by rewrite [orthogonal _ S]andbT /=; apply: (iffP allP) => ophiS ? /ophiS/eqP.
+Qed.
+Implicit Arguments orthoPl [phi S].
+
+Lemma orthogonal_sym : symmetric (@orthogonal _ G).
+Proof.
+apply: symmetric_from_pre => R S /orthogonalP oRS.
+by apply/orthogonalP=> phi psi Rpsi Sphi; rewrite cfdotC oRS ?rmorph0.
+Qed.
+
+Lemma orthoPr S psi :
+ reflect {in S, forall phi, '[phi, psi] = 0} (orthogonal S psi).
+Proof.
+rewrite orthogonal_sym.
+by apply: (iffP orthoPl) => oSpsi phi Sphi; rewrite cfdotC oSpsi ?conjC0.
+Qed.
+
+Lemma eq_orthogonal R1 R2 S1 S2 :
+ R1 =i R2 -> S1 =i S2 -> orthogonal R1 S1 = orthogonal R2 S2.
+Proof.
+move=> eqR eqS; rewrite [orthogonal _ _](eq_all_r eqR).
+by apply: eq_all => psi /=; exact: eq_all_r.
+Qed.
+
+Lemma orthogonal_catl R1 R2 S :
+ orthogonal (R1 ++ R2) S = orthogonal R1 S && orthogonal R2 S.
+Proof. exact: all_cat. Qed.
+
+Lemma orthogonal_catr R S1 S2 :
+ orthogonal R (S1 ++ S2) = orthogonal R S1 && orthogonal R S2.
+Proof. by rewrite !(orthogonal_sym R) orthogonal_catl. Qed.
+
+Lemma span_orthogonal S1 S2 phi1 phi2 :
+ orthogonal S1 S2 -> phi1 \in <<S1>>%VS -> phi2 \in <<S2>>%VS ->
+ '[phi1, phi2] = 0.
+Proof.
+move/orthogonalP=> oS12; do 2!move/(@coord_span _ _ _ (in_tuple _))->.
+rewrite cfdot_suml big1 // => i _; rewrite cfdot_sumr big1 // => j _.
+by rewrite cfdotZl cfdotZr oS12 ?mem_nth ?mulr0.
+Qed.
+
+Lemma orthogonal_split S beta :
+ {X : 'CF(G) & X \in <<S>>%VS &
+ {Y | [/\ beta = X + Y, '[X, Y] = 0 & orthogonal Y S]}}.
+Proof.
+suffices [X S_X [Y -> oYS]]:
+ {X : _ & X \in <<S>>%VS & {Y | beta = X + Y & orthogonal Y S}}.
+- exists X => //; exists Y.
+ by rewrite cfdotC (span_orthogonal oYS) ?memv_span1 ?conjC0.
+elim: S beta => [|phi S IHS] beta.
+ by exists 0; last exists beta; rewrite ?mem0v ?add0r.
+have [[U S_U [V -> oVS]] [X S_X [Y -> oYS]]] := (IHS phi, IHS beta).
+pose Z := '[Y, V] / '[V] *: V; exists (X + Z).
+ rewrite /Z -{4}(addKr U V) scalerDr scalerN addrA addrC span_cons.
+ by rewrite memv_add ?memvB ?memvZ ?memv_line.
+exists (Y - Z); first by rewrite addrCA !addrA addrK addrC.
+apply/orthoPl=> psi; rewrite !inE => /predU1P[-> | Spsi]; last first.
+ by rewrite cfdotBl cfdotZl (orthoPl oVS _ Spsi) mulr0 subr0 (orthoPl oYS).
+rewrite cfdotBl !cfdotDr (span_orthogonal oYS) // ?memv_span ?mem_head //.
+rewrite !cfdotZl (span_orthogonal oVS _ S_U) ?mulr0 ?memv_span ?mem_head //.
+have [-> | nzV] := eqVneq V 0; first by rewrite cfdot0r !mul0r subrr.
+by rewrite divfK ?cfnorm_eq0 ?subrr.
+Qed.
+
+Lemma map_orthogonal M (nu : 'CF(G) -> 'CF(M)) S R (A : pred 'CF(G)) :
+ {in A &, isometry nu} -> {subset S <= A} -> {subset R <= A} ->
+ orthogonal (map nu S) (map nu R) = orthogonal S R.
+Proof.
+move=> Inu sSA sRA; rewrite [orthogonal _ _]all_map.
+apply: eq_in_all => phi Sphi; rewrite /= all_map.
+by apply: eq_in_all => psi Rpsi; rewrite /= Inu ?(sSA phi) ?(sRA psi).
+Qed.
+
+Lemma orthogonal_oppr S R : orthogonal S (map -%R R) = orthogonal S R.
+Proof.
+wlog suffices IH: S R / orthogonal S R -> orthogonal S (map -%R R).
+ apply/idP/idP=> /IH; rewrite ?mapK //; exact: opprK.
+move/orthogonalP=> oSR; apply/orthogonalP=> xi1 _ Sxi1 /mapP[xi2 Rxi2 ->].
+by rewrite cfdotNr oSR ?oppr0.
+Qed.
+
+Lemma orthogonal_oppl S R : orthogonal (map -%R S) R = orthogonal S R.
+Proof. by rewrite -!(orthogonal_sym R) orthogonal_oppr. Qed.
+
+Lemma pairwise_orthogonalP S :
+ reflect (uniq (0 :: S)
+ /\ {in S &, forall phi psi, phi != psi -> '[phi, psi] = 0})
+ (pairwise_orthogonal S).
+Proof.
+rewrite /pairwise_orthogonal /=; case notS0: (~~ _); last by right; case.
+elim: S notS0 => [|phi S IH] /=; first by left.
+rewrite inE eq_sym andbT => /norP[nz_phi /IH{IH}IH].
+have [opS | not_opS] := allP; last first.
+ right=> [[/andP[notSp _] opS]]; case: not_opS => psi Spsi /=.
+ by rewrite opS ?mem_head 1?mem_behead // (memPnC notSp).
+rewrite (contra (opS _)) /= ?cfnorm_eq0 //.
+apply: (iffP IH) => [] [uniqS oSS]; last first.
+ by split=> //; apply: sub_in2 oSS => psi Spsi; exact: mem_behead.
+split=> // psi xi; rewrite !inE => /predU1P[-> // | Spsi].
+ by case/predU1P=> [-> | /opS] /eqP.
+case/predU1P=> [-> _ | Sxi /oSS-> //].
+by apply/eqP; rewrite cfdotC conjC_eq0 [_ == 0]opS.
+Qed.
+
+Lemma pairwise_orthogonal_cat R S :
+ pairwise_orthogonal (R ++ S) =
+ [&& pairwise_orthogonal R, pairwise_orthogonal S & orthogonal R S].
+Proof.
+rewrite /pairwise_orthogonal mem_cat negb_or -!andbA; do !bool_congr.
+elim: R => [|phi R /= ->]; rewrite ?andbT // orthogonal_cons all_cat -!andbA /=.
+by do !bool_congr.
+Qed.
+
+Lemma eq_pairwise_orthogonal R S :
+ perm_eq R S -> pairwise_orthogonal R = pairwise_orthogonal S.
+Proof.
+apply: catCA_perm_subst R S => R S S'.
+rewrite !pairwise_orthogonal_cat !orthogonal_catr (orthogonal_sym R S) -!andbA.
+by do !bool_congr.
+Qed.
+
+Lemma sub_pairwise_orthogonal S1 S2 :
+ {subset S1 <= S2} -> uniq S1 ->
+ pairwise_orthogonal S2 -> pairwise_orthogonal S1.
+Proof.
+move=> sS12 uniqS1 /pairwise_orthogonalP[/andP[notS2_0 _] oS2].
+apply/pairwise_orthogonalP; rewrite /= (contra (sS12 0)) //.
+by split=> //; exact: sub_in2 oS2.
+Qed.
+
+Lemma orthogonal_free S : pairwise_orthogonal S -> free S.
+Proof.
+case/pairwise_orthogonalP=> [/=/andP[notS0 uniqS] oSS].
+rewrite -(in_tupleE S); apply/freeP => a aS0 i.
+have S_i: S`_i \in S by exact: mem_nth.
+have /eqP: '[S`_i, 0]_G = 0 := cfdot0r _.
+rewrite -{2}aS0 raddf_sum /= (bigD1 i) //= big1 => [|j neq_ji]; last 1 first.
+ by rewrite cfdotZr oSS ?mulr0 ?mem_nth // eq_sym nth_uniq.
+rewrite addr0 cfdotZr mulf_eq0 conjC_eq0 cfnorm_eq0.
+by case/pred2P=> // Si0; rewrite -Si0 S_i in notS0.
+Qed.
+
+Lemma filter_pairwise_orthogonal S p :
+ pairwise_orthogonal S -> pairwise_orthogonal (filter p S).
+Proof.
+move=> orthoS; apply: sub_pairwise_orthogonal (orthoS).
+ exact: mem_subseq (filter_subseq p S).
+exact/filter_uniq/free_uniq/orthogonal_free.
+Qed.
+
+Lemma orthonormal_not0 S : orthonormal S -> 0 \notin S.
+Proof.
+by case/andP=> /allP S1 _; rewrite (contra (S1 _)) //= cfdot0r eq_sym oner_eq0.
+Qed.
+
+Lemma orthonormalE S :
+ orthonormal S = all [pred phi | '[phi] == 1] S && pairwise_orthogonal S.
+Proof. by rewrite -(andb_idl (@orthonormal_not0 S)) andbCA. Qed.
+
+Lemma orthonormal_orthogonal S : orthonormal S -> pairwise_orthogonal S.
+Proof. by rewrite orthonormalE => /andP[_]. Qed.
+
+Lemma orthonormal_cat R S :
+ orthonormal (R ++ S) = [&& orthonormal R, orthonormal S & orthogonal R S].
+Proof.
+rewrite !orthonormalE pairwise_orthogonal_cat all_cat -!andbA.
+by do !bool_congr.
+Qed.
+
+Lemma eq_orthonormal R S : perm_eq R S -> orthonormal R = orthonormal S.
+Proof.
+move=> eqRS; rewrite !orthonormalE (eq_all_r (perm_eq_mem eqRS)).
+by rewrite (eq_pairwise_orthogonal eqRS).
+Qed.
+
+Lemma orthonormal_free S : orthonormal S -> free S.
+Proof. by move/orthonormal_orthogonal/orthogonal_free. Qed.
+
+Lemma orthonormalP S :
+ reflect (uniq S /\ {in S &, forall phi psi, '[phi, psi]_G = (phi == psi)%:R})
+ (orthonormal S).
+Proof.
+rewrite orthonormalE; have [/= normS | not_normS] := allP; last first.
+ by right=> [[_ o1S]]; case: not_normS => phi Sphi; rewrite /= o1S ?eqxx.
+apply: (iffP (pairwise_orthogonalP S)) => [] [uniqS oSS].
+ split=> // [|phi psi]; first by case/andP: uniqS.
+ by have [-> _ /normS/eqP | /oSS] := altP eqP.
+split=> // [|phi psi Sphi Spsi /negbTE]; last by rewrite oSS // => ->.
+by rewrite /= (contra (normS _)) // cfdot0r eq_sym oner_eq0.
+Qed.
+
+Lemma sub_orthonormal S1 S2 :
+ {subset S1 <= S2} -> uniq S1 -> orthonormal S2 -> orthonormal S1.
+Proof.
+move=> sS12 uniqS1 /orthonormalP[_ oS1].
+by apply/orthonormalP; split; last exact: sub_in2 sS12 _ _.
+Qed.
+
+Lemma orthonormal2P phi psi :
+ reflect [/\ '[phi, psi] = 0, '[phi] = 1 & '[psi] = 1]
+ (orthonormal [:: phi; psi]).
+Proof.
+rewrite /orthonormal /= !andbT andbC.
+by apply: (iffP and3P) => [] []; do 3!move/eqP->.
+Qed.
+
+Lemma conjC_pair_orthogonal S chi :
+ conjC_closed S -> ~~ has cfReal S -> pairwise_orthogonal S -> chi \in S ->
+ pairwise_orthogonal (chi :: chi^*%CF).
+Proof.
+move=> ccS /hasPn nrS oSS Schi; apply: sub_pairwise_orthogonal oSS.
+ by apply/allP; rewrite /= Schi ccS.
+by rewrite /= inE eq_sym nrS.
+Qed.
+
+Lemma cfdot_real_conjC phi psi : cfReal phi -> '[phi, psi^*]_G = '[phi, psi]^*.
+Proof. by rewrite -cfdot_conjC => /eqcfP->. Qed.
+
+(* Note: other isometry lemmas, and the dot product lemmas for orthogonal *)
+(* and orthonormal sequences are in vcharacter, because we need the 'Z[S] *)
+(* notation for the isometry domains. Alternatively, this could be moved to *)
+(* cfun. *)
+
+End DotProduct.
+
+Implicit Arguments orthoP [gT G phi psi].
+Implicit Arguments orthoPl [gT G phi S].
+Implicit Arguments orthoPr [gT G S psi].
+Implicit Arguments orthogonalP [gT G R S].
+Implicit Arguments pairwise_orthogonalP [gT G S].
+Implicit Arguments orthonormalP [gT G S].
+
+Section CfunOrder.
+
+Variables (gT : finGroupType) (G : {group gT}) (phi : 'CF(G)).
+
+Lemma dvdn_cforderP n :
+ reflect {in G, forall x, phi x ^+ n = 1} (#[phi]%CF %| n)%N.
+Proof.
+apply: (iffP (dvdn_biglcmP _ _ _)); rewrite genGid => phiG1 x Gx.
+ by apply/eqP; rewrite -dvdn_orderC phiG1.
+by rewrite dvdn_orderC phiG1.
+Qed.
+
+Lemma dvdn_cforder n : (#[phi]%CF %| n) = (phi ^+ n == 1).
+Proof.
+apply/dvdn_cforderP/eqP=> phi_n_1 => [|x Gx].
+ by apply/cfun_inP=> x Gx; rewrite exp_cfunE // cfun1E Gx phi_n_1.
+by rewrite -exp_cfunE // phi_n_1 // cfun1E Gx.
+Qed.
+
+Lemma exp_cforder : phi ^+ #[phi]%CF = 1.
+Proof. by apply/eqP; rewrite -dvdn_cforder. Qed.
+
+End CfunOrder.
+
+Implicit Arguments dvdn_cforderP [gT G phi n].
+
+Section MorphOrder.
+
+Variables (aT rT : finGroupType) (G : {group aT}) (R : {group rT}).
+Variable f : {rmorphism 'CF(G) -> 'CF(R)}.
+
+Lemma cforder_rmorph phi : #[f phi]%CF %| #[phi]%CF.
+Proof. by rewrite dvdn_cforder -rmorphX exp_cforder rmorph1. Qed.
+
+Lemma cforder_inj_rmorph phi : injective f -> #[f phi]%CF = #[phi]%CF.
+Proof.
+move=> inj_f; apply/eqP; rewrite eqn_dvd cforder_rmorph dvdn_cforder /=.
+by rewrite -(rmorph_eq1 _ inj_f) rmorphX exp_cforder.
+Qed.
+
+End MorphOrder.
+
+Section BuildIsometries.
+
+Variable (gT : finGroupType) (L G : {group gT}).
+Implicit Types (phi psi xi : 'CF(L)) (R S : seq 'CF(L)).
+Implicit Types (U : pred 'CF(L)) (W : pred 'CF(G)).
+
+Lemma sub_iso_to U1 U2 W1 W2 tau :
+ {subset U2 <= U1} -> {subset W1 <= W2} ->
+ {in U1, isometry tau, to W1} -> {in U2, isometry tau, to W2}.
+Proof.
+by move=> sU sW [Itau Wtau]; split=> [|u /sU/Wtau/sW //]; exact: sub_in2 Itau.
+Qed.
+
+Lemma isometry_of_cfnorm S tauS :
+ pairwise_orthogonal S -> pairwise_orthogonal tauS ->
+ map cfnorm tauS = map cfnorm S ->
+ {tau : {linear 'CF(L) -> 'CF(G)} | map tau S = tauS
+ & {in <<S>>%VS &, isometry tau}}.
+Proof.
+move=> oS oT eq_nST; have freeS := orthogonal_free oS.
+have eq_sz: size tauS = size S by have:= congr1 size eq_nST; rewrite !size_map.
+have [tau /(_ freeS eq_sz) defT] := linear_of_free S tauS.
+rewrite -[S]/(tval (in_tuple S)).
+exists tau => // u v /coord_span-> /coord_span->; rewrite !raddf_sum /=.
+apply: eq_bigr => i _ /=; rewrite linearZ !cfdotZr !cfdot_suml; congr (_ * _).
+apply: eq_bigr => j _ /=; rewrite linearZ !cfdotZl; congr (_ * _).
+rewrite -!((nth_map _ 0) tau) // defT; have [-> | neq_ji] := eqVneq j i.
+ by rewrite -!['[_]]((nth_map _ 0) cfnorm) ?eq_sz // eq_nST.
+have{oS} [/=/andP[_ uS] oS] := pairwise_orthogonalP oS.
+have{oT} [/=/andP[_ uT] oT] := pairwise_orthogonalP oT.
+by rewrite oS ?oT ?mem_nth ? nth_uniq ?eq_sz.
+Qed.
+
+Lemma isometry_raddf_inj U (tau : {additive 'CF(L) -> 'CF(G)}) :
+ {in U &, isometry tau} -> {in U &, forall u v, u - v \in U} ->
+ {in U &, injective tau}.
+Proof.
+move=> Itau linU phi psi Uphi Upsi /eqP; rewrite -subr_eq0 -raddfB.
+by rewrite -cfnorm_eq0 Itau ?linU // cfnorm_eq0 subr_eq0 => /eqP.
+Qed.
+
+Lemma opp_isometry : @isometry _ _ G G -%R.
+Proof. by move=> x y; rewrite cfdotNl cfdotNr opprK. Qed.
+
+End BuildIsometries.
+
+Section Restrict.
+
+Variables (gT : finGroupType) (A B : {set gT}).
+Local Notation H := <<A>>.
+Local Notation G := <<B>>.
+
+Fact cfRes_subproof (phi : 'CF(B)) :
+ is_class_fun H [ffun x => phi (if H \subset G then x else 1%g) *+ (x \in H)].
+Proof.
+apply: intro_class_fun => /= [x y Hx Hy | x /negbTE/=-> //].
+by rewrite Hx (groupJ Hx) //; case: subsetP => // sHG; rewrite cfunJgen ?sHG.
+Qed.
+Definition cfRes phi := Cfun 1 (cfRes_subproof phi).
+
+Lemma cfResE phi : A \subset B -> {in A, cfRes phi =1 phi}.
+Proof. by move=> sAB x Ax; rewrite cfunElock mem_gen ?genS. Qed.
+
+Lemma cfRes1 phi : cfRes phi 1%g = phi 1%g.
+Proof. by rewrite cfunElock if_same group1. Qed.
+
+Lemma cfRes_is_linear : linear cfRes.
+Proof.
+by move=> a phi psi; apply/cfunP=> x; rewrite !cfunElock mulrnAr mulrnDl.
+Qed.
+Canonical cfRes_additive := Additive cfRes_is_linear.
+Canonical cfRes_linear := Linear cfRes_is_linear.
+
+Lemma cfRes_cfun1 : cfRes 1 = 1.
+Proof.
+apply: cfun_in_genP => x Hx; rewrite cfunElock Hx !cfun1Egen Hx.
+by case: subsetP => [-> // | _]; rewrite group1.
+Qed.
+
+Lemma cfRes_is_multiplicative : multiplicative cfRes.
+Proof.
+split=> [phi psi|]; [apply/cfunP=> x | exact: cfRes_cfun1].
+by rewrite !cfunElock mulrnAr mulrnAl -mulrnA mulnb andbb.
+Qed.
+Canonical cfRes_rmorphism := AddRMorphism cfRes_is_multiplicative.
+Canonical cfRes_lrmorphism := [lrmorphism of cfRes].
+
+End Restrict.
+
+Arguments Scope cfRes [_ group_scope group_scope cfun_scope].
+Notation "''Res[' H , G ]" := (@cfRes _ H G) (only parsing) : ring_scope.
+Notation "''Res[' H ]" := 'Res[H, _] : ring_scope.
+Notation "''Res'" := 'Res[_] (only parsing) : ring_scope.
+
+Section MoreRestrict.
+
+Variables (gT : finGroupType) (G H : {group gT}).
+Implicit Types (A : {set gT}) (phi : 'CF(G)).
+
+Lemma cfResEout phi : ~~ (H \subset G) -> 'Res[H] phi = (phi 1%g)%:A.
+Proof.
+move/negPf=> not_sHG; apply/cfunP=> x.
+by rewrite cfunE cfun1E mulr_natr cfunElock !genGid not_sHG.
+Qed.
+
+Lemma cfResRes A phi :
+ A \subset H -> H \subset G -> 'Res[A] ('Res[H] phi) = 'Res[A] phi.
+Proof.
+move=> sAH sHG; apply/cfunP=> x; rewrite !cfunElock !genGid !gen_subG sAH sHG.
+by rewrite (subset_trans sAH) // -mulrnA mulnb -in_setI (setIidPr _) ?gen_subG.
+Qed.
+
+Lemma cfRes_id A psi : 'Res[A] psi = psi.
+Proof. by apply/cfun_in_genP=> x Ax; rewrite cfunElock Ax subxx. Qed.
+
+Lemma sub_cfker_Res A phi :
+ A \subset H -> A \subset cfker phi -> A \subset cfker ('Res[H, G] phi).
+Proof.
+move=> sAH kerA; apply/subsetP=> x Ax; have Hx := subsetP sAH x Ax.
+rewrite inE Hx; apply/forallP=> y; rewrite !cfunElock !genGid groupMl //.
+by rewrite !(fun_if phi) cfkerMl // (subsetP kerA).
+Qed.
+
+Lemma eq_cfker_Res phi : H \subset cfker phi -> cfker ('Res[H, G] phi) = H.
+Proof. by move=> kH; apply/eqP; rewrite eqEsubset cfker_sub sub_cfker_Res. Qed.
+
+Lemma cfRes_sub_ker phi : H \subset cfker phi -> 'Res[H, G] phi = (phi 1%g)%:A.
+Proof.
+move=> kerHphi; have sHG := subset_trans kerHphi (cfker_sub phi).
+apply/cfun_inP=> x Hx; have ker_x := subsetP kerHphi x Hx.
+by rewrite cfResE // cfunE cfun1E Hx mulr1 cfker1.
+Qed.
+
+Lemma cforder_Res phi : #['Res[H] phi]%CF %| #[phi]%CF.
+Proof. exact: cforder_rmorph. Qed.
+
+End MoreRestrict.
+
+Section Morphim.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+
+Section Main.
+
+Variable G : {group aT}.
+Implicit Type phi : 'CF(f @* G).
+
+Fact cfMorph_subproof phi :
+ is_class_fun <<G>>
+ [ffun x => phi (if G \subset D then f x else 1%g) *+ (x \in G)].
+Proof.
+rewrite genGid; apply: intro_class_fun => [x y Gx Gy | x /negPf-> //].
+rewrite Gx groupJ //; case subsetP => // sGD.
+by rewrite morphJ ?cfunJ ?mem_morphim ?sGD.
+Qed.
+Definition cfMorph phi := Cfun 1 (cfMorph_subproof phi).
+
+Lemma cfMorphE phi x : G \subset D -> x \in G -> cfMorph phi x = phi (f x).
+Proof. by rewrite cfunElock => -> ->. Qed.
+
+Lemma cfMorph1 phi : cfMorph phi 1%g = phi 1%g.
+Proof. by rewrite cfunElock morph1 if_same group1. Qed.
+
+Lemma cfMorphEout phi : ~~ (G \subset D) -> cfMorph phi = (phi 1%g)%:A.
+Proof.
+move/negPf=> not_sGD; apply/cfunP=> x; rewrite cfunE cfun1E mulr_natr.
+by rewrite cfunElock not_sGD.
+Qed.
+
+Lemma cfMorph_cfun1 : cfMorph 1 = 1.
+Proof.
+apply/cfun_inP=> x Gx; rewrite cfunElock !cfun1E Gx.
+by case: subsetP => [sGD | _]; rewrite ?group1 // mem_morphim ?sGD.
+Qed.
+
+Fact cfMorph_is_linear : linear cfMorph.
+Proof.
+by move=> a phi psi; apply/cfunP=> x; rewrite !cfunElock mulrnAr -mulrnDl.
+Qed.
+Canonical cfMorph_additive := Additive cfMorph_is_linear.
+Canonical cfMorph_linear := Linear cfMorph_is_linear.
+
+Fact cfMorph_is_multiplicative : multiplicative cfMorph.
+Proof.
+split=> [phi psi|]; [apply/cfunP=> x | exact: cfMorph_cfun1].
+by rewrite !cfunElock mulrnAr mulrnAl -mulrnA mulnb andbb.
+Qed.
+Canonical cfMorph_rmorphism := AddRMorphism cfMorph_is_multiplicative.
+Canonical cfMorph_lrmorphism := [lrmorphism of cfMorph].
+
+Hypothesis sGD : G \subset D.
+
+Lemma cfMorph_inj : injective cfMorph.
+Proof.
+move=> phi1 phi2 eq_phi; apply/cfun_inP=> _ /morphimP[x Dx Gx ->].
+by rewrite -!cfMorphE // eq_phi.
+Qed.
+
+Lemma cfMorph_eq1 phi : (cfMorph phi == 1) = (phi == 1).
+Proof. by apply: rmorph_eq1; apply: cfMorph_inj. Qed.
+
+Lemma cfker_morph phi : cfker (cfMorph phi) = G :&: f @*^-1 (cfker phi).
+Proof.
+apply/setP=> x; rewrite !inE; apply: andb_id2l => Gx.
+have Dx := subsetP sGD x Gx; rewrite Dx mem_morphim //=.
+apply/forallP/forallP=> Kx y.
+ have [{y} /morphimP[y Dy Gy ->] | fG'y] := boolP (y \in f @* G).
+ by rewrite -morphM // -!(cfMorphE phi) ?groupM.
+ by rewrite !cfun0 ?groupMl // mem_morphim.
+have [Gy | G'y] := boolP (y \in G); last by rewrite !cfun0 ?groupMl.
+by rewrite !cfMorphE ?groupM ?morphM // (subsetP sGD).
+Qed.
+
+Lemma cfker_morph_im phi : f @* cfker (cfMorph phi) = cfker phi.
+Proof. by rewrite cfker_morph // morphim_setIpre (setIidPr (cfker_sub _)). Qed.
+
+Lemma sub_cfker_morph phi (A : {set aT}) :
+ (A \subset cfker (cfMorph phi)) = (A \subset G) && (f @* A \subset cfker phi).
+Proof.
+rewrite cfker_morph // subsetI; apply: andb_id2l => sAG.
+by rewrite sub_morphim_pre // (subset_trans sAG).
+Qed.
+
+Lemma sub_morphim_cfker phi (A : {set aT}) :
+ A \subset G -> (f @* A \subset cfker phi) = (A \subset cfker (cfMorph phi)).
+Proof. by move=> sAG; rewrite sub_cfker_morph ?sAG. Qed.
+
+Lemma cforder_morph phi : #[cfMorph phi]%CF = #[phi]%CF.
+Proof. by apply: cforder_inj_rmorph; apply: cfMorph_inj. Qed.
+
+End Main.
+
+Lemma cfResMorph (G H : {group aT}) (phi : 'CF(f @* G)) :
+ H \subset G -> G \subset D -> 'Res (cfMorph phi) = cfMorph ('Res[f @* H] phi).
+Proof.
+move=> sHG sGD; have sHD := subset_trans sHG sGD.
+apply/cfun_inP=> x Hx; have [Gx Dx] := (subsetP sHG x Hx, subsetP sHD x Hx).
+by rewrite !(cfMorphE, cfResE) ?morphimS ?mem_morphim //.
+Qed.
+
+End Morphim.
+
+Prenex Implicits cfMorph.
+
+Section Isomorphism.
+
+Variables (aT rT : finGroupType) (G : {group aT}) (f : {morphism G >-> rT}).
+Variable R : {group rT}.
+
+Hypothesis isoGR : isom G R f.
+
+Let defR := isom_im isoGR.
+Local Notation G1 := (isom_inv isoGR @* R).
+Let defG : G1 = G := isom_im (isom_sym isoGR).
+
+Fact cfIsom_key : unit. Proof. by []. Qed.
+Definition cfIsom :=
+ locked_with cfIsom_key (cfMorph \o 'Res[G1] : 'CF(G) -> 'CF(R)).
+Canonical cfIsom_unlockable := [unlockable of cfIsom].
+
+Lemma cfIsomE phi x : x \in G -> cfIsom phi (f x) = phi x.
+Proof.
+move=> Gx; rewrite unlock cfMorphE //= /restrm ?defG ?cfRes_id ?invmE //.
+by rewrite -defR mem_morphim.
+Qed.
+
+Lemma cfIsom1 phi : cfIsom phi 1%g = phi 1%g.
+Proof. by rewrite -(morph1 f) cfIsomE. Qed.
+
+Canonical cfIsom_additive := [additive of cfIsom].
+Canonical cfIsom_linear := [linear of cfIsom].
+Canonical cfIsom_rmorphism := [rmorphism of cfIsom].
+Canonical cfIsom_lrmorphism := [lrmorphism of cfIsom].
+Lemma cfIsom_cfun1 : cfIsom 1 = 1. Proof. exact: rmorph1. Qed.
+
+Lemma cfker_isom phi : cfker (cfIsom phi) = f @* cfker phi.
+Proof.
+rewrite unlock cfker_morph // defG cfRes_id morphpre_restrm morphpre_invm.
+by rewrite -defR !morphimIim.
+Qed.
+
+End Isomorphism.
+
+Prenex Implicits cfIsom.
+
+Section InvMorphism.
+
+Variables (aT rT : finGroupType) (G : {group aT}) (f : {morphism G >-> rT}).
+Variable R : {group rT}.
+
+Hypothesis isoGR : isom G R f.
+
+Lemma cfIsomK : cancel (cfIsom isoGR) (cfIsom (isom_sym isoGR)).
+Proof.
+move=> phi; apply/cfun_inP=> x Gx; rewrite -{1}(invmE (isom_inj isoGR) Gx).
+by rewrite !cfIsomE // -(isom_im isoGR) mem_morphim.
+Qed.
+
+Lemma cfIsomKV : cancel (cfIsom (isom_sym isoGR)) (cfIsom isoGR).
+Proof.
+move=> phi; apply/cfun_inP=> y Ry; pose injGR := isom_inj isoGR.
+rewrite -{1}[y](invmK injGR) ?(isom_im isoGR) //.
+suffices /morphpreP[fGy Gf'y]: y \in invm injGR @*^-1 G by rewrite !cfIsomE.
+by rewrite morphpre_invm (isom_im isoGR).
+Qed.
+
+Lemma cfIsom_inj : injective (cfIsom isoGR). Proof. exact: can_inj cfIsomK. Qed.
+
+Lemma cfIsom_eq1 phi : (cfIsom isoGR phi == 1) = (phi == 1).
+Proof. by apply: rmorph_eq1; apply: cfIsom_inj. Qed.
+
+Lemma cforder_isom phi : #[cfIsom isoGR phi]%CF = #[phi]%CF.
+Proof. exact: cforder_inj_rmorph cfIsom_inj. Qed.
+
+End InvMorphism.
+
+Implicit Arguments cfIsom_inj [aT rT G R f x1 x2].
+
+Section Coset.
+
+Variables (gT : finGroupType) (G : {group gT}) (B : {set gT}).
+Implicit Type rT : finGroupType.
+Local Notation H := <<B>>%g.
+
+Definition cfMod : 'CF(G / B) -> 'CF(G) := cfMorph.
+
+Definition ffun_Quo (phi : 'CF(G)) :=
+ [ffun Hx : coset_of B =>
+ phi (if B \subset cfker phi then repr Hx else 1%g) *+ (Hx \in G / B)%g].
+Fact cfQuo_subproof phi : is_class_fun <<G / B>> (ffun_Quo phi).
+Proof.
+rewrite genGid; apply: intro_class_fun => [|Hx /negPf-> //].
+move=> _ _ /morphimP[x Nx Gx ->] /morphimP[z Nz Gz ->].
+rewrite -morphJ ?mem_morphim ?val_coset_prim ?groupJ //= -gen_subG.
+case: subsetP => // KphiH; do 2!case: repr_rcosetP => _ /KphiH/cfkerMl->.
+by rewrite cfunJ.
+Qed.
+Definition cfQuo phi := Cfun 1 (cfQuo_subproof phi).
+
+Local Notation "phi / 'B'" := (cfQuo phi) (at level 40) : cfun_scope.
+Local Notation "phi %% 'B'" := (cfMod phi) (at level 40) : cfun_scope.
+
+(* We specialize the cfMorph lemmas to cfMod by strengthening the domain *)
+(* condition G \subset 'N(H) to H <| G; the cfMorph lemmas can be used if the *)
+(* stronger results are needed. *)
+
+Lemma cfModE phi x : B <| G -> x \in G -> (phi %% B)%CF x = phi (coset B x).
+Proof. by move/normal_norm=> nBG; exact: cfMorphE. Qed.
+
+Lemma cfMod1 phi : (phi %% B)%CF 1%g = phi 1%g. Proof. exact: cfMorph1. Qed.
+
+Canonical cfMod_additive := [additive of cfMod].
+Canonical cfMod_rmorphism := [rmorphism of cfMod].
+Canonical cfMod_linear := [linear of cfMod].
+Canonical cfMod_lrmorphism := [lrmorphism of cfMod].
+
+Lemma cfMod_cfun1 : (1 %% B)%CF = 1. Proof. exact: rmorph1. Qed.
+
+Lemma cfker_mod phi : B <| G -> B \subset cfker (phi %% B).
+Proof.
+case/andP=> sBG nBG; rewrite cfker_morph // subsetI sBG.
+apply: subset_trans _ (ker_sub_pre _ _); rewrite ker_coset_prim subsetI.
+by rewrite (subset_trans sBG nBG) sub_gen.
+Qed.
+
+(* Note that cfQuo is nondegenerate even when G does not normalize B. *)
+
+Lemma cfQuoEnorm (phi : 'CF(G)) x :
+ B \subset cfker phi -> x \in 'N_G(B) -> (phi / B)%CF (coset B x) = phi x.
+Proof.
+rewrite cfunElock -gen_subG => sHK /setIP[Gx nHx]; rewrite sHK /=.
+rewrite mem_morphim // val_coset_prim //.
+by case: repr_rcosetP => _ /(subsetP sHK)/cfkerMl->.
+Qed.
+
+Lemma cfQuoE (phi : 'CF(G)) x :
+ B <| G -> B \subset cfker phi -> x \in G -> (phi / B)%CF (coset B x) = phi x.
+Proof. by case/andP=> _ nBG sBK Gx; rewrite cfQuoEnorm // (setIidPl _). Qed.
+
+Lemma cfQuo1 (phi : 'CF(G)) : (phi / B)%CF 1%g = phi 1%g.
+Proof. by rewrite cfunElock repr_coset1 group1 if_same. Qed.
+
+Lemma cfQuoEout (phi : 'CF(G)) :
+ ~~ (B \subset cfker phi) -> (phi / B)%CF = (phi 1%g)%:A.
+Proof.
+move/negPf=> not_kerB; apply/cfunP=> x; rewrite cfunE cfun1E mulr_natr.
+by rewrite cfunElock not_kerB.
+Qed.
+
+(* cfQuo is only linear on the class functions that have H in their kernel. *)
+
+Lemma cfQuo_cfun1 : (1 / B)%CF = 1.
+Proof.
+apply/cfun_inP=> Hx G_Hx; rewrite cfunElock !cfun1E G_Hx cfker_cfun1 -gen_subG.
+have [x nHx Gx ->] := morphimP G_Hx.
+case: subsetP=> [sHG | _]; last by rewrite group1.
+by rewrite val_coset_prim //; case: repr_rcosetP => y /sHG/groupM->.
+Qed.
+
+(* Cancellation properties *)
+
+Lemma cfModK : B <| G -> cancel cfMod cfQuo.
+Proof.
+move=> nsBG phi; apply/cfun_inP=> _ /morphimP[x Nx Gx ->] //.
+by rewrite cfQuoE ?cfker_mod ?cfModE.
+Qed.
+
+Lemma cfQuoK :
+ B <| G -> forall phi, B \subset cfker phi -> (phi / B %% B)%CF = phi.
+Proof.
+by move=> nsHG phi sHK; apply/cfun_inP=> x Gx; rewrite cfModE ?cfQuoE.
+Qed.
+
+Lemma cfMod_eq1 psi : B <| G -> (psi %% B == 1)%CF = (psi == 1).
+Proof. by move/cfModK/can_eq <-; rewrite rmorph1. Qed.
+
+Lemma cfQuo_eq1 phi :
+ B <| G -> B \subset cfker phi -> (phi / B == 1)%CF = (phi == 1).
+Proof. by move=> nsBG kerH; rewrite -cfMod_eq1 // cfQuoK. Qed.
+
+End Coset.
+
+Arguments Scope cfQuo [_ Group_scope group_scope cfun_scope].
+Arguments Scope cfMod [_ Group_scope group_scope cfun_scope].
+Prenex Implicits cfMod.
+Notation "phi / H" := (cfQuo H phi) : cfun_scope.
+Notation "phi %% H" := (@cfMod _ _ H phi) : cfun_scope.
+
+Section MoreCoset.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Implicit Types (H K : {group gT}) (phi : 'CF(G)).
+
+Lemma cfResMod H K (psi : 'CF(G / K)) :
+ H \subset G -> K <| G -> ('Res (psi %% K) = 'Res[H / K] psi %% K)%CF.
+Proof. by move=> sHG /andP[_]; apply: cfResMorph. Qed.
+
+Lemma quotient_cfker_mod (A : {set gT}) K (psi : 'CF(G / K)) :
+ K <| G -> (cfker (psi %% K) / K)%g = cfker psi.
+Proof. by case/andP=> _ /cfker_morph_im <-. Qed.
+
+Lemma sub_cfker_mod (A : {set gT}) K (psi : 'CF(G / K)) :
+ K <| G -> A \subset 'N(K) ->
+ (A \subset cfker (psi %% K)) = (A / K \subset cfker psi)%g.
+Proof.
+by move=> nsKG nKA; rewrite -(quotientSGK nKA) ?quotient_cfker_mod ?cfker_mod.
+Qed.
+
+Lemma cfker_quo H phi :
+ H <| G -> H \subset cfker (phi) -> cfker (phi / H) = (cfker phi / H)%g.
+Proof.
+move=> nsHG /cfQuoK {2}<- //; have [sHG nHG] := andP nsHG.
+by rewrite cfker_morph 1?quotientGI // cosetpreK (setIidPr _) ?cfker_sub.
+Qed.
+
+Lemma cfQuoEker phi x :
+ x \in G -> (phi / cfker phi)%CF (coset (cfker phi) x) = phi x.
+Proof. by move/cfQuoE->; rewrite ?cfker_normal. Qed.
+
+Lemma cfaithful_quo phi : cfaithful (phi / cfker phi).
+Proof. by rewrite cfaithfulE cfker_quo ?cfker_normal ?trivg_quotient. Qed.
+
+(* Note that there is no requirement that K be normal in H or G. *)
+Lemma cfResQuo H K phi :
+ K \subset cfker phi -> K \subset H -> H \subset G ->
+ ('Res[H / K] (phi / K) = 'Res[H] phi / K)%CF.
+Proof.
+move=> kerK sKH sHG; apply/cfun_inP=> xb Hxb; rewrite cfResE ?quotientS //.
+have{xb Hxb} [x nKx Hx ->] := morphimP Hxb.
+by rewrite !cfQuoEnorm ?cfResE ?sub_cfker_Res // inE ?Hx ?(subsetP sHG).
+Qed.
+
+Lemma cfQuoInorm K phi :
+ K \subset cfker phi -> (phi / K)%CF = 'Res ('Res['N_G(K)] phi / K)%CF.
+Proof.
+move=> kerK; rewrite -cfResQuo ?subsetIl ?quotientInorm ?cfRes_id //.
+by rewrite subsetI normG (subset_trans kerK) ?cfker_sub.
+Qed.
+
+Lemma cforder_mod H (psi : 'CF(G / H)) : H <| G -> #[psi %% H]%CF = #[psi]%CF.
+Proof. by move/cfModK/can_inj/cforder_inj_rmorph->. Qed.
+
+Lemma cforder_quo H phi :
+ H <| G -> H \subset cfker phi -> #[phi / H]%CF = #[phi]%CF.
+Proof. by move=> nsHG kerHphi; rewrite -cforder_mod ?cfQuoK. Qed.
+
+End MoreCoset.
+
+Section Product.
+
+Variable (gT : finGroupType) (G : {group gT}).
+
+Lemma cfunM_onI A B phi psi :
+ phi \in 'CF(G, A) -> psi \in 'CF(G, B) -> phi * psi \in 'CF(G, A :&: B).
+Proof.
+rewrite !cfun_onE => Aphi Bpsi; apply/subsetP=> x; rewrite !inE cfunE mulf_eq0.
+by case/norP=> /(subsetP Aphi)-> /(subsetP Bpsi).
+Qed.
+
+Lemma cfunM_on A phi psi :
+ phi \in 'CF(G, A) -> psi \in 'CF(G, A) -> phi * psi \in 'CF(G, A).
+Proof. by move=> Aphi Bpsi; rewrite -[A]setIid cfunM_onI. Qed.
+
+End Product.
+
+Section SDproduct.
+
+Variables (gT : finGroupType) (G K H : {group gT}).
+Hypothesis defG : K ><| H = G.
+
+Fact cfSdprodKey : unit. Proof. by []. Qed.
+
+Definition cfSdprod :=
+ locked_with cfSdprodKey
+ (cfMorph \o cfIsom (tagged (sdprod_isom defG)) : 'CF(H) -> 'CF(G)).
+Canonical cfSdprod_unlockable := [unlockable of cfSdprod].
+
+Canonical cfSdprod_additive := [additive of cfSdprod].
+Canonical cfSdprod_linear := [linear of cfSdprod].
+Canonical cfSdprod_rmorphism := [rmorphism of cfSdprod].
+Canonical cfSdprod_lrmorphism := [lrmorphism of cfSdprod].
+
+Lemma cfSdprod1 phi : cfSdprod phi 1%g = phi 1%g.
+Proof. by rewrite unlock /= cfMorph1 cfIsom1. Qed.
+
+Let nsKG : K <| G. Proof. by have [] := sdprod_context defG. Qed.
+Let sHG : H \subset G. Proof. by have [] := sdprod_context defG. Qed.
+Let sKG : K \subset G. Proof. by have [] := andP nsKG. Qed.
+
+Lemma cfker_sdprod phi : K \subset cfker (cfSdprod phi).
+Proof. by rewrite unlock_with cfker_mod. Qed.
+
+Lemma cfSdprodEr phi : {in H, cfSdprod phi =1 phi}.
+Proof. by move=> y Hy; rewrite unlock cfModE ?cfIsomE ?(subsetP sHG). Qed.
+
+Lemma cfSdprodE phi : {in K & H, forall x y, cfSdprod phi (x * y)%g = phi y}.
+Proof.
+by move=> x y Kx Hy; rewrite /= cfkerMl ?(subsetP (cfker_sdprod _)) ?cfSdprodEr.
+Qed.
+
+Lemma cfSdprodK : cancel cfSdprod 'Res[H].
+Proof. by move=> phi; apply/cfun_inP=> x Hx; rewrite cfResE ?cfSdprodEr. Qed.
+
+Lemma cfSdprod_inj : injective cfSdprod. Proof. exact: can_inj cfSdprodK. Qed.
+
+Lemma cfSdprod_eq1 phi : (cfSdprod phi == 1) = (phi == 1).
+Proof. exact: rmorph_eq1 cfSdprod_inj. Qed.
+
+Lemma cfRes_sdprodK phi : K \subset cfker phi -> cfSdprod ('Res[H] phi) = phi.
+Proof.
+move=> kerK; apply/cfun_inP=> _ /(mem_sdprod defG)[x [y [Kx Hy -> _]]].
+by rewrite cfSdprodE // cfResE // cfkerMl ?(subsetP kerK).
+Qed.
+
+Lemma sdprod_cfker phi : K ><| cfker phi = cfker (cfSdprod phi).
+Proof.
+have [skerH [_ _ nKH tiKH]] := (cfker_sub phi, sdprodP defG).
+rewrite unlock cfker_morph ?normal_norm // cfker_isom restrmEsub //=.
+rewrite -(sdprod_modl defG) ?sub_cosetpre //=; congr (_ ><| _).
+by rewrite quotientK ?(subset_trans skerH) // -group_modr //= setIC tiKH mul1g.
+Qed.
+
+Lemma cforder_sdprod phi : #[cfSdprod phi]%CF = #[phi]%CF.
+Proof. by apply: cforder_inj_rmorph cfSdprod_inj. Qed.
+
+End SDproduct.
+
+Section DProduct.
+
+Variables (gT : finGroupType) (G K H : {group gT}).
+Hypothesis KxH : K \x H = G.
+
+Lemma reindex_dprod R idx (op : Monoid.com_law idx) (F : gT -> R) :
+ \big[op/idx]_(g in G) F g =
+ \big[op/idx]_(k in K) \big[op/idx]_(h in H) F (k * h)%g.
+Proof.
+have /mulgmP/misomP[fM /isomP[injf im_f]] := KxH.
+rewrite pair_big_dep -im_f morphimEdom big_imset; last exact/injmP.
+by apply: eq_big => [][x y]; rewrite ?inE.
+Qed.
+
+Definition cfDprodr := cfSdprod (dprodWsd KxH).
+Definition cfDprodl := cfSdprod (dprodWsdC KxH).
+Definition cfDprod phi psi := cfDprodl phi * cfDprodr psi.
+
+Canonical cfDprodl_additive := [additive of cfDprodl].
+Canonical cfDprodl_linear := [linear of cfDprodl].
+Canonical cfDprodl_rmorphism := [rmorphism of cfDprodl].
+Canonical cfDprodl_lrmorphism := [lrmorphism of cfDprodl].
+Canonical cfDprodr_additive := [additive of cfDprodr].
+Canonical cfDprodr_linear := [linear of cfDprodr].
+Canonical cfDprodr_rmorphism := [rmorphism of cfDprodr].
+Canonical cfDprodr_lrmorphism := [lrmorphism of cfDprodr].
+
+Lemma cfDprodl1 phi : cfDprodl phi 1%g = phi 1%g. Proof. exact: cfSdprod1. Qed.
+Lemma cfDprodr1 psi : cfDprodr psi 1%g = psi 1%g. Proof. exact: cfSdprod1. Qed.
+Lemma cfDprod1 phi psi : cfDprod phi psi 1%g = phi 1%g * psi 1%g.
+Proof. by rewrite cfunE /= !cfSdprod1. Qed.
+
+Lemma cfDprodl_eq1 phi : (cfDprodl phi == 1) = (phi == 1).
+Proof. exact: cfSdprod_eq1. Qed.
+Lemma cfDprodr_eq1 psi : (cfDprodr psi == 1) = (psi == 1).
+Proof. exact: cfSdprod_eq1. Qed.
+
+Lemma cfDprod_cfun1r phi : cfDprod phi 1 = cfDprodl phi.
+Proof. by rewrite /cfDprod rmorph1 mulr1. Qed.
+Lemma cfDprod_cfun1l psi : cfDprod 1 psi = cfDprodr psi.
+Proof. by rewrite /cfDprod rmorph1 mul1r. Qed.
+Lemma cfDprod_cfun1 : cfDprod 1 1 = 1.
+Proof. by rewrite cfDprod_cfun1l rmorph1. Qed.
+Lemma cfDprod_split phi psi : cfDprod phi psi = cfDprod phi 1 * cfDprod 1 psi.
+Proof. by rewrite cfDprod_cfun1l cfDprod_cfun1r. Qed.
+
+Let nsKG : K <| G. Proof. by have [] := dprod_normal2 KxH. Qed.
+Let nsHG : H <| G. Proof. by have [] := dprod_normal2 KxH. Qed.
+Let cKH : H \subset 'C(K). Proof. by have [] := dprodP KxH. Qed.
+Let sKG := normal_sub nsKG.
+Let sHG := normal_sub nsHG.
+
+Lemma cfDprodlK : cancel cfDprodl 'Res[K]. Proof. exact: cfSdprodK. Qed.
+Lemma cfDprodrK : cancel cfDprodr 'Res[H]. Proof. exact: cfSdprodK. Qed.
+
+Lemma cfker_dprodl phi : cfker phi \x H = cfker (cfDprodl phi).
+Proof.
+by rewrite dprodC -sdprod_cfker dprodEsd // centsC (centsS (cfker_sub _)).
+Qed.
+
+Lemma cfker_dprodr psi : K \x cfker psi = cfker (cfDprodr psi).
+Proof. by rewrite -sdprod_cfker dprodEsd // (subset_trans (cfker_sub _)). Qed.
+
+Lemma cfDprodEl phi : {in K & H, forall k h, cfDprodl phi (k * h)%g = phi k}.
+Proof. by move=> k h Kk Hh /=; rewrite -(centsP cKH) // cfSdprodE. Qed.
+
+Lemma cfDprodEr psi : {in K & H, forall k h, cfDprodr psi (k * h)%g = psi h}.
+Proof. exact: cfSdprodE. Qed.
+
+Lemma cfDprodE phi psi :
+ {in K & H, forall h k, cfDprod phi psi (h * k)%g = phi h * psi k}.
+Proof. by move=> k h Kk Hh /=; rewrite cfunE cfDprodEl ?cfDprodEr. Qed.
+
+Lemma cfDprod_Resl phi psi : 'Res[K] (cfDprod phi psi) = psi 1%g *: phi.
+Proof.
+by apply/cfun_inP=> x Kx; rewrite cfunE cfResE // -{1}[x]mulg1 mulrC cfDprodE.
+Qed.
+
+Lemma cfDprod_Resr phi psi : 'Res[H] (cfDprod phi psi) = phi 1%g *: psi.
+Proof.
+by apply/cfun_inP=> y Hy; rewrite cfunE cfResE // -{1}[y]mul1g cfDprodE.
+Qed.
+
+Lemma cfDprodKl (psi : 'CF(H)) : psi 1%g = 1 -> cancel (cfDprod^~ psi) 'Res.
+Proof. by move=> psi1 phi; rewrite cfDprod_Resl psi1 scale1r. Qed.
+
+Lemma cfDprodKr (phi : 'CF(K)) : phi 1%g = 1 -> cancel (cfDprod phi) 'Res.
+Proof. by move=> phi1 psi; rewrite cfDprod_Resr phi1 scale1r. Qed.
+
+(* Note that equality holds here iff either cfker phi = K and cfker psi = H, *)
+(* or else phi != 0, psi != 0 and coprime #|K : cfker phi| #|H : cfker phi|. *)
+Lemma cfker_dprod phi psi :
+ cfker phi <*> cfker psi \subset cfker (cfDprod phi psi).
+Proof.
+rewrite -genM_join gen_subG; apply/subsetP=> _ /mulsgP[x y kKx kHy ->] /=.
+have [[Kx _] [Hy _]] := (setIdP kKx, setIdP kHy).
+have Gxy: (x * y)%g \in G by rewrite -(dprodW KxH) mem_mulg.
+rewrite inE Gxy; apply/forallP=> g.
+have [Gg | G'g] := boolP (g \in G); last by rewrite !cfun0 1?groupMl.
+have{g Gg} [k [h [Kk Hh -> _]]] := mem_dprod KxH Gg.
+rewrite mulgA -(mulgA x) (centsP cKH y) // mulgA -mulgA !cfDprodE ?groupM //.
+by rewrite !cfkerMl.
+Qed.
+
+Lemma cfdot_dprod phi1 phi2 psi1 psi2 :
+ '[cfDprod phi1 psi1, cfDprod phi2 psi2] = '[phi1, phi2] * '[psi1, psi2].
+Proof.
+rewrite !cfdotE mulrCA -mulrA mulrCA mulrA -invfM -natrM (dprod_card KxH).
+congr (_ * _); rewrite big_distrl reindex_dprod /=; apply: eq_bigr => k Kk.
+rewrite big_distrr; apply: eq_bigr => h Hh /=.
+by rewrite mulrCA -mulrA -rmorphM mulrCA mulrA !cfDprodE.
+Qed.
+
+Lemma cfDprodl_iso : isometry cfDprodl.
+Proof.
+by move=> phi1 phi2; rewrite -!cfDprod_cfun1r cfdot_dprod cfnorm1 mulr1.
+Qed.
+
+Lemma cfDprodr_iso : isometry cfDprodr.
+Proof.
+by move=> psi1 psi2; rewrite -!cfDprod_cfun1l cfdot_dprod cfnorm1 mul1r.
+Qed.
+
+Lemma cforder_dprodl phi : #[cfDprodl phi]%CF = #[phi]%CF.
+Proof. exact: cforder_sdprod. Qed.
+
+Lemma cforder_dprodr psi : #[cfDprodr psi]%CF = #[psi]%CF.
+Proof. exact: cforder_sdprod. Qed.
+
+End DProduct.
+
+Lemma cfDprodC (gT : finGroupType) (G K H : {group gT})
+ (KxH : K \x H = G) (HxK : H \x K = G) chi psi :
+ cfDprod KxH chi psi = cfDprod HxK psi chi.
+Proof.
+rewrite /cfDprod mulrC.
+by congr (_ * _); congr (cfSdprod _ _); apply: eq_irrelevance.
+Qed.
+
+Section Bigdproduct.
+
+Variables (gT : finGroupType) (I : finType) (P : pred I).
+Variables (A : I -> {group gT}) (G : {group gT}).
+Hypothesis defG : \big[dprod/1%g]_(i | P i) A i = G.
+
+Let sAG i : P i -> A i \subset G.
+Proof. by move=> Pi; rewrite -(bigdprodWY defG) (bigD1 i) ?joing_subl. Qed.
+
+Fact cfBigdprodi_subproof i :
+ gval (if P i then A i else 1%G) \x <<\bigcup_(j | P j && (j != i)) A j>> = G.
+Proof.
+have:= defG; rewrite fun_if big_mkcond (bigD1 i) // -big_mkcondl /= => defGi.
+by have [[_ Gi' _ defGi']] := dprodP defGi; rewrite (bigdprodWY defGi') -defGi'.
+Qed.
+Definition cfBigdprodi i := cfDprodl (cfBigdprodi_subproof i) \o 'Res[_, A i].
+
+Canonical cfBigdprodi_additive i := [additive of @cfBigdprodi i].
+Canonical cfBigdprodi_linear i := [linear of @cfBigdprodi i].
+Canonical cfBigdprodi_rmorphism i := [rmorphism of @cfBigdprodi i].
+Canonical cfBigdprodi_lrmorphism i := [lrmorphism of @cfBigdprodi i].
+
+Lemma cfBigdprodi1 i (phi : 'CF(A i)) : cfBigdprodi phi 1%g = phi 1%g.
+Proof. by rewrite cfDprodl1 cfRes1. Qed.
+
+Lemma cfBigdprodi_eq1 i (phi : 'CF(A i)) :
+ P i -> (cfBigdprodi phi == 1) = (phi == 1).
+Proof. by move=> Pi; rewrite cfSdprod_eq1 Pi cfRes_id. Qed.
+
+Lemma cfBigdprodiK i : P i -> cancel (@cfBigdprodi i) 'Res[A i].
+Proof.
+move=> Pi phi; have:= cfDprodlK (cfBigdprodi_subproof i) ('Res phi).
+by rewrite -[cfDprodl _ _]/(cfBigdprodi phi) Pi cfRes_id.
+Qed.
+
+Lemma cfBigdprodi_inj i : P i -> injective (@cfBigdprodi i).
+Proof. by move/cfBigdprodiK; apply: can_inj. Qed.
+
+Lemma cfBigdprodEi i (phi : 'CF(A i)) x :
+ P i -> (forall j, P j -> x j \in A j) ->
+ cfBigdprodi phi (\prod_(j | P j) x j)%g = phi (x i).
+Proof.
+set r := enum P => Pi /forall_inP; have r_i: i \in r by rewrite mem_enum.
+have:= bigdprodWcp defG; rewrite -big_andE -!(big_filter _ P) filter_index_enum.
+rewrite -/r big_all => defGr /allP Ax.
+rewrite (perm_bigcprod defGr Ax (perm_to_rem r_i)) big_cons cfDprodEl ?Pi //.
+- by rewrite cfRes_id.
+- by rewrite Ax.
+rewrite big_seq group_prod // => j; rewrite mem_rem_uniq ?enum_uniq //.
+case/andP=> i'j /= r_j; apply/mem_gen/bigcupP; exists j; last exact: Ax.
+by rewrite -[P j](mem_enum P) r_j.
+Qed.
+
+Lemma cfBigdprodi_iso i : P i -> isometry (@cfBigdprodi i).
+Proof. by move=> Pi phi psi; rewrite cfDprodl_iso Pi !cfRes_id. Qed.
+
+Definition cfBigdprod (phi : forall i, 'CF(A i)) :=
+ \prod_(i | P i) cfBigdprodi (phi i).
+
+Lemma cfBigdprodE phi x :
+ (forall i, P i -> x i \in A i) ->
+ cfBigdprod phi (\prod_(i | P i) x i)%g = \prod_(i | P i) phi i (x i).
+Proof.
+move=> Ax; rewrite prod_cfunE; last by rewrite -(bigdprodW defG) mem_prodg.
+by apply: eq_bigr => i Pi; rewrite cfBigdprodEi.
+Qed.
+
+Lemma cfBigdprod1 phi : cfBigdprod phi 1%g = \prod_(i | P i) phi i 1%g.
+Proof. by rewrite prod_cfunE //; apply/eq_bigr=> i _; apply: cfBigdprodi1. Qed.
+
+Lemma cfBigdprodK phi (Phi := cfBigdprod phi) i (a := phi i 1%g / Phi 1%g) :
+ Phi 1%g != 0 -> P i -> a != 0 /\ a *: 'Res[A i] Phi = phi i.
+Proof.
+move=> nzPhi Pi; split.
+ rewrite mulf_neq0 ?invr_eq0 // (contraNneq _ nzPhi) // => phi_i0.
+ by rewrite cfBigdprod1 (bigD1 i) //= phi_i0 mul0r.
+apply/cfun_inP=> x Aix; rewrite cfunE cfResE ?sAG // mulrAC.
+have {1}->: x = (\prod_(j | P j) (if j == i then x else 1))%g.
+ rewrite -big_mkcondr (big_pred1 i) ?eqxx // => j /=.
+ by apply: andb_idl => /eqP->.
+rewrite cfBigdprodE => [|j _]; last by case: eqP => // ->.
+apply: canLR (mulfK nzPhi) _; rewrite cfBigdprod1 !(bigD1 i Pi) /= eqxx.
+by rewrite mulrCA !mulrA; congr (_ * _); apply: eq_bigr => j /andP[_ /negPf->].
+Qed.
+
+Lemma cfdot_bigdprod phi psi :
+ '[cfBigdprod phi, cfBigdprod psi] = \prod_(i | P i) '[phi i, psi i].
+Proof.
+apply: canLR (mulKf (neq0CG G)) _; rewrite -(bigdprod_card defG).
+rewrite (big_morph _ (@natrM _) (erefl _)) -big_split /=.
+rewrite (eq_bigr _ (fun i _ => mulVKf (neq0CG _) _)) (big_distr_big_dep 1%g) /=.
+set F := pfamily _ _ _; pose h (f : {ffun I -> gT}) := (\prod_(i | P i) f i)%g.
+pose is_hK x f := forall f1, (f1 \in F) && (h f1 == x) = (f == f1).
+have /fin_all_exists[h1 Dh1] x: exists f, x \in G -> is_hK x f.
+ case Gx: (x \in G); last by exists [ffun _ => x].
+ have [f [Af fK Uf]] := mem_bigdprod defG Gx.
+ exists [ffun i => if P i then f i else 1%g] => _ f1.
+ apply/andP/eqP=> [[/pfamilyP[Pf1 Af1] /eqP Dx] | <-].
+ by apply/ffunP=> i; rewrite ffunE; case: ifPn => [/Uf-> | /(supportP Pf1)].
+ split; last by rewrite fK; apply/eqP/eq_bigr=> i Pi; rewrite ffunE Pi.
+ by apply/familyP=> i; rewrite ffunE !unfold_in; case: ifP => //= /Af.
+rewrite (reindex_onto h h1) /= => [|x /Dh1/(_ (h1 x))]; last first.
+ by rewrite eqxx => /andP[_ /eqP].
+apply/eq_big => [f | f /andP[/Dh1<- /andP[/pfamilyP[_ Af] _]]]; last first.
+ by rewrite !cfBigdprodE // rmorph_prod -big_split /=.
+apply/idP/idP=> [/andP[/Dh1<-] | Ff]; first by rewrite eqxx andbT.
+have /pfamilyP[_ Af] := Ff; suffices Ghf: h f \in G by rewrite -Dh1 ?Ghf ?Ff /=.
+by apply/group_prod=> i Pi; rewrite (subsetP (sAG Pi)) ?Af.
+Qed.
+
+End Bigdproduct.
+
+Section MorphIsometry.
+
+Variable gT : finGroupType.
+Implicit Types (D G H K : {group gT}) (aT rT : finGroupType).
+
+Lemma cfMorph_iso aT rT (G D : {group aT}) (f : {morphism D >-> rT}) :
+ G \subset D -> isometry (cfMorph : 'CF(f @* G) -> 'CF(G)).
+Proof.
+move=> sGD phi psi; rewrite !cfdotE card_morphim (setIidPr sGD).
+rewrite -(LagrangeI G ('ker f)) /= mulnC natrM invfM -mulrA.
+congr (_ * _); apply: (canLR (mulKf (neq0CG _))).
+rewrite mulr_sumr (partition_big_imset f) /= -morphimEsub //.
+apply: eq_bigr => _ /morphimP[x Dx Gx ->].
+rewrite -(card_rcoset _ x) mulr_natl -sumr_const.
+apply/eq_big => [y | y /andP[Gy /eqP <-]]; last by rewrite !cfMorphE.
+rewrite mem_rcoset inE groupMr ?groupV // -mem_rcoset.
+by apply: andb_id2l => /(subsetP sGD) Dy; exact: sameP eqP (rcoset_kerP f _ _).
+Qed.
+
+Lemma cfIsom_iso rT G (R : {group rT}) (f : {morphism G >-> rT}) :
+ forall isoG : isom G R f, isometry (cfIsom isoG).
+Proof.
+move=> isoG phi psi; rewrite unlock cfMorph_iso //; set G1 := _ @* R.
+by rewrite -(isom_im (isom_sym isoG)) -/G1 in phi psi *; rewrite !cfRes_id.
+Qed.
+
+Lemma cfMod_iso H G : H <| G -> isometry (@cfMod _ G H).
+Proof. by case/andP=> _; apply: cfMorph_iso. Qed.
+
+Lemma cfQuo_iso H G :
+ H <| G -> {in [pred phi | H \subset cfker phi] &, isometry (@cfQuo _ G H)}.
+Proof.
+by move=> nsHG phi psi sHkphi sHkpsi; rewrite -(cfMod_iso nsHG) !cfQuoK.
+Qed.
+
+Lemma cfnorm_quo H G phi :
+ H <| G -> H \subset cfker phi -> '[phi / H] = '[phi]_G.
+Proof. by move=> nsHG sHker; apply: cfQuo_iso. Qed.
+
+Lemma cfSdprod_iso K H G (defG : K ><| H = G) : isometry (cfSdprod defG).
+Proof.
+move=> phi psi; have [/andP[_ nKG] _ _ _ _] := sdprod_context defG.
+by rewrite [cfSdprod _]locked_withE cfMorph_iso ?cfIsom_iso.
+Qed.
+
+End MorphIsometry.
+
+Section Induced.
+
+Variable gT : finGroupType.
+
+Section Def.
+
+Variables B A : {set gT}.
+Local Notation G := <<B>>.
+Local Notation H := <<A>>.
+
+(* The defalut value for the ~~ (H \subset G) case matches the one for cfRes *)
+(* so that Frobenius reciprocity holds even in this degenerate case. *)
+Definition ffun_cfInd (phi : 'CF(A)) :=
+ [ffun x => if H \subset G then #|A|%:R^-1 * (\sum_(y in G) phi (x ^ y))
+ else #|G|%:R * '[phi, 1] *+ (x == 1%g)].
+
+Fact cfInd_subproof phi : is_class_fun G (ffun_cfInd phi).
+Proof.
+apply: intro_class_fun => [x y Gx Gy | x H'x]; last first.
+ case: subsetP => [sHG | _]; last by rewrite (negPf (group1_contra H'x)).
+ rewrite big1 ?mulr0 // => y Gy; rewrite cfun0gen ?(contra _ H'x) //= => /sHG.
+ by rewrite memJ_norm ?(subsetP (normG _)).
+rewrite conjg_eq1 (reindex_inj (mulgI y^-1)%g); congr (if _ then _ * _ else _).
+by apply: eq_big => [z | z Gz]; rewrite ?groupMl ?groupV // -conjgM mulKVg.
+Qed.
+Definition cfInd phi := Cfun 1 (cfInd_subproof phi).
+
+Lemma cfInd_is_linear : linear cfInd.
+Proof.
+move=> c phi psi; apply/cfunP=> x; rewrite !cfunElock; case: ifP => _.
+ rewrite mulrCA -mulrDr [c * _]mulr_sumr -big_split /=.
+ by congr (_ * _); apply: eq_bigr => y _; rewrite !cfunE.
+rewrite mulrnAr -mulrnDl !(mulrCA c) -!mulrDr [c * _]mulr_sumr -big_split /=.
+by congr (_ * (_ * _) *+ _); apply: eq_bigr => y; rewrite !cfunE mulrA mulrDl.
+Qed.
+Canonical cfInd_additive := Additive cfInd_is_linear.
+Canonical cfInd_linear := Linear cfInd_is_linear.
+
+End Def.
+
+Local Notation "''Ind[' B , A ]" := (@cfInd B A) : ring_scope.
+Local Notation "''Ind[' B ]" := 'Ind[B, _] : ring_scope.
+
+Lemma cfIndE (G H : {group gT}) phi x :
+ H \subset G -> 'Ind[G, H] phi x = #|H|%:R^-1 * (\sum_(y in G) phi (x ^ y)).
+Proof. by rewrite cfunElock !genGid => ->. Qed.
+
+Variables G K H : {group gT}.
+Implicit Types (phi : 'CF(H)) (psi : 'CF(G)).
+
+Lemma cfIndEout phi :
+ ~~ (H \subset G) -> 'Ind[G] phi = (#|G|%:R * '[phi, 1]) *: '1_1%G.
+Proof.
+move/negPf=> not_sHG; apply/cfunP=> x; rewrite cfunE cfuniE ?normal1 // inE.
+by rewrite mulr_natr cfunElock !genGid not_sHG.
+Qed.
+
+Lemma cfIndEsdprod (phi : 'CF(K)) x :
+ K ><| H = G -> 'Ind[G] phi x = \sum_(w in H) phi (x ^ w)%g.
+Proof.
+move=> defG; have [/andP[sKG _] _ mulKH nKH _] := sdprod_context defG.
+rewrite cfIndE //; apply: canLR (mulKf (neq0CG _)) _; rewrite -mulKH mulr_sumr.
+rewrite (set_partition_big _ (rcosets_partition_mul H K)) ?big_imset /=.
+ apply: eq_bigr => y Hy; rewrite rcosetE norm_rlcoset ?(subsetP nKH) //.
+ rewrite -lcosetE mulr_natl big_imset /=; last exact: in2W (mulgI _).
+ by rewrite -sumr_const; apply: eq_bigr => z Kz; rewrite conjgM cfunJ.
+have [{nKH}nKH /isomP[injf _]] := sdprod_isom defG.
+apply: can_in_inj (fun Ky => invm injf (coset K (repr Ky))) _ => y Hy.
+by rewrite rcosetE -val_coset ?(subsetP nKH) // coset_reprK invmE.
+Qed.
+
+Lemma cfInd_on A phi :
+ H \subset G -> phi \in 'CF(H, A) -> 'Ind[G] phi \in 'CF(G, class_support A G).
+Proof.
+move=> sHG Af; apply/cfun_onP=> g AG'g; rewrite cfIndE ?big1 ?mulr0 // => h Gh.
+apply: (cfun_on0 Af); apply: contra AG'g => Agh.
+by rewrite -[g](conjgK h) memJ_class_support // groupV.
+Qed.
+
+Lemma cfInd_id phi : 'Ind[H] phi = phi.
+Proof.
+apply/cfun_inP=> x Hx; rewrite cfIndE // (eq_bigr _ (cfunJ phi x)) sumr_const.
+by rewrite -[phi x *+ _]mulr_natl mulKf ?neq0CG.
+Qed.
+
+Lemma cfInd_normal phi : H <| G -> 'Ind[G] phi \in 'CF(G, H).
+Proof.
+case/andP=> sHG nHG; apply: (cfun_onS (class_support_sub_norm (subxx _) nHG)).
+by rewrite cfInd_on ?cfun_onG.
+Qed.
+
+Lemma cfInd1 phi : H \subset G -> 'Ind[G] phi 1%g = #|G : H|%:R * phi 1%g.
+Proof.
+move=> sHG; rewrite cfIndE // natf_indexg // -mulrA mulrCA; congr (_ * _).
+by rewrite mulr_natl -sumr_const; apply: eq_bigr => x; rewrite conj1g.
+Qed.
+
+Lemma cfInd_cfun1 : H <| G -> 'Ind[G, H] 1 = #|G : H|%:R *: '1_H.
+Proof.
+move=> nsHG; have [sHG nHG] := andP nsHG; rewrite natf_indexg // mulrC.
+apply/cfunP=> x; rewrite cfIndE ?cfunE ?cfuniE // -mulrA; congr (_ * _).
+rewrite mulr_natl -sumr_const; apply: eq_bigr => y Gy.
+by rewrite cfun1E -{1}(normsP nHG y Gy) memJ_conjg.
+Qed.
+
+Lemma cfnorm_Ind_cfun1 : H <| G -> '['Ind[G, H] 1] = #|G : H|%:R.
+Proof.
+move=> nsHG; rewrite cfInd_cfun1 // cfnormZ normr_nat cfdot_cfuni // setIid.
+by rewrite expr2 {2}natf_indexg ?normal_sub // !mulrA divfK ?mulfK ?neq0CG.
+Qed.
+
+Lemma cfIndInd phi :
+ K \subset G -> H \subset K -> 'Ind[G] ('Ind[K] phi) = 'Ind[G] phi.
+Proof.
+move=> sKG sHK; apply/cfun_inP=> x Gx; rewrite !cfIndE ?(subset_trans sHK) //.
+apply: canLR (mulKf (neq0CG K)) _; rewrite mulr_sumr mulr_natl.
+transitivity (\sum_(y in G) \sum_(z in K) #|H|%:R^-1 * phi ((x ^ y) ^ z)).
+ by apply: eq_bigr => y Gy; rewrite cfIndE // -mulr_sumr.
+symmetry; rewrite exchange_big /= -sumr_const; apply: eq_bigr => z Kz.
+rewrite (reindex_inj (mulIg z)).
+by apply: eq_big => [y | y _]; rewrite ?conjgM // groupMr // (subsetP sKG).
+Qed.
+
+(* This is Isaacs, Lemma (5.2). *)
+Lemma Frobenius_reciprocity phi psi : '[phi, 'Res[H] psi] = '['Ind[G] phi, psi].
+Proof.
+have [sHG | not_sHG] := boolP (H \subset G); last first.
+ rewrite cfResEout // cfIndEout // cfdotZr cfdotZl mulrAC; congr (_ * _).
+ rewrite (cfdotEl _ (cfuni_on _ _)) mulVKf ?neq0CG // big_set1.
+ by rewrite cfuniE ?normal1 ?set11 ?mul1r.
+transitivity (#|H|%:R^-1 * \sum_(x in G) phi x * (psi x)^*).
+ rewrite (big_setID H) /= (setIidPr sHG) addrC big1 ?add0r; last first.
+ by move=> x /setDP[_ /cfun0->]; rewrite mul0r.
+ by congr (_ * _); apply: eq_bigr => x Hx; rewrite cfResE.
+set h' := _^-1; apply: canRL (mulKf (neq0CG G)) _.
+transitivity (h' * \sum_(y in G) \sum_(x in G) phi (x ^ y) * (psi (x ^ y))^*).
+ rewrite mulrCA mulr_natl -sumr_const; congr (_ * _); apply: eq_bigr => y Gy.
+ by rewrite (reindex_acts 'J _ Gy) ?astabsJ ?normG.
+rewrite exchange_big mulr_sumr; apply: eq_bigr => x _; rewrite cfIndE //=.
+by rewrite -mulrA mulr_suml; congr (_ * _); apply: eq_bigr => y /(cfunJ psi)->.
+Qed.
+Definition cfdot_Res_r := Frobenius_reciprocity.
+
+Lemma cfdot_Res_l psi phi : '['Res[H] psi, phi] = '[psi, 'Ind[G] phi].
+Proof. by rewrite cfdotC cfdot_Res_r -cfdotC. Qed.
+
+Lemma cfIndM phi psi: H \subset G ->
+ 'Ind[G] (phi * ('Res[H] psi)) = 'Ind[G] phi * psi.
+Proof.
+move=> HsG; apply/cfun_inP=> x Gx; rewrite !cfIndE // !cfunE !cfIndE // -mulrA.
+congr (_ * _); rewrite mulr_suml; apply: eq_bigr=> i iG; rewrite !cfunE.
+case:(boolP (x^i \in H))=> xJi; last by rewrite cfun0gen ?mul0r ?genGid.
+by rewrite !cfResE //; congr (_*_); rewrite cfunJgen ?genGid.
+Qed.
+
+End Induced.
+
+Arguments Scope cfInd [_ group_scope group_scope cfun_scope].
+Notation "''Ind[' G , H ]" := (@cfInd _ G H) (only parsing) : ring_scope.
+Notation "''Ind[' G ]" := 'Ind[G, _] : ring_scope.
+Notation "''Ind'" := 'Ind[_] (only parsing) : ring_scope.
+
+Section MorphInduced.
+
+Variables (aT rT : finGroupType) (D G H : {group aT}) (R S : {group rT}).
+
+Lemma cfIndMorph (f : {morphism D >-> rT}) (phi : 'CF(f @* H)) :
+ 'ker f \subset H -> H \subset G -> G \subset D ->
+ 'Ind[G] (cfMorph phi) = cfMorph ('Ind[f @* G] phi).
+Proof.
+move=> sKH sHG sGD; have [sHD inD] := (subset_trans sHG sGD, subsetP sGD).
+apply/cfun_inP=> /= x Gx; have [Dx sKG] := (inD x Gx, subset_trans sKH sHG).
+rewrite cfMorphE ?cfIndE ?morphimS // (partition_big_imset f) -morphimEsub //=.
+rewrite card_morphim (setIidPr sHD) natf_indexg // invfM invrK -mulrA.
+congr (_ * _); rewrite mulr_sumr; apply: eq_bigr => _ /morphimP[y Dy Gy ->].
+rewrite -(card_rcoset _ y) mulr_natl -sumr_const.
+apply: eq_big => [z | z /andP[Gz /eqP <-]].
+ have [Gz | G'z] := boolP (z \in G).
+ by rewrite (sameP eqP (rcoset_kerP _ _ _)) ?inD.
+ by case: rcosetP G'z => // [[t Kt ->]]; rewrite groupM // (subsetP sKG).
+have [Dz Dxz] := (inD z Gz, inD (x ^ z) (groupJ Gx Gz)); rewrite -morphJ //.
+have [Hxz | notHxz] := boolP (x ^ z \in H); first by rewrite cfMorphE.
+by rewrite !cfun0 // -sub1set -morphim_set1 // morphimSGK ?sub1set.
+Qed.
+
+Variables (g : {morphism G >-> rT}) (h : {morphism H >-> rT}).
+Hypotheses (isoG : isom G R g) (isoH : isom H S h) (eq_hg : {in H, h =1 g}).
+Hypothesis sHG : H \subset G.
+
+Lemma cfResIsom phi : 'Res[S] (cfIsom isoG phi) = cfIsom isoH ('Res[H] phi).
+Proof.
+have [[injg defR] [injh defS]] := (isomP isoG, isomP isoH).
+rewrite !morphimEdom in defS defR; apply/cfun_inP=> s.
+rewrite -{1}defS => /imsetP[x Hx ->] {s}; have Gx := subsetP sHG x Hx.
+rewrite {1}eq_hg ?(cfResE, cfIsomE) // -defS -?eq_hg ?mem_imset // -defR.
+by rewrite (eq_in_imset eq_hg) imsetS.
+Qed.
+
+Lemma cfIndIsom phi : 'Ind[R] (cfIsom isoH phi) = cfIsom isoG ('Ind[G] phi).
+Proof.
+have [[injg defR] [_ defS]] := (isomP isoG, isomP isoH).
+rewrite morphimEdom (eq_in_imset eq_hg) -morphimEsub // in defS.
+apply/cfun_inP=> s; rewrite -{1}defR => /morphimP[x _ Gx ->]{s}.
+rewrite cfIsomE ?cfIndE // -defR -{1}defS ?morphimS ?card_injm // morphimEdom.
+congr (_ * _); rewrite big_imset //=; last exact/injmP.
+apply: eq_bigr => y Gy; rewrite -morphJ //.
+have [Hxy | H'xy] := boolP (x ^ y \in H); first by rewrite -eq_hg ?cfIsomE.
+by rewrite !cfun0 -?defS // -sub1set -morphim_set1 ?injmSK ?sub1set // groupJ.
+Qed.
+
+End MorphInduced.
+
+Section FieldAutomorphism.
+
+Variables (u : {rmorphism algC -> algC}) (gT rT : finGroupType).
+Variables (G K H : {group gT}) (f : {morphism G >-> rT}) (R : {group rT}).
+Implicit Types (phi : 'CF(G)) (S : seq 'CF(G)).
+Local Notation "phi ^u" := (cfAut u phi) (at level 3, format "phi ^u").
+
+Lemma cfAutZ_nat n phi : (n%:R *: phi)^u = n%:R *: phi^u.
+Proof. exact: raddfZnat. Qed.
+
+Lemma cfAutZ_Cnat z phi : z \in Cnat -> (z *: phi)^u = z *: phi^u.
+Proof. exact: raddfZ_Cnat. Qed.
+
+Lemma cfAutZ_Cint z phi : z \in Cint -> (z *: phi)^u = z *: phi^u.
+Proof. exact: raddfZ_Cint. Qed.
+
+Lemma cfAut_inj : injective (@cfAut gT G u).
+Proof.
+move=> phi psi /cfunP eqfg; apply/cfunP=> x.
+by have := eqfg x; rewrite !cfunE => /fmorph_inj.
+Qed.
+
+Lemma cfAut_eq1 phi : (cfAut u phi == 1) = (phi == 1).
+Proof. by rewrite rmorph_eq1 //; apply: cfAut_inj. Qed.
+
+Lemma support_cfAut phi : support phi^u =i support phi.
+Proof. by move=> x; rewrite !inE cfunE fmorph_eq0. Qed.
+
+Lemma map_cfAut_free S : cfAut_closed u S -> free S -> free (map (cfAut u) S).
+Proof.
+set Su := map _ S => sSuS freeS; have uniqS := free_uniq freeS.
+have uniqSu: uniq Su by rewrite (map_inj_uniq cfAut_inj).
+have{sSuS} sSuS: {subset Su <= S} by move=> _ /mapP[phi Sphi ->]; apply: sSuS.
+have [|eqSuS _] := leq_size_perm uniqSu sSuS; first by rewrite size_map.
+by rewrite (perm_free (uniq_perm_eq uniqSu uniqS eqSuS)).
+Qed.
+
+Lemma cfAut_on A phi : (phi^u \in 'CF(G, A)) = (phi \in 'CF(G, A)).
+Proof. by rewrite !cfun_onE (eq_subset (support_cfAut phi)). Qed.
+
+Lemma cfker_aut phi : cfker phi^u = cfker phi.
+Proof.
+apply/setP=> x; rewrite !inE; apply: andb_id2l => Gx.
+by apply/forallP/forallP=> Kx y;
+ have:= Kx y; rewrite !cfunE (inj_eq (fmorph_inj u)).
+Qed.
+
+Lemma cfAut_cfuni A : ('1_A)^u = '1_A :> 'CF(G).
+Proof. by apply/cfunP=> x; rewrite !cfunElock rmorph_nat. Qed.
+
+Lemma cforder_aut phi : #[phi^u]%CF = #[phi]%CF.
+Proof. exact: cforder_inj_rmorph cfAut_inj. Qed.
+
+Lemma cfAutRes phi : ('Res[H] phi)^u = 'Res phi^u.
+Proof. by apply/cfunP=> x; rewrite !cfunElock rmorphMn. Qed.
+
+Lemma cfAutMorph (psi : 'CF(f @* H)) : (cfMorph psi)^u = cfMorph psi^u.
+Proof. by apply/cfun_inP=> x Hx; rewrite !cfunElock Hx. Qed.
+
+Lemma cfAutIsom (isoGR : isom G R f) phi :
+ (cfIsom isoGR phi)^u = cfIsom isoGR phi^u.
+Proof.
+apply/cfun_inP=> y; have [_ {1}<-] := isomP isoGR => /morphimP[x _ Gx ->{y}].
+by rewrite !(cfunE, cfIsomE).
+Qed.
+
+Lemma cfAutQuo phi : (phi / H)^u = (phi^u / H)%CF.
+Proof. by apply/cfunP=> Hx; rewrite !cfunElock cfker_aut rmorphMn. Qed.
+
+Lemma cfAutMod (psi : 'CF(G / H)) : (psi %% H)^u = (psi^u %% H)%CF.
+Proof. by apply/cfunP=> x; rewrite !cfunElock rmorphMn. Qed.
+
+Lemma cfAutInd (psi : 'CF(H)) : ('Ind[G] psi)^u = 'Ind psi^u.
+Proof.
+have [sHG | not_sHG] := boolP (H \subset G).
+ apply/cfunP=> x; rewrite !(cfunE, cfIndE) // rmorphM fmorphV rmorph_nat.
+ by congr (_ * _); rewrite rmorph_sum; apply: eq_bigr => y; rewrite !cfunE.
+rewrite !cfIndEout // linearZ /= cfAut_cfuni rmorphM rmorph_nat.
+rewrite -cfdot_cfAut ?rmorph1 // => _ /imageP[x Hx ->].
+by rewrite cfun1E Hx !rmorph1.
+Qed.
+
+Hypothesis KxH : K \x H = G.
+
+Lemma cfAutDprodl (phi : 'CF(K)) : (cfDprodl KxH phi)^u = cfDprodl KxH phi^u.
+Proof.
+apply/cfun_inP=> _ /(mem_dprod KxH)[x [y [Kx Hy -> _]]].
+by rewrite !(cfunE, cfDprodEl).
+Qed.
+
+Lemma cfAutDprodr (psi : 'CF(H)) : (cfDprodr KxH psi)^u = cfDprodr KxH psi^u.
+Proof.
+apply/cfun_inP=> _ /(mem_dprod KxH)[x [y [Kx Hy -> _]]].
+by rewrite !(cfunE, cfDprodEr).
+Qed.
+
+Lemma cfAutDprod (phi : 'CF(K)) (psi : 'CF(H)) :
+ (cfDprod KxH phi psi)^u = cfDprod KxH phi^u psi^u.
+Proof. by rewrite rmorphM /= cfAutDprodl cfAutDprodr. Qed.
+
+End FieldAutomorphism.
+
+Implicit Arguments cfAut_inj [gT G x1 x2].
+
+Definition conj_cfRes := cfAutRes conjC.
+Definition cfker_conjC := cfker_aut conjC.
+Definition conj_cfQuo := cfAutQuo conjC.
+Definition conj_cfMod := cfAutMod conjC.
+Definition conj_cfInd := cfAutInd conjC.
+Definition cfconjC_eq1 := cfAut_eq1 conjC.
+
diff --git a/mathcomp/character/inertia.v b/mathcomp/character/inertia.v
new file mode 100644
index 0000000..fc33e1a
--- /dev/null
+++ b/mathcomp/character/inertia.v
@@ -0,0 +1,1607 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq path choice div.
+Require Import fintype tuple finfun bigop prime ssralg ssrnum finset fingroup.
+Require Import morphism perm automorphism quotient action zmodp cyclic center.
+Require Import gproduct commutator gseries nilpotent pgroup sylow maximal.
+Require Import frobenius.
+Require Import matrix mxalgebra mxrepresentation vector algC classfun character.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope GRing.Theory Num.Theory.
+Local Open Scope ring_scope.
+
+(******************************************************************************)
+(* This file contains the definitions and properties of inertia groups: *)
+(* (phi ^ y)%CF == the y-conjugate of phi : 'CF(G), i.e., the class *)
+(* function mapping x ^ y to phi x provided y normalises G. *)
+(* We take (phi ^ y)%CF = phi when y \notin 'N(G). *)
+(* (phi ^: G)%CF == the sequence of all distinct conjugates of phi : 'CF(H) *)
+(* by all y in G. *)
+(* 'I[phi] == the inertia group of phi : CF(H), i.e., the set of y *)
+(* such that (phi ^ y)%CF = phi AND H :^ y = y. *)
+(* 'I_G[phi] == the inertia group of phi in G, i.e., G :&: 'I[phi]. *)
+(* conjg_Iirr i y == the index j : Iirr G such that ('chi_i ^ y)%CF = 'chi_j. *)
+(* cfclass_Iirr G i == the image of G under conjg_Iirr i, i.e., the set of j *)
+(* such that 'chi_j \in ('chi_i ^: G)%CF. *)
+(* mul_Iirr i j == the index k such that 'chi_j * 'chi_i = 'chi[G]_k, *)
+(* or 0 if 'chi_j * 'chi_i is reducible. *)
+(* mul_mod_Iirr i j := mul_Iirr i (mod_Iirr j), for j : Iirr (G / H). *)
+(******************************************************************************)
+
+Reserved Notation "''I[' phi ]"
+ (at level 8, format "''I[' phi ]").
+Reserved Notation "''I_' G [ phi ]"
+ (at level 8, G at level 2, format "''I_' G [ phi ]").
+
+Section ConjDef.
+
+Variables (gT : finGroupType) (B : {set gT}) (y : gT) (phi : 'CF(B)).
+Local Notation G := <<B>>.
+
+Fact cfConjg_subproof :
+ is_class_fun G [ffun x => phi (if y \in 'N(G) then x ^ y^-1 else x)].
+Proof.
+apply: intro_class_fun => [x z _ Gz | x notGx].
+ have [nGy | _] := ifP; last by rewrite cfunJgen.
+ by rewrite -conjgM conjgC conjgM cfunJgen // memJ_norm ?groupV.
+by rewrite cfun0gen //; case: ifP => // nGy; rewrite memJ_norm ?groupV.
+Qed.
+Definition cfConjg := Cfun 1 cfConjg_subproof.
+
+End ConjDef.
+
+Prenex Implicits cfConjg.
+Notation "f ^ y" := (cfConjg y f) : cfun_scope.
+
+Section Conj.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Implicit Type phi : 'CF(G).
+
+Lemma cfConjgE phi y x : y \in 'N(G) -> (phi ^ y)%CF x = phi (x ^ y^-1)%g.
+Proof. by rewrite cfunElock genGid => ->. Qed.
+
+Lemma cfConjgEJ phi y x : y \in 'N(G) -> (phi ^ y)%CF (x ^ y) = phi x.
+Proof. by move/cfConjgE->; rewrite conjgK. Qed.
+
+Lemma cfConjgEout phi y : y \notin 'N(G) -> (phi ^ y = phi)%CF.
+Proof.
+by move/negbTE=> notNy; apply/cfunP=> x; rewrite !cfunElock genGid notNy.
+Qed.
+
+Lemma cfConjgEin phi y (nGy : y \in 'N(G)) :
+ (phi ^ y)%CF = cfIsom (norm_conj_isom nGy) phi.
+Proof.
+apply/cfun_inP=> x Gx.
+by rewrite cfConjgE // -{2}[x](conjgKV y) cfIsomE ?memJ_norm ?groupV.
+Qed.
+
+Lemma cfConjgMnorm phi :
+ {in 'N(G) &, forall y z, phi ^ (y * z) = (phi ^ y) ^ z}%CF.
+Proof.
+move=> y z nGy nGz.
+by apply/cfunP=> x; rewrite !cfConjgE ?groupM // invMg conjgM.
+Qed.
+
+Lemma cfConjg_id phi y : y \in G -> (phi ^ y)%CF = phi.
+Proof.
+move=> Gy; apply/cfunP=> x; have nGy := subsetP (normG G) y Gy.
+by rewrite -(cfunJ _ _ Gy) cfConjgEJ.
+Qed.
+
+(* Isaacs' 6.1.b *)
+Lemma cfConjgM L phi :
+ G <| L -> {in L &, forall y z, phi ^ (y * z) = (phi ^ y) ^ z}%CF.
+Proof. by case/andP=> _ /subsetP nGL; exact: sub_in2 (cfConjgMnorm phi). Qed.
+
+Lemma cfConjgJ1 phi : (phi ^ 1)%CF = phi.
+Proof. by apply/cfunP=> x; rewrite cfConjgE ?group1 // invg1 conjg1. Qed.
+
+Lemma cfConjgK y : cancel (cfConjg y) (cfConjg y^-1 : 'CF(G) -> 'CF(G)).
+Proof.
+move=> phi; apply/cfunP=> x; rewrite !cfunElock groupV /=.
+by case: ifP => -> //; rewrite conjgKV.
+Qed.
+
+Lemma cfConjgKV y : cancel (cfConjg y^-1) (cfConjg y : 'CF(G) -> 'CF(G)).
+Proof. by move=> phi /=; rewrite -{1}[y]invgK cfConjgK. Qed.
+
+Lemma cfConjg1 phi y : (phi ^ y)%CF 1%g = phi 1%g.
+Proof. by rewrite cfunElock conj1g if_same. Qed.
+
+Fact cfConjg_is_linear y : linear (cfConjg y : 'CF(G) -> 'CF(G)).
+Proof. by move=> a phi psi; apply/cfunP=> x; rewrite !cfunElock. Qed.
+Canonical cfConjg_additive y := Additive (cfConjg_is_linear y).
+Canonical cfConjg_linear y := AddLinear (cfConjg_is_linear y).
+
+Lemma cfConjg_cfuniJ A y : y \in 'N(G) -> ('1_A ^ y)%CF = '1_(A :^ y) :> 'CF(G).
+Proof.
+move=> nGy; apply/cfunP=> x; rewrite !cfunElock genGid nGy -sub_conjgV.
+by rewrite -class_lcoset -class_rcoset norm_rlcoset ?memJ_norm ?groupV.
+Qed.
+
+Lemma cfConjg_cfuni A y : y \in 'N(A) -> ('1_A ^ y)%CF = '1_A :> 'CF(G).
+Proof.
+by have [/cfConjg_cfuniJ-> /normP-> | /cfConjgEout] := boolP (y \in 'N(G)).
+Qed.
+
+Lemma cfConjg_cfun1 y : (1 ^ y)%CF = 1 :> 'CF(G).
+Proof.
+by rewrite -cfuniG; have [/cfConjg_cfuni|/cfConjgEout] := boolP (y \in 'N(G)).
+Qed.
+
+Fact cfConjg_is_multiplicative y : multiplicative (cfConjg y : _ -> 'CF(G)).
+Proof.
+split=> [phi psi|]; last exact: cfConjg_cfun1.
+by apply/cfunP=> x; rewrite !cfunElock.
+Qed.
+Canonical cfConjg_rmorphism y := AddRMorphism (cfConjg_is_multiplicative y).
+Canonical cfConjg_lrmorphism y := [lrmorphism of cfConjg y].
+
+Lemma cfConjg_eq1 phi y : ((phi ^ y)%CF == 1) = (phi == 1).
+Proof. by apply: rmorph_eq1; apply: can_inj (cfConjgK y). Qed.
+
+Lemma cfAutConjg phi u y : cfAut u (phi ^ y) = (cfAut u phi ^ y)%CF.
+Proof. by apply/cfunP=> x; rewrite !cfunElock. Qed.
+
+Lemma conj_cfConjg phi y : (phi ^ y)^*%CF = (phi^* ^ y)%CF.
+Proof. exact: cfAutConjg. Qed.
+
+Lemma cfker_conjg phi y : y \in 'N(G) -> cfker (phi ^ y) = cfker phi :^ y.
+Proof.
+move=> nGy; rewrite cfConjgEin // cfker_isom.
+by rewrite morphim_conj (setIidPr (cfker_sub _)).
+Qed.
+
+Lemma cfDetConjg phi y : cfDet (phi ^ y) = (cfDet phi ^ y)%CF.
+Proof.
+have [nGy | not_nGy] := boolP (y \in 'N(G)); last by rewrite !cfConjgEout.
+by rewrite !cfConjgEin cfDetIsom.
+Qed.
+
+End Conj.
+
+Section Inertia.
+
+Variable gT : finGroupType.
+
+Definition inertia (B : {set gT}) (phi : 'CF(B)) :=
+ [set y in 'N(B) | (phi ^ y)%CF == phi].
+
+Local Notation "''I[' phi ]" := (inertia phi) : group_scope.
+Local Notation "''I_' G [ phi ]" := (G%g :&: 'I[phi]) : group_scope.
+
+Fact group_set_inertia (H : {group gT}) phi : group_set 'I[phi : 'CF(H)].
+Proof.
+apply/group_setP; split; first by rewrite inE group1 /= cfConjgJ1.
+move=> y z /setIdP[nHy /eqP n_phi_y] /setIdP[nHz n_phi_z].
+by rewrite inE groupM //= cfConjgMnorm ?n_phi_y.
+Qed.
+Canonical inertia_group H phi := Group (@group_set_inertia H phi).
+
+Local Notation "''I[' phi ]" := (inertia_group phi) : Group_scope.
+Local Notation "''I_' G [ phi ]" := (G :&: 'I[phi])%G : Group_scope.
+
+Variables G H : {group gT}.
+Implicit Type phi : 'CF(H).
+
+Lemma inertiaJ phi y : y \in 'I[phi] -> (phi ^ y)%CF = phi.
+Proof. by case/setIdP=> _ /eqP->. Qed.
+
+Lemma inertia_valJ phi x y : y \in 'I[phi] -> phi (x ^ y)%g = phi x.
+Proof. by case/setIdP=> nHy /eqP {1}<-; rewrite cfConjgEJ. Qed.
+
+(* To disambiguate basic inclucion lemma names we capitalize Inertia for *)
+(* lemmas concerning the localized inertia group 'I_G[phi]. *)
+Lemma Inertia_sub phi : 'I_G[phi] \subset G.
+Proof. exact: subsetIl. Qed.
+
+Lemma norm_inertia phi : 'I[phi] \subset 'N(H).
+Proof. by rewrite ['I[_]]setIdE subsetIl. Qed.
+
+Lemma sub_inertia phi : H \subset 'I[phi].
+Proof.
+by apply/subsetP=> y Hy; rewrite inE cfConjg_id ?(subsetP (normG H)) /=.
+Qed.
+
+Lemma normal_inertia phi : H <| 'I[phi].
+Proof. by rewrite /normal sub_inertia norm_inertia. Qed.
+
+Lemma sub_Inertia phi : H \subset G -> H \subset 'I_G[phi].
+Proof. by rewrite subsetI sub_inertia andbT. Qed.
+
+Lemma norm_Inertia phi : 'I_G[phi] \subset 'N(H).
+Proof. by rewrite setIC subIset ?norm_inertia. Qed.
+
+Lemma normal_Inertia phi : H \subset G -> H <| 'I_G[phi].
+Proof. by rewrite /normal norm_Inertia andbT; apply: sub_Inertia. Qed.
+
+Lemma cfConjg_eqE phi :
+ H <| G ->
+ {in G &, forall y z, (phi ^ y == phi ^ z)%CF = (z \in 'I_G[phi] :* y)}.
+Proof.
+case/andP=> _ nHG y z Gy; rewrite -{1 2}[z](mulgKV y) groupMr // mem_rcoset.
+move: {z}(z * _)%g => z Gz; rewrite 2!inE Gz cfConjgMnorm ?(subsetP nHG) //=.
+by rewrite eq_sym (can_eq (cfConjgK y)).
+Qed.
+
+Lemma cent_sub_inertia phi : 'C(H) \subset 'I[phi].
+Proof.
+apply/subsetP=> y cHy; have nHy := subsetP (cent_sub H) y cHy.
+rewrite inE nHy; apply/eqP/cfun_inP=> x Hx; rewrite cfConjgE //.
+by rewrite /conjg invgK mulgA (centP cHy) ?mulgK.
+Qed.
+
+Lemma cent_sub_Inertia phi : 'C_G(H) \subset 'I_G[phi].
+Proof. exact: setIS (cent_sub_inertia phi). Qed.
+
+Lemma center_sub_Inertia phi : H \subset G -> 'Z(G) \subset 'I_G[phi].
+Proof.
+by move/centS=> sHG; rewrite setIS // (subset_trans sHG) // cent_sub_inertia.
+Qed.
+
+Lemma conjg_inertia phi y : y \in 'N(H) -> 'I[phi] :^ y = 'I[phi ^ y].
+Proof.
+move=> nHy; apply/setP=> z; rewrite !['I[_]]setIdE conjIg conjGid // !in_setI.
+apply/andb_id2l=> nHz; rewrite mem_conjg !inE.
+by rewrite !cfConjgMnorm ?in_group ?(can2_eq (cfConjgKV y) (cfConjgK y)) ?invgK.
+Qed.
+
+Lemma inertia0 : 'I[0 : 'CF(H)] = 'N(H).
+Proof. by apply/setP=> x; rewrite !inE linear0 eqxx andbT. Qed.
+
+Lemma inertia_add phi psi : 'I[phi] :&: 'I[psi] \subset 'I[phi + psi].
+Proof.
+rewrite !['I[_]]setIdE -setIIr setIS //.
+by apply/subsetP=> x; rewrite !inE linearD /= => /andP[/eqP-> /eqP->].
+Qed.
+
+Lemma inertia_sum I r (P : pred I) (Phi : I -> 'CF(H)) :
+ 'N(H) :&: \bigcap_(i <- r | P i) 'I[Phi i]
+ \subset 'I[\sum_(i <- r | P i) Phi i].
+Proof.
+elim/big_rec2: _ => [|i K psi Pi sK_Ipsi]; first by rewrite setIT inertia0.
+by rewrite setICA; apply: subset_trans (setIS _ sK_Ipsi) (inertia_add _ _).
+Qed.
+
+Lemma inertia_scale a phi : 'I[phi] \subset 'I[a *: phi].
+Proof.
+apply/subsetP=> x /setIdP[nHx /eqP Iphi_x].
+by rewrite inE nHx linearZ /= Iphi_x.
+Qed.
+
+Lemma inertia_scale_nz a phi : a != 0 -> 'I[a *: phi] = 'I[phi].
+Proof.
+move=> nz_a; apply/eqP.
+by rewrite eqEsubset -{2}(scalerK nz_a phi) !inertia_scale.
+Qed.
+
+Lemma inertia_opp phi : 'I[- phi] = 'I[phi].
+Proof. by rewrite -scaleN1r inertia_scale_nz // oppr_eq0 oner_eq0. Qed.
+
+Lemma inertia1 : 'I[1 : 'CF(H)] = 'N(H).
+Proof. by apply/setP=> x; rewrite inE rmorph1 eqxx andbT. Qed.
+
+Lemma Inertia1 : H <| G -> 'I_G[1 : 'CF(H)] = G.
+Proof. by rewrite inertia1 => /normal_norm/setIidPl. Qed.
+
+Lemma inertia_mul phi psi : 'I[phi] :&: 'I[psi] \subset 'I[phi * psi].
+Proof.
+rewrite !['I[_]]setIdE -setIIr setIS //.
+by apply/subsetP=> x; rewrite !inE rmorphM /= => /andP[/eqP-> /eqP->].
+Qed.
+
+Lemma inertia_prod I r (P : pred I) (Phi : I -> 'CF(H)) :
+ 'N(H) :&: \bigcap_(i <- r | P i) 'I[Phi i]
+ \subset 'I[\prod_(i <- r | P i) Phi i].
+Proof.
+elim/big_rec2: _ => [|i K psi Pi sK_psi]; first by rewrite inertia1 setIT.
+by rewrite setICA; apply: subset_trans (setIS _ sK_psi) (inertia_mul _ _).
+Qed.
+
+Lemma inertia_injective (chi : 'CF(H)) :
+ {in H &, injective chi} -> 'I[chi] = 'C(H).
+Proof.
+move=> inj_chi; apply/eqP; rewrite eqEsubset cent_sub_inertia andbT.
+apply/subsetP=> y Ichi_y; have /setIdP[nHy _] := Ichi_y.
+apply/centP=> x Hx; apply/esym/commgP/conjg_fixP.
+by apply/inj_chi; rewrite ?memJ_norm ?(inertia_valJ _ Ichi_y).
+Qed.
+
+Lemma inertia_irr_prime p i :
+ #|H| = p -> prime p -> i != 0 -> 'I['chi[H]_i] = 'C(H).
+Proof. by move=> <- pr_H /(irr_prime_injP pr_H); apply: inertia_injective. Qed.
+
+Lemma inertia_irr0 : 'I['chi[H]_0] = 'N(H).
+Proof. by rewrite irr0 inertia1. Qed.
+
+(* Isaacs' 6.1.c *)
+Lemma cfConjg_iso y : isometry (cfConjg y : 'CF(H) -> 'CF(H)).
+Proof.
+move=> phi psi; congr (_ * _).
+have [nHy | not_nHy] := boolP (y \in 'N(H)); last by rewrite !cfConjgEout.
+rewrite (reindex_astabs 'J y) ?astabsJ //=.
+by apply: eq_bigr=> x _; rewrite !cfConjgEJ.
+Qed.
+
+(* Isaacs' 6.1.d *)
+Lemma cfdot_Res_conjg psi phi y :
+ y \in G -> '['Res[H, G] psi, phi ^ y] = '['Res[H] psi, phi].
+Proof.
+move=> Gy; rewrite -(cfConjg_iso y _ phi); congr '[_, _]; apply/cfunP=> x.
+rewrite !cfunElock !genGid; case nHy: (y \in 'N(H)) => //.
+by rewrite !(fun_if psi) cfunJ ?memJ_norm ?groupV.
+Qed.
+
+(* Isaac's 6.1.e *)
+Lemma cfConjg_char (chi : 'CF(H)) y :
+ chi \is a character -> (chi ^ y)%CF \is a character.
+Proof.
+have [nHy Nchi | /cfConjgEout-> //] := boolP (y \in 'N(H)).
+by rewrite cfConjgEin cfIsom_char.
+Qed.
+
+Lemma cfConjg_lin_char (chi : 'CF(H)) y :
+ chi \is a linear_char -> (chi ^ y)%CF \is a linear_char.
+Proof. by case/andP=> Nchi chi1; rewrite qualifE cfConjg1 cfConjg_char. Qed.
+
+Lemma cfConjg_irr y chi : chi \in irr H -> (chi ^ y)%CF \in irr H.
+Proof. by rewrite !irrEchar cfConjg_iso => /andP[/cfConjg_char->]. Qed.
+
+Definition conjg_Iirr i y := cfIirr ('chi[H]_i ^ y)%CF.
+
+Lemma conjg_IirrE i y : 'chi_(conjg_Iirr i y) = ('chi_i ^ y)%CF.
+Proof. by rewrite cfIirrE ?cfConjg_irr ?mem_irr. Qed.
+
+Lemma conjg_IirrK y : cancel (conjg_Iirr^~ y) (conjg_Iirr^~ y^-1%g).
+Proof. by move=> i; apply/irr_inj; rewrite !conjg_IirrE cfConjgK. Qed.
+
+Lemma conjg_IirrKV y : cancel (conjg_Iirr^~ y^-1%g) (conjg_Iirr^~ y).
+Proof. by rewrite -{2}[y]invgK; apply: conjg_IirrK. Qed.
+
+Lemma conjg_Iirr_inj y : injective (conjg_Iirr^~ y).
+Proof. exact: can_inj (conjg_IirrK y). Qed.
+
+Lemma conjg_Iirr_eq0 i y : (conjg_Iirr i y == 0) = (i == 0).
+Proof. by rewrite -!irr_eq1 conjg_IirrE cfConjg_eq1. Qed.
+
+Lemma conjg_Iirr0 x : conjg_Iirr 0 x = 0.
+Proof. by apply/eqP; rewrite conjg_Iirr_eq0. Qed.
+
+Lemma cfdot_irr_conjg i y :
+ H <| G -> y \in G -> '['chi_i, 'chi_i ^ y]_H = (y \in 'I_G['chi_i])%:R.
+Proof.
+move=> nsHG Gy; rewrite -conjg_IirrE cfdot_irr -(inj_eq irr_inj) conjg_IirrE.
+by rewrite -{1}['chi_i]cfConjgJ1 cfConjg_eqE ?mulg1.
+Qed.
+
+Definition cfclass (A : {set gT}) (phi : 'CF(A)) (B : {set gT}) :=
+ [seq (phi ^ repr Tx)%CF | Tx in rcosets 'I_B[phi] B].
+
+Local Notation "phi ^: G" := (cfclass phi G) : cfun_scope.
+
+Lemma size_cfclass i : size ('chi[H]_i ^: G)%CF = #|G : 'I_G['chi_i]|.
+Proof. by rewrite size_map -cardE. Qed.
+
+Lemma cfclassP (A : {group gT}) phi psi :
+ reflect (exists2 y, y \in A & psi = phi ^ y)%CF (psi \in phi ^: A)%CF.
+Proof.
+apply: (iffP imageP) => [[_ /rcosetsP[y Ay ->] ->] | [y Ay ->]].
+ by case: repr_rcosetP => z /setIdP[Az _]; exists (z * y)%g; rewrite ?groupM.
+without loss nHy: y Ay / y \in 'N(H).
+ have [nHy | /cfConjgEout->] := boolP (y \in 'N(H)); first exact.
+ by move/(_ 1%g); rewrite !group1 !cfConjgJ1; exact.
+exists ('I_A[phi] :* y); first by rewrite -rcosetE mem_imset.
+case: repr_rcosetP => z /setIP[_ /setIdP[nHz /eqP Tz]].
+by rewrite cfConjgMnorm ?Tz.
+Qed.
+
+Lemma cfclassInorm phi : (phi ^: 'N_G(H) =i phi ^: G)%CF.
+Proof.
+move=> xi; apply/cfclassP/cfclassP=> [[x /setIP[Gx _] ->] | [x Gx ->]].
+ by exists x.
+have [Nx | /cfConjgEout-> //] := boolP (x \in 'N(H)).
+ by exists x; first exact/setIP.
+by exists 1%g; rewrite ?group1 ?cfConjgJ1.
+Qed.
+
+Lemma cfclass_refl phi : phi \in (phi ^: G)%CF.
+Proof. by apply/cfclassP; exists 1%g => //; rewrite cfConjgJ1. Qed.
+
+Lemma cfclass_transl phi psi :
+ (psi \in phi ^: G)%CF -> (phi ^: G =i psi ^: G)%CF.
+Proof.
+rewrite -cfclassInorm; case/cfclassP=> x Gx -> xi; rewrite -!cfclassInorm.
+have nHN: {subset 'N_G(H) <= 'N(H)} by apply/subsetP; exact: subsetIr.
+apply/cfclassP/cfclassP=> [[y Gy ->] | [y Gy ->]].
+ by exists (x^-1 * y)%g; rewrite -?cfConjgMnorm ?groupM ?groupV ?nHN // mulKVg.
+by exists (x * y)%g; rewrite -?cfConjgMnorm ?groupM ?nHN.
+Qed.
+
+Lemma cfclass_sym phi psi : (psi \in phi ^: G)%CF = (phi \in psi ^: G)%CF.
+Proof. by apply/idP/idP=> /cfclass_transl <-; exact: cfclass_refl. Qed.
+
+Lemma cfclass_uniq phi : H <| G -> uniq (phi ^: G)%CF.
+Proof.
+move=> nsHG; rewrite map_inj_in_uniq ?enum_uniq // => Ty Tz; rewrite !mem_enum.
+move=> {Ty}/rcosetsP[y Gy ->] {Tz}/rcosetsP[z Gz ->] /eqP.
+case: repr_rcosetP => u Iphi_u; case: repr_rcosetP => v Iphi_v.
+have [[Gu _] [Gv _]] := (setIdP Iphi_u, setIdP Iphi_v).
+rewrite cfConjg_eqE ?groupM // => /rcoset_transl.
+by rewrite !rcosetM (rcoset_id Iphi_v) (rcoset_id Iphi_u).
+Qed.
+
+Lemma cfclass_invariant phi : G \subset 'I[phi] -> (phi ^: G)%CF = phi.
+Proof.
+move/setIidPl=> IGphi; rewrite /cfclass IGphi // rcosets_id.
+by rewrite /(image _ _) enum_set1 /= repr_group cfConjgJ1.
+Qed.
+
+Lemma cfclass1 : H <| G -> (1 ^: G)%CF = [:: 1 : 'CF(H)].
+Proof. by move/normal_norm=> nHG; rewrite cfclass_invariant ?inertia1. Qed.
+
+Definition cfclass_Iirr (A : {set gT}) i := conjg_Iirr i @: A.
+
+Lemma cfclass_IirrE i j :
+ (j \in cfclass_Iirr G i) = ('chi_j \in 'chi_i ^: G)%CF.
+Proof.
+apply/imsetP/cfclassP=> [[y Gy ->] | [y]]; exists y; rewrite ?conjg_IirrE //.
+by apply: irr_inj; rewrite conjg_IirrE.
+Qed.
+
+Lemma eq_cfclass_IirrE i j :
+ (cfclass_Iirr G j == cfclass_Iirr G i) = (j \in cfclass_Iirr G i).
+Proof.
+apply/eqP/idP=> [<- | iGj]; first by rewrite cfclass_IirrE cfclass_refl.
+by apply/setP=> k; rewrite !cfclass_IirrE in iGj *; apply/esym/cfclass_transl.
+Qed.
+
+Lemma im_cfclass_Iirr i :
+ H <| G -> perm_eq [seq 'chi_j | j in cfclass_Iirr G i] ('chi_i ^: G)%CF.
+Proof.
+move=> nsHG; have UchiG := cfclass_uniq 'chi_i nsHG.
+apply: uniq_perm_eq; rewrite ?(map_inj_uniq irr_inj) ?enum_uniq // => phi.
+apply/imageP/idP=> [[j iGj ->] | /cfclassP[y]]; first by rewrite -cfclass_IirrE.
+by exists (conjg_Iirr i y); rewrite ?mem_imset ?conjg_IirrE.
+Qed.
+
+Lemma card_cfclass_Iirr i : H <| G -> #|cfclass_Iirr G i| = #|G : 'I_G['chi_i]|.
+Proof.
+move=> nsHG; rewrite -size_cfclass -(perm_eq_size (im_cfclass_Iirr i nsHG)).
+by rewrite size_map -cardE.
+Qed.
+
+Lemma reindex_cfclass R idx (op : Monoid.com_law idx) (F : 'CF(H) -> R) i :
+ H <| G ->
+ \big[op/idx]_(chi <- ('chi_i ^: G)%CF) F chi
+ = \big[op/idx]_(j | 'chi_j \in ('chi_i ^: G)%CF) F 'chi_j.
+Proof.
+move/im_cfclass_Iirr/(eq_big_perm _) <-; rewrite big_map big_filter /=.
+by apply: eq_bigl => j; rewrite cfclass_IirrE.
+Qed.
+
+Lemma cfResInd j:
+ H <| G ->
+ 'Res[H] ('Ind[G] 'chi_j) = #|H|%:R^-1 *: (\sum_(y in G) 'chi_j ^ y)%CF.
+Proof.
+case/andP=> [sHG /subsetP nHG].
+rewrite (reindex_inj invg_inj); apply/cfun_inP=> x Hx.
+rewrite cfResE // cfIndE // ?cfunE ?sum_cfunE; congr (_ * _).
+by apply: eq_big => [y | y Gy]; rewrite ?cfConjgE ?groupV ?invgK ?nHG.
+Qed.
+
+(* This is Isaacs, Theorem (6.2) *)
+Lemma Clifford_Res_sum_cfclass i j :
+ H <| G -> j \in irr_constt ('Res[H, G] 'chi_i) ->
+ 'Res[H] 'chi_i =
+ '['Res[H] 'chi_i, 'chi_j] *: (\sum_(chi <- ('chi_j ^: G)%CF) chi).
+Proof.
+move=> nsHG chiHj; have [sHG /subsetP nHG] := andP nsHG.
+rewrite reindex_cfclass //= big_mkcond.
+rewrite {1}['Res _]cfun_sum_cfdot linear_sum /=; apply: eq_bigr => k _.
+have [[y Gy ->] | ] := altP (cfclassP _ _ _); first by rewrite cfdot_Res_conjg.
+apply: contraNeq; rewrite scaler0 scaler_eq0 orbC => /norP[_ chiHk].
+have{chiHk chiHj}: '['Res[H] ('Ind[G] 'chi_j), 'chi_k] != 0.
+ rewrite !inE !cfdot_Res_l in chiHj chiHk *.
+ apply: contraNneq chiHk; rewrite cfdot_sum_irr => /psumr_eq0P/(_ i isT)/eqP.
+ rewrite -cfdotC cfdotC mulf_eq0 conjC_eq0 (negbTE chiHj) /= => -> // i1.
+ by rewrite -cfdotC Cnat_ge0 // rpredM ?Cnat_cfdot_char ?cfInd_char ?irr_char.
+rewrite cfResInd // cfdotZl mulf_eq0 cfdot_suml => /norP[_].
+apply: contraR => chiGk'j; rewrite big1 // => x Gx; apply: contraNeq chiGk'j.
+rewrite -conjg_IirrE cfdot_irr pnatr_eq0; case: (_ =P k) => // <- _.
+by rewrite conjg_IirrE; apply/cfclassP; exists x.
+Qed.
+
+Lemma cfRes_Ind_invariant psi :
+ H <| G -> G \subset 'I[psi] -> 'Res ('Ind[G, H] psi) = #|G : H|%:R *: psi.
+Proof.
+case/andP=> sHG _ /subsetP IGpsi; apply/cfun_inP=> x Hx.
+rewrite cfResE ?cfIndE ?natf_indexg // cfunE -mulrA mulrCA; congr (_ * _).
+by rewrite mulr_natl -sumr_const; apply: eq_bigr => y /IGpsi/inertia_valJ->.
+Qed.
+
+(* This is Isaacs, Corollary (6.7). *)
+Corollary constt0_Res_cfker i :
+ H <| G -> 0 \in irr_constt ('Res[H] 'chi[G]_i) -> H \subset cfker 'chi[G]_i.
+Proof.
+move=> nsHG /(Clifford_Res_sum_cfclass nsHG); have [sHG nHG] := andP nsHG.
+rewrite irr0 cfdot_Res_l cfclass1 // big_seq1 cfInd_cfun1 //.
+rewrite cfdotZr conjC_nat => def_chiH.
+apply/subsetP=> x Hx; rewrite cfkerEirr inE -!(cfResE _ sHG) //.
+by rewrite def_chiH !cfunE cfun11 cfun1E Hx.
+Qed.
+
+(* This is Isaacs, Lemma (6.8). *)
+Lemma dvdn_constt_Res1_irr1 i j :
+ H <| G -> j \in irr_constt ('Res[H, G] 'chi_i) ->
+ exists n, 'chi_i 1%g = n%:R * 'chi_j 1%g.
+Proof.
+move=> nsHG chiHj; have [sHG nHG] := andP nsHG; rewrite -(cfResE _ sHG) //.
+rewrite {1}(Clifford_Res_sum_cfclass nsHG chiHj) cfunE sum_cfunE.
+have /CnatP[n ->]: '['Res[H] 'chi_i, 'chi_j] \in Cnat.
+ by rewrite Cnat_cfdot_char ?cfRes_char ?irr_char.
+exists (n * size ('chi_j ^: G)%CF)%N; rewrite natrM -mulrA; congr (_ * _).
+rewrite mulr_natl -[size _]card_ord big_tnth -sumr_const; apply: eq_bigr => k _.
+by have /cfclassP[y Gy ->]:= mem_tnth k (in_tuple _); rewrite cfConjg1.
+Qed.
+
+Lemma cfclass_Ind phi psi :
+ H <| G -> psi \in (phi ^: G)%CF -> 'Ind[G] phi = 'Ind[G] psi.
+Proof.
+move=> nsHG /cfclassP[y Gy ->]; have [sHG /subsetP nHG] := andP nsHG.
+apply/cfun_inP=> x Hx; rewrite !cfIndE //; congr (_ * _).
+rewrite (reindex_acts 'R _ (groupVr Gy)) ?astabsR //=.
+by apply: eq_bigr => z Gz; rewrite conjgM cfConjgE ?nHG.
+Qed.
+
+End Inertia.
+
+Arguments Scope inertia [_ group_scope cfun_scope].
+Arguments Scope cfclass [_ group_scope cfun_scope group_scope].
+Implicit Arguments conjg_Iirr_inj [gT H x1 x2].
+
+Notation "''I[' phi ] " := (inertia phi) : group_scope.
+Notation "''I[' phi ] " := (inertia_group phi) : Group_scope.
+Notation "''I_' G [ phi ] " := (G%g :&: 'I[phi]) : group_scope.
+Notation "''I_' G [ phi ] " := (G :&: 'I[phi])%G : Group_scope.
+Notation "phi ^: G" := (cfclass phi G) : cfun_scope.
+
+Section ConjRestrict.
+
+Variables (gT : finGroupType) (G H K : {group gT}).
+
+Lemma cfConjgRes_norm phi y :
+ y \in 'N(K) -> y \in 'N(H) -> ('Res[K, H] phi ^ y)%CF = 'Res (phi ^ y)%CF.
+Proof.
+move=> nKy nHy; have [sKH | not_sKH] := boolP (K \subset H); last first.
+ by rewrite !cfResEout // linearZ rmorph1 cfConjg1.
+by apply/cfun_inP=> x Kx; rewrite !(cfConjgE, cfResE) ?memJ_norm ?groupV.
+Qed.
+
+Lemma cfConjgRes phi y :
+ H <| G -> K <| G -> y \in G -> ('Res[K, H] phi ^ y)%CF = 'Res (phi ^ y)%CF.
+Proof.
+move=> /andP[_ nHG] /andP[_ nKG] Gy.
+by rewrite cfConjgRes_norm ?(subsetP nHG) ?(subsetP nKG).
+Qed.
+
+Lemma sub_inertia_Res phi :
+ G \subset 'N(K) -> 'I_G[phi] \subset 'I_G['Res[K, H] phi].
+Proof.
+move=> nKG; apply/subsetP=> y /setIP[Gy /setIdP[nHy /eqP Iphi_y]].
+by rewrite 2!inE Gy cfConjgRes_norm ?(subsetP nKG) ?Iphi_y /=.
+Qed.
+
+Lemma cfConjgInd_norm phi y :
+ y \in 'N(K) -> y \in 'N(H) -> ('Ind[H, K] phi ^ y)%CF = 'Ind (phi ^ y)%CF.
+Proof.
+move=> nKy nHy; have [sKH | not_sKH] := boolP (K \subset H).
+ by rewrite !cfConjgEin (cfIndIsom (norm_conj_isom nHy)).
+rewrite !cfIndEout // linearZ -(cfConjg_iso y) rmorph1 /=; congr (_ *: _).
+by rewrite cfConjg_cfuni ?norm1 ?inE.
+Qed.
+
+Lemma cfConjgInd phi y :
+ H <| G -> K <| G -> y \in G -> ('Ind[H, K] phi ^ y)%CF = 'Ind (phi ^ y)%CF.
+Proof.
+move=> /andP[_ nHG] /andP[_ nKG] Gy.
+by rewrite cfConjgInd_norm ?(subsetP nHG) ?(subsetP nKG).
+Qed.
+
+Lemma sub_inertia_Ind phi :
+ G \subset 'N(H) -> 'I_G[phi] \subset 'I_G['Ind[H, K] phi].
+Proof.
+move=> nHG; apply/subsetP=> y /setIP[Gy /setIdP[nKy /eqP Iphi_y]].
+by rewrite 2!inE Gy cfConjgInd_norm ?(subsetP nHG) ?Iphi_y /=.
+Qed.
+
+End ConjRestrict.
+
+Section MoreInertia.
+
+Variables (gT : finGroupType) (G H : {group gT}) (i : Iirr H).
+Let T := 'I_G['chi_i].
+
+Lemma inertia_id : 'I_T['chi_i] = T. Proof. by rewrite -setIA setIid. Qed.
+
+Lemma cfclass_inertia : ('chi[H]_i ^: T)%CF = [:: 'chi_i].
+Proof.
+rewrite /cfclass inertia_id rcosets_id /(image _ _) enum_set1 /=.
+by rewrite repr_group cfConjgJ1.
+Qed.
+
+End MoreInertia.
+
+Section ConjMorph.
+
+Variables (aT rT : finGroupType) (D G H : {group aT}) (f : {morphism D >-> rT}).
+
+Lemma cfConjgMorph (phi : 'CF(f @* H)) y :
+ y \in D -> y \in 'N(H) -> (cfMorph phi ^ y)%CF = cfMorph (phi ^ f y).
+Proof.
+move=> Dy nHy; have [sHD | not_sHD] := boolP (H \subset D); last first.
+ by rewrite !cfMorphEout // linearZ rmorph1 cfConjg1.
+apply/cfun_inP=> x Gx; rewrite !(cfConjgE, cfMorphE) ?memJ_norm ?groupV //.
+ by rewrite morphJ ?morphV ?groupV // (subsetP sHD).
+by rewrite (subsetP (morphim_norm _ _)) ?mem_morphim.
+Qed.
+
+Lemma inertia_morph_pre (phi : 'CF(f @* H)) :
+ H <| G -> G \subset D -> 'I_G[cfMorph phi] = G :&: f @*^-1 'I_(f @* G)[phi].
+Proof.
+case/andP=> sHG nHG sGD; have sHD := subset_trans sHG sGD.
+apply/setP=> y; rewrite !in_setI; apply: andb_id2l => Gy.
+have [Dy nHy] := (subsetP sGD y Gy, subsetP nHG y Gy).
+rewrite Dy inE nHy 4!inE mem_morphim // -morphimJ ?(normP nHy) // subxx /=.
+rewrite cfConjgMorph //; apply/eqP/eqP=> [Iphi_y | -> //].
+by apply/cfun_inP=> _ /morphimP[x Dx Hx ->]; rewrite -!cfMorphE ?Iphi_y.
+Qed.
+
+Lemma inertia_morph_im (phi : 'CF(f @* H)) :
+ H <| G -> G \subset D -> f @* 'I_G[cfMorph phi] = 'I_(f @* G)[phi].
+Proof.
+move=> nsHG sGD; rewrite inertia_morph_pre // morphim_setIpre.
+by rewrite (setIidPr _) ?Inertia_sub.
+Qed.
+
+Variables (R S : {group rT}).
+Variables (g : {morphism G >-> rT}) (h : {morphism H >-> rT}).
+Hypotheses (isoG : isom G R g) (isoH : isom H S h).
+Hypotheses (eq_hg : {in H, h =1 g}) (sHG : H \subset G).
+
+(* This does not depend on the (isoG : isom G R g) assumption. *)
+Lemma cfConjgIsom phi y :
+ y \in G -> y \in 'N(H) -> (cfIsom isoH phi ^ g y)%CF = cfIsom isoH (phi ^ y).
+Proof.
+move=> Gy nHy; have [_ defS] := isomP isoH.
+rewrite morphimEdom (eq_in_imset eq_hg) -morphimEsub // in defS.
+apply/cfun_inP=> gx; rewrite -{1}defS => /morphimP[x Gx Hx ->] {gx}.
+rewrite cfConjgE; last by rewrite -defS inE -morphimJ ?(normP nHy).
+by rewrite -morphV -?morphJ -?eq_hg ?cfIsomE ?cfConjgE ?memJ_norm ?groupV.
+Qed.
+
+Lemma inertia_isom phi : 'I_R[cfIsom isoH phi] = g @* 'I_G[phi].
+Proof.
+have [[_ defS] [injg <-]] := (isomP isoH, isomP isoG).
+rewrite morphimEdom (eq_in_imset eq_hg) -morphimEsub // in defS.
+rewrite /inertia !setIdE morphimIdom setIA -{1}defS -injm_norm ?injmI //.
+apply/setP=> gy; rewrite !inE; apply: andb_id2l => /morphimP[y Gy nHy ->] {gy}.
+rewrite cfConjgIsom // -sub1set -morphim_set1 // injmSK ?sub1set //= inE.
+apply/eqP/eqP=> [Iphi_y | -> //].
+by apply/cfun_inP=> x Hx; rewrite -!(cfIsomE isoH) ?Iphi_y.
+Qed.
+
+End ConjMorph.
+
+Section ConjQuotient.
+
+Variables gT : finGroupType.
+Implicit Types G H K : {group gT}.
+
+Lemma cfConjgMod_norm H K (phi : 'CF(H / K)) y :
+ y \in 'N(K) -> y \in 'N(H) -> ((phi %% K) ^ y)%CF = (phi ^ coset K y %% K)%CF.
+Proof. exact: cfConjgMorph. Qed.
+
+Lemma cfConjgMod G H K (phi : 'CF(H / K)) y :
+ H <| G -> K <| G -> y \in G ->
+ ((phi %% K) ^ y)%CF = (phi ^ coset K y %% K)%CF.
+Proof.
+move=> /andP[_ nHG] /andP[_ nKG] Gy.
+by rewrite cfConjgMod_norm ?(subsetP nHG) ?(subsetP nKG).
+Qed.
+
+Lemma cfConjgQuo_norm H K (phi : 'CF(H)) y :
+ y \in 'N(K) -> y \in 'N(H) -> ((phi / K) ^ coset K y)%CF = (phi ^ y / K)%CF.
+Proof.
+move=> nKy nHy; have keryK: (K \subset cfker (phi ^ y)) = (K \subset cfker phi).
+ by rewrite cfker_conjg // -{1}(normP nKy) conjSg.
+have [kerK | not_kerK] := boolP (K \subset cfker phi); last first.
+ by rewrite !cfQuoEout ?linearZ ?rmorph1 ?cfConjg1 ?keryK.
+apply/cfun_inP=> _ /morphimP[x nKx Hx ->].
+have nHyb: coset K y \in 'N(H / K) by rewrite inE -morphimJ ?(normP nHy).
+rewrite !(cfConjgE, cfQuoEnorm) ?keryK // ?in_setI ?Hx //.
+rewrite -morphV -?morphJ ?groupV // cfQuoEnorm //.
+by rewrite inE memJ_norm ?Hx ?groupJ ?groupV.
+Qed.
+
+Lemma cfConjgQuo G H K (phi : 'CF(H)) y :
+ H <| G -> K <| G -> y \in G ->
+ ((phi / K) ^ coset K y)%CF = (phi ^ y / K)%CF.
+Proof.
+move=> /andP[_ nHG] /andP[_ nKG] Gy.
+by rewrite cfConjgQuo_norm ?(subsetP nHG) ?(subsetP nKG).
+Qed.
+
+Lemma inertia_mod_pre G H K (phi : 'CF(H / K)) :
+ H <| G -> K <| G -> 'I_G[phi %% K] = G :&: coset K @*^-1 'I_(G / K)[phi].
+Proof. by move=> nsHG /andP[_]; apply: inertia_morph_pre. Qed.
+
+Lemma inertia_mod_quo G H K (phi : 'CF(H / K)) :
+ H <| G -> K <| G -> ('I_G[phi %% K] / K)%g = 'I_(G / K)[phi].
+Proof. by move=> nsHG /andP[_]; apply: inertia_morph_im. Qed.
+
+Lemma inertia_quo G H K (phi : 'CF(H)) :
+ H <| G -> K <| G -> K \subset cfker phi ->
+ 'I_(G / K)[phi / K] = ('I_G[phi] / K)%g.
+Proof.
+move=> nsHG nsKG kerK; rewrite -inertia_mod_quo ?cfQuoK //.
+by rewrite (normalS _ (normal_sub nsHG)) // (subset_trans _ (cfker_sub phi)).
+Qed.
+
+End ConjQuotient.
+
+Section InertiaSdprod.
+
+Variables (gT : finGroupType) (K H G : {group gT}).
+
+Hypothesis defG : K ><| H = G.
+
+Lemma cfConjgSdprod phi y :
+ y \in 'N(K) -> y \in 'N(H) ->
+ (cfSdprod defG phi ^ y = cfSdprod defG (phi ^ y))%CF.
+Proof.
+move=> nKy nHy.
+have nGy: y \in 'N(G) by rewrite -sub1set -(sdprodW defG) normsM ?sub1set.
+rewrite -{2}[phi](cfSdprodK defG) cfConjgRes_norm // cfRes_sdprodK //.
+by rewrite cfker_conjg // -{1}(normP nKy) conjSg cfker_sdprod.
+Qed.
+
+Lemma inertia_sdprod (L : {group gT}) phi :
+ L \subset 'N(K) -> L \subset 'N(H) -> 'I_L[cfSdprod defG phi] = 'I_L[phi].
+Proof.
+move=> nKL nHL; have nGL: L \subset 'N(G) by rewrite -(sdprodW defG) normsM.
+apply/setP=> z; rewrite !in_setI ![z \in 'I[_]]inE; apply: andb_id2l => Lz.
+rewrite cfConjgSdprod ?(subsetP nKL) ?(subsetP nHL) ?(subsetP nGL) //=.
+by rewrite (can_eq (cfSdprodK defG)).
+Qed.
+
+End InertiaSdprod.
+
+Section InertiaDprod.
+
+Variables (gT : finGroupType) (G K H : {group gT}).
+Implicit Type L : {group gT}.
+Hypothesis KxH : K \x H = G.
+
+Lemma cfConjgDprodl phi y :
+ y \in 'N(K) -> y \in 'N(H) ->
+ (cfDprodl KxH phi ^ y = cfDprodl KxH (phi ^ y))%CF.
+Proof. by move=> nKy nHy; apply: cfConjgSdprod. Qed.
+
+Lemma cfConjgDprodr psi y :
+ y \in 'N(K) -> y \in 'N(H) ->
+ (cfDprodr KxH psi ^ y = cfDprodr KxH (psi ^ y))%CF.
+Proof. by move=> nKy nHy; apply: cfConjgSdprod. Qed.
+
+Lemma cfConjgDprod phi psi y :
+ y \in 'N(K) -> y \in 'N(H) ->
+ (cfDprod KxH phi psi ^ y = cfDprod KxH (phi ^ y) (psi ^ y))%CF.
+Proof. by move=> nKy nHy; rewrite rmorphM /= cfConjgDprodl ?cfConjgDprodr. Qed.
+
+Lemma inertia_dprodl L phi :
+ L \subset 'N(K) -> L \subset 'N(H) -> 'I_L[cfDprodl KxH phi] = 'I_L[phi].
+Proof. by move=> nKL nHL; apply: inertia_sdprod. Qed.
+
+Lemma inertia_dprodr L psi :
+ L \subset 'N(K) -> L \subset 'N(H) -> 'I_L[cfDprodr KxH psi] = 'I_L[psi].
+Proof. by move=> nKL nHL; apply: inertia_sdprod. Qed.
+
+Lemma inertia_dprod L (phi : 'CF(K)) (psi : 'CF(H)) :
+ L \subset 'N(K) -> L \subset 'N(H) -> phi 1%g != 0 -> psi 1%g != 0 ->
+ 'I_L[cfDprod KxH phi psi] = 'I_L[phi] :&: 'I_L[psi].
+Proof.
+move=> nKL nHL nz_phi nz_psi; apply/eqP; rewrite eqEsubset subsetI.
+rewrite -{1}(inertia_scale_nz psi nz_phi) -{1}(inertia_scale_nz phi nz_psi).
+rewrite -(cfDprod_Resl KxH) -(cfDprod_Resr KxH) !sub_inertia_Res //=.
+by rewrite -inertia_dprodl -?inertia_dprodr // -setIIr setIS ?inertia_mul.
+Qed.
+
+Lemma inertia_dprod_irr L i j :
+ L \subset 'N(K) -> L \subset 'N(H) ->
+ 'I_L[cfDprod KxH 'chi_i 'chi_j] = 'I_L['chi_i] :&: 'I_L['chi_j].
+Proof. by move=> nKL nHL; rewrite inertia_dprod ?irr1_neq0. Qed.
+
+End InertiaDprod.
+
+Section InertiaBigdprod.
+
+Variables (gT : finGroupType) (I : finType) (P : pred I).
+Variables (A : I -> {group gT}) (G : {group gT}).
+Implicit Type L : {group gT}.
+Hypothesis defG : \big[dprod/1%g]_(i | P i) A i = G.
+
+Section ConjBig.
+
+Variable y : gT.
+Hypothesis nAy: forall i, P i -> y \in 'N(A i).
+
+Lemma cfConjgBigdprodi i (phi : 'CF(A i)) :
+ (cfBigdprodi defG phi ^ y = cfBigdprodi defG (phi ^ y))%CF.
+Proof.
+rewrite cfConjgDprodl; try by case: ifP => [/nAy// | _]; rewrite norm1 inE.
+ congr (cfDprodl _ _); case: ifP => [Pi | _].
+ by rewrite cfConjgRes_norm ?nAy.
+ by apply/cfun_inP=> _ /set1P->; rewrite !(cfRes1, cfConjg1).
+rewrite -sub1set norms_gen ?norms_bigcup // sub1set.
+by apply/bigcapP=> j /andP[/nAy].
+Qed.
+
+Lemma cfConjgBigdprod phi :
+ (cfBigdprod defG phi ^ y = cfBigdprod defG (fun i => phi i ^ y))%CF.
+Proof.
+by rewrite rmorph_prod /=; apply: eq_bigr => i _; apply: cfConjgBigdprodi.
+Qed.
+
+End ConjBig.
+
+Section InertiaBig.
+
+Variable L : {group gT}.
+Hypothesis nAL : forall i, P i -> L \subset 'N(A i).
+
+Lemma inertia_bigdprodi i (phi : 'CF(A i)) :
+ P i -> 'I_L[cfBigdprodi defG phi] = 'I_L[phi].
+Proof.
+move=> Pi; rewrite inertia_dprodl ?Pi ?cfRes_id ?nAL //.
+by apply/norms_gen/norms_bigcup/bigcapsP=> j /andP[/nAL].
+Qed.
+
+Lemma inertia_bigdprod phi (Phi := cfBigdprod defG phi) :
+ Phi 1%g != 0 -> 'I_L[Phi] = L :&: \bigcap_(i | P i) 'I_L[phi i].
+Proof.
+move=> nz_Phi; apply/eqP; rewrite eqEsubset; apply/andP; split.
+ rewrite subsetI Inertia_sub; apply/bigcapsP=> i Pi.
+ have [] := cfBigdprodK nz_Phi Pi; move: (_ / _) => a nz_a <-.
+ by rewrite inertia_scale_nz ?sub_inertia_Res //= ?nAL.
+rewrite subsetI subsetIl; apply: subset_trans (inertia_prod _ _ _).
+apply: setISS.
+ by rewrite -(bigdprodWY defG) norms_gen ?norms_bigcup //; apply/bigcapsP.
+apply/bigcapsP=> i Pi; rewrite (bigcap_min i) //.
+by rewrite -inertia_bigdprodi ?subsetIr.
+Qed.
+
+Lemma inertia_bigdprod_irr Iphi (phi := fun i => 'chi_(Iphi i)) :
+ 'I_L[cfBigdprod defG phi] = L :&: \bigcap_(i | P i) 'I_L[phi i].
+Proof.
+rewrite inertia_bigdprod // -[cfBigdprod _ _]cfIirrE ?irr1_neq0 //.
+by apply: cfBigdprod_irr => i _; apply: mem_irr.
+Qed.
+
+End InertiaBig.
+
+End InertiaBigdprod.
+
+Section ConsttInertiaBijection.
+
+Variables (gT : finGroupType) (H G : {group gT}) (t : Iirr H).
+Hypothesis nsHG : H <| G.
+
+Local Notation theta := 'chi_t.
+Local Notation T := 'I_G[theta]%G.
+Local Notation "` 'T'" := 'I_(gval G)[theta]
+ (at level 0, format "` 'T'") : group_scope.
+
+Let calA := irr_constt ('Ind[T] theta).
+Let calB := irr_constt ('Ind[G] theta).
+Local Notation AtoB := (Ind_Iirr G).
+
+(* This is Isaacs, Theorem (6.11). *)
+Theorem constt_Inertia_bijection :
+ [/\ (*a*) {in calA, forall s, 'Ind[G] 'chi_s \in irr G},
+ (*b*) {in calA &, injective (Ind_Iirr G)},
+ Ind_Iirr G @: calA =i calB,
+ (*c*) {in calA, forall s (psi := 'chi_s) (chi := 'Ind[G] psi),
+ [predI irr_constt ('Res chi) & calA] =i pred1 s}
+ & (*d*) {in calA, forall s (psi := 'chi_s) (chi := 'Ind[G] psi),
+ '['Res psi, theta] = '['Res chi, theta]}].
+Proof.
+have [sHG sTG]: H \subset G /\ T \subset G by rewrite subsetIl normal_sub.
+have nsHT : H <| T := normal_Inertia theta sHG; have sHT := normal_sub nsHT.
+have AtoB_P s (psi := 'chi_s) (chi := 'Ind[G] psi): s \in calA ->
+ [/\ chi \in irr G, AtoB s \in calB & '['Res psi, theta] = '['Res chi, theta]].
+- rewrite !constt_Ind_Res => sHt; have [r sGr] := constt_cfInd_irr s sTG.
+ have rTs: s \in irr_constt ('Res[T] 'chi_r) by rewrite -constt_Ind_Res.
+ have NrT: 'Res[T] 'chi_r \is a character by rewrite cfRes_char ?irr_char.
+ have rHt: t \in irr_constt ('Res[H] 'chi_r).
+ by have:= constt_Res_trans NrT rTs sHt; rewrite cfResRes.
+ pose e := '['Res[H] 'chi_r, theta]; set f := '['Res[H] psi, theta].
+ have DrH: 'Res[H] 'chi_r = e *: \sum_(xi <- (theta ^: G)%CF) xi.
+ exact: Clifford_Res_sum_cfclass.
+ have DpsiH: 'Res[H] psi = f *: theta.
+ rewrite (Clifford_Res_sum_cfclass nsHT sHt).
+ by rewrite cfclass_invariant ?subsetIr ?big_seq1.
+ have ub_chi_r: 'chi_r 1%g <= chi 1%g ?= iff ('chi_r == chi).
+ have Nchi: chi \is a character by rewrite cfInd_char ?irr_char.
+ have [chi1 Nchi1->] := constt_charP _ Nchi sGr.
+ rewrite addrC cfunE -lerif_subLR subrr eq_sym -subr_eq0 addrK.
+ by split; rewrite ?char1_ge0 // eq_sym char1_eq0.
+ have lb_chi_r: chi 1%g <= 'chi_r 1%g ?= iff (f == e).
+ rewrite cfInd1 // -(cfRes1 H) DpsiH -(cfRes1 H 'chi_r) DrH !cfunE sum_cfunE.
+ rewrite (eq_big_seq (fun _ => theta 1%g)) => [|i]; last first.
+ by case/cfclassP=> y _ ->; rewrite cfConjg1.
+ rewrite reindex_cfclass //= sumr_const -(eq_card (cfclass_IirrE _ _)).
+ rewrite mulr_natl mulrnAr card_cfclass_Iirr //.
+ rewrite (mono_lerif (ler_pmuln2r (indexg_gt0 G T))).
+ rewrite (mono_lerif (ler_pmul2r (irr1_gt0 t))); apply: lerif_eq.
+ by rewrite /e -(cfResRes _ sHT) ?cfdot_Res_ge_constt.
+ have [_ /esym] := lerif_trans ub_chi_r lb_chi_r; rewrite eqxx.
+ by case/andP=> /eqP Dchi /eqP->;rewrite cfIirrE -/chi -?Dchi ?mem_irr.
+have part_c: {in calA, forall s (chi := 'Ind[G] 'chi_s),
+ [predI irr_constt ('Res[T] chi) & calA] =i pred1 s}.
+- move=> s As chi s1; have [irr_chi _ /eqP Dchi_theta] := AtoB_P s As.
+ have chiTs: s \in irr_constt ('Res[T] chi).
+ by rewrite irr_consttE cfdot_Res_l irrWnorm ?oner_eq0.
+ apply/andP/eqP=> [[/= chiTs1 As1] | -> //].
+ apply: contraTeq Dchi_theta => s's1; rewrite ltr_eqF // -/chi.
+ have [|phi Nphi DchiT] := constt_charP _ _ chiTs.
+ by rewrite cfRes_char ?cfInd_char ?irr_char.
+ have [|phi1 Nphi1 Dphi] := constt_charP s1 Nphi _.
+ rewrite irr_consttE -(canLR (addKr _) DchiT) addrC cfdotBl cfdot_irr.
+ by rewrite mulrb ifN_eqC ?subr0.
+ rewrite -(cfResRes chi sHT sTG) DchiT Dphi !rmorphD !cfdotDl /=.
+ rewrite -ltr_subl_addl subrr ltr_paddr ?ltr_def //;
+ rewrite Cnat_ge0 ?Cnat_cfdot_char ?cfRes_char ?irr_char //.
+ by rewrite andbT -irr_consttE -constt_Ind_Res.
+do [split=> //; try by move=> s /AtoB_P[]] => [s1 s2 As1 As2 | r].
+ have [[irr_s1G _ _] [irr_s2G _ _]] := (AtoB_P _ As1, AtoB_P _ As2).
+ move/(congr1 (tnth (irr G))); rewrite !cfIirrE // => eq_s12_G.
+ apply/eqP; rewrite -[_ == _]part_c // inE /= As1 -eq_s12_G.
+ by rewrite -As1 [_ && _]part_c // inE /=.
+apply/imsetP/idP=> [[s /AtoB_P[_ BsG _] -> //] | Br].
+have /exists_inP[s rTs As]: [exists s in irr_constt ('Res 'chi_r), s \in calA].
+ rewrite -negb_forall_in; apply: contra Br => /eqfun_inP => o_tT_rT.
+ rewrite -(cfIndInd _ sTG sHT) -cfdot_Res_r ['Res _]cfun_sum_constt.
+ by rewrite cfdot_sumr big1 // => i rTi; rewrite cfdotZr o_tT_rT ?mulr0.
+exists s => //; have [/irrP[r1 DsG] _ _] := AtoB_P s As.
+by apply/eqP; rewrite /AtoB -constt_Ind_Res DsG irrK constt_irr in rTs *.
+Qed.
+
+End ConsttInertiaBijection.
+
+Section ExtendInvariantIrr.
+
+Variable gT : finGroupType.
+Implicit Types G H K L M N : {group gT}.
+
+Section ConsttIndExtendible.
+
+Variables (G N : {group gT}) (t : Iirr N) (c : Iirr G).
+Let theta := 'chi_t.
+Let chi := 'chi_c.
+
+Definition mul_Iirr b := cfIirr ('chi_b * chi).
+Definition mul_mod_Iirr (b : Iirr (G / N)) := mul_Iirr (mod_Iirr b).
+
+Hypotheses (nsNG : N <| G) (cNt : 'Res[N] chi = theta).
+Let sNG : N \subset G. Proof. exact: normal_sub. Qed.
+Let nNG : G \subset 'N(N). Proof. exact: normal_norm. Qed.
+
+Lemma extendible_irr_invariant : G \subset 'I[theta].
+Proof.
+apply/subsetP=> y Gy; have nNy := subsetP nNG y Gy.
+rewrite inE nNy; apply/eqP/cfun_inP=> x Nx; rewrite cfConjgE // -cNt.
+by rewrite !cfResE ?memJ_norm ?cfunJ ?groupV.
+Qed.
+Let IGtheta := extendible_irr_invariant.
+
+(* This is Isaacs, Theorem (6.16) *)
+Theorem constt_Ind_mul_ext f (phi := 'chi_f) (psi := phi * theta) :
+ G \subset 'I[phi] -> psi \in irr N ->
+ let calS := irr_constt ('Ind phi) in
+ [/\ {in calS, forall b, 'chi_b * chi \in irr G},
+ {in calS &, injective mul_Iirr},
+ irr_constt ('Ind psi) =i [seq mul_Iirr b | b in calS]
+ & 'Ind psi = \sum_(b in calS) '['Ind phi, 'chi_b] *: 'chi_(mul_Iirr b)].
+Proof.
+move=> IGphi irr_psi calS.
+have IGpsi: G \subset 'I[psi].
+ by rewrite (subset_trans _ (inertia_mul _ _)) // subsetI IGphi.
+pose e b := '['Ind[G] phi, 'chi_b]; pose d b g := '['chi_b * chi, 'chi_g * chi].
+have Ne b: e b \in Cnat by rewrite Cnat_cfdot_char ?cfInd_char ?irr_char.
+have egt0 b: b \in calS -> e b > 0 by rewrite Cnat_gt0.
+have DphiG: 'Ind phi = \sum_(b in calS) e b *: 'chi_b := cfun_sum_constt _.
+have DpsiG: 'Ind psi = \sum_(b in calS) e b *: 'chi_b * chi.
+ by rewrite /psi -cNt cfIndM // DphiG mulr_suml.
+pose d_delta := [forall b in calS, forall g in calS, d b g == (b == g)%:R].
+have charMchi b: 'chi_b * chi \is a character by rewrite rpredM ?irr_char.
+have [_]: '['Ind[G] phi] <= '['Ind[G] psi] ?= iff d_delta.
+ pose sum_delta := \sum_(b in calS) e b * \sum_(g in calS) e g * (b == g)%:R.
+ pose sum_d := \sum_(b in calS) e b * \sum_(g in calS) e g * d b g.
+ have ->: '['Ind[G] phi] = sum_delta.
+ rewrite DphiG cfdot_suml; apply: eq_bigr => b _; rewrite cfdotZl cfdot_sumr.
+ by congr (_ * _); apply: eq_bigr => g; rewrite cfdotZr cfdot_irr conj_Cnat.
+ have ->: '['Ind[G] psi] = sum_d.
+ rewrite DpsiG cfdot_suml; apply: eq_bigr => b _.
+ rewrite -scalerAl cfdotZl cfdot_sumr; congr (_ * _).
+ by apply: eq_bigr => g _; rewrite -scalerAl cfdotZr conj_Cnat.
+ have eMmono := mono_lerif (ler_pmul2l (egt0 _ _)).
+ apply: lerif_sum => b /eMmono->; apply: lerif_sum => g /eMmono->.
+ split; last exact: eq_sym.
+ have /CnatP[n Dd]: d b g \in Cnat by rewrite Cnat_cfdot_char.
+ have [Db | _] := eqP; rewrite Dd leC_nat // -ltC_nat -Dd Db cfnorm_gt0.
+ by rewrite -char1_eq0 // cfunE mulf_neq0 ?irr1_neq0.
+rewrite -!cfdot_Res_l ?cfRes_Ind_invariant // !cfdotZl cfnorm_irr irrWnorm //.
+rewrite eqxx => /esym/forall_inP/(_ _ _)/eqfun_inP; rewrite /d /= => Dd.
+have irrMchi: {in calS, forall b, 'chi_b * chi \in irr G}.
+ by move=> b Sb; rewrite /= irrEchar charMchi Dd ?eqxx.
+have injMchi: {in calS &, injective mul_Iirr}.
+ move=> b g Sb Sg /(congr1 (fun s => '['chi_s, 'chi_(mul_Iirr g)]))/eqP.
+ by rewrite cfnorm_irr !cfIirrE ?irrMchi ?Dd // pnatr_eq1; case: (b =P g).
+have{DpsiG} ->: 'Ind psi = \sum_(b in calS) e b *: 'chi_(mul_Iirr b).
+ by rewrite DpsiG; apply: eq_bigr => b Sb; rewrite -scalerAl cfIirrE ?irrMchi.
+split=> // i; rewrite irr_consttE cfdot_suml;
+apply/idP/idP=> [|/imageP[b Sb ->]].
+ apply: contraR => N'i; rewrite big1 // => b Sb.
+ rewrite cfdotZl cfdot_irr mulrb ifN_eqC ?mulr0 //.
+ by apply: contraNneq N'i => ->; apply: image_f.
+rewrite gtr_eqF // (bigD1 b) //= cfdotZl cfnorm_irr mulr1 ltr_paddr ?egt0 //.
+apply: sumr_ge0 => g /andP[Sg _]; rewrite cfdotZl cfdot_irr.
+by rewrite mulr_ge0 ?ler0n ?Cnat_ge0.
+Qed.
+
+(* This is Isaacs, Corollary (6.17) (due to Gallagher). *)
+Corollary constt_Ind_ext :
+ [/\ forall b : Iirr (G / N), 'chi_(mod_Iirr b) * chi \in irr G,
+ injective mul_mod_Iirr,
+ irr_constt ('Ind theta) =i codom mul_mod_Iirr
+ & 'Ind theta = \sum_b 'chi_b 1%g *: 'chi_(mul_mod_Iirr b)].
+Proof.
+have IHchi0: G \subset 'I['chi[N]_0] by rewrite inertia_irr0.
+have [] := constt_Ind_mul_ext IHchi0; rewrite irr0 ?mul1r ?mem_irr //.
+set psiG := 'Ind 1 => irrMchi injMchi constt_theta {2}->.
+have dot_psiG b: '[psiG, 'chi_(mod_Iirr b)] = 'chi[G / N]_b 1%g.
+ rewrite mod_IirrE // -cfdot_Res_r cfRes_sub_ker ?cfker_mod //.
+ by rewrite cfdotZr cfnorm1 mulr1 conj_Cnat ?cfMod1 ?Cnat_irr1.
+have mem_psiG (b : Iirr (G / N)): mod_Iirr b \in irr_constt psiG.
+ by rewrite irr_consttE dot_psiG irr1_neq0.
+have constt_psiG b: (b \in irr_constt psiG) = (N \subset cfker 'chi_b).
+ apply/idP/idP=> [psiGb | /quo_IirrK <- //].
+ by rewrite constt0_Res_cfker // -constt_Ind_Res irr0.
+split=> [b | b g /injMchi/(can_inj (mod_IirrK nsNG))-> // | b0 | ].
+- exact: irrMchi.
+- rewrite constt_theta.
+ apply/imageP/imageP=> [][b psiGb ->]; last by exists (mod_Iirr b).
+ by exists (quo_Iirr N b) => //; rewrite /mul_mod_Iirr quo_IirrK -?constt_psiG.
+rewrite (reindex_onto _ _ (in1W (mod_IirrK nsNG))) /=.
+apply/esym/eq_big => b; first by rewrite constt_psiG quo_IirrKeq.
+by rewrite -dot_psiG /mul_mod_Iirr => /eqP->.
+Qed.
+
+End ConsttIndExtendible.
+
+(* This is Isaacs, Theorem (6.19). *)
+Theorem invariant_chief_irr_cases G K L s (theta := 'chi[K]_s) :
+ chief_factor G L K -> abelian (K / L) -> G \subset 'I[theta] ->
+ let t := #|K : L| in
+ [\/ 'Res[L] theta \in irr L,
+ exists2 e, exists p, 'Res[L] theta = e%:R *: 'chi_p & (e ^ 2)%N = t
+ | exists2 p, injective p & 'Res[L] theta = \sum_(i < t) 'chi_(p i)].
+Proof.
+case/andP=> /maxgroupP[/andP[ltLK nLG] maxL] nsKG abKbar IGtheta t.
+have [sKG nKG] := andP nsKG; have sLG := subset_trans (proper_sub ltLK) sKG.
+have nsLG: L <| G by apply/andP.
+have nsLK := normalS (proper_sub ltLK) sKG nsLG; have [sLK nLK] := andP nsLK.
+have [p0 sLp0] := constt_cfRes_irr L s; rewrite -/theta in sLp0.
+pose phi := 'chi_p0; pose T := 'I_G[phi].
+have sTG: T \subset G := subsetIl G _.
+have /eqP mulKT: (K * T)%g == G.
+ rewrite eqEcard mulG_subG sKG sTG -LagrangeMr -indexgI -(Lagrange sTG) /= -/T.
+ rewrite mulnC leq_mul // setIA (setIidPl sKG) -!size_cfclass // -/phi.
+ rewrite uniq_leq_size ?cfclass_uniq // => _ /cfclassP[x Gx ->].
+ have: conjg_Iirr p0 x \in irr_constt ('Res theta).
+ have /inertiaJ <-: x \in 'I[theta] := subsetP IGtheta x Gx.
+ by rewrite -(cfConjgRes _ nsKG) // irr_consttE conjg_IirrE // cfConjg_iso.
+ apply: contraR; rewrite -conjg_IirrE // => not_sLp0x.
+ rewrite (Clifford_Res_sum_cfclass nsLK sLp0) cfdotZl cfdot_suml.
+ rewrite big1_seq ?mulr0 // => _ /cfclassP[y Ky ->]; rewrite -conjg_IirrE //.
+ rewrite cfdot_irr mulrb ifN_eq ?(contraNneq _ not_sLp0x) // => <-.
+ by rewrite conjg_IirrE //; apply/cfclassP; exists y.
+have nsKT_G: K :&: T <| G.
+ rewrite /normal subIset ?sKG // -mulKT setIA (setIidPl sKG) mulG_subG.
+ rewrite normsIG // sub_der1_norm ?subsetIl //.
+ exact: subset_trans (der1_min nLK abKbar) (sub_Inertia _ sLK).
+have [e DthL]: exists e, 'Res theta = e%:R *: \sum_(xi <- (phi ^: K)%CF) xi.
+ rewrite (Clifford_Res_sum_cfclass nsLK sLp0) -/phi; set e := '[_, _].
+ by exists (truncC e); rewrite truncCK ?Cnat_cfdot_char ?cfRes_char ?irr_char.
+have [defKT | ltKT_K] := eqVneq (K :&: T) K; last first.
+ have defKT: K :&: T = L.
+ apply: maxL; last by rewrite subsetI sLK sub_Inertia.
+ by rewrite normal_norm // properEneq ltKT_K subsetIl.
+ have t_cast: size (phi ^: K)%CF = t.
+ by rewrite size_cfclass //= -{2}(setIidPl sKG) -setIA defKT.
+ pose phiKt := Tuple (introT eqP t_cast); pose p i := cfIirr (tnth phiKt i).
+ have pK i: 'chi_(p i) = (phi ^: K)%CF`_i.
+ rewrite cfIirrE; first by rewrite (tnth_nth 0).
+ by have /cfclassP[y _ ->] := mem_tnth i phiKt; rewrite cfConjg_irr ?mem_irr.
+ constructor 3; exists p => [i j /(congr1 (tnth (irr L)))/eqP| ].
+ by apply: contraTeq; rewrite !pK !nth_uniq ?t_cast ?cfclass_uniq.
+ have{DthL} DthL: 'Res theta = e%:R *: \sum_(i < t) (phi ^: K)%CF`_i.
+ by rewrite DthL (big_nth 0) big_mkord t_cast.
+ suffices /eqP e1: e == 1%N by rewrite DthL e1 scale1r; apply: eq_bigr.
+ have Dth1: theta 1%g = e%:R * t%:R * phi 1%g.
+ rewrite -[t]card_ord -mulrA -(cfRes1 L) DthL cfunE; congr (_ * _).
+ rewrite mulr_natl -sumr_const sum_cfunE -t_cast; apply: eq_bigr => i _.
+ by have /cfclassP[y _ ->] := mem_nth 0 (valP i); rewrite cfConjg1.
+ rewrite eqn_leq lt0n (contraNneq _ (irr1_neq0 s)); last first.
+ by rewrite Dth1 => ->; rewrite !mul0r.
+ rewrite -leC_nat -(ler_pmul2r (gt0CiG K L)) -/t -(ler_pmul2r (irr1_gt0 p0)).
+ rewrite mul1r -Dth1 -cfInd1 //.
+ by rewrite char1_ge_constt ?cfInd_char ?irr_char ?constt_Ind_Res.
+have IKphi: 'I_K[phi] = K by rewrite -{1}(setIidPl sKG) -setIA.
+have{DthL} DthL: 'Res[L] theta = e%:R *: phi.
+ by rewrite DthL -[rhs in (_ ^: rhs)%CF]IKphi cfclass_inertia big_seq1.
+pose mmLth := @mul_mod_Iirr K L s.
+have linKbar := char_abelianP _ abKbar.
+have LmodL i: ('chi_i %% L)%CF \is a linear_char := cfMod_lin_char (linKbar i).
+have mmLthE i: 'chi_(mmLth i) = ('chi_i %% L)%CF * theta.
+ by rewrite cfIirrE ?mod_IirrE // mul_lin_irr ?mem_irr.
+have mmLthL i: 'Res[L] 'chi_(mmLth i) = 'Res[L] theta.
+ rewrite mmLthE rmorphM /= cfRes_sub_ker ?cfker_mod ?lin_char1 //.
+ by rewrite scale1r mul1r.
+have [inj_Mphi | /injectivePn[i [j i'j eq_mm_ij]]] := boolP (injectiveb mmLth).
+ suffices /eqP e1: e == 1%N by constructor 1; rewrite DthL e1 scale1r mem_irr.
+ rewrite eqn_leq lt0n (contraNneq _ (irr1_neq0 s)); last first.
+ by rewrite -(cfRes1 L) DthL cfunE => ->; rewrite !mul0r.
+ rewrite -leq_sqr -leC_nat natrX -(ler_pmul2r (irr1_gt0 p0)) -mulrA mul1r.
+ have ->: e%:R * 'chi_p0 1%g = 'Res[L] theta 1%g by rewrite DthL cfunE.
+ rewrite cfRes1 -(ler_pmul2l (gt0CiG K L)) -cfInd1 // -/phi.
+ rewrite -card_quotient // -card_Iirr_abelian // mulr_natl.
+ rewrite ['Ind phi]cfun_sum_cfdot sum_cfunE (bigID (mem (codom mmLth))) /=.
+ rewrite ler_paddr ?sumr_ge0 // => [i _|].
+ by rewrite char1_ge0 ?rpredZ_Cnat ?Cnat_cfdot_char ?cfInd_char ?irr_char.
+ rewrite -big_uniq //= big_map big_filter -sumr_const ler_sum // => i _.
+ rewrite cfunE -[in rhs in _ <= rhs](cfRes1 L) -cfdot_Res_r mmLthL cfRes1.
+ by rewrite DthL cfdotZr rmorph_nat cfnorm_irr mulr1.
+constructor 2; exists e; first by exists p0.
+pose mu := (('chi_i / 'chi_j)%R %% L)%CF; pose U := cfker mu.
+have lin_mu: mu \is a linear_char by rewrite cfMod_lin_char ?rpred_div.
+have Uj := lin_char_unitr (linKbar j).
+have ltUK: U \proper K.
+ rewrite /proper cfker_sub /U; have /irrP[k Dmu] := lin_char_irr lin_mu.
+ rewrite Dmu subGcfker -irr_eq1 -Dmu cfMod_eq1 //.
+ by rewrite (can2_eq (divrK Uj) (mulrK Uj)) mul1r (inj_eq irr_inj).
+suffices: theta \in 'CF(K, L).
+ rewrite -cfnorm_Res_lerif // DthL cfnormZ !cfnorm_irr !mulr1 normr_nat.
+ by rewrite -natrX eqC_nat => /eqP.
+have <-: gcore U G = L.
+ apply: maxL; last by rewrite sub_gcore ?cfker_mod.
+ by rewrite gcore_norm (sub_proper_trans (gcore_sub _ _)).
+apply/cfun_onP=> x; apply: contraNeq => nz_th_x.
+apply/bigcapP=> y /(subsetP IGtheta)/setIdP[nKy /eqP th_y].
+apply: contraR nz_th_x; rewrite mem_conjg -{}th_y cfConjgE {nKy}//.
+move: {x y}(x ^ _) => x U'x; have [Kx | /cfun0-> //] := boolP (x \in K).
+have /eqP := congr1 (fun k => (('chi_j %% L)%CF^-1 * 'chi_k) x) eq_mm_ij.
+rewrite -rmorphV // !mmLthE !mulrA -!rmorphM mulVr //= rmorph1 !cfunE.
+rewrite (mulrC _^-1) -/mu -subr_eq0 -mulrBl cfun1E Kx mulf_eq0 => /orP[]//.
+rewrite mulrb subr_eq0 -(lin_char1 lin_mu) [_ == _](contraNF _ U'x) //.
+by rewrite /U cfkerEchar ?lin_charW // inE Kx.
+Qed.
+
+(* This is Isaacs, Corollary (6.19). *)
+Corollary cfRes_prime_irr_cases G N s p (chi := 'chi[G]_s) :
+ N <| G -> #|G : N| = p -> prime p ->
+ [\/ 'Res[N] chi \in irr N
+ | exists2 c, injective c & 'Res[N] chi = \sum_(i < p) 'chi_(c i)].
+Proof.
+move=> /andP[sNG nNG] iGN pr_p.
+have chiefGN: chief_factor G N G.
+ apply/andP; split=> //; apply/maxgroupP.
+ split=> [|M /andP[/andP[sMG ltMG] _] sNM].
+ by rewrite /proper sNG -indexg_gt1 iGN prime_gt1.
+ apply/esym/eqP; rewrite eqEsubset sNM -indexg_eq1 /= eq_sym.
+ rewrite -(eqn_pmul2l (indexg_gt0 G M)) muln1 Lagrange_index // iGN.
+ by apply/eqP/prime_nt_dvdP; rewrite ?indexg_eq1 // -iGN indexgS.
+have abGbar: abelian (G / N).
+ by rewrite cyclic_abelian ?prime_cyclic ?card_quotient ?iGN.
+have IGchi: G \subset 'I[chi] by apply: sub_inertia.
+have [] := invariant_chief_irr_cases chiefGN abGbar IGchi; first by left.
+ case=> e _ /(congr1 (fun m => odd (logn p m)))/eqP/idPn[].
+ by rewrite lognX mul2n odd_double iGN logn_prime // eqxx.
+by rewrite iGN; right.
+Qed.
+
+(* This is Isaacs, Corollary (6.20). *)
+Corollary prime_invariant_irr_extendible G N s p :
+ N <| G -> #|G : N| = p -> prime p -> G \subset 'I['chi_s] ->
+ {t | 'Res[N, G] 'chi_t = 'chi_s}.
+Proof.
+move=> nsNG iGN pr_p IGchi.
+have [t sGt] := constt_cfInd_irr s (normal_sub nsNG); exists t.
+have [e DtN]: exists e, 'Res 'chi_t = e%:R *: 'chi_s.
+ rewrite constt_Ind_Res in sGt.
+ rewrite (Clifford_Res_sum_cfclass nsNG sGt); set e := '[_, _].
+ rewrite cfclass_invariant // big_seq1.
+ by exists (truncC e); rewrite truncCK ?Cnat_cfdot_char ?cfRes_char ?irr_char.
+have [/irrWnorm/eqP | [c injc DtNc]] := cfRes_prime_irr_cases t nsNG iGN pr_p.
+ rewrite DtN cfnormZ cfnorm_irr normr_nat mulr1 -natrX pnatr_eq1.
+ by rewrite muln_eq1 andbb => /eqP->; rewrite scale1r.
+have nz_e: e != 0%N.
+ have: 'Res[N] 'chi_t != 0 by rewrite cfRes_eq0 // ?irr_char ?irr_neq0.
+ by rewrite DtN; apply: contraNneq => ->; rewrite scale0r.
+have [i s'ci]: exists i, c i != s.
+ pose i0 := Ordinal (prime_gt0 pr_p); pose i1 := Ordinal (prime_gt1 pr_p).
+ have [<- | ] := eqVneq (c i0) s; last by exists i0.
+ by exists i1; rewrite (inj_eq injc).
+have /esym/eqP/idPn[] := congr1 (cfdotr 'chi_(c i)) DtNc; rewrite {1}DtN /=.
+rewrite cfdot_suml cfdotZl cfdot_irr mulrb ifN_eqC // mulr0.
+rewrite (bigD1 i) //= cfnorm_irr big1 ?addr0 ?oner_eq0 // => j i'j.
+by rewrite cfdot_irr mulrb ifN_eq ?(inj_eq injc).
+Qed.
+
+(* This is Isaacs, Lemma (6.24). *)
+Lemma extend_to_cfdet G N s c0 u :
+ let theta := 'chi_s in let lambda := cfDet theta in let mu := 'chi_u in
+ N <| G -> coprime #|G : N| (truncC (theta 1%g)) ->
+ 'Res[N, G] 'chi_c0 = theta -> 'Res[N, G] mu = lambda ->
+ exists2 c, 'Res 'chi_c = theta /\ cfDet 'chi_c = mu
+ & forall c1, 'Res 'chi_c1 = theta -> cfDet 'chi_c1 = mu -> c1 = c.
+Proof.
+move=> theta lambda mu nsNG; set e := #|G : N|; set f := truncC _.
+set eta := 'chi_c0 => co_e_f etaNth muNlam; have [sNG nNG] := andP nsNG.
+have fE: f%:R = theta 1%g by rewrite truncCK ?Cnat_irr1.
+pose nu := cfDet eta; have lin_nu: nu \is a linear_char := cfDet_lin_char _.
+have nuNlam: 'Res nu = lambda by rewrite -cfDetRes ?irr_char ?etaNth.
+have lin_lam: lambda \is a linear_char := cfDet_lin_char _.
+have lin_mu: mu \is a linear_char.
+ by have:= lin_lam; rewrite -muNlam; apply: cfRes_lin_lin; apply: irr_char.
+have [Unu Ulam] := (lin_char_unitr lin_nu, lin_char_unitr lin_lam).
+pose alpha := mu / nu.
+have alphaN_1: 'Res[N] alpha = 1 by rewrite rmorph_div //= muNlam nuNlam divrr.
+have lin_alpha: alpha \is a linear_char by apply: rpred_div.
+have alpha_e: alpha ^+ e = 1.
+ have kerNalpha: N \subset cfker alpha.
+ by rewrite -subsetIidl -cfker_Res ?lin_charW // alphaN_1 cfker_cfun1.
+ apply/eqP; rewrite -(cfQuoK nsNG kerNalpha) -rmorphX cfMod_eq1 //.
+ rewrite -dvdn_cforder /e -card_quotient //.
+ by rewrite cforder_lin_char_dvdG ?cfQuo_lin_char.
+have det_alphaXeta b: cfDet (alpha ^+ b * eta) = alpha ^+ (b * f) * nu.
+ by rewrite cfDet_mul_lin ?rpredX ?irr_char // -exprM -(cfRes1 N) etaNth.
+have [b bf_mod_e]: exists b, b * f = 1 %[mod e].
+ rewrite -(chinese_modl co_e_f 1 0) /chinese !mul0n addn0 !mul1n mulnC.
+ by exists (egcdn f e).1.
+have alpha_bf: alpha ^+ (b * f) = alpha.
+ by rewrite -(expr_mod _ alpha_e) bf_mod_e expr_mod.
+have /irrP[c Dc]: alpha ^+ b * eta \in irr G.
+ by rewrite mul_lin_irr ?rpredX ?mem_irr.
+have chiN: 'Res 'chi_c = theta.
+ by rewrite -Dc rmorphM rmorphX /= alphaN_1 expr1n mul1r.
+have det_chi: cfDet 'chi_c = mu by rewrite -Dc det_alphaXeta alpha_bf divrK.
+exists c => // c2 c2Nth det_c2_mu; apply: irr_inj.
+have [irrMc _ imMc _] := constt_Ind_ext nsNG chiN.
+have /codomP[s2 Dc2]: c2 \in codom (@mul_mod_Iirr G N c).
+ by rewrite -imMc constt_Ind_Res c2Nth constt_irr ?inE.
+have{Dc2} Dc2: 'chi_c2 = ('chi_s2 %% N)%CF * 'chi_c.
+ by rewrite Dc2 cfIirrE // mod_IirrE.
+have s2_lin: 'chi_s2 \is a linear_char.
+ rewrite qualifE irr_char; apply/eqP/(mulIf (irr1_neq0 c)).
+ rewrite mul1r -[in rhs in _ = rhs](cfRes1 N) chiN -c2Nth cfRes1.
+ by rewrite Dc2 cfunE cfMod1.
+have s2Xf_1: 'chi_s2 ^+ f = 1.
+ apply/(can_inj (cfModK nsNG))/(mulIr (lin_char_unitr lin_mu))/esym.
+ rewrite rmorph1 rmorphX /= mul1r -{1}det_c2_mu Dc2 -det_chi.
+ by rewrite cfDet_mul_lin ?cfMod_lin_char ?irr_char // -(cfRes1 N) chiN.
+suffices /eqP s2_1: 'chi_s2 == 1 by rewrite Dc2 s2_1 rmorph1 mul1r.
+rewrite -['chi_s2]expr1 -dvdn_cforder -(eqnP co_e_f) dvdn_gcd.
+by rewrite /e -card_quotient ?cforder_lin_char_dvdG //= dvdn_cforder s2Xf_1.
+Qed.
+
+(* This is Isaacs, Theorem (6.25). *)
+Theorem solvable_irr_extendible_from_det G N s (theta := 'chi[N]_s) :
+ N <| G -> solvable (G / N) ->
+ G \subset 'I[theta] -> coprime #|G : N| (truncC (theta 1%g)) ->
+ [exists c, 'Res 'chi[G]_c == theta]
+ = [exists u, 'Res 'chi[G]_u == cfDet theta].
+Proof.
+set e := #|G : N|; set f := truncC _ => nsNG solG IGtheta co_e_f.
+apply/exists_eqP/exists_eqP=> [[c cNth] | [u uNdth]].
+ have /lin_char_irr/irrP[u Du] := cfDet_lin_char 'chi_c.
+ by exists u; rewrite -Du -cfDetRes ?irr_char ?cNth.
+move: {2}e.+1 (ltnSn e) => m.
+elim: m => // m IHm in G u e nsNG solG IGtheta co_e_f uNdth *.
+rewrite ltnS => le_e; have [sNG nNG] := andP nsNG.
+have [<- | ltNG] := eqsVneq N G; first by exists s; rewrite cfRes_id.
+have [G0 maxG0 sNG0]: {G0 | maxnormal (gval G0) G G & N \subset G0}.
+ by apply: maxgroup_exists; rewrite properEneq ltNG sNG.
+have [/andP[ltG0G nG0G] maxG0_P] := maxgroupP maxG0.
+set mu := 'chi_u in uNdth; have lin_mu: mu \is a linear_char.
+ by rewrite qualifE irr_char -(cfRes1 N) uNdth /= lin_char1 ?cfDet_lin_char.
+have sG0G := proper_sub ltG0G; have nsNG0 := normalS sNG0 sG0G nsNG.
+have nsG0G: G0 <| G by apply/andP.
+have /lin_char_irr/irrP[u0 Du0] := cfRes_lin_char G0 lin_mu.
+have u0Ndth: 'Res 'chi_u0 = cfDet theta by rewrite -Du0 cfResRes.
+have IG0theta: G0 \subset 'I[theta].
+ by rewrite (subset_trans sG0G) // -IGtheta subsetIr.
+have coG0f: coprime #|G0 : N| f by rewrite (coprime_dvdl _ co_e_f) ?indexSg.
+have{m IHm le_e} [c0 c0Ns]: exists c0, 'Res 'chi[G0]_c0 = theta.
+ have solG0: solvable (G0 / N) := solvableS (quotientS N sG0G) solG.
+ apply: IHm nsNG0 solG0 IG0theta coG0f u0Ndth (leq_trans _ le_e).
+ by rewrite -(ltn_pmul2l (cardG_gt0 N)) !Lagrange ?proper_card.
+have{c0 c0Ns} [c0 [c0Ns dc0_u0] Uc0] := extend_to_cfdet nsNG0 coG0f c0Ns u0Ndth.
+have IGc0: G \subset 'I['chi_c0].
+ apply/subsetP=> x Gx; rewrite inE (subsetP nG0G) //= -conjg_IirrE.
+ apply/eqP; congr 'chi__; apply: Uc0; rewrite conjg_IirrE.
+ by rewrite -(cfConjgRes _ nsG0G nsNG) // c0Ns inertiaJ ?(subsetP IGtheta).
+ by rewrite cfDetConjg dc0_u0 -Du0 (cfConjgRes _ _ nsG0G) // cfConjg_id.
+have prG0G: prime #|G : G0|.
+ have [h injh im_h] := third_isom sNG0 nsNG nsG0G.
+ rewrite -card_quotient // -im_h // card_injm //.
+ rewrite simple_sol_prime 1?quotient_sol //.
+ by rewrite /simple -(injm_minnormal injh) // im_h // maxnormal_minnormal.
+have [t tG0c0] := prime_invariant_irr_extendible nsG0G (erefl _) prG0G IGc0.
+by exists t; rewrite /theta -c0Ns -tG0c0 cfResRes.
+Qed.
+
+(* This is Isaacs, Theorem (6.26). *)
+Theorem extend_linear_char_from_Sylow G N (lambda : 'CF(N)) :
+ N <| G -> lambda \is a linear_char -> G \subset 'I[lambda] ->
+ (forall p, p \in \pi('o(lambda)%CF) ->
+ exists2 Hp : {group gT},
+ [/\ N \subset Hp, Hp \subset G & p.-Sylow(G / N) (Hp / N)%g]
+ & exists u, 'Res 'chi[Hp]_u = lambda) ->
+ exists u, 'Res[N, G] 'chi_u = lambda.
+Proof.
+set m := 'o(lambda)%CF => nsNG lam_lin IGlam p_ext_lam.
+have [sNG nNG] := andP nsNG; have linN := @cfRes_lin_lin _ _ N.
+wlog [p p_lam]: lambda @m lam_lin IGlam p_ext_lam /
+ exists p : nat, \pi(m) =i (p : nat_pred).
+- move=> IHp; have [linG [cf [inj_cf _ lin_cf onto_cf]]] := lin_char_group N.
+ case=> cf1 cfM cfX _ cf_order; have [lam cf_lam] := onto_cf _ lam_lin.
+ pose mu p := cf lam.`_p; pose pi_m p := p \in \pi(m).
+ have Dm: m = #[lam] by rewrite /m cfDet_order_lin // cf_lam cf_order.
+ have Dlambda: lambda = \prod_(p < m.+1 | pi_m p) mu p.
+ rewrite -(big_morph cf cfM cf1) big_mkcond cf_lam /pi_m Dm; congr (cf _).
+ rewrite -{1}[lam]prod_constt big_mkord; apply: eq_bigr => p _.
+ by case: ifPn => // p'lam; apply/constt1P; rewrite /p_elt p'natEpi.
+ have lin_mu p: mu p \is a linear_char by rewrite /mu cfX -cf_lam rpredX.
+ suffices /fin_all_exists [u uNlam] (p : 'I_m.+1):
+ exists u, pi_m p -> 'Res[N, G] 'chi_u = mu p.
+ - pose nu := \prod_(p < m.+1 | pi_m p) 'chi_(u p).
+ have lin_nu: nu \is a linear_char.
+ by apply: rpred_prod => p m_p; rewrite linN ?irr_char ?uNlam.
+ have /irrP[u1 Dnu] := lin_char_irr lin_nu.
+ by exists u1; rewrite Dlambda -Dnu rmorph_prod; apply: eq_bigr.
+ have [m_p | _] := boolP (pi_m p); last by exists 0.
+ have o_mu: \pi('o(mu p)%CF) =i (p : nat_pred).
+ rewrite cfDet_order_lin // cf_order orderE /=.
+ have [|pr_p _ [k ->]] := pgroup_pdiv (p_elt_constt p lam).
+ by rewrite cycle_eq1 (sameP eqP constt1P) /p_elt p'natEpi // negbK -Dm.
+ by move=> q; rewrite pi_of_exp // pi_of_prime.
+ have IGmu: G \subset 'I[mu p].
+ rewrite (subset_trans IGlam) // /mu cfX -cf_lam.
+ elim: (chinese _ _ _ _) => [|k IHk]; first by rewrite inertia1 norm_inertia.
+ by rewrite exprS (subset_trans _ (inertia_mul _ _)) // subsetIidl.
+ have [q||u] := IHp _ (lin_mu p) IGmu; [ | by exists p | by exists u].
+ rewrite o_mu => /eqnP-> {q}.
+ have [Hp sylHp [u uNlam]] := p_ext_lam p m_p; exists Hp => //.
+ rewrite /mu cfX -cf_lam -uNlam -rmorphX /=; set nu := _ ^+ _.
+ have /lin_char_irr/irrP[v ->]: nu \is a linear_char; last by exists v.
+ by rewrite rpredX // linN ?irr_char ?uNlam.
+have pi_m_p: p \in \pi(m) by rewrite p_lam !inE.
+have [pr_p mgt0]: prime p /\ (m > 0)%N.
+ by have:= pi_m_p; rewrite mem_primes => /and3P[].
+have p_m: p.-nat m by rewrite -(eq_pnat _ p_lam) pnat_pi.
+have{p_ext_lam} [H [sNH sHG sylHbar] [v vNlam]] := p_ext_lam p pi_m_p.
+have co_p_GH: coprime p #|G : H|.
+ rewrite -(index_quotient_eq _ sHG nNG) ?subIset ?sNH ?orbT //.
+ by rewrite (pnat_coprime (pnat_id pr_p)) //; have [] := and3P sylHbar.
+have lin_v: 'chi_v \is a linear_char by rewrite linN ?irr_char ?vNlam.
+pose nuG := 'Ind[G] 'chi_v.
+have [c vGc co_p_f]: exists2 c, c \in irr_constt nuG & ~~ (p %| 'chi_c 1%g)%C.
+ apply/exists_inP; rewrite -negb_forall_in.
+ apply: contraL co_p_GH => /forall_inP p_dv_v1.
+ rewrite prime_coprime // negbK -dvdC_nat -[rhs in (_ %| rhs)%C]mulr1.
+ rewrite -(lin_char1 lin_v) -cfInd1 // ['Ind _]cfun_sum_constt /=.
+ rewrite sum_cfunE rpred_sum // => i /p_dv_v1 p_dv_chi1i.
+ rewrite cfunE dvdC_mull // rpred_Cnat //.
+ by rewrite Cnat_cfdot_char ?cfInd_char ?irr_char.
+pose f := truncC ('chi_c 1%g); pose b := (egcdn f m).1.
+have fK: f%:R = 'chi_c 1%g by rewrite truncCK ?Cnat_irr1.
+have fb_mod_m: f * b = 1 %[mod m].
+ have co_m_f: coprime m f.
+ by rewrite (pnat_coprime p_m) ?p'natE // -dvdC_nat CdivE fK.
+ by rewrite -(chinese_modl co_m_f 1 0) /chinese !mul0n addn0 mul1n.
+have /irrP[s Dlam] := lin_char_irr lam_lin.
+have cHv: v \in irr_constt ('Res[H] 'chi_c) by rewrite -constt_Ind_Res.
+have{cHv} cNs: s \in irr_constt ('Res[N] 'chi_c).
+ rewrite -(cfResRes _ sNH) ?(constt_Res_trans _ cHv) ?cfRes_char ?irr_char //.
+ by rewrite vNlam Dlam constt_irr !inE.
+have DcN: 'Res[N] 'chi_c = lambda *+ f.
+ have:= Clifford_Res_sum_cfclass nsNG cNs.
+ rewrite cfclass_invariant -Dlam // big_seq1 Dlam => DcN.
+ have:= cfRes1 N 'chi_c; rewrite DcN cfunE -Dlam lin_char1 // mulr1 => ->.
+ by rewrite -scaler_nat fK.
+have /lin_char_irr/irrP[d Dd]: cfDet 'chi_c ^+ b \is a linear_char.
+ by rewrite rpredX // cfDet_lin_char.
+exists d; rewrite -{}Dd rmorphX /= -cfDetRes ?irr_char // DcN.
+rewrite cfDetMn ?lin_charW // -exprM cfDet_id //.
+rewrite -(expr_mod _ (exp_cforder _)) -cfDet_order_lin // -/m.
+by rewrite fb_mod_m /m cfDet_order_lin // expr_mod ?exp_cforder.
+Qed.
+
+(* This is Isaacs, Corollary (6.27). *)
+Corollary extend_coprime_linear_char G N (lambda : 'CF(N)) :
+ N <| G -> lambda \is a linear_char -> G \subset 'I[lambda] ->
+ coprime #|G : N| 'o(lambda)%CF ->
+ exists u, [/\ 'Res 'chi[G]_u = lambda, 'o('chi_u)%CF = 'o(lambda)%CF
+ & forall v,
+ 'Res 'chi_v = lambda -> coprime #|G : N| 'o('chi_v)%CF ->
+ v = u].
+Proof.
+set e := #|G : N| => nsNG lam_lin IGlam co_e_lam; have [sNG nNG] := andP nsNG.
+have [p lam_p | v vNlam] := extend_linear_char_from_Sylow nsNG lam_lin IGlam.
+ exists N; last first.
+ by have /irrP[u ->] := lin_char_irr lam_lin; exists u; rewrite cfRes_id.
+ split=> //; rewrite trivg_quotient /pHall sub1G pgroup1 indexg1.
+ rewrite card_quotient //= -/e (pi'_p'nat _ lam_p) //.
+ rewrite -coprime_pi' ?indexg_gt0 1?coprime_sym //.
+ by have:= lam_p; rewrite mem_primes => /and3P[].
+set nu := 'chi_v in vNlam.
+have lin_nu: nu \is a linear_char.
+ by rewrite (@cfRes_lin_lin _ _ N) ?vNlam ?irr_char.
+have [b be_mod_lam]: exists b, b * e = 1 %[mod 'o(lambda)%CF].
+ rewrite -(chinese_modr co_e_lam 0 1) /chinese !mul0n !mul1n mulnC.
+ by set b := _.1; exists b.
+have /irrP[u Du]: nu ^+ (b * e) \in irr G by rewrite lin_char_irr ?rpredX.
+exists u; set mu := 'chi_u in Du *.
+have uNlam: 'Res mu = lambda.
+ rewrite cfDet_order_lin // in be_mod_lam.
+ rewrite -Du rmorphX /= vNlam -(expr_mod _ (exp_cforder _)) //.
+ by rewrite be_mod_lam expr_mod ?exp_cforder.
+have lin_mu: mu \is a linear_char by rewrite -Du rpredX.
+have o_mu: ('o(mu) = 'o(lambda))%CF.
+ have dv_o_lam_mu: 'o(lambda)%CF %| 'o(mu)%CF.
+ by rewrite !cfDet_order_lin // -uNlam cforder_Res.
+ have kerNnu_olam: N \subset cfker (nu ^+ 'o(lambda)%CF).
+ rewrite -subsetIidl -cfker_Res ?rpredX ?irr_char //.
+ by rewrite rmorphX /= vNlam cfDet_order_lin // exp_cforder cfker_cfun1.
+ apply/eqP; rewrite eqn_dvd dv_o_lam_mu andbT cfDet_order_lin //.
+ rewrite dvdn_cforder -Du exprAC -dvdn_cforder dvdn_mull //.
+ rewrite -(cfQuoK nsNG kerNnu_olam) cforder_mod // /e -card_quotient //.
+ by rewrite cforder_lin_char_dvdG ?cfQuo_lin_char ?rpredX.
+split=> // t tNlam co_e_t.
+have lin_t: 'chi_t \is a linear_char.
+ by rewrite (@cfRes_lin_lin _ _ N) ?tNlam ?irr_char.
+have Ut := lin_char_unitr lin_t.
+have kerN_mu_t: N \subset cfker (mu / 'chi_t)%R.
+ rewrite -subsetIidl -cfker_Res ?lin_charW ?rpred_div ?rmorph_div //.
+ by rewrite /= uNlam tNlam divrr ?lin_char_unitr ?cfker_cfun1.
+have co_e_mu_t: coprime e #[(mu / 'chi_t)%R]%CF.
+ suffices dv_o_mu_t: #[(mu / 'chi_t)%R]%CF %| 'o(mu)%CF * 'o('chi_t)%CF.
+ by rewrite (coprime_dvdr dv_o_mu_t) // coprime_mulr o_mu co_e_lam.
+ rewrite !cfDet_order_lin //; apply/dvdn_cforderP=> x Gx.
+ rewrite invr_lin_char // !cfunE exprMn -rmorphX {2}mulnC.
+ by rewrite !(dvdn_cforderP _) ?conjC1 ?mulr1 // dvdn_mulr.
+have /eqP mu_t_1: mu / 'chi_t == 1.
+ rewrite -(dvdn_cforder (_ / _)%R 1) -(eqnP co_e_mu_t) dvdn_gcd dvdnn andbT.
+ rewrite -(cfQuoK nsNG kerN_mu_t) cforder_mod // /e -card_quotient //.
+ by rewrite cforder_lin_char_dvdG ?cfQuo_lin_char ?rpred_div.
+by apply: irr_inj; rewrite -['chi_t]mul1r -mu_t_1 divrK.
+Qed.
+
+(* This is Isaacs, Corollary (6.28). *)
+Corollary extend_solvable_coprime_irr G N t (theta := 'chi[N]_t) :
+ N <| G -> solvable (G / N) -> G \subset 'I[theta] ->
+ coprime #|G : N| ('o(theta)%CF * truncC (theta 1%g)) ->
+ exists c, [/\ 'Res 'chi[G]_c = theta, 'o('chi_c)%CF = 'o(theta)%CF
+ & forall d,
+ 'Res 'chi_d = theta -> coprime #|G : N| 'o('chi_d)%CF ->
+ d = c].
+Proof.
+set e := #|G : N|; set f := truncC _ => nsNG solG IGtheta.
+rewrite coprime_mulr => /andP[co_e_th co_e_f].
+have [sNG nNG] := andP nsNG; pose lambda := cfDet theta.
+have lin_lam: lambda \is a linear_char := cfDet_lin_char theta.
+have IGlam: G \subset 'I[lambda].
+ apply/subsetP=> y /(subsetP IGtheta)/setIdP[nNy /eqP th_y].
+ by rewrite inE nNy /= -cfDetConjg th_y.
+have co_e_lam: coprime e 'o(lambda)%CF by rewrite cfDet_order_lin.
+have [//|u [uNlam o_u Uu]] := extend_coprime_linear_char nsNG lin_lam IGlam.
+have /exists_eqP[c cNth]: [exists c, 'Res 'chi[G]_c == theta].
+ rewrite solvable_irr_extendible_from_det //.
+ by apply/exists_eqP; exists u.
+have{c cNth} [c [cNth det_c] Uc] := extend_to_cfdet nsNG co_e_f cNth uNlam.
+have lin_u: 'chi_u \is a linear_char by rewrite -det_c cfDet_lin_char.
+exists c; split=> // [|c0 c0Nth co_e_c0].
+ by rewrite !cfDet_order_lin // -det_c in o_u.
+have lin_u0: cfDet 'chi_c0 \is a linear_char := cfDet_lin_char 'chi_c0.
+have /irrP[u0 Du0] := lin_char_irr lin_u0.
+have co_e_u0: coprime e 'o('chi_u0)%CF by rewrite -Du0 cfDet_order_lin.
+have eq_u0u: u0 = u by apply: Uu; rewrite // -Du0 -cfDetRes ?irr_char ?c0Nth.
+by apply: Uc; rewrite // Du0 eq_u0u.
+Qed.
+
+End ExtendInvariantIrr.
+
+Section Frobenius.
+
+Variables (gT : finGroupType) (G K : {group gT}).
+
+(* Because he only defines Frobenius groups in chapter 7, Isaacs does not *)
+(* state these theorems using the Frobenius property. *)
+Hypothesis frobGK : [Frobenius G with kernel K].
+
+(* This is Isaacs, Theorem 6.34(a1). *)
+Theorem inertia_Frobenius_ker i : i != 0 -> 'I_G['chi[K]_i] = K.
+Proof.
+have [_ _ nsKG regK] := Frobenius_kerP frobGK; have [sKG nKG] := andP nsKG.
+move=> nzi; apply/eqP; rewrite eqEsubset sub_Inertia // andbT.
+apply/subsetP=> x /setIP[Gx /setIdP[nKx /eqP x_stab_i]].
+have actIirrK: is_action G (@conjg_Iirr _ K).
+ split=> [y j k eq_jk | j y z Gy Gz].
+ by apply/irr_inj/(can_inj (cfConjgK y)); rewrite -!conjg_IirrE eq_jk.
+ by apply: irr_inj; rewrite !conjg_IirrE (cfConjgM _ nsKG).
+pose ito := Action actIirrK; pose cto := ('Js \ (subsetT G))%act.
+have acts_Js : [acts G, on classes K | 'Js].
+ apply/subsetP=> y Gy; have nKy := subsetP nKG y Gy.
+ rewrite !inE; apply/subsetP=> _ /imsetP[z Gz ->]; rewrite !inE /=.
+ rewrite -class_rcoset norm_rlcoset // class_lcoset.
+ by apply: mem_imset; rewrite memJ_norm.
+have acts_cto : [acts G, on classes K | cto] by rewrite astabs_ract subsetIidl.
+pose m := #|'Fix_(classes K | cto)[x]|.
+have def_m: #|'Fix_ito[x]| = m.
+ apply: card_afix_irr_classes => // j y _ Ky /imsetP[_ /imsetP[z Kz ->] ->].
+ by rewrite conjg_IirrE cfConjgEJ // cfunJ.
+have: (m != 1)%N.
+ rewrite -def_m (cardD1 (0 : Iirr K)) (cardD1 i) !(inE, sub1set) /=.
+ by rewrite conjg_Iirr0 nzi eqxx -(inj_eq irr_inj) conjg_IirrE x_stab_i eqxx.
+apply: contraR => notKx; apply/cards1P; exists 1%g; apply/esym/eqP.
+rewrite eqEsubset !(sub1set, inE) classes1 /= conjs1g eqxx /=.
+apply/subsetP=> _ /setIP[/imsetP[y Ky ->] /afix1P /= cyKx].
+have /imsetP[z Kz def_yx]: y ^ x \in y ^: K.
+ by rewrite -cyKx; apply: mem_imset; exact: class_refl.
+rewrite inE classG_eq1; apply: contraR notKx => nty.
+rewrite -(groupMr x (groupVr Kz)).
+apply: (subsetP (regK y _)); first exact/setD1P.
+rewrite !inE groupMl // groupV (subsetP sKG) //=.
+by rewrite conjg_set1 conjgM def_yx conjgK.
+Qed.
+
+(* This is Isaacs, Theorem 6.34(a2) *)
+Theorem irr_induced_Frobenius_ker i : i != 0 -> 'Ind[G, K] 'chi_i \in irr G.
+Proof.
+move/inertia_Frobenius_ker/group_inj=> defK.
+have [_ _ nsKG _] := Frobenius_kerP frobGK.
+have [] := constt_Inertia_bijection i nsKG; rewrite defK cfInd_id => -> //.
+by rewrite constt_irr !inE.
+Qed.
+
+(* This is Isaacs, Theorem 6.34(b) *)
+Theorem Frobenius_Ind_irrP j :
+ reflect (exists2 i, i != 0 & 'chi_j = 'Ind[G, K] 'chi_i)
+ (~~ (K \subset cfker 'chi_j)).
+Proof.
+have [_ _ nsKG _] := Frobenius_kerP frobGK; have [sKG nKG] := andP nsKG.
+apply: (iffP idP) => [not_chijK1 | [i nzi ->]]; last first.
+ by rewrite cfker_Ind_irr ?sub_gcore // subGcfker.
+have /neq0_has_constt[i chijKi]: 'Res[K] 'chi_j != 0 by exact: Res_irr_neq0.
+have nz_i: i != 0.
+ by apply: contraNneq not_chijK1 => i0; rewrite constt0_Res_cfker // -i0.
+have /irrP[k def_chik] := irr_induced_Frobenius_ker nz_i.
+have: '['chi_j, 'chi_k] != 0 by rewrite -def_chik -cfdot_Res_l.
+by rewrite cfdot_irr pnatr_eq0; case: (j =P k) => // ->; exists i.
+Qed.
+
+End Frobenius.
diff --git a/mathcomp/character/integral_char.v b/mathcomp/character/integral_char.v
new file mode 100644
index 0000000..d327ac3
--- /dev/null
+++ b/mathcomp/character/integral_char.v
@@ -0,0 +1,708 @@
+(* (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.
+Require Import fingroup morphism perm automorphism quotient action finalg zmodp.
+Require Import commutator cyclic center pgroup sylow gseries nilpotent abelian.
+Require Import ssrnum ssrint polydiv rat matrix mxalgebra intdiv mxpoly.
+Require Import vector falgebra fieldext separable galois algC cyclotomic algnum.
+Require Import mxrepresentation classfun character.
+
+(******************************************************************************)
+(* This file provides some standard results based on integrality properties *)
+(* of characters, such as theorem asserting that the degree of an irreducible *)
+(* character of G divides the order of G (Isaacs 3.11), or the famous p^a.q^b *)
+(* solvability theorem of Burnside. *)
+(* Defined here: *)
+(* 'K_k == the kth class sum in gring F G, where k : 'I_#|classes G|, and *)
+(* F is inferred from the context. *)
+(* := gset_mx F G (enum_val k) (see mxrepresentation.v). *)
+(* --> The 'K_k form a basis of 'Z(group_ring F G)%MS. *)
+(* gring_classM_coef i j k == the coordinate of 'K_i *m 'K_j on 'K_k; this *)
+(* is usually abbreviated as a i j k. *)
+(* gring_classM_coef_set A B z == the set of all (x, y) in setX A B such *)
+(* that x * y = z; if A and B are respectively the ith and jth *)
+(* conjugacy class of G, and z is in the kth conjugacy class, then *)
+(* gring_classM_coef i j k is exactly the cadinal of this set. *)
+(* 'omega_i[A] == the mode of 'chi[G]_i on (A \in 'Z(group_ring algC G))%MS, *)
+(* i.e., the z such that gring_op 'Chi_i A = z%:M. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope GRing.Theory Num.Theory.
+Local Open Scope ring_scope.
+
+Lemma group_num_field_exists (gT : finGroupType) (G : {group gT}) :
+ {Qn : splittingFieldType rat & galois 1 {:Qn} &
+ {QnC : {rmorphism Qn -> algC}
+ & forall nuQn : argumentType (mem ('Gal({:Qn}%VS / 1%VS))),
+ {nu : {rmorphism algC -> algC} |
+ {morph QnC: a / nuQn a >-> nu a}}
+ & {w : Qn & #|G|.-primitive_root w /\ <<1; w>>%VS = fullv
+ & forall (hT : finGroupType) (H : {group hT}) (phi : 'CF(H)),
+ phi \is a character ->
+ forall x, (#[x] %| #|G|)%N -> {a | QnC a = phi x}}}}.
+Proof.
+have [z prim_z] := C_prim_root_exists (cardG_gt0 G); set n := #|G| in prim_z *.
+have [Qn [QnC [[|w []] // [Dz] genQn]]] := num_field_exists [:: z].
+have prim_w: n.-primitive_root w by rewrite -Dz fmorph_primitive_root in prim_z.
+have Q_Xn1: ('X^n - 1 : {poly Qn}) \is a polyOver 1%AS.
+ by rewrite rpredB ?rpred1 ?rpredX //= polyOverX.
+have splitXn1: splittingFieldFor 1 ('X^n - 1) {:Qn}.
+ pose r := codom (fun i : 'I_n => w ^+ i).
+ have Dr: 'X^n - 1 = \prod_(y <- r) ('X - y%:P).
+ by rewrite -(factor_Xn_sub_1 prim_w) big_mkord big_map enumT.
+ exists r; first by rewrite -Dr eqpxx.
+ apply/eqP; rewrite eqEsubv subvf -genQn adjoin_seqSr //; apply/allP=> /=.
+ by rewrite andbT -root_prod_XsubC -Dr; apply/unity_rootP/prim_expr_order.
+have Qn_ax : SplittingField.axiom Qn by exists ('X^n - 1).
+exists (SplittingFieldType _ _ Qn_ax).
+ apply/splitting_galoisField.
+ exists ('X^n - 1); split => //.
+ apply: separable_Xn_sub_1; rewrite -(fmorph_eq0 QnC) rmorph_nat.
+ by rewrite pnatr_eq0 -lt0n cardG_gt0.
+exists QnC => [// nuQn|].
+ by exact: (extend_algC_subfield_aut QnC [rmorphism of nuQn]).
+rewrite span_seq1 in genQn.
+exists w => // hT H phi Nphi x x_dv_n.
+apply: sig_eqW; have [rH ->] := char_reprP Nphi.
+have [Hx | /cfun0->] := boolP (x \in H); last by exists 0; rewrite rmorph0.
+have [e [_ [enx1 _] [-> _] _]] := repr_rsim_diag rH Hx.
+have /fin_all_exists[k Dk] i: exists k, e 0 i = z ^+ k.
+ have [|k ->] := (prim_rootP prim_z) (e 0 i); last by exists k.
+ by have /dvdnP[q ->] := x_dv_n; rewrite mulnC exprM enx1 expr1n.
+exists (\sum_i w ^+ k i); rewrite rmorph_sum; apply/eq_bigr => i _.
+by rewrite rmorphX Dz Dk.
+Qed.
+
+Section GenericClassSums.
+
+(* This is Isaacs, Theorem (2.4), generalized to an arbitrary field, and with *)
+(* the combinatorial definition of the coeficients exposed. *)
+(* This part could move to mxrepresentation.*)
+
+Variable (gT : finGroupType) (G : {group gT}) (F : fieldType).
+
+Definition gring_classM_coef_set (Ki Kj : {set gT}) g :=
+ [set xy in [predX Ki & Kj] | let: (x, y) := xy in x * y == g]%g.
+
+Definition gring_classM_coef (i j k : 'I_#|classes G|) :=
+ #|gring_classM_coef_set (enum_val i) (enum_val j) (repr (enum_val k))|.
+
+Definition gring_class_sum (i : 'I_#|classes G|) := gset_mx F G (enum_val i).
+
+Local Notation "''K_' i" := (gring_class_sum i)
+ (at level 8, i at level 2, format "''K_' i") : ring_scope.
+Local Notation a := gring_classM_coef.
+
+Lemma gring_class_sum_central i : ('K_i \in 'Z(group_ring F G))%MS.
+Proof. by rewrite -classg_base_center (eq_row_sub i) // rowK. Qed.
+
+Lemma set_gring_classM_coef (i j k : 'I_#|classes G|) g :
+ g \in enum_val k ->
+ a i j k = #|gring_classM_coef_set (enum_val i) (enum_val j) g|.
+Proof.
+rewrite /a; have /repr_classesP[] := enum_valP k; move: (repr _) => g1 Gg1 ->.
+have [/imsetP[zi Gzi ->] /imsetP[zj Gzj ->]] := (enum_valP i, enum_valP j).
+move=> g1Gg; have Gg := subsetP (class_subG Gg1 (subxx _)) _ g1Gg.
+set Aij := gring_classM_coef_set _ _.
+without loss suffices IH: g g1 Gg Gg1 g1Gg / (#|Aij g1| <= #|Aij g|)%N.
+ by apply/eqP; rewrite eqn_leq !IH // class_sym.
+have [w Gw Dg] := imsetP g1Gg; pose J2 (v : gT) xy := (xy.1 ^ v, xy.2 ^ v)%g.
+have J2inj: injective (J2 w).
+ by apply: can_inj (J2 w^-1)%g _ => [[x y]]; rewrite /J2 /= !conjgK.
+rewrite -(card_imset _ J2inj) subset_leq_card //; apply/subsetP.
+move=> _ /imsetP[[x y] /setIdP[/andP[/= x1Gx y1Gy] Dxy1] ->]; rewrite !inE /=.
+rewrite !(class_sym _ (_ ^ _)) !classGidl // class_sym x1Gx class_sym y1Gy.
+by rewrite -conjMg (eqP Dxy1) /= -Dg.
+Qed.
+
+Theorem gring_classM_expansion i j : 'K_i *m 'K_j = \sum_k (a i j k)%:R *: 'K_k.
+Proof.
+have [/imsetP[zi Gzi dKi] /imsetP[zj Gzj dKj]] := (enum_valP i, enum_valP j).
+pose aG := regular_repr F G; have sKG := subsetP (class_subG _ (subxx G)).
+transitivity (\sum_(x in zi ^: G) \sum_(y in zj ^: G) aG (x * y)%g).
+ rewrite mulmx_suml -/aG dKi; apply: eq_bigr => x /sKG Gx.
+ rewrite mulmx_sumr -/aG dKj; apply: eq_bigr => y /sKG Gy.
+ by rewrite repr_mxM ?Gx ?Gy.
+pose h2 xy : gT := (xy.1 * xy.2)%g.
+pose h1 xy := enum_rank_in (classes1 G) (h2 xy ^: G).
+rewrite pair_big (partition_big h1 xpredT) //=; apply: eq_bigr => k _.
+rewrite (partition_big h2 (mem (enum_val k))) /= => [|[x y]]; last first.
+ case/andP=> /andP[/= /sKG Gx /sKG Gy] /eqP <-.
+ by rewrite enum_rankK_in ?class_refl ?mem_classes ?groupM ?Gx ?Gy.
+rewrite scaler_sumr; apply: eq_bigr => g Kk_g; rewrite scaler_nat.
+rewrite (set_gring_classM_coef _ _ Kk_g) -sumr_const; apply: eq_big => [] [x y].
+ rewrite !inE /= dKi dKj /h1 /h2 /=; apply: andb_id2r => /eqP ->.
+ have /imsetP[zk Gzk dKk] := enum_valP k; rewrite dKk in Kk_g.
+ by rewrite (class_transr Kk_g) -dKk enum_valK_in eqxx andbT.
+by rewrite /h2 /= => /andP[_ /eqP->].
+Qed.
+
+Fact gring_irr_mode_key : unit. Proof. by []. Qed.
+Definition gring_irr_mode_def (i : Iirr G) := ('chi_i 1%g)^-1 *: 'chi_i.
+Definition gring_irr_mode := locked_with gring_irr_mode_key gring_irr_mode_def.
+Canonical gring_irr_mode_unlockable := [unlockable fun gring_irr_mode].
+
+End GenericClassSums.
+
+Arguments Scope gring_irr_mode [_ Group_scope ring_scope group_scope].
+
+Notation "''K_' i" := (gring_class_sum _ i)
+ (at level 8, i at level 2, format "''K_' i") : ring_scope.
+
+Notation "''omega_' i [ A ]" := (xcfun (gring_irr_mode i) A)
+ (at level 8, i at level 2, format "''omega_' i [ A ]") : ring_scope.
+
+Section IntegralChar.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+(* This is Isaacs, Corollary (3.6). *)
+Lemma Aint_char (chi : 'CF(G)) x : chi \is a character -> chi x \in Aint.
+Proof.
+have [Gx /char_reprP[rG ->] {chi} | /cfun0->//] := boolP (x \in G).
+have [e [_ [unit_e _] [-> _] _]] := repr_rsim_diag rG Gx.
+rewrite rpred_sum // => i _; apply: (@Aint_unity_root #[x]) => //.
+exact/unity_rootP.
+Qed.
+
+Lemma Aint_irr i x : 'chi[G]_i x \in Aint.
+Proof. by apply: Aint_char; exact: irr_char. Qed.
+
+Local Notation R_G := (group_ring algCfield G).
+Local Notation a := gring_classM_coef.
+
+(* This is Isaacs (2.25). *)
+Lemma mx_irr_gring_op_center_scalar n (rG : mx_representation algCfield G n) A :
+ mx_irreducible rG -> (A \in 'Z(R_G))%MS -> is_scalar_mx (gring_op rG A).
+Proof.
+move/groupC=> irrG /center_mxP[R_A cGA].
+apply: mx_abs_irr_cent_scalar irrG _ _; apply/centgmxP => x Gx.
+by rewrite -(gring_opG rG Gx) -!gring_opM ?cGA // envelop_mx_id.
+Qed.
+
+Section GringIrrMode.
+
+Variable i : Iirr G.
+
+Let n := irr_degree (socle_of_Iirr i).
+Let mxZn_inj: injective (@scalar_mx algCfield n).
+Proof. by rewrite -[n]prednK ?irr_degree_gt0 //; apply: fmorph_inj. Qed.
+
+Lemma cfRepr_gring_center n1 (rG : mx_representation algCfield G n1) A :
+ cfRepr rG = 'chi_i -> (A \in 'Z(R_G))%MS -> gring_op rG A = 'omega_i[A]%:M.
+Proof.
+move=> def_rG Z_A; rewrite unlock xcfunZl -{2}def_rG xcfun_repr.
+have irr_rG: mx_irreducible rG.
+ have sim_rG: mx_rsim 'Chi_i rG by apply: cfRepr_inj; rewrite irrRepr.
+ exact: mx_rsim_irr sim_rG (socle_irr _).
+have /is_scalar_mxP[e ->] := mx_irr_gring_op_center_scalar irr_rG Z_A.
+congr _%:M; apply: (canRL (mulKf (irr1_neq0 i))).
+by rewrite mulrC -def_rG cfunE repr_mx1 group1 -mxtraceZ scalemx1.
+Qed.
+
+Lemma irr_gring_center A :
+ (A \in 'Z(R_G))%MS -> gring_op 'Chi_i A = 'omega_i[A]%:M.
+Proof. exact: cfRepr_gring_center (irrRepr i). Qed.
+
+Lemma gring_irr_modeM A B :
+ (A \in 'Z(R_G))%MS -> (B \in 'Z(R_G))%MS ->
+ 'omega_i[A *m B] = 'omega_i[A] * 'omega_i[B].
+Proof.
+move=> Z_A Z_B; have [[R_A cRA] [R_B cRB]] := (center_mxP Z_A, center_mxP Z_B).
+apply: mxZn_inj; rewrite scalar_mxM -!irr_gring_center ?gring_opM //.
+apply/center_mxP; split=> [|C R_C]; first exact: envelop_mxM.
+by rewrite mulmxA cRA // -!mulmxA cRB.
+Qed.
+
+Lemma gring_mode_class_sum_eq (k : 'I_#|classes G|) g :
+ g \in enum_val k -> 'omega_i['K_k] = #|g ^: G|%:R * 'chi_i g / 'chi_i 1%g.
+Proof.
+have /imsetP[x Gx DxG] := enum_valP k; rewrite DxG => /imsetP[u Gu ->{g}].
+rewrite unlock classGidl ?cfunJ {u Gu}// mulrC mulr_natl.
+rewrite xcfunZl raddf_sum DxG -sumr_const /=; congr (_ * _).
+by apply: eq_bigr => _ /imsetP[u Gu ->]; rewrite xcfunG ?groupJ ?cfunJ.
+Qed.
+
+(* This is Isaacs, Theorem (3.7). *)
+Lemma Aint_gring_mode_class_sum k : 'omega_i['K_k] \in Aint.
+Proof.
+move: k; pose X := [tuple 'omega_i['K_k] | k < #|classes G| ].
+have memX k: 'omega_i['K_k] \in X by apply: map_f; exact: mem_enum.
+have S_P := Cint_spanP X; set S := Cint_span X in S_P.
+have S_X: {subset X <= S} by exact: mem_Cint_span.
+have S_1: 1 \in S.
+ apply: S_X; apply/codomP; exists (enum_rank_in (classes1 G) 1%g).
+ rewrite (@gring_mode_class_sum_eq _ 1%g) ?enum_rankK_in ?classes1 //.
+ by rewrite mulfK ?irr1_neq0 // class1G cards1.
+suffices Smul: mulr_closed S.
+ by move=> k; apply: fin_Csubring_Aint S_P _ _; rewrite ?S_X.
+split=> // _ _ /S_P[x ->] /S_P[y ->].
+rewrite mulr_sumr rpred_sum // => j _.
+rewrite mulrzAr mulr_suml rpredMz ?rpred_sum // => k _.
+rewrite mulrzAl rpredMz {x y}// !nth_mktuple.
+rewrite -gring_irr_modeM ?gring_class_sum_central //.
+rewrite gring_classM_expansion raddf_sum rpred_sum // => jk _.
+by rewrite scaler_nat raddfMn rpredMn ?S_X ?memX.
+Qed.
+
+(* A more usable reformulation that does not involve the class sums. *)
+Corollary Aint_class_div_irr1 x :
+ x \in G -> #|x ^: G|%:R * 'chi_i x / 'chi_i 1%g \in Aint.
+Proof.
+move=> Gx; have clGxG := mem_classes Gx; pose k := enum_rank_in clGxG (x ^: G).
+have k_x: x \in enum_val k by rewrite enum_rankK_in // class_refl.
+by rewrite -(gring_mode_class_sum_eq k_x) Aint_gring_mode_class_sum.
+Qed.
+
+(* This is Isaacs, Theorem (3.8). *)
+Theorem coprime_degree_support_cfcenter g :
+ coprime (truncC ('chi_i 1%g)) #|g ^: G| -> g \notin ('Z('chi_i))%CF ->
+ 'chi_i g = 0.
+Proof.
+set m := truncC _ => co_m_gG notZg.
+have [Gg | /cfun0-> //] := boolP (g \in G).
+have Dm: 'chi_i 1%g = m%:R by rewrite truncCK ?Cnat_irr1.
+have m_gt0: (0 < m)%N by rewrite -ltC_nat -Dm irr1_gt0.
+have nz_m: m%:R != 0 :> algC by rewrite pnatr_eq0 -lt0n.
+pose alpha := 'chi_i g / m%:R.
+have a_lt1: `|alpha| < 1.
+ rewrite normrM normfV normr_nat -{2}(divff nz_m).
+ rewrite ltr_def (can_eq (mulfVK nz_m)) eq_sym -{1}Dm -irr_cfcenterE // notZg.
+ by rewrite ler_pmul2r ?invr_gt0 ?ltr0n // -Dm char1_ge_norm ?irr_char.
+have Za: alpha \in Aint.
+ have [u _ /dvdnP[v eq_uv]] := Bezoutl #|g ^: G| m_gt0.
+ suffices ->: alpha = v%:R * 'chi_i g - u%:R * (alpha * #|g ^: G|%:R).
+ rewrite rpredB // rpredM ?rpred_nat ?Aint_irr //.
+ by rewrite mulrC mulrA -Dm Aint_class_div_irr1.
+ rewrite -mulrCA -[v%:R](mulfK nz_m) -!natrM -eq_uv (eqnP co_m_gG).
+ by rewrite mulrAC -mulrA -/alpha mulr_natl mulr_natr mulrS addrK.
+have [Qn galQn [QnC gQnC [_ _ Qn_g]]] := group_num_field_exists <[g]>.
+have{Qn_g} [a Da]: exists a, QnC a = alpha.
+ rewrite /alpha; have [a <-] := Qn_g _ G _ (irr_char i) g (dvdnn _).
+ by exists (a / m%:R); rewrite fmorph_div rmorph_nat.
+have Za_nu nu: sval (gQnC nu) alpha \in Aint by rewrite Aint_aut.
+have norm_a_nu nu: `|sval (gQnC nu) alpha| <= 1.
+ move: {nu}(sval _) => nu; rewrite fmorph_div rmorph_nat normrM normfV.
+ rewrite normr_nat -Dm -(ler_pmul2r (irr1_gt0 (aut_Iirr nu i))) mul1r.
+ congr (_ <= _): (char1_ge_norm g (irr_char (aut_Iirr nu i))).
+ by rewrite !aut_IirrE !cfunE Dm rmorph_nat divfK.
+pose beta := QnC (galNorm 1 {:Qn} a).
+have Dbeta: beta = \prod_(nu in 'Gal({:Qn} / 1)) sval (gQnC nu) alpha.
+ rewrite /beta rmorph_prod. apply: eq_bigr => nu _.
+ by case: (gQnC nu) => f /= ->; rewrite Da.
+have Zbeta: beta \in Cint.
+ apply: Cint_rat_Aint; last by rewrite Dbeta rpred_prod.
+ rewrite /beta; have /vlineP[/= c ->] := mem_galNorm galQn (memvf a).
+ by rewrite alg_num_field fmorph_rat rpred_rat.
+have [|nz_a] := boolP (alpha == 0).
+ by rewrite (can2_eq (divfK _) (mulfK _)) // mul0r => /eqP.
+have: beta != 0 by rewrite Dbeta; apply/prodf_neq0 => nu _; rewrite fmorph_eq0.
+move/(norm_Cint_ge1 Zbeta); rewrite ltr_geF //; apply: ler_lt_trans a_lt1.
+rewrite -[`|alpha|]mulr1 Dbeta (bigD1 1%g) ?group1 //= -Da.
+case: (gQnC _) => /= _ <-; rewrite gal_id normrM.
+rewrite -subr_ge0 -mulrBr mulr_ge0 ?normr_ge0 // Da subr_ge0.
+elim/big_rec: _ => [|nu c _]; first by rewrite normr1 lerr.
+apply: ler_trans; rewrite -subr_ge0 -{1}[`|c|]mul1r normrM -mulrBl.
+by rewrite mulr_ge0 ?normr_ge0 // subr_ge0 norm_a_nu.
+Qed.
+
+End GringIrrMode.
+
+(* This is Isaacs, Theorem (3.9). *)
+Theorem primes_class_simple_gt1 C :
+ simple G -> ~~ abelian G -> C \in (classes G)^# -> (size (primes #|C|) > 1)%N.
+Proof.
+move=> simpleG not_cGG /setD1P[ntC /imsetP[g Gg defC]].
+have{ntC} nt_g: g != 1%g by rewrite defC classG_eq1 in ntC.
+rewrite ltnNge {C}defC; set m := #|_|; apply/negP=> p_natC.
+have{p_natC} [p p_pr [a Dm]]: {p : nat & prime p & {a | m = p ^ a}%N}.
+ have /prod_prime_decomp->: (0 < m)%N by rewrite /m -index_cent1.
+ rewrite prime_decompE; case Dpr: (primes _) p_natC => [|p []] // _.
+ by exists 2 => //; rewrite big_nil; exists 0%N.
+ rewrite big_seq1; exists p; last by exists (logn p m).
+ by have:= mem_primes p m; rewrite Dpr mem_head => /esym/and3P[].
+have{simpleG} [ntG minG] := simpleP _ simpleG.
+pose p_dv1 i := (p %| 'chi[G]_i 1%g)%C.
+have p_dvd_supp_g i: ~~ p_dv1 i && (i != 0) -> 'chi_i g = 0.
+ rewrite /p_dv1 irr1_degree dvdC_nat -prime_coprime // => /andP[co_p_i1 nz_i].
+ have fful_i: cfker 'chi_i = [1].
+ have /minG[//|/eqP] := cfker_normal 'chi_i.
+ by rewrite eqEsubset subGcfker (negPf nz_i) andbF.
+ have trivZ: 'Z(G) = [1] by have /minG[|/center_idP/idPn] := center_normal G.
+ have trivZi: ('Z('chi_i))%CF = [1].
+ apply/trivgP; rewrite -quotient_sub1 ?norms1 //= -fful_i cfcenter_eq_center.
+ rewrite fful_i subG1 -(isog_eq1 (isog_center (quotient1_isog G))) /=.
+ by rewrite trivZ.
+ rewrite coprime_degree_support_cfcenter ?trivZi ?inE //.
+ by rewrite -/m Dm irr1_degree natCK coprime_sym coprime_expl.
+pose alpha := \sum_(i | p_dv1 i && (i != 0)) 'chi_i 1%g / p%:R * 'chi_i g.
+have nz_p: p%:R != 0 :> algC by rewrite pnatr_eq0 -lt0n prime_gt0.
+have Dalpha: alpha = - 1 / p%:R.
+ apply/(canRL (mulfK nz_p))/eqP; rewrite -addr_eq0 addrC; apply/eqP/esym.
+ transitivity (cfReg G g); first by rewrite cfRegE (negPf nt_g).
+ rewrite cfReg_sum sum_cfunE (bigD1 0) //= irr0 !cfunE cfun11 cfun1E Gg.
+ rewrite mulr1; congr (1 + _); rewrite (bigID p_dv1) /= addrC big_andbC.
+ rewrite big1 => [|i /p_dvd_supp_g chig0]; last by rewrite cfunE chig0 mulr0.
+ rewrite add0r big_andbC mulr_suml; apply: eq_bigr => i _.
+ by rewrite mulrAC divfK // cfunE.
+suffices: (p %| 1)%C by rewrite (dvdC_nat p 1) dvdn1 -(subnKC (prime_gt1 p_pr)).
+rewrite unfold_in (negPf nz_p).
+rewrite Cint_rat_Aint ?rpred_div ?rpred1 ?rpred_nat //.
+rewrite -rpredN // -mulNr -Dalpha rpred_sum // => i /andP[/dvdCP[c Zc ->] _].
+by rewrite mulfK // rpredM ?Aint_irr ?Aint_Cint.
+Qed.
+
+End IntegralChar.
+
+Section MoreIntegralChar.
+
+Implicit Type gT : finGroupType.
+
+(* This is Burnside's famous p^a.q^b theorem (Isaacs, Theorem (3.10)). *)
+Theorem Burnside_p_a_q_b gT (G : {group gT}) :
+ (size (primes #|G|) <= 2)%N -> solvable G.
+Proof.
+move: {2}_.+1 (ltnSn #|G|) => n; elim: n => // n IHn in gT G *.
+rewrite ltnS => leGn piGle2; have [simpleG | ] := boolP (simple G); last first.
+ rewrite negb_forall_in => /exists_inP[N sNG]; rewrite eq_sym.
+ have [-> | ] := altP (N =P G).
+ rewrite groupP /= genGid normG andbT eqb_id negbK => /eqP->.
+ exact: solvable1.
+ rewrite [N == G]eqEproper sNG eqbF_neg !negbK => ltNG /and3P[grN].
+ case/isgroupP: grN => {N}N -> in sNG ltNG *; rewrite /= genGid => ntN nNG.
+ have nsNG: N <| G by exact/andP.
+ have dv_le_pi m: (m %| #|G| -> size (primes m) <= 2)%N.
+ move=> m_dv_G; apply: leq_trans piGle2.
+ by rewrite uniq_leq_size ?primes_uniq //; apply: pi_of_dvd.
+ rewrite (series_sol nsNG) !IHn ?dv_le_pi ?cardSg ?dvdn_quotient //.
+ by apply: leq_trans leGn; apply: ltn_quotient.
+ by apply: leq_trans leGn; apply: proper_card.
+have [->|[p p_pr p_dv_G]] := trivgVpdiv G; first exact: solvable1.
+have piGp: p \in \pi(G) by rewrite mem_primes p_pr cardG_gt0.
+have [P sylP] := Sylow_exists p G; have [sPG pP p'GP] := and3P sylP.
+have ntP: P :!=: 1%g by rewrite -rank_gt0 (rank_Sylow sylP) p_rank_gt0.
+have /trivgPn[g /setIP[Pg cPg] nt_g]: 'Z(P) != 1%g.
+ by rewrite center_nil_eq1 // (pgroup_nil pP).
+apply: abelian_sol; have: (size (primes #|g ^: G|) <= 1)%N.
+ rewrite -ltnS -[_.+1]/(size (p :: _)) (leq_trans _ piGle2) //.
+ rewrite -index_cent1 uniq_leq_size // => [/= | q].
+ rewrite primes_uniq -p'natEpi ?(pnat_dvd _ p'GP) ?indexgS //.
+ by rewrite subsetI sPG sub_cent1.
+ by rewrite inE => /predU1P[-> // |]; apply: pi_of_dvd; rewrite ?dvdn_indexg.
+rewrite leqNgt; apply: contraR => /primes_class_simple_gt1-> //.
+by rewrite !inE classG_eq1 nt_g mem_classes // (subsetP sPG).
+Qed.
+
+(* This is Isaacs, Theorem (3.11). *)
+Theorem dvd_irr1_cardG gT (G : {group gT}) i : ('chi[G]_i 1%g %| #|G|)%C.
+Proof.
+rewrite unfold_in -if_neg irr1_neq0 Cint_rat_Aint //=.
+ by rewrite rpred_div ?rpred_nat // rpred_Cnat ?Cnat_irr1.
+rewrite -[n in n / _]/(_ *+ true) -(eqxx i) -mulr_natr.
+rewrite -first_orthogonality_relation mulVKf ?neq0CG //.
+rewrite sum_by_classes => [|x y Gx Gy]; rewrite -?conjVg ?cfunJ //.
+rewrite mulr_suml rpred_sum // => K /repr_classesP[Gx {1}->].
+by rewrite !mulrA mulrAC rpredM ?Aint_irr ?Aint_class_div_irr1.
+Qed.
+
+(* This is Isaacs, Theorem (3.12). *)
+Theorem dvd_irr1_index_center gT (G : {group gT}) i :
+ ('chi[G]_i 1%g %| #|G : 'Z('chi_i)%CF|)%C.
+Proof.
+without loss fful: gT G i / cfaithful 'chi_i.
+ rewrite -{2}[i](quo_IirrK _ (subxx _)) ?mod_IirrE ?cfModE ?cfker_normal //.
+ rewrite morph1; set i1 := quo_Iirr _ i => /(_ _ _ i1) IH.
+ have fful_i1: cfaithful 'chi_i1.
+ by rewrite quo_IirrE ?cfker_normal ?cfaithful_quo.
+ have:= IH fful_i1; rewrite cfcenter_fful_irr // -cfcenter_eq_center.
+ rewrite index_quotient_eq ?cfcenter_sub ?cfker_norm //.
+ by rewrite setIC subIset // normal_sub ?cfker_center_normal.
+have [lambda lin_lambda Dlambda] := cfcenter_Res 'chi_i.
+have DchiZ: {in G & 'Z(G), forall x y, 'chi_i (x * y)%g = 'chi_i x * lambda y}.
+ rewrite -(cfcenter_fful_irr fful) => x y Gx Zy.
+ apply: (mulfI (irr1_neq0 i)); rewrite mulrCA.
+ transitivity ('chi_i x * ('chi_i 1%g *: lambda) y); last by rewrite !cfunE.
+ rewrite -Dlambda cfResE ?cfcenter_sub //.
+ rewrite -irrRepr cfcenter_repr !cfunE in Zy *.
+ case/setIdP: Zy => Gy /is_scalar_mxP[e De].
+ rewrite repr_mx1 group1 (groupM Gx Gy) (repr_mxM _ Gx Gy) Gx Gy De.
+ by rewrite mul_mx_scalar mxtraceZ mulrCA mulrA mulrC -mxtraceZ scalemx1.
+have inj_lambda: {in 'Z(G) &, injective lambda}.
+ rewrite -(cfcenter_fful_irr fful) => x y Zx Zy eq_xy.
+ apply/eqP; rewrite eq_mulVg1 -in_set1 (subsetP fful) // cfkerEirr inE.
+ apply/eqP; transitivity ('Res['Z('chi_i)%CF] 'chi_i (x^-1 * y)%g).
+ by rewrite cfResE ?cfcenter_sub // groupM ?groupV.
+ rewrite Dlambda !cfunE lin_charM ?groupV // -eq_xy -lin_charM ?groupV //.
+ by rewrite mulrC mulVg lin_char1 ?mul1r.
+rewrite unfold_in -if_neg irr1_neq0 Cint_rat_Aint //.
+ by rewrite rpred_div ?rpred_nat // rpred_Cnat ?Cnat_irr1.
+rewrite (cfcenter_fful_irr fful) nCdivE natf_indexg ?center_sub //=.
+have ->: #|G|%:R = \sum_(x in G) 'chi_i x * 'chi_i (x^-1)%g.
+ rewrite -[_%:R]mulr1; apply: canLR (mulVKf (neq0CG G)) _.
+ by rewrite first_orthogonality_relation eqxx.
+rewrite (big_setID [set x | 'chi_i x == 0]) /= -setIdE.
+rewrite big1 ?add0r => [| x /setIdP[_ /eqP->]]; last by rewrite mul0r.
+pose h x := (x ^: G * 'Z(G))%g; rewrite (partition_big_imset h).
+rewrite !mulr_suml rpred_sum //= => _ /imsetP[x /setDP[Gx nz_chi_x] ->].
+have: #|x ^: G|%:R * ('chi_i x * 'chi_i x^-1%g) / 'chi_i 1%g \in Aint.
+ by rewrite !mulrA mulrAC rpredM ?Aint_irr ?Aint_class_div_irr1.
+congr 2 (_ * _ \in Aint); apply: canRL (mulfK (neq0CG _)) _.
+rewrite inE in nz_chi_x.
+transitivity ('chi_i x * 'chi_i (x^-1)%g *+ #|h x|); last first.
+ rewrite -sumr_const.
+ apply: eq_big => [y | _ /mulsgP[_ z /imsetP[u Gu ->] Zz] ->].
+ rewrite !inE -andbA; apply/idP/and3P=> [|[_ _ /eqP <-]]; last first.
+ by rewrite -{1}[y]mulg1 mem_mulg ?class_refl.
+ case/mulsgP=> _ z /imsetP[u Gu ->] Zz ->; have /centerP[Gz cGz] := Zz.
+ rewrite groupM 1?DchiZ ?groupJ ?cfunJ //; split=> //.
+ by rewrite mulf_neq0 // lin_char_neq0 /= ?cfcenter_fful_irr.
+ rewrite -[z](mulKg u) -cGz // -conjMg /h classGidl {u Gu}//.
+ apply/eqP/setP=> w; apply/mulsgP/mulsgP=> [][_ z1 /imsetP[v Gv ->] Zz1 ->].
+ exists (x ^ v)%g (z * z1)%g; rewrite ?mem_imset ?groupM //.
+ by rewrite conjMg -mulgA /(z ^ v)%g cGz // mulKg.
+ exists ((x * z) ^ v)%g (z^-1 * z1)%g; rewrite ?mem_imset ?groupM ?groupV //.
+ by rewrite conjMg -mulgA /(z ^ v)%g cGz // mulKg mulKVg.
+ rewrite !irr_inv DchiZ ?groupJ ?cfunJ // rmorphM mulrACA -!normCK -exprMn.
+ by rewrite (normC_lin_char lin_lambda) ?mulr1 //= cfcenter_fful_irr.
+rewrite mulrAC -natrM mulr_natl; congr (_ *+ _).
+symmetry; rewrite /h /mulg /= /set_mulg [in _ @2: (_, _)]unlock cardsE.
+rewrite -cardX card_in_image // => [] [y1 z1] [y2 z2] /=.
+move=> /andP[/=/imsetP[u1 Gu1 ->] Zz1] /andP[/=/imsetP[u2 Gu2 ->] Zz2] {y1 y2}.
+move=> eq12; have /eqP := congr1 'chi_i eq12.
+rewrite !(cfunJ, DchiZ) ?groupJ // (can_eq (mulKf nz_chi_x)).
+rewrite (inj_in_eq inj_lambda) // => /eqP eq_z12; rewrite eq_z12 in eq12 *.
+by rewrite (mulIg _ _ _ eq12).
+Qed.
+
+(* This is Isaacs, Problem (3.7). *)
+Lemma gring_classM_coef_sum_eq gT (G : {group gT}) j1 j2 k g1 g2 g :
+ let a := @gring_classM_coef gT G j1 j2 in let a_k := a k in
+ g1 \in enum_val j1 -> g2 \in enum_val j2 -> g \in enum_val k ->
+ let sum12g := \sum_i 'chi[G]_i g1 * 'chi_i g2 * ('chi_i g)^* / 'chi_i 1%g in
+ a_k%:R = (#|enum_val j1| * #|enum_val j2|)%:R / #|G|%:R * sum12g.
+Proof.
+move=> a /= Kg1 Kg2 Kg; rewrite mulrAC; apply: canRL (mulfK (neq0CG G)) _.
+transitivity (\sum_j (#|G| * a j)%:R *+ (j == k) : algC).
+ by rewrite (bigD1 k) //= eqxx -natrM mulnC big1 ?addr0 // => j /negPf->.
+have defK (j : 'I_#|classes G|) x: x \in enum_val j -> enum_val j = x ^: G.
+ by have /imsetP[y Gy ->] := enum_valP j => /class_transr.
+have Gg: g \in G.
+ by case/imsetP: (enum_valP k) Kg => x Gx -> /imsetP[y Gy ->]; apply: groupJ.
+transitivity (\sum_j \sum_i 'omega_i['K_j] * 'chi_i 1%g * ('chi_i g)^* *+ a j).
+ apply: eq_bigr => j _; have /imsetP[z Gz Dj] := enum_valP j.
+ have Kz: z \in enum_val j by rewrite Dj class_refl.
+ rewrite -(Lagrange (subsetIl G 'C[z])) index_cent1 -mulnA natrM -mulrnAl.
+ have ->: (j == k) = (z \in enum_val k).
+ by rewrite -(inj_eq enum_val_inj); apply/eqP/idP=> [<-|/defK->].
+ rewrite (defK _ g) // -second_orthogonality_relation // mulr_suml.
+ apply: eq_bigr=> i _; rewrite natrM mulrA mulr_natr mulrC mulrA.
+ by rewrite (gring_mode_class_sum_eq i Kz) divfK ?irr1_neq0.
+rewrite exchange_big /= mulr_sumr; apply: eq_bigr => i _.
+transitivity ('omega_i['K_j1 *m 'K_j2] * 'chi_i 1%g * ('chi_i g)^*).
+ rewrite gring_classM_expansion -/a raddf_sum !mulr_suml /=.
+ by apply: eq_bigr => j _; rewrite xcfunZr -!mulrA mulr_natl.
+rewrite !mulrA 2![_ / _]mulrAC (defK _ _ Kg1) (defK _ _ Kg2); congr (_ * _).
+rewrite gring_irr_modeM ?gring_class_sum_central // mulnC natrM.
+rewrite (gring_mode_class_sum_eq i Kg2) !mulrA divfK ?irr1_neq0 //.
+by congr (_ * _); rewrite [_ * _]mulrC (gring_mode_class_sum_eq i Kg1) !mulrA.
+Qed.
+
+(* This is Isaacs, Problem (2.16). *)
+Lemma index_support_dvd_degree gT (G H : {group gT}) chi :
+ H \subset G -> chi \is a character -> chi \in 'CF(G, H) ->
+ (H :==: 1%g) || abelian G ->
+ (#|G : H| %| chi 1%g)%C.
+Proof.
+move=> sHG Nchi Hchi ZHG.
+suffices: (#|G : H| %| 'Res[H] chi 1%g)%C by rewrite cfResE ?group1.
+rewrite ['Res _]cfun_sum_cfdot sum_cfunE rpred_sum // => i _.
+rewrite cfunE dvdC_mulr ?Cint_Cnat ?Cnat_irr1 //.
+have [j ->]: exists j, 'chi_i = 'Res 'chi[G]_j.
+ case/predU1P: ZHG => [-> | cGG] in i *.
+ suffices ->: i = 0 by exists 0; rewrite !irr0 cfRes_cfun1 ?sub1G.
+ apply/val_inj; case: i => [[|i] //=]; rewrite ltnNge NirrE.
+ by rewrite (@leq_trans 1) // leqNgt classes_gt1 eqxx.
+ have linG := char_abelianP G cGG; have linG1 j := eqP (proj2 (andP (linG j))).
+ have /fin_all_exists[rH DrH] j: exists k, 'Res[H, G] 'chi_j = 'chi_k.
+ apply/irrP/lin_char_irr/andP.
+ by rewrite cfRes_char ?irr_char // cfRes1 ?linG1.
+ suffices{i} all_rH: codom rH =i Iirr H.
+ by exists (iinv (all_rH i)); rewrite DrH f_iinv.
+ apply/subset_cardP; last exact/subsetP; apply/esym/eqP.
+ rewrite card_Iirr_abelian ?(abelianS sHG) //.
+ rewrite -(eqn_pmul2r (indexg_gt0 G H)) Lagrange //; apply/eqP.
+ rewrite -sum_nat_const -card_Iirr_abelian // -sum1_card.
+ rewrite (partition_big rH (mem (codom rH))) /=; last exact: image_f.
+ have nsHG: H <| G by rewrite -sub_abelian_normal.
+ apply: eq_bigr => _ /codomP[i ->]; rewrite -card_quotient ?normal_norm //.
+ rewrite -card_Iirr_abelian ?quotient_abelian //.
+ have Mlin j1 j2: exists k, 'chi_j1 * 'chi_j2 = 'chi[G]_k.
+ exact/irrP/lin_char_irr/rpredM.
+ have /fin_all_exists[rQ DrQ] (j : Iirr (G / H)) := Mlin i (mod_Iirr j).
+ have mulJi: ('chi[G]_i)^*%CF * 'chi_i = 1.
+ apply/cfun_inP=> x Gx; rewrite !cfunE -lin_charV_conj ?linG // cfun1E Gx.
+ by rewrite lin_charV ?mulVf ?lin_char_neq0 ?linG.
+ have inj_rQ: injective rQ.
+ move=> j1 j2 /(congr1 (fun k => (('chi_i)^*%CF * 'chi_k) / H)%CF).
+ by rewrite -!DrQ !mulrA mulJi !mul1r !mod_IirrE ?cfModK // => /irr_inj.
+ rewrite -(card_imset _ inj_rQ) -sum1_card; apply: eq_bigl => j.
+ rewrite -(inj_eq irr_inj) -!DrH; apply/eqP/imsetP=> [eq_ij | [k _ ->]].
+ have [k Dk] := Mlin (conjC_Iirr i) j; exists (quo_Iirr H k) => //.
+ apply/irr_inj; rewrite -DrQ quo_IirrK //.
+ by rewrite -Dk conjC_IirrE mulrCA mulrA mulJi mul1r.
+ apply/subsetP=> x Hx; have Gx := subsetP sHG x Hx.
+ rewrite cfkerEirr inE linG1 -Dk conjC_IirrE; apply/eqP.
+ transitivity ((1 : 'CF(G)) x); last by rewrite cfun1E Gx.
+ by rewrite -mulJi !cfunE -!(cfResE _ sHG Hx) eq_ij.
+ rewrite -DrQ; apply/cfun_inP=> x Hx; rewrite !cfResE // cfunE mulrC.
+ by rewrite cfker1 ?linG1 ?mul1r ?(subsetP _ x Hx) // mod_IirrE ?cfker_mod.
+have: (#|G : H| %| #|G : H|%:R * '[chi, 'chi_j])%C.
+ by rewrite dvdC_mulr ?Cint_Cnat ?Cnat_cfdot_char_irr.
+congr (_ %| _)%C; rewrite (cfdotEl _ Hchi) -(Lagrange sHG) mulnC natrM.
+rewrite invfM -mulrA mulVKf ?neq0CiG //; congr (_ * _).
+by apply: eq_bigr => x Hx; rewrite !cfResE.
+Qed.
+
+(* This is Isaacs, Theorem (3.13). *)
+Theorem faithful_degree_p_part gT (p : nat) (G P : {group gT}) i :
+ cfaithful 'chi[G]_i -> p.-nat (truncC ('chi_i 1%g)) ->
+ p.-Sylow(G) P -> abelian P ->
+ 'chi_i 1%g = (#|G : 'Z(G)|`_p)%:R.
+Proof.
+have [p_pr | pr'p] := boolP (prime p); last first.
+ have p'n n: (n > 0)%N -> p^'.-nat n.
+ by move/p'natEpi->; rewrite mem_primes (negPf pr'p).
+ rewrite irr1_degree natCK => _ /pnat_1-> => [_ _|].
+ by rewrite part_p'nat ?p'n.
+ by rewrite p'n ?irr_degree_gt0.
+move=> fful_i /p_natP[a Dchi1] sylP cPP.
+have Dchi1C: 'chi_i 1%g = (p ^ a)%:R by rewrite -Dchi1 irr1_degree natCK.
+have pa_dv_ZiG: (p ^ a %| #|G : 'Z(G)|)%N.
+ rewrite -dvdC_nat -[pa in (pa %| _)%C]Dchi1C -(cfcenter_fful_irr fful_i).
+ exact: dvd_irr1_index_center.
+have [sPG pP p'PiG] := and3P sylP.
+have ZchiP: 'Res[P] 'chi_i \in 'CF(P, P :&: 'Z(G)).
+ apply/cfun_onP=> x; rewrite inE; have [Px | /cfun0->//] := boolP (x \in P).
+ rewrite /= -(cfcenter_fful_irr fful_i) cfResE //.
+ apply: coprime_degree_support_cfcenter.
+ rewrite Dchi1 coprime_expl // prime_coprime // -p'natE //.
+ apply: pnat_dvd p'PiG; rewrite -index_cent1 indexgS // subsetI sPG.
+ by rewrite sub_cent1 (subsetP cPP).
+have /andP[_ nZG] := center_normal G; have nZP := subset_trans sPG nZG.
+apply/eqP; rewrite Dchi1C eqr_nat eqn_dvd -{1}(pfactorK a p_pr) -p_part.
+rewrite partn_dvd //= -dvdC_nat -[pa in (_ %| pa)%C]Dchi1C -card_quotient //=.
+rewrite -(card_Hall (quotient_pHall nZP sylP)) card_quotient // -indexgI.
+rewrite -(cfResE _ sPG) // index_support_dvd_degree ?subsetIl ?cPP ?orbT //.
+by rewrite cfRes_char ?irr_char.
+Qed.
+
+(* This is Isaacs, Lemma (3.14). *)
+(* Note that the assumption that G be cyclic is unnecessary, as S will be *)
+(* empty if this is not the case. *)
+Lemma sum_norm2_char_generators gT (G : {group gT}) (chi : 'CF(G)) :
+ let S := [pred s | generator G s] in
+ chi \is a character -> {in S, forall s, chi s != 0} ->
+ \sum_(s in S) `|chi s| ^+ 2 >= #|S|%:R.
+Proof.
+move=> S Nchi nz_chi_S; pose n := #|G|.
+have [g Sg | S_0] := pickP (generator G); last first.
+ by rewrite eq_card0 // big_pred0 ?lerr.
+have defG: <[g]> = G by apply/esym/eqP.
+have [cycG Gg]: cyclic G /\ g \in G by rewrite -defG cycle_cyclic cycle_id.
+pose I := {k : 'I_n | coprime n k}; pose ItoS (k : I) := (g ^+ sval k)%g.
+have imItoS: codom ItoS =i S.
+ move=> s; rewrite inE /= /ItoS /I /n /S -defG -orderE.
+ apply/codomP/idP=> [[[i cogi] ->] | Ss]; first by rewrite generator_coprime.
+ have [m ltmg Ds] := cyclePmin (cycle_generator Ss).
+ by rewrite Ds generator_coprime in Ss; apply: ex_intro (Sub (Sub m _) _) _.
+have /injectiveP injItoS: injective ItoS.
+ move=> k1 k2 /eqP; apply: contraTeq.
+ by rewrite eq_expg_mod_order orderE defG -/n !modn_small.
+have [Qn galQn [QnC gQnC [eps [pr_eps defQn] QnG]]] := group_num_field_exists G.
+have{QnG} QnGg := QnG _ G _ _ g (order_dvdG Gg).
+pose calG := 'Gal({:Qn} / 1).
+have /fin_all_exists2[ItoQ inItoQ defItoQ] (k : I):
+ exists2 nu, nu \in calG & nu eps = eps ^+ val k.
+- case: k => [[m _] /=]; rewrite coprime_sym => /Qn_aut_exists[nuC DnuC].
+ have [nuQ DnuQ] := restrict_aut_to_normal_num_field QnC nuC.
+ have hom_nu: kHom 1 {:Qn} (linfun nuQ).
+ rewrite k1HomE; apply/ahom_inP.
+ by split=> [u v | ]; rewrite !lfunE ?rmorphM ?rmorph1.
+ have [|nu cGnu Dnu] := kHom_to_gal _ (normalFieldf 1) hom_nu.
+ by rewrite !subvf.
+ exists nu => //; apply: (fmorph_inj QnC).
+ rewrite -Dnu ?memvf // lfunE DnuQ rmorphX DnuC //.
+ by rewrite prim_expr_order // fmorph_primitive_root.
+have{defQn} imItoQ: calG = ItoQ @: {:I}.
+ apply/setP=> nu; apply/idP/imsetP=> [cGnu | [k _ ->] //].
+ have pr_nu_e: n.-primitive_root (nu eps) by rewrite fmorph_primitive_root.
+ have [i Dnue] := prim_rootP pr_eps (prim_expr_order pr_nu_e).
+ rewrite Dnue prim_root_exp_coprime // coprime_sym in pr_nu_e.
+ apply: ex_intro2 (Sub i _) _ _ => //; apply/eqP.
+ rewrite /calG /= -defQn in ItoQ inItoQ defItoQ nu cGnu Dnue *.
+ by rewrite gal_adjoin_eq // defItoQ -Dnue.
+have injItoQ: {in {:I} &, injective ItoQ}.
+ move=> k1 k2 _ _ /(congr1 (fun nu : gal_of _ => nu eps))/eqP.
+ by apply: contraTeq; rewrite !defItoQ (eq_prim_root_expr pr_eps) !modn_small.
+pose pi1 := \prod_(s in S) chi s; pose pi2 := \prod_(s in S) `|chi s| ^+ 2.
+have Qpi1: pi1 \in Crat.
+ have [a Da] := QnGg _ Nchi; suffices ->: pi1 = QnC (galNorm 1 {:Qn} a).
+ have /vlineP[q ->] := mem_galNorm galQn (memvf a).
+ by rewrite rmorphZ_num rmorph1 mulr1 Crat_rat.
+ rewrite /galNorm rmorph_prod -/calG imItoQ big_imset //=.
+ rewrite /pi1 -(eq_bigl _ _ imItoS) -big_uniq // big_map big_filter /=.
+ apply: eq_bigr => k _; have [nuC DnuC] := gQnC (ItoQ k); rewrite DnuC Da.
+ have [r ->] := char_sum_irr Nchi; rewrite !sum_cfunE rmorph_sum.
+ apply: eq_bigr => i _; have /QnGg[b Db] := irr_char i.
+ have Lchi_i: 'chi_i \is a linear_char by rewrite irr_cyclic_lin.
+ have /(prim_rootP pr_eps)[m Dem]: b ^+ n = 1.
+ apply/eqP; rewrite -(fmorph_eq1 QnC) rmorphX Db -lin_charX //.
+ by rewrite -expg_mod_order orderE defG modnn lin_char1.
+ rewrite -Db -DnuC Dem rmorphX /= defItoQ exprAC -{m}Dem rmorphX {b}Db.
+ by rewrite lin_charX.
+clear I ItoS imItoS injItoS ItoQ inItoQ defItoQ imItoQ injItoQ.
+clear Qn galQn QnC gQnC eps pr_eps QnGg calG.
+have{Qpi1} Zpi1: pi1 \in Cint.
+ by rewrite Cint_rat_Aint // rpred_prod // => s _; apply: Aint_char.
+have{pi1 Zpi1} pi2_ge1: 1 <= pi2.
+ have ->: pi2 = `|pi1| ^+ 2.
+ by rewrite (big_morph Num.norm (@normrM _) (@normr1 _)) -prodrXl.
+ by rewrite Cint_normK // sqr_Cint_ge1 //; exact/prodf_neq0.
+have Sgt0: (#|S| > 0)%N by rewrite (cardD1 g) [g \in S]Sg.
+rewrite -mulr_natr -ler_pdivl_mulr ?ltr0n //.
+have n2chi_ge0 s: s \in S -> 0 <= `|chi s| ^+ 2 by rewrite exprn_ge0 ?normr_ge0.
+rewrite -(expr_ge1 Sgt0); last by rewrite divr_ge0 ?ler0n ?sumr_ge0.
+by rewrite (ler_trans pi2_ge1) // lerif_AGM.
+Qed.
+
+(* This is Burnside's vanishing theorem (Isaacs, Theorem (3.15)). *)
+Theorem nonlinear_irr_vanish gT (G : {group gT}) i :
+ 'chi[G]_i 1%g > 1 -> exists2 x, x \in G & 'chi_i x = 0.
+Proof.
+move=> chi1gt1; apply/exists_eq_inP; apply: contraFT (ltr_geF chi1gt1).
+rewrite negb_exists_in => /forall_inP nz_chi.
+rewrite -(norm_Cnat (Cnat_irr1 i)) -(@expr_le1 _ 2) ?normr_ge0 //.
+rewrite -(ler_add2r (#|G|%:R * '['chi_i])) {1}cfnorm_irr mulr1.
+rewrite (cfnormE (cfun_onG _)) mulVKf ?neq0CG // (big_setD1 1%g) //=.
+rewrite addrCA ler_add2l (cardsD1 1%g) group1 mulrS ler_add2l.
+rewrite -sumr_const !(partition_big_imset (fun s => <[s]>)) /=.
+apply: ler_sum => _ /imsetP[g /setD1P[ntg Gg] ->].
+have sgG: <[g]> \subset G by rewrite cycle_subG.
+pose S := [pred s | generator <[g]> s]; pose chi := 'Res[<[g]>] 'chi_i.
+have defS: [pred s in G^# | <[s]> == <[g]>] =i S.
+ move=> s; rewrite inE /= eq_sym andb_idl // !inE -cycle_eq1 -cycle_subG.
+ by move/eqP <-; rewrite cycle_eq1 ntg.
+have resS: {in S, 'chi_i =1 chi}.
+ by move=> s /cycle_generator=> g_s; rewrite cfResE ?cycle_subG.
+rewrite !(eq_bigl _ _ defS) sumr_const.
+rewrite (eq_bigr (fun s => `|chi s| ^+ 2)) => [|s /resS-> //].
+apply: sum_norm2_char_generators => [|s Ss].
+ by rewrite cfRes_char ?irr_char.
+by rewrite -resS // nz_chi ?(subsetP sgG) ?cycle_generator.
+Qed.
+
+End MoreIntegralChar. \ No newline at end of file
diff --git a/mathcomp/character/mxabelem.v b/mathcomp/character/mxabelem.v
new file mode 100644
index 0000000..188000d
--- /dev/null
+++ b/mathcomp/character/mxabelem.v
@@ -0,0 +1,1057 @@
+(* (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.
+Require Import fingroup morphism perm automorphism quotient gproduct action.
+Require Import finalg zmodp commutator cyclic center pgroup gseries nilpotent.
+Require Import sylow maximal abelian matrix mxalgebra mxrepresentation.
+
+(******************************************************************************)
+(* This file completes the theory developed in mxrepresentation.v with the *)
+(* construction and properties of linear representations over finite fields, *)
+(* and in particular the correspondance between internal action on a (normal) *)
+(* elementary abelian p-subgroup and a linear representation on an Fp-module. *)
+(* We provide the following next constructions for a finite field F: *)
+(* 'Zm%act == the action of {unit F} on 'M[F]_(m, n). *)
+(* rowg A == the additive group of 'rV[F]_n spanned by the row space *)
+(* of the matrix A. *)
+(* rowg_mx L == the partial inverse to rowg; for any 'Zm-stable group L *)
+(* of 'rV[F]_n we have rowg (rowg_mx L) = L. *)
+(* GLrepr F n == the natural, faithful representation of 'GL_n[F]. *)
+(* reprGLm rG == the morphism G >-> 'GL_n[F] equivalent to the *)
+(* representation r of G (with rG : mx_repr r G). *)
+(* ('MR rG)%act == the action of G on 'rV[F]_n equivalent to the *)
+(* representation r of G (with rG : mx_repr r G). *)
+(* The second set of constructions defines the interpretation of a normal *)
+(* non-trivial elementary abelian p-subgroup as an 'F_p module. We assume *)
+(* abelE : p.-abelem E and ntE : E != 1, throughout, as these are needed to *)
+(* build the isomorphism between E and a nontrivial 'rV['F_p]_n. *)
+(* 'rV(E) == the type of row vectors of the 'F_p module equivalent *)
+(* to E when E is a non-trivial p.-abelem group. *)
+(* 'M(E) == the type of matrices corresponding to E. *)
+(* 'dim E == the width of vectors/matrices in 'rV(E) / 'M(E). *)
+(* abelem_rV abelE ntE == the one-to-one injection of E onto 'rV(E). *)
+(* rVabelem abelE ntE == the one-to-one projection of 'rV(E) onto E. *)
+(* abelem_repr abelE ntE nEG == the representation of G on 'rV(E) that is *)
+(* equivalent to conjugation by G in E; here abelE, ntE are *)
+(* as above, and G \subset 'N(E). *)
+(* This file end with basic results on p-modular representations of p-groups, *)
+(* and theorems giving the structure of the representation of extraspecial *)
+(* groups; these results use somewhat more advanced group theory than the *)
+(* rest of mxrepresentation, in particular, results of sylow.v and maximal.v. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope GRing.Theory FinRing.Theory.
+Local Open Scope ring_scope.
+
+(* Special results for representations on a finite field. In this case, the *)
+(* representation is equivalent to a morphism into the general linear group *)
+(* 'GL_n[F]. It is furthermore equivalent to a group action on the finite *)
+(* additive group of the corresponding row space 'rV_n. In addition, row *)
+(* spaces of matrices in 'M[F]_n correspond to subgroups of that vector group *)
+(* (this is only surjective when F is a prime field 'F_p), with moduleules *)
+(* corresponding to subgroups stabilized by the external action. *)
+
+Section FinRingRepr.
+
+Variable (R : finComUnitRingType) (gT : finGroupType).
+Variables (G : {group gT}) (n : nat) (rG : mx_representation R G n).
+
+Definition mx_repr_act (u : 'rV_n) x := u *m rG (val (subg G x)).
+
+Lemma mx_repr_actE u x : x \in G -> mx_repr_act u x = u *m rG x.
+Proof. by move=> Gx; rewrite /mx_repr_act /= subgK. Qed.
+
+Fact mx_repr_is_action : is_action G mx_repr_act.
+Proof.
+split=> [x | u x y Gx Gy]; first exact: can_inj (repr_mxK _ (subgP _)).
+by rewrite !mx_repr_actE ?groupM // -mulmxA repr_mxM.
+Qed.
+Canonical Structure mx_repr_action := Action mx_repr_is_action.
+
+Fact mx_repr_is_groupAction : is_groupAction [set: 'rV[R]_n] mx_repr_action.
+Proof.
+move=> x Gx /=; rewrite !inE.
+apply/andP; split; first by apply/subsetP=> u; rewrite !inE.
+by apply/morphicP=> /= u v _ _; rewrite !actpermE /= /mx_repr_act mulmxDl.
+Qed.
+Canonical Structure mx_repr_groupAction := GroupAction mx_repr_is_groupAction.
+
+End FinRingRepr.
+
+Notation "''MR' rG" := (mx_repr_action rG)
+ (at level 10, rG at level 8) : action_scope.
+Notation "''MR' rG" := (mx_repr_groupAction rG) : groupAction_scope.
+
+Section FinFieldRepr.
+
+Variable F : finFieldType.
+
+(* The external group action (by scaling) of the multiplicative unit group *)
+(* of the finite field, and the correspondence between additive subgroups *)
+(* of row vectors that are stable by this action, and the matrix row spaces. *)
+Section ScaleAction.
+
+Variables m n : nat.
+
+Definition scale_act (A : 'M[F]_(m, n)) (a : {unit F}) := val a *: A.
+Lemma scale_actE A a : scale_act A a = val a *: A. Proof. by []. Qed.
+Fact scale_is_action : is_action setT scale_act.
+Proof.
+apply: is_total_action=> [A | A a b]; rewrite /scale_act ?scale1r //.
+by rewrite ?scalerA mulrC.
+Qed.
+Canonical scale_action := Action scale_is_action.
+Fact scale_is_groupAction : is_groupAction setT scale_action.
+Proof.
+move=> a _ /=; rewrite inE; apply/andP.
+split; first by apply/subsetP=> A; rewrite !inE.
+by apply/morphicP=> u A _ _ /=; rewrite !actpermE /= /scale_act scalerDr.
+Qed.
+Canonical scale_groupAction := GroupAction scale_is_groupAction.
+
+Lemma astab1_scale_act A : A != 0 -> 'C[A | scale_action] = 1%g.
+Proof.
+rewrite -mxrank_eq0=> nzA; apply/trivgP/subsetP=> a; apply: contraLR.
+rewrite !inE -val_eqE -subr_eq0 sub1set !inE => nz_a1.
+by rewrite -subr_eq0 -scaleN1r -scalerDl -mxrank_eq0 eqmx_scale.
+Qed.
+
+End ScaleAction.
+
+Local Notation "'Zm" := (scale_action _ _) (at level 8) : action_scope.
+
+Section RowGroup.
+
+Variable n : nat.
+Local Notation rVn := 'rV[F]_n.
+
+Definition rowg m (A : 'M[F]_(m, n)) : {set rVn} := [set u | u <= A]%MS.
+
+Lemma mem_rowg m A v : (v \in @rowg m A) = (v <= A)%MS.
+Proof. by rewrite inE. Qed.
+
+Fact rowg_group_set m A : group_set (@rowg m A).
+Proof.
+by apply/group_setP; split=> [|u v]; rewrite !inE ?sub0mx //; exact: addmx_sub.
+Qed.
+Canonical rowg_group m A := Group (@rowg_group_set m A).
+
+Lemma rowg_stable m (A : 'M_(m, n)) : [acts setT, on rowg A | 'Zm].
+Proof. by apply/actsP=> a _ v; rewrite !inE eqmx_scale // -unitfE (valP a). Qed.
+
+Lemma rowgS m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (rowg A \subset rowg B) = (A <= B)%MS.
+Proof.
+apply/subsetP/idP=> sAB => [| u].
+ by apply/row_subP=> i; have:= sAB (row i A); rewrite !inE row_sub => ->.
+by rewrite !inE => suA; exact: submx_trans sAB.
+Qed.
+
+Lemma eq_rowg m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A :=: B)%MS -> rowg A = rowg B.
+Proof. by move=> eqAB; apply/eqP; rewrite eqEsubset !rowgS !eqAB andbb. Qed.
+
+Lemma rowg0 m : rowg (0 : 'M_(m, n)) = 1%g.
+Proof. by apply/trivgP/subsetP=> v; rewrite !inE eqmx0 submx0. Qed.
+
+Lemma rowg1 : rowg 1%:M = setT.
+Proof. by apply/setP=> x; rewrite !inE submx1. Qed.
+
+Lemma trivg_rowg m (A : 'M_(m, n)) : (rowg A == 1%g) = (A == 0).
+Proof. by rewrite -submx0 -rowgS rowg0 (sameP trivgP eqP). Qed.
+
+Definition rowg_mx (L : {set rVn}) := <<\matrix_(i < #|L|) enum_val i>>%MS.
+
+Lemma rowgK m (A : 'M_(m, n)) : (rowg_mx (rowg A) :=: A)%MS.
+Proof.
+apply/eqmxP; rewrite !genmxE; apply/andP; split.
+ by apply/row_subP=> i; rewrite rowK; have:= enum_valP i; rewrite /= inE.
+apply/row_subP=> i; set v := row i A.
+have Av: v \in rowg A by rewrite inE row_sub.
+by rewrite (eq_row_sub (enum_rank_in Av v)) // rowK enum_rankK_in.
+Qed.
+
+Lemma rowg_mxS (L M : {set 'rV[F]_n}) :
+ L \subset M -> (rowg_mx L <= rowg_mx M)%MS.
+Proof.
+move/subsetP=> sLM; rewrite !genmxE; apply/row_subP=> i.
+rewrite rowK; move: (enum_val i) (sLM _ (enum_valP i)) => v Mv.
+by rewrite (eq_row_sub (enum_rank_in Mv v)) // rowK enum_rankK_in.
+Qed.
+
+Lemma sub_rowg_mx (L : {set rVn}) : L \subset rowg (rowg_mx L).
+Proof.
+apply/subsetP=> v Lv; rewrite inE genmxE.
+by rewrite (eq_row_sub (enum_rank_in Lv v)) // rowK enum_rankK_in.
+Qed.
+
+Lemma stable_rowg_mxK (L : {group rVn}) :
+ [acts setT, on L | 'Zm] -> rowg (rowg_mx L) = L.
+Proof.
+move=> linL; apply/eqP; rewrite eqEsubset sub_rowg_mx andbT.
+apply/subsetP=> v; rewrite inE genmxE => /submxP[u ->{v}].
+rewrite mulmx_sum_row group_prod // => i _.
+rewrite rowK; move: (enum_val i) (enum_valP i) => v Lv.
+case: (eqVneq (u 0 i) 0) => [->|]; first by rewrite scale0r group1.
+by rewrite -unitfE => aP; rewrite ((actsP linL) (FinRing.Unit _ aP)) ?inE.
+Qed.
+
+Lemma rowg_mx1 : rowg_mx 1%g = 0.
+Proof. by apply/eqP; rewrite -submx0 -(rowg0 0) rowgK sub0mx. Qed.
+
+Lemma rowg_mx_eq0 (L : {group rVn}) : (rowg_mx L == 0) = (L :==: 1%g).
+Proof.
+rewrite -trivg_rowg; apply/idP/idP=> [|/eqP->]; last by rewrite rowg_mx1 rowg0.
+by rewrite !(sameP eqP trivgP); apply: subset_trans; exact: sub_rowg_mx.
+Qed.
+
+Lemma rowgI m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ rowg (A :&: B)%MS = rowg A :&: rowg B.
+Proof. by apply/setP=> u; rewrite !inE sub_capmx. Qed.
+
+Lemma card_rowg m (A : 'M_(m, n)) : #|rowg A| = (#|F| ^ \rank A)%N.
+Proof.
+rewrite -[\rank A]mul1n -card_matrix.
+have injA: injective (mulmxr (row_base A)).
+ have /row_freeP[A' A'K] := row_base_free A.
+ by move=> ?; apply: can_inj (mulmxr A') _ => u; rewrite /= -mulmxA A'K mulmx1.
+rewrite -(card_image (injA _)); apply: eq_card => v.
+by rewrite inE -(eq_row_base A) (sameP submxP codomP).
+Qed.
+
+Lemma rowgD m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ rowg (A + B)%MS = (rowg A * rowg B)%g.
+Proof.
+apply/eqP; rewrite eq_sym eqEcard mulG_subG /= !rowgS.
+rewrite addsmxSl addsmxSr -(@leq_pmul2r #|rowg A :&: rowg B|) ?cardG_gt0 //=.
+by rewrite -mul_cardG -rowgI !card_rowg -!expnD mxrank_sum_cap.
+Qed.
+
+Lemma cprod_rowg m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ rowg A \* rowg B = rowg (A + B)%MS.
+Proof. by rewrite rowgD cprodE // (sub_abelian_cent2 (zmod_abelian setT)). Qed.
+
+Lemma dprod_rowg m1 m2 (A : 'M[F]_(m1, n)) (B : 'M[F]_(m2, n)) :
+ mxdirect (A + B) -> rowg A \x rowg B = rowg (A + B)%MS.
+Proof.
+rewrite (sameP mxdirect_addsP eqP) -trivg_rowg rowgI => /eqP tiAB.
+by rewrite -cprod_rowg dprodEcp.
+Qed.
+
+Lemma bigcprod_rowg m I r (P : pred I) (A : I -> 'M[F]_n) (B : 'M[F]_(m, n)) :
+ (\sum_(i <- r | P i) A i :=: B)%MS ->
+ \big[cprod/1%g]_(i <- r | P i) rowg (A i) = rowg B.
+Proof.
+by move/eq_rowg <-; apply/esym/big_morph=> [? ?|]; rewrite (rowg0, cprod_rowg).
+Qed.
+
+Lemma bigdprod_rowg m (I : finType) (P : pred I) A (B : 'M[F]_(m, n)) :
+ let S := (\sum_(i | P i) A i)%MS in (S :=: B)%MS -> mxdirect S ->
+ \big[dprod/1%g]_(i | P i) rowg (A i) = rowg B.
+Proof.
+move=> S defS; rewrite mxdirectE defS /= => /eqP rankB.
+apply: bigcprod_card_dprod (bigcprod_rowg defS) (eq_leq _).
+by rewrite card_rowg rankB expn_sum; apply: eq_bigr => i _; rewrite card_rowg.
+Qed.
+
+End RowGroup.
+
+Variables (gT : finGroupType) (G : {group gT}) (n' : nat).
+Local Notation n := n'.+1.
+Variable (rG : mx_representation F G n).
+
+Fact GL_mx_repr : mx_repr 'GL_n[F] GLval. Proof. by []. Qed.
+Canonical GLrepr := MxRepresentation GL_mx_repr.
+
+Lemma GLmx_faithful : mx_faithful GLrepr.
+Proof. by apply/subsetP=> A; rewrite !inE mul1mx. Qed.
+
+Definition reprGLm x : {'GL_n[F]} := insubd (1%g : {'GL_n[F]}) (rG x).
+
+Lemma val_reprGLm x : x \in G -> val (reprGLm x) = rG x.
+Proof. by move=> Gx; rewrite val_insubd (repr_mx_unitr rG). Qed.
+
+Lemma comp_reprGLm : {in G, GLval \o reprGLm =1 rG}.
+Proof. exact: val_reprGLm. Qed.
+
+Lemma reprGLmM : {in G &, {morph reprGLm : x y / x * y}}%g.
+Proof.
+by move=> x y Gx Gy; apply: val_inj; rewrite /= !val_reprGLm ?groupM ?repr_mxM.
+Qed.
+Canonical reprGL_morphism := Morphism reprGLmM.
+
+Lemma ker_reprGLm : 'ker reprGLm = rker rG.
+Proof.
+apply/setP=> x; rewrite !inE mul1mx; apply: andb_id2l => Gx.
+by rewrite -val_eqE val_reprGLm.
+Qed.
+
+Lemma astab_rowg_repr m (A : 'M_(m, n)) : 'C(rowg A | 'MR rG) = rstab rG A.
+Proof.
+apply/setP=> x; rewrite !inE /=; apply: andb_id2l => Gx.
+apply/subsetP/eqP=> cAx => [|u]; last first.
+ by rewrite !inE mx_repr_actE // => /submxP[u' ->]; rewrite -mulmxA cAx.
+apply/row_matrixP=> i; apply/eqP; move/implyP: (cAx (row i A)).
+by rewrite !inE row_sub mx_repr_actE //= row_mul.
+Qed.
+
+Lemma astabs_rowg_repr m (A : 'M_(m, n)) : 'N(rowg A | 'MR rG) = rstabs rG A.
+Proof.
+apply/setP=> x; rewrite !inE /=; apply: andb_id2l => Gx.
+apply/subsetP/idP=> nAx => [|u]; last first.
+ rewrite !inE mx_repr_actE // => Au; exact: (submx_trans (submxMr _ Au)).
+apply/row_subP=> i; move/implyP: (nAx (row i A)).
+by rewrite !inE row_sub mx_repr_actE //= row_mul.
+Qed.
+
+Lemma acts_rowg (A : 'M_n) : [acts G, on rowg A | 'MR rG] = mxmodule rG A.
+Proof. by rewrite astabs_rowg_repr. Qed.
+
+Lemma astab_setT_repr : 'C(setT | 'MR rG) = rker rG.
+Proof. by rewrite -rowg1 astab_rowg_repr. Qed.
+
+Lemma mx_repr_action_faithful :
+ [faithful G, on setT | 'MR rG] = mx_faithful rG.
+Proof.
+by rewrite /faithful astab_setT_repr (setIidPr _) // [rker _]setIdE subsetIl.
+Qed.
+
+Lemma afix_repr (H : {set gT}) :
+ H \subset G -> 'Fix_('MR rG)(H) = rowg (rfix_mx rG H).
+Proof.
+move/subsetP=> sHG; apply/setP=> /= u; rewrite !inE.
+apply/subsetP/rfix_mxP=> cHu x Hx; have:= cHu x Hx;
+ by rewrite !inE /= => /eqP; rewrite mx_repr_actE ?sHG.
+Qed.
+
+Lemma gacent_repr (H : {set gT}) :
+ H \subset G -> 'C_(| 'MR rG)(H) = rowg (rfix_mx rG H).
+Proof. by move=> sHG; rewrite gacentE // setTI afix_repr. Qed.
+
+End FinFieldRepr.
+
+Arguments Scope rowg_mx [_ _ group_scope].
+Notation "''Zm'" := (scale_action _ _ _) (at level 8) : action_scope.
+Notation "''Zm'" := (scale_groupAction _ _ _) : groupAction_scope.
+
+Section MatrixGroups.
+
+Implicit Types m n p q : nat.
+
+Lemma exponent_mx_group m n q :
+ m > 0 -> n > 0 -> q > 1 -> exponent [set: 'M['Z_q]_(m, n)] = q.
+Proof.
+move=> m_gt0 n_gt0 q_gt1; apply/eqP; rewrite eqn_dvd; apply/andP; split.
+ apply/exponentP=> x _; apply/matrixP=> i j; rewrite mulmxnE !mxE.
+ by rewrite -mulr_natr -Zp_nat_mod // modnn mulr0.
+pose cmx1 := const_mx 1%R : 'M['Z_q]_(m, n).
+apply: dvdn_trans (dvdn_exponent (in_setT cmx1)).
+have/matrixP/(_ (Ordinal m_gt0))/(_ (Ordinal n_gt0))/eqP := expg_order cmx1.
+by rewrite mulmxnE !mxE -order_dvdn order_Zp1 Zp_cast.
+Qed.
+
+Lemma rank_mx_group m n q : 'r([set: 'M['Z_q]_(m, n)]) = (m * n)%N.
+Proof.
+wlog q_gt1: q / q > 1 by case: q => [|[|q -> //]] /(_ 2)->.
+set G := setT; have cGG: abelian G := zmod_abelian _.
+have [mn0 | ] := posnP (m * n).
+ by rewrite [G](card1_trivg _) ?rank1 // cardsT card_matrix mn0.
+rewrite muln_gt0 => /andP[m_gt0 n_gt0].
+have expG: exponent G = q := exponent_mx_group m_gt0 n_gt0 q_gt1.
+apply/eqP; rewrite eqn_leq andbC -(leq_exp2l _ _ q_gt1) -{2}expG.
+have ->: (q ^ (m * n))%N = #|G| by rewrite cardsT card_matrix card_ord Zp_cast.
+rewrite max_card_abelian //= -grank_abelian //= -/G.
+pose B : {set 'M['Z_q]_(m, n)} := [set delta_mx ij.1 ij.2 | ij : 'I_m * 'I_n].
+suffices ->: G = <<B>>.
+ have ->: (m * n)%N = #|{: 'I_m * 'I_n}| by rewrite card_prod !card_ord.
+ exact: leq_trans (grank_min _) (leq_imset_card _ _).
+apply/setP=> v; rewrite inE (matrix_sum_delta v).
+rewrite group_prod // => i _; rewrite group_prod // => j _.
+rewrite -[v i j]natr_Zp scaler_nat groupX // mem_gen //.
+by apply/imsetP; exists (i, j).
+Qed.
+
+Lemma mx_group_homocyclic m n q : homocyclic [set: 'M['Z_q]_(m, n)].
+Proof.
+wlog q_gt1: q / q > 1 by case: q => [|[|q -> //]] /(_ 2)->.
+set G := setT; have cGG: abelian G := zmod_abelian _.
+rewrite -max_card_abelian //= rank_mx_group cardsT card_matrix card_ord -/G.
+rewrite {1}Zp_cast //; have [-> // | ] := posnP (m * n).
+by rewrite muln_gt0 => /andP[m_gt0 n_gt0]; rewrite exponent_mx_group.
+Qed.
+
+Lemma abelian_type_mx_group m n q :
+ q > 1 -> abelian_type [set: 'M['Z_q]_(m, n)] = nseq (m * n) q.
+Proof.
+rewrite (abelian_type_homocyclic (mx_group_homocyclic m n q)) rank_mx_group.
+have [-> // | ] := posnP (m * n); rewrite muln_gt0 => /andP[m_gt0 n_gt0] q_gt1.
+by rewrite exponent_mx_group.
+Qed.
+
+End MatrixGroups.
+
+Delimit Scope abelem_scope with Mg.
+Open Scope abelem_scope.
+
+Definition abelem_dim' (gT : finGroupType) (E : {set gT}) :=
+ (logn (pdiv #|E|) #|E|).-1.
+Arguments Scope abelem_dim' [_ group_scope].
+Notation "''dim' E" := (abelem_dim' E).+1
+ (at level 10, E at level 8, format "''dim' E") : abelem_scope.
+
+Notation "''rV' ( E )" := 'rV_('dim E)
+ (at level 8, format "''rV' ( E )") : abelem_scope.
+Notation "''M' ( E )" := 'M_('dim E)
+ (at level 8, format "''M' ( E )") : abelem_scope.
+Notation "''rV[' F ] ( E )" := 'rV[F]_('dim E)
+ (at level 8, only parsing) : abelem_scope.
+Notation "''M[' F ] ( E )" := 'M[F]_('dim E)
+ (at level 8, only parsing) : abelem_scope.
+
+Section AbelemRepr.
+
+Section FpMatrix.
+
+Variables p m n : nat.
+Local Notation Mmn := 'M['F_p]_(m, n).
+
+Lemma mx_Fp_abelem : prime p -> p.-abelem [set: Mmn].
+Proof.
+move=> p_pr; apply/abelemP=> //; rewrite zmod_abelian.
+split=> //= v _; rewrite zmodXgE -scaler_nat.
+by case/andP: (char_Fp p_pr) => _ /eqP->; rewrite scale0r.
+Qed.
+
+Lemma mx_Fp_stable (L : {group Mmn}) : [acts setT, on L | 'Zm].
+Proof.
+apply/subsetP=> a _; rewrite !inE; apply/subsetP=> A L_A.
+by rewrite inE /= /scale_act -[val _]natr_Zp scaler_nat groupX.
+Qed.
+
+End FpMatrix.
+
+Section FpRow.
+
+Variables p n : nat.
+Local Notation rVn := 'rV['F_p]_n.
+
+Lemma rowg_mxK (L : {group rVn}) : rowg (rowg_mx L) = L.
+Proof. by apply: stable_rowg_mxK; exact: mx_Fp_stable. Qed.
+
+Lemma rowg_mxSK (L : {set rVn}) (M : {group rVn}) :
+ (rowg_mx L <= rowg_mx M)%MS = (L \subset M).
+Proof.
+apply/idP/idP; last exact: rowg_mxS.
+by rewrite -rowgS rowg_mxK; apply: subset_trans; exact: sub_rowg_mx.
+Qed.
+
+Lemma mxrank_rowg (L : {group rVn}) :
+ prime p -> \rank (rowg_mx L) = logn p #|L|.
+Proof.
+by move=> p_pr; rewrite -{2}(rowg_mxK L) card_rowg card_Fp ?pfactorK.
+Qed.
+
+End FpRow.
+
+Variables (p : nat) (gT : finGroupType) (E : {group gT}).
+Hypotheses (abelE : p.-abelem E) (ntE : E :!=: 1%g).
+
+Let pE : p.-group E := abelem_pgroup abelE.
+Let p_pr : prime p. Proof. by have [] := pgroup_pdiv pE ntE. Qed.
+
+Local Notation n' := (abelem_dim' (gval E)).
+Local Notation n := n'.+1.
+Local Notation rVn := 'rV['F_p](gval E).
+
+Lemma dim_abelemE : n = logn p #|E|.
+Proof.
+rewrite /n'; have [_ _ [k ->]] := pgroup_pdiv pE ntE.
+by rewrite /pdiv primes_exp ?primes_prime // pfactorK.
+Qed.
+
+Lemma card_abelem_rV : #|rVn| = #|E|.
+Proof.
+by rewrite dim_abelemE card_matrix mul1n card_Fp // -p_part part_pnat_id.
+Qed.
+
+Lemma isog_abelem_rV : E \isog [set: rVn].
+Proof.
+by rewrite (isog_abelem_card _ abelE) cardsT card_abelem_rV mx_Fp_abelem /=.
+Qed.
+
+Local Notation ab_rV_P := (existsP isog_abelem_rV).
+Definition abelem_rV : gT -> rVn := xchoose ab_rV_P.
+
+Local Notation ErV := abelem_rV.
+
+Lemma abelem_rV_M : {in E &, {morph ErV : x y / (x * y)%g >-> x + y}}.
+Proof. by case/misomP: (xchooseP ab_rV_P) => fM _; move/morphicP: fM. Qed.
+
+Canonical abelem_rV_morphism := Morphism abelem_rV_M.
+
+Lemma abelem_rV_isom : isom E setT ErV.
+Proof. by case/misomP: (xchooseP ab_rV_P). Qed.
+
+Lemma abelem_rV_injm : 'injm ErV. Proof. by case/isomP: abelem_rV_isom. Qed.
+
+Lemma abelem_rV_inj : {in E &, injective ErV}.
+Proof. by apply/injmP; exact: abelem_rV_injm. Qed.
+
+Lemma im_abelem_rV : ErV @* E = setT. Proof. by case/isomP: abelem_rV_isom. Qed.
+
+Lemma mem_im_abelem_rV u : u \in ErV @* E.
+Proof. by rewrite im_abelem_rV inE. Qed.
+
+Lemma sub_im_abelem_rV mA : subset mA (mem (ErV @* E)).
+Proof. by rewrite unlock; apply/pred0P=> v /=; rewrite mem_im_abelem_rV. Qed.
+Hint Resolve mem_im_abelem_rV sub_im_abelem_rV.
+
+Lemma abelem_rV_1 : ErV 1 = 0%R. Proof. by rewrite morph1. Qed.
+
+Lemma abelem_rV_X x i : x \in E -> ErV (x ^+ i) = i%:R *: ErV x.
+Proof. by move=> Ex; rewrite morphX // scaler_nat. Qed.
+
+Lemma abelem_rV_V x : x \in E -> ErV x^-1 = - ErV x.
+Proof. by move=> Ex; rewrite morphV. Qed.
+
+Definition rVabelem : rVn -> gT := invm abelem_rV_injm.
+Canonical rVabelem_morphism := [morphism of rVabelem].
+Local Notation rV_E := rVabelem.
+
+Lemma rVabelem0 : rV_E 0 = 1%g. Proof. exact: morph1. Qed.
+
+Lemma rVabelemD : {morph rV_E : u v / u + v >-> (u * v)%g}.
+Proof. by move=> u v /=; rewrite -morphM. Qed.
+
+Lemma rVabelemN : {morph rV_E: u / - u >-> (u^-1)%g}.
+Proof. by move=> u /=; rewrite -morphV. Qed.
+
+Lemma rVabelemZ (m : 'F_p) : {morph rV_E : u / m *: u >-> (u ^+ m)%g}.
+Proof. by move=> u; rewrite /= -morphX -?[(u ^+ m)%g]scaler_nat ?natr_Zp. Qed.
+
+Lemma abelem_rV_K : {in E, cancel ErV rV_E}. Proof. exact: invmE. Qed.
+
+Lemma rVabelemK : cancel rV_E ErV. Proof. by move=> u; rewrite invmK. Qed.
+
+Lemma rVabelem_inj : injective rV_E. Proof. exact: can_inj rVabelemK. Qed.
+
+Lemma rVabelem_injm : 'injm rV_E. Proof. exact: injm_invm abelem_rV_injm. Qed.
+
+Lemma im_rVabelem : rV_E @* setT = E.
+Proof. by rewrite -im_abelem_rV im_invm. Qed.
+
+Lemma mem_rVabelem u : rV_E u \in E.
+Proof. by rewrite -im_rVabelem mem_morphim. Qed.
+
+Lemma sub_rVabelem L : rV_E @* L \subset E.
+Proof. by rewrite -[_ @* L]morphimIim im_invm subsetIl. Qed.
+Hint Resolve mem_rVabelem sub_rVabelem.
+
+Lemma card_rVabelem L : #|rV_E @* L| = #|L|.
+Proof. by rewrite card_injm ?rVabelem_injm. Qed.
+
+Lemma abelem_rV_mK (H : {set gT}) : H \subset E -> rV_E @* (ErV @* H) = H.
+Proof. exact: morphim_invm abelem_rV_injm H. Qed.
+
+Lemma rVabelem_mK L : ErV @* (rV_E @* L) = L.
+Proof. by rewrite morphim_invmE morphpreK. Qed.
+
+Lemma rVabelem_minj : injective (morphim (MorPhantom rV_E)).
+Proof. exact: can_inj rVabelem_mK. Qed.
+
+Lemma rVabelemS L M : (rV_E @* L \subset rV_E @* M) = (L \subset M).
+Proof. by rewrite injmSK ?rVabelem_injm. Qed.
+
+Lemma abelem_rV_S (H K : {set gT}) :
+ H \subset E -> (ErV @* H \subset ErV @* K) = (H \subset K).
+Proof. by move=> sHE; rewrite injmSK ?abelem_rV_injm. Qed.
+
+Lemma sub_rVabelem_im L (H : {set gT}) :
+ (rV_E @* L \subset H) = (L \subset ErV @* H).
+Proof. by rewrite sub_morphim_pre ?morphpre_invm. Qed.
+
+Lemma sub_abelem_rV_im (H : {set gT}) (L : {set 'rV['F_p]_n}) :
+ H \subset E -> (ErV @* H \subset L) = (H \subset rV_E @* L).
+Proof. by move=> sHE; rewrite sub_morphim_pre ?morphim_invmE. Qed.
+
+Section OneGroup.
+
+Variable G : {group gT}.
+Definition abelem_mx_fun (g : subg_of G) v := ErV ((rV_E v) ^ val g).
+Definition abelem_mx of G \subset 'N(E) :=
+ fun x => lin1_mx (abelem_mx_fun (subg G x)).
+
+Hypothesis nEG : G \subset 'N(E).
+Local Notation r := (abelem_mx nEG).
+
+Fact abelem_mx_linear_proof g : linear (abelem_mx_fun g).
+Proof.
+rewrite /abelem_mx_fun; case: g => x /= /(subsetP nEG) Nx /= m u v.
+rewrite rVabelemD rVabelemZ conjMg conjXg.
+by rewrite abelem_rV_M ?abelem_rV_X ?groupX ?memJ_norm // natr_Zp.
+Qed.
+Canonical abelem_mx_linear g := Linear (abelem_mx_linear_proof g).
+
+Let rVabelemJmx v x : x \in G -> rV_E (v *m r x) = (rV_E v) ^ x.
+Proof.
+move=> Gx; rewrite /= mul_rV_lin1 /= /abelem_mx_fun subgK //.
+by rewrite abelem_rV_K // memJ_norm // (subsetP nEG).
+Qed.
+
+Fact abelem_mx_repr : mx_repr G r.
+Proof.
+split=> [|x y Gx Gy]; apply/row_matrixP=> i; apply: rVabelem_inj.
+ by rewrite rowE -row1 rVabelemJmx // conjg1.
+by rewrite !rowE mulmxA !rVabelemJmx ?groupM // conjgM.
+Qed.
+Canonical abelem_repr := MxRepresentation abelem_mx_repr.
+Let rG := abelem_repr.
+
+Lemma rVabelemJ v x : x \in G -> rV_E (v *m rG x) = (rV_E v) ^ x.
+Proof. exact: rVabelemJmx. Qed.
+
+Lemma abelem_rV_J : {in E & G, forall x y, ErV (x ^ y) = ErV x *m rG y}.
+Proof.
+by move=> x y Ex Gy; rewrite -{1}(abelem_rV_K Ex) -rVabelemJ ?rVabelemK.
+Qed.
+
+Lemma abelem_rowgJ m (A : 'M_(m, n)) x :
+ x \in G -> rV_E @* rowg (A *m rG x) = (rV_E @* rowg A) :^ x.
+Proof.
+move=> Gx; apply: (canRL (conjsgKV _)); apply/setP=> y.
+rewrite mem_conjgV !morphim_invmE !inE memJ_norm ?(subsetP nEG) //=.
+apply: andb_id2l => Ey; rewrite abelem_rV_J //.
+by rewrite submxMfree // row_free_unit (repr_mx_unit rG).
+Qed.
+
+Lemma rV_abelem_sJ (L : {group gT}) x :
+ x \in G -> L \subset E -> ErV @* (L :^ x) = rowg (rowg_mx (ErV @* L) *m rG x).
+Proof.
+move=> Gx sLE; apply: rVabelem_minj; rewrite abelem_rowgJ //.
+by rewrite rowg_mxK !morphim_invm // -(normsP nEG x Gx) conjSg.
+Qed.
+
+Lemma rstab_abelem m (A : 'M_(m, n)) : rstab rG A = 'C_G(rV_E @* rowg A).
+Proof.
+apply/setP=> x; rewrite !inE /=; apply: andb_id2l => Gx.
+apply/eqP/centP=> cAx => [_ /morphimP[u _ Au ->]|].
+ move: Au; rewrite inE => /submxP[u' ->] {u}.
+ by apply/esym/commgP/conjg_fixP; rewrite -rVabelemJ -?mulmxA ?cAx.
+apply/row_matrixP=> i; apply: rVabelem_inj.
+by rewrite row_mul rVabelemJ // /conjg -cAx ?mulKg ?mem_morphim // inE row_sub.
+Qed.
+
+Lemma rstabs_abelem m (A : 'M_(m, n)) : rstabs rG A = 'N_G(rV_E @* rowg A).
+Proof.
+apply/setP=> x; rewrite !inE /=; apply: andb_id2l => Gx.
+by rewrite -rowgS -rVabelemS abelem_rowgJ.
+Qed.
+
+Lemma rstabs_abelemG (L : {group gT}) :
+ L \subset E -> rstabs rG (rowg_mx (ErV @* L)) = 'N_G(L).
+Proof. by move=> sLE; rewrite rstabs_abelem rowg_mxK morphim_invm. Qed.
+
+Lemma mxmodule_abelem m (U : 'M['F_p]_(m, n)) :
+ mxmodule rG U = (G \subset 'N(rV_E @* rowg U)).
+Proof. by rewrite -subsetIidl -rstabs_abelem. Qed.
+
+Lemma mxmodule_abelemG (L : {group gT}) :
+ L \subset E -> mxmodule rG (rowg_mx (ErV @* L)) = (G \subset 'N(L)).
+Proof. by move=> sLE; rewrite -subsetIidl -rstabs_abelemG. Qed.
+
+Lemma mxsimple_abelemP (U : 'M['F_p]_n) :
+ reflect (mxsimple rG U) (minnormal (rV_E @* rowg U) G).
+Proof.
+apply: (iffP mingroupP) => [[/andP[ntU modU] minU] | [modU ntU minU]].
+ split=> [||V modV sVU ntV]; first by rewrite mxmodule_abelem.
+ by apply: contraNneq ntU => ->; rewrite /= rowg0 morphim1.
+ rewrite -rowgS -rVabelemS [_ @* rowg V]minU //.
+ rewrite -subG1 sub_rVabelem_im morphim1 subG1 trivg_rowg ntV /=.
+ by rewrite -mxmodule_abelem.
+ by rewrite rVabelemS rowgS.
+split=> [|D /andP[ntD nDG sDU]].
+ rewrite -subG1 sub_rVabelem_im morphim1 subG1 trivg_rowg ntU /=.
+ by rewrite -mxmodule_abelem.
+apply/eqP; rewrite eqEsubset sDU sub_rVabelem_im /= -rowg_mxSK rowgK.
+have sDE: D \subset E := subset_trans sDU (sub_rVabelem _).
+rewrite minU ?mxmodule_abelemG //.
+ by rewrite -rowgS rowg_mxK sub_abelem_rV_im.
+by rewrite rowg_mx_eq0 (morphim_injm_eq1 abelem_rV_injm).
+Qed.
+
+Lemma mxsimple_abelemGP (L : {group gT}) :
+ L \subset E -> reflect (mxsimple rG (rowg_mx (ErV @* L))) (minnormal L G).
+Proof.
+move/abelem_rV_mK=> {2}<-; rewrite -{2}[_ @* L]rowg_mxK.
+exact: mxsimple_abelemP.
+Qed.
+
+Lemma abelem_mx_irrP : reflect (mx_irreducible rG) (minnormal E G).
+Proof.
+by rewrite -[E in minnormal E G]im_rVabelem -rowg1; exact: mxsimple_abelemP.
+Qed.
+
+Lemma rfix_abelem (H : {set gT}) :
+ H \subset G -> (rfix_mx rG H :=: rowg_mx (ErV @* 'C_E(H)%g))%MS.
+Proof.
+move/subsetP=> sHG; apply/eqmxP/andP; split.
+ rewrite -rowgS rowg_mxK -sub_rVabelem_im // subsetI sub_rVabelem /=.
+ apply/centsP=> y /morphimP[v _]; rewrite inE => cGv ->{y} x Gx.
+ by apply/commgP/conjg_fixP; rewrite /= -rVabelemJ ?sHG ?(rfix_mxP H _).
+rewrite genmxE; apply/rfix_mxP=> x Hx; apply/row_matrixP=> i.
+rewrite row_mul rowK; case/morphimP: (enum_valP i) => z Ez /setIP[_ cHz] ->.
+by rewrite -abelem_rV_J ?sHG // conjgE (centP cHz) ?mulKg.
+Qed.
+
+Lemma rker_abelem : rker rG = 'C_G(E).
+Proof. by rewrite /rker rstab_abelem rowg1 im_rVabelem. Qed.
+
+Lemma abelem_mx_faithful : 'C_G(E) = 1%g -> mx_faithful rG.
+Proof. by rewrite /mx_faithful rker_abelem => ->. Qed.
+
+End OneGroup.
+
+Section SubGroup.
+
+Variables G H : {group gT}.
+Hypotheses (nEG : G \subset 'N(E)) (sHG : H \subset G).
+Let nEH := subset_trans sHG nEG.
+Local Notation rG := (abelem_repr nEG).
+Local Notation rHG := (subg_repr rG sHG).
+Local Notation rH := (abelem_repr nEH).
+
+Lemma eq_abelem_subg_repr : {in H, rHG =1 rH}.
+Proof.
+move=> x Hx; apply/row_matrixP=> i; rewrite !rowE !mul_rV_lin1 /=.
+by rewrite /abelem_mx_fun !subgK ?(subsetP sHG).
+Qed.
+
+Lemma rsim_abelem_subg : mx_rsim rHG rH.
+Proof.
+exists 1%:M => // [|x Hx]; first by rewrite row_free_unit unitmx1.
+by rewrite mul1mx mulmx1 eq_abelem_subg_repr.
+Qed.
+
+Lemma mxmodule_abelem_subg m (U : 'M_(m, n)) : mxmodule rHG U = mxmodule rH U.
+Proof.
+apply: eq_subset_r => x; rewrite !inE; apply: andb_id2l => Hx.
+by rewrite eq_abelem_subg_repr.
+Qed.
+
+Lemma mxsimple_abelem_subg U : mxsimple rHG U <-> mxsimple rH U.
+Proof.
+have eq_modH := mxmodule_abelem_subg; rewrite /mxsimple eq_modH.
+by split=> [] [-> -> minU]; split=> // V; have:= minU V; rewrite eq_modH.
+Qed.
+
+End SubGroup.
+
+End AbelemRepr.
+
+Section ModularRepresentation.
+
+Variables (F : fieldType) (p : nat) (gT : finGroupType).
+Hypothesis charFp : p \in [char F].
+Implicit Types G H : {group gT}.
+
+(* This is Gorenstein, Lemma 2.6.3. *)
+Lemma rfix_pgroup_char G H n (rG : mx_representation F G n) :
+ n > 0 -> p.-group H -> H \subset G -> rfix_mx rG H != 0.
+Proof.
+move=> n_gt0 pH sHG; rewrite -(rfix_subg rG sHG).
+move: {2}_.+1 (ltnSn (n + #|H|)) {rG G sHG}(subg_repr _ _) => m.
+elim: m gT H pH => // m IHm gT' G pG in n n_gt0 *; rewrite ltnS => le_nG_m rG.
+apply/eqP=> Gregular; have irrG: mx_irreducible rG.
+ apply/mx_irrP; split=> // U modU; rewrite -mxrank_eq0 -lt0n => Unz.
+ rewrite /row_full eqn_leq rank_leq_col leqNgt; apply/negP=> ltUn.
+ have: rfix_mx (submod_repr modU) G != 0.
+ by apply: IHm => //; apply: leq_trans le_nG_m; rewrite ltn_add2r.
+ by rewrite -mxrank_eq0 (rfix_submod modU) // Gregular capmx0 linear0 mxrank0.
+have{m le_nG_m IHm} faithfulG: mx_faithful rG.
+ apply/trivgP/eqP/idPn; set C := _ rG => ntC.
+ suffices: rfix_mx (kquo_repr rG) (G / _)%g != 0.
+ by rewrite -mxrank_eq0 rfix_quo // Gregular mxrank0.
+ apply: (IHm _ _ (morphim_pgroup _ _)) => //.
+ by apply: leq_trans le_nG_m; rewrite ltn_add2l ltn_quotient // rstab_sub.
+have{Gregular} ntG: G :!=: 1%g.
+ apply: contraL n_gt0; move/eqP=> G1; rewrite -leqNgt -(mxrank1 F n).
+ rewrite -(mxrank0 F n n) -Gregular mxrankS //; apply/rfix_mxP=> x.
+ by rewrite {1}G1 mul1mx => /set1P->; rewrite repr_mx1.
+have p_pr: prime p by case/andP: charFp.
+have{ntG pG} [z]: {z | z \in 'Z(G) & #[z] = p}; last case/setIP=> Gz cGz ozp.
+ apply: Cauchy => //; apply: contraR ntG; rewrite -p'natE // => p'Z.
+ have pZ: p.-group 'Z(G) by rewrite (pgroupS (center_sub G)).
+ by rewrite (trivg_center_pgroup pG (card1_trivg (pnat_1 pZ p'Z))).
+have{cGz} cGz1: centgmx rG (rG z - 1%:M).
+ apply/centgmxP=> x Gx; rewrite mulmxBl mulmxBr mulmx1 mul1mx.
+ by rewrite -!repr_mxM // (centP cGz).
+have{irrG faithfulG cGz1} Urz1: rG z - 1%:M \in unitmx.
+ apply: (mx_Schur irrG) cGz1 _; rewrite subr_eq0.
+ move/implyP: (subsetP faithfulG z).
+ by rewrite !inE Gz mul1mx -order_eq1 ozp -implybNN neq_ltn orbC prime_gt1.
+do [case: n n_gt0 => // n' _; set n := n'.+1] in rG Urz1 *.
+have charMp: p \in [char 'M[F]_n].
+ exact: (rmorph_char (scalar_mx_rmorphism _ _)).
+have{Urz1}: Frobenius_aut charMp (rG z - 1) \in GRing.unit by rewrite unitrX.
+rewrite (Frobenius_autB_comm _ (commr1 _)) Frobenius_aut1.
+by rewrite -[_ (rG z)](repr_mxX rG) // -ozp expg_order repr_mx1 subrr unitr0.
+Qed.
+
+Variables (G : {group gT}) (n : nat) (rG : mx_representation F G n).
+
+Lemma pcore_sub_rstab_mxsimple M : mxsimple rG M -> 'O_p(G) \subset rstab rG M.
+Proof.
+case=> modM nzM simM; have sGpG := pcore_sub p G.
+rewrite rfix_mx_rstabC //; set U := rfix_mx _ _.
+have:= simM (M :&: U)%MS; rewrite sub_capmx submx_refl.
+apply; rewrite ?capmxSl //.
+ by rewrite capmx_module // normal_rfix_mx_module ?pcore_normal.
+rewrite -(in_submodK (capmxSl _ _)) val_submod_eq0 -submx0.
+rewrite -(rfix_submod modM) // submx0 rfix_pgroup_char ?pcore_pgroup //.
+by rewrite lt0n mxrank_eq0.
+Qed.
+
+Lemma pcore_sub_rker_mx_irr : mx_irreducible rG -> 'O_p(G) \subset rker rG.
+Proof. exact: pcore_sub_rstab_mxsimple. Qed.
+
+(* This is Gorenstein, Lemma 3.1.3. *)
+Lemma pcore_faithful_mx_irr :
+ mx_irreducible rG -> mx_faithful rG -> 'O_p(G) = 1%g.
+Proof.
+move=> irrG ffulG; apply/trivgP; apply: subset_trans ffulG.
+exact: pcore_sub_rstab_mxsimple.
+Qed.
+
+End ModularRepresentation.
+
+Section Extraspecial.
+
+Variables (F : fieldType) (gT : finGroupType) (S : {group gT}) (p n : nat).
+Hypotheses (pS : p.-group S) (esS : extraspecial S).
+Hypothesis oSpn : #|S| = (p ^ n.*2.+1)%N.
+Hypotheses (splitF : group_splitting_field F S) (F'S : [char F]^'.-group S).
+
+Let p_pr := extraspecial_prime pS esS.
+Let p_gt0 := prime_gt0 p_pr.
+Let p_gt1 := prime_gt1 p_pr.
+Let oZp := card_center_extraspecial pS esS.
+
+Let modIp' (i : 'I_p.-1) : (i.+1 %% p = i.+1)%N.
+Proof. by case: i => i; rewrite /= -ltnS prednK //; exact: modn_small. Qed.
+
+(* This is Aschbacher (34.9), parts (1)-(4). *)
+Theorem extraspecial_repr_structure (sS : irrType F S) :
+ [/\ #|linear_irr sS| = (p ^ n.*2)%N,
+ exists iphi : 'I_p.-1 -> sS, let phi i := irr_repr (iphi i) in
+ [/\ injective iphi,
+ codom iphi =i ~: linear_irr sS,
+ forall i, mx_faithful (phi i),
+ forall z, z \in 'Z(S)^# ->
+ exists2 w, primitive_root_of_unity p w
+ & forall i, phi i z = (w ^+ i.+1)%:M
+ & forall i, irr_degree (iphi i) = (p ^ n)%N]
+ & #|sS| = (p ^ n.*2 + p.-1)%N].
+Proof.
+have [[defPhiS defS'] prZ] := esS; set linS := linear_irr sS.
+have nb_lin: #|linS| = (p ^ n.*2)%N.
+ rewrite card_linear_irr // -divgS ?der_sub //=.
+ by rewrite oSpn defS' oZp expnS mulKn.
+have nb_irr: #|sS| = (p ^ n.*2 + p.-1)%N.
+ pose Zcl := classes S ::&: 'Z(S).
+ have cardZcl: #|Zcl| = p.
+ transitivity #|[set [set z] | z in 'Z(S)]|; last first.
+ by rewrite card_imset //; exact: set1_inj.
+ apply: eq_card => zS; apply/setIdP/imsetP=> [[] | [z]].
+ case/imsetP=> z Sz ->{zS} szSZ.
+ have Zz: z \in 'Z(S) by rewrite (subsetP szSZ) ?class_refl.
+ exists z => //; rewrite inE Sz in Zz.
+ apply/eqP; rewrite eq_sym eqEcard sub1set class_refl cards1.
+ by rewrite -index_cent1 (setIidPl _) ?indexgg // sub_cent1.
+ case/setIP=> Sz cSz ->{zS}; rewrite sub1set inE Sz; split=> //.
+ apply/imsetP; exists z; rewrite //.
+ apply/eqP; rewrite eqEcard sub1set class_refl cards1.
+ by rewrite -index_cent1 (setIidPl _) ?indexgg // sub_cent1.
+ move/eqP: (class_formula S); rewrite (bigID (mem Zcl)) /=.
+ rewrite (eq_bigr (fun _ => 1%N)) => [|zS]; last first.
+ case/andP=> _ /setIdP[/imsetP[z Sz ->{zS}] /subsetIP[_ cSzS]].
+ rewrite (setIidPl _) ?indexgg // sub_cent1 (subsetP cSzS) //.
+ exact: mem_repr (class_refl S z).
+ rewrite sum1dep_card setIdE (setIidPr _) 1?cardsE ?cardZcl; last first.
+ by apply/subsetP=> zS; rewrite 2!inE => /andP[].
+ have pn_gt0: p ^ n.*2 > 0 by rewrite expn_gt0 p_gt0.
+ rewrite card_irr // oSpn expnS -(prednK pn_gt0) mulnS eqn_add2l.
+ rewrite (eq_bigr (fun _ => p)) => [|xS]; last first.
+ case/andP=> SxS; rewrite inE SxS; case/imsetP: SxS => x Sx ->{xS} notZxS.
+ have [y Sy ->] := repr_class S x; apply: p_maximal_index => //.
+ apply: cent1_extraspecial_maximal => //; first exact: groupJ.
+ apply: contra notZxS => Zxy; rewrite -{1}(lcoset_id Sy) class_lcoset.
+ rewrite ((_ ^: _ =P [set x ^ y])%g _) ?sub1set // eq_sym eqEcard.
+ rewrite sub1set class_refl cards1 -index_cent1 (setIidPl _) ?indexgg //.
+ by rewrite sub_cent1; apply: subsetP Zxy; exact: subsetIr.
+ rewrite sum_nat_dep_const mulnC eqn_pmul2l //; move/eqP <-.
+ rewrite addSnnS prednK // -cardZcl -[card _](cardsID Zcl) /= addnC.
+ by congr (_ + _)%N; apply: eq_card => t; rewrite !inE andbC // andbAC andbb.
+have fful_nlin i: i \in ~: linS -> mx_faithful (irr_repr i).
+ rewrite !inE => nlin_phi.
+ apply/trivgP; apply: (TI_center_nil (pgroup_nil pS) (rker_normal _)).
+ rewrite setIC; apply: (prime_TIg prZ); rewrite /= -defS' der1_sub_rker //.
+ exact: socle_irr.
+have [i0 nlin_i0]: exists i0, i0 \in ~: linS.
+ by apply/card_gt0P; rewrite cardsCs setCK nb_irr nb_lin addKn -subn1 subn_gt0.
+have [z defZ]: exists z, 'Z(S) = <[z]> by apply/cyclicP; rewrite prime_cyclic.
+have Zz: z \in 'Z(S) by [rewrite defZ cycle_id]; have [Sz cSz] := setIP Zz.
+have ozp: #[z] = p by rewrite -oZp defZ.
+have ntz: z != 1%g by rewrite -order_gt1 ozp.
+pose phi := irr_repr i0; have irr_phi: mx_irreducible phi := socle_irr i0.
+pose w := irr_mode i0 z.
+have phi_z: phi z = w%:M by rewrite /phi irr_center_scalar.
+have phi_ze e: phi (z ^+ e)%g = (w ^+ e)%:M.
+ by rewrite /phi irr_center_scalar ?groupX ?irr_modeX.
+have wp1: w ^+ p = 1 by rewrite -irr_modeX // -ozp expg_order irr_mode1.
+have injw: {in 'Z(S) &, injective (irr_mode i0)}.
+ move=> x y Zx Zy /= eq_xy; have [[Sx _] [Sy _]] := (setIP Zx, setIP Zy).
+ apply: mx_faithful_inj (fful_nlin _ nlin_i0) _ _ Sx Sy _.
+ by rewrite !{1}irr_center_scalar ?eq_xy; first by split.
+have prim_w e: 0 < e < p -> p.-primitive_root (w ^+ e).
+ case/andP=> e_gt0 lt_e_p; apply/andP; split=> //.
+ apply/eqfunP=> -[d ltdp] /=; rewrite unity_rootE -exprM.
+ rewrite -(irr_mode1 i0) -irr_modeX // (inj_in_eq injw) ?groupX ?group1 //.
+ rewrite -order_dvdn ozp Euclid_dvdM // gtnNdvd //=.
+ move: ltdp; rewrite leq_eqVlt.
+ by case: eqP => [-> _ | _ ltd1p]; rewrite (dvdnn, gtnNdvd).
+have /cyclicP[a defAutZ]: cyclic (Aut 'Z(S)) by rewrite Aut_prime_cyclic ?ozp.
+have phi_unitP (i : 'I_p.-1): (i.+1%:R : 'Z_#[z]) \in GRing.unit.
+ by rewrite unitZpE ?order_gt1 // ozp prime_coprime // -lt0n !modIp'.
+pose ephi i := invm (injm_Zpm a) (Zp_unitm (FinRing.Unit _ (phi_unitP i))).
+pose j : 'Z_#[z] := val (invm (injm_Zp_unitm z) a).
+have co_j_p: coprime j p.
+ rewrite coprime_sym /j; case: (invm _ a) => /=.
+ by rewrite ozp /GRing.unit /= Zp_cast.
+have [alpha Aut_alpha alphaZ] := center_aut_extraspecial pS esS co_j_p.
+have alpha_i_z i: ((alpha ^+ ephi i) z = z ^+ i.+1)%g.
+ transitivity ((a ^+ ephi i) z)%g.
+ elim: (ephi i : nat) => // e IHe; rewrite !expgS !permM alphaZ //.
+ have Aut_a: a \in Aut 'Z(S) by rewrite defAutZ cycle_id.
+ rewrite -{2}[a](invmK (injm_Zp_unitm z)); last by rewrite im_Zp_unitm -defZ.
+ rewrite /= autE ?cycle_id // -/j /= /cyclem.
+ rewrite -(autmE (groupX _ Aut_a)) -(autmE (groupX _ Aut_alpha)).
+ by rewrite !morphX //= !autmE IHe.
+ rewrite [(a ^+ _)%g](invmK (injm_Zpm a)) /=; last first.
+ by rewrite im_Zpm -defAutZ defZ Aut_aut.
+ by rewrite autE ?cycle_id //= val_Zp_nat ozp ?modIp'.
+have rphiP i: S :==: autm (groupX (ephi i) Aut_alpha) @* S by rewrite im_autm.
+pose rphi i := morphim_repr (eqg_repr phi (rphiP i)) (subxx S).
+have rphi_irr i: mx_irreducible (rphi i).
+ by apply/morphim_mx_irr; exact/eqg_mx_irr.
+have rphi_fful i: mx_faithful (rphi i).
+ rewrite /mx_faithful rker_morphim rker_eqg.
+ by rewrite (trivgP (fful_nlin _ nlin_i0)) morphpreIdom; exact: injm_autm.
+have rphi_z i: rphi i z = (w ^+ i.+1)%:M.
+ by rewrite /rphi [phi]lock /= /morphim_mx autmE alpha_i_z -lock phi_ze.
+pose iphi i := irr_comp sS (rphi i); pose phi_ i := irr_repr (iphi i).
+have{phi_ze} phi_ze i e: phi_ i (z ^+ e)%g = (w ^+ (e * i.+1)%N)%:M.
+ rewrite /phi_ !{1}irr_center_scalar ?groupX ?irr_modeX //.
+ suffices ->: irr_mode (iphi i) z = w ^+ i.+1 by rewrite mulnC exprM.
+ have:= mx_rsim_sym (rsim_irr_comp sS F'S (rphi_irr i)).
+ case/mx_rsim_def=> B [B' _ homB]; rewrite /irr_mode homB // rphi_z.
+ rewrite -{1}scalemx1 -scalemxAr -scalemxAl -{1}(repr_mx1 (rphi i)).
+ by rewrite -homB // repr_mx1 scalemx1 mxE.
+have inj_iphi: injective iphi.
+ move=> i1 i2 eqi12; apply/eqP.
+ move/eqP: (congr1 (fun i => irr_mode i (z ^+ 1)) eqi12).
+ rewrite /irr_mode !{1}[irr_repr _ _]phi_ze !{1}mxE !mul1n.
+ by rewrite (eq_prim_root_expr (prim_w 1%N p_gt1)) !modIp'.
+have deg_phi i: irr_degree (iphi i) = irr_degree i0.
+ by case: (rsim_irr_comp sS F'S (rphi_irr i)).
+have im_iphi: codom iphi =i ~: linS.
+ apply/subset_cardP; last apply/subsetP=> _ /codomP[i ->].
+ by rewrite card_image // card_ord cardsCs setCK nb_irr nb_lin addKn.
+ by rewrite !inE /= (deg_phi i) in nlin_i0 *.
+split=> //; exists iphi; rewrite -/phi_.
+split=> // [i | ze | i].
+- have sim_i := rsim_irr_comp sS F'S (rphi_irr i).
+ by rewrite -(mx_rsim_faithful sim_i) rphi_fful.
+- rewrite {1}defZ 2!inE andbC; case/andP.
+ case/cyclePmin=> e; rewrite ozp => lt_e_p ->{ze}.
+ case: (posnP e) => [-> | e_gt0 _]; first by rewrite eqxx.
+ exists (w ^+ e) => [|i]; first by rewrite prim_w ?e_gt0.
+ by rewrite phi_ze exprM.
+rewrite deg_phi {i}; set d := irr_degree i0.
+apply/eqP; move/eqP: (sum_irr_degree sS F'S splitF).
+rewrite (bigID (mem linS)) /= -/irr_degree.
+rewrite (eq_bigr (fun _ => 1%N)) => [|i]; last by rewrite !inE; move/eqP->.
+rewrite sum1_card nb_lin.
+rewrite (eq_bigl (mem (codom iphi))) // => [|i]; last first.
+ by rewrite -in_setC -im_iphi.
+rewrite (eq_bigr (fun _ => d ^ 2))%N => [|_ /codomP[i ->]]; last first.
+ by rewrite deg_phi.
+rewrite sum_nat_const card_image // card_ord oSpn (expnS p) -{3}[p]prednK //.
+rewrite mulSn eqn_add2l eqn_pmul2l; last by rewrite -ltnS prednK.
+by rewrite -muln2 expnM eqn_sqr.
+Qed.
+
+(* This is the corolloray of the above that is actually used in the proof of *)
+(* B & G, Theorem 2.5. It encapsulates the dependency on a socle of the *)
+(* regular representation. *)
+
+Variables (m : nat) (rS : mx_representation F S m) (U : 'M[F]_m).
+Hypotheses (simU : mxsimple rS U) (ffulU : rstab rS U == 1%g).
+Let sZS := center_sub S.
+Let rZ := subg_repr rS sZS.
+
+Lemma faithful_repr_extraspecial :
+ \rank U = (p ^ n)%N /\
+ (forall V, mxsimple rS V -> mx_iso rZ U V -> mx_iso rS U V).
+Proof.
+suffices IH V: mxsimple rS V -> mx_iso rZ U V ->
+ [&& \rank U == (p ^ n)%N & mxsimple_iso rS U V].
+- split=> [|/= V simV isoUV].
+ by case/andP: (IH U simU (mx_iso_refl _ _)) => /eqP.
+ by case/andP: (IH V simV isoUV) => _ /(mxsimple_isoP simU).
+move=> simV isoUV; wlog sS: / irrType F S by exact: socle_exists.
+have [[_ defS'] prZ] := esS.
+have{prZ} ntZ: 'Z(S) :!=: 1%g by case: eqP prZ => // ->; rewrite cards1.
+have [_ [iphi]] := extraspecial_repr_structure sS.
+set phi := fun i => _ => [] [inj_phi im_phi _ phiZ dim_phi] _.
+have [modU nzU _]:= simU; pose rU := submod_repr modU.
+have nlinU: \rank U != 1%N.
+ apply/eqP=> /(rker_linear rU); apply/negP; rewrite /rker rstab_submod.
+ by rewrite (eqmx_rstab _ (val_submod1 _)) (eqP ffulU) defS' subG1.
+have irrU: mx_irreducible rU by exact/submod_mx_irr.
+have rsimU := rsim_irr_comp sS F'S irrU.
+set iU := irr_comp sS rU in rsimU; have [_ degU _ _]:= rsimU.
+have phiUP: iU \in codom iphi by rewrite im_phi !inE -degU.
+rewrite degU -(f_iinv phiUP) dim_phi eqxx /=; apply/(mxsimple_isoP simU).
+have [modV _ _]:= simV; pose rV := submod_repr modV.
+have irrV: mx_irreducible rV by exact/submod_mx_irr.
+have rsimV := rsim_irr_comp sS F'S irrV.
+set iV := irr_comp sS rV in rsimV; have [_ degV _ _]:= rsimV.
+have phiVP: iV \in codom iphi by rewrite im_phi !inE -degV -(mxrank_iso isoUV).
+pose jU := iinv phiUP; pose jV := iinv phiVP.
+have [z Zz ntz]:= trivgPn _ ntZ.
+have [|w prim_w phi_z] := phiZ z; first by rewrite 2!inE ntz.
+suffices eqjUV: jU == jV.
+ apply/(mx_rsim_iso modU modV); apply: mx_rsim_trans rsimU _.
+ by rewrite -(f_iinv phiUP) -/jU (eqP eqjUV) f_iinv; exact: mx_rsim_sym.
+have rsimUV: mx_rsim (subg_repr (phi jU) sZS) (subg_repr (phi jV) sZS).
+ have [bU _ bUfree bUhom] := mx_rsim_sym rsimU.
+ have [bV _ bVfree bVhom] := rsimV.
+ have modUZ := mxmodule_subg sZS modU; have modVZ := mxmodule_subg sZS modV.
+ case/(mx_rsim_iso modUZ modVZ): isoUV => [bZ degZ bZfree bZhom].
+ rewrite /phi !f_iinv; exists (bU *m bZ *m bV)=> [||x Zx].
+ - by rewrite -degU degZ degV.
+ - by rewrite /row_free !mxrankMfree.
+ have Sx := subsetP sZS x Zx.
+ by rewrite 2!mulmxA bUhom // -(mulmxA _ _ bZ) bZhom // -4!mulmxA bVhom.
+have{rsimUV} [B [B' _ homB]] := mx_rsim_def rsimUV.
+have:= eqxx (irr_mode (iphi jU) z); rewrite /irr_mode; set i0 := Ordinal _.
+rewrite {2}[_ z]homB // ![_ z]phi_z mxE mulr1n -scalemx1 -scalemxAr -scalemxAl.
+rewrite -(repr_mx1 (subg_repr (phi jV) sZS)) -{B B'}homB // repr_mx1 scalemx1.
+by rewrite mxE (eq_prim_root_expr prim_w) !modIp'.
+Qed.
+
+End Extraspecial.
diff --git a/mathcomp/character/mxrepresentation.v b/mathcomp/character/mxrepresentation.v
new file mode 100644
index 0000000..e8bbed3
--- /dev/null
+++ b/mathcomp/character/mxrepresentation.v
@@ -0,0 +1,5853 @@
+(* (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 polydiv finset.
+Require Import fingroup morphism perm automorphism quotient finalg action zmodp.
+Require Import commutator cyclic center pgroup matrix mxalgebra mxpoly.
+
+(******************************************************************************)
+(* This file provides linkage between classic Group Theory and commutative *)
+(* algebra -- representation theory. Since general abstract linear algebra is *)
+(* still being sorted out, we develop the required theory here on the *)
+(* assumption that all vector spaces are matrix spaces, indeed that most are *)
+(* row matrix spaces; our representation theory is specialized to the latter *)
+(* case. We provide many definitions and results of representation theory: *)
+(* enveloping algebras, reducible, irreducible and absolutely irreducible *)
+(* representations, representation centralisers, submodules and kernels, *)
+(* simple and semisimple modules, the Schur lemmas, Maschke's theorem, *)
+(* components, socles, homomorphisms and isomorphisms, the Jacobson density *)
+(* theorem, similar representations, the Jordan-Holder theorem, Clifford's *)
+(* theorem and Wedderburn components, regular representations and the *)
+(* Wedderburn structure theorem for semisimple group rings, and the *)
+(* construction of a splitting field of an irreducible representation, and of *)
+(* reduced, tensored, and factored representations. *)
+(* mx_representation F G n == the Structure type for representations of G *)
+(* with n x n matrices with coefficients in F. Note that *)
+(* rG : mx_representation F G n coerces to a function from *)
+(* the element type of G to 'M_n, and conversely all such *)
+(* functions have a Canonical mx_representation. *)
+(* mx_repr G r <-> r : gT -> 'M_n defines a (matrix) group representation *)
+(* on G : {set gT} (Prop predicate). *)
+(* enveloping_algebra_mx rG == a #|G| x (n ^ 2) matrix whose rows are the *)
+(* mxvec encodings of the image of G under rG, and whose *)
+(* row space therefore encodes the enveloping algebra of *)
+(* the representation of G. *)
+(* rker rG == the kernel of the representation of r on G, i.e., the *)
+(* subgroup of elements of G mapped to the identity by rG. *)
+(* mx_faithful rG == the representation rG of G is faithful (its kernel is *)
+(* trivial). *)
+(* rfix_mx rG H == an n x n matrix whose row space is the set of vectors *)
+(* fixed (centralised) by the representation of H by rG. *)
+(* rcent rG A == the subgroup of G whose representation via rG commutes *)
+(* with the square matrix A. *)
+(* rcenter rG == the subgroup of G whose representation via rG consists of *)
+(* scalar matrices. *)
+(* mxcentg rG f <=> f commutes with every matrix in the representation of G *)
+(* (i.e., f is a total rG-homomorphism). *)
+(* rstab rG U == the subgroup of G whose representation via r fixes all *)
+(* vectors in U, pointwise. *)
+(* rstabs rG U == the subgroup of G whose representation via r fixes the row *)
+(* space of U globally. *)
+(* mxmodule rG U <=> the row-space of the matrix U is a module (globally *)
+(* invariant) under the representation rG of G. *)
+(* max_submod rG U V <-> U < V is not a proper is a proper subset of any *)
+(* proper rG-submodule of V (if both U and V are modules, *)
+(* then U is a maximal proper submodule of V). *)
+(* mx_subseries rG Us <=> Us : seq 'M_n is a list of rG-modules *)
+(* mx_composition_series rG Us <-> Us is an increasing composition series *)
+(* for an rG-module (namely, last 0 Us). *)
+(* mxsimple rG M <-> M is a simple rG-module (i.e., minimal and nontrivial) *)
+(* This is a Prop predicate on square matrices. *)
+(* mxnonsimple rG U <-> U is constructively not a submodule, that is, U *)
+(* contains a proper nontrivial submodule. *)
+(* mxnonsimple_sat rG U == U is not a simple as an rG-module. *)
+(* This is a bool predicate, which requires a decField *)
+(* structure on the scalar field. *)
+(* mxsemisimple rG W <-> W is constructively a direct sum of simple modules. *)
+(* mxsplits rG V U <-> V splits over U in rG, i.e., U has an rG-invariant *)
+(* complement in V. *)
+(* mx_completely_reducible rG V <-> V splits over all its submodules; note *)
+(* that this is only classically equivalent to stating that *)
+(* V is semisimple. *)
+(* mx_irreducible rG <-> the representation rG is irreducible, i.e., the full *)
+(* module 1%:M of rG is simple. *)
+(* mx_absolutely_irreducible rG == the representation rG of G is absolutely *)
+(* irreducible: its enveloping algebra is the full matrix *)
+(* ring. This is only classically equivalent to the more *)
+(* standard ``rG does not reduce in any field extension''. *)
+(* group_splitting_field F G <-> F is a splitting field for the group G: *)
+(* every irreducible representation of G is absolutely *)
+(* irreducible. Any field can be embedded classically into a *)
+(* splitting field. *)
+(* group_closure_field F gT <-> F is a splitting field for every group *)
+(* G : {group gT}, and indeed for any section of such a *)
+(* group. This is a convenient constructive substitute for *)
+(* algebraic closures, that can be constructed classically. *)
+(* dom_hom_mx rG f == a square matrix encoding the set of vectors for which *)
+(* multiplication by the n x n matrix f commutes with the *)
+(* representation of G, i.e., the largest domain on which *)
+(* f is an rG homomorphism. *)
+(* mx_iso rG U V <-> U and V are (constructively) rG-isomorphic; this is *)
+(* a Prop predicate. *)
+(* mx_simple_iso rG U V == U and V are rG-isomorphic if one of them is *)
+(* simple; this is a bool predicate. *)
+(* cyclic_mx rG u == the cyclic rG-module generated by the row vector u *)
+(* annihilator_mx rG u == the annihilator of the row vector u in the *)
+(* enveloping algebra the representation rG. *)
+(* row_hom_mx rG u == the image of u by the set of all rG-homomorphisms on *)
+(* its cyclic module, or, equivalently, the null-space of the *)
+(* annihilator of u. *)
+(* component_mx rG M == when M is a simple rG-module, the component of M in *)
+(* the representation rG, i.e. the module generated by all *)
+(* the (simple) modules rG-isomorphic to M. *)
+(* socleType rG == a Structure that represents the type of all components *)
+(* of rG (more precisely, it coerces to such a type via *)
+(* socle_sort). For sG : socleType, values of type sG (to be *)
+(* exact, socle_sort sG) coerce to square matrices. For any *)
+(* representation rG we can construct sG : socleType rG *)
+(* classically; the socleType structure encapsulates this *)
+(* use of classical logic. *)
+(* DecSocleType rG == a socleType rG structure, for a representation over a *)
+(* decidable field type. *)
+(* socle_base W == for W : (sG : socleType), a simple module whose *)
+(* component is W; socle_simple W and socle_module W are *)
+(* proofs that socle_base W is a simple module. *)
+(* socle_mult W == the multiplicity of socle_base W in W : sG. *)
+(* := \rank W %/ \rank (socle_base W) *)
+(* Socle sG == the Socle of rG, given sG : socleType rG, i.e., the *)
+(* (direct) sum of all the components of rG. *)
+(* mx_rsim rG rG' <-> rG and rG' are similar representations of the same *)
+(* group G. Note that rG and rG' must then have equal, but *)
+(* not necessarily convertible, degree. *)
+(* submod_repr modU == a representation of G on 'rV_(\rank U) equivalent to *)
+(* the restriction of rG to U (here modU : mxmodule rG U). *)
+(* socle_repr W := submod_repr (socle_module W) *)
+(* val/in_submod rG U == the projections resp. from/onto 'rV_(\rank U), *)
+(* that correspond to submod_repr r G U (these work both on *)
+(* vectors and row spaces). *)
+(* factmod_repr modV == a representation of G on 'rV_(\rank (cokermx V)) that *)
+(* is equivalent to the factor module 'rV_n / V induced by V *)
+(* and rG (here modV : mxmodule rG V). *)
+(* val/in_factmod rG U == the projections for factmod_repr r G U. *)
+(* section_repr modU modV == the restriction to in_factmod V U of the factor *)
+(* representation factmod_repr modV (for modU : mxmodule rG U *)
+(* and modV : mxmodule rG V); section_repr modU modV is *)
+(* irreducible iff max_submod rG U V. *)
+(* subseries_repr modUs i == the representation for the section module *)
+(* in_factmod (0 :: Us)`_i Us`_i, where *)
+(* modUs : mx_subseries rG Us. *)
+(* series_repr compUs i == the representation for the section module *)
+(* in_factmod (0 :: Us)`_i Us`_i, where *)
+(* compUs : mx_composition_series rG Us. The Jordan-Holder *)
+(* theorem asserts the uniqueness of the set of such *)
+(* representations, up to similarity and permutation. *)
+(* regular_repr F G == the regular F-representation of the group G. *)
+(* group_ring F G == a #|G| x #|G|^2 matrix that encodes the free group *)
+(* ring of G -- that is, the enveloping algebra of the *)
+(* regular F-representation of G. *)
+(* gring_index x == the index corresponding to x \in G in the matrix *)
+(* encoding of regular_repr and group_ring. *)
+(* gring_row A == the row vector corresponding to A \in group_ring F G in *)
+(* the regular FG-module. *)
+(* gring_proj x A == the 1 x 1 matrix holding the coefficient of x \in G in *)
+(* (A \in group_ring F G)%MS. *)
+(* gring_mx rG u == the image of a row vector u of the regular FG-module, *)
+(* in the enveloping algebra of another representation rG. *)
+(* gring_op rG A == the image of a matrix of the free group ring of G, *)
+(* in the enveloping algebra of rG. *)
+(* gset_mx F G C == the group sum of C in the free group ring of G -- the *)
+(* sum of the images of all the x \in C in group_ring F G. *)
+(* classg_base F G == a #|classes G| x #|G|^2 matrix whose rows encode the *)
+(* group sums of the conjugacy classes of G -- this is a *)
+(* basis of 'Z(group_ring F G)%MS. *)
+(* irrType F G == a type indexing irreducible representations of G over a *)
+(* field F, provided its characteristic does not divide the *)
+(* order of G; it also indexes Wedderburn subrings. *)
+(* := socleType (regular_repr F G) *)
+(* irr_repr i == the irreducible representation corresponding to the *)
+(* index i : irrType sG *)
+(* := socle_repr i as i coerces to a component matrix. *)
+(* 'n_i, irr_degree i == the degree of irr_repr i; the notation is only *)
+(* active after Open Scope group_ring_scope. *)
+(* linear_irr sG == the set of sG-indices of linear irreducible *)
+(* representations of G. *)
+(* irr_comp sG rG == the sG-index of the unique irreducible representation *)
+(* similar to rG, at least when rG is irreducible and the *)
+(* characteristic is coprime. *)
+(* irr_mode i z == the unique eigenvalue of irr_repr i z, at least when *)
+(* irr_repr i z is scalar (e.g., when z \in 'Z(G)). *)
+(* [1 sG]%irr == the index of the principal representation of G, in *)
+(* sG : irrType F G. The i argument ot irr_repr, irr_degree *)
+(* and irr_mode is in the %irr scope. This notation may be *)
+(* replaced locally by an interpretation of 1%irr as [1 sG] *)
+(* for some specific irrType sG. *)
+(* 'R_i, Wedderburn_subring i == the subring (indeed, the component) of the *)
+(* free group ring of G corresponding to the component i : sG *)
+(* of the regular FG-module, where sG : irrType F g. In *)
+(* coprime characteristic the Wedderburn structure theorem *)
+(* asserts that the free group ring is the direct sum of *)
+(* these subrings; as with 'n_i above, the notation is only *)
+(* active in group_ring_scope. *)
+(* 'e_i, Wedderburn_id i == the projection of the identity matrix 1%:M on the *)
+(* Wedderburn subring of i : sG (with sG a socleType). In *)
+(* coprime characteristic this is the identity element of *)
+(* the subring, and the basis of its center if the field F is *)
+(* a splitting field. As 'R_i, 'e_i is in group_ring_scope. *)
+(* subg_repr rG sHG == the restriction to H of the representation rG of G; *)
+(* here sHG : H \subset G. *)
+(* eqg_repr rG eqHG == the representation rG of G viewed a a representation *)
+(* of H; here eqHG : G == H. *)
+(* morphpre_repr f rG == the representation of f @*^-1 G obtained by *)
+(* composing the group morphism f with rG. *)
+(* morphim_repr rGf sGD == the representation of G induced by a *)
+(* representation rGf of f @* G; here sGD : G \subset D where *)
+(* D is the domain of the group morphism f. *)
+(* rconj_repr rG uB == the conjugate representation x |-> B * rG x * B^-1; *)
+(* here uB : B \in unitmx. *)
+(* quo_repr sHK nHG == the representation of G / H induced by rG, given *)
+(* sHK : H \subset rker rG, and nHG : G \subset 'N(H). *)
+(* kquo_repr rG == the representation induced on G / rker rG by rG. *)
+(* map_repr f rG == the representation f \o rG, whose module is the tensor *)
+(* product of the module of rG with the extension field into *)
+(* which f : {rmorphism F -> Fstar} embeds F. *)
+(* 'Cl%act == the transitive action of G on the Wedderburn components of *)
+(* H, with nsGH : H <| G, given by Clifford's theorem. More *)
+(* precisely this is a total action of G on socle_sort sH, *)
+(* where sH : socleType (subg_repr rG (normal_sub sGH)). *)
+(* More involved constructions are encapsulated in two Coq submodules: *)
+(* MatrixGenField == a module that encapsulates the lengthy details of the *)
+(* construction of appropriate extension fields. We assume we *)
+(* have an irreducible representation r of a group G, and a *)
+(* non-scalar matrix A that centralises an r(G), as this data *)
+(* is readily extracted from the Jacobson density theorem. It *)
+(* then follows from Schur's lemma that the ring generated by *)
+(* A is a field on which the extension of the representation *)
+(* r of G is reducible. Note that this is equivalent to the *)
+(* more traditional quotient of the polynomial ring by an *)
+(* irreducible polynomial (the minimal polynomial of A), but *)
+(* much better suited to our needs. *)
+(* Here are the main definitions of MatrixGenField; they all have three *)
+(* proofs as arguments: rG : mx_repr r G, irrG : mx_irreducible rG, and *)
+(* cGA : mxcentg rG A, which ensure the validity of the construction and *)
+(* allow us to define Canonical instances (the ~~ is_scalar_mx A assumption *)
+(* is only needed to prove reducibility). *)
+(* + gen_of irrG cGA == the carrier type of the field generated by A. It is *)
+(* at least equipped with a fieldType structure; we also *)
+(* propagate any decFieldType/finFieldType structures on the *)
+(* original field. *)
+(* + gen irrG cGA == the morphism injecting into gen_of rG irrG cGA. *)
+(* + groot irrG cGA == the root of mxminpoly A in the gen_of field. *)
+(* + gen_repr irrG cGA == an alternative to the field extension *)
+(* representation, which consists in reconsidering the *)
+(* original module as a module over the new gen_of field, *)
+(* thereby DIVIDING the original dimension n by the degree of *)
+(* the minimal polynomial of A. This can be simpler than the *)
+(* extension method, and is actually required by the proof *)
+(* that odd groups are p-stable (B & G 6.1-2, and Appendix A) *)
+(* but is only applicable if G is the LARGEST group *)
+(* represented by rG (e.g., NOT for B & G 2.6). *)
+(* + val_gen/in_gen rG irrG cGA : the bijections from/to the module *)
+(* corresponding to gen_repr. *)
+(* + rowval_gen rG irrG cGA : the projection of row spaces in the module *)
+(* corresponding to gen_repr to row spaces in 'rV_n. *)
+(* We build on the MatrixFormula toolkit to define decision procedures for *)
+(* the reducibility property: *)
+(* + mxmodule_form rG U == a formula asserting that the interpretation of U *)
+(* is a module of the representation rG of G via r. *)
+(* + mxnonsimple_form rG U == a formula asserting that the interpretation *)
+(* of U contains a proper nontrivial rG-module. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope GRing.Theory.
+Local Open Scope ring_scope.
+
+Reserved Notation "''n_' i" (at level 8, i at level 2, format "''n_' i").
+Reserved Notation "''R_' i" (at level 8, i at level 2, format "''R_' i").
+Reserved Notation "''e_' i" (at level 8, i at level 2, format "''e_' i").
+
+Delimit Scope irrType_scope with irr.
+
+Section RingRepr.
+
+Variable R : comUnitRingType.
+
+Section OneRepresentation.
+
+Variable gT : finGroupType.
+
+Definition mx_repr (G : {set gT}) n (r : gT -> 'M[R]_n) :=
+ r 1%g = 1%:M /\ {in G &, {morph r : x y / (x * y)%g >-> x *m y}}.
+
+Structure mx_representation G n :=
+ MxRepresentation { repr_mx :> gT -> 'M_n; _ : mx_repr G repr_mx }.
+
+Variables (G : {group gT}) (n : nat) (rG : mx_representation G n).
+Arguments Scope rG [group_scope].
+
+Lemma repr_mx1 : rG 1 = 1%:M.
+Proof. by case: rG => r []. Qed.
+
+Lemma repr_mxM : {in G &, {morph rG : x y / (x * y)%g >-> x *m y}}.
+Proof. by case: rG => r []. Qed.
+
+Lemma repr_mxK m x :
+ x \in G -> cancel ((@mulmx _ m n n)^~ (rG x)) (mulmx^~ (rG x^-1)).
+Proof.
+by move=> Gx U; rewrite -mulmxA -repr_mxM ?groupV // mulgV repr_mx1 mulmx1.
+Qed.
+
+Lemma repr_mxKV m x :
+ x \in G -> cancel ((@mulmx _ m n n)^~ (rG x^-1)) (mulmx^~ (rG x)).
+Proof. by rewrite -groupV -{3}[x]invgK; exact: repr_mxK. Qed.
+
+Lemma repr_mx_unit x : x \in G -> rG x \in unitmx.
+Proof. by move=> Gx; case/mulmx1_unit: (repr_mxKV Gx 1%:M). Qed.
+
+Lemma repr_mxV : {in G, {morph rG : x / x^-1%g >-> invmx x}}.
+Proof.
+by move=> x Gx /=; rewrite -[rG x^-1](mulKmx (repr_mx_unit Gx)) mulmxA repr_mxK.
+Qed.
+
+(* This is only used in the group ring construction below, as we only have *)
+(* developped the theory of matrix subalgebras for F-algebras. *)
+Definition enveloping_algebra_mx := \matrix_(i < #|G|) mxvec (rG (enum_val i)).
+
+Section Stabiliser.
+
+Variables (m : nat) (U : 'M[R]_(m, n)).
+
+Definition rstab := [set x in G | U *m rG x == U].
+
+Lemma rstab_sub : rstab \subset G.
+Proof. by apply/subsetP=> x; case/setIdP. Qed.
+
+Lemma rstab_group_set : group_set rstab.
+Proof.
+apply/group_setP; rewrite inE group1 repr_mx1 mulmx1; split=> //= x y.
+case/setIdP=> Gx cUx; case/setIdP=> Gy cUy; rewrite inE repr_mxM ?groupM //.
+by rewrite mulmxA (eqP cUx).
+Qed.
+Canonical rstab_group := Group rstab_group_set.
+
+End Stabiliser.
+
+(* Centralizer subgroup and central homomorphisms. *)
+Section CentHom.
+
+Variable f : 'M[R]_n.
+
+Definition rcent := [set x in G | f *m rG x == rG x *m f].
+
+Lemma rcent_sub : rcent \subset G.
+Proof. by apply/subsetP=> x; case/setIdP. Qed.
+
+Lemma rcent_group_set : group_set rcent.
+Proof.
+apply/group_setP; rewrite inE group1 repr_mx1 mulmx1 mul1mx; split=> //= x y.
+case/setIdP=> Gx; move/eqP=> cfx; case/setIdP=> Gy; move/eqP=> cfy.
+by rewrite inE repr_mxM ?groupM //= -mulmxA -cfy !mulmxA cfx.
+Qed.
+Canonical rcent_group := Group rcent_group_set.
+
+Definition centgmx := G \subset rcent.
+
+Lemma centgmxP : reflect (forall x, x \in G -> f *m rG x = rG x *m f) centgmx.
+Proof.
+apply: (iffP subsetP) => cGf x Gx;
+ by have:= cGf x Gx; rewrite !inE Gx /=; move/eqP.
+Qed.
+
+End CentHom.
+
+(* Representation kernel, and faithful representations. *)
+
+Definition rker := rstab 1%:M.
+Canonical rker_group := Eval hnf in [group of rker].
+
+Lemma rkerP x : reflect (x \in G /\ rG x = 1%:M) (x \in rker).
+Proof. by apply: (iffP setIdP) => [] [->]; move/eqP; rewrite mul1mx. Qed.
+
+Lemma rker_norm : G \subset 'N(rker).
+Proof.
+apply/subsetP=> x Gx; rewrite inE sub_conjg; apply/subsetP=> y.
+case/rkerP=> Gy ry1; rewrite mem_conjgV !inE groupJ //=.
+by rewrite !repr_mxM ?groupM ?groupV // ry1 !mulmxA mulmx1 repr_mxKV.
+Qed.
+
+Lemma rker_normal : rker <| G.
+Proof. by rewrite /normal rstab_sub rker_norm. Qed.
+
+Definition mx_faithful := rker \subset [1].
+
+Lemma mx_faithful_inj : mx_faithful -> {in G &, injective rG}.
+Proof.
+move=> ffulG x y Gx Gy eq_rGxy; apply/eqP; rewrite eq_mulgV1 -in_set1.
+rewrite (subsetP ffulG) // inE groupM ?repr_mxM ?groupV //= eq_rGxy.
+by rewrite mulmxA repr_mxK.
+Qed.
+
+Lemma rker_linear : n = 1%N -> G^`(1)%g \subset rker.
+Proof.
+move=> n1; rewrite gen_subG; apply/subsetP=> xy; case/imset2P=> x y Gx Gy ->.
+rewrite !inE groupR //= /commg mulgA -invMg repr_mxM ?groupV ?groupM //.
+rewrite mulmxA (can2_eq (repr_mxK _) (repr_mxKV _)) ?groupM //.
+rewrite !repr_mxV ?repr_mxM ?groupM //; move: (rG x) (rG y).
+by rewrite n1 => rx ry; rewrite (mx11_scalar rx) scalar_mxC.
+Qed.
+
+(* Representation center. *)
+
+Definition rcenter := [set g in G | is_scalar_mx (rG g)].
+
+Fact rcenter_group_set : group_set rcenter.
+Proof.
+apply/group_setP; split=> [|x y].
+ by rewrite inE group1 repr_mx1 scalar_mx_is_scalar.
+move=> /setIdP[Gx /is_scalar_mxP[a defx]] /setIdP[Gy /is_scalar_mxP[b defy]].
+by rewrite !inE groupM ?repr_mxM // defx defy -scalar_mxM ?scalar_mx_is_scalar.
+Qed.
+Canonical rcenter_group := Group rcenter_group_set.
+
+Lemma rcenter_normal : rcenter <| G.
+Proof.
+rewrite /normal /rcenter {1}setIdE subsetIl; apply/subsetP=> x Gx; rewrite inE.
+apply/subsetP=> _ /imsetP[y /setIdP[Gy /is_scalar_mxP[c rGy]] ->].
+rewrite inE !repr_mxM ?groupM ?groupV //= mulmxA rGy scalar_mxC repr_mxKV //.
+exact: scalar_mx_is_scalar.
+Qed.
+
+End OneRepresentation.
+
+Implicit Arguments rkerP [gT G n rG x].
+
+Section Proper.
+
+Variables (gT : finGroupType) (G : {group gT}) (n' : nat).
+Local Notation n := n'.+1.
+Variable rG : mx_representation G n.
+
+Lemma repr_mxMr : {in G &, {morph rG : x y / (x * y)%g >-> x * y}}.
+Proof. exact: repr_mxM. Qed.
+
+Lemma repr_mxVr : {in G, {morph rG : x / (x^-1)%g >-> x^-1}}.
+Proof. exact: repr_mxV.
+ Qed.
+
+Lemma repr_mx_unitr x : x \in G -> rG x \is a GRing.unit.
+Proof. exact: repr_mx_unit. Qed.
+
+Lemma repr_mxX m : {in G, {morph rG : x / (x ^+ m)%g >-> x ^+ m}}.
+Proof.
+elim: m => [|m IHm] x Gx; rewrite /= ?repr_mx1 // expgS exprS -IHm //.
+by rewrite repr_mxM ?groupX.
+Qed.
+
+End Proper.
+
+Section ChangeGroup.
+
+Variables (gT : finGroupType) (G H : {group gT}) (n : nat).
+Variables (rG : mx_representation G n).
+
+Section SubGroup.
+
+Hypothesis sHG : H \subset G.
+
+Lemma subg_mx_repr : mx_repr H rG.
+Proof.
+by split=> [|x y Hx Hy]; rewrite (repr_mx1, repr_mxM) ?(subsetP sHG).
+Qed.
+Definition subg_repr := MxRepresentation subg_mx_repr.
+Local Notation rH := subg_repr.
+
+Lemma rcent_subg U : rcent rH U = H :&: rcent rG U.
+Proof. by apply/setP=> x; rewrite !inE andbA -in_setI (setIidPl sHG). Qed.
+
+Section Stabiliser.
+
+Variables (m : nat) (U : 'M[R]_(m, n)).
+
+Lemma rstab_subg : rstab rH U = H :&: rstab rG U.
+Proof. by apply/setP=> x; rewrite !inE andbA -in_setI (setIidPl sHG). Qed.
+
+End Stabiliser.
+
+Lemma rker_subg : rker rH = H :&: rker rG. Proof. exact: rstab_subg. Qed.
+
+Lemma subg_mx_faithful : mx_faithful rG -> mx_faithful rH.
+Proof. by apply: subset_trans; rewrite rker_subg subsetIr. Qed.
+
+End SubGroup.
+
+Section SameGroup.
+
+Hypothesis eqGH : G :==: H.
+
+Lemma eqg_repr_proof : H \subset G. Proof. by rewrite (eqP eqGH). Qed.
+
+Definition eqg_repr := subg_repr eqg_repr_proof.
+Local Notation rH := eqg_repr.
+
+Lemma rcent_eqg U : rcent rH U = rcent rG U.
+Proof. by rewrite rcent_subg -(eqP eqGH) (setIidPr _) ?rcent_sub. Qed.
+
+Section Stabiliser.
+
+Variables (m : nat) (U : 'M[R]_(m, n)).
+
+Lemma rstab_eqg : rstab rH U = rstab rG U.
+Proof. by rewrite rstab_subg -(eqP eqGH) (setIidPr _) ?rstab_sub. Qed.
+
+End Stabiliser.
+
+Lemma rker_eqg : rker rH = rker rG. Proof. exact: rstab_eqg. Qed.
+
+Lemma eqg_mx_faithful : mx_faithful rH = mx_faithful rG.
+Proof. by rewrite /mx_faithful rker_eqg. Qed.
+
+End SameGroup.
+
+End ChangeGroup.
+
+Section Morphpre.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+Variables (G : {group rT}) (n : nat) (rG : mx_representation G n).
+
+Lemma morphpre_mx_repr : mx_repr (f @*^-1 G) (rG \o f).
+Proof.
+split=> [|x y]; first by rewrite /= morph1 repr_mx1.
+case/morphpreP=> Dx Gfx; case/morphpreP=> Dy Gfy.
+by rewrite /= morphM ?repr_mxM.
+Qed.
+Canonical morphpre_repr := MxRepresentation morphpre_mx_repr.
+Local Notation rGf := morphpre_repr.
+
+Section Stabiliser.
+
+Variables (m : nat) (U : 'M[R]_(m, n)).
+
+Lemma rstab_morphpre : rstab rGf U = f @*^-1 (rstab rG U).
+Proof. by apply/setP=> x; rewrite !inE andbA. Qed.
+
+End Stabiliser.
+
+Lemma rker_morphpre : rker rGf = f @*^-1 (rker rG).
+Proof. exact: rstab_morphpre. Qed.
+
+End Morphpre.
+
+Section Morphim.
+
+Variables (aT rT : finGroupType) (G D : {group aT}) (f : {morphism D >-> rT}).
+Variables (n : nat) (rGf : mx_representation (f @* G) n).
+
+Definition morphim_mx of G \subset D := fun x => rGf (f x).
+
+Hypothesis sGD : G \subset D.
+
+Lemma morphim_mxE x : morphim_mx sGD x = rGf (f x). Proof. by []. Qed.
+
+Let sG_f'fG : G \subset f @*^-1 (f @* G).
+Proof. by rewrite -sub_morphim_pre. Qed.
+
+Lemma morphim_mx_repr : mx_repr G (morphim_mx sGD).
+Proof. exact: subg_mx_repr (morphpre_repr f rGf) sG_f'fG. Qed.
+Canonical morphim_repr := MxRepresentation morphim_mx_repr.
+Local Notation rG := morphim_repr.
+
+Section Stabiliser.
+Variables (m : nat) (U : 'M[R]_(m, n)).
+
+Lemma rstab_morphim : rstab rG U = G :&: f @*^-1 rstab rGf U.
+Proof. by rewrite -rstab_morphpre -(rstab_subg _ sG_f'fG). Qed.
+
+End Stabiliser.
+
+Lemma rker_morphim : rker rG = G :&: f @*^-1 (rker rGf).
+Proof. exact: rstab_morphim. Qed.
+
+End Morphim.
+
+Section Conjugate.
+
+Variables (gT : finGroupType) (G : {group gT}) (n : nat).
+Variables (rG : mx_representation G n) (B : 'M[R]_n).
+
+Definition rconj_mx of B \in unitmx := fun x => B *m rG x *m invmx B.
+
+Hypothesis uB : B \in unitmx.
+
+Lemma rconj_mx_repr : mx_repr G (rconj_mx uB).
+Proof.
+split=> [|x y Gx Gy]; rewrite /rconj_mx ?repr_mx1 ?mulmx1 ?mulmxV ?repr_mxM //.
+by rewrite !mulmxA mulmxKV.
+Qed.
+Canonical rconj_repr := MxRepresentation rconj_mx_repr.
+Local Notation rGB := rconj_repr.
+
+Lemma rconj_mxE x : rGB x = B *m rG x *m invmx B.
+Proof. by []. Qed.
+
+Lemma rconj_mxJ m (W : 'M_(m, n)) x : W *m rGB x *m B = W *m B *m rG x.
+Proof. by rewrite !mulmxA mulmxKV. Qed.
+
+Lemma rcent_conj A : rcent rGB A = rcent rG (invmx B *m A *m B).
+Proof.
+apply/setP=> x; rewrite !inE /= rconj_mxE !mulmxA.
+rewrite (can2_eq (mulmxKV uB) (mulmxK uB)) -!mulmxA.
+by rewrite -(can2_eq (mulKVmx uB) (mulKmx uB)).
+Qed.
+
+Lemma rstab_conj m (U : 'M_(m, n)) : rstab rGB U = rstab rG (U *m B).
+Proof.
+apply/setP=> x; rewrite !inE /= rconj_mxE !mulmxA.
+by rewrite (can2_eq (mulmxKV uB) (mulmxK uB)).
+Qed.
+
+Lemma rker_conj : rker rGB = rker rG.
+Proof.
+apply/setP=> x; rewrite !inE /= mulmxA (can2_eq (mulmxKV uB) (mulmxK uB)).
+by rewrite mul1mx -scalar_mxC (inj_eq (can_inj (mulKmx uB))) mul1mx.
+Qed.
+
+Lemma conj_mx_faithful : mx_faithful rGB = mx_faithful rG.
+Proof. by rewrite /mx_faithful rker_conj. Qed.
+
+End Conjugate.
+
+Section Quotient.
+
+Variables (gT : finGroupType) (G : {group gT}) (n : nat).
+Variable rG : mx_representation G n.
+
+Definition quo_mx (H : {set gT}) of H \subset rker rG & G \subset 'N(H) :=
+ fun Hx : coset_of H => rG (repr Hx).
+
+Section SubQuotient.
+
+Variable H : {group gT}.
+Hypotheses (krH : H \subset rker rG) (nHG : G \subset 'N(H)).
+Let nHGs := subsetP nHG.
+
+Lemma quo_mx_coset x : x \in G -> quo_mx krH nHG (coset H x) = rG x.
+Proof.
+move=> Gx; rewrite /quo_mx val_coset ?nHGs //; case: repr_rcosetP => z Hz.
+by case/rkerP: (subsetP krH z Hz) => Gz rz1; rewrite repr_mxM // rz1 mul1mx.
+Qed.
+
+Lemma quo_mx_repr : mx_repr (G / H)%g (quo_mx krH nHG).
+Proof.
+split=> [|Hx Hy]; first by rewrite /quo_mx repr_coset1 repr_mx1.
+case/morphimP=> x Nx Gx ->{Hx}; case/morphimP=> y Ny Gy ->{Hy}.
+by rewrite -morphM // !quo_mx_coset ?groupM ?repr_mxM.
+Qed.
+Canonical quo_repr := MxRepresentation quo_mx_repr.
+Local Notation rGH := quo_repr.
+
+Lemma quo_repr_coset x : x \in G -> rGH (coset H x) = rG x.
+Proof. exact: quo_mx_coset. Qed.
+
+Lemma rcent_quo A : rcent rGH A = (rcent rG A / H)%g.
+Proof.
+apply/setP=> Hx; rewrite !inE.
+apply/andP/idP=> [[]|]; case/morphimP=> x Nx Gx ->{Hx}.
+ by rewrite quo_repr_coset // => cAx; rewrite mem_morphim // inE Gx.
+by case/setIdP: Gx => Gx cAx; rewrite quo_repr_coset ?mem_morphim.
+Qed.
+
+Lemma rstab_quo m (U : 'M_(m, n)) : rstab rGH U = (rstab rG U / H)%g.
+Proof.
+apply/setP=> Hx; rewrite !inE.
+apply/andP/idP=> [[]|]; case/morphimP=> x Nx Gx ->{Hx}.
+ by rewrite quo_repr_coset // => nUx; rewrite mem_morphim // inE Gx.
+by case/setIdP: Gx => Gx nUx; rewrite quo_repr_coset ?mem_morphim.
+Qed.
+
+Lemma rker_quo : rker rGH = (rker rG / H)%g.
+Proof. exact: rstab_quo. Qed.
+
+End SubQuotient.
+
+Definition kquo_mx := quo_mx (subxx (rker rG)) (rker_norm rG).
+Lemma kquo_mxE : kquo_mx = quo_mx (subxx (rker rG)) (rker_norm rG).
+Proof. by []. Qed.
+
+Canonical kquo_repr := @MxRepresentation _ _ _ kquo_mx (quo_mx_repr _ _).
+
+Lemma kquo_repr_coset x :
+ x \in G -> kquo_repr (coset (rker rG) x) = rG x.
+Proof. exact: quo_repr_coset. Qed.
+
+Lemma kquo_mx_faithful : mx_faithful kquo_repr.
+Proof. by rewrite /mx_faithful rker_quo trivg_quotient. Qed.
+
+End Quotient.
+
+Section Regular.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Local Notation nG := #|pred_of_set (gval G)|.
+
+Definition gring_index (x : gT) := enum_rank_in (group1 G) x.
+
+Lemma gring_valK : cancel enum_val gring_index.
+Proof. exact: enum_valK_in. Qed.
+
+Lemma gring_indexK : {in G, cancel gring_index enum_val}.
+Proof. exact: enum_rankK_in. Qed.
+
+Definition regular_mx x : 'M[R]_nG :=
+ \matrix_i delta_mx 0 (gring_index (enum_val i * x)).
+
+Lemma regular_mx_repr : mx_repr G regular_mx.
+Proof.
+split=> [|x y Gx Gy]; apply/row_matrixP=> i; rewrite !rowK.
+ by rewrite mulg1 row1 gring_valK.
+by rewrite row_mul rowK -rowE rowK mulgA gring_indexK // groupM ?enum_valP.
+Qed.
+Canonical regular_repr := MxRepresentation regular_mx_repr.
+Local Notation aG := regular_repr.
+
+Definition group_ring := enveloping_algebra_mx aG.
+Local Notation R_G := group_ring.
+
+Definition gring_row : 'M[R]_nG -> 'rV_nG := row (gring_index 1).
+Canonical gring_row_linear := [linear of gring_row].
+
+Lemma gring_row_mul A B : gring_row (A *m B) = gring_row A *m B.
+Proof. exact: row_mul. Qed.
+
+Definition gring_proj x := row (gring_index x) \o trmx \o gring_row.
+Canonical gring_proj_linear x := [linear of gring_proj x].
+
+Lemma gring_projE : {in G &, forall x y, gring_proj x (aG y) = (x == y)%:R}.
+Proof.
+move=> x y Gx Gy; rewrite /gring_proj /= /gring_row rowK gring_indexK //=.
+rewrite mul1g trmx_delta rowE mul_delta_mx_cond [delta_mx 0 0]mx11_scalar !mxE.
+by rewrite /= -(inj_eq (can_inj gring_valK)) !gring_indexK.
+Qed.
+
+Lemma regular_mx_faithful : mx_faithful aG.
+Proof.
+apply/subsetP=> x /setIdP[Gx].
+rewrite mul1mx inE => /eqP/(congr1 (gring_proj 1%g)).
+rewrite -(repr_mx1 aG) !gring_projE ?group1 // eqxx eq_sym.
+by case: (x == _) => // /eqP; rewrite eq_sym oner_eq0.
+Qed.
+
+Section GringMx.
+
+Variables (n : nat) (rG : mx_representation G n).
+
+Definition gring_mx := vec_mx \o mulmxr (enveloping_algebra_mx rG).
+Canonical gring_mx_linear := [linear of gring_mx].
+
+Lemma gring_mxJ a x :
+ x \in G -> gring_mx (a *m aG x) = gring_mx a *m rG x.
+Proof.
+move=> Gx; rewrite /gring_mx /= ![a *m _]mulmx_sum_row.
+rewrite !(mulmx_suml, linear_sum); apply: eq_bigr => i _.
+rewrite linearZ -!scalemxAl linearZ /=; congr (_ *: _) => {a}.
+rewrite !rowK /= !mxvecK -rowE rowK mxvecK.
+by rewrite gring_indexK ?groupM ?repr_mxM ?enum_valP.
+Qed.
+
+End GringMx.
+
+Lemma gring_mxK : cancel (gring_mx aG) gring_row.
+Proof.
+move=> a; rewrite /gring_mx /= mulmx_sum_row !linear_sum.
+rewrite {2}[a]row_sum_delta; apply: eq_bigr => i _.
+rewrite !linearZ /= /gring_row !(rowK, mxvecK).
+by rewrite gring_indexK // mul1g gring_valK.
+Qed.
+
+Section GringOp.
+
+Variables (n : nat) (rG : mx_representation G n).
+
+Definition gring_op := gring_mx rG \o gring_row.
+Canonical gring_op_linear := [linear of gring_op].
+
+Lemma gring_opE a : gring_op a = gring_mx rG (gring_row a).
+Proof. by []. Qed.
+
+Lemma gring_opG x : x \in G -> gring_op (aG x) = rG x.
+Proof.
+move=> Gx; rewrite gring_opE /gring_row rowK gring_indexK // mul1g.
+by rewrite /gring_mx /= -rowE rowK mxvecK gring_indexK.
+Qed.
+
+Lemma gring_op1 : gring_op 1%:M = 1%:M.
+Proof. by rewrite -(repr_mx1 aG) gring_opG ?repr_mx1. Qed.
+
+Lemma gring_opJ A b :
+ gring_op (A *m gring_mx aG b) = gring_op A *m gring_mx rG b.
+Proof.
+rewrite /gring_mx /= ![b *m _]mulmx_sum_row !linear_sum.
+apply: eq_bigr => i _; rewrite !linearZ /= !rowK !mxvecK.
+by rewrite gring_opE gring_row_mul gring_mxJ ?enum_valP.
+Qed.
+
+Lemma gring_op_mx b : gring_op (gring_mx aG b) = gring_mx rG b.
+Proof. by rewrite -[_ b]mul1mx gring_opJ gring_op1 mul1mx. Qed.
+
+Lemma gring_mxA a b :
+ gring_mx rG (a *m gring_mx aG b) = gring_mx rG a *m gring_mx rG b.
+Proof.
+by rewrite -(gring_op_mx a) -gring_opJ gring_opE gring_row_mul gring_mxK.
+Qed.
+
+End GringOp.
+
+End Regular.
+
+End RingRepr.
+
+Arguments Scope mx_representation [_ _ group_scope nat_scope].
+Arguments Scope mx_repr [_ _ group_scope nat_scope _].
+Arguments Scope group_ring [_ _ group_scope].
+Arguments Scope regular_repr [_ _ group_scope].
+
+Implicit Arguments centgmxP [R gT G n rG f].
+Implicit Arguments rkerP [R gT G n rG x].
+Prenex Implicits gring_mxK.
+
+Section ChangeOfRing.
+
+Variables (aR rR : comUnitRingType) (f : {rmorphism aR -> rR}).
+Local Notation "A ^f" := (map_mx (GRing.RMorphism.apply f) A) : ring_scope.
+Variables (gT : finGroupType) (G : {group gT}).
+
+Lemma map_regular_mx x : (regular_mx aR G x)^f = regular_mx rR G x.
+Proof. by apply/matrixP=> i j; rewrite !mxE rmorph_nat. Qed.
+
+Lemma map_gring_row (A : 'M_#|G|) : (gring_row A)^f = gring_row A^f.
+Proof. by rewrite map_row. Qed.
+
+Lemma map_gring_proj x (A : 'M_#|G|) : (gring_proj x A)^f = gring_proj x A^f.
+Proof. by rewrite map_row -map_trmx map_gring_row. Qed.
+
+Section OneRepresentation.
+
+Variables (n : nat) (rG : mx_representation aR G n).
+
+Definition map_repr_mx (f0 : aR -> rR) rG0 (g : gT) : 'M_n := map_mx f0 (rG0 g).
+
+Lemma map_mx_repr : mx_repr G (map_repr_mx f rG).
+Proof.
+split=> [|x y Gx Gy]; first by rewrite /map_repr_mx repr_mx1 map_mx1.
+by rewrite -map_mxM -repr_mxM.
+Qed.
+Canonical map_repr := MxRepresentation map_mx_repr.
+Local Notation rGf := map_repr.
+
+Lemma map_reprE x : rGf x = (rG x)^f. Proof. by []. Qed.
+
+Lemma map_reprJ m (A : 'M_(m, n)) x : (A *m rG x)^f = A^f *m rGf x.
+Proof. exact: map_mxM. Qed.
+
+Lemma map_enveloping_algebra_mx :
+ (enveloping_algebra_mx rG)^f = enveloping_algebra_mx rGf.
+Proof. by apply/row_matrixP=> i; rewrite -map_row !rowK map_mxvec. Qed.
+
+Lemma map_gring_mx a : (gring_mx rG a)^f = gring_mx rGf a^f.
+Proof. by rewrite map_vec_mx map_mxM map_enveloping_algebra_mx. Qed.
+
+Lemma map_gring_op A : (gring_op rG A)^f = gring_op rGf A^f.
+Proof. by rewrite map_gring_mx map_gring_row. Qed.
+
+End OneRepresentation.
+
+Lemma map_regular_repr : map_repr (regular_repr aR G) =1 regular_repr rR G.
+Proof. exact: map_regular_mx. Qed.
+
+Lemma map_group_ring : (group_ring aR G)^f = group_ring rR G.
+Proof.
+rewrite map_enveloping_algebra_mx; apply/row_matrixP=> i.
+by rewrite !rowK map_regular_repr.
+Qed.
+
+(* Stabilisers, etc, are only mapped properly for fields. *)
+
+End ChangeOfRing.
+
+Section FieldRepr.
+
+Variable F : fieldType.
+
+Section OneRepresentation.
+
+Variable gT : finGroupType.
+
+Variables (G : {group gT}) (n : nat) (rG : mx_representation F G n).
+Arguments Scope rG [group_scope].
+
+Local Notation E_G := (enveloping_algebra_mx rG).
+
+Lemma repr_mx_free x : x \in G -> row_free (rG x).
+Proof. by move=> Gx; rewrite row_free_unit repr_mx_unit. Qed.
+
+Section Stabilisers.
+
+Variables (m : nat) (U : 'M[F]_(m, n)).
+
+Definition rstabs := [set x in G | U *m rG x <= U]%MS.
+
+Lemma rstabs_sub : rstabs \subset G.
+Proof. by apply/subsetP=> x /setIdP[]. Qed.
+
+Lemma rstabs_group_set : group_set rstabs.
+Proof.
+apply/group_setP; rewrite inE group1 repr_mx1 mulmx1.
+split=> //= x y /setIdP[Gx nUx] /setIdP[Gy]; rewrite inE repr_mxM ?groupM //.
+by apply: submx_trans; rewrite mulmxA submxMr.
+Qed.
+Canonical rstabs_group := Group rstabs_group_set.
+
+Lemma rstab_act x m1 (W : 'M_(m1, n)) :
+ x \in rstab rG U -> (W <= U)%MS -> W *m rG x = W.
+Proof. by case/setIdP=> _ /eqP cUx /submxP[w ->]; rewrite -mulmxA cUx. Qed.
+
+Lemma rstabs_act x m1 (W : 'M_(m1, n)) :
+ x \in rstabs -> (W <= U)%MS -> (W *m rG x <= U)%MS.
+Proof.
+by case/setIdP=> [_ nUx] sWU; apply: submx_trans nUx; exact: submxMr.
+Qed.
+
+Definition mxmodule := G \subset rstabs.
+
+Lemma mxmoduleP : reflect {in G, forall x, U *m rG x <= U}%MS mxmodule.
+Proof.
+by apply: (iffP subsetP) => modU x Gx; have:= modU x Gx; rewrite !inE ?Gx.
+Qed.
+
+End Stabilisers.
+Implicit Arguments mxmoduleP [m U].
+
+Lemma rstabS m1 m2 (U : 'M_(m1, n)) (V : 'M_(m2, n)) :
+ (U <= V)%MS -> rstab rG V \subset rstab rG U.
+Proof.
+case/submxP=> u ->; apply/subsetP=> x.
+by rewrite !inE => /andP[-> /= /eqP cVx]; rewrite -mulmxA cVx.
+Qed.
+
+Lemma eqmx_rstab m1 m2 (U : 'M_(m1, n)) (V : 'M_(m2, n)) :
+ (U :=: V)%MS -> rstab rG U = rstab rG V.
+Proof. by move=> eqUV; apply/eqP; rewrite eqEsubset !rstabS ?eqUV. Qed.
+
+Lemma eqmx_rstabs m1 m2 (U : 'M_(m1, n)) (V : 'M_(m2, n)) :
+ (U :=: V)%MS -> rstabs U = rstabs V.
+Proof. by move=> eqUV; apply/setP=> x; rewrite !inE eqUV (eqmxMr _ eqUV). Qed.
+
+Lemma eqmx_module m1 m2 (U : 'M_(m1, n)) (V : 'M_(m2, n)) :
+ (U :=: V)%MS -> mxmodule U = mxmodule V.
+Proof. by move=> eqUV; rewrite /mxmodule (eqmx_rstabs eqUV). Qed.
+
+Lemma mxmodule0 m : mxmodule (0 : 'M_(m, n)).
+Proof. by apply/mxmoduleP=> x _; rewrite mul0mx. Qed.
+
+Lemma mxmodule1 : mxmodule 1%:M.
+Proof. by apply/mxmoduleP=> x _; rewrite submx1. Qed.
+
+Lemma mxmodule_trans m1 m2 (U : 'M_(m1, n)) (W : 'M_(m2, n)) x :
+ mxmodule U -> x \in G -> (W <= U -> W *m rG x <= U)%MS.
+Proof.
+by move=> modU Gx sWU; apply: submx_trans (mxmoduleP modU x Gx); exact: submxMr.
+Qed.
+
+Lemma mxmodule_eigenvector m (U : 'M_(m, n)) :
+ mxmodule U -> \rank U = 1%N ->
+ {u : 'rV_n & {a | (U :=: u)%MS & {in G, forall x, u *m rG x = a x *: u}}}.
+Proof.
+move=> modU linU; set u := nz_row U; exists u.
+have defU: (U :=: u)%MS.
+ apply/eqmxP; rewrite andbC -(geq_leqif (mxrank_leqif_eq _)) ?nz_row_sub //.
+ by rewrite linU lt0n mxrank_eq0 nz_row_eq0 -mxrank_eq0 linU.
+pose a x := (u *m rG x *m pinvmx u) 0 0; exists a => // x Gx.
+by rewrite -mul_scalar_mx -mx11_scalar mulmxKpV // -defU mxmodule_trans ?defU.
+Qed.
+
+Lemma addsmx_module m1 m2 U V :
+ @mxmodule m1 U -> @mxmodule m2 V -> mxmodule (U + V)%MS.
+Proof.
+move=> modU modV; apply/mxmoduleP=> x Gx.
+by rewrite addsmxMr addsmxS ?(mxmoduleP _ x Gx).
+Qed.
+
+Lemma sumsmx_module I r (P : pred I) U :
+ (forall i, P i -> mxmodule (U i)) -> mxmodule (\sum_(i <- r | P i) U i)%MS.
+Proof.
+by move=> modU; elim/big_ind: _; [exact: mxmodule0 | exact: addsmx_module | ].
+Qed.
+
+Lemma capmx_module m1 m2 U V :
+ @mxmodule m1 U -> @mxmodule m2 V -> mxmodule (U :&: V)%MS.
+Proof.
+move=> modU modV; apply/mxmoduleP=> x Gx.
+by rewrite sub_capmx !mxmodule_trans ?capmxSl ?capmxSr.
+Qed.
+
+Lemma bigcapmx_module I r (P : pred I) U :
+ (forall i, P i -> mxmodule (U i)) -> mxmodule (\bigcap_(i <- r | P i) U i)%MS.
+Proof.
+by move=> modU; elim/big_ind: _; [exact: mxmodule1 | exact: capmx_module | ].
+Qed.
+
+(* Sub- and factor representations induced by a (sub)module. *)
+Section Submodule.
+
+Variable U : 'M[F]_n.
+
+Definition val_submod m : 'M_(m, \rank U) -> 'M_(m, n) := mulmxr (row_base U).
+Definition in_submod m : 'M_(m, n) -> 'M_(m, \rank U) :=
+ mulmxr (invmx (row_ebase U) *m pid_mx (\rank U)).
+Canonical val_submod_linear m := [linear of @val_submod m].
+Canonical in_submod_linear m := [linear of @in_submod m].
+
+Lemma val_submodE m W : @val_submod m W = W *m val_submod 1%:M.
+Proof. by rewrite mulmxA mulmx1. Qed.
+
+Lemma in_submodE m W : @in_submod m W = W *m in_submod 1%:M.
+Proof. by rewrite mulmxA mulmx1. Qed.
+
+Lemma val_submod1 : (val_submod 1%:M :=: U)%MS.
+Proof. by rewrite /val_submod /= mul1mx; exact: eq_row_base. Qed.
+
+Lemma val_submodP m W : (@val_submod m W <= U)%MS.
+Proof. by rewrite mulmx_sub ?eq_row_base. Qed.
+
+Lemma val_submodK m : cancel (@val_submod m) (@in_submod m).
+Proof.
+move=> W; rewrite /in_submod /= -!mulmxA mulKVmx ?row_ebase_unit //.
+by rewrite pid_mx_id ?rank_leq_row // pid_mx_1 mulmx1.
+Qed.
+
+Lemma val_submod_inj m : injective (@val_submod m).
+Proof. exact: can_inj (@val_submodK m). Qed.
+
+Lemma val_submodS m1 m2 (V : 'M_(m1, \rank U)) (W : 'M_(m2, \rank U)) :
+ (val_submod V <= val_submod W)%MS = (V <= W)%MS.
+Proof.
+apply/idP/idP=> sVW; last exact: submxMr.
+by rewrite -[V]val_submodK -[W]val_submodK submxMr.
+Qed.
+
+Lemma in_submodK m W : (W <= U)%MS -> val_submod (@in_submod m W) = W.
+Proof.
+case/submxP=> w ->; rewrite /val_submod /= -!mulmxA.
+congr (_ *m _); rewrite -{1}[U]mulmx_ebase !mulmxA mulmxK ?row_ebase_unit //.
+by rewrite -2!(mulmxA (col_ebase U)) !pid_mx_id ?rank_leq_row // mulmx_ebase.
+Qed.
+
+Lemma val_submod_eq0 m W : (@val_submod m W == 0) = (W == 0).
+Proof. by rewrite -!submx0 -val_submodS linear0 !(submx0, eqmx0). Qed.
+
+Lemma in_submod_eq0 m W : (@in_submod m W == 0) = (W <= U^C)%MS.
+Proof.
+apply/eqP/submxP=> [W_U0 | [w ->{W}]].
+ exists (W *m invmx (row_ebase U)).
+ rewrite mulmxA mulmxBr mulmx1 -(pid_mx_id _ _ _ (leqnn _)).
+ rewrite mulmxA -(mulmxA W) [W *m (_ *m _)]W_U0 mul0mx subr0.
+ by rewrite mulmxKV ?row_ebase_unit.
+rewrite /in_submod /= -!mulmxA mulKVmx ?row_ebase_unit //.
+by rewrite mul_copid_mx_pid ?rank_leq_row ?mulmx0.
+Qed.
+
+Lemma mxrank_in_submod m (W : 'M_(m, n)) :
+ (W <= U)%MS -> \rank (in_submod W) = \rank W.
+Proof.
+by move=> sWU; apply/eqP; rewrite eqn_leq -{3}(in_submodK sWU) !mxrankM_maxl.
+Qed.
+
+Definition val_factmod m : _ -> 'M_(m, n) :=
+ mulmxr (row_base (cokermx U) *m row_ebase U).
+Definition in_factmod m : 'M_(m, n) -> _ := mulmxr (col_base (cokermx U)).
+Canonical val_factmod_linear m := [linear of @val_factmod m].
+Canonical in_factmod_linear m := [linear of @in_factmod m].
+
+Lemma val_factmodE m W : @val_factmod m W = W *m val_factmod 1%:M.
+Proof. by rewrite mulmxA mulmx1. Qed.
+
+Lemma in_factmodE m W : @in_factmod m W = W *m in_factmod 1%:M.
+Proof. by rewrite mulmxA mulmx1. Qed.
+
+Lemma val_factmodP m W : (@val_factmod m W <= U^C)%MS.
+Proof.
+by rewrite mulmx_sub {m W}// (eqmxMr _ (eq_row_base _)) -mulmxA submxMl.
+Qed.
+
+Lemma val_factmodK m : cancel (@val_factmod m) (@in_factmod m).
+Proof.
+move=> W /=; rewrite /in_factmod /=; set Uc := cokermx U.
+apply: (row_free_inj (row_base_free Uc)); rewrite -mulmxA mulmx_base.
+rewrite /val_factmod /= 2!mulmxA -/Uc mulmxK ?row_ebase_unit //.
+have /submxP[u ->]: (row_base Uc <= Uc)%MS by rewrite eq_row_base.
+by rewrite -!mulmxA copid_mx_id ?rank_leq_row.
+Qed.
+
+Lemma val_factmod_inj m : injective (@val_factmod m).
+Proof. exact: can_inj (@val_factmodK m). Qed.
+
+Lemma val_factmodS m1 m2 (V : 'M_(m1, _)) (W : 'M_(m2, _)) :
+ (val_factmod V <= val_factmod W)%MS = (V <= W)%MS.
+Proof.
+apply/idP/idP=> sVW; last exact: submxMr.
+by rewrite -[V]val_factmodK -[W]val_factmodK submxMr.
+Qed.
+
+Lemma val_factmod_eq0 m W : (@val_factmod m W == 0) = (W == 0).
+Proof. by rewrite -!submx0 -val_factmodS linear0 !(submx0, eqmx0). Qed.
+
+Lemma in_factmod_eq0 m (W : 'M_(m, n)) : (in_factmod W == 0) = (W <= U)%MS.
+Proof.
+rewrite submxE -!mxrank_eq0 -{2}[_ U]mulmx_base mulmxA.
+by rewrite (mxrankMfree _ (row_base_free _)).
+Qed.
+
+Lemma in_factmodK m (W : 'M_(m, n)) :
+ (W <= U^C)%MS -> val_factmod (in_factmod W) = W.
+Proof.
+case/submxP=> w ->{W}; rewrite /val_factmod /= -2!mulmxA.
+congr (_ *m _); rewrite (mulmxA (col_base _)) mulmx_base -2!mulmxA.
+by rewrite mulKVmx ?row_ebase_unit // mulmxA copid_mx_id ?rank_leq_row.
+Qed.
+
+Lemma in_factmod_addsK m (W : 'M_(m, n)) :
+ (in_factmod (U + W)%MS :=: in_factmod W)%MS.
+Proof.
+apply: eqmx_trans (addsmxMr _ _ _) _.
+by rewrite ((_ *m _ =P 0) _) ?in_factmod_eq0 //; exact: adds0mx.
+Qed.
+
+Lemma add_sub_fact_mod m (W : 'M_(m, n)) :
+ val_submod (in_submod W) + val_factmod (in_factmod W) = W.
+Proof.
+rewrite /val_submod /val_factmod /= -!mulmxA -mulmxDr.
+rewrite addrC (mulmxA (pid_mx _)) pid_mx_id // (mulmxA (col_ebase _)).
+rewrite (mulmxA _ _ (row_ebase _)) mulmx_ebase.
+rewrite (mulmxA (pid_mx _)) pid_mx_id // mulmxA -mulmxDl -mulmxDr.
+by rewrite subrK mulmx1 mulmxA mulmxKV ?row_ebase_unit.
+Qed.
+
+Lemma proj_factmodS m (W : 'M_(m, n)) :
+ (val_factmod (in_factmod W) <= U + W)%MS.
+Proof.
+by rewrite -{2}[W]add_sub_fact_mod addsmx_addKl ?val_submodP ?addsmxSr.
+Qed.
+
+Lemma in_factmodsK m (W : 'M_(m, n)) :
+ (U <= W)%MS -> (U + val_factmod (in_factmod W) :=: W)%MS.
+Proof.
+move/addsmx_idPr; apply: eqmx_trans (eqmx_sym _).
+by rewrite -{1}[W]add_sub_fact_mod; apply: addsmx_addKl; exact: val_submodP.
+Qed.
+
+Lemma mxrank_in_factmod m (W : 'M_(m, n)) :
+ (\rank (in_factmod W) + \rank U)%N = \rank (U + W).
+Proof.
+rewrite -in_factmod_addsK in_factmodE; set fU := in_factmod 1%:M.
+suffices <-: ((U + W) :&: kermx fU :=: U)%MS by rewrite mxrank_mul_ker.
+apply: eqmx_trans (capmx_idPr (addsmxSl U W)).
+apply: cap_eqmx => //; apply/eqmxP/rV_eqP => u.
+by rewrite (sameP sub_kermxP eqP) -in_factmodE in_factmod_eq0.
+Qed.
+
+Definition submod_mx of mxmodule U :=
+ fun x => in_submod (val_submod 1%:M *m rG x).
+
+Definition factmod_mx of mxmodule U :=
+ fun x => in_factmod (val_factmod 1%:M *m rG x).
+
+Hypothesis Umod : mxmodule U.
+
+Lemma in_submodJ m (W : 'M_(m, n)) x :
+ (W <= U)%MS -> in_submod (W *m rG x) = in_submod W *m submod_mx Umod x.
+Proof.
+move=> sWU; rewrite mulmxA; congr (in_submod _).
+by rewrite mulmxA -val_submodE in_submodK.
+Qed.
+
+Lemma val_submodJ m (W : 'M_(m, \rank U)) x :
+ x \in G -> val_submod (W *m submod_mx Umod x) = val_submod W *m rG x.
+Proof.
+move=> Gx; rewrite 2!(mulmxA W) -val_submodE in_submodK //.
+by rewrite mxmodule_trans ?val_submodP.
+Qed.
+
+Lemma submod_mx_repr : mx_repr G (submod_mx Umod).
+Proof.
+rewrite /submod_mx; split=> [|x y Gx Gy /=].
+ by rewrite repr_mx1 mulmx1 val_submodK.
+rewrite -in_submodJ; first by rewrite repr_mxM ?mulmxA.
+by rewrite mxmodule_trans ?val_submodP.
+Qed.
+
+Canonical submod_repr := MxRepresentation submod_mx_repr.
+
+Lemma in_factmodJ m (W : 'M_(m, n)) x :
+ x \in G -> in_factmod (W *m rG x) = in_factmod W *m factmod_mx Umod x.
+Proof.
+move=> Gx; rewrite -{1}[W]add_sub_fact_mod mulmxDl linearD /=.
+apply: (canLR (subrK _)); apply: etrans (_ : 0 = _).
+ apply/eqP; rewrite in_factmod_eq0 (submx_trans _ (mxmoduleP Umod x Gx)) //.
+ by rewrite submxMr ?val_submodP.
+by rewrite /in_factmod /val_factmod /= !mulmxA mulmx1 ?subrr.
+Qed.
+
+Lemma val_factmodJ m (W : 'M_(m, \rank (cokermx U))) x :
+ x \in G ->
+ val_factmod (W *m factmod_mx Umod x) =
+ val_factmod (in_factmod (val_factmod W *m rG x)).
+Proof. by move=> Gx; rewrite -{1}[W]val_factmodK -in_factmodJ. Qed.
+
+Lemma factmod_mx_repr : mx_repr G (factmod_mx Umod).
+Proof.
+split=> [|x y Gx Gy /=].
+ by rewrite /factmod_mx repr_mx1 mulmx1 val_factmodK.
+by rewrite -in_factmodJ // -mulmxA -repr_mxM.
+Qed.
+Canonical factmod_repr := MxRepresentation factmod_mx_repr.
+
+(* For character theory. *)
+Lemma mxtrace_sub_fact_mod x :
+ \tr (submod_repr x) + \tr (factmod_repr x) = \tr (rG x).
+Proof.
+rewrite -[submod_repr x]mulmxA mxtrace_mulC -val_submodE addrC.
+rewrite -[factmod_repr x]mulmxA mxtrace_mulC -val_factmodE addrC.
+by rewrite -mxtraceD add_sub_fact_mod.
+Qed.
+
+End Submodule.
+
+(* Properties of enveloping algebra as a subspace of 'rV_(n ^ 2). *)
+
+Lemma envelop_mx_id x : x \in G -> (rG x \in E_G)%MS.
+Proof.
+by move=> Gx; rewrite (eq_row_sub (enum_rank_in Gx x)) // rowK enum_rankK_in.
+Qed.
+
+Lemma envelop_mx1 : (1%:M \in E_G)%MS.
+Proof. by rewrite -(repr_mx1 rG) envelop_mx_id. Qed.
+
+Lemma envelop_mxP A :
+ reflect (exists a, A = \sum_(x in G) a x *: rG x) (A \in E_G)%MS.
+Proof.
+have G_1 := group1 G; have bijG := enum_val_bij_in G_1.
+set h := enum_val in bijG; have Gh: h _ \in G by exact: enum_valP.
+apply: (iffP submxP) => [[u defA] | [a ->]].
+ exists (fun x => u 0 (enum_rank_in G_1 x)); apply: (can_inj mxvecK).
+ rewrite defA mulmx_sum_row linear_sum (reindex h) //=.
+ by apply: eq_big => [i | i _]; rewrite ?Gh // rowK linearZ enum_valK_in.
+exists (\row_i a (h i)); rewrite mulmx_sum_row linear_sum (reindex h) //=.
+by apply: eq_big => [i | i _]; rewrite ?Gh // mxE rowK linearZ.
+Qed.
+
+Lemma envelop_mxM A B : (A \in E_G -> B \in E_G -> A *m B \in E_G)%MS.
+Proof.
+case/envelop_mxP=> a ->{A}; case/envelop_mxP=> b ->{B}.
+rewrite mulmx_suml !linear_sum summx_sub //= => x Gx.
+rewrite !linear_sum summx_sub //= => y Gy.
+rewrite -scalemxAl !(linearZ, scalemx_sub) //= -repr_mxM //.
+by rewrite envelop_mx_id ?groupM.
+Qed.
+
+Lemma mxmodule_envelop m1 m2 (U : 'M_(m1, n)) (W : 'M_(m2, n)) A :
+ (mxmodule U -> mxvec A <= E_G -> W <= U -> W *m A <= U)%MS.
+Proof.
+move=> modU /envelop_mxP[a ->] sWU; rewrite linear_sum summx_sub // => x Gx.
+by rewrite linearZ scalemx_sub ?mxmodule_trans.
+Qed.
+
+(* Module homomorphisms; any square matrix f defines a module homomorphism *)
+(* over some domain, namely, dom_hom_mx f. *)
+
+Definition dom_hom_mx f : 'M_n :=
+ kermx (lin1_mx (mxvec \o mulmx (cent_mx_fun E_G f) \o lin_mul_row)).
+
+Lemma hom_mxP m f (W : 'M_(m, n)) :
+ reflect (forall x, x \in G -> W *m rG x *m f = W *m f *m rG x)
+ (W <= dom_hom_mx f)%MS.
+Proof.
+apply: (iffP row_subP) => [cGf x Gx | cGf i].
+ apply/row_matrixP=> i; apply/eqP; rewrite -subr_eq0 -!mulmxA -!linearB /=.
+ have:= sub_kermxP (cGf i); rewrite mul_rV_lin1 /=.
+ move/(canRL mxvecK)/row_matrixP/(_ (enum_rank_in Gx x))/eqP; rewrite !linear0.
+ by rewrite !row_mul rowK mul_vec_lin /= mul_vec_lin_row enum_rankK_in.
+apply/sub_kermxP; rewrite mul_rV_lin1 /=; apply: (canLR vec_mxK).
+apply/row_matrixP=> j; rewrite !row_mul rowK mul_vec_lin /= mul_vec_lin_row.
+by rewrite -!row_mul mulmxBr !mulmxA cGf ?enum_valP // subrr !linear0.
+Qed.
+Implicit Arguments hom_mxP [m f W].
+
+Lemma hom_envelop_mxC m f (W : 'M_(m, n)) A :
+ (W <= dom_hom_mx f -> A \in E_G -> W *m A *m f = W *m f *m A)%MS.
+Proof.
+move/hom_mxP=> cWfG /envelop_mxP[a ->]; rewrite !linear_sum mulmx_suml.
+by apply: eq_bigr => x Gx; rewrite !linearZ -scalemxAl /= cWfG.
+Qed.
+
+Lemma dom_hom_invmx f :
+ f \in unitmx -> (dom_hom_mx (invmx f) :=: dom_hom_mx f *m f)%MS.
+Proof.
+move=> injf; set U := dom_hom_mx _; apply/eqmxP.
+rewrite -{1}[U](mulmxKV injf) submxMr; apply/hom_mxP=> x Gx.
+ by rewrite -[_ *m rG x](hom_mxP _) ?mulmxK.
+by rewrite -[_ *m rG x](hom_mxP _) ?mulmxKV.
+Qed.
+
+Lemma dom_hom_mx_module f : mxmodule (dom_hom_mx f).
+Proof.
+apply/mxmoduleP=> x Gx; apply/hom_mxP=> y Gy.
+rewrite -[_ *m rG y]mulmxA -repr_mxM // 2?(hom_mxP _) ?groupM //.
+by rewrite repr_mxM ?mulmxA.
+Qed.
+
+Lemma hom_mxmodule m (U : 'M_(m, n)) f :
+ (U <= dom_hom_mx f)%MS -> mxmodule U -> mxmodule (U *m f).
+Proof.
+move/hom_mxP=> cGfU modU; apply/mxmoduleP=> x Gx.
+by rewrite -cGfU // submxMr // (mxmoduleP modU).
+Qed.
+
+Lemma kermx_hom_module m (U : 'M_(m, n)) f :
+ (U <= dom_hom_mx f)%MS -> mxmodule U -> mxmodule (U :&: kermx f)%MS.
+Proof.
+move=> homUf modU; apply/mxmoduleP=> x Gx.
+rewrite sub_capmx mxmodule_trans ?capmxSl //=.
+apply/sub_kermxP; rewrite (hom_mxP _) ?(submx_trans (capmxSl _ _)) //.
+by rewrite (sub_kermxP (capmxSr _ _)) mul0mx.
+Qed.
+
+Lemma scalar_mx_hom a m (U : 'M_(m, n)) : (U <= dom_hom_mx a%:M)%MS.
+Proof. by apply/hom_mxP=> x Gx; rewrite -!mulmxA scalar_mxC. Qed.
+
+Lemma proj_mx_hom (U V : 'M_n) :
+ (U :&: V = 0)%MS -> mxmodule U -> mxmodule V ->
+ (U + V <= dom_hom_mx (proj_mx U V))%MS.
+Proof.
+move=> dxUV modU modV; apply/hom_mxP=> x Gx.
+rewrite -{1}(add_proj_mx dxUV (submx_refl _)) !mulmxDl addrC.
+rewrite {1}[_ *m _]proj_mx_0 ?add0r //; last first.
+ by rewrite mxmodule_trans ?proj_mx_sub.
+by rewrite [_ *m _](proj_mx_id dxUV) // mxmodule_trans ?proj_mx_sub.
+Qed.
+
+(* The subspace fixed by a subgroup H of G; it is a module if H <| G. *)
+(* The definition below is extensionally equivalent to the straightforward *)
+(* \bigcap_(x in H) kermx (rG x - 1%:M) *)
+(* but it avoids the dependency on the choice function; this allows it to *)
+(* commute with ring morphisms. *)
+
+Definition rfix_mx (H : {set gT}) :=
+ let commrH := \matrix_(i < #|H|) mxvec (rG (enum_val i) - 1%:M) in
+ kermx (lin1_mx (mxvec \o mulmx commrH \o lin_mul_row)).
+
+Lemma rfix_mxP m (W : 'M_(m, n)) (H : {set gT}) :
+ reflect (forall x, x \in H -> W *m rG x = W) (W <= rfix_mx H)%MS.
+Proof.
+rewrite /rfix_mx; set C := \matrix_i _.
+apply: (iffP row_subP) => [cHW x Hx | cHW j].
+ apply/row_matrixP=> j; apply/eqP; rewrite -subr_eq0 row_mul.
+ move/sub_kermxP: {cHW}(cHW j); rewrite mul_rV_lin1 /=; move/(canRL mxvecK).
+ move/row_matrixP/(_ (enum_rank_in Hx x)); rewrite row_mul rowK !linear0.
+ by rewrite enum_rankK_in // mul_vec_lin_row mulmxBr mulmx1 => ->.
+apply/sub_kermxP; rewrite mul_rV_lin1 /=; apply: (canLR vec_mxK).
+apply/row_matrixP=> i; rewrite row_mul rowK mul_vec_lin_row -row_mul.
+by rewrite mulmxBr mulmx1 cHW ?enum_valP // subrr !linear0.
+Qed.
+Implicit Arguments rfix_mxP [m W].
+
+Lemma rfix_mx_id (H : {set gT}) x : x \in H -> rfix_mx H *m rG x = rfix_mx H.
+Proof. exact/rfix_mxP. Qed.
+
+Lemma rfix_mxS (H K : {set gT}) : H \subset K -> (rfix_mx K <= rfix_mx H)%MS.
+Proof.
+by move=> sHK; apply/rfix_mxP=> x Hx; exact: rfix_mxP (subsetP sHK x Hx).
+Qed.
+
+Lemma rfix_mx_conjsg (H : {set gT}) x :
+ x \in G -> H \subset G -> (rfix_mx (H :^ x) :=: rfix_mx H *m rG x)%MS.
+Proof.
+move=> Gx sHG; pose rf y := rfix_mx (H :^ y).
+suffices{x Gx} IH: {in G &, forall y z, rf y *m rG z <= rf (y * z)%g}%MS.
+ apply/eqmxP; rewrite -/(rf x) -[H]conjsg1 -/(rf 1%g).
+ rewrite -{4}[x] mul1g -{1}[rf x](repr_mxKV rG Gx) -{1}(mulgV x).
+ by rewrite submxMr IH ?groupV.
+move=> x y Gx Gy; apply/rfix_mxP=> zxy; rewrite actM => /imsetP[zx Hzx ->].
+have Gzx: zx \in G by apply: subsetP Hzx; rewrite conj_subG.
+rewrite -mulmxA -repr_mxM ?groupM ?groupV // -conjgC repr_mxM // mulmxA.
+by rewrite rfix_mx_id.
+Qed.
+
+Lemma norm_sub_rstabs_rfix_mx (H : {set gT}) :
+ H \subset G -> 'N_G(H) \subset rstabs (rfix_mx H).
+Proof.
+move=> sHG; apply/subsetP=> x /setIP[Gx nHx]; rewrite inE Gx.
+apply/rfix_mxP=> y Hy; have Gy := subsetP sHG y Hy.
+have Hyx: (y ^ x^-1)%g \in H by rewrite memJ_norm ?groupV.
+rewrite -mulmxA -repr_mxM // conjgCV repr_mxM ?(subsetP sHG _ Hyx) // mulmxA.
+by rewrite (rfix_mx_id Hyx).
+Qed.
+
+Lemma normal_rfix_mx_module H : H <| G -> mxmodule (rfix_mx H).
+Proof.
+case/andP=> sHG nHG.
+by rewrite /mxmodule -{1}(setIidPl nHG) norm_sub_rstabs_rfix_mx.
+Qed.
+
+Lemma rfix_mx_module : mxmodule (rfix_mx G).
+Proof. exact: normal_rfix_mx_module. Qed.
+
+Lemma rfix_mx_rstabC (H : {set gT}) m (U : 'M[F]_(m, n)) :
+ H \subset G -> (H \subset rstab rG U) = (U <= rfix_mx H)%MS.
+Proof.
+move=> sHG; apply/subsetP/rfix_mxP=> cHU x Hx.
+ by rewrite (rstab_act (cHU x Hx)).
+by rewrite !inE (subsetP sHG) //= cHU.
+Qed.
+
+(* The cyclic module generated by a single vector. *)
+Definition cyclic_mx u := <<E_G *m lin_mul_row u>>%MS.
+
+Lemma cyclic_mxP u v :
+ reflect (exists2 A, A \in E_G & v = u *m A)%MS (v <= cyclic_mx u)%MS.
+Proof.
+rewrite genmxE; apply: (iffP submxP) => [[a] | [A /submxP[a defA]]] -> {v}.
+ exists (vec_mx (a *m E_G)); last by rewrite mulmxA mul_rV_lin1.
+ by rewrite vec_mxK submxMl.
+by exists a; rewrite mulmxA mul_rV_lin1 /= -defA mxvecK.
+Qed.
+Implicit Arguments cyclic_mxP [u v].
+
+Lemma cyclic_mx_id u : (u <= cyclic_mx u)%MS.
+Proof. by apply/cyclic_mxP; exists 1%:M; rewrite ?mulmx1 ?envelop_mx1. Qed.
+
+Lemma cyclic_mx_eq0 u : (cyclic_mx u == 0) = (u == 0).
+Proof.
+rewrite -!submx0; apply/idP/idP.
+ by apply: submx_trans; exact: cyclic_mx_id.
+move/submx0null->; rewrite genmxE; apply/row_subP=> i.
+by rewrite row_mul mul_rV_lin1 /= mul0mx ?sub0mx.
+Qed.
+
+Lemma cyclic_mx_module u : mxmodule (cyclic_mx u).
+Proof.
+apply/mxmoduleP=> x Gx; apply/row_subP=> i; rewrite row_mul.
+have [A E_A ->{i}] := @cyclic_mxP u _ (row_sub i _); rewrite -mulmxA.
+by apply/cyclic_mxP; exists (A *m rG x); rewrite ?envelop_mxM ?envelop_mx_id.
+Qed.
+
+Lemma cyclic_mx_sub m u (W : 'M_(m, n)) :
+ mxmodule W -> (u <= W)%MS -> (cyclic_mx u <= W)%MS.
+Proof.
+move=> modU Wu; rewrite genmxE; apply/row_subP=> i.
+by rewrite row_mul mul_rV_lin1 /= mxmodule_envelop // vec_mxK row_sub.
+Qed.
+
+Lemma hom_cyclic_mx u f :
+ (u <= dom_hom_mx f)%MS -> (cyclic_mx u *m f :=: cyclic_mx (u *m f))%MS.
+Proof.
+move=> domf_u; apply/eqmxP; rewrite !(eqmxMr _ (genmxE _)).
+apply/genmxP; rewrite genmx_id; congr <<_>>%MS; apply/row_matrixP=> i.
+by rewrite !row_mul !mul_rV_lin1 /= hom_envelop_mxC // vec_mxK row_sub.
+Qed.
+
+(* The annihilator of a single vector. *)
+
+Definition annihilator_mx u := (E_G :&: kermx (lin_mul_row u))%MS.
+
+Lemma annihilator_mxP u A :
+ reflect (A \in E_G /\ u *m A = 0)%MS (A \in annihilator_mx u)%MS.
+Proof.
+rewrite sub_capmx; apply: (iffP andP) => [[-> /sub_kermxP]|[-> uA0]].
+ by rewrite mul_rV_lin1 /= mxvecK.
+by split=> //; apply/sub_kermxP; rewrite mul_rV_lin1 /= mxvecK.
+Qed.
+
+(* The subspace of homomorphic images of a row vector. *)
+
+Definition row_hom_mx u :=
+ (\bigcap_j kermx (vec_mx (row j (annihilator_mx u))))%MS.
+
+Lemma row_hom_mxP u v :
+ reflect (exists2 f, u <= dom_hom_mx f & u *m f = v)%MS (v <= row_hom_mx u)%MS.
+Proof.
+apply: (iffP sub_bigcapmxP) => [iso_uv | [f hom_uf <-] i _].
+ have{iso_uv} uv0 A: (A \in E_G)%MS /\ u *m A = 0 -> v *m A = 0.
+ move/annihilator_mxP=> /submxP[a defA].
+ rewrite -[A]mxvecK {A}defA [a *m _]mulmx_sum_row !linear_sum big1 // => i _.
+ by rewrite !linearZ /= (sub_kermxP _) ?scaler0 ?iso_uv.
+ pose U := E_G *m lin_mul_row u; pose V := E_G *m lin_mul_row v.
+ pose f := pinvmx U *m V.
+ have hom_uv_f x: x \in G -> u *m rG x *m f = v *m rG x.
+ move=> Gx; apply/eqP; rewrite 2!mulmxA mul_rV_lin1 -subr_eq0 -mulmxBr.
+ rewrite uv0 // 2!linearB /= vec_mxK; split.
+ by rewrite addmx_sub ?submxMl // eqmx_opp envelop_mx_id.
+ have Uux: (u *m rG x <= U)%MS.
+ by rewrite -(genmxE U) mxmodule_trans ?cyclic_mx_id ?cyclic_mx_module.
+ by rewrite -{2}(mulmxKpV Uux) [_ *m U]mulmxA mul_rV_lin1 subrr.
+ have def_uf: u *m f = v.
+ by rewrite -[u]mulmx1 -[v]mulmx1 -(repr_mx1 rG) hom_uv_f.
+ by exists f => //; apply/hom_mxP=> x Gx; rewrite def_uf hom_uv_f.
+apply/sub_kermxP; set A := vec_mx _.
+have: (A \in annihilator_mx u)%MS by rewrite vec_mxK row_sub.
+by case/annihilator_mxP => E_A uA0; rewrite -hom_envelop_mxC // uA0 mul0mx.
+Qed.
+
+(* Sub-, isomorphic, simple, semisimple and completely reducible modules. *)
+(* All these predicates are intuitionistic (since, e.g., testing simplicity *)
+(* requires a splitting algorithm fo r the mas field). They are all *)
+(* specialized to square matrices, to avoid spurrious height parameters. *)
+
+(* Module isomorphism is an intentional property in general, but it can be *)
+(* decided when one of the two modules is known to be simple. *)
+
+CoInductive mx_iso (U V : 'M_n) : Prop :=
+ MxIso f of f \in unitmx & (U <= dom_hom_mx f)%MS & (U *m f :=: V)%MS.
+
+Lemma eqmx_iso U V : (U :=: V)%MS -> mx_iso U V.
+Proof.
+by move=> eqUV; exists 1%:M; rewrite ?unitmx1 ?scalar_mx_hom ?mulmx1.
+Qed.
+
+Lemma mx_iso_refl U : mx_iso U U.
+Proof. exact: eqmx_iso. Qed.
+
+Lemma mx_iso_sym U V : mx_iso U V -> mx_iso V U.
+Proof.
+case=> f injf homUf defV; exists (invmx f); first by rewrite unitmx_inv.
+ by rewrite dom_hom_invmx // -defV submxMr.
+by rewrite -[U](mulmxK injf); exact: eqmxMr (eqmx_sym _).
+Qed.
+
+Lemma mx_iso_trans U V W : mx_iso U V -> mx_iso V W -> mx_iso U W.
+Proof.
+case=> f injf homUf defV [g injg homVg defW].
+exists (f *m g); first by rewrite unitmx_mul injf.
+ by apply/hom_mxP=> x Gx; rewrite !mulmxA 2?(hom_mxP _) ?defV.
+by rewrite mulmxA; exact: eqmx_trans (eqmxMr g defV) defW.
+Qed.
+
+Lemma mxrank_iso U V : mx_iso U V -> \rank U = \rank V.
+Proof. by case=> f injf _ <-; rewrite mxrankMfree ?row_free_unit. Qed.
+
+Lemma mx_iso_module U V : mx_iso U V -> mxmodule U -> mxmodule V.
+Proof.
+by case=> f _ homUf defV; rewrite -(eqmx_module defV); exact: hom_mxmodule.
+Qed.
+
+(* Simple modules (we reserve the term "irreducible" for representations). *)
+
+Definition mxsimple (V : 'M_n) :=
+ [/\ mxmodule V, V != 0 &
+ forall U : 'M_n, mxmodule U -> (U <= V)%MS -> U != 0 -> (V <= U)%MS].
+
+Definition mxnonsimple (U : 'M_n) :=
+ exists V : 'M_n, [&& mxmodule V, (V <= U)%MS, V != 0 & \rank V < \rank U].
+
+Lemma mxsimpleP U :
+ [/\ mxmodule U, U != 0 & ~ mxnonsimple U] <-> mxsimple U.
+Proof.
+do [split => [] [modU nzU simU]; split] => // [V modV sVU nzV | [V]].
+ apply/idPn; rewrite -(ltn_leqif (mxrank_leqif_sup sVU)) => ltVU.
+ by case: simU; exists V; exact/and4P.
+by case/and4P=> modV sVU nzV; apply/negP; rewrite -leqNgt mxrankS ?simU.
+Qed.
+
+Lemma mxsimple_module U : mxsimple U -> mxmodule U.
+Proof. by case. Qed.
+
+Lemma mxsimple_exists m (U : 'M_(m, n)) :
+ mxmodule U -> U != 0 -> classically (exists2 V, mxsimple V & V <= U)%MS.
+Proof.
+move=> modU nzU [] // simU; move: {2}_.+1 (ltnSn (\rank U)) => r leUr.
+elim: r => // r IHr in m U leUr modU nzU simU.
+have genU := genmxE U; apply simU; exists <<U>>%MS; last by rewrite genU.
+apply/mxsimpleP; split; rewrite ?(eqmx_eq0 genU) ?(eqmx_module genU) //.
+case=> V; rewrite !genU=> /and4P[modV sVU nzV ltVU]; case: notF.
+apply: IHr nzV _ => // [|[W simW sWV]]; first exact: leq_trans ltVU _.
+by apply: simU; exists W => //; exact: submx_trans sWV sVU.
+Qed.
+
+Lemma mx_iso_simple U V : mx_iso U V -> mxsimple U -> mxsimple V.
+Proof.
+move=> isoUV [modU nzU simU]; have [f injf homUf defV] := isoUV.
+split=> [||W modW sWV nzW]; first by rewrite (mx_iso_module isoUV).
+ by rewrite -(eqmx_eq0 defV) -(mul0mx n f) (can_eq (mulmxK injf)).
+rewrite -defV -[W](mulmxKV injf) submxMr //; set W' := W *m _.
+have sW'U: (W' <= U)%MS by rewrite -[U](mulmxK injf) submxMr ?defV.
+rewrite (simU W') //; last by rewrite -(can_eq (mulmxK injf)) mul0mx mulmxKV.
+rewrite hom_mxmodule ?dom_hom_invmx // -[W](mulmxKV injf) submxMr //.
+exact: submx_trans sW'U homUf.
+Qed.
+
+Lemma mxsimple_cyclic u U :
+ mxsimple U -> u != 0 -> (u <= U)%MS -> (U :=: cyclic_mx u)%MS.
+Proof.
+case=> [modU _ simU] nz_u Uu; apply/eqmxP; set uG := cyclic_mx u.
+have s_uG_U: (uG <= U)%MS by rewrite cyclic_mx_sub.
+by rewrite simU ?cyclic_mx_eq0 ?submx_refl // cyclic_mx_module.
+Qed.
+
+(* The surjective part of Schur's lemma. *)
+Lemma mx_Schur_onto m (U : 'M_(m, n)) V f :
+ mxmodule U -> mxsimple V -> (U <= dom_hom_mx f)%MS ->
+ (U *m f <= V)%MS -> U *m f != 0 -> (U *m f :=: V)%MS.
+Proof.
+move=> modU [modV _ simV] homUf sUfV nzUf.
+apply/eqmxP; rewrite sUfV -(genmxE (U *m f)).
+rewrite simV ?(eqmx_eq0 (genmxE _)) ?genmxE //.
+by rewrite (eqmx_module (genmxE _)) hom_mxmodule.
+Qed.
+
+(* The injective part of Schur's lemma. *)
+Lemma mx_Schur_inj U f :
+ mxsimple U -> (U <= dom_hom_mx f)%MS -> U *m f != 0 -> (U :&: kermx f)%MS = 0.
+Proof.
+case=> [modU _ simU] homUf nzUf; apply/eqP; apply: contraR nzUf => nz_ker.
+rewrite (sameP eqP sub_kermxP) (sameP capmx_idPl eqmxP) simU ?capmxSl //.
+exact: kermx_hom_module.
+Qed.
+
+(* The injectve part of Schur's lemma, stated as isomorphism with the image. *)
+Lemma mx_Schur_inj_iso U f :
+ mxsimple U -> (U <= dom_hom_mx f)%MS -> U *m f != 0 -> mx_iso U (U *m f).
+Proof.
+move=> simU homUf nzUf; have [modU _ _] := simU.
+have eqUfU: \rank (U *m f) = \rank U by apply/mxrank_injP; rewrite mx_Schur_inj.
+have{eqUfU} [g invg defUf] := complete_unitmx eqUfU.
+suffices homUg: (U <= dom_hom_mx g)%MS by exists g; rewrite ?defUf.
+apply/hom_mxP=> x Gx; have [ux defUx] := submxP (mxmoduleP modU x Gx).
+by rewrite -defUf -(hom_mxP homUf) // defUx -!(mulmxA ux) defUf.
+Qed.
+
+(* The isomorphism part of Schur's lemma. *)
+Lemma mx_Schur_iso U V f :
+ mxsimple U -> mxsimple V -> (U <= dom_hom_mx f)%MS ->
+ (U *m f <= V)%MS -> U *m f != 0 -> mx_iso U V.
+Proof.
+move=> simU simV homUf sUfV nzUf; have [modU _ _] := simU.
+have [g invg homUg defUg] := mx_Schur_inj_iso simU homUf nzUf.
+exists g => //; apply: mx_Schur_onto; rewrite ?defUg //.
+by rewrite -!submx0 defUg in nzUf *.
+Qed.
+
+(* A boolean test for module isomorphism that is only valid for simple *)
+(* modules; this is the only case that matters in practice. *)
+
+Lemma nz_row_mxsimple U : mxsimple U -> nz_row U != 0.
+Proof. by case=> _ nzU _; rewrite nz_row_eq0. Qed.
+
+Definition mxsimple_iso (U V : 'M_n) :=
+ [&& mxmodule V, (V :&: row_hom_mx (nz_row U))%MS != 0 & \rank V <= \rank U].
+
+Lemma mxsimple_isoP U V :
+ mxsimple U -> reflect (mx_iso U V) (mxsimple_iso U V).
+Proof.
+move=> simU; pose u := nz_row U.
+have [Uu nz_u]: (u <= U)%MS /\ u != 0 by rewrite nz_row_sub nz_row_mxsimple.
+apply: (iffP and3P) => [[modV] | isoUV]; last first.
+ split; last by rewrite (mxrank_iso isoUV).
+ by case: (mx_iso_simple isoUV simU).
+ have [f injf homUf defV] := isoUV; apply/rowV0Pn; exists (u *m f).
+ rewrite sub_capmx -defV submxMr //.
+ by apply/row_hom_mxP; exists f; first exact: (submx_trans Uu).
+ by rewrite -(mul0mx _ f) (can_eq (mulmxK injf)) nz_u.
+case/rowV0Pn=> v; rewrite sub_capmx => /andP[Vv].
+case/row_hom_mxP => f homMf def_v nz_v eqrUV.
+pose uG := cyclic_mx u; pose vG := cyclic_mx v.
+have def_vG: (uG *m f :=: vG)%MS by rewrite /vG -def_v; exact: hom_cyclic_mx.
+have defU: (U :=: uG)%MS by exact: mxsimple_cyclic.
+have mod_uG: mxmodule uG by rewrite cyclic_mx_module.
+have homUf: (U <= dom_hom_mx f)%MS.
+ by rewrite defU cyclic_mx_sub ?dom_hom_mx_module.
+have isoUf: mx_iso U (U *m f).
+ apply: mx_Schur_inj_iso => //; apply: contra nz_v; rewrite -!submx0.
+ by rewrite (eqmxMr f defU) def_vG; exact: submx_trans (cyclic_mx_id v).
+apply: mx_iso_trans (isoUf) (eqmx_iso _); apply/eqmxP.
+have sUfV: (U *m f <= V)%MS by rewrite (eqmxMr f defU) def_vG cyclic_mx_sub.
+by rewrite -mxrank_leqif_eq ?eqn_leq 1?mxrankS // -(mxrank_iso isoUf).
+Qed.
+
+Lemma mxsimple_iso_simple U V :
+ mxsimple_iso U V -> mxsimple U -> mxsimple V.
+Proof.
+by move=> isoUV simU; apply: mx_iso_simple (simU); exact/mxsimple_isoP.
+Qed.
+
+(* For us, "semisimple" means "sum of simple modules"; this is classically, *)
+(* but not intuitionistically, equivalent to the "completely reducible" *)
+(* alternate characterization. *)
+
+Implicit Type I : finType.
+
+CoInductive mxsemisimple (V : 'M_n) :=
+ MxSemisimple I U (W := (\sum_(i : I) U i)%MS) of
+ forall i, mxsimple (U i) & (W :=: V)%MS & mxdirect W.
+
+(* This is a slight generalization of Aschbacher 12.5 for finite sets. *)
+Lemma sum_mxsimple_direct_compl m I W (U : 'M_(m, n)) :
+ let V := (\sum_(i : I) W i)%MS in
+ (forall i : I, mxsimple (W i)) -> mxmodule U -> (U <= V)%MS ->
+ {J : {set I} | let S := U + \sum_(i in J) W i in S :=: V /\ mxdirect S}%MS.
+Proof.
+move=> V simW modU sUV; pose V_ (J : {set I}) := (\sum_(i in J) W i)%MS.
+pose dxU (J : {set I}) := mxdirect (U + V_ J).
+have [J maxJ]: {J | maxset dxU J}; last case/maxsetP: maxJ => dxUVJ maxJ.
+ apply: ex_maxset; exists set0.
+ by rewrite /dxU mxdirectE /V_ /= !big_set0 addn0 addsmx0 /=.
+have modWJ: mxmodule (V_ J) by apply: sumsmx_module => i _; case: (simW i).
+exists J; split=> //; apply/eqmxP; rewrite addsmx_sub sUV; apply/andP; split.
+ by apply/sumsmx_subP=> i Ji; rewrite (sumsmx_sup i).
+rewrite -/(V_ J); apply/sumsmx_subP=> i _.
+case Ji: (i \in J).
+ by apply: submx_trans (addsmxSr _ _); exact: (sumsmx_sup i).
+have [modWi nzWi simWi] := simW i.
+rewrite (sameP capmx_idPl eqmxP) simWi ?capmxSl ?capmx_module ?addsmx_module //.
+apply: contraFT (Ji); rewrite negbK => dxWiUVJ.
+rewrite -(maxJ (i |: J)) ?setU11 ?subsetUr // /dxU.
+rewrite mxdirectE /= !big_setU1 ?Ji //=.
+rewrite addnCA addsmxA (addsmxC U) -addsmxA -mxdirectE /=.
+by rewrite mxdirect_addsE /= mxdirect_trivial -/(dxU _) dxUVJ.
+Qed.
+
+Lemma sum_mxsimple_direct_sub I W (V : 'M_n) :
+ (forall i : I, mxsimple (W i)) -> (\sum_i W i :=: V)%MS ->
+ {J : {set I} | let S := \sum_(i in J) W i in S :=: V /\ mxdirect S}%MS.
+Proof.
+move=> simW defV.
+have [|J [defS dxS]] := sum_mxsimple_direct_compl simW (mxmodule0 n).
+ exact: sub0mx.
+exists J; split; last by rewrite mxdirectE /= adds0mx mxrank0 in dxS.
+by apply: eqmx_trans defV; rewrite adds0mx_id in defS.
+Qed.
+
+Lemma mxsemisimple0 : mxsemisimple 0.
+Proof.
+exists [finType of 'I_0] (fun _ => 0); [by case | by rewrite big_ord0 | ].
+by rewrite mxdirectE /= !big_ord0 mxrank0.
+Qed.
+
+Lemma intro_mxsemisimple (I : Type) r (P : pred I) W V :
+ (\sum_(i <- r | P i) W i :=: V)%MS ->
+ (forall i, P i -> W i != 0 -> mxsimple (W i)) ->
+ mxsemisimple V.
+Proof.
+move=> defV simW; pose W_0 := [pred i | W i == 0].
+have [-> | nzV] := eqVneq V 0; first exact: mxsemisimple0.
+case def_r: r => [| i0 r'] => [|{r' def_r}].
+ by rewrite -mxrank_eq0 -defV def_r big_nil mxrank0 in nzV.
+move: defV; rewrite (bigID W_0) /= addsmxC -big_filter !(big_nth i0) !big_mkord.
+rewrite addsmxC big1 ?adds0mx_id => [|i /andP[_ /eqP] //].
+set tI := 'I_(_); set r_ := nth _ _ => defV.
+have{simW} simWr (i : tI) : mxsimple (W (r_ i)).
+ case: i => m /=; set Pr := fun i => _ => lt_m_r /=.
+ suffices: (Pr (r_ m)) by case/andP; exact: simW.
+ apply: all_nthP m lt_m_r; apply/all_filterP.
+ by rewrite -filter_predI; apply: eq_filter => i; rewrite /= andbb.
+have [J []] := sum_mxsimple_direct_sub simWr defV.
+case: (set_0Vmem J) => [-> V0 | [j0 Jj0]].
+ by rewrite -mxrank_eq0 -V0 big_set0 mxrank0 in nzV.
+pose K := {j | j \in J}; pose k0 : K := Sub j0 Jj0.
+have bij_KJ: {on J, bijective (sval : K -> _)}.
+ by exists (insubd k0) => [k _ | j Jj]; rewrite ?valKd ?insubdK.
+have J_K (k : K) : sval k \in J by exact: valP k.
+rewrite mxdirectE /= !(reindex _ bij_KJ) !(eq_bigl _ _ J_K) -mxdirectE /= -/tI.
+exact: MxSemisimple.
+Qed.
+
+Lemma mxsimple_semisimple U : mxsimple U -> mxsemisimple U.
+Proof.
+move=> simU; apply: (intro_mxsemisimple (_ : \sum_(i < 1) U :=: U))%MS => //.
+by rewrite big_ord1.
+Qed.
+
+Lemma addsmx_semisimple U V :
+ mxsemisimple U -> mxsemisimple V -> mxsemisimple (U + V)%MS.
+Proof.
+case=> [I W /= simW defU _] [J T /= simT defV _].
+have defUV: (\sum_ij sum_rect (fun _ => 'M_n) W T ij :=: U + V)%MS.
+ by rewrite big_sumType /=; exact: adds_eqmx.
+by apply: intro_mxsemisimple defUV _; case=> /=.
+Qed.
+
+Lemma sumsmx_semisimple (I : finType) (P : pred I) V :
+ (forall i, P i -> mxsemisimple (V i)) -> mxsemisimple (\sum_(i | P i) V i)%MS.
+Proof.
+move=> ssimV; elim/big_ind: _ => //; first exact: mxsemisimple0.
+exact: addsmx_semisimple.
+Qed.
+
+Lemma eqmx_semisimple U V : (U :=: V)%MS -> mxsemisimple U -> mxsemisimple V.
+Proof.
+by move=> eqUV [I W S simW defU dxS]; exists I W => //; exact: eqmx_trans eqUV.
+Qed.
+
+Lemma hom_mxsemisimple (V f : 'M_n) :
+ mxsemisimple V -> (V <= dom_hom_mx f)%MS -> mxsemisimple (V *m f).
+Proof.
+case=> I W /= simW defV _; rewrite -defV => /sumsmx_subP homWf.
+have{defV} defVf: (\sum_i W i *m f :=: V *m f)%MS.
+ by apply: eqmx_trans (eqmx_sym _) (eqmxMr f defV); exact: sumsmxMr.
+apply: (intro_mxsemisimple defVf) => i _ nzWf.
+by apply: mx_iso_simple (simW i); apply: mx_Schur_inj_iso; rewrite ?homWf.
+Qed.
+
+Lemma mxsemisimple_module U : mxsemisimple U -> mxmodule U.
+Proof.
+case=> I W /= simW defU _.
+by rewrite -(eqmx_module defU) sumsmx_module // => i _; case: (simW i).
+Qed.
+
+(* Completely reducible modules, and Maeschke's Theorem. *)
+
+CoInductive mxsplits (V U : 'M_n) :=
+ MxSplits (W : 'M_n) of mxmodule W & (U + W :=: V)%MS & mxdirect (U + W).
+
+Definition mx_completely_reducible V :=
+ forall U, mxmodule U -> (U <= V)%MS -> mxsplits V U.
+
+Lemma mx_reducibleS U V :
+ mxmodule U -> (U <= V)%MS ->
+ mx_completely_reducible V -> mx_completely_reducible U.
+Proof.
+move=> modU sUV redV U1 modU1 sU1U.
+have [W modW defV dxU1W] := redV U1 modU1 (submx_trans sU1U sUV).
+exists (W :&: U)%MS; first exact: capmx_module.
+ by apply/eqmxP; rewrite !matrix_modl // capmxSr sub_capmx defV sUV /=.
+by apply/mxdirect_addsP; rewrite capmxA (mxdirect_addsP dxU1W) cap0mx.
+Qed.
+
+Lemma mx_Maschke : [char F]^'.-group G -> mx_completely_reducible 1%:M.
+Proof.
+rewrite /pgroup charf'_nat; set nG := _%:R => nzG U => /mxmoduleP Umod _.
+pose phi := nG^-1 *: (\sum_(x in G) rG x^-1 *m pinvmx U *m U *m rG x).
+have phiG x: x \in G -> phi *m rG x = rG x *m phi.
+ move=> Gx; rewrite -scalemxAl -scalemxAr; congr (_ *: _).
+ rewrite {2}(reindex_acts 'R _ Gx) ?astabsR //= mulmx_suml mulmx_sumr.
+ apply: eq_bigr => y Gy; rewrite !mulmxA -repr_mxM ?groupV ?groupM //.
+ by rewrite invMg mulKVg repr_mxM ?mulmxA.
+have Uphi: U *m phi = U.
+ rewrite -scalemxAr mulmx_sumr (eq_bigr (fun _ => U)) => [|x Gx].
+ by rewrite sumr_const -scaler_nat !scalerA mulVf ?scale1r.
+ by rewrite 3!mulmxA mulmxKpV ?repr_mxKV ?Umod ?groupV.
+have tiUker: (U :&: kermx phi = 0)%MS.
+ apply/eqP/rowV0P=> v; rewrite sub_capmx => /andP[/submxP[u ->] /sub_kermxP].
+ by rewrite -mulmxA Uphi.
+exists (kermx phi); last exact/mxdirect_addsP.
+ apply/mxmoduleP=> x Gx; apply/sub_kermxP.
+ by rewrite -mulmxA -phiG // mulmxA mulmx_ker mul0mx.
+apply/eqmxP; rewrite submx1 sub1mx.
+rewrite /row_full mxrank_disjoint_sum //= mxrank_ker.
+suffices ->: (U :=: phi)%MS by rewrite subnKC ?rank_leq_row.
+apply/eqmxP; rewrite -{1}Uphi submxMl scalemx_sub //.
+by rewrite summx_sub // => x Gx; rewrite -mulmxA mulmx_sub ?Umod.
+Qed.
+
+Lemma mxsemisimple_reducible V : mxsemisimple V -> mx_completely_reducible V.
+Proof.
+case=> [I W /= simW defV _] U modU sUV; rewrite -defV in sUV.
+have [J [defV' dxV]] := sum_mxsimple_direct_compl simW modU sUV.
+exists (\sum_(i in J) W i)%MS.
+- by apply: sumsmx_module => i _; case: (simW i).
+- exact: eqmx_trans defV' defV.
+by rewrite mxdirect_addsE (sameP eqP mxdirect_addsP) /= in dxV; case/and3P: dxV.
+Qed.
+
+Lemma mx_reducible_semisimple V :
+ mxmodule V -> mx_completely_reducible V -> classically (mxsemisimple V).
+Proof.
+move=> modV redV [] // nssimV; move: {-1}_.+1 (ltnSn (\rank V)) => r leVr.
+elim: r => // r IHr in V leVr modV redV nssimV.
+have [V0 | nzV] := eqVneq V 0.
+ by rewrite nssimV ?V0 //; exact: mxsemisimple0.
+apply (mxsimple_exists modV nzV) => [[U simU sUV]]; have [modU nzU _] := simU.
+have [W modW defUW dxUW] := redV U modU sUV.
+have sWV: (W <= V)%MS by rewrite -defUW addsmxSr.
+apply: IHr (mx_reducibleS modW sWV redV) _ => // [|ssimW].
+ rewrite ltnS -defUW (mxdirectP dxUW) /= in leVr; apply: leq_trans leVr.
+ by rewrite -add1n leq_add2r lt0n mxrank_eq0.
+apply: nssimV (eqmx_semisimple defUW (addsmx_semisimple _ ssimW)).
+exact: mxsimple_semisimple.
+Qed.
+
+Lemma mxsemisimpleS U V :
+ mxmodule U -> (U <= V)%MS -> mxsemisimple V -> mxsemisimple U.
+Proof.
+move=> modU sUV ssimV.
+have [W modW defUW dxUW]:= mxsemisimple_reducible ssimV modU sUV.
+move/mxdirect_addsP: dxUW => dxUW.
+have defU : (V *m proj_mx U W :=: U)%MS.
+ by apply/eqmxP; rewrite proj_mx_sub -{1}[U](proj_mx_id dxUW) ?submxMr.
+apply: eqmx_semisimple defU _; apply: hom_mxsemisimple ssimV _.
+by rewrite -defUW proj_mx_hom.
+Qed.
+
+Lemma hom_mxsemisimple_iso I P U W f :
+ let V := (\sum_(i : I | P i) W i)%MS in
+ mxsimple U -> (forall i, P i -> W i != 0 -> mxsimple (W i)) ->
+ (V <= dom_hom_mx f)%MS -> (U <= V *m f)%MS ->
+ {i | P i & mx_iso (W i) U}.
+Proof.
+move=> V simU simW homVf sUVf; have [modU nzU _] := simU.
+have ssimVf: mxsemisimple (V *m f).
+ exact: hom_mxsemisimple (intro_mxsemisimple (eqmx_refl V) simW) homVf.
+have [U' modU' defVf] := mxsemisimple_reducible ssimVf modU sUVf.
+move/mxdirect_addsP=> dxUU'; pose p := f *m proj_mx U U'.
+case: (pickP (fun i => P i && (W i *m p != 0))) => [i /andP[Pi nzWip] | no_i].
+ have sWiV: (W i <= V)%MS by rewrite (sumsmx_sup i).
+ have sWipU: (W i *m p <= U)%MS by rewrite mulmxA proj_mx_sub.
+ exists i => //; apply: (mx_Schur_iso (simW i Pi _) simU _ sWipU nzWip).
+ by apply: contraNneq nzWip => ->; rewrite mul0mx.
+ apply: (submx_trans sWiV); apply/hom_mxP=> x Gx.
+ by rewrite mulmxA [_ *m p]mulmxA 2?(hom_mxP _) -?defVf ?proj_mx_hom.
+case/negP: nzU; rewrite -submx0 -[U](proj_mx_id dxUU') //.
+rewrite (submx_trans (submxMr _ sUVf)) // -mulmxA -/p sumsmxMr.
+by apply/sumsmx_subP=> i Pi; move/negbT: (no_i i); rewrite Pi negbK submx0.
+Qed.
+
+(* The component associated to a given irreducible module. *)
+
+Section Components.
+
+Fact component_mx_key : unit. Proof. by []. Qed.
+Definition component_mx_expr (U : 'M[F]_n) :=
+ (\sum_i cyclic_mx (row i (row_hom_mx (nz_row U))))%MS.
+Definition component_mx := locked_with component_mx_key component_mx_expr.
+Canonical component_mx_unfoldable := [unlockable fun component_mx].
+
+Variable U : 'M[F]_n.
+Hypothesis simU : mxsimple U.
+
+Let u := nz_row U.
+Let iso_u := row_hom_mx u.
+Let nz_u : u != 0 := nz_row_mxsimple simU.
+Let Uu : (u <= U)%MS := nz_row_sub U.
+Let defU : (U :=: cyclic_mx u)%MS := mxsimple_cyclic simU nz_u Uu.
+Local Notation compU := (component_mx U).
+
+Lemma component_mx_module : mxmodule compU.
+Proof.
+by rewrite unlock sumsmx_module // => i; rewrite cyclic_mx_module.
+Qed.
+
+Lemma genmx_component : <<compU>>%MS = compU.
+Proof.
+by rewrite [in compU]unlock genmx_sums; apply: eq_bigr => i; rewrite genmx_id.
+Qed.
+
+Lemma component_mx_def : {I : finType & {W : I -> 'M_n |
+ forall i, mx_iso U (W i) & compU = \sum_i W i}}%MS.
+Proof.
+pose r i := row i iso_u; pose r_nz i := r i != 0; pose I := {i | r_nz i}.
+exists [finType of I]; exists (fun i => cyclic_mx (r (sval i))) => [i|].
+ apply/mxsimple_isoP=> //; apply/and3P.
+ split; first by rewrite cyclic_mx_module.
+ apply/rowV0Pn; exists (r (sval i)); last exact: (svalP i).
+ by rewrite sub_capmx cyclic_mx_id row_sub.
+ have [f hom_u_f <-] := @row_hom_mxP u (r (sval i)) (row_sub _ _).
+ by rewrite defU -hom_cyclic_mx ?mxrankM_maxl.
+rewrite -(eq_bigr _ (fun _ _ => genmx_id _)) -genmx_sums -genmx_component.
+rewrite [in compU]unlock; apply/genmxP/andP; split; last first.
+ by apply/sumsmx_subP => i _; rewrite (sumsmx_sup (sval i)).
+apply/sumsmx_subP => i _.
+case i0: (r_nz i); first by rewrite (sumsmx_sup (Sub i i0)).
+by move/negbFE: i0; rewrite -cyclic_mx_eq0 => /eqP->; exact: sub0mx.
+Qed.
+
+Lemma component_mx_semisimple : mxsemisimple compU.
+Proof.
+have [I [W isoUW ->]] := component_mx_def.
+apply: intro_mxsemisimple (eqmx_refl _) _ => i _ _.
+exact: mx_iso_simple (isoUW i) simU.
+Qed.
+
+Lemma mx_iso_component V : mx_iso U V -> (V <= compU)%MS.
+Proof.
+move=> isoUV; have [f injf homUf defV] := isoUV.
+have simV := mx_iso_simple isoUV simU.
+have hom_u_f := submx_trans Uu homUf.
+have ->: (V :=: cyclic_mx (u *m f))%MS.
+ apply: eqmx_trans (hom_cyclic_mx hom_u_f).
+ exact: eqmx_trans (eqmx_sym defV) (eqmxMr _ defU).
+have iso_uf: (u *m f <= iso_u)%MS by apply/row_hom_mxP; exists f.
+rewrite genmxE; apply/row_subP=> j; rewrite row_mul mul_rV_lin1 /=.
+set a := vec_mx _; apply: submx_trans (submxMr _ iso_uf) _.
+apply/row_subP=> i; rewrite row_mul [in compU]unlock (sumsmx_sup i) //.
+by apply/cyclic_mxP; exists a; rewrite // vec_mxK row_sub.
+Qed.
+
+Lemma component_mx_id : (U <= compU)%MS.
+Proof. exact: mx_iso_component (mx_iso_refl U). Qed.
+
+Lemma hom_component_mx_iso f V :
+ mxsimple V -> (compU <= dom_hom_mx f)%MS -> (V <= compU *m f)%MS ->
+ mx_iso U V.
+Proof.
+have [I [W isoUW ->]] := component_mx_def => simV homWf sVWf.
+have [i _ _|i _ ] := hom_mxsemisimple_iso simV _ homWf sVWf.
+ exact: mx_iso_simple (simU).
+exact: mx_iso_trans.
+Qed.
+
+Lemma component_mx_iso V : mxsimple V -> (V <= compU)%MS -> mx_iso U V.
+Proof.
+move=> simV; rewrite -[compU]mulmx1.
+exact: hom_component_mx_iso (scalar_mx_hom _ _).
+Qed.
+
+Lemma hom_component_mx f :
+ (compU <= dom_hom_mx f)%MS -> (compU *m f <= compU)%MS.
+Proof.
+move=> hom_f.
+have [I W /= simW defW _] := hom_mxsemisimple component_mx_semisimple hom_f.
+rewrite -defW; apply/sumsmx_subP=> i _; apply: mx_iso_component.
+by apply: hom_component_mx_iso hom_f _ => //; rewrite -defW (sumsmx_sup i).
+Qed.
+
+End Components.
+
+Lemma component_mx_isoP U V :
+ mxsimple U -> mxsimple V ->
+ reflect (mx_iso U V) (component_mx U == component_mx V).
+Proof.
+move=> simU simV; apply: (iffP eqP) => isoUV.
+ by apply: component_mx_iso; rewrite ?isoUV ?component_mx_id.
+rewrite -(genmx_component U) -(genmx_component V); apply/genmxP.
+wlog suffices: U V simU simV isoUV / (component_mx U <= component_mx V)%MS.
+ by move=> IH; rewrite !IH //; exact: mx_iso_sym.
+have [I [W isoWU ->]] := component_mx_def simU.
+apply/sumsmx_subP => i _; apply: mx_iso_component => //.
+exact: mx_iso_trans (mx_iso_sym isoUV) (isoWU i).
+Qed.
+
+Lemma component_mx_disjoint U V :
+ mxsimple U -> mxsimple V -> component_mx U != component_mx V ->
+ (component_mx U :&: component_mx V = 0)%MS.
+Proof.
+move=> simU simV neUV; apply: contraNeq neUV => ntUV.
+apply: (mxsimple_exists _ ntUV) => [|[W simW]].
+ by rewrite capmx_module ?component_mx_module.
+rewrite sub_capmx => /andP[sWU sWV]; apply/component_mx_isoP=> //.
+by apply: mx_iso_trans (_ : mx_iso U W) (mx_iso_sym _); exact: component_mx_iso.
+Qed.
+
+Section Socle.
+
+Record socleType := EnumSocle {
+ socle_base_enum : seq 'M[F]_n;
+ _ : forall M, M \in socle_base_enum -> mxsimple M;
+ _ : forall M, mxsimple M -> has (mxsimple_iso M) socle_base_enum
+}.
+
+Lemma socle_exists : classically socleType.
+Proof.
+pose V : 'M[F]_n := 0; have: mxsemisimple V by exact: mxsemisimple0.
+have: n - \rank V < n.+1 by rewrite mxrank0 subn0.
+elim: _.+1 V => // n' IHn' V; rewrite ltnS => le_nV_n' ssimV.
+case=> // maxV; apply: (maxV); have [I /= U simU defV _] := ssimV.
+exists (codom U) => [M | M simM]; first by case/mapP=> i _ ->.
+suffices sMV: (M <= V)%MS.
+ rewrite -defV -(mulmx1 (\sum_i _)%MS) in sMV.
+ have [//| i _] := hom_mxsemisimple_iso simM _ (scalar_mx_hom _ _) sMV.
+ move/mx_iso_sym=> isoM; apply/hasP.
+ exists (U i); [exact: codom_f | exact/mxsimple_isoP].
+have ssimMV := addsmx_semisimple (mxsimple_semisimple simM) ssimV.
+apply: contraLR isT => nsMV; apply: IHn' ssimMV _ maxV.
+apply: leq_trans le_nV_n'; rewrite ltn_sub2l //.
+ rewrite ltn_neqAle rank_leq_row andbT -[_ == _]sub1mx.
+ apply: contra nsMV; apply: submx_trans; exact: submx1.
+rewrite (ltn_leqif (mxrank_leqif_sup _)) ?addsmxSr //.
+by rewrite addsmx_sub submx_refl andbT.
+Qed.
+
+Section SocleDef.
+
+Variable sG0 : socleType.
+
+Definition socle_enum := map component_mx (socle_base_enum sG0).
+
+Lemma component_socle M : mxsimple M -> component_mx M \in socle_enum.
+Proof.
+rewrite /socle_enum; case: sG0 => e0 /= sim_e mem_e simM.
+have /hasP[M' e0M' isoMM'] := mem_e M simM; apply/mapP; exists M' => //.
+by apply/eqP/component_mx_isoP; [|exact: sim_e | exact/mxsimple_isoP].
+Qed.
+
+Inductive socle_sort : predArgType := PackSocle W of W \in socle_enum.
+
+Local Notation sG := socle_sort.
+Local Notation e0 := (socle_base_enum sG0).
+
+Definition socle_base W := let: PackSocle W _ := W in e0`_(index W socle_enum).
+
+Coercion socle_val W : 'M[F]_n := component_mx (socle_base W).
+
+Definition socle_mult (W : sG) := (\rank W %/ \rank (socle_base W))%N.
+
+Lemma socle_simple W : mxsimple (socle_base W).
+Proof.
+case: W => M /=; rewrite /= /socle_enum /=; case: sG0 => e sim_e _ /= e_M.
+by apply: sim_e; rewrite mem_nth // -(size_map component_mx) index_mem.
+Qed.
+
+Definition socle_module (W : sG) := mxsimple_module (socle_simple W).
+
+Definition socle_repr W := submod_repr (socle_module W).
+
+Lemma nz_socle (W : sG) : W != 0 :> 'M_n.
+Proof.
+have simW := socle_simple W; have [_ nzW _] := simW; apply: contra nzW.
+by rewrite -!submx0; exact: submx_trans (component_mx_id simW).
+Qed.
+
+Lemma socle_mem (W : sG) : (W : 'M_n) \in socle_enum.
+Proof. exact: component_socle (socle_simple _). Qed.
+
+Lemma PackSocleK W e0W : @PackSocle W e0W = W :> 'M_n.
+Proof.
+rewrite /socle_val /= in e0W *; rewrite -(nth_map _ 0) ?nth_index //.
+by rewrite -(size_map component_mx) index_mem.
+Qed.
+
+Canonical socle_subType := SubType _ _ _ socle_sort_rect PackSocleK.
+Definition socle_eqMixin := Eval hnf in [eqMixin of sG by <:].
+Canonical socle_eqType := Eval hnf in EqType sG socle_eqMixin.
+Definition socle_choiceMixin := Eval hnf in [choiceMixin of sG by <:].
+Canonical socle_choiceType := ChoiceType sG socle_choiceMixin.
+
+Lemma socleP (W W' : sG) : reflect (W = W') (W == W')%MS.
+Proof. by rewrite (sameP genmxP eqP) !{1}genmx_component; exact: (W =P _). Qed.
+
+Fact socle_finType_subproof :
+ cancel (fun W => SeqSub (socle_mem W)) (fun s => PackSocle (valP s)).
+Proof. by move=> W /=; apply: val_inj; rewrite /= PackSocleK. Qed.
+
+Definition socle_countMixin := CanCountMixin socle_finType_subproof.
+Canonical socle_countType := CountType sG socle_countMixin.
+Canonical socle_subCountType := [subCountType of sG].
+Definition socle_finMixin := CanFinMixin socle_finType_subproof.
+Canonical socle_finType := FinType sG socle_finMixin.
+Canonical socle_subFinType := [subFinType of sG].
+
+End SocleDef.
+
+Coercion socle_sort : socleType >-> predArgType.
+
+Variable sG : socleType.
+
+Section SubSocle.
+
+Variable P : pred sG.
+Notation S := (\sum_(W : sG | P W) socle_val W)%MS.
+
+Lemma subSocle_module : mxmodule S.
+Proof. by rewrite sumsmx_module // => W _; exact: component_mx_module. Qed.
+
+Lemma subSocle_semisimple : mxsemisimple S.
+Proof.
+apply: sumsmx_semisimple => W _; apply: component_mx_semisimple.
+exact: socle_simple.
+Qed.
+Local Notation ssimS := subSocle_semisimple.
+
+Lemma subSocle_iso M :
+ mxsimple M -> (M <= S)%MS -> {W : sG | P W & mx_iso (socle_base W) M}.
+Proof.
+move=> simM sMS; have [modM nzM _] := simM.
+have [V /= modV defMV] := mxsemisimple_reducible ssimS modM sMS.
+move/mxdirect_addsP=> dxMV; pose p := proj_mx M V; pose Sp (W : sG) := W *m p.
+case: (pickP [pred i | P i & Sp i != 0]) => [/= W | Sp0]; last first.
+ case/negP: nzM; rewrite -submx0 -[M](proj_mx_id dxMV) //.
+ rewrite (submx_trans (submxMr _ sMS)) // sumsmxMr big1 // => W P_W.
+ by apply/eqP; move/negbT: (Sp0 W); rewrite /= P_W negbK.
+rewrite {}/Sp /= => /andP[P_W nzSp]; exists W => //.
+have homWp: (W <= dom_hom_mx p)%MS.
+ apply: submx_trans (proj_mx_hom dxMV modM modV).
+ by rewrite defMV (sumsmx_sup W).
+have simWP := socle_simple W; apply: hom_component_mx_iso (homWp) _ => //.
+by rewrite (mx_Schur_onto _ simM) ?proj_mx_sub ?component_mx_module.
+Qed.
+
+Lemma capmx_subSocle m (M : 'M_(m, n)) :
+ mxmodule M -> (M :&: S :=: \sum_(W : sG | P W) (M :&: W))%MS.
+Proof.
+move=> modM; apply/eqmxP/andP; split; last first.
+ by apply/sumsmx_subP=> W P_W; rewrite capmxS // (sumsmx_sup W).
+have modMS: mxmodule (M :&: S)%MS by rewrite capmx_module ?subSocle_module.
+have [J /= U simU defMS _] := mxsemisimpleS modMS (capmxSr M S) ssimS.
+rewrite -defMS; apply/sumsmx_subP=> j _.
+have [sUjV sUjS]: (U j <= M /\ U j <= S)%MS.
+ by apply/andP; rewrite -sub_capmx -defMS (sumsmx_sup j).
+have [W P_W isoWU] := subSocle_iso (simU j) sUjS.
+rewrite (sumsmx_sup W) // sub_capmx sUjV mx_iso_component //.
+exact: socle_simple.
+Qed.
+
+End SubSocle.
+
+Lemma subSocle_direct P : mxdirect (\sum_(W : sG | P W) W).
+Proof.
+apply/mxdirect_sumsP=> W _; apply/eqP.
+rewrite -submx0 capmx_subSocle ?component_mx_module //.
+apply/sumsmx_subP=> W' /andP[_ neWW'].
+by rewrite capmxC component_mx_disjoint //; exact: socle_simple.
+Qed.
+
+Definition Socle := (\sum_(W : sG) W)%MS.
+
+Lemma simple_Socle M : mxsimple M -> (M <= Socle)%MS.
+Proof.
+move=> simM; have socM := component_socle sG simM.
+by rewrite (sumsmx_sup (PackSocle socM)) // PackSocleK component_mx_id.
+Qed.
+
+Lemma semisimple_Socle U : mxsemisimple U -> (U <= Socle)%MS.
+Proof.
+by case=> I M /= simM <- _; apply/sumsmx_subP=> i _; exact: simple_Socle.
+Qed.
+
+Lemma reducible_Socle U :
+ mxmodule U -> mx_completely_reducible U -> (U <= Socle)%MS.
+Proof.
+move=> modU redU; apply: (mx_reducible_semisimple modU redU).
+exact: semisimple_Socle.
+Qed.
+
+Lemma genmx_Socle : <<Socle>>%MS = Socle.
+Proof. by rewrite genmx_sums; apply: eq_bigr => W; rewrite genmx_component. Qed.
+
+Lemma reducible_Socle1 : mx_completely_reducible 1%:M -> Socle = 1%:M.
+Proof.
+move=> redG; rewrite -genmx1 -genmx_Socle; apply/genmxP.
+by rewrite submx1 reducible_Socle ?mxmodule1.
+Qed.
+
+Lemma Socle_module : mxmodule Socle. Proof. exact: subSocle_module. Qed.
+
+Lemma Socle_semisimple : mxsemisimple Socle.
+Proof. exact: subSocle_semisimple. Qed.
+
+Lemma Socle_direct : mxdirect Socle. Proof. exact: subSocle_direct. Qed.
+
+Lemma Socle_iso M : mxsimple M -> {W : sG | mx_iso (socle_base W) M}.
+Proof.
+by move=> simM; case/subSocle_iso: (simple_Socle simM) => // W _; exists W.
+Qed.
+
+End Socle.
+
+(* Centralizer subgroup and central homomorphisms. *)
+Section CentHom.
+
+Variable f : 'M[F]_n.
+
+Lemma row_full_dom_hom : row_full (dom_hom_mx f) = centgmx rG f.
+Proof.
+by rewrite -sub1mx; apply/hom_mxP/centgmxP=> cfG x /cfG; rewrite !mul1mx.
+Qed.
+
+Lemma memmx_cent_envelop : (f \in 'C(E_G))%MS = centgmx rG f.
+Proof.
+apply/cent_rowP/centgmxP=> [cfG x Gx | cfG i].
+ by have:= cfG (enum_rank_in Gx x); rewrite rowK mxvecK enum_rankK_in.
+by rewrite rowK mxvecK /= cfG ?enum_valP.
+Qed.
+
+Lemma kermx_centg_module : centgmx rG f -> mxmodule (kermx f).
+Proof.
+move/centgmxP=> cGf; apply/mxmoduleP=> x Gx; apply/sub_kermxP.
+by rewrite -mulmxA -cGf // mulmxA mulmx_ker mul0mx.
+Qed.
+
+Lemma centgmx_hom m (U : 'M_(m, n)) : centgmx rG f -> (U <= dom_hom_mx f)%MS.
+Proof. by rewrite -row_full_dom_hom -sub1mx; exact: submx_trans (submx1 _). Qed.
+
+End CentHom.
+
+(* (Globally) irreducible, and absolutely irreducible representations. Note *)
+(* that unlike "reducible", "absolutely irreducible" can easily be decided. *)
+
+Definition mx_irreducible := mxsimple 1%:M.
+
+Lemma mx_irrP :
+ mx_irreducible <-> n > 0 /\ (forall U, @mxmodule n U -> U != 0 -> row_full U).
+Proof.
+rewrite /mx_irreducible /mxsimple mxmodule1 -mxrank_eq0 mxrank1 -lt0n.
+do [split=> [[_ -> irrG] | [-> irrG]]; split=> // U] => [modU | modU _] nzU.
+ by rewrite -sub1mx (irrG U) ?submx1.
+by rewrite sub1mx irrG.
+Qed.
+
+(* Schur's lemma for endomorphisms. *)
+Lemma mx_Schur :
+ mx_irreducible -> forall f, centgmx rG f -> f != 0 -> f \in unitmx.
+Proof.
+move/mx_Schur_onto=> irrG f.
+rewrite -row_full_dom_hom -!row_full_unit -!sub1mx => cGf nz.
+by rewrite -[f]mul1mx irrG ?submx1 ?mxmodule1 ?mul1mx.
+Qed.
+
+Definition mx_absolutely_irreducible := (n > 0) && row_full E_G.
+
+Lemma mx_abs_irrP :
+ reflect (n > 0 /\ exists a_, forall A, A = \sum_(x in G) a_ x A *: rG x)
+ mx_absolutely_irreducible.
+Proof.
+have G_1 := group1 G; have bijG := enum_val_bij_in G_1.
+set h := enum_val in bijG; have Gh : h _ \in G by exact: enum_valP.
+rewrite /mx_absolutely_irreducible; case: (n > 0); last by right; case.
+apply: (iffP row_fullP) => [[E' E'G] | [_ [a_ a_G]]].
+ split=> //; exists (fun x B => (mxvec B *m E') 0 (enum_rank_in G_1 x)) => B.
+ apply: (can_inj mxvecK); rewrite -{1}[mxvec B]mulmx1 -{}E'G mulmxA.
+ move: {B E'}(_ *m E') => u; apply/rowP=> j.
+ rewrite linear_sum (reindex h) //= mxE summxE.
+ by apply: eq_big => [k| k _]; rewrite ?Gh // enum_valK_in mxE linearZ !mxE.
+exists (\matrix_(j, i) a_ (h i) (vec_mx (row j 1%:M))).
+apply/row_matrixP=> i; rewrite -[row i 1%:M]vec_mxK {}[vec_mx _]a_G.
+apply/rowP=> j; rewrite linear_sum (reindex h) //= 2!mxE summxE.
+by apply: eq_big => [k| k _]; [rewrite Gh | rewrite linearZ !mxE].
+Qed.
+
+Lemma mx_abs_irr_cent_scalar :
+ mx_absolutely_irreducible -> forall A, centgmx rG A -> is_scalar_mx A.
+Proof.
+case/mx_abs_irrP=> n_gt0 [a_ a_G] A /centgmxP cGA.
+have{cGA a_G} cMA B: A *m B = B *m A.
+ rewrite {}[B]a_G mulmx_suml mulmx_sumr.
+ by apply: eq_bigr => x Gx; rewrite -scalemxAl -scalemxAr cGA.
+pose i0 := Ordinal n_gt0; apply/is_scalar_mxP; exists (A i0 i0).
+apply/matrixP=> i j; move/matrixP/(_ i0 j): (esym (cMA (delta_mx i0 i))).
+rewrite -[A *m _]trmxK trmx_mul trmx_delta -!(@mul_delta_mx _ n 1 n 0) -!mulmxA.
+by rewrite -!rowE !mxE !big_ord1 !mxE !eqxx !mulr_natl /= andbT eq_sym.
+Qed.
+
+Lemma mx_abs_irrW : mx_absolutely_irreducible -> mx_irreducible.
+Proof.
+case/mx_abs_irrP=> n_gt0 [a_ a_G]; apply/mx_irrP; split=> // U Umod.
+case/rowV0Pn=> u Uu; rewrite -mxrank_eq0 -lt0n row_leq_rank -sub1mx.
+case/submxP: Uu => v ->{u} /row_freeP[u' vK]; apply/row_subP=> i.
+rewrite rowE scalar_mxC -{}vK -2![_ *m _]mulmxA; move: {u' i}(u' *m _) => A.
+rewrite mulmx_sub {v}// [A]a_G linear_sum summx_sub //= => x Gx.
+by rewrite linearZ /= scalemx_sub // (mxmoduleP Umod).
+Qed.
+
+Lemma linear_mx_abs_irr : n = 1%N -> mx_absolutely_irreducible.
+Proof.
+move=> n1; rewrite /mx_absolutely_irreducible /row_full eqn_leq rank_leq_col.
+rewrite {1 2 3}n1 /= lt0n mxrank_eq0; apply: contraTneq envelop_mx1 => ->.
+by rewrite eqmx0 submx0 mxvec_eq0 -mxrank_eq0 mxrank1 n1.
+Qed.
+
+Lemma abelian_abs_irr : abelian G -> mx_absolutely_irreducible = (n == 1%N).
+Proof.
+move=> cGG; apply/idP/eqP=> [absG|]; last exact: linear_mx_abs_irr.
+have [n_gt0 _] := andP absG.
+pose M := <<delta_mx 0 (Ordinal n_gt0) : 'rV[F]_n>>%MS.
+have rM: \rank M = 1%N by rewrite genmxE mxrank_delta.
+suffices defM: (M == 1%:M)%MS by rewrite (eqmxP defM) mxrank1 in rM.
+case: (mx_abs_irrW absG) => _ _ ->; rewrite ?submx1 -?mxrank_eq0 ?rM //.
+apply/mxmoduleP=> x Gx; suffices: is_scalar_mx (rG x).
+ by case/is_scalar_mxP=> a ->; rewrite mul_mx_scalar scalemx_sub.
+apply: (mx_abs_irr_cent_scalar absG).
+by apply/centgmxP=> y Gy; rewrite -!repr_mxM // (centsP cGG).
+Qed.
+
+End OneRepresentation.
+
+Implicit Arguments mxmoduleP [gT G n rG m U].
+Implicit Arguments envelop_mxP [gT G n rG A].
+Implicit Arguments hom_mxP [gT G n rG m f W].
+Implicit Arguments rfix_mxP [gT G n rG m W].
+Implicit Arguments cyclic_mxP [gT G n rG u v].
+Implicit Arguments annihilator_mxP [gT G n rG u A].
+Implicit Arguments row_hom_mxP [gT G n rG u v].
+Implicit Arguments mxsimple_isoP [gT G n rG U V].
+Implicit Arguments socleP [gT G n rG sG0 W W'].
+Implicit Arguments mx_abs_irrP [gT G n rG].
+
+Implicit Arguments val_submod_inj [n U m].
+Implicit Arguments val_factmod_inj [n U m].
+
+Prenex Implicits val_submod_inj val_factmod_inj.
+
+Section Proper.
+
+Variables (gT : finGroupType) (G : {group gT}) (n' : nat).
+Local Notation n := n'.+1.
+Variable rG : mx_representation F G n.
+
+Lemma envelop_mx_ring : mxring (enveloping_algebra_mx rG).
+Proof.
+apply/andP; split; first by apply/mulsmx_subP; exact: envelop_mxM.
+apply/mxring_idP; exists 1%:M; split=> *; rewrite ?mulmx1 ?mul1mx //.
+ by rewrite -mxrank_eq0 mxrank1.
+exact: envelop_mx1.
+Qed.
+
+End Proper.
+
+Section JacobsonDensity.
+
+Variables (gT : finGroupType) (G : {group gT}) (n : nat).
+Variable rG : mx_representation F G n.
+Hypothesis irrG : mx_irreducible rG.
+
+Local Notation E_G := (enveloping_algebra_mx rG).
+Local Notation Hom_G := 'C(E_G)%MS.
+
+Lemma mx_Jacobson_density : ('C(Hom_G) <= E_G)%MS.
+Proof.
+apply/row_subP=> iB; rewrite -[row iB _]vec_mxK; move defB: (vec_mx _) => B.
+have{defB} cBcE: (B \in 'C(Hom_G))%MS by rewrite -defB vec_mxK row_sub.
+have rGnP: mx_repr G (fun x => lin_mx (mulmxr (rG x)) : 'A_n).
+ split=> [|x y Gx Gy]; apply/row_matrixP=> i.
+ by rewrite !rowE mul_rV_lin repr_mx1 /= !mulmx1 vec_mxK.
+ by rewrite !rowE mulmxA !mul_rV_lin repr_mxM //= mxvecK mulmxA.
+move def_rGn: (MxRepresentation rGnP) => rGn.
+pose E_Gn := enveloping_algebra_mx rGn.
+pose e1 : 'rV[F]_(n ^ 2) := mxvec 1%:M; pose U := cyclic_mx rGn e1.
+have U_e1: (e1 <= U)%MS by rewrite cyclic_mx_id.
+have modU: mxmodule rGn U by rewrite cyclic_mx_module.
+pose Bn : 'M_(n ^ 2) := lin_mx (mulmxr B).
+suffices U_e1Bn: (e1 *m Bn <= U)%MS.
+ rewrite mul_vec_lin /= mul1mx in U_e1Bn; apply: submx_trans U_e1Bn _.
+ rewrite genmxE; apply/row_subP=> i; rewrite row_mul rowK mul_vec_lin_row.
+ by rewrite -def_rGn mul_vec_lin /= mul1mx (eq_row_sub i) ?rowK.
+have{cBcE} cBncEn A: centgmx rGn A -> A *m Bn = Bn *m A.
+ rewrite -def_rGn => cAG; apply/row_matrixP; case/mxvec_indexP=> j k /=.
+ rewrite !rowE !mulmxA -mxvec_delta -(mul_delta_mx (0 : 'I_1)).
+ rewrite mul_rV_lin mul_vec_lin /= -mulmxA; apply: (canLR vec_mxK).
+ apply/row_matrixP=> i; set dj0 := delta_mx j 0.
+ pose Aij := row i \o vec_mx \o mulmxr A \o mxvec \o mulmx dj0.
+ have defAij := mul_rV_lin1 [linear of Aij]; rewrite /= {2}/Aij /= in defAij.
+ rewrite -defAij row_mul -defAij -!mulmxA (cent_mxP cBcE) {k}//.
+ rewrite memmx_cent_envelop; apply/centgmxP=> x Gx; apply/row_matrixP=> k.
+ rewrite !row_mul !rowE !{}defAij /= -row_mul mulmxA mul_delta_mx.
+ congr (row i _); rewrite -(mul_vec_lin (mulmxr_linear _ _)) -mulmxA.
+ by rewrite -(centgmxP cAG) // mulmxA mx_rV_lin.
+suffices redGn: mx_completely_reducible rGn 1%:M.
+ have [V modV defUV] := redGn _ modU (submx1 _); move/mxdirect_addsP=> dxUV.
+ rewrite -(proj_mx_id dxUV U_e1) -mulmxA {}cBncEn 1?mulmxA ?proj_mx_sub //.
+ by rewrite -row_full_dom_hom -sub1mx -defUV proj_mx_hom.
+pose W i : 'M[F]_(n ^ 2) := <<lin1_mx (mxvec \o mulmx (delta_mx i 0))>>%MS.
+have defW: (\sum_i W i :=: 1%:M)%MS.
+ apply/eqmxP; rewrite submx1; apply/row_subP; case/mxvec_indexP=> i j.
+ rewrite row1 -mxvec_delta (sumsmx_sup i) // genmxE; apply/submxP.
+ by exists (delta_mx 0 j); rewrite mul_rV_lin1 /= mul_delta_mx.
+apply: mxsemisimple_reducible; apply: (intro_mxsemisimple defW) => i _ nzWi.
+split=> // [|Vi modVi sViWi nzVi].
+ apply/mxmoduleP=> x Gx; rewrite genmxE (eqmxMr _ (genmxE _)) -def_rGn.
+ apply/row_subP=> j; rewrite rowE mulmxA !mul_rV_lin1 /= mxvecK -mulmxA.
+ by apply/submxP; move: (_ *m rG x) => v; exists v; rewrite mul_rV_lin1.
+do [rewrite !genmxE; set f := lin1_mx _] in sViWi *.
+have f_free: row_free f.
+ apply/row_freeP; exists (lin1_mx (row i \o vec_mx)); apply/row_matrixP=> j.
+ by rewrite row1 rowE mulmxA !mul_rV_lin1 /= mxvecK rowE !mul_delta_mx.
+pose V := <<Vi *m pinvmx f>>%MS; have Vidf := mulmxKpV sViWi.
+suffices: (1%:M <= V)%MS by rewrite genmxE -(submxMfree _ _ f_free) mul1mx Vidf.
+case: irrG => _ _ ->; rewrite ?submx1 //; last first.
+ by rewrite -mxrank_eq0 genmxE -(mxrankMfree _ f_free) Vidf mxrank_eq0.
+apply/mxmoduleP=> x Gx; rewrite genmxE (eqmxMr _ (genmxE _)).
+rewrite -(submxMfree _ _ f_free) Vidf.
+apply: submx_trans (mxmoduleP modVi x Gx); rewrite -{2}Vidf.
+apply/row_subP=> j; apply: (eq_row_sub j); rewrite row_mul -def_rGn.
+by rewrite !(row_mul _ _ f) !mul_rV_lin1 /= mxvecK !row_mul !mulmxA.
+Qed.
+
+Lemma cent_mx_scalar_abs_irr : \rank Hom_G <= 1 -> mx_absolutely_irreducible rG.
+Proof.
+rewrite leqNgt => /(has_non_scalar_mxP (scalar_mx_cent _ _)) scal_cE.
+apply/andP; split; first by case/mx_irrP: irrG.
+rewrite -sub1mx; apply: submx_trans mx_Jacobson_density.
+apply/memmx_subP=> B _; apply/cent_mxP=> A cGA.
+case scalA: (is_scalar_mx A); last by case: scal_cE; exists A; rewrite ?scalA.
+by case/is_scalar_mxP: scalA => a ->; rewrite scalar_mxC.
+Qed.
+
+End JacobsonDensity.
+
+Section ChangeGroup.
+
+Variables (gT : finGroupType) (G H : {group gT}) (n : nat).
+Variables (rG : mx_representation F G n).
+
+Section SubGroup.
+
+Hypothesis sHG : H \subset G.
+
+Local Notation rH := (subg_repr rG sHG).
+
+Lemma rfix_subg : rfix_mx rH = rfix_mx rG. Proof. by []. Qed.
+
+Section Stabilisers.
+
+Variables (m : nat) (U : 'M[F]_(m, n)).
+
+Lemma rstabs_subg : rstabs rH U = H :&: rstabs rG U.
+Proof. by apply/setP=> x; rewrite !inE andbA -in_setI (setIidPl sHG). Qed.
+
+Lemma mxmodule_subg : mxmodule rG U -> mxmodule rH U.
+Proof. by rewrite /mxmodule rstabs_subg subsetI subxx; exact: subset_trans. Qed.
+
+End Stabilisers.
+
+Lemma mxsimple_subg M : mxmodule rG M -> mxsimple rH M -> mxsimple rG M.
+Proof.
+by move=> modM [_ nzM minM]; split=> // U /mxmodule_subg; exact: minM.
+Qed.
+
+Lemma subg_mx_irr : mx_irreducible rH -> mx_irreducible rG.
+Proof. by apply: mxsimple_subg; exact: mxmodule1. Qed.
+
+Lemma subg_mx_abs_irr :
+ mx_absolutely_irreducible rH -> mx_absolutely_irreducible rG.
+Proof.
+rewrite /mx_absolutely_irreducible -!sub1mx => /andP[-> /submx_trans-> //].
+apply/row_subP=> i; rewrite rowK /= envelop_mx_id //.
+by rewrite (subsetP sHG) ?enum_valP.
+Qed.
+
+End SubGroup.
+
+Section SameGroup.
+
+Hypothesis eqGH : G :==: H.
+
+Local Notation rH := (eqg_repr rG eqGH).
+
+Lemma rfix_eqg : rfix_mx rH = rfix_mx rG. Proof. by []. Qed.
+
+Section Stabilisers.
+
+Variables (m : nat) (U : 'M[F]_(m, n)).
+
+Lemma rstabs_eqg : rstabs rH U = rstabs rG U.
+Proof. by rewrite rstabs_subg -(eqP eqGH) (setIidPr _) ?rstabs_sub. Qed.
+
+Lemma mxmodule_eqg : mxmodule rH U = mxmodule rG U.
+Proof. by rewrite /mxmodule rstabs_eqg -(eqP eqGH). Qed.
+
+End Stabilisers.
+
+Lemma mxsimple_eqg M : mxsimple rH M <-> mxsimple rG M.
+Proof.
+rewrite /mxsimple mxmodule_eqg.
+split=> [] [-> -> minM]; split=> // U modU;
+ by apply: minM; rewrite mxmodule_eqg in modU *.
+Qed.
+
+Lemma eqg_mx_irr : mx_irreducible rH <-> mx_irreducible rG.
+Proof. exact: mxsimple_eqg. Qed.
+
+Lemma eqg_mx_abs_irr :
+ mx_absolutely_irreducible rH = mx_absolutely_irreducible rG.
+Proof.
+by congr (_ && (_ == _)); rewrite /enveloping_algebra_mx /= -(eqP eqGH).
+Qed.
+
+End SameGroup.
+
+End ChangeGroup.
+
+Section Morphpre.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+Variables (G : {group rT}) (n : nat) (rG : mx_representation F G n).
+
+Local Notation rGf := (morphpre_repr f rG).
+
+Section Stabilisers.
+Variables (m : nat) (U : 'M[F]_(m, n)).
+
+Lemma rstabs_morphpre : rstabs rGf U = f @*^-1 (rstabs rG U).
+Proof. by apply/setP=> x; rewrite !inE andbA. Qed.
+
+Lemma mxmodule_morphpre : G \subset f @* D -> mxmodule rGf U = mxmodule rG U.
+Proof. by move=> sGf; rewrite /mxmodule rstabs_morphpre morphpreSK. Qed.
+
+End Stabilisers.
+
+Lemma rfix_morphpre (H : {set aT}) :
+ H \subset D -> (rfix_mx rGf H :=: rfix_mx rG (f @* H))%MS.
+Proof.
+move=> sHD; apply/eqmxP/andP; split.
+ by apply/rfix_mxP=> _ /morphimP[x _ Hx ->]; rewrite rfix_mx_id.
+by apply/rfix_mxP=> x Hx; rewrite rfix_mx_id ?mem_morphim ?(subsetP sHD).
+Qed.
+
+Lemma morphpre_mx_irr :
+ G \subset f @* D -> (mx_irreducible rGf <-> mx_irreducible rG).
+Proof.
+move/mxmodule_morphpre=> modG; split=> /mx_irrP[n_gt0 irrG];
+ by apply/mx_irrP; split=> // U modU; apply: irrG; rewrite modG in modU *.
+Qed.
+
+Lemma morphpre_mx_abs_irr :
+ G \subset f @* D ->
+ mx_absolutely_irreducible rGf = mx_absolutely_irreducible rG.
+Proof.
+move=> sGfD; congr (_ && (_ == _)); apply/eqP; rewrite mxrank_leqif_sup //.
+ apply/row_subP=> i; rewrite rowK.
+ case/morphimP: (subsetP sGfD _ (enum_valP i)) => x Dx _ def_i.
+ by rewrite def_i (envelop_mx_id rGf) // !inE Dx -def_i enum_valP.
+apply/row_subP=> i; rewrite rowK (envelop_mx_id rG) //.
+by case/morphpreP: (enum_valP i).
+Qed.
+
+End Morphpre.
+
+Section Morphim.
+
+Variables (aT rT : finGroupType) (G D : {group aT}) (f : {morphism D >-> rT}).
+Variables (n : nat) (rGf : mx_representation F (f @* G) n).
+
+Hypothesis sGD : G \subset D.
+
+Let sG_f'fG : G \subset f @*^-1 (f @* G).
+Proof. by rewrite -sub_morphim_pre. Qed.
+
+Local Notation rG := (morphim_repr rGf sGD).
+
+Section Stabilisers.
+Variables (m : nat) (U : 'M[F]_(m, n)).
+
+Lemma rstabs_morphim : rstabs rG U = G :&: f @*^-1 rstabs rGf U.
+Proof. by rewrite -rstabs_morphpre -(rstabs_subg _ sG_f'fG). Qed.
+
+Lemma mxmodule_morphim : mxmodule rG U = mxmodule rGf U.
+Proof. by rewrite /mxmodule rstabs_morphim subsetI subxx -sub_morphim_pre. Qed.
+
+End Stabilisers.
+
+Lemma rfix_morphim (H : {set aT}) :
+ H \subset D -> (rfix_mx rG H :=: rfix_mx rGf (f @* H))%MS.
+Proof. exact: rfix_morphpre. Qed.
+
+Lemma mxsimple_morphim M : mxsimple rG M <-> mxsimple rGf M.
+Proof.
+rewrite /mxsimple mxmodule_morphim.
+split=> [] [-> -> minM]; split=> // U modU;
+ by apply: minM; rewrite mxmodule_morphim in modU *.
+Qed.
+
+Lemma morphim_mx_irr : (mx_irreducible rG <-> mx_irreducible rGf).
+Proof. exact: mxsimple_morphim. Qed.
+
+Lemma morphim_mx_abs_irr :
+ mx_absolutely_irreducible rG = mx_absolutely_irreducible rGf.
+Proof.
+have fG_onto: f @* G \subset restrm sGD f @* G.
+ by rewrite morphim_restrm setIid.
+rewrite -(morphpre_mx_abs_irr _ fG_onto); congr (_ && (_ == _)).
+by rewrite /enveloping_algebra_mx /= morphpre_restrm (setIidPl _).
+Qed.
+
+End Morphim.
+
+Section Submodule.
+
+Variables (gT : finGroupType) (G : {group gT}) (n : nat).
+Variables (rG : mx_representation F G n) (U : 'M[F]_n) (Umod : mxmodule rG U).
+Local Notation rU := (submod_repr Umod).
+Local Notation rU' := (factmod_repr Umod).
+
+Lemma rfix_submod (H : {set gT}) :
+ H \subset G -> (rfix_mx rU H :=: in_submod U (U :&: rfix_mx rG H))%MS.
+Proof.
+move=> sHG; apply/eqmxP/andP; split; last first.
+ apply/rfix_mxP=> x Hx; rewrite -in_submodJ ?capmxSl //.
+ by rewrite (rfix_mxP H _) ?capmxSr.
+rewrite -val_submodS in_submodK ?capmxSl // sub_capmx val_submodP //=.
+apply/rfix_mxP=> x Hx.
+by rewrite -(val_submodJ Umod) ?(subsetP sHG) ?rfix_mx_id.
+Qed.
+
+Lemma rfix_factmod (H : {set gT}) :
+ H \subset G -> (in_factmod U (rfix_mx rG H) <= rfix_mx rU' H)%MS.
+Proof.
+move=> sHG; apply/rfix_mxP=> x Hx.
+by rewrite -(in_factmodJ Umod) ?(subsetP sHG) ?rfix_mx_id.
+Qed.
+
+Lemma rstab_submod m (W : 'M_(m, \rank U)) :
+ rstab rU W = rstab rG (val_submod W).
+Proof.
+apply/setP=> x; rewrite !inE; apply: andb_id2l => Gx.
+by rewrite -(inj_eq val_submod_inj) val_submodJ.
+Qed.
+
+Lemma rstabs_submod m (W : 'M_(m, \rank U)) :
+ rstabs rU W = rstabs rG (val_submod W).
+Proof.
+apply/setP=> x; rewrite !inE; apply: andb_id2l => Gx.
+by rewrite -val_submodS val_submodJ.
+Qed.
+
+Lemma val_submod_module m (W : 'M_(m, \rank U)) :
+ mxmodule rG (val_submod W) = mxmodule rU W.
+Proof. by rewrite /mxmodule rstabs_submod. Qed.
+
+Lemma in_submod_module m (V : 'M_(m, n)) :
+ (V <= U)%MS -> mxmodule rU (in_submod U V) = mxmodule rG V.
+Proof. by move=> sVU; rewrite -val_submod_module in_submodK. Qed.
+
+Lemma rstab_factmod m (W : 'M_(m, n)) :
+ rstab rG W \subset rstab rU' (in_factmod U W).
+Proof.
+by apply/subsetP=> x /setIdP[Gx /eqP cUW]; rewrite inE Gx -in_factmodJ //= cUW.
+Qed.
+
+Lemma rstabs_factmod m (W : 'M_(m, \rank (cokermx U))) :
+ rstabs rU' W = rstabs rG (U + val_factmod W)%MS.
+Proof.
+apply/setP=> x; rewrite !inE; apply: andb_id2l => Gx.
+rewrite addsmxMr addsmx_sub (submx_trans (mxmoduleP Umod x Gx)) ?addsmxSl //.
+rewrite -val_factmodS val_factmodJ //= val_factmodS; apply/idP/idP=> nWx.
+ rewrite (submx_trans (addsmxSr U _)) // -(in_factmodsK (addsmxSl U _)) //.
+ by rewrite addsmxS // val_factmodS in_factmod_addsK.
+rewrite in_factmodE (submx_trans (submxMr _ nWx)) // -in_factmodE.
+by rewrite in_factmod_addsK val_factmodK.
+Qed.
+
+Lemma val_factmod_module m (W : 'M_(m, \rank (cokermx U))) :
+ mxmodule rG (U + val_factmod W)%MS = mxmodule rU' W.
+Proof. by rewrite /mxmodule rstabs_factmod. Qed.
+
+Lemma in_factmod_module m (V : 'M_(m, n)) :
+ mxmodule rU' (in_factmod U V) = mxmodule rG (U + V)%MS.
+Proof.
+rewrite -(eqmx_module _ (in_factmodsK (addsmxSl U V))).
+by rewrite val_factmod_module (eqmx_module _ (in_factmod_addsK _ _)).
+Qed.
+
+Lemma rker_submod : rker rU = rstab rG U.
+Proof. by rewrite /rker rstab_submod; exact: eqmx_rstab (val_submod1 U). Qed.
+
+Lemma rstab_norm : G \subset 'N(rstab rG U).
+Proof. by rewrite -rker_submod rker_norm. Qed.
+
+Lemma rstab_normal : rstab rG U <| G.
+Proof. by rewrite -rker_submod rker_normal. Qed.
+
+Lemma submod_mx_faithful : mx_faithful rU -> mx_faithful rG.
+Proof. by apply: subset_trans; rewrite rker_submod rstabS ?submx1. Qed.
+
+Lemma rker_factmod : rker rG \subset rker rU'.
+Proof.
+apply/subsetP=> x /rkerP[Gx cVx].
+by rewrite inE Gx /= /factmod_mx cVx mul1mx mulmx1 val_factmodK.
+Qed.
+
+Lemma factmod_mx_faithful : mx_faithful rU' -> mx_faithful rG.
+Proof. exact: subset_trans rker_factmod. Qed.
+
+Lemma submod_mx_irr : mx_irreducible rU <-> mxsimple rG U.
+Proof.
+split=> [] [_ nzU simU].
+ rewrite -mxrank_eq0 mxrank1 mxrank_eq0 in nzU; split=> // V modV sVU nzV.
+ rewrite -(in_submodK sVU) -val_submod1 val_submodS.
+ rewrite -(genmxE (in_submod U V)) simU ?genmxE ?submx1 //=.
+ by rewrite (eqmx_module _ (genmxE _)) in_submod_module.
+ rewrite -submx0 genmxE -val_submodS in_submodK //.
+ by rewrite linear0 eqmx0 submx0.
+apply/mx_irrP; rewrite lt0n mxrank_eq0; split=> // V modV.
+rewrite -(inj_eq val_submod_inj) linear0 -(eqmx_eq0 (genmxE _)) => nzV.
+rewrite -sub1mx -val_submodS val_submod1 -(genmxE (val_submod V)).
+rewrite simU ?genmxE ?val_submodP //=.
+by rewrite (eqmx_module _ (genmxE _)) val_submod_module.
+Qed.
+
+End Submodule.
+
+Section Conjugate.
+
+Variables (gT : finGroupType) (G : {group gT}) (n : nat).
+Variables (rG : mx_representation F G n) (B : 'M[F]_n).
+
+Hypothesis uB : B \in unitmx.
+
+Local Notation rGB := (rconj_repr rG uB).
+
+Lemma rfix_conj (H : {set gT}) :
+ (rfix_mx rGB H :=: B *m rfix_mx rG H *m invmx B)%MS.
+Proof.
+apply/eqmxP/andP; split.
+ rewrite -mulmxA (eqmxMfull (_ *m _)) ?row_full_unit //.
+ rewrite -[rfix_mx rGB H](mulmxK uB) submxMr //; apply/rfix_mxP=> x Hx.
+ apply: (canRL (mulmxKV uB)); rewrite -(rconj_mxJ _ uB) mulmxK //.
+ by rewrite rfix_mx_id.
+apply/rfix_mxP=> x Gx; rewrite -3!mulmxA; congr (_ *m _).
+by rewrite !mulmxA mulmxKV // rfix_mx_id.
+Qed.
+
+Lemma rstabs_conj m (U : 'M_(m, n)) : rstabs rGB U = rstabs rG (U *m B).
+Proof.
+apply/setP=> x; rewrite !inE rconj_mxE !mulmxA.
+by rewrite -{2}[U](mulmxK uB) submxMfree // row_free_unit unitmx_inv.
+Qed.
+
+Lemma mxmodule_conj m (U : 'M_(m, n)) : mxmodule rGB U = mxmodule rG (U *m B).
+Proof. by rewrite /mxmodule rstabs_conj. Qed.
+
+Lemma conj_mx_irr : mx_irreducible rGB <-> mx_irreducible rG.
+Proof.
+have Bfree: row_free B by rewrite row_free_unit.
+split => /mx_irrP[n_gt0 irrG]; apply/mx_irrP; split=> // U.
+ rewrite -[U](mulmxKV uB) -mxmodule_conj -mxrank_eq0 /row_full mxrankMfree //.
+ by rewrite mxrank_eq0; exact: irrG.
+rewrite -mxrank_eq0 /row_full -(mxrankMfree _ Bfree) mxmodule_conj mxrank_eq0.
+exact: irrG.
+Qed.
+
+End Conjugate.
+
+Section Quotient.
+
+Variables (gT : finGroupType) (G : {group gT}) (n : nat).
+Variables (rG : mx_representation F G n) (H : {group gT}).
+Hypotheses (krH : H \subset rker rG) (nHG : G \subset 'N(H)).
+Let nHGs := subsetP nHG.
+
+Local Notation rGH := (quo_repr krH nHG).
+
+Local Notation E_ r := (enveloping_algebra_mx r).
+Lemma quo_mx_quotient : (E_ rGH :=: E_ rG)%MS.
+Proof.
+apply/eqmxP/andP; split; apply/row_subP=> i.
+ rewrite rowK; case/morphimP: (enum_valP i) => x _ Gx ->{i}.
+ rewrite quo_repr_coset // (eq_row_sub (enum_rank_in Gx x)) // rowK.
+ by rewrite enum_rankK_in.
+rewrite rowK -(quo_mx_coset krH nHG) ?enum_valP //; set Hx := coset H _.
+have GHx: Hx \in (G / H)%g by rewrite mem_quotient ?enum_valP.
+by rewrite (eq_row_sub (enum_rank_in GHx Hx)) // rowK enum_rankK_in.
+Qed.
+
+Lemma rfix_quo (K : {group gT}) :
+ K \subset G -> (rfix_mx rGH (K / H)%g :=: rfix_mx rG K)%MS.
+Proof.
+move=> sKG; apply/eqmxP/andP; (split; apply/rfix_mxP) => [x Kx | Hx].
+ have Gx := subsetP sKG x Kx; rewrite -(quo_mx_coset krH nHG) // rfix_mx_id //.
+ by rewrite mem_morphim ?(subsetP nHG).
+case/morphimP=> x _ Kx ->; have Gx := subsetP sKG x Kx.
+by rewrite quo_repr_coset ?rfix_mx_id.
+Qed.
+
+Lemma rstabs_quo m (U : 'M_(m, n)) : rstabs rGH U = (rstabs rG U / H)%g.
+Proof.
+apply/setP=> Hx; rewrite !inE; apply/andP/idP=> [[]|] /morphimP[x Nx Gx ->{Hx}].
+ by rewrite quo_repr_coset // => nUx; rewrite mem_morphim // inE Gx.
+by case/setIdP: Gx => Gx nUx; rewrite quo_repr_coset ?mem_morphim.
+Qed.
+
+Lemma mxmodule_quo m (U : 'M_(m, n)) : mxmodule rGH U = mxmodule rG U.
+Proof.
+rewrite /mxmodule rstabs_quo quotientSGK // ?(subset_trans krH) //.
+apply/subsetP=> x; rewrite !inE mul1mx => /andP[-> /eqP->].
+by rewrite /= mulmx1.
+Qed.
+
+Lemma quo_mx_irr : mx_irreducible rGH <-> mx_irreducible rG.
+Proof.
+split; case/mx_irrP=> n_gt0 irrG; apply/mx_irrP; split=> // U modU;
+ by apply: irrG; rewrite mxmodule_quo in modU *.
+Qed.
+
+End Quotient.
+
+Section SplittingField.
+
+Implicit Type gT : finGroupType.
+
+Definition group_splitting_field gT (G : {group gT}) :=
+ forall n (rG : mx_representation F G n),
+ mx_irreducible rG -> mx_absolutely_irreducible rG.
+
+Definition group_closure_field gT :=
+ forall G : {group gT}, group_splitting_field G.
+
+Lemma quotient_splitting_field gT (G : {group gT}) (H : {set gT}) :
+ G \subset 'N(H) -> group_splitting_field G -> group_splitting_field (G / H).
+Proof.
+move=> nHG splitG n rGH irrGH.
+by rewrite -(morphim_mx_abs_irr _ nHG) splitG //; exact/morphim_mx_irr.
+Qed.
+
+Lemma coset_splitting_field gT (H : {set gT}) :
+ group_closure_field gT -> group_closure_field (coset_groupType H).
+Proof.
+move=> split_gT Gbar; have ->: Gbar = (coset H @*^-1 Gbar / H)%G.
+ by apply: val_inj; rewrite /= /quotient morphpreK ?sub_im_coset.
+by apply: quotient_splitting_field; [exact: subsetIl | exact: split_gT].
+Qed.
+
+End SplittingField.
+
+Section Abelian.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Lemma mx_faithful_irr_center_cyclic n (rG : mx_representation F G n) :
+ mx_faithful rG -> mx_irreducible rG -> cyclic 'Z(G).
+Proof.
+case: n rG => [|n] rG injG irrG; first by case/mx_irrP: irrG.
+move/trivgP: injG => KrG1; pose rZ := subg_repr rG (center_sub _).
+apply: (div_ring_mul_group_cyclic (repr_mx1 rZ)) (repr_mxM rZ) _ _; last first.
+ exact: center_abelian.
+move=> x; rewrite -[[set _]]KrG1 !inE mul1mx -subr_eq0 andbC; set U := _ - _.
+do 2![case/andP]=> Gx cGx; rewrite Gx /=; apply: (mx_Schur irrG).
+apply/centgmxP=> y Gy; rewrite mulmxBl mulmxBr mulmx1 mul1mx.
+by rewrite -!repr_mxM // (centP cGx).
+Qed.
+
+Lemma mx_faithful_irr_abelian_cyclic n (rG : mx_representation F G n) :
+ mx_faithful rG -> mx_irreducible rG -> abelian G -> cyclic G.
+Proof.
+move=> injG irrG cGG; rewrite -(setIidPl cGG).
+exact: mx_faithful_irr_center_cyclic injG irrG.
+Qed.
+
+Hypothesis splitG : group_splitting_field G.
+
+Lemma mx_irr_abelian_linear n (rG : mx_representation F G n) :
+ mx_irreducible rG -> abelian G -> n = 1%N.
+Proof.
+by move=> irrG cGG; apply/eqP; rewrite -(abelian_abs_irr rG) ?splitG.
+Qed.
+
+Lemma mxsimple_abelian_linear n (rG : mx_representation F G n) M :
+ abelian G -> mxsimple rG M -> \rank M = 1%N.
+Proof.
+move=> cGG simM; have [modM _ _] := simM.
+by move/(submod_mx_irr modM)/mx_irr_abelian_linear: simM => ->.
+Qed.
+
+Lemma linear_mxsimple n (rG : mx_representation F G n) (M : 'M_n) :
+ mxmodule rG M -> \rank M = 1%N -> mxsimple rG M.
+Proof.
+move=> modM rM1; apply/(submod_mx_irr modM).
+by apply: mx_abs_irrW; rewrite linear_mx_abs_irr.
+Qed.
+
+End Abelian.
+
+Section AbelianQuotient.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Variables (n : nat) (rG : mx_representation F G n).
+
+Lemma center_kquo_cyclic : mx_irreducible rG -> cyclic 'Z(G / rker rG)%g.
+Proof.
+move=> irrG; apply: mx_faithful_irr_center_cyclic (kquo_mx_faithful rG) _.
+exact/quo_mx_irr.
+Qed.
+
+Lemma der1_sub_rker :
+ group_splitting_field G -> mx_irreducible rG ->
+ (G^`(1) \subset rker rG)%g = (n == 1)%N.
+Proof.
+move=> splitG irrG; apply/idP/idP; last by move/eqP; exact: rker_linear.
+move/sub_der1_abelian; move/(abelian_abs_irr (kquo_repr rG))=> <-.
+by apply: (quotient_splitting_field (rker_norm _) splitG); exact/quo_mx_irr.
+Qed.
+
+End AbelianQuotient.
+
+Section Similarity.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Local Notation reprG := (mx_representation F G).
+
+CoInductive mx_rsim n1 (rG1 : reprG n1) n2 (rG2 : reprG n2) : Prop :=
+ MxReprSim B of n1 = n2 & row_free B
+ & forall x, x \in G -> rG1 x *m B = B *m rG2 x.
+
+Lemma mxrank_rsim n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) :
+ mx_rsim rG1 rG2 -> n1 = n2.
+Proof. by case. Qed.
+
+Lemma mx_rsim_refl n (rG : reprG n) : mx_rsim rG rG.
+Proof.
+exists 1%:M => // [|x _]; first by rewrite row_free_unit unitmx1.
+by rewrite mulmx1 mul1mx.
+Qed.
+
+Lemma mx_rsim_sym n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) :
+ mx_rsim rG1 rG2 -> mx_rsim rG2 rG1.
+Proof.
+case=> B def_n1; rewrite def_n1 in rG1 B *.
+rewrite row_free_unit => injB homB; exists (invmx B) => // [|x Gx].
+ by rewrite row_free_unit unitmx_inv.
+by apply: canRL (mulKmx injB) _; rewrite mulmxA -homB ?mulmxK.
+Qed.
+
+Lemma mx_rsim_trans n1 n2 n3
+ (rG1 : reprG n1) (rG2 : reprG n2) (rG3 : reprG n3) :
+ mx_rsim rG1 rG2 -> mx_rsim rG2 rG3 -> mx_rsim rG1 rG3.
+Proof.
+case=> [B1 defn1 freeB1 homB1] [B2 defn2 freeB2 homB2].
+exists (B1 *m B2); rewrite /row_free ?mxrankMfree 1?defn1 // => x Gx.
+by rewrite mulmxA homB1 // -!mulmxA homB2.
+Qed.
+
+Lemma mx_rsim_def n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) :
+ mx_rsim rG1 rG2 ->
+ exists B, exists2 B', B' *m B = 1%:M &
+ forall x, x \in G -> rG1 x = B *m rG2 x *m B'.
+Proof.
+case=> B def_n1; rewrite def_n1 in rG1 B *; rewrite row_free_unit => injB homB.
+by exists B, (invmx B) => [|x Gx]; rewrite ?mulVmx // -homB // mulmxK.
+Qed.
+
+Lemma mx_rsim_iso n (rG : reprG n) (U V : 'M_n)
+ (modU : mxmodule rG U) (modV : mxmodule rG V) :
+ mx_rsim (submod_repr modU) (submod_repr modV) <-> mx_iso rG U V.
+Proof.
+split=> [[B eqrUV injB homB] | [f injf homf defV]].
+ have: \rank (U *m val_submod (in_submod U 1%:M *m B)) = \rank U.
+ do 2!rewrite mulmxA mxrankMfree ?row_base_free //.
+ by rewrite -(eqmxMr _ (val_submod1 U)) -in_submodE val_submodK mxrank1.
+ case/complete_unitmx => f injf defUf; exists f => //.
+ apply/hom_mxP=> x Gx; rewrite -defUf -2!mulmxA -(val_submodJ modV) //.
+ rewrite -(mulmxA _ B) -homB // val_submodE 3!(mulmxA U) (mulmxA _ _ B).
+ rewrite -in_submodE -in_submodJ //.
+ have [u ->] := submxP (mxmoduleP modU x Gx).
+ by rewrite in_submodE -mulmxA -defUf !mulmxA mulmx1.
+ apply/eqmxP; rewrite -mxrank_leqif_eq.
+ by rewrite mxrankMfree ?eqrUV ?row_free_unit.
+ by rewrite -defUf mulmxA val_submodP.
+have eqrUV: \rank U = \rank V by rewrite -defV mxrankMfree ?row_free_unit.
+exists (in_submod V (val_submod 1%:M *m f)) => // [|x Gx].
+ rewrite /row_free {6}eqrUV -[_ == _]sub1mx -val_submodS.
+ rewrite in_submodK; last by rewrite -defV submxMr ?val_submodP.
+ by rewrite val_submod1 -defV submxMr ?val_submod1.
+rewrite -in_submodJ; last by rewrite -defV submxMr ?val_submodP.
+rewrite -(hom_mxP (submx_trans (val_submodP _) homf)) //.
+by rewrite -(val_submodJ modU) // mul1mx 2!(mulmxA ((submod_repr _) x)) -val_submodE.
+Qed.
+
+Lemma mx_rsim_irr n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) :
+ mx_rsim rG1 rG2 -> mx_irreducible rG1 -> mx_irreducible rG2.
+Proof.
+case/mx_rsim_sym=> f def_n2; rewrite {n2}def_n2 in f rG2 * => injf homf.
+case/mx_irrP=> n1_gt0 minG; apply/mx_irrP; split=> // U modU nzU.
+rewrite /row_full -(mxrankMfree _ injf) -genmxE.
+apply: minG; last by rewrite -mxrank_eq0 genmxE mxrankMfree // mxrank_eq0.
+rewrite (eqmx_module _ (genmxE _)); apply/mxmoduleP=> x Gx.
+by rewrite -mulmxA -homf // mulmxA submxMr // (mxmoduleP modU).
+Qed.
+
+Lemma mx_rsim_abs_irr n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) :
+ mx_rsim rG1 rG2 ->
+ mx_absolutely_irreducible rG1 = mx_absolutely_irreducible rG2.
+Proof.
+case=> f def_n2; rewrite -{n2}def_n2 in f rG2 *.
+rewrite row_free_unit => injf homf; congr (_ && (_ == _)).
+pose Eg (g : 'M[F]_n1) := lin_mx (mulmxr (invmx g) \o mulmx g).
+have free_Ef: row_free (Eg f).
+ apply/row_freeP; exists (Eg (invmx f)); apply/row_matrixP=> i.
+ rewrite rowE row1 mulmxA mul_rV_lin mx_rV_lin /=.
+ by rewrite invmxK !{1}mulmxA mulmxKV // -mulmxA mulKmx // vec_mxK.
+symmetry; rewrite -(mxrankMfree _ free_Ef); congr (\rank _).
+apply/row_matrixP=> i; rewrite row_mul !rowK mul_vec_lin /=.
+by rewrite -homf ?enum_valP // mulmxK.
+Qed.
+
+Lemma rker_mx_rsim n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) :
+ mx_rsim rG1 rG2 -> rker rG1 = rker rG2.
+Proof.
+case=> f def_n2; rewrite -{n2}def_n2 in f rG2 *.
+rewrite row_free_unit => injf homf.
+apply/setP=> x; rewrite !inE !mul1mx; apply: andb_id2l => Gx.
+by rewrite -(can_eq (mulmxK injf)) homf // -scalar_mxC (can_eq (mulKmx injf)).
+Qed.
+
+Lemma mx_rsim_faithful n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) :
+ mx_rsim rG1 rG2 -> mx_faithful rG1 = mx_faithful rG2.
+Proof. by move=> simG12; rewrite /mx_faithful (rker_mx_rsim simG12). Qed.
+
+Lemma mx_rsim_factmod n (rG : reprG n) U V
+ (modU : mxmodule rG U) (modV : mxmodule rG V) :
+ (U + V :=: 1%:M)%MS -> mxdirect (U + V) ->
+ mx_rsim (factmod_repr modV) (submod_repr modU).
+Proof.
+move=> addUV dxUV.
+have eqUV: \rank U = \rank (cokermx V).
+ by rewrite mxrank_coker -{3}(mxrank1 F n) -addUV (mxdirectP dxUV) addnK.
+have{dxUV} dxUV: (U :&: V = 0)%MS by exact/mxdirect_addsP.
+exists (in_submod U (val_factmod 1%:M *m proj_mx U V)) => // [|x Gx].
+ rewrite /row_free -{6}eqUV -[_ == _]sub1mx -val_submodS val_submod1.
+ rewrite in_submodK ?proj_mx_sub // -{1}[U](proj_mx_id dxUV) //.
+ rewrite -{1}(add_sub_fact_mod V U) mulmxDl proj_mx_0 ?val_submodP // add0r.
+ by rewrite submxMr // val_factmodS submx1.
+rewrite -in_submodJ ?proj_mx_sub // -(hom_mxP _) //; last first.
+ by apply: submx_trans (submx1 _) _; rewrite -addUV proj_mx_hom.
+rewrite mulmxA; congr (_ *m _); rewrite mulmxA -val_factmodE; apply/eqP.
+rewrite eq_sym -subr_eq0 -mulmxBl proj_mx_0 //.
+by rewrite -[_ *m rG x](add_sub_fact_mod V) addrK val_submodP.
+Qed.
+
+Lemma mxtrace_rsim n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) :
+ mx_rsim rG1 rG2 -> {in G, forall x, \tr (rG1 x) = \tr (rG2 x)}.
+Proof.
+case/mx_rsim_def=> B [B' B'B def_rG1] x Gx.
+by rewrite def_rG1 // mxtrace_mulC mulmxA B'B mul1mx.
+Qed.
+
+Lemma mx_rsim_scalar n1 n2 (rG1 : reprG n1) (rG2 : reprG n2) x c :
+ x \in G -> mx_rsim rG1 rG2 -> rG1 x = c%:M -> rG2 x = c%:M.
+Proof.
+move=> Gx /mx_rsim_sym[B _ Bfree rG2_B] rG1x.
+by apply: (row_free_inj Bfree); rewrite rG2_B // rG1x scalar_mxC.
+Qed.
+
+End Similarity.
+
+Section Socle.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Variables (n : nat) (rG : mx_representation F G n) (sG : socleType rG).
+
+Lemma socle_irr (W : sG) : mx_irreducible (socle_repr W).
+Proof. by apply/submod_mx_irr; exact: socle_simple. Qed.
+
+Lemma socle_rsimP (W1 W2 : sG) :
+ reflect (mx_rsim (socle_repr W1) (socle_repr W2)) (W1 == W2).
+Proof.
+have [simW1 simW2] := (socle_simple W1, socle_simple W2).
+by apply: (iffP (component_mx_isoP simW1 simW2)); move/mx_rsim_iso; exact.
+Qed.
+
+Local Notation mG U := (mxmodule rG U).
+Local Notation sr modV := (submod_repr modV).
+
+Lemma mx_rsim_in_submod U V (modU : mG U) (modV : mG V) :
+ let U' := <<in_submod V U>>%MS in
+ (U <= V)%MS ->
+ exists modU' : mxmodule (sr modV) U', mx_rsim (sr modU) (sr modU').
+Proof.
+move=> U' sUV; have modU': mxmodule (sr modV) U'.
+ by rewrite (eqmx_module _ (genmxE _)) in_submod_module.
+have rankU': \rank U = \rank U' by rewrite genmxE mxrank_in_submod.
+pose v1 := val_submod 1%:M; pose U1 := v1 _ U.
+have sU1V: (U1 <= V)%MS by rewrite val_submod1.
+have sU1U': (in_submod V U1 <= U')%MS by rewrite genmxE submxMr ?val_submod1.
+exists modU', (in_submod U' (in_submod V U1)) => // [|x Gx].
+ apply/row_freeP; exists (v1 _ _ *m v1 _ _ *m in_submod U 1%:M).
+ by rewrite 2!mulmxA -in_submodE -!val_submodE !in_submodK ?val_submodK.
+rewrite -!in_submodJ // -(val_submodJ modU) // mul1mx.
+by rewrite 2!{1}in_submodE mulmxA (mulmxA _ U1) -val_submodE -!in_submodE.
+Qed.
+
+Lemma rsim_submod1 U (modU : mG U) : (U :=: 1%:M)%MS -> mx_rsim (sr modU) rG.
+Proof.
+move=> U1; exists (val_submod 1%:M) => [||x Gx]; first by rewrite U1 mxrank1.
+ by rewrite /row_free val_submod1.
+by rewrite -(val_submodJ modU) // mul1mx -val_submodE.
+Qed.
+
+Lemma mxtrace_submod1 U (modU : mG U) :
+ (U :=: 1%:M)%MS -> {in G, forall x, \tr (sr modU x) = \tr (rG x)}.
+Proof. by move=> defU; exact: mxtrace_rsim (rsim_submod1 modU defU). Qed.
+
+Lemma mxtrace_dadd_mod U V W (modU : mG U) (modV : mG V) (modW : mG W) :
+ (U + V :=: W)%MS -> mxdirect (U + V) ->
+ {in G, forall x, \tr (sr modU x) + \tr (sr modV x) = \tr (sr modW x)}.
+Proof.
+move=> defW dxW x Gx; have [sUW sVW]: (U <= W)%MS /\ (V <= W)%MS.
+ by apply/andP; rewrite -addsmx_sub defW.
+pose U' := <<in_submod W U>>%MS; pose V' := <<in_submod W V>>%MS.
+have addUV': (U' + V' :=: 1%:M)%MS.
+ apply/eqmxP; rewrite submx1 /= (adds_eqmx (genmxE _) (genmxE _)).
+ by rewrite -addsmxMr -val_submodS val_submod1 in_submodK ?defW.
+have dxUV': mxdirect (U' + V').
+ apply/eqnP; rewrite /= addUV' mxrank1 !genmxE !mxrank_in_submod //.
+ by rewrite -(mxdirectP dxW) /= defW.
+have [modU' simU] := mx_rsim_in_submod modU modW sUW.
+have [modV' simV] := mx_rsim_in_submod modV modW sVW.
+rewrite (mxtrace_rsim simU) // (mxtrace_rsim simV) //.
+rewrite -(mxtrace_sub_fact_mod modV') addrC; congr (_ + _).
+by rewrite (mxtrace_rsim (mx_rsim_factmod modU' modV' addUV' dxUV')).
+Qed.
+
+Lemma mxtrace_dsum_mod (I : finType) (P : pred I) U W
+ (modU : forall i, mG (U i)) (modW : mG W) :
+ let S := (\sum_(i | P i) U i)%MS in (S :=: W)%MS -> mxdirect S ->
+ {in G, forall x, \sum_(i | P i) \tr (sr (modU i) x) = \tr (sr modW x)}.
+Proof.
+move=> /= sumS dxS x Gx.
+elim: {P}_.+1 {-2}P (ltnSn #|P|) => // m IHm P lePm in W modW sumS dxS *.
+have [j /= Pj | P0] := pickP P; last first.
+ case: sumS (_ x); rewrite !big_pred0 // mxrank0 => <- _ rWx.
+ by rewrite [rWx]flatmx0 linear0.
+rewrite ltnS (cardD1x Pj) in lePm.
+rewrite mxdirectE /= !(bigD1 j Pj) -mxdirectE mxdirect_addsE /= in dxS sumS *.
+have [_ dxW' dxW] := and3P dxS; rewrite (sameP eqP mxdirect_addsP) in dxW.
+rewrite (IHm _ _ _ (sumsmx_module _ (fun i _ => modU i)) (eqmx_refl _)) //.
+exact: mxtrace_dadd_mod.
+Qed.
+
+Lemma mxtrace_component U (simU : mxsimple rG U) :
+ let V := component_mx rG U in
+ let modV := component_mx_module rG U in let modU := mxsimple_module simU in
+ {in G, forall x, \tr (sr modV x) = \tr (sr modU x) *+ (\rank V %/ \rank U)}.
+Proof.
+move=> V modV modU x Gx.
+have [I W S simW defV dxV] := component_mx_semisimple simU.
+rewrite -(mxtrace_dsum_mod (fun i => mxsimple_module (simW i)) modV defV) //.
+have rankU_gt0: \rank U > 0 by rewrite lt0n mxrank_eq0; case simU.
+have isoW i: mx_iso rG U (W i).
+ by apply: component_mx_iso; rewrite ?simU // -defV (sumsmx_sup i).
+have ->: (\rank V %/ \rank U)%N = #|I|.
+ symmetry; rewrite -(mulnK #|I| rankU_gt0); congr (_ %/ _)%N.
+ rewrite -defV (mxdirectP dxV) /= -sum_nat_const.
+ by apply: eq_bigr => i _; exact: mxrank_iso.
+rewrite -sumr_const; apply: eq_bigr => i _; symmetry.
+by apply: mxtrace_rsim Gx; apply/mx_rsim_iso; exact: isoW.
+Qed.
+
+Lemma mxtrace_Socle : let modS := Socle_module sG in
+ {in G, forall x,
+ \tr (sr modS x) = \sum_(W : sG) \tr (socle_repr W x) *+ socle_mult W}.
+Proof.
+move=> /= x Gx /=; pose modW (W : sG) := component_mx_module rG (socle_base W).
+rewrite -(mxtrace_dsum_mod modW _ (eqmx_refl _) (Socle_direct sG)) //.
+by apply: eq_bigr => W _; rewrite (mxtrace_component (socle_simple W)).
+Qed.
+
+End Socle.
+
+Section Clifford.
+
+Variables (gT : finGroupType) (G H : {group gT}).
+Hypothesis nsHG : H <| G.
+Variables (n : nat) (rG : mx_representation F G n).
+Let sHG := normal_sub nsHG.
+Let nHG := normal_norm nsHG.
+Let rH := subg_repr rG sHG.
+
+Lemma Clifford_simple M x : mxsimple rH M -> x \in G -> mxsimple rH (M *m rG x).
+Proof.
+have modmG m U y: y \in G -> (mxmodule rH) m U -> mxmodule rH (U *m rG y).
+ move=> Gy modU; apply/mxmoduleP=> h Hh; have Gh := subsetP sHG h Hh.
+ rewrite -mulmxA -repr_mxM // conjgCV repr_mxM ?groupJ ?groupV // mulmxA.
+ by rewrite submxMr ?(mxmoduleP modU) // -mem_conjg (normsP nHG).
+have nzmG m y (U : 'M_(m, n)): y \in G -> (U *m rG y == 0) = (U == 0).
+ by move=> Gy; rewrite -{1}(mul0mx m (rG y)) (can_eq (repr_mxK rG Gy)).
+case=> [modM nzM simM] Gx; have Gx' := groupVr Gx.
+split=> [||U modU sUMx nzU]; rewrite ?modmG ?nzmG //.
+rewrite -(repr_mxKV rG Gx U) submxMr //.
+by rewrite (simM (U *m _)) ?modmG ?nzmG // -(repr_mxK rG Gx M) submxMr.
+Qed.
+
+Lemma Clifford_hom x m (U : 'M_(m, n)) :
+ x \in 'C_G(H) -> (U <= dom_hom_mx rH (rG x))%MS.
+Proof.
+case/setIP=> Gx cHx; apply/rV_subP=> v _{U}.
+apply/hom_mxP=> h Hh; have Gh := subsetP sHG h Hh.
+by rewrite -!mulmxA /= -!repr_mxM // (centP cHx).
+Qed.
+
+Lemma Clifford_iso x U : x \in 'C_G(H) -> mx_iso rH U (U *m rG x).
+Proof.
+move=> cHx; have [Gx _] := setIP cHx.
+by exists (rG x); rewrite ?repr_mx_unit ?Clifford_hom.
+Qed.
+
+Lemma Clifford_iso2 x U V :
+ mx_iso rH U V -> x \in G -> mx_iso rH (U *m rG x) (V *m rG x).
+Proof.
+case=> [f injf homUf defV] Gx; have Gx' := groupVr Gx.
+pose fx := rG (x^-1)%g *m f *m rG x; exists fx; last 1 first.
+- by rewrite !mulmxA repr_mxK //; exact: eqmxMr.
+- by rewrite !unitmx_mul andbC !repr_mx_unit.
+apply/hom_mxP=> h Hh; have Gh := subsetP sHG h Hh.
+rewrite -(mulmxA U) -repr_mxM // conjgCV repr_mxM ?groupJ // !mulmxA.
+rewrite !repr_mxK // (hom_mxP homUf) -?mem_conjg ?(normsP nHG) //=.
+by rewrite !repr_mxM ?invgK ?groupM // !mulmxA repr_mxKV.
+Qed.
+
+Lemma Clifford_componentJ M x :
+ mxsimple rH M -> x \in G ->
+ (component_mx rH (M *m rG x) :=: component_mx rH M *m rG x)%MS.
+Proof.
+set simH := mxsimple rH; set cH := component_mx rH.
+have actG: {in G, forall y M, simH M -> cH M *m rG y <= cH (M *m rG y)}%MS.
+ move=> {M} y Gy /= M simM; have [I [U isoU def_cHM]] := component_mx_def simM.
+ rewrite /cH def_cHM sumsmxMr; apply/sumsmx_subP=> i _.
+ by apply: mx_iso_component; [exact: Clifford_simple | exact: Clifford_iso2].
+move=> simM Gx; apply/eqmxP; rewrite actG // -/cH.
+rewrite -{1}[cH _](repr_mxKV rG Gx) submxMr // -{2}[M](repr_mxK rG Gx).
+by rewrite actG ?groupV //; exact: Clifford_simple.
+Qed.
+
+Hypothesis irrG : mx_irreducible rG.
+
+Lemma Clifford_basis M : mxsimple rH M ->
+ {X : {set gT} | X \subset G &
+ let S := \sum_(x in X) M *m rG x in S :=: 1%:M /\ mxdirect S}%MS.
+Proof.
+move=> simM. have simMG (g : [subg G]) : mxsimple rH (M *m rG (val g)).
+ by case: g => x Gx; exact: Clifford_simple.
+have [|XG [defX1 dxX1]] := sum_mxsimple_direct_sub simMG (_ : _ :=: 1%:M)%MS.
+ apply/eqmxP; case irrG => _ _ ->; rewrite ?submx1 //; last first.
+ rewrite -submx0; apply/sumsmx_subP; move/(_ 1%g (erefl _)); apply: negP.
+ by rewrite submx0 repr_mx1 mulmx1; case simM.
+ apply/mxmoduleP=> x Gx; rewrite sumsmxMr; apply/sumsmx_subP=> [[y Gy]] /= _.
+ by rewrite (sumsmx_sup (subg G (y * x))) // subgK ?groupM // -mulmxA repr_mxM.
+exists (val @: XG); first by apply/subsetP=> ?; case/imsetP=> [[x Gx]] _ ->.
+have bij_val: {on val @: XG, bijective (@sgval _ G)}.
+ exists (subg G) => [g _ | x]; first exact: sgvalK.
+ by case/imsetP=> [[x' Gx]] _ ->; rewrite subgK.
+have defXG g: (val g \in val @: XG) = (g \in XG).
+ by apply/imsetP/idP=> [[h XGh] | XGg]; [move/val_inj-> | exists g].
+by rewrite /= mxdirectE /= !(reindex _ bij_val) !(eq_bigl _ _ defXG).
+Qed.
+
+Variable sH : socleType rH.
+
+Definition Clifford_act (W : sH) x :=
+ let Gx := subgP (subg G x) in
+ PackSocle (component_socle sH (Clifford_simple (socle_simple W) Gx)).
+
+Let valWact W x : (Clifford_act W x :=: W *m rG (sgval (subg G x)))%MS.
+Proof.
+rewrite PackSocleK; apply: Clifford_componentJ (subgP _).
+exact: socle_simple.
+Qed.
+
+Fact Clifford_is_action : is_action G Clifford_act.
+Proof.
+split=> [x W W' eqWW' | W x y Gx Gy].
+ pose Gx := subgP (subg G x); apply/socleP; apply/eqmxP.
+ rewrite -(repr_mxK rG Gx W) -(repr_mxK rG Gx W'); apply: eqmxMr.
+ apply: eqmx_trans (eqmx_sym _) (valWact _ _); rewrite -eqWW'; exact: valWact.
+apply/socleP; rewrite !{1}valWact 2!{1}(eqmxMr _ (valWact _ _)).
+by rewrite !subgK ?groupM ?repr_mxM ?mulmxA ?andbb.
+Qed.
+
+Definition Clifford_action := Action Clifford_is_action.
+
+Local Notation "'Cl" := Clifford_action (at level 8) : action_scope.
+
+Lemma val_Clifford_act W x : x \in G -> ('Cl%act W x :=: W *m rG x)%MS.
+Proof. by move=> Gx; apply: eqmx_trans (valWact _ _) _; rewrite subgK. Qed.
+
+Lemma Clifford_atrans : [transitive G, on [set: sH] | 'Cl].
+Proof.
+have [_ nz1 _] := irrG.
+apply: mxsimple_exists (mxmodule1 rH) nz1 _ _ => [[M simM _]].
+pose W1 := PackSocle (component_socle sH simM).
+have [X sXG [def1 _]] := Clifford_basis simM; move/subsetP: sXG => sXG.
+apply/imsetP; exists W1; first by rewrite inE.
+symmetry; apply/setP=> W; rewrite inE; have simW := socle_simple W.
+have:= submx1 (socle_base W); rewrite -def1 -[(\sum_(x in X) _)%MS]mulmx1.
+case/(hom_mxsemisimple_iso simW) => [x Xx _ | | x Xx isoMxW].
+- by apply: Clifford_simple; rewrite ?sXG.
+- exact: scalar_mx_hom.
+have Gx := sXG x Xx; apply/imsetP; exists x => //; apply/socleP/eqmxP/eqmx_sym.
+apply: eqmx_trans (val_Clifford_act _ Gx) _; rewrite PackSocleK.
+apply: eqmx_trans (eqmx_sym (Clifford_componentJ simM Gx)) _.
+apply/eqmxP; rewrite (sameP genmxP eqP) !{1}genmx_component.
+by apply/component_mx_isoP=> //; exact: Clifford_simple.
+Qed.
+
+Lemma Clifford_Socle1 : Socle sH = 1%:M.
+Proof.
+case/imsetP: Clifford_atrans => W _ _; have simW := socle_simple W.
+have [X sXG [def1 _]] := Clifford_basis simW.
+rewrite reducible_Socle1 //; apply: mxsemisimple_reducible.
+apply: intro_mxsemisimple def1 _ => x /(subsetP sXG) Gx _.
+exact: Clifford_simple.
+Qed.
+
+Lemma Clifford_rank_components (W : sH) : (#|sH| * \rank W)%N = n.
+Proof.
+rewrite -{9}(mxrank1 F n) -Clifford_Socle1.
+rewrite (mxdirectP (Socle_direct sH)) /= -sum_nat_const.
+apply: eq_bigr => W1 _; have [W0 _ W0G] := imsetP Clifford_atrans.
+have{W0G} W0G W': W' \in orbit 'Cl G W0 by rewrite -W0G inE.
+have [/orbitP[x Gx <-] /orbitP[y Gy <-]] := (W0G W, W0G W1).
+by rewrite !{1}val_Clifford_act // !mxrankMfree // !repr_mx_free.
+Qed.
+
+Theorem Clifford_component_basis M : mxsimple rH M ->
+ {t : nat & {x_ : sH -> 'I_t -> gT |
+ forall W, let sW := (\sum_j M *m rG (x_ W j))%MS in
+ [/\ forall j, x_ W j \in G, (sW :=: W)%MS & mxdirect sW]}}.
+Proof.
+move=> simM; pose t := (n %/ #|sH| %/ \rank M)%N; exists t.
+have [X /subsetP sXG [defX1 dxX1]] := Clifford_basis simM.
+pose sMv (W : sH) x := (M *m rG x <= W)%MS; pose Xv := [pred x in X | sMv _ x].
+have sXvG W: {subset Xv W <= G} by move=> x /andP[/sXG].
+have defW W: (\sum_(x in Xv W) M *m rG x :=: W)%MS.
+ apply/eqmxP; rewrite -(geq_leqif (mxrank_leqif_eq _)); last first.
+ by apply/sumsmx_subP=> x /andP[].
+ rewrite -(leq_add2r (\sum_(W' | W' != W) \rank W')) -((bigD1 W) predT) //=.
+ rewrite -(mxdirectP (Socle_direct sH)) /= -/(Socle _) Clifford_Socle1 -defX1.
+ apply: leq_trans (mxrankS _) (mxrank_sum_leqif _).1 => /=.
+ rewrite (bigID (sMv W))%MS addsmxS //=.
+ apply/sumsmx_subP=> x /andP[Xx notW_Mx]; have Gx := sXG x Xx.
+ have simMx := Clifford_simple simM Gx.
+ pose Wx := PackSocle (component_socle sH simMx).
+ have sMxWx: (M *m rG x <= Wx)%MS by rewrite PackSocleK component_mx_id.
+ by rewrite (sumsmx_sup Wx) //; apply: contra notW_Mx => /eqP <-.
+have dxXv W: mxdirect (\sum_(x in Xv W) M *m rG x).
+ move: dxX1; rewrite !mxdirectE /= !(bigID (sMv W) (mem X)) /=.
+ by rewrite -mxdirectE mxdirect_addsE /= => /andP[].
+have def_t W: #|Xv W| = t.
+ rewrite /t -{1}(Clifford_rank_components W) mulKn 1?(cardD1 W) //.
+ rewrite -defW (mxdirectP (dxXv W)) /= (eq_bigr (fun _ => \rank M)) => [|x].
+ rewrite sum_nat_const mulnK //; last by rewrite lt0n mxrank_eq0; case simM.
+ by move/sXvG=> Gx; rewrite mxrankMfree // row_free_unit repr_mx_unit.
+exists (fun W i => enum_val (cast_ord (esym (def_t W)) i)) => W.
+case: {def_t}t / (def_t W) => sW.
+case: (pickP (Xv W)) => [x0 XvWx0 | XvW0]; last first.
+ by case/negP: (nz_socle W); rewrite -submx0 -defW big_pred0.
+have{x0 XvWx0} reXv := reindex _ (enum_val_bij_in XvWx0).
+have def_sW: (sW :=: W)%MS.
+ apply: eqmx_trans (defW W); apply/eqmxP; apply/genmxP; congr <<_>>%MS.
+ rewrite reXv /=; apply: eq_big => [j | j _]; first by have:= enum_valP j.
+ by rewrite cast_ord_id.
+split=> // [j|]; first by rewrite (sXvG W) ?enum_valP.
+apply/mxdirectP; rewrite def_sW -(defW W) /= (mxdirectP (dxXv W)) /= reXv /=.
+by apply: eq_big => [j | j _]; [move: (enum_valP j) | rewrite cast_ord_id].
+Qed.
+
+Lemma Clifford_astab : H <*> 'C_G(H) \subset 'C([set: sH] | 'Cl).
+Proof.
+rewrite join_subG !subsetI sHG subsetIl /=; apply/andP; split.
+ apply/subsetP=> h Hh; have Gh := subsetP sHG h Hh; rewrite inE.
+ apply/subsetP=> W _; have simW := socle_simple W; have [modW _ _] := simW.
+ have simWh: mxsimple rH (socle_base W *m rG h) by exact: Clifford_simple.
+ rewrite inE -val_eqE /= PackSocleK eq_sym.
+ apply/component_mx_isoP; rewrite ?subgK //; apply: component_mx_iso => //.
+ by apply: submx_trans (component_mx_id simW); move/mxmoduleP: modW => ->.
+apply/subsetP=> z cHz; have [Gz _] := setIP cHz; rewrite inE.
+apply/subsetP=> W _; have simW := socle_simple W; have [modW _ _] := simW.
+have simWz: mxsimple rH (socle_base W *m rG z) by exact: Clifford_simple.
+rewrite inE -val_eqE /= PackSocleK eq_sym.
+by apply/component_mx_isoP; rewrite ?subgK //; exact: Clifford_iso.
+Qed.
+
+Lemma Clifford_astab1 (W : sH) : 'C[W | 'Cl] = rstabs rG W.
+Proof.
+apply/setP=> x; rewrite !inE; apply: andb_id2l => Gx.
+rewrite sub1set inE (sameP eqP socleP) !val_Clifford_act //.
+rewrite andb_idr // => sWxW; rewrite -mxrank_leqif_sup //.
+by rewrite mxrankMfree ?repr_mx_free.
+Qed.
+
+Lemma Clifford_rstabs_simple (W : sH) :
+ mxsimple (subg_repr rG (rstabs_sub rG W)) W.
+Proof.
+split => [||U modU sUW nzU]; last 2 [exact: nz_socle].
+ by rewrite /mxmodule rstabs_subg setIid.
+have modUH: mxmodule rH U.
+ apply/mxmoduleP=> h Hh; rewrite (mxmoduleP modU) //.
+ rewrite /= -Clifford_astab1 !(inE, sub1set) (subsetP sHG) //.
+ rewrite (astab_act (subsetP Clifford_astab h _)) ?inE //=.
+ by rewrite mem_gen // inE Hh.
+apply: (mxsimple_exists modUH nzU) => [[M simM sMU]].
+have [t [x_ /(_ W)[Gx_ defW _]]] := Clifford_component_basis simM.
+rewrite -defW; apply/sumsmx_subP=> j _; set x := x_ W j.
+have{Gx_} Gx: x \in G by rewrite Gx_.
+apply: submx_trans (submxMr _ sMU) _; apply: (mxmoduleP modU).
+rewrite inE -val_Clifford_act Gx //; set Wx := 'Cl%act W x.
+have [-> //= | neWxW] := eqVneq Wx W.
+case: (simM) => _ /negP[]; rewrite -submx0.
+rewrite (canF_eq (actKin 'Cl Gx)) in neWxW.
+rewrite -(component_mx_disjoint _ _ neWxW); try exact: socle_simple.
+rewrite sub_capmx {1}(submx_trans sMU sUW) val_Clifford_act ?groupV //.
+by rewrite -(eqmxMr _ defW) sumsmxMr (sumsmx_sup j) ?repr_mxK.
+Qed.
+
+End Clifford.
+
+Section JordanHolder.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Variables (n : nat) (rG : mx_representation F G n).
+Local Notation modG := ((mxmodule rG) n).
+
+Lemma section_module (U V : 'M_n) (modU : modG U) (modV : modG V) :
+ mxmodule (factmod_repr modU) <<in_factmod U V>>%MS.
+Proof.
+by rewrite (eqmx_module _ (genmxE _)) in_factmod_module addsmx_module.
+Qed.
+
+Definition section_repr U V (modU : modG U) (modV : modG V) :=
+ submod_repr (section_module modU modV).
+
+Lemma mx_factmod_sub U modU :
+ mx_rsim (@section_repr U _ modU (mxmodule1 rG)) (factmod_repr modU).
+Proof.
+exists (val_submod 1%:M) => [||x Gx].
+- apply: (@addIn (\rank U)); rewrite genmxE mxrank_in_factmod mxrank_coker.
+ by rewrite (addsmx_idPr (submx1 U)) mxrank1 subnK ?rank_leq_row.
+- by rewrite /row_free val_submod1.
+by rewrite -[_ x]mul1mx -val_submodE val_submodJ.
+Qed.
+
+Definition max_submod (U V : 'M_n) :=
+ (U < V)%MS /\ (forall W, ~ [/\ modG W, U < W & W < V])%MS.
+
+Lemma max_submodP U V (modU : modG U) (modV : modG V) :
+ (U <= V)%MS -> (max_submod U V <-> mx_irreducible (section_repr modU modV)).
+Proof.
+move=> sUV; split=> [[ltUV maxU] | ].
+ apply/mx_irrP; split=> [|WU modWU nzWU].
+ by rewrite genmxE lt0n mxrank_eq0 in_factmod_eq0; case/andP: ltUV.
+ rewrite -sub1mx -val_submodS val_submod1 genmxE.
+ pose W := (U + val_factmod (val_submod WU))%MS.
+ suffices sVW: (V <= W)%MS.
+ rewrite {2}in_factmodE (submx_trans (submxMr _ sVW)) //.
+ rewrite addsmxMr -!in_factmodE val_factmodK.
+ by rewrite ((in_factmod U U =P 0) _) ?adds0mx ?in_factmod_eq0.
+ move/and3P: {maxU}(maxU W); apply: contraR; rewrite /ltmx addsmxSl => -> /=.
+ move: modWU; rewrite /mxmodule rstabs_submod rstabs_factmod => -> /=.
+ rewrite addsmx_sub submx_refl -in_factmod_eq0 val_factmodK.
+ move: nzWU; rewrite -[_ == 0](inj_eq val_submod_inj) linear0 => ->.
+ rewrite -(in_factmodsK sUV) addsmxS // val_factmodS.
+ by rewrite -(genmxE (in_factmod U V)) val_submodP.
+case/mx_irrP; rewrite lt0n {1}genmxE mxrank_eq0 in_factmod_eq0 => ltUV maxV.
+split=> // [|W [modW /andP[sUW ltUW] /andP[sWV /negP[]]]]; first exact/andP.
+rewrite -(in_factmodsK sUV) -(in_factmodsK sUW) addsmxS // val_factmodS.
+rewrite -genmxE -val_submod1; set VU := <<_>>%MS.
+have sW_VU: (in_factmod U W <= VU)%MS.
+ by rewrite genmxE -val_factmodS !submxMr.
+rewrite -(in_submodK sW_VU) val_submodS -(genmxE (in_submod _ _)).
+rewrite sub1mx maxV //.
+ rewrite (eqmx_module _ (genmxE _)) in_submod_module ?genmxE ?submxMr //.
+ by rewrite in_factmod_module addsmx_module.
+rewrite -submx0 [(_ <= 0)%MS]genmxE -val_submodS linear0 in_submodK //.
+by rewrite eqmx0 submx0 in_factmod_eq0.
+Qed.
+
+Lemma max_submod_eqmx U1 U2 V1 V2 :
+ (U1 :=: U2)%MS -> (V1 :=: V2)%MS -> max_submod U1 V1 -> max_submod U2 V2.
+Proof.
+move=> eqU12 eqV12 [ltUV1 maxU1].
+by split=> [|W]; rewrite -(lt_eqmx eqU12) -(lt_eqmx eqV12).
+Qed.
+
+Definition mx_subseries := all modG.
+
+Definition mx_composition_series V :=
+ mx_subseries V /\ (forall i, i < size V -> max_submod (0 :: V)`_i V`_i).
+Local Notation mx_series := mx_composition_series.
+
+Fact mx_subseries_module V i : mx_subseries V -> mxmodule rG V`_i.
+Proof.
+move=> modV; have [|leVi] := ltnP i (size V); first exact: all_nthP.
+by rewrite nth_default ?mxmodule0.
+Qed.
+
+Fact mx_subseries_module' V i : mx_subseries V -> mxmodule rG (0 :: V)`_i.
+Proof. by move=> modV; rewrite mx_subseries_module //= mxmodule0. Qed.
+
+Definition subseries_repr V i (modV : all modG V) :=
+ section_repr (mx_subseries_module' i modV) (mx_subseries_module i modV).
+
+Definition series_repr V i (compV : mx_composition_series V) :=
+ subseries_repr i (proj1 compV).
+
+Lemma mx_series_lt V : mx_composition_series V -> path ltmx 0 V.
+Proof. by case=> _ compV; apply/(pathP 0)=> i /compV[]. Qed.
+
+Lemma max_size_mx_series (V : seq 'M[F]_n) :
+ path ltmx 0 V -> size V <= \rank (last 0 V).
+Proof.
+rewrite -[size V]addn0 -(mxrank0 F n n); elim: V 0 => //= V1 V IHV V0.
+rewrite ltmxErank -andbA => /and3P[_ ltV01 ltV].
+by apply: leq_trans (IHV _ ltV); rewrite addSnnS leq_add2l.
+Qed.
+
+Lemma mx_series_repr_irr V i (compV : mx_composition_series V) :
+ i < size V -> mx_irreducible (series_repr i compV).
+Proof.
+case: compV => modV compV /compV maxVi; apply/max_submodP => //.
+by apply: ltmxW; case: maxVi.
+Qed.
+
+Lemma mx_series_rcons U V :
+ mx_series (rcons U V) <-> [/\ mx_series U, modG V & max_submod (last 0 U) V].
+Proof.
+rewrite /mx_series /mx_subseries all_rcons size_rcons -rcons_cons.
+split=> [ [/andP[modU modV] maxU] | [[modU maxU] modV maxV]].
+ split=> //; last first.
+ by have:= maxU _ (leqnn _); rewrite !nth_rcons leqnn ltnn eqxx -last_nth.
+ by split=> // i ltiU; have:= maxU i (ltnW ltiU); rewrite !nth_rcons leqW ltiU.
+rewrite modV; split=> // i; rewrite !nth_rcons ltnS leq_eqVlt.
+case: eqP => [-> _ | /= _ ltiU]; first by rewrite ltnn ?eqxx -last_nth.
+by rewrite ltiU; exact: maxU.
+Qed.
+
+Theorem mx_Schreier U :
+ mx_subseries U -> path ltmx 0 U ->
+ classically (exists V, [/\ mx_series V, last 0 V :=: 1%:M & subseq U V])%MS.
+Proof.
+move: U => U0; set U := {1 2}U0; have: subseq U0 U := subseq_refl U.
+pose n' := n.+1; have: n < size U + n' by rewrite leq_addl.
+elim: n' U => [|n' IH_U] U ltUn' sU0U modU incU [] // noV.
+ rewrite addn0 ltnNge in ltUn'; case/negP: ltUn'.
+ by rewrite (leq_trans (max_size_mx_series incU)) ?rank_leq_row.
+apply: (noV); exists U; split => //; first split=> // i lt_iU; last first.
+ apply/eqmxP; apply: contraT => neU1.
+ apply: {IH_U}(IH_U (rcons U 1%:M)) noV.
+ - by rewrite size_rcons addSnnS.
+ - by rewrite (subseq_trans sU0U) ?subseq_rcons.
+ - by rewrite /mx_subseries all_rcons mxmodule1.
+ by rewrite rcons_path ltmxEneq neU1 submx1 !andbT.
+set U'i := _`_i; set Ui := _`_i; have defU := cat_take_drop i U.
+have defU'i: U'i = last 0 (take i U).
+ rewrite (last_nth 0) /U'i -{1}defU -cat_cons nth_cat /=.
+ by rewrite size_take lt_iU leqnn.
+move: incU; rewrite -defU cat_path (drop_nth 0) //= -/Ui -defU'i.
+set U' := take i U; set U'' := drop _ U; case/and3P=> incU' ltUi incU''.
+split=> // W [modW ltUW ltWV]; case: notF.
+apply: {IH_U}(IH_U (U' ++ W :: Ui :: U'')) noV; last 2 first.
+- by rewrite /mx_subseries -drop_nth // all_cat /= modW -all_cat defU.
+- by rewrite cat_path /= -defU'i; exact/and4P.
+- by rewrite -drop_nth // size_cat /= addnS -size_cat defU addSnnS.
+by rewrite (subseq_trans sU0U) // -defU cat_subseq // -drop_nth ?subseq_cons.
+Qed.
+
+Lemma mx_second_rsim U V (modU : modG U) (modV : modG V) :
+ let modI := capmx_module modU modV in let modA := addsmx_module modU modV in
+ mx_rsim (section_repr modI modU) (section_repr modV modA).
+Proof.
+move=> modI modA; set nI := {1}(\rank _).
+have sIU := capmxSl U V; have sVA := addsmxSr U V.
+pose valI := val_factmod (val_submod (1%:M : 'M_nI)).
+have UvalI: (valI <= U)%MS.
+ rewrite -(addsmx_idPr sIU) (submx_trans _ (proj_factmodS _ _)) //.
+ by rewrite submxMr // val_submod1 genmxE.
+exists (valI *m in_factmod _ 1%:M *m in_submod _ 1%:M) => [||x Gx].
+- apply: (@addIn (\rank (U :&: V) + \rank V)%N); rewrite genmxE addnA addnCA.
+ rewrite /nI genmxE !{1}mxrank_in_factmod 2?(addsmx_idPr _) //.
+ by rewrite -mxrank_sum_cap addnC.
+- rewrite -kermx_eq0; apply/rowV0P=> u; rewrite (sameP sub_kermxP eqP).
+ rewrite mulmxA -in_submodE mulmxA -in_factmodE -(inj_eq val_submod_inj).
+ rewrite linear0 in_submodK ?in_factmod_eq0 => [Vvu|]; last first.
+ by rewrite genmxE addsmxC in_factmod_addsK submxMr // mulmx_sub.
+ apply: val_submod_inj; apply/eqP; rewrite linear0 -[val_submod u]val_factmodK.
+ rewrite val_submodE val_factmodE -mulmxA -val_factmodE -/valI.
+ by rewrite in_factmod_eq0 sub_capmx mulmx_sub.
+symmetry; rewrite -{1}in_submodE -{1}in_submodJ; last first.
+ by rewrite genmxE addsmxC in_factmod_addsK -in_factmodE submxMr.
+rewrite -{1}in_factmodE -{1}in_factmodJ // mulmxA in_submodE; congr (_ *m _).
+apply/eqP; rewrite mulmxA -in_factmodE -subr_eq0 -linearB in_factmod_eq0.
+apply: submx_trans (capmxSr U V); rewrite -in_factmod_eq0 linearB /=.
+rewrite subr_eq0 {1}(in_factmodJ modI) // val_factmodK eq_sym.
+rewrite /valI val_factmodE mulmxA -val_factmodE val_factmodK.
+by rewrite -[submod_mx _ _]mul1mx -val_submodE val_submodJ.
+Qed.
+
+Lemma section_eqmx_add U1 U2 V1 V2 modU1 modU2 modV1 modV2 :
+ (U1 :=: U2)%MS -> (U1 + V1 :=: U2 + V2)%MS ->
+ mx_rsim (@section_repr U1 V1 modU1 modV1) (@section_repr U2 V2 modU2 modV2).
+Proof.
+move=> eqU12 eqV12; set n1 := {1}(\rank _).
+pose v1 := val_factmod (val_submod (1%:M : 'M_n1)).
+have sv12: (v1 <= U2 + V2)%MS.
+ rewrite -eqV12 (submx_trans _ (proj_factmodS _ _)) //.
+ by rewrite submxMr // val_submod1 genmxE.
+exists (v1 *m in_factmod _ 1%:M *m in_submod _ 1%:M) => [||x Gx].
+- apply: (@addIn (\rank U1)); rewrite {2}eqU12 /n1 !{1}genmxE.
+ by rewrite !{1}mxrank_in_factmod eqV12.
+- rewrite -kermx_eq0; apply/rowV0P=> u; rewrite (sameP sub_kermxP eqP) mulmxA.
+ rewrite -in_submodE mulmxA -in_factmodE -(inj_eq val_submod_inj) linear0.
+ rewrite in_submodK ?in_factmod_eq0 -?eqU12 => [U1uv1|]; last first.
+ by rewrite genmxE -(in_factmod_addsK U2 V2) submxMr // mulmx_sub.
+ apply: val_submod_inj; apply/eqP; rewrite linear0 -[val_submod _]val_factmodK.
+ by rewrite in_factmod_eq0 val_factmodE val_submodE -mulmxA -val_factmodE.
+symmetry; rewrite -{1}in_submodE -{1}in_factmodE -{1}in_submodJ; last first.
+ by rewrite genmxE -(in_factmod_addsK U2 V2) submxMr.
+rewrite -{1}in_factmodJ // mulmxA in_submodE; congr (_ *m _); apply/eqP.
+rewrite mulmxA -in_factmodE -subr_eq0 -linearB in_factmod_eq0 -eqU12.
+rewrite -in_factmod_eq0 linearB /= subr_eq0 {1}(in_factmodJ modU1) //.
+rewrite val_factmodK /v1 val_factmodE eq_sym mulmxA -val_factmodE val_factmodK.
+by rewrite -[_ *m _]mul1mx mulmxA -val_submodE val_submodJ.
+Qed.
+
+Lemma section_eqmx U1 U2 V1 V2 modU1 modU2 modV1 modV2
+ (eqU : (U1 :=: U2)%MS) (eqV : (V1 :=: V2)%MS) :
+ mx_rsim (@section_repr U1 V1 modU1 modV1) (@section_repr U2 V2 modU2 modV2).
+Proof. by apply: section_eqmx_add => //; exact: adds_eqmx. Qed.
+
+Lemma mx_butterfly U V W modU modV modW :
+ ~~ (U == V)%MS -> max_submod U W -> max_submod V W ->
+ let modUV := capmx_module modU modV in
+ max_submod (U :&: V)%MS U
+ /\ mx_rsim (@section_repr V W modV modW) (@section_repr _ U modUV modU).
+Proof.
+move=> neUV maxU maxV modUV; have{neUV maxU} defW: (U + V :=: W)%MS.
+ wlog{neUV modUV} ltUV: U V modU modV maxU maxV / ~~ (V <= U)%MS.
+ by case/nandP: neUV => ?; first rewrite addsmxC; exact.
+ apply/eqmxP/idPn=> neUVW; case: maxU => ltUW; case/(_ (U + V)%MS).
+ rewrite addsmx_module // ltmxE ltmxEneq neUVW addsmxSl !addsmx_sub.
+ by have [ltVW _] := maxV; rewrite submx_refl andbT ltUV !ltmxW.
+have sUV_U := capmxSl U V; have sVW: (V <= W)%MS by rewrite -defW addsmxSr.
+set goal := mx_rsim _ _; suffices{maxV} simUV: goal.
+ split=> //; apply/(max_submodP modUV modU sUV_U).
+ by apply: mx_rsim_irr simUV _; exact/max_submodP.
+apply: {goal}mx_rsim_sym.
+by apply: mx_rsim_trans (mx_second_rsim modU modV) _; exact: section_eqmx.
+Qed.
+
+Lemma mx_JordanHolder_exists U V :
+ mx_composition_series U -> modG V -> max_submod V (last 0 U) ->
+ {W : seq 'M_n | mx_composition_series W & last 0 W = V}.
+Proof.
+elim/last_ind: U V => [|U Um IHU] V compU modV; first by case; rewrite ltmx0.
+rewrite last_rcons => maxV; case/mx_series_rcons: compU => compU modUm maxUm.
+case eqUV: (last 0 U == V)%MS.
+ case/lastP: U eqUV compU {maxUm IHU} => [|U' Um'].
+ by rewrite andbC; move/eqmx0P->; exists [::].
+ rewrite last_rcons; move/eqmxP=> eqU'V; case/mx_series_rcons=> compU _ maxUm'.
+ exists (rcons U' V); last by rewrite last_rcons.
+ apply/mx_series_rcons; split => //; exact: max_submod_eqmx maxUm'.
+set Um' := last 0 U in maxUm eqUV; have [modU _] := compU.
+have modUm': modG Um' by rewrite /Um' (last_nth 0) mx_subseries_module'.
+have [|||W compW lastW] := IHU (V :&: Um')%MS; rewrite ?capmx_module //.
+ by case: (mx_butterfly modUm' modV modUm); rewrite ?eqUV // {1}capmxC.
+exists (rcons W V); last by rewrite last_rcons.
+apply/mx_series_rcons; split; rewrite // lastW.
+by case: (mx_butterfly modV modUm' modUm); rewrite // andbC eqUV.
+Qed.
+
+Let rsim_rcons U V compU compUV i : i < size U ->
+ mx_rsim (@series_repr U i compU) (@series_repr (rcons U V) i compUV).
+Proof.
+by move=> ltiU; apply: section_eqmx; rewrite -?rcons_cons nth_rcons ?leqW ?ltiU.
+Qed.
+
+Let last_mod U (compU : mx_series U) : modG (last 0 U).
+Proof.
+by case: compU => modU _; rewrite (last_nth 0) (mx_subseries_module' _ modU).
+Qed.
+
+Let rsim_last U V modUm modV compUV :
+ mx_rsim (@section_repr (last 0 U) V modUm modV)
+ (@series_repr (rcons U V) (size U) compUV).
+Proof.
+apply: section_eqmx; last by rewrite nth_rcons ltnn eqxx.
+by rewrite -rcons_cons nth_rcons leqnn -last_nth.
+Qed.
+Local Notation rsimT := mx_rsim_trans.
+Local Notation rsimC := mx_rsim_sym.
+
+Lemma mx_JordanHolder U V compU compV :
+ let m := size U in (last 0 U :=: last 0 V)%MS ->
+ m = size V /\ (exists p : 'S_m, forall i : 'I_m,
+ mx_rsim (@series_repr U i compU) (@series_repr V (p i) compV)).
+Proof.
+elim: {U}(size U) {-2}U V (eqxx (size U)) compU compV => /= [|r IHr] U V.
+ move/nilP->; case/lastP: V => [|V Vm] /= ? compVm; rewrite ?last_rcons => Vm0.
+ by split=> //; exists 1%g; case.
+ by case/mx_series_rcons: (compVm) => _ _ []; rewrite -(lt_eqmx Vm0) ltmx0.
+case/lastP: U => // [U Um]; rewrite size_rcons eqSS => szUr compUm.
+case/mx_series_rcons: (compUm); set Um' := last 0 U => compU modUm maxUm.
+case/lastP: V => [|V Vm] compVm; rewrite ?last_rcons ?size_rcons /= => eqUVm.
+ by case/mx_series_rcons: (compUm) => _ _ []; rewrite (lt_eqmx eqUVm) ltmx0.
+case/mx_series_rcons: (compVm); set Vm' := last 0 V => compV modVm maxVm.
+have [modUm' modVm']: modG Um' * modG Vm' := (last_mod compU, last_mod compV).
+pose i_m := @ord_max (size U).
+have [eqUVm' | neqUVm'] := altP (@eqmxP _ _ _ _ Um' Vm').
+ have [szV [p sim_p]] := IHr U V szUr compU compV eqUVm'.
+ split; first by rewrite szV.
+ exists (lift_perm i_m i_m p) => i; case: (unliftP i_m i) => [j|] ->{i}.
+ apply: rsimT (rsimC _) (rsimT (sim_p j) _).
+ by rewrite lift_max; exact: rsim_rcons.
+ by rewrite lift_perm_lift lift_max; apply: rsim_rcons; rewrite -szV.
+ have simUVm := section_eqmx modUm' modVm' modUm modVm eqUVm' eqUVm.
+ apply: rsimT (rsimC _) (rsimT simUVm _); first exact: rsim_last.
+ by rewrite lift_perm_id /= szV; exact: rsim_last.
+have maxVUm: max_submod Vm' Um by exact: max_submod_eqmx (eqmx_sym _) maxVm.
+have:= mx_butterfly modUm' modVm' modUm neqUVm' maxUm maxVUm.
+move: (capmx_module _ _); set Wm := (Um' :&: Vm')%MS => modWm [maxWUm simWVm].
+have:= mx_butterfly modVm' modUm' modUm _ maxVUm maxUm.
+move: (capmx_module _ _); rewrite andbC capmxC -/Wm => modWmV [// | maxWVm].
+rewrite {modWmV}(bool_irrelevance modWmV modWm) => simWUm.
+have [W compW lastW] := mx_JordanHolder_exists compU modWm maxWUm.
+have compWU: mx_series (rcons W Um') by apply/mx_series_rcons; rewrite lastW.
+have compWV: mx_series (rcons W Vm') by apply/mx_series_rcons; rewrite lastW.
+have [|szW [pU pUW]] := IHr U _ szUr compU compWU; first by rewrite last_rcons.
+rewrite size_rcons in szW; have ltWU: size W < size U by rewrite -szW.
+have{IHr} := IHr _ V _ compWV compV; rewrite last_rcons size_rcons -szW.
+case=> {r szUr}// szV [pV pWV]; split; first by rewrite szV.
+pose j_m := Ordinal ltWU; pose i_m' := lift i_m j_m.
+exists (lift_perm i_m i_m pU * tperm i_m i_m' * lift_perm i_m i_m pV)%g => i.
+rewrite !permM; case: (unliftP i_m i) => [j {simWUm}|] ->{i}; last first.
+ rewrite lift_perm_id tpermL lift_perm_lift lift_max {simWVm}.
+ apply: rsimT (rsimT (pWV j_m) _); last by apply: rsim_rcons; rewrite -szV.
+ apply: rsimT (rsimC _) {simWUm}(rsimT simWUm _); first exact: rsim_last.
+ by rewrite -lastW in modWm *; exact: rsim_last.
+apply: rsimT (rsimC _) {pUW}(rsimT (pUW j) _).
+ by rewrite lift_max; exact: rsim_rcons.
+rewrite lift_perm_lift; case: (unliftP j_m (pU j)) => [k|] ->{j pU}.
+ rewrite tpermD ?(inj_eq (@lift_inj _ _)) ?neq_lift //.
+ rewrite lift_perm_lift !lift_max; set j := lift j_m k.
+ have ltjW: j < size W by have:= ltn_ord k; rewrite -(lift_max k) /= {1 3}szW.
+ apply: rsimT (rsimT (pWV j) _); last by apply: rsim_rcons; rewrite -szV.
+ by apply: rsimT (rsimC _) (rsim_rcons compW _ _); first exact: rsim_rcons.
+apply: rsimT {simWVm}(rsimC (rsimT simWVm _)) _.
+ by rewrite -lastW in modWm *; exact: rsim_last.
+rewrite tpermR lift_perm_id /= szV.
+by apply: rsimT (rsim_last modVm' modVm _); exact: section_eqmx.
+Qed.
+
+Lemma mx_JordanHolder_max U (m := size U) V compU modV :
+ (last 0 U :=: 1%:M)%MS -> mx_irreducible (@factmod_repr _ G n rG V modV) ->
+ exists i : 'I_m, mx_rsim (factmod_repr modV) (@series_repr U i compU).
+Proof.
+rewrite {}/m; set Um := last 0 U => Um1 irrV.
+have modUm: modG Um := last_mod compU; have simV := rsimC (mx_factmod_sub modV).
+have maxV: max_submod V Um.
+ move/max_submodP: (mx_rsim_irr simV irrV) => /(_ (submx1 _)).
+ by apply: max_submod_eqmx; last exact: eqmx_sym.
+have [W compW lastW] := mx_JordanHolder_exists compU modV maxV.
+have compWU: mx_series (rcons W Um) by apply/mx_series_rcons; rewrite lastW.
+have:= mx_JordanHolder compU compWU; rewrite last_rcons size_rcons.
+case=> // szW [p pUW]; have ltWU: size W < size U by rewrite szW.
+pose i := Ordinal ltWU; exists ((p^-1)%g i).
+apply: rsimT simV (rsimT _ (rsimC (pUW _))); rewrite permKV.
+apply: rsimT (rsimC _) (rsim_last (last_mod compW) modUm _).
+by apply: section_eqmx; rewrite ?lastW.
+Qed.
+
+End JordanHolder.
+
+Bind Scope irrType_scope with socle_sort.
+
+Section Regular.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Local Notation nG := #|pred_of_set (gval G)|.
+
+Local Notation rF := (GRing.Field.comUnitRingType F) (only parsing).
+Local Notation aG := (regular_repr rF G).
+Local Notation R_G := (group_ring rF G).
+
+Lemma gring_free : row_free R_G.
+Proof.
+apply/row_freeP; exists (lin1_mx (row (gring_index G 1) \o vec_mx)).
+apply/row_matrixP=> i; rewrite row_mul rowK mul_rV_lin1 /= mxvecK rowK row1.
+by rewrite gring_indexK // mul1g gring_valK.
+Qed.
+
+Lemma gring_op_id A : (A \in R_G)%MS -> gring_op aG A = A.
+Proof.
+case/envelop_mxP=> a ->{A}; rewrite linear_sum.
+by apply: eq_bigr => x Gx; rewrite linearZ /= gring_opG.
+Qed.
+
+Lemma gring_rowK A : (A \in R_G)%MS -> gring_mx aG (gring_row A) = A.
+Proof. exact: gring_op_id. Qed.
+
+Lemma mem_gring_mx m a (M : 'M_(m, nG)) :
+ (gring_mx aG a \in M *m R_G)%MS = (a <= M)%MS.
+Proof. by rewrite vec_mxK submxMfree ?gring_free. Qed.
+
+Lemma mem_sub_gring m A (M : 'M_(m, nG)) :
+ (A \in M *m R_G)%MS = (A \in R_G)%MS && (gring_row A <= M)%MS.
+Proof.
+rewrite -(andb_idl (memmx_subP (submxMl _ _) A)); apply: andb_id2l => R_A.
+by rewrite -mem_gring_mx gring_rowK.
+Qed.
+
+Section GringMx.
+
+Variables (n : nat) (rG : mx_representation F G n).
+
+Lemma gring_mxP a : (gring_mx rG a \in enveloping_algebra_mx rG)%MS.
+Proof. by rewrite vec_mxK submxMl. Qed.
+
+Lemma gring_opM A B :
+ (B \in R_G)%MS -> gring_op rG (A *m B) = gring_op rG A *m gring_op rG B.
+Proof. by move=> R_B; rewrite -gring_opJ gring_rowK. Qed.
+
+Hypothesis irrG : mx_irreducible rG.
+
+Lemma rsim_regular_factmod :
+ {U : 'M_nG & {modU : mxmodule aG U & mx_rsim rG (factmod_repr modU)}}.
+Proof.
+pose v : 'rV[F]_n := nz_row 1%:M.
+pose fU := lin1_mx (mulmx v \o gring_mx rG); pose U := kermx fU.
+have modU: mxmodule aG U.
+ apply/mxmoduleP => x Gx; apply/sub_kermxP/row_matrixP=> i.
+ rewrite 2!row_mul row0; move: (row i U) (sub_kermxP (row_sub i U)) => u.
+ by rewrite !mul_rV_lin1 /= gring_mxJ // mulmxA => ->; rewrite mul0mx.
+have def_n: \rank (cokermx U) = n.
+ apply/eqP; rewrite mxrank_coker mxrank_ker subKn ?rank_leq_row // -genmxE.
+ rewrite -[_ == _]sub1mx; have [_ _ ->] := irrG; rewrite ?submx1 //.
+ rewrite (eqmx_module _ (genmxE _)); apply/mxmoduleP=> x Gx.
+ apply/row_subP=> i; apply: eq_row_sub (gring_index G (enum_val i * x)) _.
+ rewrite !rowE mulmxA !mul_rV_lin1 /= -mulmxA -gring_mxJ //.
+ by rewrite -rowE rowK.
+ rewrite (eqmx_eq0 (genmxE _)); apply/rowV0Pn.
+ exists v; last exact: (nz_row_mxsimple irrG).
+ apply/submxP; exists (gring_row (aG 1%g)); rewrite mul_rV_lin1 /=.
+ by rewrite -gring_opE gring_opG // repr_mx1 mulmx1.
+exists U; exists modU; apply: mx_rsim_sym.
+exists (val_factmod 1%:M *m fU) => // [|x Gx].
+ rewrite /row_free eqn_leq rank_leq_row /= -subn_eq0 -mxrank_ker mxrank_eq0.
+ apply/rowV0P=> u /sub_kermxP; rewrite mulmxA => /sub_kermxP.
+ by rewrite -/U -in_factmod_eq0 mulmxA mulmx1 val_factmodK => /eqP.
+rewrite mulmxA -val_factmodE (canRL (addKr _) (add_sub_fact_mod U _)).
+rewrite mulmxDl mulNmx (sub_kermxP (val_submodP _)) oppr0 add0r.
+apply/row_matrixP=> i; move: (val_factmod _) => zz.
+by rewrite !row_mul !mul_rV_lin1 /= gring_mxJ // mulmxA.
+Qed.
+
+Lemma rsim_regular_series U (compU : mx_composition_series aG U) :
+ (last 0 U :=: 1%:M)%MS ->
+ exists i : 'I_(size U), mx_rsim rG (series_repr i compU).
+Proof.
+move=> lastU; have [V [modV simGV]] := rsim_regular_factmod.
+have irrV := mx_rsim_irr simGV irrG.
+have [i simVU] := mx_JordanHolder_max compU lastU irrV.
+exists i; exact: mx_rsim_trans simGV simVU.
+Qed.
+
+Hypothesis F'G : [char F]^'.-group G.
+
+Lemma rsim_regular_submod :
+ {U : 'M_nG & {modU : mxmodule aG U & mx_rsim rG (submod_repr modU)}}.
+Proof.
+have [V [modV eqG'V]] := rsim_regular_factmod.
+have [U modU defVU dxVU] := mx_Maschke F'G modV (submx1 V).
+exists U; exists modU; apply: mx_rsim_trans eqG'V _.
+by apply: mx_rsim_factmod; rewrite ?mxdirectE /= addsmxC // addnC.
+Qed.
+
+End GringMx.
+
+Definition gset_mx (A : {set gT}) := \sum_(x in A) aG x.
+
+Local Notation tG := #|pred_of_set (classes (gval G))|.
+
+Definition classg_base := \matrix_(k < tG) mxvec (gset_mx (enum_val k)).
+
+Let groupCl : {in G, forall x, {subset x ^: G <= G}}.
+Proof. by move=> x Gx; apply: subsetP; exact: class_subG. Qed.
+
+Lemma classg_base_free : row_free classg_base.
+Proof.
+rewrite -kermx_eq0; apply/rowV0P=> v /sub_kermxP; rewrite mulmx_sum_row => v0.
+apply/rowP=> k; rewrite mxE.
+have [x Gx def_k] := imsetP (enum_valP k).
+transitivity (@gring_proj F _ G x (vec_mx 0) 0 0); last first.
+ by rewrite !linear0 !mxE.
+rewrite -{}v0 !linear_sum (bigD1 k) //= !linearZ /= rowK mxvecK def_k.
+rewrite linear_sum (bigD1 x) ?class_refl //= gring_projE // eqxx.
+rewrite !big1 ?addr0 ?mxE ?mulr1 // => [k' | y /andP[xGy ne_yx]]; first 1 last.
+ by rewrite gring_projE ?(groupCl Gx xGy) // eq_sym (negPf ne_yx).
+rewrite rowK !linearZ /= mxvecK -(inj_eq enum_val_inj) def_k eq_sym.
+have [z Gz ->] := imsetP (enum_valP k').
+move/eqP=> not_Gxz; rewrite linear_sum big1 ?scaler0 //= => y zGy.
+rewrite gring_projE ?(groupCl Gz zGy) //.
+by case: eqP zGy => // <- /class_transr.
+Qed.
+
+Lemma classg_base_center : (classg_base :=: 'Z(R_G))%MS.
+Proof.
+apply/eqmxP/andP; split.
+ apply/row_subP=> k; rewrite rowK /gset_mx sub_capmx {1}linear_sum.
+ have [x Gx ->{k}] := imsetP (enum_valP k); have sxGG := groupCl Gx.
+ rewrite summx_sub => [|y xGy]; last by rewrite envelop_mx_id ?sxGG.
+ rewrite memmx_cent_envelop; apply/centgmxP=> y Gy.
+ rewrite {2}(reindex_acts 'J _ Gy) ?astabsJ ?class_norm //=.
+ rewrite mulmx_suml mulmx_sumr; apply: eq_bigr => z; move/sxGG=> Gz.
+ by rewrite -!repr_mxM ?groupJ -?conjgC.
+apply/memmx_subP=> A; rewrite sub_capmx memmx_cent_envelop.
+case/andP=> /envelop_mxP[a ->{A}] cGa.
+rewrite (partition_big_imset (class^~ G)) -/(classes G) /=.
+rewrite linear_sum summx_sub //= => xG GxG; have [x Gx def_xG] := imsetP GxG.
+apply: submx_trans (scalemx_sub (a x) (submx_refl _)).
+rewrite (eq_row_sub (enum_rank_in GxG xG)) // linearZ /= rowK enum_rankK_in //.
+rewrite !linear_sum {xG GxG}def_xG; apply: eq_big => [y | xy] /=.
+ apply/idP/andP=> [| [_ xGy]]; last by rewrite -(eqP xGy) class_refl.
+ by case/imsetP=> z Gz ->; rewrite groupJ // classGidl.
+case/imsetP=> y Gy ->{xy}; rewrite linearZ; congr (_ *: _).
+move/(canRL (repr_mxK aG Gy)): (centgmxP cGa y Gy); have Gy' := groupVr Gy.
+move/(congr1 (gring_proj x)); rewrite -mulmxA mulmx_suml !linear_sum.
+rewrite (bigD1 x Gx) big1 => [|z /andP[Gz]]; rewrite !linearZ /=; last first.
+ by rewrite eq_sym gring_projE // => /negPf->; rewrite scaler0.
+rewrite gring_projE // eqxx scalemx1 (bigD1 (x ^ y)%g) ?groupJ //=.
+rewrite big1 => [|z /andP[Gz]]; rewrite -scalemxAl !linearZ /=.
+ rewrite !addr0 -!repr_mxM ?groupM // mulgA mulKVg mulgK => /rowP/(_ 0).
+ by rewrite gring_projE // eqxx scalemx1 !mxE.
+rewrite eq_sym -(can_eq (conjgKV y)) conjgK conjgE invgK.
+by rewrite -!repr_mxM ?gring_projE ?groupM // => /negPf->; rewrite scaler0.
+Qed.
+
+Lemma regular_module_ideal m (M : 'M_(m, nG)) :
+ mxmodule aG M = right_mx_ideal R_G (M *m R_G).
+Proof.
+apply/idP/idP=> modM.
+ apply/mulsmx_subP=> A B; rewrite !mem_sub_gring => /andP[R_A M_A] R_B.
+ by rewrite envelop_mxM // gring_row_mul (mxmodule_envelop modM).
+apply/mxmoduleP=> x Gx; apply/row_subP=> i; rewrite row_mul -mem_gring_mx.
+rewrite gring_mxJ // (mulsmx_subP modM) ?envelop_mx_id //.
+by rewrite mem_gring_mx row_sub.
+Qed.
+
+Definition irrType := socleType aG.
+Identity Coercion type_of_irrType : irrType >-> socleType.
+
+Variable sG : irrType.
+
+Definition irr_degree (i : sG) := \rank (socle_base i).
+Local Notation "'n_ i" := (irr_degree i) : group_ring_scope.
+Local Open Scope group_ring_scope.
+
+Lemma irr_degreeE i : 'n_i = \rank (socle_base i). Proof. by []. Qed.
+Lemma irr_degree_gt0 i : 'n_i > 0.
+Proof. by rewrite lt0n mxrank_eq0; case: (socle_simple i). Qed.
+
+Definition irr_repr i : mx_representation F G 'n_i := socle_repr i.
+Lemma irr_reprE i x : irr_repr i x = submod_mx (socle_module i) x.
+Proof. by []. Qed.
+
+Lemma rfix_regular : (rfix_mx aG G :=: gring_row (gset_mx G))%MS.
+Proof.
+apply/eqmxP/andP; split; last first.
+ apply/rfix_mxP => x Gx; rewrite -gring_row_mul; congr gring_row.
+ rewrite {2}/gset_mx (reindex_astabs 'R x) ?astabsR //= mulmx_suml.
+ by apply: eq_bigr => y Gy; rewrite repr_mxM.
+apply/rV_subP=> v /rfix_mxP cGv.
+have /envelop_mxP[a def_v]: (gring_mx aG v \in R_G)%MS.
+ by rewrite vec_mxK submxMl.
+suffices ->: v = a 1%g *: gring_row (gset_mx G) by rewrite scalemx_sub.
+rewrite -linearZ scaler_sumr -[v]gring_mxK def_v; congr (gring_row _).
+apply: eq_bigr => x Gx; congr (_ *: _).
+move/rowP/(_ 0): (congr1 (gring_proj x \o gring_mx aG) (cGv x Gx)).
+rewrite /= gring_mxJ // def_v mulmx_suml !linear_sum (bigD1 1%g) //=.
+rewrite repr_mx1 -scalemxAl mul1mx linearZ /= gring_projE // eqxx scalemx1.
+rewrite big1 ?addr0 ?mxE /= => [ | y /andP[Gy nt_y]]; last first.
+ rewrite -scalemxAl linearZ -repr_mxM //= gring_projE ?groupM //.
+ by rewrite eq_sym eq_mulgV1 mulgK (negPf nt_y) scaler0.
+rewrite (bigD1 x) //= linearZ /= gring_projE // eqxx scalemx1.
+rewrite big1 ?addr0 ?mxE // => y /andP[Gy ne_yx].
+by rewrite linearZ /= gring_projE // eq_sym (negPf ne_yx) scaler0.
+Qed.
+
+Lemma principal_comp_subproof : mxsimple aG (rfix_mx aG G).
+Proof.
+apply: linear_mxsimple; first exact: rfix_mx_module.
+apply/eqP; rewrite rfix_regular eqn_leq rank_leq_row lt0n mxrank_eq0.
+apply/eqP => /(congr1 (gring_proj 1 \o gring_mx aG)); apply/eqP.
+rewrite /= -[gring_mx _ _]/(gring_op _ _) !linear0 !linear_sum (bigD1 1%g) //=.
+rewrite gring_opG ?gring_projE // eqxx big1 ?addr0 ?oner_eq0 // => x.
+by case/andP=> Gx nt_x; rewrite gring_opG // gring_projE // eq_sym (negPf nt_x).
+Qed.
+
+Fact principal_comp_key : unit. Proof. by []. Qed.
+Definition principal_comp_def :=
+ PackSocle (component_socle sG principal_comp_subproof).
+Definition principal_comp := locked_with principal_comp_key principal_comp_def.
+Local Notation "1" := principal_comp : irrType_scope.
+
+Lemma irr1_rfix : (1%irr :=: rfix_mx aG G)%MS.
+Proof.
+rewrite [1%irr]unlock PackSocleK; apply/eqmxP.
+rewrite (component_mx_id principal_comp_subproof) andbT.
+have [I [W isoW ->]] := component_mx_def principal_comp_subproof.
+apply/sumsmx_subP=> i _; have [f _ hom_f <-]:= isoW i.
+by apply/rfix_mxP=> x Gx; rewrite -(hom_mxP hom_f) // (rfix_mxP G _).
+Qed.
+
+Lemma rank_irr1 : \rank 1%irr = 1%N.
+Proof.
+apply/eqP; rewrite eqn_leq lt0n mxrank_eq0 nz_socle andbT.
+by rewrite irr1_rfix rfix_regular rank_leq_row.
+Qed.
+
+Lemma degree_irr1 : 'n_1 = 1%N.
+Proof.
+apply/eqP; rewrite eqn_leq irr_degree_gt0 -rank_irr1.
+by rewrite mxrankS ?component_mx_id //; exact: socle_simple.
+Qed.
+
+Definition Wedderburn_subring (i : sG) := <<i *m R_G>>%MS.
+
+Local Notation "''R_' i" := (Wedderburn_subring i) : group_ring_scope.
+
+Let sums_R : (\sum_i 'R_i :=: Socle sG *m R_G)%MS.
+Proof.
+apply/eqmxP; set R_S := (_ <= _)%MS.
+have sRS: R_S by apply/sumsmx_subP=> i; rewrite genmxE submxMr ?(sumsmx_sup i).
+rewrite sRS -(mulmxKpV sRS) mulmxA submxMr //; apply/sumsmx_subP=> i _.
+rewrite -(submxMfree _ _ gring_free) -(mulmxA _ _ R_G) mulmxKpV //.
+by rewrite (sumsmx_sup i) ?genmxE.
+Qed.
+
+Lemma Wedderburn_ideal i : mx_ideal R_G 'R_i.
+Proof.
+apply/andP; split; last first.
+ rewrite /right_mx_ideal genmxE (muls_eqmx (genmxE _) (eqmx_refl _)).
+ by rewrite -[(_ <= _)%MS]regular_module_ideal component_mx_module.
+apply/mulsmx_subP=> A B R_A; rewrite !genmxE !mem_sub_gring => /andP[R_B SiB].
+rewrite envelop_mxM {R_A}// gring_row_mul -{R_B}(gring_rowK R_B).
+pose f := mulmx (gring_row A) \o gring_mx aG.
+rewrite -[_ *m _](mul_rV_lin1 [linear of f]).
+suffices: (i *m lin1_mx f <= i)%MS by apply: submx_trans; rewrite submxMr.
+apply: hom_component_mx; first exact: socle_simple.
+apply/rV_subP=> v _; apply/hom_mxP=> x Gx.
+by rewrite !mul_rV_lin1 /f /= gring_mxJ ?mulmxA.
+Qed.
+
+Lemma Wedderburn_direct : mxdirect (\sum_i 'R_i)%MS.
+Proof.
+apply/mxdirectP; rewrite /= sums_R mxrankMfree ?gring_free //.
+rewrite (mxdirectP (Socle_direct sG)); apply: eq_bigr=> i _ /=.
+by rewrite genmxE mxrankMfree ?gring_free.
+Qed.
+
+Lemma Wedderburn_disjoint i j : i != j -> ('R_i :&: 'R_j)%MS = 0.
+Proof.
+move=> ne_ij; apply/eqP; rewrite -submx0 capmxC.
+by rewrite -(mxdirect_sumsP Wedderburn_direct j) // capmxS // (sumsmx_sup i).
+Qed.
+
+Lemma Wedderburn_annihilate i j : i != j -> ('R_i * 'R_j)%MS = 0.
+Proof.
+move=> ne_ij; apply/eqP; rewrite -submx0 -(Wedderburn_disjoint ne_ij).
+rewrite sub_capmx; apply/andP; split.
+ case/andP: (Wedderburn_ideal i) => _; apply: submx_trans.
+ by rewrite mulsmxS // genmxE submxMl.
+case/andP: (Wedderburn_ideal j) => idlRj _; apply: submx_trans idlRj.
+by rewrite mulsmxS // genmxE submxMl.
+Qed.
+
+Lemma Wedderburn_mulmx0 i j A B :
+ i != j -> (A \in 'R_i)%MS -> (B \in 'R_j)%MS -> A *m B = 0.
+Proof.
+move=> ne_ij RiA RjB; apply: memmx0.
+by rewrite -(Wedderburn_annihilate ne_ij) mem_mulsmx.
+Qed.
+
+Hypothesis F'G : [char F]^'.-group G.
+
+Lemma irr_mx_sum : (\sum_(i : sG) i = 1%:M)%MS.
+Proof. by apply: reducible_Socle1; exact: mx_Maschke. Qed.
+
+Lemma Wedderburn_sum : (\sum_i 'R_i :=: R_G)%MS.
+Proof. by apply: eqmx_trans sums_R _; rewrite /Socle irr_mx_sum mul1mx. Qed.
+
+Definition Wedderburn_id i :=
+ vec_mx (mxvec 1%:M *m proj_mx 'R_i (\sum_(j | j != i) 'R_j)%MS).
+
+Local Notation "''e_' i" := (Wedderburn_id i) : group_ring_scope.
+
+Lemma Wedderburn_sum_id : \sum_i 'e_i = 1%:M.
+Proof.
+rewrite -linear_sum; apply: canLR mxvecK _.
+have: (1%:M \in R_G)%MS := envelop_mx1 aG.
+rewrite -Wedderburn_sum; case/(sub_dsumsmx Wedderburn_direct) => e Re -> _.
+apply: eq_bigr => i _; have dxR := mxdirect_sumsP Wedderburn_direct i (erefl _).
+rewrite (bigD1 i) // mulmxDl proj_mx_id ?Re // proj_mx_0 ?addr0 //=.
+by rewrite summx_sub // => j ne_ji; rewrite (sumsmx_sup j) ?Re.
+Qed.
+
+Lemma Wedderburn_id_mem i : ('e_i \in 'R_i)%MS.
+Proof. by rewrite vec_mxK proj_mx_sub. Qed.
+
+Lemma Wedderburn_is_id i : mxring_id 'R_i 'e_i.
+Proof.
+have ideRi A: (A \in 'R_i)%MS -> 'e_i *m A = A.
+ move=> RiA; rewrite -{2}[A]mul1mx -Wedderburn_sum_id mulmx_suml.
+ rewrite (bigD1 i) //= big1 ?addr0 // => j ne_ji.
+ by rewrite (Wedderburn_mulmx0 ne_ji) ?Wedderburn_id_mem.
+ split=> // [||A RiA]; first 2 [exact: Wedderburn_id_mem].
+ apply: contraNneq (nz_socle i) => e0.
+ apply/rowV0P=> v; rewrite -mem_gring_mx -(genmxE (i *m _)) => /ideRi.
+ by rewrite e0 mul0mx => /(canLR gring_mxK); rewrite linear0.
+rewrite -{2}[A]mulmx1 -Wedderburn_sum_id mulmx_sumr (bigD1 i) //=.
+rewrite big1 ?addr0 // => j; rewrite eq_sym => ne_ij.
+by rewrite (Wedderburn_mulmx0 ne_ij) ?Wedderburn_id_mem.
+Qed.
+
+Lemma Wedderburn_closed i : ('R_i * 'R_i = 'R_i)%MS.
+Proof.
+rewrite -{3}['R_i]genmx_id -/'R_i -genmx_muls; apply/genmxP.
+have [idlRi idrRi] := andP (Wedderburn_ideal i).
+apply/andP; split.
+ by apply: submx_trans idrRi; rewrite mulsmxS // genmxE submxMl.
+have [_ Ri_e ideRi _] := Wedderburn_is_id i.
+by apply/memmx_subP=> A RiA; rewrite -[A]ideRi ?mem_mulsmx.
+Qed.
+
+Lemma Wedderburn_is_ring i : mxring 'R_i.
+Proof.
+rewrite /mxring /left_mx_ideal Wedderburn_closed submx_refl.
+by apply/mxring_idP; exists 'e_i; exact: Wedderburn_is_id.
+Qed.
+
+Lemma Wedderburn_min_ideal m i (E : 'A_(m, nG)) :
+ E != 0 -> (E <= 'R_i)%MS -> mx_ideal R_G E -> (E :=: 'R_i)%MS.
+Proof.
+move=> nzE sE_Ri /andP[idlE idrE]; apply/eqmxP; rewrite sE_Ri.
+pose M := E *m pinvmx R_G; have defE: E = M *m R_G.
+ by rewrite mulmxKpV // (submx_trans sE_Ri) // genmxE submxMl.
+have modM: mxmodule aG M by rewrite regular_module_ideal -defE.
+have simSi := socle_simple i; set Si := socle_base i in simSi.
+have [I [W isoW defW]]:= component_mx_def simSi.
+rewrite /'R_i /socle_val /= defW genmxE defE submxMr //.
+apply/sumsmx_subP=> j _.
+have simW := mx_iso_simple (isoW j) simSi; have [modW _ minW] := simW.
+have [{minW}dxWE | nzWE] := eqVneq (W j :&: M)%MS 0; last first.
+ by rewrite (sameP capmx_idPl eqmxP) minW ?capmxSl ?capmx_module.
+have [_ Rei ideRi _] := Wedderburn_is_id i.
+have:= nzE; rewrite -submx0 => /memmx_subP[A E_A].
+rewrite -(ideRi _ (memmx_subP sE_Ri _ E_A)).
+have:= E_A; rewrite defE mem_sub_gring => /andP[R_A M_A].
+have:= Rei; rewrite genmxE mem_sub_gring => /andP[Re].
+rewrite -{2}(gring_rowK Re) /socle_val defW => /sub_sumsmxP[e ->].
+rewrite !(linear_sum, mulmx_suml) summx_sub //= => k _.
+rewrite -(gring_rowK R_A) -gring_mxA -mulmxA gring_rowK //.
+rewrite ((W k *m _ =P 0) _) ?linear0 ?sub0mx //.
+have [f _ homWf defWk] := mx_iso_trans (mx_iso_sym (isoW j)) (isoW k).
+rewrite -submx0 -{k defWk}(eqmxMr _ defWk) -(hom_envelop_mxC homWf) //.
+rewrite -(mul0mx _ f) submxMr {f homWf}// -dxWE sub_capmx.
+rewrite (mxmodule_envelop modW) //=; apply/row_subP=> k.
+rewrite row_mul -mem_gring_mx -(gring_rowK R_A) gring_mxA gring_rowK //.
+by rewrite -defE (memmx_subP idlE) // mem_mulsmx ?gring_mxP.
+Qed.
+
+Section IrrComponent.
+
+(* The component of the socle of the regular module that is associated to an *)
+(* irreducible representation. *)
+
+Variables (n : nat) (rG : mx_representation F G n).
+Local Notation E_G := (enveloping_algebra_mx rG).
+
+Let not_rsim_op0 (iG j : sG) A :
+ mx_rsim rG (socle_repr iG) -> iG != j -> (A \in 'R_j)%MS ->
+ gring_op rG A = 0.
+Proof.
+case/mx_rsim_def=> f [f' _ hom_f] ne_iG_j RjA.
+transitivity (f *m in_submod _ (val_submod 1%:M *m A) *m f').
+ have{RjA}: (A \in R_G)%MS by rewrite -Wedderburn_sum (sumsmx_sup j).
+ case/envelop_mxP=> a ->{A}; rewrite !(linear_sum, mulmx_suml).
+ by apply: eq_bigr => x Gx; rewrite !linearZ /= -scalemxAl -hom_f ?gring_opG.
+rewrite (_ : _ *m A = 0) ?(linear0, mul0mx) //.
+apply/row_matrixP=> i; rewrite row_mul row0 -[row _ _]gring_mxK -gring_row_mul.
+rewrite (Wedderburn_mulmx0 ne_iG_j) ?linear0 // genmxE mem_gring_mx.
+by rewrite (row_subP _) // val_submod1 component_mx_id //; exact: socle_simple.
+Qed.
+
+Definition irr_comp := odflt 1%irr [pick i | gring_op rG 'e_i != 0].
+Local Notation iG := irr_comp.
+
+Hypothesis irrG : mx_irreducible rG.
+
+Lemma rsim_irr_comp : mx_rsim rG (irr_repr iG).
+Proof.
+have [M [modM rsimM]] := rsim_regular_submod irrG F'G.
+have simM: mxsimple aG M.
+ case/mx_irrP: irrG => n_gt0 minG.
+ have [f def_n injf homf] := mx_rsim_sym rsimM.
+ apply/(submod_mx_irr modM)/mx_irrP.
+ split=> [|U modU nzU]; first by rewrite def_n.
+ rewrite /row_full -(mxrankMfree _ injf) -genmxE {4}def_n.
+ apply: minG; last by rewrite -mxrank_eq0 genmxE mxrankMfree // mxrank_eq0.
+ rewrite (eqmx_module _ (genmxE _)); apply/mxmoduleP=> x Gx.
+ by rewrite -mulmxA -homf // mulmxA submxMr // (mxmoduleP modU).
+pose i := PackSocle (component_socle sG simM).
+have{modM rsimM} rsimM: mx_rsim rG (socle_repr i).
+ apply: mx_rsim_trans rsimM (mx_rsim_sym _); apply/mx_rsim_iso.
+ apply: (component_mx_iso (socle_simple _)) => //.
+ by rewrite [component_mx _ _]PackSocleK component_mx_id.
+have [<- // | ne_i_iG] := eqVneq i iG.
+suffices {i M simM ne_i_iG rsimM}: gring_op rG 'e_iG != 0.
+ by rewrite (not_rsim_op0 rsimM ne_i_iG) ?Wedderburn_id_mem ?eqxx.
+rewrite /iG; case: pickP => //= G0.
+suffices: rG 1%g == 0.
+ by case/idPn; rewrite -mxrank_eq0 repr_mx1 mxrank1 -lt0n; case/mx_irrP: irrG.
+rewrite -gring_opG // repr_mx1 -Wedderburn_sum_id linear_sum big1 // => j _.
+by move/eqP: (G0 j).
+Qed.
+
+Lemma irr_comp'_op0 j A : j != iG -> (A \in 'R_j)%MS -> gring_op rG A = 0.
+Proof. by rewrite eq_sym; exact: not_rsim_op0 rsim_irr_comp. Qed.
+
+Lemma irr_comp_envelop : ('R_iG *m lin_mx (gring_op rG) :=: E_G)%MS.
+Proof.
+apply/eqmxP/andP; split; apply/row_subP=> i.
+ by rewrite row_mul mul_rV_lin gring_mxP.
+rewrite rowK /= -gring_opG ?enum_valP // -mul_vec_lin -gring_opG ?enum_valP //.
+rewrite vec_mxK /= -mulmxA mulmx_sub {i}//= -(eqmxMr _ Wedderburn_sum).
+rewrite (bigD1 iG) //= addsmxMr addsmxC [_ *m _](sub_kermxP _) ?adds0mx //=.
+apply/sumsmx_subP => j ne_j_iG; apply/memmx_subP=> A RjA; apply/sub_kermxP.
+by rewrite mul_vec_lin /= (irr_comp'_op0 ne_j_iG RjA) linear0.
+Qed.
+
+Lemma ker_irr_comp_op : ('R_iG :&: kermx (lin_mx (gring_op rG)))%MS = 0.
+Proof.
+apply/eqP; rewrite -submx0; apply/memmx_subP=> A.
+rewrite sub_capmx /= submx0 mxvec_eq0 => /andP[R_A].
+rewrite (sameP sub_kermxP eqP) mul_vec_lin mxvec_eq0 /= => opA0.
+have [_ Re ideR _] := Wedderburn_is_id iG; rewrite -[A]ideR {ideR}//.
+move: Re; rewrite genmxE mem_sub_gring /socle_val => /andP[Re].
+rewrite -{2}(gring_rowK Re) -submx0.
+pose simMi := socle_simple iG; have [J [M isoM ->]] := component_mx_def simMi.
+case/sub_sumsmxP=> e ->; rewrite linear_sum mulmx_suml summx_sub // => j _.
+rewrite -(in_submodK (submxMl _ (M j))); move: (in_submod _ _) => v.
+have modMj: mxmodule aG (M j) by apply: mx_iso_module (isoM j) _; case: simMi.
+have rsimMj: mx_rsim rG (submod_repr modMj).
+ by apply: mx_rsim_trans rsim_irr_comp _; exact/mx_rsim_iso.
+have [f [f' _ hom_f]] := mx_rsim_def (mx_rsim_sym rsimMj); rewrite submx0.
+have <-: (gring_mx aG (val_submod (v *m (f *m gring_op rG A *m f')))) = 0.
+ by rewrite (eqP opA0) !(mul0mx, linear0).
+have: (A \in R_G)%MS by rewrite -Wedderburn_sum (sumsmx_sup iG).
+case/envelop_mxP=> a ->; rewrite !(linear_sum, mulmx_suml) /=; apply/eqP.
+apply: eq_bigr=> x Gx; rewrite !linearZ -scalemxAl !linearZ /=.
+by rewrite gring_opG // -hom_f // val_submodJ // gring_mxJ.
+Qed.
+
+Lemma regular_op_inj :
+ {in [pred A | (A \in 'R_iG)%MS] &, injective (gring_op rG)}.
+Proof.
+move=> A B RnA RnB /= eqAB; apply/eqP; rewrite -subr_eq0 -mxvec_eq0 -submx0.
+rewrite -ker_irr_comp_op sub_capmx (sameP sub_kermxP eqP) mul_vec_lin.
+by rewrite 2!linearB /= eqAB subrr linear0 addmx_sub ?eqmx_opp /=.
+Qed.
+
+Lemma rank_irr_comp : \rank 'R_iG = \rank E_G.
+Proof.
+symmetry; rewrite -{1}irr_comp_envelop; apply/mxrank_injP.
+by rewrite ker_irr_comp_op.
+Qed.
+
+End IrrComponent.
+
+Lemma irr_comp_rsim n1 n2 rG1 rG2 :
+ @mx_rsim _ G n1 rG1 n2 rG2 -> irr_comp rG1 = irr_comp rG2.
+Proof.
+case=> f eq_n12; rewrite -eq_n12 in rG2 f * => inj_f hom_f.
+congr (odflt _ _); apply: eq_pick => i; rewrite -!mxrank_eq0.
+rewrite -(mxrankMfree _ inj_f); symmetry; rewrite -(eqmxMfull _ inj_f).
+have /envelop_mxP[e ->{i}]: ('e_i \in R_G)%MS.
+ by rewrite -Wedderburn_sum (sumsmx_sup i) ?Wedderburn_id_mem.
+congr (\rank _ != _); rewrite !(mulmx_suml, linear_sum); apply: eq_bigr => x Gx.
+by rewrite !linearZ -scalemxAl /= !gring_opG ?hom_f.
+Qed.
+
+Lemma irr_reprK i : irr_comp (irr_repr i) = i.
+Proof.
+apply/eqP; apply/component_mx_isoP; try exact: socle_simple.
+by move/mx_rsim_iso: (rsim_irr_comp (socle_irr i)); exact: mx_iso_sym.
+Qed.
+
+Lemma irr_repr'_op0 i j A :
+ j != i -> (A \in 'R_j)%MS -> gring_op (irr_repr i) A = 0.
+Proof.
+by move=> neq_ij /irr_comp'_op0-> //; [exact: socle_irr | rewrite irr_reprK].
+Qed.
+
+Lemma op_Wedderburn_id i : gring_op (irr_repr i) 'e_i = 1%:M.
+Proof.
+rewrite -(gring_op1 (irr_repr i)) -Wedderburn_sum_id.
+rewrite linear_sum (bigD1 i) //= addrC big1 ?add0r // => j neq_ji.
+exact: irr_repr'_op0 (Wedderburn_id_mem j).
+Qed.
+
+Lemma irr_comp_id (M : 'M_nG) (modM : mxmodule aG M) (iM : sG) :
+ mxsimple aG M -> (M <= iM)%MS -> irr_comp (submod_repr modM) = iM.
+Proof.
+move=> simM sMiM; rewrite -[iM]irr_reprK.
+apply/esym/irr_comp_rsim/mx_rsim_iso/component_mx_iso => //.
+exact: socle_simple.
+Qed.
+
+Lemma irr1_repr x : x \in G -> irr_repr 1 x = 1%:M.
+Proof.
+move=> Gx; suffices: x \in rker (irr_repr 1) by case/rkerP.
+apply: subsetP x Gx; rewrite rker_submod rfix_mx_rstabC // -irr1_rfix.
+by apply: component_mx_id; exact: socle_simple.
+Qed.
+
+Hypothesis splitG : group_splitting_field G.
+
+Lemma rank_Wedderburn_subring i : \rank 'R_i = ('n_i ^ 2)%N.
+Proof.
+apply/eqP; rewrite -{1}[i]irr_reprK; have irrSi := socle_irr i.
+by case/andP: (splitG irrSi) => _; rewrite rank_irr_comp.
+Qed.
+
+Lemma sum_irr_degree : (\sum_i 'n_i ^ 2 = nG)%N.
+Proof.
+apply: etrans (eqnP gring_free).
+rewrite -Wedderburn_sum (mxdirectP Wedderburn_direct) /=.
+by apply: eq_bigr => i _; rewrite rank_Wedderburn_subring.
+Qed.
+
+Lemma irr_mx_mult i : socle_mult i = 'n_i.
+Proof.
+rewrite /socle_mult -(mxrankMfree _ gring_free) -genmxE.
+by rewrite rank_Wedderburn_subring mulKn ?irr_degree_gt0.
+Qed.
+
+Lemma mxtrace_regular :
+ {in G, forall x, \tr (aG x) = \sum_i \tr (socle_repr i x) *+ 'n_i}.
+Proof.
+move=> x Gx; have soc1: (Socle sG :=: 1%:M)%MS by rewrite -irr_mx_sum.
+rewrite -(mxtrace_submod1 (Socle_module sG) soc1) // mxtrace_Socle //.
+by apply: eq_bigr => i _; rewrite irr_mx_mult.
+Qed.
+
+Definition linear_irr := [set i | 'n_i == 1%N].
+
+Lemma irr_degree_abelian : abelian G -> forall i, 'n_i = 1%N.
+Proof. by move=> cGG i; exact: mxsimple_abelian_linear (socle_simple i). Qed.
+
+Lemma linear_irr_comp i : 'n_i = 1%N -> (i :=: socle_base i)%MS.
+Proof.
+move=> ni1; apply/eqmxP; rewrite andbC -mxrank_leqif_eq -/'n_i.
+ by rewrite -(mxrankMfree _ gring_free) -genmxE rank_Wedderburn_subring ni1.
+exact: component_mx_id (socle_simple i).
+Qed.
+
+Lemma Wedderburn_subring_center i : ('Z('R_i) :=: mxvec 'e_i)%MS.
+Proof.
+have [nz_e Re ideR idRe] := Wedderburn_is_id i.
+have Ze: (mxvec 'e_i <= 'Z('R_i))%MS.
+ rewrite sub_capmx [(_ <= _)%MS]Re.
+ by apply/cent_mxP=> A R_A; rewrite ideR // idRe.
+pose irrG := socle_irr i; set rG := socle_repr i in irrG.
+pose E_G := enveloping_algebra_mx rG; have absG := splitG irrG.
+apply/eqmxP; rewrite andbC -(geq_leqif (mxrank_leqif_eq Ze)).
+have ->: \rank (mxvec 'e_i) = (0 + 1)%N.
+ by apply/eqP; rewrite eqn_leq rank_leq_row lt0n mxrank_eq0 mxvec_eq0.
+rewrite -(mxrank_mul_ker _ (lin_mx (gring_op rG))) addnC leq_add //.
+ rewrite leqn0 mxrank_eq0 -submx0 -(ker_irr_comp_op irrG) capmxS //.
+ by rewrite irr_reprK capmxSl.
+apply: leq_trans (mxrankS _) (rank_leq_row (mxvec 1%:M)).
+apply/memmx_subP=> Ar; case/submxP=> a ->{Ar}.
+rewrite mulmxA mul_rV_lin /=; set A := vec_mx _.
+rewrite memmx1 (mx_abs_irr_cent_scalar absG) // -memmx_cent_envelop.
+apply/cent_mxP=> Br; rewrite -(irr_comp_envelop irrG) irr_reprK.
+case/submxP=> b /(canRL mxvecK) ->{Br}; rewrite mulmxA mx_rV_lin /=.
+set B := vec_mx _; have RiB: (B \in 'R_i)%MS by rewrite vec_mxK submxMl.
+have sRiR: ('R_i <= R_G)%MS by rewrite -Wedderburn_sum (sumsmx_sup i).
+have: (A \in 'Z('R_i))%MS by rewrite vec_mxK submxMl.
+rewrite sub_capmx => /andP[RiA /cent_mxP cRiA].
+by rewrite -!gring_opM ?(memmx_subP sRiR) 1?cRiA.
+Qed.
+
+Lemma Wedderburn_center :
+ ('Z(R_G) :=: \matrix_(i < #|sG|) mxvec 'e_(enum_val i))%MS.
+Proof.
+have:= mxdirect_sums_center Wedderburn_sum Wedderburn_direct Wedderburn_ideal.
+move/eqmx_trans; apply; apply/eqmxP/andP; split.
+ apply/sumsmx_subP=> i _; rewrite Wedderburn_subring_center.
+ by apply: (eq_row_sub (enum_rank i)); rewrite rowK enum_rankK.
+apply/row_subP=> i; rewrite rowK -Wedderburn_subring_center.
+by rewrite (sumsmx_sup (enum_val i)).
+Qed.
+
+Lemma card_irr : #|sG| = tG.
+Proof.
+rewrite -(eqnP classg_base_free) classg_base_center.
+have:= mxdirect_sums_center Wedderburn_sum Wedderburn_direct Wedderburn_ideal.
+move->; rewrite (mxdirectP _) /=; last first.
+ apply/mxdirect_sumsP=> i _; apply/eqP; rewrite -submx0.
+ rewrite -{2}(mxdirect_sumsP Wedderburn_direct i) // capmxS ?capmxSl //=.
+ by apply/sumsmx_subP=> j neji; rewrite (sumsmx_sup j) ?capmxSl.
+rewrite -sum1_card; apply: eq_bigr => i _; apply/eqP.
+rewrite Wedderburn_subring_center eqn_leq rank_leq_row lt0n mxrank_eq0.
+by rewrite andbT mxvec_eq0; case: (Wedderburn_is_id i).
+Qed.
+
+Section CenterMode.
+
+Variable i : sG.
+
+Let i0 := Ordinal (irr_degree_gt0 i).
+
+Definition irr_mode x := irr_repr i x i0 i0.
+
+Lemma irr_mode1 : irr_mode 1 = 1.
+Proof. by rewrite /irr_mode repr_mx1 mxE eqxx. Qed.
+
+Lemma irr_center_scalar : {in 'Z(G), forall x, irr_repr i x = (irr_mode x)%:M}.
+Proof.
+rewrite /irr_mode => x /setIP[Gx cGx].
+suffices [a ->]: exists a, irr_repr i x = a%:M by rewrite mxE eqxx.
+apply/is_scalar_mxP; apply: (mx_abs_irr_cent_scalar (splitG (socle_irr i))).
+by apply/centgmxP=> y Gy; rewrite -!{1}repr_mxM 1?(centP cGx).
+Qed.
+
+Lemma irr_modeM : {in 'Z(G) &, {morph irr_mode : x y / (x * y)%g >-> x * y}}.
+Proof.
+move=> x y Zx Zy; rewrite {1}/irr_mode repr_mxM ?(subsetP (center_sub G)) //.
+by rewrite !irr_center_scalar // -scalar_mxM mxE eqxx.
+Qed.
+
+Lemma irr_modeX n : {in 'Z(G), {morph irr_mode : x / (x ^+ n)%g >-> x ^+ n}}.
+Proof.
+elim: n => [|n IHn] x Zx; first exact: irr_mode1.
+by rewrite expgS irr_modeM ?groupX // exprS IHn.
+Qed.
+
+Lemma irr_mode_unit : {in 'Z(G), forall x, irr_mode x \is a GRing.unit}.
+Proof.
+move=> x Zx /=; have:= unitr1 F.
+by rewrite -irr_mode1 -(mulVg x) irr_modeM ?groupV // unitrM; case/andP=> _.
+Qed.
+
+Lemma irr_mode_neq0 : {in 'Z(G), forall x, irr_mode x != 0}.
+Proof. by move=> x /irr_mode_unit; rewrite unitfE. Qed.
+
+Lemma irr_modeV : {in 'Z(G), {morph irr_mode : x / (x^-1)%g >-> x^-1}}.
+Proof.
+move=> x Zx /=; rewrite -[_^-1]mul1r; apply: canRL (mulrK (irr_mode_unit Zx)) _.
+by rewrite -irr_modeM ?groupV // mulVg irr_mode1.
+Qed.
+
+End CenterMode.
+
+Lemma irr1_mode x : x \in G -> irr_mode 1 x = 1.
+Proof. by move=> Gx; rewrite /irr_mode irr1_repr ?mxE. Qed.
+
+End Regular.
+
+Local Notation "[ 1 sG ]" := (principal_comp sG) : irrType_scope.
+
+Section LinearIrr.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Lemma card_linear_irr (sG : irrType G) :
+ [char F]^'.-group G -> group_splitting_field G ->
+ #|linear_irr sG| = #|G : G^`(1)|%g.
+Proof.
+move=> F'G splitG; apply/eqP.
+wlog sGq: / irrType (G / G^`(1))%G by exact: socle_exists.
+have [_ nG'G] := andP (der_normal 1 G); apply/eqP; rewrite -card_quotient //.
+have cGqGq: abelian (G / G^`(1))%g by exact: sub_der1_abelian.
+have F'Gq: [char F]^'.-group (G / G^`(1))%g by exact: morphim_pgroup.
+have splitGq: group_splitting_field (G / G^`(1))%G.
+ exact: quotient_splitting_field.
+rewrite -(sum_irr_degree sGq) // -sum1_card.
+pose rG (j : sGq) := morphim_repr (socle_repr j) nG'G.
+have irrG j: mx_irreducible (rG j) by apply/morphim_mx_irr; exact: socle_irr.
+rewrite (reindex (fun j => irr_comp sG (rG j))) /=.
+ apply: eq_big => [j | j _]; last by rewrite irr_degree_abelian.
+ have [_ lin_j _ _] := rsim_irr_comp sG F'G (irrG j).
+ by rewrite inE -lin_j -irr_degreeE irr_degree_abelian.
+pose sGlin := {i | i \in linear_irr sG}.
+have sG'k (i : sGlin) : G^`(1)%g \subset rker (irr_repr (val i)).
+ by case: i => i /=; rewrite !inE => lin; rewrite rker_linear //=; exact/eqP.
+pose h' u := irr_comp sGq (quo_repr (sG'k u) nG'G).
+have irrGq u: mx_irreducible (quo_repr (sG'k u) nG'G).
+ by apply/quo_mx_irr; exact: socle_irr.
+exists (fun i => oapp h' [1 sGq]%irr (insub i)) => [j | i] lin_i.
+ rewrite (insubT (mem _) lin_i) /=; apply/esym/eqP/socle_rsimP.
+ apply: mx_rsim_trans (rsim_irr_comp sGq F'Gq (irrGq _)).
+ have [g lin_g inj_g hom_g] := rsim_irr_comp sG F'G (irrG j).
+ exists g => [||G'x]; last 1 [case/morphimP=> x _ Gx ->] || by [].
+ by rewrite quo_repr_coset ?hom_g.
+rewrite (insubT (mem _) lin_i) /=; apply/esym/eqP/socle_rsimP.
+set u := exist _ _ _; apply: mx_rsim_trans (rsim_irr_comp sG F'G (irrG _)).
+have [g lin_g inj_g hom_g] := rsim_irr_comp sGq F'Gq (irrGq u).
+exists g => [||x Gx]; last 1 [have:= hom_g (coset _ x)] || by [].
+by rewrite quo_repr_coset; first by apply; rewrite mem_quotient.
+Qed.
+
+Lemma primitive_root_splitting_abelian (z : F) :
+ #|G|.-primitive_root z -> abelian G -> group_splitting_field G.
+Proof.
+move=> ozG cGG [|n] rG irrG; first by case/mx_irrP: irrG.
+case: (pickP [pred x in G | ~~ is_scalar_mx (rG x)]) => [x | scalG].
+ case/andP=> Gx nscal_rGx; have: horner_mx (rG x) ('X^#|G| - 1) == 0.
+ rewrite rmorphB rmorphX /= horner_mx_C horner_mx_X.
+ rewrite -repr_mxX ?inE // ((_ ^+ _ =P 1)%g _) ?repr_mx1 ?subrr //.
+ by rewrite -order_dvdn order_dvdG.
+ case/idPn; rewrite -mxrank_eq0 -(factor_Xn_sub_1 ozG).
+ elim: #|G| => [|i IHi]; first by rewrite big_nil horner_mx_C mxrank1.
+ rewrite big_nat_recr //= rmorphM mxrankMfree {IHi}//.
+ rewrite row_free_unit rmorphB /= horner_mx_X horner_mx_C.
+ rewrite (mx_Schur irrG) ?subr_eq0 //; last first.
+ by apply: contraNneq nscal_rGx => ->; exact: scalar_mx_is_scalar.
+ rewrite -memmx_cent_envelop linearB.
+ rewrite addmx_sub ?eqmx_opp ?scalar_mx_cent //= memmx_cent_envelop.
+ by apply/centgmxP=> j Zh_j; rewrite -!repr_mxM // (centsP cGG).
+pose M := <<delta_mx 0 0 : 'rV[F]_n.+1>>%MS.
+have linM: \rank M = 1%N by rewrite genmxE mxrank_delta.
+have modM: mxmodule rG M.
+ apply/mxmoduleP=> x Gx; move/idPn: (scalG x); rewrite /= Gx negbK.
+ by case/is_scalar_mxP=> ? ->; rewrite scalar_mxC submxMl.
+apply: linear_mx_abs_irr; apply/eqP; rewrite eq_sym -linM.
+by case/mx_irrP: irrG => _; apply; rewrite // -mxrank_eq0 linM.
+Qed.
+
+Lemma cycle_repr_structure x (sG : irrType G) :
+ G :=: <[x]> -> [char F]^'.-group G -> group_splitting_field G ->
+ exists2 w : F, #|G|.-primitive_root w &
+ exists iphi : 'I_#|G| -> sG,
+ [/\ bijective iphi,
+ #|sG| = #|G|,
+ forall i, irr_mode (iphi i) x = w ^+ i
+ & forall i, irr_repr (iphi i) x = (w ^+ i)%:M].
+Proof.
+move=> defG; rewrite {defG}(group_inj defG) -/#[x] in sG * => F'X splitF.
+have Xx := cycle_id x; have cXX := cycle_abelian x.
+have card_sG: #|sG| = #[x].
+ by rewrite card_irr //; apply/eqP; rewrite -card_classes_abelian.
+have linX := irr_degree_abelian splitF cXX (_ : sG).
+pose r (W : sG) := irr_mode W x.
+have scalX W: irr_repr W x = (r W)%:M.
+ by apply: irr_center_scalar; rewrite ?(center_idP _).
+have inj_r: injective r.
+ move=> V W eqVW; rewrite -(irr_reprK F'X V) -(irr_reprK F'X W).
+ move: (irr_repr V) (irr_repr W) (scalX V) (scalX W).
+ rewrite !linX {}eqVW => rV rW <- rWx; apply: irr_comp_rsim => //.
+ exists 1%:M; rewrite ?row_free_unit ?unitmx1 // => xk; case/cycleP=> k ->{xk}.
+ by rewrite mulmx1 mul1mx !repr_mxX // rWx.
+have rx1 W: r W ^+ #[x] = 1.
+ by rewrite -irr_modeX ?(center_idP _) // expg_order irr_mode1.
+have /hasP[w _ prim_w]: has #[x].-primitive_root (map r (enum sG)).
+ rewrite has_prim_root 1?map_inj_uniq ?enum_uniq //; first 1 last.
+ by rewrite size_map -cardE card_sG.
+ by apply/allP=> _ /mapP[W _ ->]; rewrite unity_rootE rx1.
+have iphi'P := prim_rootP prim_w (rx1 _); pose iphi' := sval (iphi'P _).
+have def_r W: r W = w ^+ iphi' W by exact: svalP (iphi'P W).
+have inj_iphi': injective iphi'.
+ by move=> i j eq_ij; apply: inj_r; rewrite !def_r eq_ij.
+have iphiP: codom iphi' =i 'I_#[x].
+ by apply/subset_cardP; rewrite ?subset_predT // card_ord card_image.
+pose iphi i := iinv (iphiP i); exists w => //; exists iphi.
+have iphiK: cancel iphi iphi' by move=> i; exact: f_iinv.
+have r_iphi i: r (iphi i) = w ^+ i by rewrite def_r iphiK.
+split=> // [|i]; last by rewrite scalX r_iphi.
+by exists iphi' => // W; rewrite /iphi iinv_f.
+Qed.
+
+Lemma splitting_cyclic_primitive_root :
+ cyclic G -> [char F]^'.-group G -> group_splitting_field G ->
+ classically {z : F | #|G|.-primitive_root z}.
+Proof.
+case/cyclicP=> x defG F'G splitF; case=> // IH.
+wlog sG: / irrType G by exact: socle_exists.
+have [w prim_w _] := cycle_repr_structure sG defG F'G splitF.
+by apply: IH; exists w.
+Qed.
+
+End LinearIrr.
+
+End FieldRepr.
+
+Arguments Scope rfix_mx [_ _ group_scope nat_scope _ group_scope].
+Arguments Scope gset_mx [_ _ group_scope group_scope].
+Arguments Scope classg_base [_ _ group_scope group_scope].
+Arguments Scope irrType [_ _ group_scope group_scope].
+
+Implicit Arguments mxmoduleP [F gT G n rG m U].
+Implicit Arguments envelop_mxP [F gT G n rG A].
+Implicit Arguments hom_mxP [F gT G n rG m f W].
+Implicit Arguments mx_Maschke [F gT G n U].
+Implicit Arguments rfix_mxP [F gT G n rG m W].
+Implicit Arguments cyclic_mxP [F gT G n rG u v].
+Implicit Arguments annihilator_mxP [F gT G n rG u A].
+Implicit Arguments row_hom_mxP [F gT G n rG u v].
+Implicit Arguments mxsimple_isoP [F gT G n rG U V].
+Implicit Arguments socle_exists [F gT G n].
+Implicit Arguments socleP [F gT G n rG sG0 W W'].
+Implicit Arguments mx_abs_irrP [F gT G n rG].
+Implicit Arguments socle_rsimP [F gT G n rG sG W1 W2].
+
+Implicit Arguments val_submod_inj [F n U m].
+Implicit Arguments val_factmod_inj [F n U m].
+Prenex Implicits val_submod_inj val_factmod_inj.
+
+Notation "'Cl" := (Clifford_action _) : action_scope.
+
+Bind Scope irrType_scope with socle_sort.
+Notation "[ 1 sG ]" := (principal_comp sG) : irrType_scope.
+Arguments Scope irr_degree [_ _ Group_scope _ irrType_scope].
+Arguments Scope irr_repr [_ _ Group_scope _ irrType_scope group_scope].
+Arguments Scope irr_mode [_ _ Group_scope _ irrType_scope group_scope].
+Notation "''n_' i" := (irr_degree i) : group_ring_scope.
+Notation "''R_' i" := (Wedderburn_subring i) : group_ring_scope.
+Notation "''e_' i" := (Wedderburn_id i) : group_ring_scope.
+
+Section DecideRed.
+
+Import MatrixFormula.
+Local Notation term := GRing.term.
+Local Notation True := GRing.True.
+Local Notation And := GRing.And (only parsing).
+Local Notation morphAnd f := ((big_morph f) true andb).
+Local Notation eval := GRing.eval.
+Local Notation holds := GRing.holds.
+Local Notation qf_form := GRing.qf_form.
+Local Notation qf_eval := GRing.qf_eval.
+
+Section Definitions.
+
+Variables (F : fieldType) (gT : finGroupType) (G : {group gT}) (n : nat).
+Variable rG : mx_representation F G n.
+
+Definition mxmodule_form (U : 'M[term F]_n) :=
+ \big[And/True]_(x in G) submx_form (mulmx_term U (mx_term (rG x))) U.
+
+Lemma mxmodule_form_qf U : qf_form (mxmodule_form U).
+Proof.
+by rewrite (morphAnd (@qf_form _)) ?big1 //= => x _; rewrite submx_form_qf.
+Qed.
+
+Lemma eval_mxmodule U e :
+ qf_eval e (mxmodule_form U) = mxmodule rG (eval_mx e U).
+Proof.
+rewrite (morphAnd (qf_eval e)) //= big_andE /=.
+apply/forallP/mxmoduleP=> Umod x; move/implyP: (Umod x);
+ by rewrite eval_submx eval_mulmx eval_mx_term.
+Qed.
+
+Definition mxnonsimple_form (U : 'M[term F]_n) :=
+ let V := vec_mx (row_var F (n * n) 0) in
+ let nzV := (~ mxrank_form 0 V)%T in
+ let properVU := (submx_form V U /\ ~ submx_form U V)%T in
+ (Exists_row_form (n * n) 0 (mxmodule_form V /\ nzV /\ properVU))%T.
+
+End Definitions.
+
+Variables (F : decFieldType) (gT : finGroupType) (G : {group gT}) (n : nat).
+Variable rG : mx_representation F G n.
+
+Definition mxnonsimple_sat U :=
+ GRing.sat (@row_env _ (n * n) [::]) (mxnonsimple_form rG (mx_term U)).
+
+Lemma mxnonsimpleP U :
+ U != 0 -> reflect (mxnonsimple rG U) (mxnonsimple_sat U).
+Proof.
+rewrite /mxnonsimple_sat {1}/mxnonsimple_form; set Vt := vec_mx _ => /= nzU.
+pose nsim V := [&& mxmodule rG V, (V <= U)%MS, V != 0 & \rank V < \rank U].
+set nsimUt := (_ /\ _)%T; have: qf_form nsimUt.
+ by rewrite /= mxmodule_form_qf !mxrank_form_qf !submx_form_qf.
+move/GRing.qf_evalP; set qev := @GRing.qf_eval _ => qevP.
+have qev_nsim u: qev (row_env [:: u]) nsimUt = nsim n (vec_mx u).
+ rewrite /nsim -mxrank_eq0 /qev /= eval_mxmodule eval_mxrank.
+ rewrite !eval_submx eval_mx_term eval_vec_mx eval_row_var /=.
+ do 2!bool_congr; apply: andb_id2l => sUV.
+ by rewrite ltn_neqAle andbC !mxrank_leqif_sup.
+have n2gt0: n ^ 2 > 0.
+ by move: nzU; rewrite muln_gt0 -mxrank_eq0; case: posnP (U) => // ->.
+apply: (iffP satP) => [|[V nsimV]].
+ by case/Exists_rowP=> // v; move/qevP; rewrite qev_nsim; exists (vec_mx v).
+apply/Exists_rowP=> //; exists (mxvec V); apply/qevP.
+by rewrite qev_nsim mxvecK.
+Qed.
+
+Lemma dec_mxsimple_exists (U : 'M_n) :
+ mxmodule rG U -> U != 0 -> {V | mxsimple rG V & V <= U}%MS.
+Proof.
+elim: {U}_.+1 {-2}U (ltnSn (\rank U)) => // m IHm U leUm modU nzU.
+have [nsimU | simU] := mxnonsimpleP nzU; last first.
+ by exists U; first exact/mxsimpleP.
+move: (xchooseP nsimU); move: (xchoose _) => W /and4P[modW sWU nzW ltWU].
+case: (IHm W) => // [|V simV sVW]; first exact: leq_trans ltWU _.
+by exists V; last exact: submx_trans sVW sWU.
+Qed.
+
+Lemma dec_mx_reducible_semisimple U :
+ mxmodule rG U -> mx_completely_reducible rG U -> mxsemisimple rG U.
+Proof.
+elim: {U}_.+1 {-2}U (ltnSn (\rank U)) => // m IHm U leUm modU redU.
+have [U0 | nzU] := eqVneq U 0.
+ have{U0} U0: (\sum_(i < 0) 0 :=: U)%MS by rewrite big_ord0 U0.
+ by apply: (intro_mxsemisimple U0); case.
+have [V simV sVU] := dec_mxsimple_exists modU nzU; have [modV nzV _] := simV.
+have [W modW defVW dxVW] := redU V modV sVU.
+have [||I W_ /= simW defW _] := IHm W _ modW.
+- rewrite ltnS in leUm; apply: leq_trans leUm.
+ by rewrite -defVW (mxdirectP dxVW) /= -add1n leq_add2r lt0n mxrank_eq0.
+- by apply: mx_reducibleS redU; rewrite // -defVW addsmxSr.
+suffices defU: (\sum_i oapp W_ V i :=: U)%MS.
+ by apply: (intro_mxsemisimple defU) => [] [|i] //=.
+apply: eqmx_trans defVW; rewrite (bigD1 None) //=; apply/eqmxP.
+have [i0 _ | I0] := pickP I.
+ by rewrite (reindex some) ?addsmxS ?defW //; exists (odflt i0) => //; case.
+rewrite big_pred0 //; last by case => // /I0.
+by rewrite !addsmxS ?sub0mx // -defW big_pred0.
+Qed.
+
+Lemma DecSocleType : socleType rG.
+Proof.
+have [n0 | n_gt0] := posnP n.
+ by exists [::] => // M [_]; rewrite -mxrank_eq0 -leqn0 -n0 rank_leq_row.
+have n2_gt0: n ^ 2 > 0 by rewrite muln_gt0 n_gt0.
+pose span Ms := (\sum_(M <- Ms) component_mx rG M)%MS.
+have: {in [::], forall M, mxsimple rG M} by [].
+elim: _.+1 {-2}nil (ltnSn (n - \rank (span nil))) => // m IHm Ms Ms_ge_n simMs.
+rewrite ltnS in Ms_ge_n; pose V := span Ms; pose Vt := mx_term V.
+pose Ut i := vec_mx (row_var F (n * n) i); pose Zt := mx_term (0 : 'M[F]_n).
+pose exU i f := Exists_row_form (n * n) i (~ submx_form (Ut i) Zt /\ f (Ut i)).
+pose meetUVf U := exU 1%N (fun W => submx_form W Vt /\ submx_form W U)%T.
+pose mx_sat := GRing.sat (@row_env F (n * n) [::]).
+have ev_sub0 := GRing.qf_evalP _ (submx_form_qf _ Zt).
+have ev_mod := GRing.qf_evalP _ (mxmodule_form_qf rG _).
+pose ev := (eval_mxmodule, eval_submx, eval_vec_mx, eval_row_var, eval_mx_term).
+case haveU: (mx_sat (exU 0%N (fun U => mxmodule_form rG U /\ ~ meetUVf _ U)%T)).
+ have [U modU]: {U : 'M_n | mxmodule rG U & (U != 0) && ((U :&: V)%MS == 0)}.
+ apply: sig2W; case/Exists_rowP: (satP haveU) => //= u [nzU [modU tiUV]].
+ exists (vec_mx u); first by move/ev_mod: modU; rewrite !ev.
+ set W := (_ :&: V)%MS; move/ev_sub0: nzU; rewrite !ev -!submx0 => -> /=.
+ apply/idPn=> nzW; case: tiUV; apply/Exists_rowP=> //; exists (mxvec W).
+ apply/GRing.qf_evalP; rewrite /= ?submx_form_qf // !ev mxvecK nzW /=.
+ by rewrite andbC -sub_capmx.
+ case/andP=> nzU tiUV; have [M simM sMU] := dec_mxsimple_exists modU nzU.
+ apply: (IHm (M :: Ms)) => [|M']; last first.
+ by case/predU1P=> [-> //|]; exact: simMs.
+ have [_ nzM _] := simM.
+ suffices ltVMV: \rank V < \rank (span (M :: Ms)).
+ rewrite (leq_trans _ Ms_ge_n) // ltn_sub2l ?(leq_trans ltVMV) //.
+ exact: rank_leq_row.
+ rewrite /span big_cons (ltn_leqif (mxrank_leqif_sup (addsmxSr _ _))).
+ apply: contra nzM; rewrite addsmx_sub -submx0 -(eqP tiUV) sub_capmx sMU.
+ by case/andP=> sMV _; rewrite (submx_trans _ sMV) ?component_mx_id.
+exists Ms => // M simM; have [modM nzM minM] := simM.
+have sMV: (M <= V)%MS.
+ apply: contraFT haveU => not_sMV; apply/satP/Exists_rowP=> //.
+ exists (mxvec M); split; first by apply/ev_sub0; rewrite !ev mxvecK submx0.
+ split; first by apply/ev_mod; rewrite !ev mxvecK.
+ apply/Exists_rowP=> // [[w]].
+ apply/GRing.qf_evalP; rewrite /= ?submx_form_qf // !ev /= mxvecK submx0.
+ rewrite -nz_row_eq0 -(cyclic_mx_eq0 rG); set W := cyclic_mx _ _.
+ apply: contra not_sMV => /and3P[nzW Vw Mw].
+ have{Vw Mw} [sWV sWM]: (W <= V /\ W <= M)%MS.
+ rewrite !cyclic_mx_sub ?(submx_trans (nz_row_sub _)) //.
+ by rewrite sumsmx_module // => M' _; exact: component_mx_module.
+ by rewrite (submx_trans _ sWV) // minM ?cyclic_mx_module.
+wlog sG: / socleType rG by exact: socle_exists.
+have sVS: (V <= \sum_(W : sG | has (fun Mi => Mi <= W) Ms) W)%MS.
+ rewrite [V](big_nth 0) big_mkord; apply/sumsmx_subP=> i _.
+ set Mi := Ms`_i; have MsMi: Mi \in Ms by exact: mem_nth.
+ have simMi := simMs _ MsMi; have S_Mi := component_socle sG simMi.
+ rewrite (sumsmx_sup (PackSocle S_Mi)) ?PackSocleK //.
+ by apply/hasP; exists Mi; rewrite ?component_mx_id.
+have [W MsW isoWM] := subSocle_iso simM (submx_trans sMV sVS).
+have [Mi MsMi sMiW] := hasP MsW; apply/hasP; exists Mi => //.
+have [simMi simW] := (simMs _ MsMi, socle_simple W); apply/mxsimple_isoP=> //.
+exact: mx_iso_trans (mx_iso_sym isoWM) (component_mx_iso simW simMi sMiW).
+Qed.
+
+End DecideRed.
+
+(* Change of representation field (by tensoring) *)
+Section ChangeOfField.
+
+Variables (aF rF : fieldType) (f : {rmorphism aF -> rF}).
+Local Notation "A ^f" := (map_mx (GRing.RMorphism.apply f) A) : ring_scope.
+Variables (gT : finGroupType) (G : {group gT}).
+
+Section OneRepresentation.
+
+Variables (n : nat) (rG : mx_representation aF G n).
+Local Notation rGf := (map_repr f rG).
+
+Lemma map_rfix_mx H : (rfix_mx rG H)^f = rfix_mx rGf H.
+Proof.
+rewrite map_kermx //; congr (kermx _); apply: map_lin1_mx => //= v.
+rewrite map_mxvec map_mxM; congr (mxvec (_ *m _)); last first.
+ by apply: map_lin1_mx => //= u; rewrite map_mxM map_vec_mx.
+apply/row_matrixP=> i.
+by rewrite -map_row !rowK map_mxvec map_mx_sub map_mx1.
+Qed.
+
+Lemma rcent_map A : rcent rGf A^f = rcent rG A.
+Proof.
+by apply/setP=> x; rewrite !inE -!map_mxM inj_eq //; exact: map_mx_inj.
+Qed.
+
+Lemma rstab_map m (U : 'M_(m, n)) : rstab rGf U^f = rstab rG U.
+Proof.
+by apply/setP=> x; rewrite !inE -!map_mxM inj_eq //; exact: map_mx_inj.
+Qed.
+
+Lemma rstabs_map m (U : 'M_(m, n)) : rstabs rGf U^f = rstabs rG U.
+Proof. by apply/setP=> x; rewrite !inE -!map_mxM ?map_submx. Qed.
+
+Lemma centgmx_map A : centgmx rGf A^f = centgmx rG A.
+Proof. by rewrite /centgmx rcent_map. Qed.
+
+Lemma mxmodule_map m (U : 'M_(m, n)) : mxmodule rGf U^f = mxmodule rG U.
+Proof. by rewrite /mxmodule rstabs_map. Qed.
+
+Lemma mxsimple_map (U : 'M_n) : mxsimple rGf U^f -> mxsimple rG U.
+Proof.
+case; rewrite map_mx_eq0 // mxmodule_map // => modU nzU minU.
+split=> // V modV sVU nzV; rewrite -(map_submx f).
+by rewrite (minU V^f) //= ?mxmodule_map ?map_mx_eq0 // map_submx.
+Qed.
+
+Lemma mx_irr_map : mx_irreducible rGf -> mx_irreducible rG.
+Proof. by move=> irrGf; apply: mxsimple_map; rewrite map_mx1. Qed.
+
+Lemma rker_map : rker rGf = rker rG.
+Proof. by rewrite /rker -rstab_map map_mx1. Qed.
+
+Lemma map_mx_faithful : mx_faithful rGf = mx_faithful rG.
+Proof. by rewrite /mx_faithful rker_map. Qed.
+
+Lemma map_mx_abs_irr :
+ mx_absolutely_irreducible rGf = mx_absolutely_irreducible rG.
+Proof.
+by rewrite /mx_absolutely_irreducible -map_enveloping_algebra_mx row_full_map.
+Qed.
+
+End OneRepresentation.
+
+Lemma mx_rsim_map n1 n2 rG1 rG2 :
+ @mx_rsim _ _ G n1 rG1 n2 rG2 -> mx_rsim (map_repr f rG1) (map_repr f rG2).
+Proof.
+case=> g eqn12 inj_g hom_g.
+by exists g^f => // [|x Gx]; rewrite ?row_free_map // -!map_mxM ?hom_g.
+Qed.
+
+Lemma map_section_repr n (rG : mx_representation aF G n) rGf U V
+ (modU : mxmodule rG U) (modV : mxmodule rG V)
+ (modUf : mxmodule rGf U^f) (modVf : mxmodule rGf V^f) :
+ map_repr f rG =1 rGf ->
+ mx_rsim (map_repr f (section_repr modU modV)) (section_repr modUf modVf).
+Proof.
+move=> def_rGf; set VU := <<_>>%MS.
+pose valUV := val_factmod (val_submod (1%:M : 'M[aF]_(\rank VU))).
+have sUV_Uf: (valUV^f <= U^f + V^f)%MS.
+ rewrite -map_addsmx map_submx; apply: submx_trans (proj_factmodS _ _).
+ by rewrite val_factmodS val_submod1 genmxE.
+exists (in_submod _ (in_factmod U^f valUV^f)) => [||x Gx].
+- rewrite !genmxE -(mxrank_map f) map_mxM map_col_base.
+ by case: (\rank (cokermx U)) / (mxrank_map _ _); rewrite map_cokermx.
+- rewrite -kermx_eq0 -submx0; apply/rV_subP=> u.
+ rewrite (sameP sub_kermxP eqP) submx0 -val_submod_eq0.
+ rewrite val_submodE -mulmxA -val_submodE in_submodK; last first.
+ by rewrite genmxE -(in_factmod_addsK _ V^f) submxMr.
+ rewrite in_factmodE mulmxA -in_factmodE in_factmod_eq0.
+ move/(submxMr (in_factmod U 1%:M *m in_submod VU 1%:M)^f).
+ rewrite -mulmxA -!map_mxM //; do 2!rewrite mulmxA -in_factmodE -in_submodE.
+ rewrite val_factmodK val_submodK map_mx1 mulmx1.
+ have ->: in_factmod U U = 0 by apply/eqP; rewrite in_factmod_eq0.
+ by rewrite linear0 map_mx0 eqmx0 submx0.
+rewrite {1}in_submodE mulmxA -in_submodE -in_submodJ; last first.
+ by rewrite genmxE -(in_factmod_addsK _ V^f) submxMr.
+congr (in_submod _ _); rewrite -in_factmodJ // in_factmodE mulmxA -in_factmodE.
+apply/eqP; rewrite -subr_eq0 -def_rGf -!map_mxM -linearB in_factmod_eq0.
+rewrite -map_mx_sub map_submx -in_factmod_eq0 linearB.
+rewrite /= (in_factmodJ modU) // val_factmodK.
+rewrite [valUV]val_factmodE mulmxA -val_factmodE val_factmodK.
+rewrite -val_submodE in_submodK ?subrr //.
+by rewrite mxmodule_trans ?section_module // val_submod1.
+Qed.
+
+Lemma map_regular_subseries U i (modU : mx_subseries (regular_repr aF G) U)
+ (modUf : mx_subseries (regular_repr rF G) [seq M^f | M <- U]) :
+ mx_rsim (map_repr f (subseries_repr i modU)) (subseries_repr i modUf).
+Proof.
+set mf := map _ in modUf *; rewrite /subseries_repr.
+do 2!move: (mx_subseries_module' _ _) (mx_subseries_module _ _).
+have mf_i V: nth 0^f (mf V) i = (V`_i)^f.
+ case: (ltnP i (size V)) => [ltiV | leVi]; first exact: nth_map.
+ by rewrite !nth_default ?size_map.
+rewrite -(map_mx0 f) mf_i (mf_i (0 :: U)) => modUi'f modUif modUi' modUi.
+by apply: map_section_repr; exact: map_regular_repr.
+Qed.
+
+Lemma extend_group_splitting_field :
+ group_splitting_field aF G -> group_splitting_field rF G.
+Proof.
+move=> splitG n rG irrG.
+have modU0: all ((mxmodule (regular_repr aF G)) #|G|) [::] by [].
+apply: (mx_Schreier modU0 _) => // [[U [compU lastU _]]]; have [modU _]:= compU.
+pose Uf := map ((map_mx f) _ _) U.
+have{lastU} lastUf: (last 0 Uf :=: 1%:M)%MS.
+ by rewrite -(map_mx0 f) -(map_mx1 f) last_map; exact/map_eqmx.
+have modUf: mx_subseries (regular_repr rF G) Uf.
+ rewrite /mx_subseries all_map; apply: etrans modU; apply: eq_all => Ui /=.
+ rewrite -mxmodule_map; apply: eq_subset_r => x.
+ by rewrite !inE map_regular_repr.
+have absUf i: i < size U -> mx_absolutely_irreducible (subseries_repr i modUf).
+ move=> lt_i_U; rewrite -(mx_rsim_abs_irr (map_regular_subseries i modU _)).
+ rewrite map_mx_abs_irr; apply: splitG.
+ by apply: mx_rsim_irr (mx_series_repr_irr compU lt_i_U); exact: section_eqmx.
+have compUf: mx_composition_series (regular_repr rF G) Uf.
+ split=> // i; rewrite size_map => ltiU.
+ move/max_submodP: (mx_abs_irrW (absUf i ltiU)); apply.
+ rewrite -{2}(map_mx0 f) -map_cons !(nth_map 0) ?leqW //.
+ by rewrite map_submx // ltmxW // (pathP _ (mx_series_lt compU)).
+have [[i ltiU] simUi] := rsim_regular_series irrG compUf lastUf.
+have{simUi} simUi: mx_rsim rG (subseries_repr i modUf).
+ by apply: mx_rsim_trans simUi _; exact: section_eqmx.
+by rewrite (mx_rsim_abs_irr simUi) absUf; rewrite size_map in ltiU.
+Qed.
+
+End ChangeOfField.
+
+(* Construction of a splitting field FA of an irreducible representation, for *)
+(* a matrix A in the centraliser of the representation. FA is the row-vector *)
+(* space of the matrix algebra generated by A with basis 1, A, ..., A ^+ d.-1 *)
+(* or, equivalently, the polynomials in {poly F} taken mod the (irreducible) *)
+(* minimal polynomial pA of A (of degree d). *)
+(* The details of the construction of FA are encapsulated in a submodule. *)
+Module Import MatrixGenField.
+
+Section GenField.
+
+Variables (F : fieldType) (gT : finGroupType) (G : {group gT}) (n' : nat).
+Local Notation n := n'.+1.
+Variables (rG : mx_representation F G n) (A : 'M[F]_n).
+
+Local Notation d := (degree_mxminpoly A).
+Local Notation Ad := (powers_mx A d).
+Local Notation pA := (mxminpoly A).
+Let d_gt0 := mxminpoly_nonconstant A.
+Local Notation irr := mx_irreducible.
+
+Record gen_of (irrG : irr rG) (cGA : centgmx rG A) := Gen {rVval : 'rV[F]_d}.
+Prenex Implicits rVval.
+
+Hypotheses (irrG : irr rG) (cGA : centgmx rG A).
+
+Notation FA := (gen_of irrG cGA).
+Let inFA := Gen irrG cGA.
+
+Canonical gen_subType := Eval hnf in [newType for @rVval irrG cGA].
+Definition gen_eqMixin := Eval hnf in [eqMixin of FA by <:].
+Canonical gen_eqType := Eval hnf in EqType FA gen_eqMixin.
+Definition gen_choiceMixin := [choiceMixin of FA by <:].
+Canonical gen_choiceType := Eval hnf in ChoiceType FA gen_choiceMixin.
+
+Definition gen0 := inFA 0.
+Definition genN (x : FA) := inFA (- val x).
+Definition genD (x y : FA) := inFA (val x + val y).
+
+Lemma gen_addA : associative genD.
+Proof. by move=> x y z; apply: val_inj; rewrite /= addrA. Qed.
+
+Lemma gen_addC : commutative genD.
+Proof. by move=> x y; apply: val_inj; rewrite /= addrC. Qed.
+
+Lemma gen_add0r : left_id gen0 genD.
+Proof. by move=> x; apply: val_inj; rewrite /= add0r. Qed.
+
+Lemma gen_addNr : left_inverse gen0 genN genD.
+Proof. by move=> x; apply: val_inj; rewrite /= addNr. Qed.
+
+Definition gen_zmodMixin := ZmodMixin gen_addA gen_addC gen_add0r gen_addNr.
+Canonical gen_zmodType := Eval hnf in ZmodType FA gen_zmodMixin.
+
+Definition pval (x : FA) := rVpoly (val x).
+
+Definition mxval (x : FA) := horner_mx A (pval x).
+
+Definition gen (x : F) := inFA (poly_rV x%:P).
+
+Lemma genK x : mxval (gen x) = x%:M.
+Proof.
+by rewrite /mxval [pval _]poly_rV_K ?horner_mx_C // size_polyC; case: (x != 0).
+Qed.
+
+Lemma mxval_inj : injective mxval.
+Proof. exact: inj_comp (@horner_rVpoly_inj _ _ A) val_inj. Qed.
+
+Lemma mxval0 : mxval 0 = 0.
+Proof. by rewrite /mxval [pval _]raddf0 rmorph0. Qed.
+
+Lemma mxvalN : {morph mxval : x / - x}.
+Proof. by move=> x; rewrite /mxval [pval _]raddfN rmorphN. Qed.
+
+Lemma mxvalD : {morph mxval : x y / x + y}.
+Proof. by move=> x y; rewrite /mxval [pval _]raddfD rmorphD. Qed.
+
+Definition mxval_sum := big_morph mxval mxvalD mxval0.
+
+Definition gen1 := inFA (poly_rV 1).
+Definition genM x y := inFA (poly_rV (pval x * pval y %% pA)).
+Definition genV x := inFA (poly_rV (mx_inv_horner A (mxval x)^-1)).
+
+Lemma mxval_gen1 : mxval gen1 = 1%:M.
+Proof. by rewrite /mxval [pval _]poly_rV_K ?size_poly1 // horner_mx_C. Qed.
+
+Lemma mxval_genM : {morph mxval : x y / genM x y >-> x *m y}.
+Proof.
+move=> x y; rewrite /mxval [pval _]poly_rV_K ?size_mod_mxminpoly //.
+by rewrite -horner_mxK mx_inv_hornerK ?horner_mx_mem // rmorphM.
+Qed.
+
+Lemma mxval_genV : {morph mxval : x / genV x >-> invmx x}.
+Proof.
+move=> x; rewrite /mxval [pval _]poly_rV_K ?size_poly ?mx_inv_hornerK //.
+pose m B : 'M[F]_(n * n) := lin_mx (mulmxr B); set B := mxval x.
+case uB: (B \is a GRing.unit); last by rewrite invr_out ?uB ?horner_mx_mem.
+have defAd: Ad = Ad *m m B *m m B^-1.
+ apply/row_matrixP=> i.
+ by rewrite !row_mul mul_rV_lin /= mx_rV_lin /= mulmxK ?vec_mxK.
+rewrite -[B^-1]mul1mx -(mul_vec_lin (mulmxr_linear _ _)) defAd submxMr //.
+rewrite -mxval_gen1 (submx_trans (horner_mx_mem _ _)) // {1}defAd.
+rewrite -(geq_leqif (mxrank_leqif_sup _)) ?mxrankM_maxl // -{}defAd.
+apply/row_subP=> i; rewrite row_mul rowK mul_vec_lin /= -{2}[A]horner_mx_X.
+by rewrite -rmorphX mulmxE -rmorphM horner_mx_mem.
+Qed.
+
+Lemma gen_mulA : associative genM.
+Proof. by move=> x y z; apply: mxval_inj; rewrite !mxval_genM mulmxA. Qed.
+
+Lemma gen_mulC : commutative genM.
+Proof. by move=> x y; rewrite /genM mulrC. Qed.
+
+Lemma gen_mul1r : left_id gen1 genM.
+Proof. by move=> x; apply: mxval_inj; rewrite mxval_genM mxval_gen1 mul1mx. Qed.
+
+Lemma gen_mulDr : left_distributive genM +%R.
+Proof.
+by move=> x y z; apply: mxval_inj; rewrite !(mxvalD, mxval_genM) mulmxDl.
+Qed.
+
+Lemma gen_ntriv : gen1 != 0.
+Proof. by rewrite -(inj_eq mxval_inj) mxval_gen1 mxval0 oner_eq0. Qed.
+
+Definition gen_ringMixin :=
+ ComRingMixin gen_mulA gen_mulC gen_mul1r gen_mulDr gen_ntriv.
+Canonical gen_ringType := Eval hnf in RingType FA gen_ringMixin.
+Canonical gen_comRingType := Eval hnf in ComRingType FA gen_mulC.
+
+Lemma mxval1 : mxval 1 = 1%:M. Proof. exact: mxval_gen1. Qed.
+
+Lemma mxvalM : {morph mxval : x y / x * y >-> x *m y}.
+Proof. exact: mxval_genM. Qed.
+
+Lemma mxval_sub : additive mxval.
+Proof. by move=> x y; rewrite mxvalD mxvalN. Qed.
+Canonical mxval_additive := Additive mxval_sub.
+
+Lemma mxval_is_multiplicative : multiplicative mxval.
+Proof. by split; [exact: mxvalM | exact: mxval1]. Qed.
+Canonical mxval_rmorphism := AddRMorphism mxval_is_multiplicative.
+
+Lemma mxval_centg x : centgmx rG (mxval x).
+Proof.
+rewrite [mxval _]horner_rVpoly -memmx_cent_envelop vec_mxK {x}mulmx_sub //.
+apply/row_subP=> k; rewrite rowK memmx_cent_envelop; apply/centgmxP => g Gg /=.
+by rewrite !mulmxE commrX // /GRing.comm -mulmxE (centgmxP cGA).
+Qed.
+
+Lemma gen_mulVr : GRing.Field.axiom genV.
+Proof.
+move=> x; rewrite -(inj_eq mxval_inj) mxval0.
+move/(mx_Schur irrG (mxval_centg x)) => u_x.
+by apply: mxval_inj; rewrite mxvalM mxval_genV mxval1 mulVmx.
+Qed.
+
+Lemma gen_invr0 : genV 0 = 0.
+Proof. by apply: mxval_inj; rewrite mxval_genV !mxval0 -{2}invr0. Qed.
+
+Definition gen_unitRingMixin := FieldUnitMixin gen_mulVr gen_invr0.
+Canonical gen_unitRingType := Eval hnf in UnitRingType FA gen_unitRingMixin.
+Canonical gen_comUnitRingType := Eval hnf in [comUnitRingType of FA].
+Definition gen_fieldMixin :=
+ @FieldMixin _ _ _ _ : GRing.Field.mixin_of gen_unitRingType.
+Definition gen_idomainMixin := FieldIdomainMixin gen_fieldMixin.
+Canonical gen_idomainType := Eval hnf in IdomainType FA gen_idomainMixin.
+Canonical gen_fieldType := Eval hnf in FieldType FA gen_fieldMixin.
+
+Lemma mxvalV : {morph mxval : x / x^-1 >-> invmx x}.
+Proof. exact: mxval_genV. Qed.
+
+Lemma gen_is_rmorphism : rmorphism gen.
+Proof.
+split=> [x y|]; first by apply: mxval_inj; rewrite genK !rmorphB /= !genK.
+by split=> // x y; apply: mxval_inj; rewrite genK !rmorphM /= !genK.
+Qed.
+Canonical gen_additive := Additive gen_is_rmorphism.
+Canonical gen_rmorphism := RMorphism gen_is_rmorphism.
+
+(* The generated field contains a root of the minimal polynomial (in some *)
+(* cases we want to use the construction solely for that purpose). *)
+
+Definition groot := inFA (poly_rV ('X %% pA)).
+
+Lemma mxval_groot : mxval groot = A.
+Proof.
+rewrite /mxval [pval _]poly_rV_K ?size_mod_mxminpoly // -horner_mxK.
+by rewrite mx_inv_hornerK ?horner_mx_mem // horner_mx_X.
+Qed.
+
+Lemma mxval_grootX k : mxval (groot ^+ k) = A ^+ k.
+Proof. by rewrite rmorphX /= mxval_groot. Qed.
+
+Lemma map_mxminpoly_groot : (map_poly gen pA).[groot] = 0.
+Proof. (* The [_ groot] prevents divergence of simpl. *)
+apply: mxval_inj; rewrite -horner_map [_ groot]/= mxval_groot mxval0.
+rewrite -(mx_root_minpoly A); congr ((_ : {poly _}).[A]).
+by apply/polyP=> i; rewrite 3!coef_map; exact: genK.
+Qed.
+
+(* Plugging the extension morphism gen into the ext_repr construction *)
+(* yields a (reducible) tensored representation. *)
+
+Lemma non_linear_gen_reducible :
+ d > 1 -> mxnonsimple (map_repr gen_rmorphism rG) 1%:M.
+Proof.
+rewrite ltnNge mxminpoly_linear_is_scalar => Anscal.
+pose Af := map_mx gen A; exists (kermx (Af - groot%:M)).
+rewrite submx1 kermx_centg_module /=; last first.
+ apply/centgmxP=> z Gz; rewrite mulmxBl mulmxBr scalar_mxC.
+ by rewrite -!map_mxM 1?(centgmxP cGA).
+rewrite andbC mxrank_ker -subn_gt0 mxrank1 subKn ?rank_leq_row // lt0n.
+rewrite mxrank_eq0 subr_eq0; case: eqP => [defAf | _].
+ rewrite -(map_mx_is_scalar gen_rmorphism) -/Af in Anscal.
+ by case/is_scalar_mxP: Anscal; exists groot.
+rewrite -mxrank_eq0 mxrank_ker subn_eq0 row_leq_rank.
+apply/row_freeP=> [[XA' XAK]].
+have pAf0: (mxminpoly Af).[groot] == 0.
+ by rewrite mxminpoly_map ?map_mxminpoly_groot.
+have{pAf0} [q def_pAf]:= factor_theorem _ _ pAf0.
+have q_nz: q != 0.
+ case: eqP (congr1 (fun p : {poly _} => size p) def_pAf) => // ->.
+ by rewrite size_mxminpoly mul0r size_poly0.
+have qAf0: horner_mx Af q = 0.
+ rewrite -[_ q]mulr1 -[1]XAK mulrA -{2}(horner_mx_X Af) -(horner_mx_C Af).
+ by rewrite -rmorphB -rmorphM -def_pAf /= mx_root_minpoly mul0r.
+have{qAf0} := dvdp_leq q_nz (mxminpoly_min qAf0); rewrite def_pAf.
+by rewrite size_Mmonic ?monicXsubC // polyseqXsubC addn2 ltnn.
+Qed.
+
+(* An alternative to the above, used in the proof of the p-stability of *)
+(* groups of odd order, is to reconsider the original vector space as a *)
+(* vector space of dimension n / e over FA. This is applicable only if G is *)
+(* the largest group represented on the original vector space (i.e., if we *)
+(* are not studying a representation of G induced by one of a larger group, *)
+(* as in B & G Theorem 2.6 for instance). We can't fully exploit one of the *)
+(* benefits of this approach -- that the type domain for the vector space can *)
+(* remain unchanged -- because we're restricting ourselves to row matrices; *)
+(* we have to use explicit bijections to convert between the two views. *)
+
+Definition subbase m (B : 'rV_m) : 'M_(m * d, n) :=
+ \matrix_ik mxvec (\matrix_(i, k) (row (B 0 i) (A ^+ k))) 0 ik.
+
+Lemma gen_dim_ex_proof : exists m, [exists B : 'rV_m, row_free (subbase B)].
+Proof. by exists 0%N; apply/existsP; exists 0. Qed.
+
+Lemma gen_dim_ub_proof m :
+ [exists B : 'rV_m, row_free (subbase B)] -> (m <= n)%N.
+Proof.
+case/existsP=> B /eqnP def_md.
+by rewrite (leq_trans _ (rank_leq_col (subbase B))) // def_md leq_pmulr.
+Qed.
+
+Definition gen_dim := ex_maxn gen_dim_ex_proof gen_dim_ub_proof.
+Notation m := gen_dim.
+
+Definition gen_base : 'rV_m := odflt 0 [pick B | row_free (subbase B)].
+Definition base := subbase gen_base.
+
+Lemma base_free : row_free base.
+Proof.
+rewrite /base /gen_base /m; case: pickP => //; case: ex_maxnP => m_max.
+by case/existsP=> B Bfree _ no_free; rewrite no_free in Bfree.
+Qed.
+
+Lemma base_full : row_full base.
+Proof.
+rewrite /row_full (eqnP base_free) /m; case: ex_maxnP => m.
+case/existsP=> /= B /eqnP Bfree m_max; rewrite -Bfree eqn_leq rank_leq_col.
+rewrite -{1}(mxrank1 F n) mxrankS //; apply/row_subP=> j; set u := row _ _.
+move/implyP: {m_max}(m_max m.+1); rewrite ltnn implybF.
+apply: contraR => nBj; apply/existsP.
+exists (row_mx (const_mx j : 'M_1) B); rewrite -row_leq_rank.
+pose Bj := Ad *m lin1_mx (mulmx u \o vec_mx).
+have rBj: \rank Bj = d.
+ apply/eqP; rewrite eqn_leq rank_leq_row -subn_eq0 -mxrank_ker mxrank_eq0 /=.
+ apply/rowV0P=> v /sub_kermxP; rewrite mulmxA mul_rV_lin1 /=.
+ rewrite -horner_rVpoly; pose x := inFA v; rewrite -/(mxval x).
+ have [[] // | nzx /(congr1 (mulmx^~ (mxval x^-1)))] := eqVneq x 0.
+ rewrite mul0mx /= -mulmxA -mxvalM divff // mxval1 mulmx1.
+ by move/rowP/(_ j)/eqP; rewrite !mxE !eqxx oner_eq0.
+rewrite {1}mulSn -Bfree -{1}rBj {rBj} -mxrank_disjoint_sum.
+ rewrite mxrankS // addsmx_sub -[m.+1]/(1 + m)%N; apply/andP; split.
+ apply/row_subP=> k; rewrite row_mul mul_rV_lin1 /=.
+ apply: eq_row_sub (mxvec_index (lshift _ 0) k) _.
+ by rewrite !rowK mxvecK mxvecE mxE row_mxEl mxE -row_mul mul1mx.
+ apply/row_subP; case/mxvec_indexP=> i k.
+ apply: eq_row_sub (mxvec_index (rshift 1 i) k) _.
+ by rewrite !rowK !mxvecE 2!mxE row_mxEr.
+apply/eqP/rowV0P=> v; rewrite sub_capmx => /andP[/submxP[w]].
+set x := inFA w; rewrite {Bj}mulmxA mul_rV_lin1 /= -horner_rVpoly -/(mxval x).
+have [-> | nzx ->] := eqVneq x 0; first by rewrite mxval0 mulmx0.
+move/(submxMr (mxval x^-1)); rewrite -mulmxA -mxvalM divff {nzx}//.
+rewrite mxval1 mulmx1 => Bx'j; rewrite (submx_trans Bx'j) in nBj => {Bx'j} //.
+apply/row_subP; case/mxvec_indexP=> i k.
+rewrite row_mul rowK mxvecE mxE rowE -mulmxA.
+have ->: A ^+ k *m mxval x^-1 = mxval (groot ^+ k / x).
+ by rewrite mxvalM rmorphX /= mxval_groot.
+rewrite [mxval _]horner_rVpoly; move: {k u x}(val _) => u.
+rewrite (mulmx_sum_row u) !linear_sum summx_sub //= => k _.
+rewrite !linearZ scalemx_sub //= rowK mxvecK -rowE.
+by apply: eq_row_sub (mxvec_index i k) _; rewrite rowK mxvecE mxE.
+Qed.
+
+Lemma gen_dim_factor : (m * d)%N = n.
+Proof. by rewrite -(eqnP base_free) (eqnP base_full). Qed.
+
+Lemma gen_dim_gt0 : m > 0.
+Proof. by case: posnP gen_dim_factor => // ->. Qed.
+
+Section Bijection.
+
+Variable m1 : nat.
+
+Definition in_gen (W : 'M[F]_(m1, n)) : 'M[FA]_(m1, m) :=
+ \matrix_(i, j) inFA (row j (vec_mx (row i W *m pinvmx base))).
+
+Definition val_gen (W : 'M[FA]_(m1, m)) : 'M[F]_(m1, n) :=
+ \matrix_i (mxvec (\matrix_j val (W i j)) *m base).
+
+Lemma in_genK : cancel in_gen val_gen.
+Proof.
+move=> W; apply/row_matrixP=> i; rewrite rowK; set w := row i W.
+have b_w: (w <= base)%MS by rewrite submx_full ?base_full.
+rewrite -{b_w}(mulmxKpV b_w); congr (_ *m _).
+by apply/rowP; case/mxvec_indexP=> j k; rewrite mxvecE !mxE.
+Qed.
+
+Lemma val_genK : cancel val_gen in_gen.
+Proof.
+move=> W; apply/matrixP=> i j; apply: val_inj; rewrite mxE /= rowK.
+case/row_freeP: base_free => B' BB'; rewrite -[_ *m _]mulmx1 -BB' mulmxA.
+by rewrite mulmxKpV ?submxMl // -mulmxA BB' mulmx1 mxvecK rowK.
+Qed.
+
+Lemma in_gen0 : in_gen 0 = 0.
+Proof. by apply/matrixP=> i j; rewrite !mxE !(mul0mx, linear0). Qed.
+
+Lemma val_gen0 : val_gen 0 = 0.
+Proof. by apply: (canLR in_genK); rewrite in_gen0. Qed.
+
+Lemma in_genN : {morph in_gen : W / - W}.
+Proof.
+move=> W; apply/matrixP=> i j; apply: val_inj.
+by rewrite !mxE !(mulNmx, linearN).
+Qed.
+
+Lemma val_genN : {morph val_gen : W / - W}.
+Proof. by move=> W; apply: (canLR in_genK); rewrite in_genN val_genK. Qed.
+
+Lemma in_genD : {morph in_gen : U V / U + V}.
+Proof.
+move=> U V; apply/matrixP=> i j; apply: val_inj.
+by rewrite !mxE !(mulmxDl, linearD).
+Qed.
+
+Lemma val_genD : {morph val_gen : U V / U + V}.
+Proof. by move=> U V; apply: (canLR in_genK); rewrite in_genD !val_genK. Qed.
+
+Definition in_gen_sum := big_morph in_gen in_genD in_gen0.
+Definition val_gen_sum := big_morph val_gen val_genD val_gen0.
+
+Lemma in_genZ a : {morph in_gen : W / a *: W >-> gen a *: W}.
+Proof.
+move=> W; apply/matrixP=> i j; apply: mxval_inj.
+rewrite !mxE mxvalM genK ![mxval _]horner_rVpoly /=.
+by rewrite mul_scalar_mx !(I, scalemxAl, linearZ).
+Qed.
+
+End Bijection.
+
+Prenex Implicits val_genK in_genK.
+
+Lemma val_gen_rV (w : 'rV_m) :
+ val_gen w = mxvec (\matrix_j val (w 0 j)) *m base.
+Proof. by apply/rowP=> j; rewrite mxE. Qed.
+
+Section Bijection2.
+
+Variable m1 : nat.
+
+Lemma val_gen_row W (i : 'I_m1) : val_gen (row i W) = row i (val_gen W).
+Proof.
+rewrite val_gen_rV rowK; congr (mxvec _ *m _).
+by apply/matrixP=> j k; rewrite !mxE.
+Qed.
+
+Lemma in_gen_row W (i : 'I_m1) : in_gen (row i W) = row i (in_gen W).
+Proof. by apply: (canLR val_genK); rewrite val_gen_row in_genK. Qed.
+
+Lemma row_gen_sum_mxval W (i : 'I_m1) :
+ row i (val_gen W) = \sum_j row (gen_base 0 j) (mxval (W i j)).
+Proof.
+rewrite -val_gen_row [row i W]row_sum_delta val_gen_sum.
+apply: eq_bigr => /= j _; rewrite mxE; move: {W i}(W i j) => x.
+have ->: x = \sum_k gen (val x 0 k) * inFA (delta_mx 0 k).
+ case: x => u; apply: mxval_inj; rewrite {1}[u]row_sum_delta.
+ rewrite mxval_sum [mxval _]horner_rVpoly mulmx_suml linear_sum /=.
+ apply: eq_bigr => k _; rewrite mxvalM genK [mxval _]horner_rVpoly /=.
+ by rewrite mul_scalar_mx -scalemxAl linearZ.
+rewrite scaler_suml val_gen_sum mxval_sum linear_sum; apply: eq_bigr => k _.
+rewrite mxvalM genK mul_scalar_mx linearZ [mxval _]horner_rVpoly /=.
+rewrite -scalerA; apply: (canLR in_genK); rewrite in_genZ; congr (_ *: _).
+apply: (canRL val_genK); transitivity (row (mxvec_index j k) base); last first.
+ by rewrite -rowE rowK mxvecE mxE rowK mxvecK.
+rewrite rowE -mxvec_delta -[val_gen _](row_id 0) rowK /=; congr (mxvec _ *m _).
+apply/row_matrixP=> j'; rewrite rowK !mxE mulr_natr rowE mul_delta_mx_cond.
+by rewrite !mulrb (fun_if rVval).
+Qed.
+
+Lemma val_genZ x : {morph @val_gen m1 : W / x *: W >-> W *m mxval x}.
+Proof.
+move=> W; apply/row_matrixP=> i; rewrite row_mul !row_gen_sum_mxval.
+by rewrite mulmx_suml; apply: eq_bigr => j _; rewrite mxE mulrC mxvalM row_mul.
+Qed.
+
+End Bijection2.
+
+Lemma submx_in_gen m1 m2 (U : 'M_(m1, n)) (V : 'M_(m2, n)) :
+ (U <= V -> in_gen U <= in_gen V)%MS.
+Proof.
+move=> sUV; apply/row_subP=> i; rewrite -in_gen_row.
+case/submxP: (row_subP sUV i) => u ->{i}.
+rewrite mulmx_sum_row in_gen_sum summx_sub // => j _.
+by rewrite in_genZ in_gen_row scalemx_sub ?row_sub.
+Qed.
+
+Lemma submx_in_gen_eq m1 m2 (U : 'M_(m1, n)) (V : 'M_(m2, n)) :
+ (V *m A <= V -> (in_gen U <= in_gen V) = (U <= V))%MS.
+Proof.
+move=> sVA_V; apply/idP/idP=> siUV; last exact: submx_in_gen.
+apply/row_subP=> i; rewrite -[row i U]in_genK in_gen_row.
+case/submxP: (row_subP siUV i) => u ->{i U siUV}.
+rewrite mulmx_sum_row val_gen_sum summx_sub // => j _.
+rewrite val_genZ val_gen_row in_genK rowE -mulmxA mulmx_sub //.
+rewrite [mxval _]horner_poly mulmx_sumr summx_sub // => [[k _]] _ /=.
+rewrite mulmxA mul_mx_scalar -scalemxAl scalemx_sub {u j}//.
+elim: k => [|k IHk]; first by rewrite mulmx1.
+by rewrite exprSr mulmxA (submx_trans (submxMr A IHk)).
+Qed.
+
+Definition gen_mx g := \matrix_i in_gen (row (gen_base 0 i) (rG g)).
+
+Let val_genJmx m :
+ {in G, forall g, {morph @val_gen m : W / W *m gen_mx g >-> W *m rG g}}.
+Proof.
+move=> g Gg /= W; apply/row_matrixP=> i; rewrite -val_gen_row !row_mul.
+rewrite mulmx_sum_row val_gen_sum row_gen_sum_mxval mulmx_suml.
+apply: eq_bigr => /= j _; rewrite val_genZ rowK in_genK mxE -!row_mul.
+by rewrite (centgmxP (mxval_centg _)).
+Qed.
+
+Lemma gen_mx_repr : mx_repr G gen_mx.
+Proof.
+split=> [|g h Gg Gh]; apply: (can_inj val_genK).
+ by rewrite -[gen_mx 1]mul1mx val_genJmx // repr_mx1 mulmx1.
+rewrite {1}[val_gen]lock -[gen_mx g]mul1mx !val_genJmx // -mulmxA -repr_mxM //.
+by rewrite -val_genJmx ?groupM ?mul1mx -?lock.
+Qed.
+Canonical gen_repr := MxRepresentation gen_mx_repr.
+Local Notation rGA := gen_repr.
+
+Lemma val_genJ m :
+ {in G, forall g, {morph @val_gen m : W / W *m rGA g >-> W *m rG g}}.
+Proof. exact: val_genJmx. Qed.
+
+Lemma in_genJ m :
+ {in G, forall g, {morph @in_gen m : v / v *m rG g >-> v *m rGA g}}.
+Proof.
+by move=> g Gg /= v; apply: (canLR val_genK); rewrite val_genJ ?in_genK.
+Qed.
+
+Lemma rfix_gen (H : {set gT}) :
+ H \subset G -> (rfix_mx rGA H :=: in_gen (rfix_mx rG H))%MS.
+Proof.
+move/subsetP=> sHG; apply/eqmxP/andP; split; last first.
+ by apply/rfix_mxP=> g Hg; rewrite -in_genJ ?sHG ?rfix_mx_id.
+rewrite -[rfix_mx rGA H]val_genK; apply: submx_in_gen.
+by apply/rfix_mxP=> g Hg; rewrite -val_genJ ?rfix_mx_id ?sHG.
+Qed.
+
+Definition rowval_gen m1 U :=
+ <<\matrix_ik
+ mxvec (\matrix_(i < m1, k < d) (row i (val_gen U) *m A ^+ k)) 0 ik>>%MS.
+
+Lemma submx_rowval_gen m1 m2 (U : 'M_(m1, n)) (V : 'M_(m2, m)) :
+ (U <= rowval_gen V)%MS = (in_gen U <= V)%MS.
+Proof.
+rewrite genmxE; apply/idP/idP=> sUV.
+ apply: submx_trans (submx_in_gen sUV) _.
+ apply/row_subP; case/mxvec_indexP=> i k; rewrite -in_gen_row rowK mxvecE mxE.
+ rewrite -mxval_grootX -val_gen_row -val_genZ val_genK scalemx_sub //.
+ exact: row_sub.
+rewrite -[U]in_genK; case/submxP: sUV => u ->{U}.
+apply/row_subP=> i0; rewrite -val_gen_row row_mul; move: {i0 u}(row _ u) => u.
+rewrite mulmx_sum_row val_gen_sum summx_sub // => i _.
+rewrite val_genZ [mxval _]horner_rVpoly [_ *m Ad]mulmx_sum_row.
+rewrite !linear_sum summx_sub // => k _.
+rewrite !linearZ scalemx_sub {u}//= rowK mxvecK val_gen_row.
+by apply: (eq_row_sub (mxvec_index i k)); rewrite rowK mxvecE mxE.
+Qed.
+
+Lemma rowval_genK m1 (U : 'M_(m1, m)) : (in_gen (rowval_gen U) :=: U)%MS.
+Proof.
+apply/eqmxP; rewrite -submx_rowval_gen submx_refl /=.
+by rewrite -{1}[U]val_genK submx_in_gen // submx_rowval_gen val_genK.
+Qed.
+
+Lemma rowval_gen_stable m1 (U : 'M_(m1, m)) :
+ (rowval_gen U *m A <= rowval_gen U)%MS.
+Proof.
+rewrite -[A]mxval_groot -{1}[_ U]in_genK -val_genZ.
+by rewrite submx_rowval_gen val_genK scalemx_sub // rowval_genK.
+Qed.
+
+Lemma rstab_in_gen m1 (U : 'M_(m1, n)) : rstab rGA (in_gen U) = rstab rG U.
+Proof.
+apply/setP=> x; rewrite !inE; case Gx: (x \in G) => //=.
+by rewrite -in_genJ // (inj_eq (can_inj in_genK)).
+Qed.
+
+Lemma rstabs_in_gen m1 (U : 'M_(m1, n)) :
+ rstabs rG U \subset rstabs rGA (in_gen U).
+Proof.
+apply/subsetP=> x; rewrite !inE => /andP[Gx nUx].
+by rewrite -in_genJ Gx // submx_in_gen.
+Qed.
+
+Lemma rstabs_rowval_gen m1 (U : 'M_(m1, m)) :
+ rstabs rG (rowval_gen U) = rstabs rGA U.
+Proof.
+apply/setP=> x; rewrite !inE; case Gx: (x \in G) => //=.
+by rewrite submx_rowval_gen in_genJ // (eqmxMr _ (rowval_genK U)).
+Qed.
+
+Lemma mxmodule_rowval_gen m1 (U : 'M_(m1, m)) :
+ mxmodule rG (rowval_gen U) = mxmodule rGA U.
+Proof. by rewrite /mxmodule rstabs_rowval_gen. Qed.
+
+Lemma gen_mx_irr : mx_irreducible rGA.
+Proof.
+apply/mx_irrP; split=> [|U Umod nzU]; first exact: gen_dim_gt0.
+rewrite -sub1mx -rowval_genK -submx_rowval_gen submx_full //.
+case/mx_irrP: irrG => _; apply; first by rewrite mxmodule_rowval_gen.
+rewrite -(inj_eq (can_inj in_genK)) in_gen0.
+by rewrite -mxrank_eq0 rowval_genK mxrank_eq0.
+Qed.
+
+Lemma rker_gen : rker rGA = rker rG.
+Proof.
+apply/setP=> g; rewrite !inE !mul1mx; case Gg: (g \in G) => //=.
+apply/eqP/eqP=> g1; apply/row_matrixP=> i.
+ by apply: (can_inj in_genK); rewrite rowE in_genJ //= g1 mulmx1 row1.
+by apply: (can_inj val_genK); rewrite rowE val_genJ //= g1 mulmx1 row1.
+Qed.
+
+Lemma gen_mx_faithful : mx_faithful rGA = mx_faithful rG.
+Proof. by rewrite /mx_faithful rker_gen. Qed.
+
+End GenField.
+
+Section DecideGenField.
+
+Import MatrixFormula.
+
+Variable F : decFieldType.
+
+Local Notation False := GRing.False.
+Local Notation True := GRing.True.
+Local Notation Bool b := (GRing.Bool b%bool).
+Local Notation term := (GRing.term F).
+Local Notation form := (GRing.formula F).
+
+Local Notation morphAnd f := ((big_morph f) true andb).
+
+Variables (gT : finGroupType) (G : {group gT}) (n' : nat).
+Local Notation n := n'.+1.
+Variables (rG : mx_representation F G n) (A : 'M[F]_n).
+Hypotheses (irrG : mx_irreducible rG) (cGA : centgmx rG A).
+Local Notation FA := (gen_of irrG cGA).
+Local Notation inFA := (Gen irrG cGA).
+
+Local Notation d := (degree_mxminpoly A).
+Let d_gt0 : d > 0 := mxminpoly_nonconstant A.
+Local Notation Ad := (powers_mx A d).
+
+Let mxT (u : 'rV_d) := vec_mx (mulmx_term u (mx_term Ad)).
+
+Let eval_mxT e u : eval_mx e (mxT u) = mxval (inFA (eval_mx e u)).
+Proof.
+by rewrite eval_vec_mx eval_mulmx eval_mx_term [mxval _]horner_rVpoly.
+Qed.
+
+Let Ad'T := mx_term (pinvmx Ad).
+Let mulT (u v : 'rV_d) := mulmx_term (mxvec (mulmx_term (mxT u) (mxT v))) Ad'T.
+
+Lemma eval_mulT e u v :
+ eval_mx e (mulT u v) = val (inFA (eval_mx e u) * inFA (eval_mx e v)).
+Proof.
+rewrite !(eval_mulmx, eval_mxvec) !eval_mxT eval_mx_term.
+by apply: (can_inj (@rVpolyK _ _)); rewrite -mxvalM [rVpoly _]horner_rVpolyK.
+Qed.
+
+Fixpoint gen_term t := match t with
+| 'X_k => row_var _ d k
+| x%:T => mx_term (val (x : FA))
+| n1%:R => mx_term (val (n1%:R : FA))%R
+| t1 + t2 => \row_i (gen_term t1 0%R i + gen_term t2 0%R i)
+| - t1 => \row_i (- gen_term t1 0%R i)
+| t1 *+ n1 => mulmx_term (mx_term n1%:R%:M)%R (gen_term t1)
+| t1 * t2 => mulT (gen_term t1) (gen_term t2)
+| t1^-1 => gen_term t1
+| t1 ^+ n1 => iter n1 (mulT (gen_term t1)) (mx_term (val (1%R : FA)))
+end%T.
+
+Definition gen_env (e : seq FA) := row_env (map val e).
+
+Lemma nth_map_rVval (e : seq FA) j : (map val e)`_j = val e`_j.
+Proof.
+case: (ltnP j (size e)) => [| leej]; first exact: (nth_map 0 0).
+by rewrite !nth_default ?size_map.
+Qed.
+
+Lemma set_nth_map_rVval (e : seq FA) j v :
+ set_nth 0 (map val e) j v = map val (set_nth 0 e j (inFA v)).
+Proof.
+apply: (@eq_from_nth _ 0) => [|k _]; first by rewrite !(size_set_nth, size_map).
+by rewrite !(nth_map_rVval, nth_set_nth) /= nth_map_rVval [rVval _]fun_if.
+Qed.
+
+Lemma eval_gen_term e t :
+ GRing.rterm t -> eval_mx (gen_env e) (gen_term t) = val (GRing.eval e t).
+Proof.
+elim: t => //=.
+- by move=> k _; apply/rowP=> i; rewrite !mxE /= nth_row_env nth_map_rVval.
+- by move=> x _; rewrite eval_mx_term.
+- by move=> x _; rewrite eval_mx_term.
+- move=> t1 IH1 t2 IH2 /andP[rt1 rt2]; rewrite -{}IH1 // -{}IH2 //.
+ by apply/rowP=> k; rewrite !mxE.
+- by move=> t1 IH1 rt1; rewrite -{}IH1 //; apply/rowP=> k; rewrite !mxE.
+- move=> t1 IH1 n1 rt1; rewrite eval_mulmx eval_mx_term mul_scalar_mx.
+ by rewrite scaler_nat {}IH1 //; elim: n1 => //= n1 IHn1; rewrite !mulrS IHn1.
+- by move=> t1 IH1 t2 IH2 /andP[rt1 rt2]; rewrite eval_mulT IH1 ?IH2.
+move=> t1 IH1 n1 /IH1 {IH1}IH1.
+elim: n1 => [|n1 IHn1] /=; first by rewrite eval_mx_term.
+by rewrite eval_mulT exprS IH1 IHn1.
+Qed.
+
+(* WARNING: Coq will core dump if the Notation Bool is used in the match *)
+(* pattern here. *)
+Fixpoint gen_form f := match f with
+| GRing.Bool b => Bool b
+| t1 == t2 => mxrank_form 0 (gen_term (t1 - t2))
+| GRing.Unit t1 => mxrank_form 1 (gen_term t1)
+| f1 /\ f2 => gen_form f1 /\ gen_form f2
+| f1 \/ f2 => gen_form f1 \/ gen_form f2
+| f1 ==> f2 => gen_form f1 ==> gen_form f2
+| ~ f1 => ~ gen_form f1
+| ('exists 'X_k, f1) => Exists_row_form d k (gen_form f1)
+| ('forall 'X_k, f1) => ~ Exists_row_form d k (~ (gen_form f1))
+end%T.
+
+Lemma sat_gen_form e f : GRing.rformula f ->
+ reflect (GRing.holds e f) (GRing.sat (gen_env e) (gen_form f)).
+Proof.
+have ExP := Exists_rowP; have set_val := set_nth_map_rVval.
+elim: f e => //.
+- by move=> b e _; exact: (iffP satP).
+- rewrite /gen_form => t1 t2 e rt_t; set t := (_ - _)%T.
+ have:= GRing.qf_evalP (gen_env e) (mxrank_form_qf 0 (gen_term t)).
+ rewrite eval_mxrank mxrank_eq0 eval_gen_term // => tP.
+ by rewrite (sameP satP tP) /= subr_eq0 val_eqE; exact: eqP.
+- move=> f1 IH1 f2 IH2 s /= /andP[/(IH1 s)f1P /(IH2 s)f2P].
+ by apply: (iffP satP) => [[/satP/f1P ? /satP/f2P] | [/f1P/satP ? /f2P/satP]].
+- move=> f1 IH1 f2 IH2 s /= /andP[/(IH1 s)f1P /(IH2 s)f2P].
+ by apply: (iffP satP) => /= [] [];
+ try move/satP; do [move/f1P | move/f2P]; try move/satP; auto.
+- move=> f1 IH1 f2 IH2 s /= /andP[/(IH1 s)f1P /(IH2 s)f2P].
+ by apply: (iffP satP) => /= implP;
+ try move/satP; move/f1P; try move/satP; move/implP;
+ try move/satP; move/f2P; try move/satP.
+- move=> f1 IH1 s /= /(IH1 s) f1P.
+ by apply: (iffP satP) => /= notP; try move/satP; move/f1P; try move/satP.
+- move=> k f1 IHf1 s /IHf1 f1P; apply: (iffP satP) => /= [|[[v f1v]]].
+ by case/ExP=> // x /satP; rewrite set_val => /f1P; exists (inFA x).
+ by apply/ExP=> //; exists v; rewrite set_val; apply/satP/f1P.
+move=> i f1 IHf1 s /IHf1 f1P; apply: (iffP satP) => /= allf1 => [[v]|].
+ apply/f1P; case: satP => // notf1x; case: allf1; apply/ExP=> //.
+ by exists v; rewrite set_val.
+by case/ExP=> //= v []; apply/satP; rewrite set_val; apply/f1P.
+Qed.
+
+Definition gen_sat e f := GRing.sat (gen_env e) (gen_form (GRing.to_rform f)).
+
+Lemma gen_satP : GRing.DecidableField.axiom gen_sat.
+Proof.
+move=> e f; have [tor rto] := GRing.to_rformP e f.
+exact: (iffP (sat_gen_form e (GRing.to_rform_rformula f))).
+Qed.
+
+Definition gen_decFieldMixin := DecFieldMixin gen_satP.
+
+Canonical gen_decFieldType := Eval hnf in DecFieldType FA gen_decFieldMixin.
+
+End DecideGenField.
+
+Section FiniteGenField.
+
+Variables (F : finFieldType) (gT : finGroupType) (G : {group gT}) (n' : nat).
+Local Notation n := n'.+1.
+Variables (rG : mx_representation F G n) (A : 'M[F]_n).
+Hypotheses (irrG : mx_irreducible rG) (cGA : centgmx rG A).
+Notation FA := (gen_of irrG cGA).
+
+(* This should be [countMixin of FA by <:]*)
+Definition gen_countMixin := (sub_countMixin (gen_subType irrG cGA)).
+Canonical gen_countType := Eval hnf in CountType FA gen_countMixin.
+Canonical gen_subCountType := Eval hnf in [subCountType of FA].
+Definition gen_finMixin := [finMixin of FA by <:].
+Canonical gen_finType := Eval hnf in FinType FA gen_finMixin.
+Canonical gen_subFinType := Eval hnf in [subFinType of FA].
+Canonical gen_finZmodType := Eval hnf in [finZmodType of FA].
+Canonical gen_baseFinGroupType := Eval hnf in [baseFinGroupType of FA for +%R].
+Canonical gen_finGroupType := Eval hnf in [finGroupType of FA for +%R].
+Canonical gen_finRingType := Eval hnf in [finRingType of FA].
+Canonical gen_finComRingType := Eval hnf in [finComRingType of FA].
+Canonical gen_finUnitRingType := Eval hnf in [finUnitRingType of FA].
+Canonical gen_finComUnitRingType := Eval hnf in [finComUnitRingType of FA].
+Canonical gen_finIdomainType := Eval hnf in [finIdomainType of FA].
+Canonical gen_finFieldType := Eval hnf in [finFieldType of FA].
+
+Lemma card_gen : #|{:FA}| = (#|F| ^ degree_mxminpoly A)%N.
+Proof. by rewrite card_sub card_matrix mul1n. Qed.
+
+End FiniteGenField.
+
+End MatrixGenField.
+
+Canonical gen_subType.
+Canonical gen_eqType.
+Canonical gen_choiceType.
+Canonical gen_countType.
+Canonical gen_subCountType.
+Canonical gen_finType.
+Canonical gen_subFinType.
+Canonical gen_zmodType.
+Canonical gen_finZmodType.
+Canonical gen_baseFinGroupType.
+Canonical gen_finGroupType.
+Canonical gen_ringType.
+Canonical gen_finRingType.
+Canonical gen_comRingType.
+Canonical gen_finComRingType.
+Canonical gen_unitRingType.
+Canonical gen_finUnitRingType.
+Canonical gen_comUnitRingType.
+Canonical gen_finComUnitRingType.
+Canonical gen_idomainType.
+Canonical gen_finIdomainType.
+Canonical gen_fieldType.
+Canonical gen_finFieldType.
+Canonical gen_decFieldType.
+
+(* Classical splitting and closure field constructions provide convenient *)
+(* packaging for the pointwise construction. *)
+Section BuildSplittingField.
+
+Implicit Type gT : finGroupType.
+Implicit Type F : fieldType.
+
+Lemma group_splitting_field_exists gT (G : {group gT}) F :
+ classically {Fs : fieldType & {rmorphism F -> Fs}
+ & group_splitting_field Fs G}.
+Proof.
+move: F => F0 [] // nosplit; pose nG := #|G|; pose aG F := regular_repr F G.
+pose m := nG.+1; pose F := F0; pose U : seq 'M[F]_nG := [::].
+suffices: size U + m <= nG by rewrite ltnn.
+have: mx_subseries (aG F) U /\ path ltmx 0 U by [].
+pose f : {rmorphism F0 -> F} := [rmorphism of idfun].
+elim: m F U f => [|m IHm] F U f [modU ltU].
+ by rewrite addn0 (leq_trans (max_size_mx_series ltU)) ?rank_leq_row.
+rewrite addnS ltnNge -implybF; apply/implyP=> le_nG_Um; apply nosplit.
+exists F => //; case=> [|n] rG irrG; first by case/mx_irrP: irrG.
+apply/idPn=> nabsG; pose cG := ('C(enveloping_algebra_mx rG))%MS.
+have{nabsG} [A]: exists2 A, (A \in cG)%MS & ~~ is_scalar_mx A.
+ apply/has_non_scalar_mxP; rewrite ?scalar_mx_cent // ltnNge.
+ by apply: contra nabsG; exact: cent_mx_scalar_abs_irr.
+rewrite {cG}memmx_cent_envelop -mxminpoly_linear_is_scalar -ltnNge => cGA.
+move/(non_linear_gen_reducible irrG cGA).
+set F' := gen_fieldType _ _; set rG' := @map_repr _ F' _ _ _ _ rG.
+move: F' (gen_rmorphism _ _ : {rmorphism F -> F'}) => F' f' in rG' * => irrG'.
+pose U' := [seq map_mx f' Ui | Ui <- U].
+have modU': mx_subseries (aG F') U'.
+ apply: etrans modU; rewrite /mx_subseries all_map; apply: eq_all => Ui.
+ rewrite -(mxmodule_map f'); apply: eq_subset_r => x.
+ by rewrite !inE map_regular_repr.
+case: notF; apply: (mx_Schreier modU ltU) => [[V [compV lastV sUV]]].
+have{lastV} [] := rsim_regular_series irrG compV lastV.
+have{sUV} defV: V = U.
+ apply/eqP; rewrite eq_sym -(geq_leqif (size_subseq_leqif sUV)).
+ rewrite -(leq_add2r m); apply: leq_trans le_nG_Um.
+ by apply: IHm f _; rewrite (mx_series_lt compV); case: compV.
+rewrite {V}defV in compV * => i rsimVi.
+apply: (mx_Schreier modU') => [|[V' [compV' _ sUV']]].
+ rewrite {modU' compV modU i le_nG_Um rsimVi}/U' -(map_mx0 f').
+ by apply: etrans ltU; elim: U 0 => //= Ui U IHU Ui'; rewrite IHU map_ltmx.
+have{sUV'} defV': V' = U'; last rewrite {V'}defV' in compV'.
+ apply/eqP; rewrite eq_sym -(geq_leqif (size_subseq_leqif sUV')) size_map.
+ rewrite -(leq_add2r m); apply: leq_trans le_nG_Um.
+ apply: IHm [rmorphism of f' \o f] _.
+ by rewrite (mx_series_lt compV'); case: compV'.
+suffices{irrG'}: mx_irreducible rG' by case/mxsimpleP=> _ _ [].
+have ltiU': i < size U' by rewrite size_map.
+apply: mx_rsim_irr (mx_rsim_sym _ ) (mx_series_repr_irr compV' ltiU').
+apply: mx_rsim_trans (mx_rsim_map f' rsimVi) _; exact: map_regular_subseries.
+Qed.
+
+Lemma group_closure_field_exists gT F :
+ classically {Fs : fieldType & {rmorphism F -> Fs}
+ & group_closure_field Fs gT}.
+Proof.
+set n := #|{group gT}|.
+suffices: classically {Fs : fieldType & {rmorphism F -> Fs}
+ & forall G : {group gT}, enum_rank G < n -> group_splitting_field Fs G}.
+- apply: classic_bind => [[Fs f splitFs]] _ -> //.
+ by exists Fs => // G; exact: splitFs.
+elim: (n) => [|i IHi]; first by move=> _ -> //; exists F => //; exists id.
+apply: classic_bind IHi => [[F' f splitF']].
+have [le_n_i _ -> // | lt_i_n] := leqP n i.
+ by exists F' => // G _; apply: splitF'; exact: leq_trans le_n_i.
+have:= @group_splitting_field_exists _ (enum_val (Ordinal lt_i_n)) F'.
+apply: classic_bind => [[Fs f' splitFs]] _ -> //.
+exists Fs => [|G]; first exact: [rmorphism of (f' \o f)].
+rewrite ltnS leq_eqVlt -{1}[i]/(val (Ordinal lt_i_n)) val_eqE.
+case/predU1P=> [defG | ltGi]; first by rewrite -[G]enum_rankK defG.
+by apply: (extend_group_splitting_field f'); exact: splitF'.
+Qed.
+
+Lemma group_closure_closed_field (F : closedFieldType) gT :
+ group_closure_field F gT.
+Proof.
+move=> G [|n] rG irrG; first by case/mx_irrP: irrG.
+apply: cent_mx_scalar_abs_irr => //; rewrite leqNgt.
+apply/(has_non_scalar_mxP (scalar_mx_cent _ _)) => [[A cGA nscalA]].
+have [a]: exists a, eigenvalue A a.
+ pose P := mxminpoly A; pose d := degree_mxminpoly A.
+ have Pd1: P`_d = 1.
+ by rewrite -(eqP (mxminpoly_monic A)) /lead_coef size_mxminpoly.
+ have d_gt0: d > 0 := mxminpoly_nonconstant A.
+ have [a def_ad] := solve_monicpoly (nth 0 (- P)) d_gt0.
+ exists a; rewrite eigenvalue_root_min -/P /root -oppr_eq0 -hornerN.
+ rewrite horner_coef size_opp size_mxminpoly -/d big_ord_recr -def_ad.
+ by rewrite coefN Pd1 mulN1r /= subrr.
+case/negP; rewrite kermx_eq0 row_free_unit (mx_Schur irrG) ?subr_eq0 //.
+ by rewrite -memmx_cent_envelop -raddfN linearD addmx_sub ?scalar_mx_cent.
+by apply: contraNneq nscalA => ->; exact: scalar_mx_is_scalar.
+Qed.
+
+End BuildSplittingField.
diff --git a/mathcomp/character/vcharacter.v b/mathcomp/character/vcharacter.v
new file mode 100644
index 0000000..3ef364d
--- /dev/null
+++ b/mathcomp/character/vcharacter.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.
+Require Import fingroup morphism perm automorphism quotient finalg action.
+Require Import gproduct zmodp commutator cyclic center pgroup sylow frobenius.
+Require Import vector ssrnum ssrint intdiv algC algnum.
+Require Import classfun character integral_char.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope GRing.Theory Num.Theory.
+Local Open Scope ring_scope.
+
+(******************************************************************************)
+(* This file provides basic notions of virtual character theory: *)
+(* 'Z[S, A] == collective predicate for the phi that are Z-linear *)
+(* combinations of elements of S : seq 'CF(G) and have *)
+(* support in A : {set gT}. *)
+(* 'Z[S] == collective predicate for the Z-linear combinations of *)
+(* elements of S. *)
+(* 'Z[irr G] == the collective predicate for virtual characters. *)
+(* dirr G == the collective predicate for normal virtual characters, *)
+(* i.e., virtual characters of norm 1: *)
+(* mu \in dirr G <=> m \in 'Z[irr G] and '[mu] = 1 *)
+(* <=> mu or - mu \in irr G. *)
+(* --> othonormal subsets of 'Z[irr G] are contained in dirr G. *)
+(* dIirr G == an index type for normal virtual characters. *)
+(* dchi i == the normal virtual character of index i. *)
+(* of_irr i == the (unique) irreducible constituent of dchi i: *)
+(* dchi i = 'chi_(of_irr i) or - 'chi_(of_irr i). *)
+(* ndirr i == the index of - dchi i. *)
+(* dirr1 G == the normal virtual character index of 1 : 'CF(G), the *)
+(* principal character. *)
+(* dirr_dIirr j f == the index i (or dirr1 G if it does not exist) such that *)
+(* dchi i = f j. *)
+(* dirr_constt phi == the normal virtual character constituents of phi: *)
+(* i \in dirr_constt phi <=> [dchi i, phi] > 0. *)
+(* to_dirr phi i == the normal virtual character constituent of phi with an *)
+(* irreducible constituent i, when i \in irr_constt phi. *)
+(******************************************************************************)
+
+Section Basics.
+
+Variables (gT : finGroupType) (B : {set gT}) (S : seq 'CF(B)) (A : {set gT}).
+
+Definition Zchar : pred_class :=
+ [pred phi in 'CF(B, A) | dec_Cint_span (in_tuple S) phi].
+Fact Zchar_key : pred_key Zchar. Proof. by []. Qed.
+Canonical Zchar_keyed := KeyedPred Zchar_key.
+
+Lemma cfun0_zchar : 0 \in Zchar.
+Proof.
+rewrite inE mem0v; apply/sumboolP; exists 0.
+by rewrite big1 // => i _; rewrite ffunE.
+Qed.
+
+Fact Zchar_zmod : zmod_closed Zchar.
+Proof.
+split; first exact: cfun0_zchar.
+move=> phi xi /andP[Aphi /sumboolP[a Da]] /andP[Axi /sumboolP[b Db]].
+rewrite inE rpredB // Da Db -sumrB; apply/sumboolP; exists (a - b).
+by apply: eq_bigr => i _; rewrite -mulrzBr !ffunE.
+Qed.
+Canonical Zchar_opprPred := OpprPred Zchar_zmod.
+Canonical Zchar_addrPred := AddrPred Zchar_zmod.
+Canonical Zchar_zmodPred := ZmodPred Zchar_zmod.
+
+Lemma scale_zchar a phi : a \in Cint -> phi \in Zchar -> a *: phi \in Zchar.
+Proof. by case/CintP=> m -> Zphi; rewrite scaler_int rpredMz. Qed.
+
+End Basics.
+
+Notation "''Z[' S , A ]" := (Zchar S A)
+ (at level 8, format "''Z[' S , A ]") : group_scope.
+Notation "''Z[' S ]" := 'Z[S, setT]
+ (at level 8, format "''Z[' S ]") : group_scope.
+
+Section Zchar.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Implicit Types (A B : {set gT}) (S : seq 'CF(G)).
+
+Lemma zchar_split S A phi :
+ phi \in 'Z[S, A] = (phi \in 'Z[S]) && (phi \in 'CF(G, A)).
+Proof. by rewrite !inE cfun_onT andbC. Qed.
+
+Lemma zcharD1E phi S : (phi \in 'Z[S, G^#]) = (phi \in 'Z[S]) && (phi 1%g == 0).
+Proof. by rewrite zchar_split cfunD1E. Qed.
+
+Lemma zcharD1 phi S A :
+ (phi \in 'Z[S, A^#]) = (phi \in 'Z[S, A]) && (phi 1%g == 0).
+Proof. by rewrite zchar_split cfun_onD1 andbA -zchar_split. Qed.
+
+Lemma zcharW S A : {subset 'Z[S, A] <= 'Z[S]}.
+Proof. by move=> phi; rewrite zchar_split => /andP[]. Qed.
+
+Lemma zchar_on S A : {subset 'Z[S, A] <= 'CF(G, A)}.
+Proof. by move=> phi /andP[]. Qed.
+
+Lemma zchar_onS A B S : A \subset B -> {subset 'Z[S, A] <= 'Z[S, B]}.
+Proof.
+move=> sAB phi; rewrite zchar_split (zchar_split _ B) => /andP[->].
+exact: cfun_onS.
+Qed.
+
+Lemma zchar_onG S : 'Z[S, G] =i 'Z[S].
+Proof. by move=> phi; rewrite zchar_split cfun_onG andbT. Qed.
+
+Lemma irr_vchar_on A : {subset 'Z[irr G, A] <= 'CF(G, A)}.
+Proof. exact: zchar_on. Qed.
+
+Lemma support_zchar S A phi : phi \in 'Z[S, A] -> support phi \subset A.
+Proof. by move/zchar_on; rewrite cfun_onE. Qed.
+
+Lemma mem_zchar_on S A phi :
+ phi \in 'CF(G, A) -> phi \in S -> phi \in 'Z[S, A].
+Proof.
+move=> Aphi /(@tnthP _ _ (in_tuple S))[i Dphi]; rewrite inE /= {}Aphi {phi}Dphi.
+apply/sumboolP; exists [ffun j => (j == i)%:Z].
+rewrite (bigD1 i) //= ffunE eqxx (tnth_nth 0) big1 ?addr0 // => j i'j.
+by rewrite ffunE (negPf i'j).
+Qed.
+
+(* A special lemma is needed because trivial fails to use the cfun_onT Hint. *)
+Lemma mem_zchar S phi : phi \in S -> phi \in 'Z[S].
+Proof. by move=> Sphi; rewrite mem_zchar_on ?cfun_onT. Qed.
+
+Lemma zchar_nth_expansion S A phi :
+ phi \in 'Z[S, A] ->
+ {z | forall i, z i \in Cint & phi = \sum_(i < size S) z i *: S`_i}.
+Proof.
+case/andP=> _ /sumboolP/sig_eqW[/= z ->].
+exists (intr \o z) => [i|]; first exact: Cint_int.
+by apply: eq_bigr => i _; rewrite scaler_int.
+Qed.
+
+Lemma zchar_tuple_expansion n (S : n.-tuple 'CF(G)) A phi :
+ phi \in 'Z[S, A] ->
+ {z | forall i, z i \in Cint & phi = \sum_(i < n) z i *: S`_i}.
+Proof. by move/zchar_nth_expansion; rewrite size_tuple. Qed.
+
+(* A pure seq version with the extra hypothesis of S's unicity. *)
+Lemma zchar_expansion S A phi : uniq S ->
+ phi \in 'Z[S, A] ->
+ {z | forall xi, z xi \in Cint & phi = \sum_(xi <- S) z xi *: xi}.
+Proof.
+move=> Suniq /zchar_nth_expansion[z Zz ->] /=.
+pose zS xi := oapp z 0 (insub (index xi S)).
+exists zS => [xi | ]; rewrite {}/zS; first by case: (insub _).
+rewrite (big_nth 0) big_mkord; apply: eq_bigr => i _; congr (_ *: _).
+by rewrite index_uniq // valK.
+Qed.
+
+Lemma zchar_span S A : {subset 'Z[S, A] <= <<S>>%VS}.
+Proof.
+move=> _ /zchar_nth_expansion[z Zz ->] /=.
+by apply: rpred_sum => i _; rewrite rpredZ // memv_span ?mem_nth.
+Qed.
+
+Lemma zchar_trans S1 S2 A B :
+ {subset S1 <= 'Z[S2, B]} -> {subset 'Z[S1, A] <= 'Z[S2, A]}.
+Proof.
+move=> sS12 phi; rewrite !(zchar_split _ A) andbC => /andP[->]; rewrite andbT.
+case/zchar_nth_expansion=> z Zz ->; apply: rpred_sum => i _.
+by rewrite scale_zchar // (@zcharW _ B) ?sS12 ?mem_nth.
+Qed.
+
+Lemma zchar_trans_on S1 S2 A :
+ {subset S1 <= 'Z[S2, A]} -> {subset 'Z[S1] <= 'Z[S2, A]}.
+Proof.
+move=> sS12 _ /zchar_nth_expansion[z Zz ->]; apply: rpred_sum => i _.
+by rewrite scale_zchar // sS12 ?mem_nth.
+Qed.
+
+Lemma zchar_sub_irr S A :
+ {subset S <= 'Z[irr G]} -> {subset 'Z[S, A] <= 'Z[irr G, A]}.
+Proof. exact: zchar_trans. Qed.
+
+Lemma zchar_subset S1 S2 A :
+ {subset S1 <= S2} -> {subset 'Z[S1, A] <= 'Z[S2, A]}.
+Proof.
+move=> sS12; apply: zchar_trans setT _ => // f /sS12 S2f.
+by rewrite mem_zchar.
+Qed.
+
+Lemma zchar_subseq S1 S2 A :
+ subseq S1 S2 -> {subset 'Z[S1, A] <= 'Z[S2, A]}.
+Proof. move=> sS12; exact: zchar_subset (mem_subseq sS12). Qed.
+
+Lemma zchar_filter S A (p : pred 'CF(G)) :
+ {subset 'Z[filter p S, A] <= 'Z[S, A]}.
+Proof. by apply: zchar_subset=> f; rewrite mem_filter => /andP[]. Qed.
+
+End Zchar.
+
+Section VChar.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Implicit Types (A B : {set gT}) (phi chi : 'CF(G)) (S : seq 'CF(G)).
+
+Lemma char_vchar chi : chi \is a character -> chi \in 'Z[irr G].
+Proof.
+case/char_sum_irr=> r ->; apply: rpred_sum => i _.
+by rewrite mem_zchar ?mem_tnth.
+Qed.
+
+Lemma irr_vchar i : 'chi[G]_i \in 'Z[irr G].
+Proof. exact/char_vchar/irr_char. Qed.
+
+Lemma cfun1_vchar : 1 \in 'Z[irr G]. Proof. by rewrite -irr0 irr_vchar. Qed.
+
+Lemma vcharP phi :
+ reflect (exists2 chi1, chi1 \is a character
+ & exists2 chi2, chi2 \is a character & phi = chi1 - chi2)
+ (phi \in 'Z[irr G]).
+Proof.
+apply: (iffP idP) => [| [a Na [b Nb ->]]]; last by rewrite rpredB ?char_vchar.
+case/zchar_tuple_expansion=> z Zz ->; rewrite (bigID (fun i => 0 <= z i)) /=.
+set chi1 := \sum_(i | _) _; set nchi2 := \sum_(i | _) _.
+exists chi1; last exists (- nchi2); last by rewrite opprK.
+ apply: rpred_sum => i zi_ge0; rewrite -tnth_nth rpredZ_Cnat ?irr_char //.
+ by rewrite CnatEint Zz.
+rewrite -sumrN rpred_sum // => i zi_lt0; rewrite -scaleNr -tnth_nth.
+rewrite rpredZ_Cnat ?irr_char // CnatEint rpredN Zz oppr_ge0 ltrW //.
+by rewrite real_ltrNge ?Creal_Cint.
+Qed.
+
+Lemma Aint_vchar phi x : phi \in 'Z[irr G] -> phi x \in Aint.
+Proof.
+case/vcharP=> [chi1 Nchi1 [chi2 Nchi2 ->]].
+by rewrite !cfunE rpredB ?Aint_char.
+Qed.
+
+Lemma Cint_vchar1 phi : phi \in 'Z[irr G] -> phi 1%g \in Cint.
+Proof.
+case/vcharP=> phi1 Nphi1 [phi2 Nphi2 ->].
+by rewrite !cfunE rpredB // rpred_Cnat ?Cnat_char1.
+Qed.
+
+Lemma Cint_cfdot_vchar_irr i phi : phi \in 'Z[irr G] -> '[phi, 'chi_i] \in Cint.
+Proof.
+case/vcharP=> chi1 Nchi1 [chi2 Nchi2 ->].
+by rewrite cfdotBl rpredB // rpred_Cnat ?Cnat_cfdot_char_irr.
+Qed.
+
+Lemma cfdot_vchar_r phi psi :
+ psi \in 'Z[irr G] -> '[phi, psi] = \sum_i '[phi, 'chi_i] * '[psi, 'chi_i].
+Proof.
+move=> Zpsi; rewrite cfdot_sum_irr; apply: eq_bigr => i _; congr (_ * _).
+by rewrite aut_Cint ?Cint_cfdot_vchar_irr.
+Qed.
+
+Lemma Cint_cfdot_vchar : {in 'Z[irr G] &, forall phi psi, '[phi, psi] \in Cint}.
+Proof.
+move=> phi psi Zphi Zpsi; rewrite /= cfdot_vchar_r // rpred_sum // => k _.
+by rewrite rpredM ?Cint_cfdot_vchar_irr.
+Qed.
+
+Lemma Cnat_cfnorm_vchar : {in 'Z[irr G], forall phi, '[phi] \in Cnat}.
+Proof.
+by move=> phi Zphi; rewrite /= CnatEint cfnorm_ge0 Cint_cfdot_vchar.
+Qed.
+
+Fact vchar_mulr_closed : mulr_closed 'Z[irr G].
+Proof.
+split; first exact: cfun1_vchar.
+move=> _ _ /vcharP[xi1 Nxi1 [xi2 Nxi2 ->]] /vcharP[xi3 Nxi3 [xi4 Nxi4 ->]].
+by rewrite mulrBl !mulrBr !(rpredB, rpredD) // char_vchar ?rpredM.
+Qed.
+Canonical vchar_mulrPred := MulrPred vchar_mulr_closed.
+Canonical vchar_smulrPred := SmulrPred vchar_mulr_closed.
+Canonical vchar_semiringPred := SemiringPred vchar_mulr_closed.
+Canonical vchar_subringPred := SubringPred vchar_mulr_closed.
+
+Lemma mul_vchar A :
+ {in 'Z[irr G, A] &, forall phi psi, phi * psi \in 'Z[irr G, A]}.
+Proof.
+move=> phi psi; rewrite zchar_split => /andP[Zphi Aphi] /zcharW Zpsi.
+rewrite zchar_split rpredM //; apply/cfun_onP=> x A'x.
+by rewrite cfunE (cfun_onP Aphi) ?mul0r.
+Qed.
+
+Section CfdotPairwiseOrthogonal.
+
+Variables (M : {group gT}) (S : seq 'CF(G)) (nu : 'CF(G) -> 'CF(M)).
+Hypotheses (Inu : {in 'Z[S] &, isometry nu}) (oSS : pairwise_orthogonal S).
+
+Let freeS := orthogonal_free oSS.
+Let uniqS : uniq S := free_uniq freeS.
+Let Z_S : {subset S <= 'Z[S]}. Proof. by move=> phi; exact: mem_zchar. Qed.
+Let notS0 : 0 \notin S. Proof. by case/andP: oSS. Qed.
+Let dotSS := proj2 (pairwise_orthogonalP oSS).
+
+Lemma map_pairwise_orthogonal : pairwise_orthogonal (map nu S).
+Proof.
+have inj_nu: {in S &, injective nu}.
+ move=> phi psi Sphi Spsi /= eq_nu; apply: contraNeq (memPn notS0 _ Sphi).
+ by rewrite -cfnorm_eq0 -Inu ?Z_S // {2}eq_nu Inu ?Z_S // => /dotSS->.
+have notSnu0: 0 \notin map nu S.
+ apply: contra notS0 => /mapP[phi Sphi /esym/eqP].
+ by rewrite -cfnorm_eq0 Inu ?Z_S // cfnorm_eq0 => /eqP <-.
+apply/pairwise_orthogonalP; split; first by rewrite /= notSnu0 map_inj_in_uniq.
+move=>_ _ /mapP[phi Sphi ->] /mapP[psi Spsi ->].
+by rewrite (inj_in_eq inj_nu) // Inu ?Z_S //; exact: dotSS.
+Qed.
+
+Lemma cfproj_sum_orthogonal P z phi :
+ phi \in S ->
+ '[\sum_(xi <- S | P xi) z xi *: nu xi, nu phi]
+ = if P phi then z phi * '[phi] else 0.
+Proof.
+move=> Sphi; have defS := perm_to_rem Sphi.
+rewrite cfdot_suml (eq_big_perm _ defS) big_cons /= cfdotZl Inu ?Z_S //.
+rewrite big1_seq ?addr0 // => xi; rewrite mem_rem_uniq ?inE //.
+by case/and3P=> _ neq_xi Sxi; rewrite cfdotZl Inu ?Z_S // dotSS ?mulr0.
+Qed.
+
+Lemma cfdot_sum_orthogonal z1 z2 :
+ '[\sum_(xi <- S) z1 xi *: nu xi, \sum_(xi <- S) z2 xi *: nu xi]
+ = \sum_(xi <- S) z1 xi * (z2 xi)^* * '[xi].
+Proof.
+rewrite cfdot_sumr; apply: eq_big_seq => phi Sphi.
+by rewrite cfdotZr cfproj_sum_orthogonal // mulrCA mulrA.
+Qed.
+
+Lemma cfnorm_sum_orthogonal z :
+ '[\sum_(xi <- S) z xi *: nu xi] = \sum_(xi <- S) `|z xi| ^+ 2 * '[xi].
+Proof.
+by rewrite cfdot_sum_orthogonal; apply: eq_bigr => xi _; rewrite normCK.
+Qed.
+
+Lemma cfnorm_orthogonal : '[\sum_(xi <- S) nu xi] = \sum_(xi <- S) '[xi].
+Proof.
+rewrite -(eq_bigr _ (fun _ _ => scale1r _)) cfnorm_sum_orthogonal.
+by apply: eq_bigr => xi; rewrite normCK conjC1 !mul1r.
+Qed.
+
+End CfdotPairwiseOrthogonal.
+
+Lemma orthogonal_span S phi :
+ pairwise_orthogonal S -> phi \in <<S>>%VS ->
+ {z | z = fun xi => '[phi, xi] / '[xi] & phi = \sum_(xi <- S) z xi *: xi}.
+Proof.
+move=> oSS /free_span[|c -> _]; first exact: orthogonal_free.
+set z := fun _ => _ : algC; exists z => //; apply: eq_big_seq => u Su.
+rewrite /z cfproj_sum_orthogonal // mulfK // cfnorm_eq0.
+by rewrite (memPn _ u Su); case/andP: oSS.
+Qed.
+
+Section CfDotOrthonormal.
+
+Variables (M : {group gT}) (S : seq 'CF(G)) (nu : 'CF(G) -> 'CF(M)).
+Hypotheses (Inu : {in 'Z[S] &, isometry nu}) (onS : orthonormal S).
+Let oSS := orthonormal_orthogonal onS.
+Let freeS := orthogonal_free oSS.
+Let nS1 : {in S, forall phi, '[phi] = 1}.
+Proof. by move=> phi Sphi; case/orthonormalP: onS => _ -> //; rewrite eqxx. Qed.
+
+Lemma map_orthonormal : orthonormal (map nu S).
+Proof.
+rewrite !orthonormalE map_pairwise_orthogonal // andbT.
+by apply/allP=> _ /mapP[xi Sxi ->]; rewrite /= Inu ?nS1 // mem_zchar.
+Qed.
+
+Lemma cfproj_sum_orthonormal z phi :
+ phi \in S -> '[\sum_(xi <- S) z xi *: nu xi, nu phi] = z phi.
+Proof. by move=> Sphi; rewrite cfproj_sum_orthogonal // nS1 // mulr1. Qed.
+
+Lemma cfdot_sum_orthonormal z1 z2 :
+ '[\sum_(xi <- S) z1 xi *: xi, \sum_(xi <- S) z2 xi *: xi]
+ = \sum_(xi <- S) z1 xi * (z2 xi)^*.
+Proof.
+rewrite cfdot_sum_orthogonal //; apply: eq_big_seq => phi /nS1->.
+by rewrite mulr1.
+Qed.
+
+Lemma cfnorm_sum_orthonormal z :
+ '[\sum_(xi <- S) z xi *: nu xi] = \sum_(xi <- S) `|z xi| ^+ 2.
+Proof.
+rewrite cfnorm_sum_orthogonal //.
+by apply: eq_big_seq => xi /nS1->; rewrite mulr1.
+Qed.
+
+Lemma cfnorm_map_orthonormal : '[\sum_(xi <- S) nu xi] = (size S)%:R.
+Proof.
+by rewrite cfnorm_orthogonal // (eq_big_seq _ nS1) big_tnth sumr_const card_ord.
+Qed.
+
+Lemma orthonormal_span phi :
+ phi \in <<S>>%VS ->
+ {z | z = fun xi => '[phi, xi] & phi = \sum_(xi <- S) z xi *: xi}.
+Proof.
+case/orthogonal_span=> // _ -> {2}->; set z := fun _ => _ : algC.
+by exists z => //; apply: eq_big_seq => xi /nS1->; rewrite divr1.
+Qed.
+
+End CfDotOrthonormal.
+
+Lemma cfnorm_orthonormal S :
+ orthonormal S -> '[\sum_(xi <- S) xi] = (size S)%:R.
+Proof. exact: cfnorm_map_orthonormal. Qed.
+
+Lemma zchar_orthonormalP S :
+ {subset S <= 'Z[irr G]} ->
+ reflect (exists I : {set Iirr G}, exists b : Iirr G -> bool,
+ perm_eq S [seq (-1) ^+ b i *: 'chi_i | i in I])
+ (orthonormal S).
+Proof.
+move=> vcS; apply: (equivP orthonormalP).
+split=> [[uniqS oSS] | [I [b defS]]]; last first.
+ split=> [|xi1 xi2]; rewrite ?(perm_eq_mem defS).
+ rewrite (perm_eq_uniq defS) map_inj_uniq ?enum_uniq // => i j /eqP.
+ by rewrite eq_signed_irr => /andP[_ /eqP].
+ case/mapP=> [i _ ->] /mapP[j _ ->]; rewrite eq_signed_irr.
+ rewrite cfdotZl cfdotZr rmorph_sign mulrA cfdot_irr -signr_addb mulr_natr.
+ by rewrite mulrb andbC; case: eqP => //= ->; rewrite addbb eqxx.
+pose I := [set i | ('chi_i \in S) || (- 'chi_i \in S)].
+pose b i := - 'chi_i \in S; exists I, b.
+apply: uniq_perm_eq => // [|xi].
+ rewrite map_inj_uniq ?enum_uniq // => i j /eqP.
+ by rewrite eq_signed_irr => /andP[_ /eqP].
+apply/idP/mapP=> [Sxi | [i Ii ->{xi}]]; last first.
+ move: Ii; rewrite mem_enum inE orbC -/(b i).
+ by case b_i: (b i); rewrite (scale1r, scaleN1r).
+have: '[xi] = 1 by rewrite oSS ?eqxx.
+have vc_xi := vcS _ Sxi; rewrite cfdot_sum_irr.
+case/Cnat_sum_eq1 => [i _ | i [_ /eqP norm_xi_i xi_i'_0]].
+ by rewrite -normCK rpredX // Cnat_norm_Cint ?Cint_cfdot_vchar_irr.
+suffices def_xi: xi = (-1) ^+ b i *: 'chi_i.
+ exists i; rewrite // mem_enum inE -/(b i) orbC.
+ by case: (b i) def_xi Sxi => // ->; rewrite scale1r.
+move: Sxi; rewrite [xi]cfun_sum_cfdot (bigD1 i) //.
+rewrite big1 //= ?addr0 => [|j ne_ji]; last first.
+ apply/eqP; rewrite scaler_eq0 -normr_eq0 -[_ == 0](expf_eq0 _ 2) normCK.
+ by rewrite xi_i'_0 ?eqxx.
+have:= norm_xi_i; rewrite (aut_Cint _ (Cint_cfdot_vchar_irr _ _)) //.
+rewrite -subr_eq0 subr_sqr_1 mulf_eq0 subr_eq0 addr_eq0 /b scaler_sign.
+case/pred2P=> ->; last by rewrite scaleN1r => ->.
+rewrite scale1r => Sxi; case: ifP => // SNxi.
+have:= oSS _ _ Sxi SNxi; rewrite cfdotNr cfdot_irr eqxx; case: eqP => // _.
+by move/eqP; rewrite oppr_eq0 oner_eq0.
+Qed.
+
+Lemma vchar_norm1P phi :
+ phi \in 'Z[irr G] -> '[phi] = 1 ->
+ exists b : bool, exists i : Iirr G, phi = (-1) ^+ b *: 'chi_i.
+Proof.
+move=> Zphi phiN1.
+have: orthonormal phi by rewrite /orthonormal/= phiN1 eqxx.
+case/zchar_orthonormalP=> [xi /predU1P[->|] // | I [b def_phi]].
+have: phi \in (phi : seq _) := mem_head _ _.
+by rewrite (perm_eq_mem def_phi) => /mapP[i _ ->]; exists (b i), i.
+Qed.
+
+Lemma zchar_small_norm phi n :
+ phi \in 'Z[irr G] -> '[phi] = n%:R -> (n < 4)%N ->
+ {S : n.-tuple 'CF(G) |
+ [/\ orthonormal S, {subset S <= 'Z[irr G]} & phi = \sum_(xi <- S) xi]}.
+Proof.
+move=> Zphi def_n lt_n_4.
+pose S := [seq '[phi, 'chi_i] *: 'chi_i | i in irr_constt phi].
+have def_phi: phi = \sum_(xi <- S) xi.
+ rewrite big_map /= big_filter big_mkcond {1}[phi]cfun_sum_cfdot.
+ by apply: eq_bigr => i _; rewrite if_neg; case: eqP => // ->; rewrite scale0r.
+have orthS: orthonormal S.
+ apply/orthonormalP; split=> [|_ _ /mapP[i phi_i ->] /mapP[j _ ->]].
+ rewrite map_inj_in_uniq ?enum_uniq // => i j; rewrite mem_enum => phi_i _.
+ by move/eqP; rewrite eq_scaled_irr (negbTE phi_i) => /andP[_ /= /eqP].
+ rewrite eq_scaled_irr cfdotZl cfdotZr cfdot_irr mulrA mulr_natr mulrb.
+ rewrite mem_enum in phi_i; rewrite (negbTE phi_i) andbC; case: eqP => // <-.
+ have /CnatP[m def_m] := Cnat_norm_Cint (Cint_cfdot_vchar_irr i Zphi).
+ apply/eqP; rewrite eqxx /= -normCK def_m -natrX eqr_nat eqn_leq lt0n.
+ rewrite expn_eq0 andbT -eqC_nat -def_m normr_eq0 [~~ _]phi_i andbT.
+ rewrite (leq_exp2r _ 1) // -ltnS -(@ltn_exp2r _ _ 2) //.
+ apply: leq_ltn_trans lt_n_4; rewrite -leC_nat -def_n natrX.
+ rewrite cfdot_sum_irr (bigD1 i) //= -normCK def_m addrC -subr_ge0 addrK.
+ by rewrite sumr_ge0 // => ? _; exact: mul_conjC_ge0.
+have <-: size S = n.
+ by apply/eqP; rewrite -eqC_nat -def_n def_phi cfnorm_orthonormal.
+exists (in_tuple S); split=> // _ /mapP[i _ ->].
+by rewrite scale_zchar ?irr_vchar // Cint_cfdot_vchar_irr.
+Qed.
+
+Lemma vchar_norm2 phi :
+ phi \in 'Z[irr G, G^#] -> '[phi] = 2%:R ->
+ exists i, exists2 j, j != i & phi = 'chi_i - 'chi_j.
+Proof.
+rewrite zchar_split cfunD1E => /andP[Zphi phi1_0].
+case/zchar_small_norm => // [[[|chi [|xi [|?]]] //= S2]].
+case=> /andP[/and3P[Nchi Nxi _] /= ochi] /allP/and3P[Zchi Zxi _].
+rewrite big_cons big_seq1 => def_phi.
+have [b [i def_chi]] := vchar_norm1P Zchi (eqP Nchi).
+have [c [j def_xi]] := vchar_norm1P Zxi (eqP Nxi).
+have neq_ji: j != i.
+ apply: contraTneq ochi; rewrite !andbT def_chi def_xi => ->.
+ rewrite cfdotZl cfdotZr rmorph_sign cfnorm_irr mulr1 -signr_addb.
+ by rewrite signr_eq0.
+have neq_bc: b != c.
+ apply: contraTneq phi1_0; rewrite def_phi def_chi def_xi => ->.
+ rewrite -scalerDr !cfunE mulf_eq0 signr_eq0 eqr_le ltr_geF //.
+ by rewrite ltr_paddl ?ltrW ?irr1_gt0.
+rewrite {}def_phi {}def_chi {}def_xi !scaler_sign.
+case: b c neq_bc => [|] [|] // _; last by exists i, j.
+by exists j, i; rewrite 1?eq_sym // addrC.
+Qed.
+
+End VChar.
+
+Section Isometries.
+
+Variables (gT : finGroupType) (L G : {group gT}) (S : seq 'CF(L)).
+Implicit Type nu : {additive 'CF(L) -> 'CF(G)}.
+
+Lemma Zisometry_of_cfnorm (tauS : seq 'CF(G)) :
+ pairwise_orthogonal S -> pairwise_orthogonal tauS ->
+ map cfnorm tauS = map cfnorm S -> {subset tauS <= 'Z[irr G]} ->
+ {tau : {linear 'CF(L) -> 'CF(G)} | map tau S = tauS
+ & {in 'Z[S], isometry tau, to 'Z[irr G]}}.
+Proof.
+move=> oSS oTT /isometry_of_cfnorm[||tau defT Itau] // Z_T; exists tau => //.
+split=> [|_ /zchar_nth_expansion[u Zu ->]].
+ by apply: sub_in2 Itau; apply: zchar_span.
+rewrite big_seq linear_sum rpred_sum // => xi Sxi.
+by rewrite linearZ scale_zchar ?Z_T // -defT map_f ?mem_nth.
+Qed.
+
+Lemma Zisometry_of_iso f :
+ pairwise_orthogonal S -> {in S, isometry f, to 'Z[irr G]} ->
+ {tau : {linear 'CF(L) -> 'CF(G)} | {in S, tau =1 f}
+ & {in 'Z[S], isometry tau, to 'Z[irr G]}}.
+Proof.
+move=> oS [If Zf]; have [/=/andP[S'0 uS] oSS] := pairwise_orthogonalP oS.
+have injf: {in S &, injective f}.
+ move=> xi1 xi2 Sxi1 Sxi2 /=/(congr1 (cfdot (f xi1)))/eqP; rewrite !If //.
+ by apply: contraTeq => /oSS-> //; rewrite cfnorm_eq0 (memPn S'0).
+have{injf} oSf: pairwise_orthogonal (map f S).
+ apply/pairwise_orthogonalP; split=> /=.
+ rewrite map_inj_in_uniq // uS (contra _ S'0) // => /mapP[chi Schi /eqP].
+ by rewrite eq_sym -cfnorm_eq0 If // cfnorm_eq0 => /eqP <-.
+ move=> _ _ /mapP[xi1 Xxi1 ->] /mapP[xi2 Xxi2 ->].
+ by rewrite If ?(inj_in_eq injf) // => /oSS->.
+have{If} nSf: map cfnorm (map f S) = map cfnorm S.
+ by rewrite -map_comp; apply/eq_in_map=> xi Sxi; rewrite /= If.
+have{Zf} ZSf: {subset map f S <= 'Z[irr G]} by move=> _ /mapP[xi /Zf Zfxi ->].
+by have [tau /eq_in_map] := Zisometry_of_cfnorm oS oSf nSf ZSf; exists tau.
+Qed.
+
+Lemma Zisometry_inj A nu :
+ {in 'Z[S, A] &, isometry nu} -> {in 'Z[S, A] &, injective nu}.
+Proof. by move/isometry_raddf_inj; apply; apply: rpredB. Qed.
+
+Lemma isometry_in_zchar nu : {in S &, isometry nu} -> {in 'Z[S] &, isometry nu}.
+Proof.
+move=> Inu _ _ /zchar_nth_expansion[u Zu ->] /zchar_nth_expansion[v Zv ->].
+rewrite !raddf_sum; apply: eq_bigr => j _ /=.
+rewrite !cfdot_suml; apply: eq_bigr => i _.
+by rewrite !raddfZ_Cint //= !cfdotZl !cfdotZr Inu ?mem_nth.
+Qed.
+
+End Isometries.
+
+Section AutVchar.
+
+Variables (u : {rmorphism algC -> algC}) (gT : finGroupType) (G : {group gT}).
+Local Notation "alpha ^u" := (cfAut u alpha).
+Implicit Type (S : seq 'CF(G)) (phi chi : 'CF(G)).
+
+Lemma cfAut_zchar S A psi :
+ cfAut_closed u S -> psi \in 'Z[S, A] -> psi^u \in 'Z[S, A].
+Proof.
+rewrite zchar_split => SuS /andP[/zchar_nth_expansion[z Zz Dpsi] Apsi].
+rewrite zchar_split cfAut_on {}Apsi {psi}Dpsi rmorph_sum rpred_sum //= => i _.
+by rewrite cfAutZ_Cint // scale_zchar // mem_zchar ?SuS ?mem_nth.
+Qed.
+
+Lemma cfAut_vchar A psi : psi \in 'Z[irr G, A] -> psi^u \in 'Z[irr G, A].
+Proof. by apply: cfAut_zchar; exact: irr_aut_closed. Qed.
+
+Lemma sub_aut_zchar S A psi :
+ {subset S <= 'Z[irr G]} -> psi \in 'Z[S, A] -> psi^u \in 'Z[S, A] ->
+ psi - psi^u \in 'Z[S, A^#].
+Proof.
+move=> Z_S Spsi Spsi_u; rewrite zcharD1 !cfunE subr_eq0 rpredB //=.
+by rewrite aut_Cint // Cint_vchar1 // (zchar_trans Z_S) ?(zcharW Spsi).
+Qed.
+
+Lemma conjC_vcharAut chi x : chi \in 'Z[irr G] -> (u (chi x))^* = u (chi x)^*.
+Proof.
+case/vcharP=> chi1 Nchi1 [chi2 Nchi2 ->].
+by rewrite !cfunE !rmorphB !conjC_charAut.
+Qed.
+
+Lemma cfdot_aut_vchar phi chi :
+ chi \in 'Z[irr G] -> '[phi^u , chi^u] = u '[phi, chi].
+Proof.
+case/vcharP=> chi1 Nchi1 [chi2 Nchi2 ->].
+by rewrite !raddfB /= !cfdot_aut_char.
+Qed.
+
+Lemma vchar_aut A chi : (chi^u \in 'Z[irr G, A]) = (chi \in 'Z[irr G, A]).
+Proof.
+rewrite !(zchar_split _ A) cfAut_on; congr (_ && _).
+apply/idP/idP=> [Zuchi|]; last exact: cfAut_vchar.
+rewrite [chi]cfun_sum_cfdot rpred_sum // => i _.
+rewrite scale_zchar ?irr_vchar //.
+by rewrite -(Cint_aut u) -cfdot_aut_irr -aut_IirrE Cint_cfdot_vchar_irr.
+Qed.
+
+End AutVchar.
+
+Definition cfConjC_vchar := cfAut_vchar conjC.
+
+Section MoreVchar.
+
+Variables (gT : finGroupType) (G H : {group gT}).
+
+Lemma cfRes_vchar phi : phi \in 'Z[irr G] -> 'Res[H] phi \in 'Z[irr H].
+Proof.
+case/vcharP=> xi1 Nx1 [xi2 Nxi2 ->].
+by rewrite raddfB rpredB ?char_vchar ?cfRes_char.
+Qed.
+
+Lemma cfRes_vchar_on A phi :
+ H \subset G -> phi \in 'Z[irr G, A] -> 'Res[H] phi \in 'Z[irr H, A].
+Proof.
+rewrite zchar_split => sHG /andP[Zphi Aphi]; rewrite zchar_split cfRes_vchar //.
+apply/cfun_onP=> x /(cfun_onP Aphi); rewrite !cfunElock !genGid sHG => ->.
+exact: mul0rn.
+Qed.
+
+Lemma cfInd_vchar phi : phi \in 'Z[irr H] -> 'Ind[G] phi \in 'Z[irr G].
+Proof.
+move=> /vcharP[xi1 Nx1 [xi2 Nxi2 ->]].
+by rewrite raddfB rpredB ?char_vchar ?cfInd_char.
+Qed.
+
+Lemma sub_conjC_vchar A phi :
+ phi \in 'Z[irr G, A] -> phi - (phi^*)%CF \in 'Z[irr G, A^#].
+Proof.
+move=> Zphi; rewrite sub_aut_zchar ?cfAut_zchar // => _ /irrP[i ->].
+ exact: irr_vchar.
+exact: cfConjC_irr.
+Qed.
+
+Lemma Frobenius_kernel_exists :
+ [Frobenius G with complement H] -> {K : {group gT} | [Frobenius G = K ><| H]}.
+Proof.
+move=> frobG; have [_ ntiHG] := andP frobG.
+have [[_ sHG regGH][_ tiHG /eqP defNH]] := (normedTI_memJ_P ntiHG, and3P ntiHG).
+suffices /sigW[K defG]: exists K, gval K ><| H == G by exists K; apply/andP.
+pose K1 := G :\: cover (H^# :^: G).
+have oK1: #|K1| = #|G : H|.
+ rewrite cardsD (setIidPr _); last first.
+ rewrite cover_imset; apply/bigcupsP=> x Gx.
+ by rewrite sub_conjg conjGid ?groupV // (subset_trans (subsetDl _ _)).
+ rewrite (cover_partition (partition_normedTI ntiHG)) -(Lagrange sHG).
+ by rewrite (card_support_normedTI ntiHG) (cardsD1 1%g) group1 mulSn addnK.
+suffices extG i: {j | {in H, 'chi[G]_j =1 'chi[H]_i} & K1 \subset cfker 'chi_j}.
+ pose K := [group of \bigcap_i cfker 'chi_(s2val (extG i))].
+ have nKH: H \subset 'N(K).
+ by apply/norms_bigcap/bigcapsP=> i _; apply: subset_trans (cfker_norm _).
+ have tiKH: K :&: H = 1%g.
+ apply/trivgP; rewrite -(TI_cfker_irr H) /= setIC; apply/bigcapsP=> i _.
+ apply/subsetP=> x /setIP[Hx /bigcapP/(_ i isT)/=]; rewrite !cfkerEirr !inE.
+ by case: (extG i) => /= j def_j _; rewrite !def_j.
+ exists K; rewrite sdprodE // eqEcard TI_cardMg // mul_subG //=; last first.
+ by rewrite (bigcap_min (0 : Iirr H)) ?cfker_sub.
+ rewrite -(Lagrange sHG) mulnC leq_pmul2r // -oK1 subset_leq_card //.
+ by apply/bigcapsP=> i _; case: (extG i).
+case i0: (i == 0).
+ exists 0 => [x Hx|]; last by rewrite irr0 cfker_cfun1 subsetDl.
+ by rewrite (eqP i0) !irr0 !cfun1E // (subsetP sHG) ?Hx.
+have ochi1: '['chi_i, 1] = 0 by rewrite -irr0 cfdot_irr i0.
+pose a := 'chi_i 1%g; have Za: a \in Cint by rewrite CintE Cnat_irr1.
+pose theta := 'chi_i - a%:A; pose phi := 'Ind[G] theta + a%:A.
+have /cfun_onP theta0: theta \in 'CF(H, H^#).
+ by rewrite cfunD1E !cfunE cfun11 mulr1 subrr.
+have RItheta: 'Res ('Ind[G] theta) = theta.
+ apply/cfun_inP=> x Hx; rewrite cfResE ?cfIndE // (big_setID H) /= addrC.
+ apply: canLR (mulKf (neq0CG H)) _; rewrite (setIidPr sHG) mulr_natl.
+ rewrite big1 ?add0r => [|y /setDP[/regGH tiHy H'y]]; last first.
+ have [-> | ntx] := eqVneq x 1%g; first by rewrite conj1g theta0 ?inE ?eqxx.
+ by rewrite theta0 ?tiHy // !inE ntx.
+ by rewrite -sumr_const; apply: eq_bigr => y Hy; rewrite cfunJ.
+have ophi1: '[phi, 1] = 0.
+ rewrite cfdotDl -cfdot_Res_r cfRes_cfun1 // cfdotBl !cfdotZl !cfnorm1.
+ by rewrite ochi1 add0r addNr.
+have{ochi1} n1phi: '[phi] = 1.
+ have: '[phi - a%:A] = '[theta] by rewrite addrK -cfdot_Res_l RItheta.
+ rewrite !cfnormBd ?cfnormZ ?cfdotZr ?ophi1 ?ochi1 ?mulr0 //.
+ by rewrite !cfnorm1 cfnorm_irr => /addIr.
+have Zphi: phi \in 'Z[irr G].
+ by rewrite rpredD ?cfInd_vchar ?rpredB ?irr_vchar // scale_zchar ?rpred1.
+have def_phi: {in H, phi =1 'chi_i}.
+ move=> x Hx /=; rewrite !cfunE -[_ x](cfResE _ sHG) ?RItheta //.
+ by rewrite !cfunE !cfun1E ?(subsetP sHG) ?Hx ?subrK.
+have [j def_chi_j]: {j | 'chi_j = phi}.
+ apply/sig_eqW; have [[] [j]] := vchar_norm1P Zphi n1phi; last first.
+ by rewrite scale1r; exists j.
+ move/cfunP/(_ 1%g)/eqP; rewrite scaleN1r def_phi // cfunE -addr_eq0 eqr_le.
+ by rewrite ltr_geF // ltr_paddl ?ltrW ?irr1_gt0.
+exists j; rewrite ?cfkerEirr def_chi_j //; apply/subsetP => x /setDP[Gx notHx].
+rewrite inE cfunE def_phi // cfunE -/a cfun1E // Gx mulr1 cfIndE //.
+rewrite big1 ?mulr0 ?add0r // => y Gy; apply/theta0/(contra _ notHx) => Hxy.
+by rewrite -(conjgK y x) cover_imset -class_supportEr mem_imset2 ?groupV.
+Qed.
+
+End MoreVchar.
+
+Definition dirr (gT : finGroupType) (B : {set gT}) : pred_class :=
+ [pred f : 'CF(B) | (f \in irr B) || (- f \in irr B)].
+Implicit Arguments dirr [[gT]].
+
+Section Norm1vchar.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Fact dirr_key : pred_key (dirr G). Proof. by []. Qed.
+Canonical dirr_keyed := KeyedPred dirr_key.
+
+Fact dirr_oppr_closed : oppr_closed (dirr G).
+Proof. by move=> xi; rewrite !inE opprK orbC. Qed.
+Canonical dirr_opprPred := OpprPred dirr_oppr_closed.
+
+Lemma dirr_opp v : (- v \in dirr G) = (v \in dirr G). Proof. exact: rpredN. Qed.
+Lemma dirr_sign n v : ((-1)^+ n *: v \in dirr G) = (v \in dirr G).
+Proof. exact: rpredZsign. Qed.
+
+Lemma irr_dirr i : 'chi_i \in dirr G.
+Proof. by rewrite !inE mem_irr. Qed.
+
+Lemma dirrP f :
+ reflect (exists b : bool, exists i, f = (-1) ^+ b *: 'chi_i) (f \in dirr G).
+Proof.
+apply: (iffP idP) => [| [b [i ->]]]; last by rewrite dirr_sign irr_dirr.
+case/orP=> /irrP[i Hf]; first by exists false, i; rewrite scale1r.
+by exists true, i; rewrite scaleN1r -Hf opprK.
+Qed.
+
+(* This should perhaps be the definition of dirr. *)
+Lemma dirrE phi : phi \in dirr G = (phi \in 'Z[irr G]) && ('[phi] == 1).
+Proof.
+apply/dirrP/andP=> [[b [i ->]] | [Zphi /eqP/vchar_norm1P]]; last exact.
+by rewrite rpredZsign irr_vchar cfnorm_sign cfnorm_irr.
+Qed.
+
+Lemma cfdot_dirr f g : f \in dirr G -> g \in dirr G ->
+ '[f, g] = (if f == - g then -1 else (f == g)%:R).
+Proof.
+case/dirrP=> [b1 [i1 ->]] /dirrP[b2 [i2 ->]].
+rewrite cfdotZl cfdotZr rmorph_sign mulrA -signr_addb cfdot_irr.
+rewrite -scaleNr -signrN !eq_scaled_irr signr_eq0 !(inj_eq (@signr_inj _)) /=.
+by rewrite -!negb_add addbN mulr_sign -mulNrn mulrb; case: ifP.
+Qed.
+
+Lemma dirr_norm1 phi : phi \in 'Z[irr G] -> '[phi] = 1 -> phi \in dirr G.
+Proof. by rewrite dirrE => -> -> /=. Qed.
+
+Lemma dirr_aut u phi : (cfAut u phi \in dirr G) = (phi \in dirr G).
+Proof.
+rewrite !dirrE vchar_aut; apply: andb_id2l => /cfdot_aut_vchar->.
+exact: fmorph_eq1.
+Qed.
+
+Definition dIirr (B : {set gT}) := (bool * (Iirr B))%type.
+
+Definition dirr1 (B : {set gT}) : dIirr B := (false, 0).
+
+Definition ndirr (B : {set gT}) (i : dIirr B) : dIirr B :=
+ (~~ i.1, i.2).
+
+Lemma ndirr_diff (i : dIirr G) : ndirr i != i.
+Proof. by case: i => [] [|] i. Qed.
+
+Lemma ndirrK : involutive (@ndirr G).
+Proof. by move=> [b i]; rewrite /ndirr /= negbK. Qed.
+
+Lemma ndirr_inj : injective (@ndirr G).
+Proof. exact: (inv_inj ndirrK). Qed.
+
+Definition dchi (B : {set gT}) (i : dIirr B) : 'CF(B) :=
+ (-1)^+ i.1 *: 'chi_i.2.
+
+Lemma dchi1 : dchi (dirr1 G) = 1.
+Proof. by rewrite /dchi scale1r irr0. Qed.
+
+Lemma dirr_dchi i : dchi i \in dirr G.
+Proof. by apply/dirrP; exists i.1; exists i.2. Qed.
+
+Lemma dIrrP (phi : 'CF(G)) :
+ reflect (exists i , phi = dchi i) (phi \in dirr G).
+Proof.
+by apply: (iffP idP)=> [/dirrP [b [i ->]]| [i ->]];
+ [exists (b, i) | exact: dirr_dchi].
+Qed.
+
+Lemma dchi_ndirrE (i : dIirr G) : dchi (ndirr i) = - dchi i.
+Proof. by case: i => [b i]; rewrite /ndirr /dchi signrN scaleNr. Qed.
+
+Lemma cfdot_dchi (i j : dIirr G) :
+ '[dchi i, dchi j] = (i == j)%:R - (i == ndirr j)%:R.
+Proof.
+case: i => bi i; case: j => bj j; rewrite cfdot_dirr ?dirr_dchi // !xpair_eqE.
+rewrite -dchi_ndirrE !eq_scaled_irr signr_eq0 !(inj_eq (@signr_inj _)) /=.
+by rewrite -!negb_add addbN negbK; case: andP => [[->]|]; rewrite ?subr0 ?add0r.
+Qed.
+
+Lemma dchi_vchar i : dchi i \in 'Z[irr G].
+Proof. by case: i => b i; rewrite rpredZsign irr_vchar. Qed.
+
+Lemma cfnorm_dchi (i : dIirr G) : '[dchi i] = 1.
+Proof. by case: i => b i; rewrite cfnorm_sign cfnorm_irr. Qed.
+
+Lemma dirr_inj : injective (@dchi G).
+Proof.
+case=> b1 i1 [b2 i2] /eqP; rewrite eq_scaled_irr (inj_eq (@signr_inj _)) /=.
+by rewrite signr_eq0 -xpair_eqE => /eqP.
+Qed.
+
+Definition dirr_dIirr (B : {set gT}) J (f : J -> 'CF(B)) j : dIirr B :=
+ odflt (dirr1 B) [pick i | dchi i == f j].
+
+Lemma dirr_dIirrPE J (f : J -> 'CF(G)) (P : pred J) :
+ (forall j, P j -> f j \in dirr G) ->
+ forall j, P j -> dchi (dirr_dIirr f j) = f j.
+Proof.
+rewrite /dirr_dIirr => dirrGf j Pj; case: pickP => [i /eqP //|].
+by have /dIrrP[i-> /(_ i)/eqP] := dirrGf j Pj.
+Qed.
+
+Lemma dirr_dIirrE J (f : J -> 'CF(G)) :
+ (forall j, f j \in dirr G) -> forall j, dchi (dirr_dIirr f j) = f j.
+Proof. by move=> dirrGf j; exact: (@dirr_dIirrPE _ _ xpredT). Qed.
+
+Definition dirr_constt (B : {set gT}) (phi: 'CF(B)) : {set (dIirr B)} :=
+ [set i | 0 < '[phi, dchi i]].
+
+Lemma dirr_consttE (phi : 'CF(G)) (i : dIirr G) :
+ (i \in dirr_constt phi) = (0 < '[phi, dchi i]).
+Proof. by rewrite inE. Qed.
+
+Lemma Cnat_dirr (phi : 'CF(G)) i :
+ phi \in 'Z[irr G] -> i \in dirr_constt phi -> '[phi, dchi i] \in Cnat.
+Proof.
+move=> PiZ; rewrite CnatEint dirr_consttE andbC => /ltrW -> /=.
+by case: i => b i; rewrite cfdotZr rmorph_sign rpredMsign Cint_cfdot_vchar_irr.
+Qed.
+
+Lemma dirr_constt_oppr (i : dIirr G) (phi : 'CF(G)) :
+ (i \in dirr_constt (-phi)) = (ndirr i \in dirr_constt phi).
+Proof. by rewrite !dirr_consttE dchi_ndirrE cfdotNl cfdotNr. Qed.
+
+Lemma dirr_constt_oppI (phi: 'CF(G)) :
+ dirr_constt phi :&: dirr_constt (-phi) = set0.
+Proof.
+apply/setP=> i; rewrite inE !dirr_consttE cfdotNl inE.
+apply/idP=> /andP [L1 L2]; have := ltr_paddl (ltrW L1) L2.
+by rewrite subrr ltr_def eqxx.
+Qed.
+
+Lemma dirr_constt_oppl (phi: 'CF(G)) i :
+ i \in dirr_constt phi -> (ndirr i) \notin dirr_constt phi.
+Proof.
+rewrite !dirr_consttE dchi_ndirrE cfdotNr oppr_gt0.
+by move/ltrW=> /ler_gtF ->.
+Qed.
+
+Definition to_dirr (B : {set gT}) (phi : 'CF(B)) (i : Iirr B) : dIirr B :=
+ ('[phi, 'chi_i] < 0, i).
+
+Definition of_irr (B : {set gT}) (i : dIirr B) : Iirr B := i.2.
+
+Lemma irr_constt_to_dirr (phi: 'CF(G)) i : phi \in 'Z[irr G] ->
+ (i \in irr_constt phi) = (to_dirr phi i \in dirr_constt phi).
+Proof.
+move=> Zphi; rewrite irr_consttE dirr_consttE cfdotZr rmorph_sign /=.
+by rewrite -real_normrEsign ?normr_gt0 ?Creal_Cint // Cint_cfdot_vchar_irr.
+Qed.
+
+Lemma to_dirrK (phi: 'CF(G)) : cancel (to_dirr phi) (@of_irr G).
+Proof. by []. Qed.
+
+Lemma of_irrK (phi: 'CF(G)) :
+ {in dirr_constt phi, cancel (@of_irr G) (to_dirr phi)}.
+Proof.
+case=> b i; rewrite dirr_consttE cfdotZr rmorph_sign /= /to_dirr mulr_sign.
+by rewrite fun_if oppr_gt0; case: b => [|/ltrW/ler_gtF] ->.
+Qed.
+
+Lemma cfdot_todirrE (phi: 'CF(G)) i (phi_i := dchi (to_dirr phi i)) :
+ '[phi, phi_i] *: phi_i = '[phi, 'chi_i] *: 'chi_i.
+Proof. by rewrite cfdotZr rmorph_sign mulrC -scalerA signrZK. Qed.
+
+Lemma cfun_sum_dconstt (phi : 'CF(G)) :
+ phi \in 'Z[irr G] ->
+ phi = \sum_(i in dirr_constt phi) '[phi, dchi i] *: dchi i.
+Proof.
+(* GG -- rewrite pattern fails in trunk
+ move=> PiZ; rewrite [X in X = _]cfun_sum_constt. *)
+move=> PiZ; rewrite {1}[phi]cfun_sum_constt.
+rewrite (reindex (to_dirr phi))=> [/= |]; last first.
+ by exists (@of_irr _)=> //; exact: of_irrK .
+by apply: eq_big=> i; rewrite ?irr_constt_to_dirr // cfdot_todirrE.
+Qed.
+
+Lemma cnorm_dconstt (phi : 'CF(G)) :
+ phi \in 'Z[irr G] ->
+ '[phi] = \sum_(i in dirr_constt phi) '[phi, dchi i] ^+ 2.
+Proof.
+move=> PiZ; rewrite {1 2}(cfun_sum_dconstt PiZ).
+rewrite cfdot_suml; apply: eq_bigr=> i IiD.
+rewrite cfdot_sumr (bigD1 i) //= big1 ?addr0 => [|j /andP [JiD IdJ]].
+ rewrite cfdotZr cfdotZl cfdot_dchi eqxx eq_sym (negPf (ndirr_diff i)).
+ by rewrite subr0 mulr1 aut_Cnat ?Cnat_dirr.
+rewrite cfdotZr cfdotZl cfdot_dchi eq_sym (negPf IdJ) -natrB ?mulr0 //.
+by rewrite (negPf (contraNneq _ (dirr_constt_oppl JiD))) => // <-.
+Qed.
+
+Lemma dirr_small_norm (phi : 'CF(G)) n :
+ phi \in 'Z[irr G] -> '[phi] = n%:R -> (n < 4)%N ->
+ [/\ #|dirr_constt phi| = n, dirr_constt phi :&: dirr_constt (- phi) = set0 &
+ phi = \sum_(i in dirr_constt phi) dchi i].
+Proof.
+move=> PiZ Pln; rewrite ltnNge -leC_nat => Nl4.
+suffices Fd i: i \in dirr_constt phi -> '[phi, dchi i] = 1.
+ split; last 2 [by apply/setP=> u; rewrite !inE cfdotNl oppr_gt0 ltr_asym].
+ apply/eqP; rewrite -eqC_nat -sumr_const -Pln (cnorm_dconstt PiZ).
+ by apply/eqP/eq_bigr=> i Hi; rewrite Fd // expr1n.
+ rewrite {1}[phi]cfun_sum_dconstt //.
+ by apply: eq_bigr => i /Fd->; rewrite scale1r.
+move=> IiD; apply: contraNeq Nl4 => phi_i_neq1.
+rewrite -Pln cnorm_dconstt // (bigD1 i) ?ler_paddr ?sumr_ge0 //=.
+ by move=> j /andP[JiD _]; rewrite exprn_ge0 ?Cnat_ge0 ?Cnat_dirr.
+have /CnatP[m Dm] := Cnat_dirr PiZ IiD; rewrite Dm -natrX ler_nat (leq_sqr 2).
+by rewrite ltn_neqAle eq_sym -eqC_nat -ltC_nat -Dm phi_i_neq1 -dirr_consttE.
+Qed.
+
+Lemma cfdot_sum_dchi (phi1 phi2 : 'CF(G)) :
+ '[\sum_(i in dirr_constt phi1) dchi i,
+ \sum_(i in dirr_constt phi2) dchi i] =
+ #|dirr_constt phi1 :&: dirr_constt phi2|%:R -
+ #|dirr_constt phi1 :&: dirr_constt (- phi2)|%:R.
+Proof.
+rewrite addrC (big_setID (dirr_constt (- phi2))) /= cfdotDl; congr (_ + _).
+ rewrite cfdot_suml -sumr_const -sumrN; apply: eq_bigr => i /setIP[p1i p2i].
+ rewrite cfdot_sumr (bigD1 (ndirr i)) -?dirr_constt_oppr //= dchi_ndirrE.
+ rewrite cfdotNr cfnorm_dchi big1 ?addr0 // => j /andP[p2j i'j].
+ rewrite cfdot_dchi -(inv_eq ndirrK) [in rhs in - rhs]eq_sym (negPf i'j) subr0.
+ rewrite (negPf (contraTneq _ p2i)) // => ->.
+ by rewrite dirr_constt_oppr dirr_constt_oppl.
+rewrite cfdot_sumr (big_setID (dirr_constt phi1)) setIC /= addrC.
+rewrite big1 ?add0r => [|j /setDP[p2j p1'j]]; last first.
+ rewrite cfdot_suml big1 // => i /setDP[p1i p2'i].
+ rewrite cfdot_dchi (negPf (contraTneq _ p1i)) => [|-> //].
+ rewrite (negPf (contraNneq _ p2'i)) ?subrr // => ->.
+ by rewrite dirr_constt_oppr ndirrK.
+rewrite -sumr_const; apply: eq_bigr => i /setIP[p1i p2i]; rewrite cfdot_suml.
+rewrite (bigD1 i) /=; last by rewrite inE dirr_constt_oppr dirr_constt_oppl.
+rewrite cfnorm_dchi big1 ?addr0 // => j /andP[/setDP[p1j _] i'j].
+rewrite cfdot_dchi (negPf i'j) (negPf (contraTneq _ p1j)) ?subrr // => ->.
+exact: dirr_constt_oppl.
+Qed.
+
+Lemma cfdot_dirr_eq1 :
+ {in dirr G &, forall phi psi, ('[phi, psi] == 1) = (phi == psi)}.
+Proof.
+move=> _ _ /dirrP[b1 [i1 ->]] /dirrP[b2 [i2 ->]].
+rewrite eq_signed_irr cfdotZl cfdotZr rmorph_sign cfdot_irr mulrA -signr_addb.
+rewrite pmulrn -rmorphMsign (eqr_int _ _ 1) -negb_add.
+by case: (b1 (+) b2) (i1 == i2) => [] [].
+Qed.
+
+Lemma cfdot_add_dirr_eq1 :
+ {in dirr G & &, forall phi1 phi2 psi,
+ '[phi1 + phi2, psi] = 1 -> psi = phi1 \/ psi = phi2}.
+Proof.
+move=> _ _ _ /dirrP[b1 [i1 ->]] /dirrP[b2 [i2 ->]] /dirrP[c [j ->]] /eqP.
+rewrite cfdotDl !cfdotZl !cfdotZr !rmorph_sign !cfdot_irr !mulrA -!signr_addb.
+rewrite 2!{1}signrE !mulrBl !mul1r -!natrM addrCA -subr_eq0 -!addrA.
+rewrite -!opprD addrA subr_eq0 -mulrSr -!natrD eqr_nat => eq_phi_psi.
+apply/pred2P; rewrite /= !eq_signed_irr -!negb_add !(eq_sym j) !(addbC c).
+by case: (i1 == j) eq_phi_psi; case: (i2 == j); do 2!case: (_ (+) c).
+Qed.
+
+End Norm1vchar.