aboutsummaryrefslogtreecommitdiff
path: root/mathcomp/fingroup
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/fingroup
Initial commit
Diffstat (limited to 'mathcomp/fingroup')
-rw-r--r--mathcomp/fingroup/action.v2719
-rw-r--r--mathcomp/fingroup/all.v10
-rw-r--r--mathcomp/fingroup/automorphism.v489
-rw-r--r--mathcomp/fingroup/cyclic.v865
-rw-r--r--mathcomp/fingroup/fingroup.v3096
-rw-r--r--mathcomp/fingroup/gproduct.v1703
-rw-r--r--mathcomp/fingroup/morphism.v1539
-rw-r--r--mathcomp/fingroup/perm.v576
-rw-r--r--mathcomp/fingroup/presentation.v254
-rw-r--r--mathcomp/fingroup/quotient.v972
-rw-r--r--mathcomp/fingroup/zmodp.v362
11 files changed, 12585 insertions, 0 deletions
diff --git a/mathcomp/fingroup/action.v b/mathcomp/fingroup/action.v
new file mode 100644
index 0000000..902e9ca
--- /dev/null
+++ b/mathcomp/fingroup/action.v
@@ -0,0 +1,2719 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat div seq fintype.
+Require Import bigop finset fingroup morphism perm automorphism quotient.
+
+(******************************************************************************)
+(* Group action: orbits, stabilisers, transitivity. *)
+(* is_action D to == the function to : T -> aT -> T defines an action *)
+(* of D : {set aT} on T. *)
+(* action D T == structure for a function defining an action of D. *)
+(* act_dom to == the domain D of to : action D rT. *)
+(* {action: aT &-> T} == structure for a total action. *)
+(* := action [set: aT] T *)
+(* TotalAction to1 toM == the constructor for total actions; to1 and toM *)
+(* are the proofs of the action identities for 1 and *)
+(* a * b, respectively. *)
+(* is_groupAction R to == to is a group action on range R: for all a in D, *)
+(* the permutation induced by to a is in Aut R. Thus *)
+(* the action of D must be trivial outside R. *)
+(* groupAction D R == the structure for group actions of D on R. This *)
+(* is a telescope on action D rT. *)
+(* gact_range to == the range R of to : groupAction D R. *)
+(* GroupAction toAut == construct a groupAction for action to from *)
+(* toAut : actm to @* D \subset Aut R (actm to is *)
+(* the morphism to {perm rT} associated to 'to'). *)
+(* orbit to A x == the orbit of x under the action of A via to. *)
+(* orbit_transversal to A S == a transversal of the partition orbit to A @: S *)
+(* of S, provided A acts on S via to. *)
+(* amove to A x y == the set of a in A whose action send x to y. *)
+(* 'C_A[x | to] == the stabiliser of x : rT in A :&: D. *)
+(* 'C_A(S | to) == the point-wise stabiliser of S : {set rT} in D :&: A. *)
+(* 'N_A(S | to) == the global stabiliser of S : {set rT} in D :&: A. *)
+(* 'Fix_(S | to)[a] == the set of fixpoints of a in S. *)
+(* 'Fix_(S | to)(A) == the set of fixpoints of A in S. *)
+(* In the first three _A can be omitted and defaults to the domain D of to; *)
+(* In the last two S can be omitted and defaults to [set: T], so 'Fix_to[a] *)
+(* is the set of all fixpoints of a. *)
+(* The domain restriction ensures that stabilisers have a canonical group *)
+(* structure, but note that 'Fix sets are generally not groups. Indeed, we *)
+(* provide alternative definitions when to is a group action on R: *)
+(* 'C_(G | to)(A) == the centraliser in R :&: G of the group action of *)
+(* D :&: A via to *)
+(* 'C_(G | to)[a] == the centraliser in R :&: G of a \in D, via to. *)
+(* These sets are groups when G is. G can be omitted: 'C(|to)(A) is the *)
+(* centraliser in R of the action of D :&: A via to. *)
+(* [acts A, on S | to] == A \subset D acts on the set S via to. *)
+(* {acts A, on S | to} == A acts on the set S (Prop statement). *)
+(* {acts A, on group G | to} == [acts A, on S | to] /\ G \subset R, i.e., *)
+(* A \subset D acts on G \subset R, via *)
+(* to : groupAction D R. *)
+(* [transitive A, on S | to] == A acts transitively on S. *)
+(* [faithful A, on S | to] == A acts faithfully on S. *)
+(* acts_irreducibly to A G == A acts irreducibly via the groupAction to *)
+(* on the nontrivial group G, i.e., A does *)
+(* not act on any nontrivial subgroup of G. *)
+(* Important caveat: the definitions of orbit, amove, 'Fix_(S | to)(A), *)
+(* transitive and faithful assume that A is a subset of the domain D. As most *)
+(* of the permutation actions we consider are total this is usually harmless. *)
+(* (Note that the theory of partial actions is only partially developed.) *)
+(* In all of the above, to is expected to be the actual action structure, *)
+(* not merely the function. There is a special scope %act for actions, and *)
+(* constructions and notations for many classical actions: *)
+(* 'P == natural action of a permutation group via aperm. *)
+(* 'J == internal group action (conjugation) via conjg (_ ^ _). *)
+(* 'R == regular group action (right translation) via mulg (_ * _). *)
+(* (but, to limit ambiguity, _ * _ is NOT a canonical action) *)
+(* to^* == the action induced by to on {set rT} via to^* (== setact to). *)
+(* 'Js == the internal action on subsets via _ :^ _, equivalent to 'J^*. *)
+(* 'Rs == the regular action on subsets via rcoset, equivalent to 'R^*. *)
+(* 'JG == the conjugation action on {group rT} via (_ :^ _)%G. *)
+(* to / H == the action induced by to on coset_of H via qact to H, and *)
+(* restricted to qact_dom to H == 'N(rcosets H 'N(H) | to^* ). *)
+(* 'Q == the action induced to cosets by conjugation; the domain is *)
+(* qact_dom 'J H, which is provably equal to 'N(H). *)
+(* to %% A == the action of coset_of A via modact to A, with domain D / A *)
+(* and support restricted to 'C(D :&: A | to). *)
+(* to \ sAD == the action of A via ract to sAD == to, if sAD : A \subset D. *)
+(* [Aut G] == the permutation action restricted to Aut G, via autact G. *)
+(* <[nRA]> == the action of A on R via actby nRA == to in A and on R, and *)
+(* the trivial action elsewhere; here nRA : [acts A, on R | to] *)
+(* or nRA : {acts A, on group R | to}. *)
+(* to^? == the action induced by to on sT : @subType rT P, via subact to *)
+(* with domain subact_dom P to == 'N([set x | P x] | to). *)
+(* <<phi>> == the action of phi : D >-> {perm rT}, via mact phi. *)
+(* to \o f == the composite action (with domain f @*^-1 D) of the action to *)
+(* with f : {morphism G >-> aT}, via comp_act to f. Here f must *)
+(* be the actual morphism object (e.g., coset_morphism H), not *)
+(* the underlying function (e.g., coset H). *)
+(* The explicit application of an action to is usually written (to%act x a), *)
+(* where the %act omitted if to is an abstract action or a set action to0^*. *)
+(* Note that this form will simplify and expose the acting function. *)
+(* There is a %gact scope for group actions; the notations above are *)
+(* recognised in %gact when they denote canonical group actions. *)
+(* Actions can be used to define morphisms: *)
+(* actperm to == the morphism D >-> {perm rT} induced by to. *)
+(* actm to a == if a \in D the function on D induced by the action to, else *)
+(* the identity function. If to is a group action with range R *)
+(* then actm to a is canonically a morphism on R. *)
+(* We also define here the restriction operation on permutations (the domain *)
+(* of this operations is a stabiliser), and local automorphpism groups: *)
+(* restr_perm S p == if p acts on S, the permutation with support in S that *)
+(* coincides with p on S; else the identity. Note that *)
+(* restr_perm is a permutation group morphism that maps *)
+(* Aut G to Aut S when S is a subgroup of G. *)
+(* Aut_in A G == the local permutation group 'N_A(G | 'P) / 'C_A(G | 'P) *)
+(* Usually A is an automorphism group, and then Aut_in A G *)
+(* is isomorphic to a subgroup of Aut G, specifically *)
+(* restr_perm @* A. *)
+(* Finally, gproduct.v will provide a semi-direct group construction that *)
+(* maps an external group action to an internal one; the theory of morphisms *)
+(* between such products makes use of the following definition: *)
+(* morph_act to to' f fA <=> the action of to' on the images of f and fA is *)
+(* the image of the action of to, i.e., for all x and a we *)
+(* have f (to x a) = to' (f x) (fA a). Note that there is *)
+(* no mention of the domains of to and to'; if needed, this *)
+(* predicate should be restricted via the {in ...} notation *)
+(* and domain conditions should be added. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section ActionDef.
+
+Variables (aT : finGroupType) (D : {set aT}) (rT : Type).
+Implicit Types a b : aT.
+Implicit Type x : rT.
+
+Definition act_morph to x := forall a b, to x (a * b) = to (to x a) b.
+
+Definition is_action to :=
+ left_injective to /\ forall x, {in D &, act_morph to x}.
+
+Record action := Action {act :> rT -> aT -> rT; _ : is_action act}.
+
+Definition clone_action to :=
+ let: Action _ toP := to return {type of Action for to} -> action in
+ fun k => k toP.
+
+End ActionDef.
+
+(* Need to close the Section here to avoid re-declaring all Argument Scopes *)
+Delimit Scope action_scope with act.
+Bind Scope action_scope with action.
+Arguments Scope act_morph [_ group_scope _ _ group_scope].
+Arguments Scope is_action [_ group_scope _ _].
+Arguments Scope act
+ [_ group_scope type_scope action_scope group_scope group_scope].
+Arguments Scope clone_action [_ group_scope type_scope action_scope _].
+
+Notation "{ 'action' aT &-> T }" := (action [set: aT] T)
+ (at level 0, format "{ 'action' aT &-> T }") : type_scope.
+
+Notation "[ 'action' 'of' to ]" := (clone_action (@Action _ _ _ to))
+ (at level 0, format "[ 'action' 'of' to ]") : form_scope.
+
+Definition act_dom aT D rT of @action aT D rT := D.
+
+Section TotalAction.
+
+Variables (aT : finGroupType) (rT : Type) (to : rT -> aT -> rT).
+Hypotheses (to1 : to^~ 1 =1 id) (toM : forall x, act_morph to x).
+
+Lemma is_total_action : is_action setT to.
+Proof.
+split=> [a | x a b _ _] /=; last by rewrite toM.
+by apply: can_inj (to^~ a^-1) _ => x; rewrite -toM ?mulgV.
+Qed.
+
+Definition TotalAction := Action is_total_action.
+
+End TotalAction.
+
+Section ActionDefs.
+
+Variables (aT aT' : finGroupType) (D : {set aT}) (D' : {set aT'}).
+
+Definition morph_act rT rT' (to : action D rT) (to' : action D' rT') f fA :=
+ forall x a, f (to x a) = to' (f x) (fA a).
+
+Variable rT : finType. (* Most definitions require a finType structure on rT *)
+Implicit Type to : action D rT.
+Implicit Type A : {set aT}.
+Implicit Type S : {set rT}.
+
+Definition actm to a := if a \in D then to^~ a else id.
+
+Definition setact to S a := [set to x a | x in S].
+
+Definition orbit to A x := to x @: A.
+
+Definition amove to A x y := [set a in A | to x a == y].
+
+Definition afix to A := [set x | A \subset [set a | to x a == x]].
+
+Definition astab S to := D :&: [set a | S \subset [set x | to x a == x]].
+
+Definition astabs S to := D :&: [set a | S \subset to^~ a @^-1: S].
+
+Definition acts_on A S to := {in A, forall a x, (to x a \in S) = (x \in S)}.
+
+Definition atrans A S to := S \in orbit to A @: S.
+
+Definition faithful A S to := A :&: astab S to \subset [1].
+
+End ActionDefs.
+
+Arguments Scope setact [_ group_scope _ action_scope group_scope group_scope].
+Arguments Scope orbit [_ group_scope _ action_scope group_scope group_scope].
+Arguments Scope amove
+ [_ group_scope _ action_scope group_scope group_scope group_scope].
+Arguments Scope afix [_ group_scope _ action_scope group_scope].
+Arguments Scope astab [_ group_scope _ group_scope action_scope].
+Arguments Scope astabs [_ group_scope _ group_scope action_scope].
+Arguments Scope acts_on [_ group_scope _ group_scope group_scope action_scope].
+Arguments Scope atrans [_ group_scope _ group_scope group_scope action_scope].
+Arguments Scope faithful [_ group_scope _ group_scope group_scope action_scope].
+
+Notation "to ^*" := (setact to) (at level 2, format "to ^*") : fun_scope.
+
+Prenex Implicits orbit amove.
+
+Notation "''Fix_' to ( A )" := (afix to A)
+ (at level 8, to at level 2, format "''Fix_' to ( A )") : group_scope.
+
+(* camlp4 grammar factoring *)
+Notation "''Fix_' ( to ) ( A )" := 'Fix_to(A)
+ (at level 8, only parsing) : group_scope.
+
+Notation "''Fix_' ( S | to ) ( A )" := (S :&: 'Fix_to(A))
+ (at level 8, format "''Fix_' ( S | to ) ( A )") : group_scope.
+
+Notation "''Fix_' to [ a ]" := ('Fix_to([set a]))
+ (at level 8, to at level 2, format "''Fix_' to [ a ]") : group_scope.
+
+Notation "''Fix_' ( S | to ) [ a ]" := (S :&: 'Fix_to[a])
+ (at level 8, format "''Fix_' ( S | to ) [ a ]") : group_scope.
+
+Notation "''C' ( S | to )" := (astab S to)
+ (at level 8, format "''C' ( S | to )") : group_scope.
+
+Notation "''C_' A ( S | to )" := (A :&: 'C(S | to))
+ (at level 8, A at level 2, format "''C_' A ( S | to )") : group_scope.
+Notation "''C_' ( A ) ( S | to )" := 'C_A(S | to)
+ (at level 8, only parsing) : group_scope.
+
+Notation "''C' [ x | to ]" := ('C([set x] | to))
+ (at level 8, format "''C' [ x | to ]") : group_scope.
+
+Notation "''C_' A [ x | to ]" := (A :&: 'C[x | to])
+ (at level 8, A at level 2, format "''C_' A [ x | to ]") : group_scope.
+Notation "''C_' ( A ) [ x | to ]" := 'C_A[x | to]
+ (at level 8, only parsing) : group_scope.
+
+Notation "''N' ( S | to )" := (astabs S to)
+ (at level 8, format "''N' ( S | to )") : group_scope.
+
+Notation "''N_' A ( S | to )" := (A :&: 'N(S | to))
+ (at level 8, A at level 2, format "''N_' A ( S | to )") : group_scope.
+
+Notation "[ 'acts' A , 'on' S | to ]" := (A \subset pred_of_set 'N(S | to))
+ (at level 0, format "[ 'acts' A , 'on' S | to ]") : form_scope.
+
+Notation "{ 'acts' A , 'on' S | to }" := (acts_on A S to)
+ (at level 0, format "{ 'acts' A , 'on' S | to }") : form_scope.
+
+Notation "[ 'transitive' A , 'on' S | to ]" := (atrans A S to)
+ (at level 0, format "[ 'transitive' A , 'on' S | to ]") : form_scope.
+
+Notation "[ 'faithful' A , 'on' S | to ]" := (faithful A S to)
+ (at level 0, format "[ 'faithful' A , 'on' S | to ]") : form_scope.
+
+Section RawAction.
+(* Lemmas that do not require the group structure on the action domain. *)
+(* Some lemmas like actMin would be actually be valid for arbitrary rT, *)
+(* e.g., for actions on a function type, but would be difficult to use *)
+(* as a view due to the confusion between parameters and assumptions. *)
+
+Variables (aT : finGroupType) (D : {set aT}) (rT : finType) (to : action D rT).
+
+Implicit Types (a : aT) (x y : rT) (A B : {set aT}) (S T : {set rT}).
+
+Lemma act_inj : left_injective to. Proof. by case: to => ? []. Qed.
+Implicit Arguments act_inj [].
+
+Lemma actMin x : {in D &, act_morph to x}.
+Proof. by case: to => ? []. Qed.
+
+Lemma actmEfun a : a \in D -> actm to a = to^~ a.
+Proof. by rewrite /actm => ->. Qed.
+
+Lemma actmE a : a \in D -> actm to a =1 to^~ a.
+Proof. by move=> Da; rewrite actmEfun. Qed.
+
+Lemma setactE S a : to^* S a = [set to x a | x in S].
+Proof. by []. Qed.
+
+Lemma mem_setact S a x : x \in S -> to x a \in to^* S a.
+Proof. exact: mem_imset. Qed.
+
+Lemma card_setact S a : #|to^* S a| = #|S|.
+Proof. by apply: card_imset; exact: act_inj. Qed.
+
+Lemma setact_is_action : is_action D to^*.
+Proof.
+split=> [a R S eqRS | a b Da Db S]; last first.
+ rewrite /setact /= -imset_comp; apply: eq_imset => x; exact: actMin.
+apply/setP=> x; apply/idP/idP=> /(mem_setact a).
+ by rewrite eqRS => /imsetP[y Sy /act_inj->].
+by rewrite -eqRS => /imsetP[y Sy /act_inj->].
+Qed.
+
+Canonical set_action := Action setact_is_action.
+
+Lemma orbitE A x : orbit to A x = to x @: A. Proof. by []. Qed.
+
+Lemma orbitP A x y :
+ reflect (exists2 a, a \in A & to x a = y) (y \in orbit to A x).
+Proof. by apply: (iffP imsetP) => [] [a]; exists a. Qed.
+
+Lemma mem_orbit A x a : a \in A -> to x a \in orbit to A x.
+Proof. exact: mem_imset. Qed.
+
+Lemma afixP A x : reflect (forall a, a \in A -> to x a = x) (x \in 'Fix_to(A)).
+Proof.
+rewrite inE; apply: (iffP subsetP) => [xfix a /xfix | xfix a Aa].
+ by rewrite inE => /eqP.
+by rewrite inE xfix.
+Qed.
+
+Lemma afixS A B : A \subset B -> 'Fix_to(B) \subset 'Fix_to(A).
+Proof. by move=> sAB; apply/subsetP=> u; rewrite !inE; exact: subset_trans. Qed.
+
+Lemma afixU A B : 'Fix_to(A :|: B) = 'Fix_to(A) :&: 'Fix_to(B).
+Proof. by apply/setP=> x; rewrite !inE subUset. Qed.
+
+Lemma afix1P a x : reflect (to x a = x) (x \in 'Fix_to[a]).
+Proof. by rewrite inE sub1set inE; exact: eqP. Qed.
+
+Lemma astabIdom S : 'C_D(S | to) = 'C(S | to).
+Proof. by rewrite setIA setIid. Qed.
+
+Lemma astab_dom S : {subset 'C(S | to) <= D}.
+Proof. by move=> a /setIP[]. Qed.
+
+Lemma astab_act S a x : a \in 'C(S | to) -> x \in S -> to x a = x.
+Proof.
+rewrite 2!inE => /andP[_ cSa] Sx; apply/eqP.
+by have:= subsetP cSa x Sx; rewrite inE.
+Qed.
+
+Lemma astabS S1 S2 : S1 \subset S2 -> 'C(S2 | to) \subset 'C(S1 | to).
+Proof.
+move=> sS12; apply/subsetP=> x; rewrite !inE => /andP[->].
+exact: subset_trans.
+Qed.
+
+Lemma astabsIdom S : 'N_D(S | to) = 'N(S | to).
+Proof. by rewrite setIA setIid. Qed.
+
+Lemma astabs_dom S : {subset 'N(S | to) <= D}.
+Proof. by move=> a /setIdP[]. Qed.
+
+Lemma astabs_act S a x : a \in 'N(S | to) -> (to x a \in S) = (x \in S).
+Proof.
+rewrite 2!inE subEproper properEcard => /andP[_].
+rewrite (card_preimset _ (act_inj _)) ltnn andbF orbF => /eqP{2}->.
+by rewrite inE.
+Qed.
+
+Lemma astab_sub S : 'C(S | to) \subset 'N(S | to).
+Proof.
+apply/subsetP=> a cSa; rewrite !inE (astab_dom cSa).
+by apply/subsetP=> x Sx; rewrite inE (astab_act cSa).
+Qed.
+
+Lemma astabsC S : 'N(~: S | to) = 'N(S | to).
+Proof.
+apply/setP=> a; apply/idP/idP=> nSa; rewrite !inE (astabs_dom nSa).
+ by rewrite -setCS -preimsetC; apply/subsetP=> x; rewrite inE astabs_act.
+by rewrite preimsetC setCS; apply/subsetP=> x; rewrite inE astabs_act.
+Qed.
+
+Lemma astabsI S T : 'N(S | to) :&: 'N(T | to) \subset 'N(S :&: T | to).
+Proof.
+apply/subsetP=> a; rewrite !inE -!andbA preimsetI => /and4P[-> nSa _ nTa] /=.
+by rewrite setISS.
+Qed.
+
+Lemma astabs_setact S a : a \in 'N(S | to) -> to^* S a = S.
+Proof.
+move=> nSa; apply/eqP; rewrite eqEcard card_setact leqnn andbT.
+by apply/subsetP=> _ /imsetP[x Sx ->]; rewrite astabs_act.
+Qed.
+
+Lemma astab1_set S : 'C[S | set_action] = 'N(S | to).
+Proof.
+apply/setP=> a; apply/idP/idP=> nSa.
+ case/setIdP: nSa => Da; rewrite !inE Da sub1set inE => /eqP defS.
+ by apply/subsetP=> x Sx; rewrite inE -defS mem_setact.
+by rewrite !inE (astabs_dom nSa) sub1set inE /= astabs_setact.
+Qed.
+
+Lemma astabs_set1 x : 'N([set x] | to) = 'C[x | to].
+Proof.
+apply/eqP; rewrite eqEsubset astab_sub andbC setIS //.
+by apply/subsetP=> a; rewrite ?(inE,sub1set).
+Qed.
+
+Lemma acts_dom A S : [acts A, on S | to] -> A \subset D.
+Proof. by move=> nSA; rewrite (subset_trans nSA) ?subsetIl. Qed.
+
+Lemma acts_act A S : [acts A, on S | to] -> {acts A, on S | to}.
+Proof. by move=> nAS a Aa x; rewrite astabs_act ?(subsetP nAS). Qed.
+
+Lemma astabCin A S :
+ A \subset D -> (A \subset 'C(S | to)) = (S \subset 'Fix_to(A)).
+Proof.
+move=> sAD; apply/subsetP/subsetP=> [sAC x xS | sSF a aA].
+ by apply/afixP=> a aA; exact: astab_act (sAC _ aA) xS.
+rewrite !inE (subsetP sAD _ aA); apply/subsetP=> x xS.
+by move/afixP/(_ _ aA): (sSF _ xS); rewrite inE => ->.
+Qed.
+
+Section ActsSetop.
+
+Variables (A : {set aT}) (S T : {set rT}).
+Hypotheses (AactS : [acts A, on S | to]) (AactT : [acts A, on T | to]).
+
+Lemma astabU : 'C(S :|: T | to) = 'C(S | to) :&: 'C(T | to).
+Proof. by apply/setP=> a; rewrite !inE subUset; case: (a \in D). Qed.
+
+Lemma astabsU : 'N(S | to) :&: 'N(T | to) \subset 'N(S :|: T | to).
+Proof.
+by rewrite -(astabsC S) -(astabsC T) -(astabsC (S :|: T)) setCU astabsI.
+Qed.
+
+Lemma astabsD : 'N(S | to) :&: 'N(T | to) \subset 'N(S :\: T| to).
+Proof. by rewrite setDE -(astabsC T) astabsI. Qed.
+
+Lemma actsI : [acts A, on S :&: T | to].
+Proof. by apply: subset_trans (astabsI S T); rewrite subsetI AactS. Qed.
+
+Lemma actsU : [acts A, on S :|: T | to].
+Proof. by apply: subset_trans astabsU; rewrite subsetI AactS. Qed.
+
+Lemma actsD : [acts A, on S :\: T | to].
+Proof. by apply: subset_trans astabsD; rewrite subsetI AactS. Qed.
+
+End ActsSetop.
+
+Lemma acts_in_orbit A S x y :
+ [acts A, on S | to] -> y \in orbit to A x -> x \in S -> y \in S.
+Proof.
+by move=> nSA/imsetP[a Aa ->{y}] Sx; rewrite (astabs_act _ (subsetP nSA a Aa)).
+Qed.
+
+Lemma subset_faithful A B S :
+ B \subset A -> [faithful A, on S | to] -> [faithful B, on S | to].
+Proof. by move=> sAB; apply: subset_trans; exact: setSI. Qed.
+
+Section Reindex.
+
+Variables (vT : Type) (idx : vT) (op : Monoid.com_law idx) (S : {set rT}).
+
+Lemma reindex_astabs a F : a \in 'N(S | to) ->
+ \big[op/idx]_(i in S) F i = \big[op/idx]_(i in S) F (to i a).
+Proof.
+move=> nSa; rewrite (reindex_inj (act_inj a)); apply: eq_bigl => x.
+exact: astabs_act.
+Qed.
+
+Lemma reindex_acts A a F : [acts A, on S | to] -> a \in A ->
+ \big[op/idx]_(i in S) F i = \big[op/idx]_(i in S) F (to i a).
+Proof. by move=> nSA /(subsetP nSA); exact: reindex_astabs. Qed.
+
+End Reindex.
+
+End RawAction.
+
+(* Warning: this directive depends on names of bound variables in the *)
+(* definition of injective, in ssrfun.v. *)
+Implicit Arguments act_inj [[aT] [D] [rT] x1 x2].
+
+Notation "to ^*" := (set_action to) : action_scope.
+
+Implicit Arguments orbitP [aT D rT to A x y].
+Implicit Arguments afixP [aT D rT to A x].
+Implicit Arguments afix1P [aT D rT to a x].
+Prenex Implicits orbitP afixP afix1P.
+
+Implicit Arguments reindex_astabs [aT D rT vT idx op S F].
+Implicit Arguments reindex_acts [aT D rT vT idx op S A a F].
+
+Section PartialAction.
+(* Lemmas that require a (partial) group domain. *)
+
+Variables (aT : finGroupType) (D : {group aT}) (rT : finType).
+Variable to : action D rT.
+
+Implicit Types a : aT.
+Implicit Types x y : rT.
+Implicit Types A B : {set aT}.
+Implicit Types G H : {group aT}.
+Implicit Types S : {set rT}.
+
+Lemma act1 x : to x 1 = x.
+Proof. by apply: (act_inj to 1); rewrite -actMin ?mulg1. Qed.
+
+Lemma actKin : {in D, right_loop invg to}.
+Proof. by move=> a Da /= x; rewrite -actMin ?groupV // mulgV act1. Qed.
+
+Lemma actKVin : {in D, rev_right_loop invg to}.
+Proof. by move=> a Da /= x; rewrite -{2}(invgK a) actKin ?groupV. Qed.
+
+Lemma setactVin S a : a \in D -> to^* S a^-1 = to^~ a @^-1: S.
+Proof.
+by move=> Da; apply: can2_imset_pre; [exact: actKVin | exact: actKin].
+Qed.
+
+Lemma actXin x a i : a \in D -> to x (a ^+ i) = iter i (to^~ a) x.
+Proof.
+move=> Da; elim: i => /= [|i <-]; first by rewrite act1.
+by rewrite expgSr actMin ?groupX.
+Qed.
+
+Lemma afix1 : 'Fix_to(1) = setT.
+Proof. by apply/setP=> x; rewrite !inE sub1set inE act1 eqxx. Qed.
+
+Lemma afixD1 G : 'Fix_to(G^#) = 'Fix_to(G).
+Proof. by rewrite -{2}(setD1K (group1 G)) afixU afix1 setTI. Qed.
+
+Lemma orbit_refl G x : x \in orbit to G x.
+Proof. by rewrite -{1}[x]act1 mem_orbit. Qed.
+
+Local Notation orbit_rel A := (fun x y => y \in orbit to A x).
+
+Lemma contra_orbit G x y : x \notin orbit to G y -> x != y.
+Proof. by apply: contraNneq => ->; exact: orbit_refl. Qed.
+
+Lemma orbit_in_sym G : G \subset D -> symmetric (orbit_rel G).
+Proof.
+move=> sGD; apply: symmetric_from_pre => x y /imsetP[a Ga].
+by move/(canLR (actKin (subsetP sGD a Ga))) <-; rewrite mem_orbit ?groupV.
+Qed.
+
+Lemma orbit_in_trans G : G \subset D -> transitive (orbit_rel G).
+Proof.
+move=> sGD _ x _ /imsetP[a Ga ->] /imsetP[b Gb ->].
+by rewrite -actMin ?mem_orbit ?groupM // (subsetP sGD).
+Qed.
+
+Lemma orbit_in_transl G x y :
+ G \subset D -> y \in orbit to G x -> orbit to G y = orbit to G x.
+Proof.
+move=> sGD Gxy; apply/setP=> z.
+by apply/idP/idP; apply: orbit_in_trans; rewrite // orbit_in_sym.
+Qed.
+
+Lemma orbit_in_transr G x y z :
+ G \subset D -> y \in orbit to G x ->
+ (y \in orbit to G z) = (x \in orbit to G z).
+Proof.
+by move=> sGD Gxy; rewrite !(orbit_in_sym _ z) ?(orbit_in_transl _ Gxy).
+Qed.
+
+Lemma orbit_act_in x a G :
+ G \subset D -> a \in G -> orbit to G (to x a) = orbit to G x.
+Proof. by move=> sGD /mem_orbit/orbit_in_transl->. Qed.
+
+Lemma orbit_actr_in x a G y :
+ G \subset D -> a \in G -> (to y a \in orbit to G x) = (y \in orbit to G x).
+Proof. by move=> sGD /mem_orbit/orbit_in_transr->. Qed.
+
+Lemma orbit_inv_in A x y :
+ A \subset D -> (y \in orbit to A^-1 x) = (x \in orbit to A y).
+Proof.
+move/subsetP=> sAD; apply/imsetP/imsetP=> [] [a Aa ->].
+ by exists a^-1; rewrite -?mem_invg ?actKin // -groupV sAD -?mem_invg.
+by exists a^-1; rewrite ?memV_invg ?actKin // sAD.
+Qed.
+
+Lemma orbit_lcoset_in A a x :
+ A \subset D -> a \in D ->
+ orbit to (a *: A) x = orbit to A (to x a).
+Proof.
+move/subsetP=> sAD Da; apply/setP=> y; apply/imsetP/imsetP=> [] [b Ab ->{y}].
+ by exists (a^-1 * b); rewrite -?actMin ?mulKVg // ?sAD -?mem_lcoset.
+by exists (a * b); rewrite ?mem_mulg ?set11 ?actMin // sAD.
+Qed.
+
+Lemma orbit_rcoset_in A a x y :
+ A \subset D -> a \in D ->
+ (to y a \in orbit to (A :* a) x) = (y \in orbit to A x).
+Proof.
+move=> sAD Da; rewrite -orbit_inv_in ?mul_subG ?sub1set // invMg.
+by rewrite invg_set1 orbit_lcoset_in ?inv_subG ?groupV ?actKin ?orbit_inv_in.
+Qed.
+
+Lemma orbit_conjsg_in A a x y :
+ A \subset D -> a \in D ->
+ (to y a \in orbit to (A :^ a) (to x a)) = (y \in orbit to A x).
+Proof.
+move=> sAD Da; rewrite conjsgE.
+by rewrite orbit_lcoset_in ?groupV ?mul_subG ?sub1set ?actKin ?orbit_rcoset_in.
+Qed.
+
+Lemma orbit1P G x : reflect (orbit to G x = [set x]) (x \in 'Fix_to(G)).
+Proof.
+apply: (iffP afixP) => [xfix | xfix a Ga].
+ apply/eqP; rewrite eq_sym eqEsubset sub1set -{1}[x]act1 mem_imset //=.
+ by apply/subsetP=> y; case/imsetP=> a Ga ->; rewrite inE xfix.
+by apply/set1P; rewrite -xfix mem_imset.
+Qed.
+
+Lemma card_orbit1 G x : #|orbit to G x| = 1%N -> orbit to G x = [set x].
+Proof.
+move=> orb1; apply/eqP; rewrite eq_sym eqEcard {}orb1 cards1.
+by rewrite sub1set orbit_refl.
+Qed.
+
+Lemma orbit_partition G S :
+ [acts G, on S | to] -> partition (orbit to G @: S) S.
+Proof.
+move=> actsGS; have sGD := acts_dom actsGS.
+have eqiG: {in S & &, equivalence_rel [rel x y | y \in orbit to G x]}.
+ by move=> x y z * /=; rewrite orbit_refl; split=> // /orbit_in_transl->.
+congr (partition _ _): (equivalence_partitionP eqiG).
+apply: eq_in_imset => x Sx; apply/setP=> y.
+by rewrite inE /= andb_idl // => /acts_in_orbit->.
+Qed.
+
+Definition orbit_transversal A S := transversal (orbit to A @: S) S.
+
+Lemma orbit_transversalP G S (P := orbit to G @: S)
+ (X := orbit_transversal G S) :
+ [acts G, on S | to] ->
+ [/\ is_transversal X P S, X \subset S,
+ {in X &, forall x y, (y \in orbit to G x) = (x == y)}
+ & forall x, x \in S -> exists2 a, a \in G & to x a \in X].
+Proof.
+move/orbit_partition; rewrite -/P => partP.
+have [/eqP defS tiP _] := and3P partP.
+have trXP: is_transversal X P S := transversalP partP.
+have sXS: X \subset S := transversal_sub trXP.
+split=> // [x y Xx Xy /= | x Sx].
+ have Sx := subsetP sXS x Xx.
+ rewrite -(inj_in_eq (pblock_inj trXP)) // eq_pblock ?defS //.
+ by rewrite (def_pblock tiP (mem_imset _ Sx)) ?orbit_refl.
+have /imsetP[y Xy defxG]: orbit to G x \in pblock P @: X.
+ by rewrite (pblock_transversal trXP) ?mem_imset.
+suffices /orbitP[a Ga def_y]: y \in orbit to G x by exists a; rewrite ?def_y.
+by rewrite defxG mem_pblock defS (subsetP sXS).
+Qed.
+
+Lemma group_set_astab S : group_set 'C(S | to).
+Proof.
+apply/group_setP; split=> [|a b cSa cSb].
+ by rewrite !inE group1; apply/subsetP=> x _; rewrite inE act1.
+rewrite !inE groupM ?(@astab_dom _ _ _ to S) //; apply/subsetP=> x Sx.
+by rewrite inE actMin ?(@astab_dom _ _ _ to S) ?(astab_act _ Sx).
+Qed.
+
+Canonical astab_group S := group (group_set_astab S).
+
+Lemma afix_gen_in A : A \subset D -> 'Fix_to(<<A>>) = 'Fix_to(A).
+Proof.
+move=> sAD; apply/eqP; rewrite eqEsubset afixS ?sub_gen //=.
+by rewrite -astabCin gen_subG ?astabCin.
+Qed.
+
+Lemma afix_cycle_in a : a \in D -> 'Fix_to(<[a]>) = 'Fix_to[a].
+Proof. by move=> Da; rewrite afix_gen_in ?sub1set. Qed.
+
+Lemma afixYin A B :
+ A \subset D -> B \subset D -> 'Fix_to(A <*> B) = 'Fix_to(A) :&: 'Fix_to(B).
+Proof. by move=> sAD sBD; rewrite afix_gen_in ?afixU // subUset sAD. Qed.
+
+Lemma afixMin G H :
+ G \subset D -> H \subset D -> 'Fix_to(G * H) = 'Fix_to(G) :&: 'Fix_to(H).
+Proof.
+by move=> sGD sHD; rewrite -afix_gen_in ?mul_subG // genM_join afixYin.
+Qed.
+
+Lemma sub_astab1_in A x :
+ A \subset D -> (A \subset 'C[x | to]) = (x \in 'Fix_to(A)).
+Proof. by move=> sAD; rewrite astabCin ?sub1set. Qed.
+
+Lemma group_set_astabs S : group_set 'N(S | to).
+Proof.
+apply/group_setP; split=> [|a b cSa cSb].
+ by rewrite !inE group1; apply/subsetP=> x Sx; rewrite inE act1.
+rewrite !inE groupM ?(@astabs_dom _ _ _ to S) //; apply/subsetP=> x Sx.
+by rewrite inE actMin ?(@astabs_dom _ _ _ to S) ?astabs_act.
+Qed.
+
+Canonical astabs_group S := group (group_set_astabs S).
+
+Lemma astab_norm S : 'N(S | to) \subset 'N('C(S | to)).
+Proof.
+apply/subsetP=> a nSa; rewrite inE sub_conjg; apply/subsetP=> b cSb.
+have [Da Db] := (astabs_dom nSa, astab_dom cSb).
+rewrite mem_conjgV !inE groupJ //; apply/subsetP=> x Sx.
+rewrite inE !actMin ?groupM ?groupV //.
+by rewrite (astab_act cSb) ?actKVin ?astabs_act ?groupV.
+Qed.
+
+Lemma astab_normal S : 'C(S | to) <| 'N(S | to).
+Proof. by rewrite /normal astab_sub astab_norm. Qed.
+
+Lemma acts_sub_orbit G S x :
+ [acts G, on S | to] -> (orbit to G x \subset S) = (x \in S).
+Proof.
+move/acts_act=> GactS.
+apply/subsetP/idP=> [| Sx y]; first by apply; exact: orbit_refl.
+by case/orbitP=> a Ga <-{y}; rewrite GactS.
+Qed.
+
+Lemma acts_orbit G x : G \subset D -> [acts G, on orbit to G x | to].
+Proof.
+move/subsetP=> sGD; apply/subsetP=> a Ga; rewrite !inE sGD //.
+apply/subsetP=> _ /imsetP[b Gb ->].
+by rewrite inE -actMin ?sGD // mem_imset ?groupM.
+Qed.
+
+Lemma acts_subnorm_fix A : [acts 'N_D(A), on 'Fix_to(D :&: A) | to].
+Proof.
+apply/subsetP=> a nAa; have [Da _] := setIP nAa; rewrite !inE Da.
+apply/subsetP=> x Cx; rewrite inE; apply/afixP=> b DAb.
+have [Db _]:= setIP DAb; rewrite -actMin // conjgCV actMin ?groupJ ?groupV //.
+by rewrite /= (afixP Cx) // memJ_norm // groupV (subsetP (normsGI _ _) _ nAa).
+Qed.
+
+Lemma atrans_orbit G x : [transitive G, on orbit to G x | to].
+Proof. by apply: mem_imset; exact: orbit_refl. Qed.
+
+Section OrbitStabilizer.
+
+Variables (G : {group aT}) (x : rT).
+Hypothesis sGD : G \subset D.
+Let ssGD := subsetP sGD.
+
+Lemma amove_act a : a \in G -> amove to G x (to x a) = 'C_G[x | to] :* a.
+Proof.
+move=> Ga; apply/setP=> b; have Da := ssGD Ga.
+rewrite mem_rcoset !(inE, sub1set) !groupMr ?groupV //.
+by case Gb: (b \in G); rewrite //= actMin ?groupV ?ssGD ?(canF_eq (actKVin Da)).
+Qed.
+
+Lemma amove_orbit : amove to G x @: orbit to G x = rcosets 'C_G[x | to] G.
+Proof.
+apply/setP => Ha; apply/imsetP/rcosetsP=> [[y] | [a Ga ->]].
+ by case/imsetP=> b Gb -> ->{Ha y}; exists b => //; rewrite amove_act.
+by rewrite -amove_act //; exists (to x a); first exact: mem_orbit.
+Qed.
+
+Lemma amoveK :
+ {in orbit to G x, cancel (amove to G x) (fun Ca => to x (repr Ca))}.
+Proof.
+move=> _ /orbitP[a Ga <-]; rewrite amove_act //= -[G :&: _]/(gval _).
+case: repr_rcosetP => b; rewrite !(inE, sub1set)=> /and3P[Gb _ xbx].
+by rewrite actMin ?ssGD ?(eqP xbx).
+Qed.
+
+Lemma orbit_stabilizer :
+ orbit to G x = [set to x (repr Ca) | Ca in rcosets 'C_G[x | to] G].
+Proof.
+rewrite -amove_orbit -imset_comp /=; apply/setP=> z.
+by apply/idP/imsetP=> [xGz | [y xGy ->]]; first exists z; rewrite /= ?amoveK.
+Qed.
+
+Lemma act_reprK :
+ {in rcosets 'C_G[x | to] G, cancel (to x \o repr) (amove to G x)}.
+Proof.
+move=> _ /rcosetsP[a Ga ->] /=; rewrite amove_act ?rcoset_repr //.
+rewrite -[G :&: _]/(gval _); case: repr_rcosetP => b /setIP[Gb _].
+exact: groupM.
+Qed.
+
+End OrbitStabilizer.
+
+Lemma card_orbit_in G x : G \subset D -> #|orbit to G x| = #|G : 'C_G[x | to]|.
+Proof.
+move=> sGD; rewrite orbit_stabilizer 1?card_in_imset //.
+exact: can_in_inj (act_reprK _).
+Qed.
+
+Lemma card_orbit_in_stab G x :
+ G \subset D -> (#|orbit to G x| * #|'C_G[x | to]|)%N = #|G|.
+Proof. by move=> sGD; rewrite mulnC card_orbit_in ?Lagrange ?subsetIl. Qed.
+
+Lemma acts_sum_card_orbit G S :
+ [acts G, on S | to] -> \sum_(T in orbit to G @: S) #|T| = #|S|.
+Proof. by move/orbit_partition/card_partition. Qed.
+
+Lemma astab_setact_in S a : a \in D -> 'C(to^* S a | to) = 'C(S | to) :^ a.
+Proof.
+move=> Da; apply/setP=> b; rewrite mem_conjg !inE -mem_conjg conjGid //.
+apply: andb_id2l => Db; rewrite sub_imset_pre; apply: eq_subset_r => x.
+by rewrite !inE !actMin ?groupM ?groupV // invgK (canF_eq (actKVin Da)).
+Qed.
+
+Lemma astab1_act_in x a : a \in D -> 'C[to x a | to] = 'C[x | to] :^ a.
+Proof. by move=> Da; rewrite -astab_setact_in // /setact imset_set1. Qed.
+
+Theorem Frobenius_Cauchy G S : [acts G, on S | to] ->
+ \sum_(a in G) #|'Fix_(S | to)[a]| = (#|orbit to G @: S| * #|G|)%N.
+Proof.
+move=> GactS; have sGD := acts_dom GactS.
+transitivity (\sum_(a in G) \sum_(x in 'Fix_(S | to)[a]) 1%N).
+ by apply: eq_bigr => a _; rewrite -sum1_card.
+rewrite (exchange_big_dep (mem S)) /= => [|a x _]; last by case/setIP.
+rewrite (set_partition_big _ (orbit_partition GactS)) -sum_nat_const /=.
+apply: eq_bigr => _ /imsetP[x Sx ->].
+rewrite -(card_orbit_in_stab x sGD) -sum_nat_const.
+apply: eq_bigr => y; rewrite orbit_in_sym // => /imsetP[a Ga defx].
+rewrite defx astab1_act_in ?(subsetP sGD) //.
+rewrite -{2}(conjGid Ga) -conjIg cardJg -sum1_card setIA (setIidPl sGD).
+by apply: eq_bigl => b; rewrite !(sub1set, inE) -(acts_act GactS Ga) -defx Sx.
+Qed.
+
+Lemma atrans_dvd_index_in G S :
+ G \subset D -> [transitive G, on S | to] -> #|S| %| #|G : 'C_G(S | to)|.
+Proof.
+move=> sGD /imsetP[x Sx {1}->]; rewrite card_orbit_in //.
+by rewrite indexgS // setIS // astabS // sub1set.
+Qed.
+
+Lemma atrans_dvd_in G S :
+ G \subset D -> [transitive G, on S | to] -> #|S| %| #|G|.
+Proof.
+move=> sGD transG; apply: dvdn_trans (atrans_dvd_index_in sGD transG) _.
+exact: dvdn_indexg.
+Qed.
+
+Lemma atransPin G S :
+ G \subset D -> [transitive G, on S | to] ->
+ forall x, x \in S -> orbit to G x = S.
+Proof. by move=> sGD /imsetP[y _ ->] x; exact: orbit_in_transl. Qed.
+
+Lemma atransP2in G S :
+ G \subset D -> [transitive G, on S | to] ->
+ {in S &, forall x y, exists2 a, a \in G & y = to x a}.
+Proof. by move=> sGD transG x y /(atransPin sGD transG) <- /imsetP. Qed.
+
+Lemma atrans_acts_in G S :
+ G \subset D -> [transitive G, on S | to] -> [acts G, on S | to].
+Proof.
+move=> sGD transG; apply/subsetP=> a Ga; rewrite !inE (subsetP sGD) //.
+by apply/subsetP=> x /(atransPin sGD transG) <-; rewrite inE mem_imset.
+Qed.
+
+Lemma subgroup_transitivePin G H S x :
+ x \in S -> H \subset G -> G \subset D -> [transitive G, on S | to] ->
+ reflect ('C_G[x | to] * H = G) [transitive H, on S | to].
+Proof.
+move=> Sx sHG sGD trG; have sHD := subset_trans sHG sGD.
+apply: (iffP idP) => [trH | defG].
+ rewrite group_modr //; apply/setIidPl/subsetP=> a Ga.
+ have Sxa: to x a \in S by rewrite (acts_act (atrans_acts_in sGD trG)).
+ have [b Hb xab]:= atransP2in sHD trH Sxa Sx.
+ have Da := subsetP sGD a Ga; have Db := subsetP sHD b Hb.
+ rewrite -(mulgK b a) mem_mulg ?groupV // !inE groupM //= sub1set inE.
+ by rewrite actMin -?xab.
+apply/imsetP; exists x => //; apply/setP=> y; rewrite -(atransPin sGD trG Sx).
+apply/imsetP/imsetP=> [] [a]; last by exists a; first exact: (subsetP sHG).
+rewrite -defG => /imset2P[c b /setIP[_ cxc] Hb ->] ->.
+exists b; rewrite ?actMin ?(astab_dom cxc) ?(subsetP sHD) //.
+by rewrite (astab_act cxc) ?inE.
+Qed.
+
+End PartialAction.
+
+Arguments Scope orbit_transversal
+ [_ group_scope _ action_scope group_scope group_scope].
+Implicit Arguments orbit1P [aT D rT to G x].
+Implicit Arguments contra_orbit [aT D rT x y].
+Prenex Implicits orbit1P.
+
+Notation "''C' ( S | to )" := (astab_group to S) : Group_scope.
+Notation "''C_' A ( S | to )" := (setI_group A 'C(S | to)) : Group_scope.
+Notation "''C_' ( A ) ( S | to )" := (setI_group A 'C(S | to))
+ (only parsing) : Group_scope.
+Notation "''C' [ x | to ]" := (astab_group to [set x%g]) : Group_scope.
+Notation "''C_' A [ x | to ]" := (setI_group A 'C[x | to]) : Group_scope.
+Notation "''C_' ( A ) [ x | to ]" := (setI_group A 'C[x | to])
+ (only parsing) : Group_scope.
+Notation "''N' ( S | to )" := (astabs_group to S) : Group_scope.
+Notation "''N_' A ( S | to )" := (setI_group A 'N(S | to)) : Group_scope.
+
+Section TotalActions.
+(* These lemmas are only established for total actions (domain = [set: rT]) *)
+
+Variable (aT : finGroupType) (rT : finType).
+
+Variable to : {action aT &-> rT}.
+
+Implicit Types (a b : aT) (x y z : rT) (A B : {set aT}) (G H : {group aT}).
+Implicit Type S : {set rT}.
+
+Lemma actM x a b : to x (a * b) = to (to x a) b.
+Proof. by rewrite actMin ?inE. Qed.
+
+Lemma actK : right_loop invg to.
+Proof. by move=> a; apply: actKin; rewrite inE. Qed.
+
+Lemma actKV : rev_right_loop invg to.
+Proof. by move=> a; apply: actKVin; rewrite inE. Qed.
+
+Lemma actX x a n : to x (a ^+ n) = iter n (to^~ a) x.
+Proof. by elim: n => [|n /= <-]; rewrite ?act1 // -actM expgSr. Qed.
+
+Lemma actCJ a b x : to (to x a) b = to (to x b) (a ^ b).
+Proof. by rewrite !actM actK. Qed.
+
+Lemma actCJV a b x : to (to x a) b = to (to x (b ^ a^-1)) a.
+Proof. by rewrite (actCJ _ a) conjgKV. Qed.
+
+Lemma orbit_sym G x y : (y \in orbit to G x) = (x \in orbit to G y).
+Proof. by apply: orbit_in_sym; exact: subsetT. Qed.
+
+Lemma orbit_trans G x y z :
+ y \in orbit to G x -> z \in orbit to G y -> z \in orbit to G x.
+Proof. by apply: orbit_in_trans; exact: subsetT. Qed.
+
+Lemma orbit_transl G x y : y \in orbit to G x -> orbit to G y = orbit to G x.
+Proof.
+move=> Gxy; apply/setP=> z; apply/idP/idP; apply: orbit_trans => //.
+by rewrite orbit_sym.
+Qed.
+
+Lemma orbit_transr G x y z :
+ y \in orbit to G x -> (y \in orbit to G z) = (x \in orbit to G z).
+Proof. by move=> Gxy; rewrite orbit_sym (orbit_transl Gxy) orbit_sym. Qed.
+
+Lemma orbit_act G a x: a \in G -> orbit to G (to x a) = orbit to G x.
+Proof. by move/mem_orbit/orbit_transl; exact. Qed.
+
+Lemma orbit_actr G a x y :
+ a \in G -> (to y a \in orbit to G x) = (y \in orbit to G x).
+Proof. by move/mem_orbit/orbit_transr; exact. Qed.
+
+Lemma orbit_eq_mem G x y :
+ (orbit to G x == orbit to G y) = (x \in orbit to G y).
+Proof. by apply/eqP/idP=> [<-|]; [exact: orbit_refl | exact: orbit_transl]. Qed.
+
+Lemma orbit_inv A x y : (y \in orbit to A^-1 x) = (x \in orbit to A y).
+Proof. by rewrite orbit_inv_in ?subsetT. Qed.
+
+Lemma orbit_lcoset A a x : orbit to (a *: A) x = orbit to A (to x a).
+Proof. by rewrite orbit_lcoset_in ?subsetT ?inE. Qed.
+
+Lemma orbit_rcoset A a x y :
+ (to y a \in orbit to (A :* a) x) = (y \in orbit to A x).
+Proof. by rewrite orbit_rcoset_in ?subsetT ?inE. Qed.
+
+Lemma orbit_conjsg A a x y :
+ (to y a \in orbit to (A :^ a) (to x a)) = (y \in orbit to A x).
+Proof. by rewrite orbit_conjsg_in ?subsetT ?inE. Qed.
+
+Lemma astabP S a : reflect (forall x, x \in S -> to x a = x) (a \in 'C(S | to)).
+Proof.
+apply: (iffP idP) => [cSa x|cSa]; first exact: astab_act.
+by rewrite !inE; apply/subsetP=> x Sx; rewrite inE cSa.
+Qed.
+
+Lemma astab1P x a : reflect (to x a = x) (a \in 'C[x | to]).
+Proof. by rewrite !inE sub1set inE; exact: eqP. Qed.
+
+Lemma sub_astab1 A x : (A \subset 'C[x | to]) = (x \in 'Fix_to(A)).
+Proof. by rewrite sub_astab1_in ?subsetT. Qed.
+
+Lemma astabC A S : (A \subset 'C(S | to)) = (S \subset 'Fix_to(A)).
+Proof. by rewrite astabCin ?subsetT. Qed.
+
+Lemma afix_cycle a : 'Fix_to(<[a]>) = 'Fix_to[a].
+Proof. by rewrite afix_cycle_in ?inE. Qed.
+
+Lemma afix_gen A : 'Fix_to(<<A>>) = 'Fix_to(A).
+Proof. by rewrite afix_gen_in ?subsetT. Qed.
+
+Lemma afixM G H : 'Fix_to(G * H) = 'Fix_to(G) :&: 'Fix_to(H).
+Proof. by rewrite afixMin ?subsetT. Qed.
+
+Lemma astabsP S a :
+ reflect (forall x, (to x a \in S) = (x \in S)) (a \in 'N(S | to)).
+Proof.
+apply: (iffP idP) => [nSa x|nSa]; first exact: astabs_act.
+by rewrite !inE; apply/subsetP=> x; rewrite inE nSa.
+Qed.
+
+Lemma card_orbit G x : #|orbit to G x| = #|G : 'C_G[x | to]|.
+Proof. by rewrite card_orbit_in ?subsetT. Qed.
+
+Lemma dvdn_orbit G x : #|orbit to G x| %| #|G|.
+Proof. by rewrite card_orbit dvdn_indexg. Qed.
+
+Lemma card_orbit_stab G x : (#|orbit to G x| * #|'C_G[x | to]|)%N = #|G|.
+Proof. by rewrite mulnC card_orbit Lagrange ?subsetIl. Qed.
+
+Lemma actsP A S : reflect {acts A, on S | to} [acts A, on S | to].
+Proof.
+apply: (iffP idP) => [nSA x|nSA]; first exact: acts_act.
+by apply/subsetP=> a Aa; rewrite !inE; apply/subsetP=> x; rewrite inE nSA.
+Qed.
+Implicit Arguments actsP [A S].
+
+Lemma setact_orbit A x b : to^* (orbit to A x) b = orbit to (A :^ b) (to x b).
+Proof.
+apply/setP=> y; apply/idP/idP=> /imsetP[_ /imsetP[a Aa ->] ->{y}].
+ by rewrite actCJ mem_orbit ?memJ_conjg.
+by rewrite -actCJ mem_setact ?mem_orbit.
+Qed.
+
+Lemma astab_setact S a : 'C(to^* S a | to) = 'C(S | to) :^ a.
+Proof.
+apply/setP=> b; rewrite mem_conjg.
+apply/astabP/astabP=> stab x => [Sx|].
+ by rewrite conjgE invgK !actM stab ?actK //; apply/imsetP; exists x.
+by case/imsetP=> y Sy ->{x}; rewrite -actM conjgCV actM stab.
+Qed.
+
+Lemma astab1_act x a : 'C[to x a | to] = 'C[x | to] :^ a.
+Proof. by rewrite -astab_setact /setact imset_set1. Qed.
+
+Lemma atransP G S : [transitive G, on S | to] ->
+ forall x, x \in S -> orbit to G x = S.
+Proof. by case/imsetP=> x _ -> y; exact: orbit_transl. Qed.
+
+Lemma atransP2 G S : [transitive G, on S | to] ->
+ {in S &, forall x y, exists2 a, a \in G & y = to x a}.
+Proof. by move=> GtrS x y /(atransP GtrS) <- /imsetP. Qed.
+
+Lemma atrans_acts G S : [transitive G, on S | to] -> [acts G, on S | to].
+Proof.
+move=> GtrS; apply/subsetP=> a Ga; rewrite !inE.
+by apply/subsetP=> x /(atransP GtrS) <-; rewrite inE mem_imset.
+Qed.
+
+Lemma atrans_supgroup G H S :
+ G \subset H -> [transitive G, on S | to] ->
+ [transitive H, on S | to] = [acts H, on S | to].
+Proof.
+move=> sGH trG; apply/idP/idP=> [|actH]; first exact: atrans_acts.
+case/imsetP: trG => x Sx defS; apply/imsetP; exists x => //.
+by apply/eqP; rewrite eqEsubset acts_sub_orbit ?Sx // defS imsetS.
+Qed.
+
+Lemma atrans_acts_card G S :
+ [transitive G, on S | to] =
+ [acts G, on S | to] && (#|orbit to G @: S| == 1%N).
+Proof.
+apply/idP/andP=> [GtrS | [nSG]].
+ split; first exact: atrans_acts.
+ rewrite ((_ @: S =P [set S]) _) ?cards1 // eqEsubset sub1set.
+ apply/andP; split=> //; apply/subsetP=> _ /imsetP[x Sx ->].
+ by rewrite inE (atransP GtrS).
+rewrite eqn_leq andbC lt0n => /andP[/existsP[X /imsetP[x Sx X_Gx]]].
+rewrite (cardD1 X) {X}X_Gx mem_imset // ltnS leqn0 => /eqP GtrS.
+apply/imsetP; exists x => //; apply/eqP.
+rewrite eqEsubset acts_sub_orbit // Sx andbT.
+apply/subsetP=> y Sy; have:= card0_eq GtrS (orbit to G y).
+rewrite !inE /= mem_imset // andbT => /eqP <-; exact: orbit_refl.
+Qed.
+
+Lemma atrans_dvd G S : [transitive G, on S | to] -> #|S| %| #|G|.
+Proof. by case/imsetP=> x _ ->; exact: dvdn_orbit. Qed.
+
+(* Aschbacher 5.2 *)
+Lemma acts_fix_norm A B : A \subset 'N(B) -> [acts A, on 'Fix_to(B) | to].
+Proof.
+move=> nAB; have:= acts_subnorm_fix to B; rewrite !setTI.
+exact: subset_trans.
+Qed.
+
+Lemma faithfulP A S :
+ reflect (forall a, a \in A -> {in S, to^~ a =1 id} -> a = 1)
+ [faithful A, on S | to].
+Proof.
+apply: (iffP subsetP) => [Cto1 a Aa Ca | Cto1 a].
+ apply/set1P; rewrite Cto1 // inE Aa; exact/astabP.
+case/setIP=> Aa /astabP Ca; apply/set1P; exact: Cto1.
+Qed.
+
+(* This is the first part of Aschbacher (5.7) *)
+Lemma astab_trans_gcore G S u :
+ [transitive G, on S | to] -> u \in S -> 'C(S | to) = gcore 'C[u | to] G.
+Proof.
+move=> transG Su; apply/eqP; rewrite eqEsubset.
+rewrite gcore_max ?astabS ?sub1set //=; last first.
+ exact: subset_trans (atrans_acts transG) (astab_norm _ _).
+apply/subsetP=> x cSx; apply/astabP=> uy.
+case/(atransP2 transG Su) => y Gy ->{uy}.
+by apply/astab1P; rewrite astab1_act (bigcapP cSx).
+Qed.
+
+(* Aschbacher 5.20 *)
+Theorem subgroup_transitiveP G H S x :
+ x \in S -> H \subset G -> [transitive G, on S | to] ->
+ reflect ('C_G[x | to] * H = G) [transitive H, on S | to].
+Proof. by move=> Sx sHG; exact: subgroup_transitivePin (subsetT G). Qed.
+
+(* Aschbacher 5.21 *)
+Lemma trans_subnorm_fixP x G H S :
+ let C := 'C_G[x | to] in let T := 'Fix_(S | to)(H) in
+ [transitive G, on S | to] -> x \in S -> H \subset C ->
+ reflect ((H :^: G) ::&: C = H :^: C) [transitive 'N_G(H), on T | to].
+Proof.
+move=> C T trGS Sx sHC; have actGS := acts_act (atrans_acts trGS).
+have:= sHC; rewrite subsetI sub_astab1 => /andP[sHG cHx].
+have Tx: x \in T by rewrite inE Sx.
+apply: (iffP idP) => [trN | trC].
+ apply/setP=> Ha; apply/setIdP/imsetP=> [[]|[a Ca ->{Ha}]]; last first.
+ by rewrite conj_subG //; case/setIP: Ca => Ga _; rewrite mem_imset.
+ case/imsetP=> a Ga ->{Ha}; rewrite subsetI !sub_conjg => /andP[_ sHCa].
+ have Txa: to x a^-1 \in T.
+ by rewrite inE -sub_astab1 astab1_act actGS ?Sx ?groupV.
+ have [b] := atransP2 trN Tx Txa; case/setIP=> Gb nHb cxba.
+ exists (b * a); last by rewrite conjsgM (normP nHb).
+ by rewrite inE groupM //; apply/astab1P; rewrite actM -cxba actKV.
+apply/imsetP; exists x => //; apply/setP=> y; apply/idP/idP=> [Ty|].
+ have [Sy cHy]:= setIP Ty; have [a Ga defy] := atransP2 trGS Sx Sy.
+ have: H :^ a^-1 \in H :^: C.
+ rewrite -trC inE subsetI mem_imset 1?conj_subG ?groupV // sub_conjgV.
+ by rewrite -astab1_act -defy sub_astab1.
+ case/imsetP=> b /setIP[Gb /astab1P cxb] defHb.
+ rewrite defy -{1}cxb -actM mem_orbit // inE groupM //.
+ by apply/normP; rewrite conjsgM -defHb conjsgKV.
+case/imsetP=> a /setIP[Ga nHa] ->{y}.
+by rewrite inE actGS // Sx (acts_act (acts_fix_norm _) nHa).
+Qed.
+
+End TotalActions.
+
+Implicit Arguments astabP [aT rT to S a].
+Implicit Arguments astab1P [aT rT to x a].
+Implicit Arguments astabsP [aT rT to S a].
+Implicit Arguments atransP [aT rT to G S].
+Implicit Arguments actsP [aT rT to A S].
+Implicit Arguments faithfulP [aT rT to A S].
+Prenex Implicits astabP astab1P astabsP atransP actsP faithfulP.
+
+Section Restrict.
+
+Variables (aT : finGroupType) (D : {set aT}) (rT : Type).
+Variables (to : action D rT) (A : {set aT}).
+
+Definition ract of A \subset D := act to.
+
+Variable sAD : A \subset D.
+
+Lemma ract_is_action : is_action A (ract sAD).
+Proof.
+rewrite /ract; case: to => f [injf fM].
+split=> // x; exact: (sub_in2 (subsetP sAD)).
+Qed.
+
+Canonical raction := Action ract_is_action.
+
+Lemma ractE : raction =1 to. Proof. by []. Qed.
+
+(* Other properties of raction need rT : finType; we defer them *)
+(* until after the definition of actperm. *)
+
+End Restrict.
+
+Notation "to \ sAD" := (raction to sAD) (at level 50) : action_scope.
+
+Section ActBy.
+
+Variables (aT : finGroupType) (D : {set aT}) (rT : finType).
+
+Definition actby_cond (A : {set aT}) R (to : action D rT) : Prop :=
+ [acts A, on R | to].
+
+Definition actby A R to of actby_cond A R to :=
+ fun x a => if (x \in R) && (a \in A) then to x a else x.
+
+Variables (A : {group aT}) (R : {set rT}) (to : action D rT).
+Hypothesis nRA : actby_cond A R to.
+
+Lemma actby_is_action : is_action A (actby nRA).
+Proof.
+rewrite /actby; split=> [a x y | x a b Aa Ab /=]; last first.
+ rewrite Aa Ab groupM // !andbT actMin ?(subsetP (acts_dom nRA)) //.
+ by case Rx: (x \in R); rewrite ?(acts_act nRA) ?Rx.
+case Aa: (a \in A); rewrite ?andbF ?andbT //.
+case Rx: (x \in R); case Ry: (y \in R) => // eqxy; first exact: act_inj eqxy.
+ by rewrite -eqxy (acts_act nRA Aa) Rx in Ry.
+by rewrite eqxy (acts_act nRA Aa) Ry in Rx.
+Qed.
+
+Canonical action_by := Action actby_is_action.
+Local Notation "<[nRA]>" := action_by : action_scope.
+
+Lemma actbyE x a : x \in R -> a \in A -> <[nRA]>%act x a = to x a.
+Proof. by rewrite /= /actby => -> ->. Qed.
+
+Lemma afix_actby B : 'Fix_<[nRA]>(B) = ~: R :|: 'Fix_to(A :&: B).
+Proof.
+apply/setP=> x; rewrite !inE /= /actby.
+case: (x \in R); last by apply/subsetP=> a _; rewrite !inE.
+apply/subsetP/subsetP=> [cBx a | cABx a Ba]; rewrite !inE.
+ by case/andP=> Aa /cBx; rewrite inE Aa.
+by case: ifP => //= Aa; have:= cABx a; rewrite !inE Aa => ->.
+Qed.
+
+Lemma astab_actby S : 'C(S | <[nRA]>) = 'C_A(R :&: S | to).
+Proof.
+apply/setP=> a; rewrite setIA (setIidPl (acts_dom nRA)) !inE.
+case Aa: (a \in A) => //=; apply/subsetP/subsetP=> cRSa x => [|Sx].
+ by case/setIP=> Rx /cRSa; rewrite !inE actbyE.
+by have:= cRSa x; rewrite !inE /= /actby Aa Sx; case: (x \in R) => //; apply.
+Qed.
+
+Lemma astabs_actby S : 'N(S | <[nRA]>) = 'N_A(R :&: S | to).
+Proof.
+apply/setP=> a; rewrite setIA (setIidPl (acts_dom nRA)) !inE.
+case Aa: (a \in A) => //=; apply/subsetP/subsetP=> nRSa x => [|Sx].
+ by case/setIP=> Rx /nRSa; rewrite !inE actbyE ?(acts_act nRA) ?Rx.
+have:= nRSa x; rewrite !inE /= /actby Aa Sx ?(acts_act nRA) //.
+by case: (x \in R) => //; apply.
+Qed.
+
+Lemma acts_actby (B : {set aT}) S :
+ [acts B, on S | <[nRA]>] = (B \subset A) && [acts B, on R :&: S | to].
+Proof. by rewrite astabs_actby subsetI. Qed.
+
+End ActBy.
+
+Notation "<[ nRA ] >" := (action_by nRA) : action_scope.
+
+Section SubAction.
+
+Variables (aT : finGroupType) (D : {group aT}).
+Variables (rT : finType) (sP : pred rT) (sT : subFinType sP) (to : action D rT).
+Implicit Type A : {set aT}.
+Implicit Type u : sT.
+Implicit Type S : {set sT}.
+
+Definition subact_dom := 'N([set x | sP x] | to).
+Canonical subact_dom_group := [group of subact_dom].
+
+Implicit Type Na : {a | a \in subact_dom}.
+Lemma sub_act_proof u Na : sP (to (val u) (val Na)).
+Proof. by case: Na => a /= /(astabs_act (val u)); rewrite !inE valP. Qed.
+
+Definition subact u a :=
+ if insub a is Some Na then Sub _ (sub_act_proof u Na) else u.
+
+Lemma val_subact u a :
+ val (subact u a) = if a \in subact_dom then to (val u) a else val u.
+Proof.
+by rewrite /subact -if_neg; case: insubP => [Na|] -> //=; rewrite SubK => ->.
+Qed.
+
+Lemma subact_is_action : is_action subact_dom subact.
+Proof.
+split=> [a u v eq_uv | u a b Na Nb]; apply: val_inj.
+ move/(congr1 val): eq_uv; rewrite !val_subact.
+ by case: (a \in _); first move/act_inj.
+have Da := astabs_dom Na; have Db := astabs_dom Nb.
+by rewrite !val_subact Na Nb groupM ?actMin.
+Qed.
+
+Canonical subaction := Action subact_is_action.
+
+Lemma astab_subact S : 'C(S | subaction) = subact_dom :&: 'C(val @: S | to).
+Proof.
+apply/setP=> a; rewrite inE in_setI; apply: andb_id2l => sDa.
+have [Da _] := setIP sDa; rewrite !inE Da.
+apply/subsetP/subsetP=> [cSa _ /imsetP[x Sx ->] | cSa x Sx]; rewrite !inE.
+ by have:= cSa x Sx; rewrite inE -val_eqE val_subact sDa.
+by have:= cSa _ (mem_imset val Sx); rewrite inE -val_eqE val_subact sDa.
+Qed.
+
+Lemma astabs_subact S : 'N(S | subaction) = subact_dom :&: 'N(val @: S | to).
+Proof.
+apply/setP=> a; rewrite inE in_setI; apply: andb_id2l => sDa.
+have [Da _] := setIP sDa; rewrite !inE Da.
+apply/subsetP/subsetP=> [nSa _ /imsetP[x Sx ->] | nSa x Sx]; rewrite !inE.
+ by have:= nSa x Sx; rewrite inE => /(mem_imset val); rewrite val_subact sDa.
+have:= nSa _ (mem_imset val Sx); rewrite inE => /imsetP[y Sy def_y].
+by rewrite ((_ a =P y) _) // -val_eqE val_subact sDa def_y.
+Qed.
+
+Lemma afix_subact A :
+ A \subset subact_dom -> 'Fix_subaction(A) = val @^-1: 'Fix_to(A).
+Proof.
+move/subsetP=> sAD; apply/setP=> u.
+rewrite !inE !(sameP setIidPl eqP); congr (_ == A).
+apply/setP=> a; rewrite !inE; apply: andb_id2l => Aa.
+by rewrite -val_eqE val_subact sAD.
+
+Qed.
+
+End SubAction.
+
+Notation "to ^?" := (subaction _ to)
+ (at level 2, format "to ^?") : action_scope.
+
+Section QuotientAction.
+
+Variables (aT : finGroupType) (D : {group aT}) (rT : finGroupType).
+Variables (to : action D rT) (H : {group rT}).
+
+Definition qact_dom := 'N(rcosets H 'N(H) | to^*).
+Canonical qact_dom_group := [group of qact_dom].
+
+Local Notation subdom := (subact_dom (coset_range H) to^*).
+Fact qact_subdomE : subdom = qact_dom.
+Proof. by congr 'N(_|_); apply/setP=> Hx; rewrite !inE genGid. Qed.
+Lemma qact_proof : qact_dom \subset subdom.
+Proof. by rewrite qact_subdomE. Qed.
+
+Definition qact : coset_of H -> aT -> coset_of H := act (to^*^? \ qact_proof).
+
+Canonical quotient_action := [action of qact].
+
+Lemma acts_qact_dom : [acts qact_dom, on 'N(H) | to].
+Proof.
+apply/subsetP=> a nNa; rewrite !inE (astabs_dom nNa); apply/subsetP=> x Nx.
+have: H :* x \in rcosets H 'N(H) by rewrite -rcosetE mem_imset.
+rewrite inE -(astabs_act _ nNa) => /rcosetsP[y Ny defHy].
+have: to x a \in H :* y by rewrite -defHy (mem_imset (to^~a)) ?rcoset_refl.
+by apply: subsetP; rewrite mul_subG ?sub1set ?normG.
+Qed.
+
+Lemma qactEcond x a :
+ x \in 'N(H) ->
+ quotient_action (coset H x) a =
+ (if a \in qact_dom then coset H (to x a) else coset H x).
+Proof.
+move=> Nx; apply: val_inj; rewrite val_subact //= qact_subdomE.
+have: H :* x \in rcosets H 'N(H) by rewrite -rcosetE mem_imset.
+case nNa: (a \in _); rewrite // -(astabs_act _ nNa).
+rewrite !val_coset ?(acts_act acts_qact_dom nNa) //=.
+case/rcosetsP=> y Ny defHy; rewrite defHy; apply: rcoset_transl.
+by rewrite rcoset_sym -defHy (mem_imset (_^~_)) ?rcoset_refl.
+Qed.
+
+Lemma qactE x a :
+ x \in 'N(H) -> a \in qact_dom ->
+ quotient_action (coset H x) a = coset H (to x a).
+Proof. by move=> Nx nNa; rewrite qactEcond ?nNa. Qed.
+
+Lemma acts_quotient (A : {set aT}) (B : {set rT}) :
+ A \subset 'N_qact_dom(B | to) -> [acts A, on B / H | quotient_action].
+Proof.
+move=> nBA; apply: subset_trans {A}nBA _; apply/subsetP=> a /setIP[dHa nBa].
+rewrite inE dHa inE; apply/subsetP=> _ /morphimP[x nHx Bx ->].
+rewrite inE /= qactE //.
+by rewrite mem_morphim ?(acts_act acts_qact_dom) ?(astabs_act _ nBa).
+Qed.
+
+Lemma astabs_quotient (G : {group rT}) :
+ H <| G -> 'N(G / H | quotient_action) = 'N_qact_dom(G | to).
+Proof.
+move=> nsHG; have [_ nHG] := andP nsHG.
+apply/eqP; rewrite eqEsubset acts_quotient // andbT.
+apply/subsetP=> a nGa; have dHa := astabs_dom nGa; have [Da _]:= setIdP dHa.
+rewrite inE dHa 2!inE Da; apply/subsetP=> x Gx; have nHx := subsetP nHG x Gx.
+rewrite -(quotientGK nsHG) 2!inE (acts_act acts_qact_dom) ?nHx //= inE.
+by rewrite -qactE // (astabs_act _ nGa) mem_morphim.
+Qed.
+
+End QuotientAction.
+
+Notation "to / H" := (quotient_action to H) : action_scope.
+
+Section ModAction.
+
+Variables (aT : finGroupType) (D : {group aT}) (rT : finType).
+Variable to : action D rT.
+Implicit Types (G : {group aT}) (S : {set rT}).
+
+Section GenericMod.
+
+Variable H : {group aT}.
+
+Local Notation dom := 'N_D(H).
+Local Notation range := 'Fix_to(D :&: H).
+Let acts_dom : {acts dom, on range | to} := acts_act (acts_subnorm_fix to H).
+
+Definition modact x (Ha : coset_of H) :=
+ if x \in range then to x (repr (D :&: Ha)) else x.
+
+Lemma modactEcond x a :
+ a \in dom -> modact x (coset H a) = (if x \in range then to x a else x).
+Proof.
+case/setIP=> Da Na; case: ifP => Cx; rewrite /modact Cx //.
+rewrite val_coset // -group_modr ?sub1set //.
+case: (repr _) / (repr_rcosetP (D :&: H) a) => a' Ha'.
+by rewrite actMin ?(afixP Cx _ Ha') //; case/setIP: Ha'.
+Qed.
+
+Lemma modactE x a :
+ a \in D -> a \in 'N(H) -> x \in range -> modact x (coset H a) = to x a.
+Proof. by move=> Da Na Rx; rewrite modactEcond ?Rx // inE Da. Qed.
+
+Lemma modact_is_action : is_action (D / H) modact.
+Proof.
+split=> [Ha x y | x Ha Hb]; last first.
+ case/morphimP=> a Na Da ->{Ha}; case/morphimP=> b Nb Db ->{Hb}.
+ rewrite -morphM //= !modactEcond // ?groupM ?(introT setIP _) //.
+ by case: ifP => Cx; rewrite ?(acts_dom, Cx, actMin, introT setIP _).
+case: (set_0Vmem (D :&: Ha)) => [Da0 | [a /setIP[Da NHa]]].
+ by rewrite /modact Da0 repr_set0 !act1 !if_same.
+have Na := subsetP (coset_norm _) _ NHa.
+have NDa: a \in 'N_D(H) by rewrite inE Da.
+rewrite -(coset_mem NHa) !modactEcond //.
+do 2![case: ifP]=> Cy Cx // eqxy; first exact: act_inj eqxy.
+ by rewrite -eqxy acts_dom ?Cx in Cy.
+by rewrite eqxy acts_dom ?Cy in Cx.
+Qed.
+
+Canonical mod_action := Action modact_is_action.
+
+Section Stabilizers.
+
+Variable S : {set rT}.
+Hypothesis cSH : H \subset 'C(S | to).
+
+Let fixSH : S \subset 'Fix_to(D :&: H).
+Proof. by rewrite -astabCin ?subsetIl // subIset ?cSH ?orbT. Qed.
+
+Lemma astabs_mod : 'N(S | mod_action) = 'N(S | to) / H.
+Proof.
+apply/setP=> Ha; apply/idP/morphimP=> [nSa | [a nHa nSa ->]].
+ case/morphimP: (astabs_dom nSa) => a nHa Da defHa.
+ exists a => //; rewrite !inE Da; apply/subsetP=> x Sx; rewrite !inE.
+ by have:= Sx; rewrite -(astabs_act x nSa) defHa /= modactE ?(subsetP fixSH).
+have Da := astabs_dom nSa; rewrite !inE mem_quotient //; apply/subsetP=> x Sx.
+by rewrite !inE /= modactE ?(astabs_act x nSa) ?(subsetP fixSH).
+Qed.
+
+Lemma astab_mod : 'C(S | mod_action) = 'C(S | to) / H.
+Proof.
+apply/setP=> Ha; apply/idP/morphimP=> [cSa | [a nHa cSa ->]].
+ case/morphimP: (astab_dom cSa) => a nHa Da defHa.
+ exists a => //; rewrite !inE Da; apply/subsetP=> x Sx; rewrite !inE.
+ by rewrite -{2}[x](astab_act cSa) // defHa /= modactE ?(subsetP fixSH).
+have Da := astab_dom cSa; rewrite !inE mem_quotient //; apply/subsetP=> x Sx.
+by rewrite !inE /= modactE ?(astab_act cSa) ?(subsetP fixSH).
+Qed.
+
+End Stabilizers.
+
+Lemma afix_mod G S :
+ H \subset 'C(S | to) -> G \subset 'N_D(H) ->
+ 'Fix_(S | mod_action)(G / H) = 'Fix_(S | to)(G).
+Proof.
+move=> cSH /subsetIP[sGD nHG].
+apply/eqP; rewrite eqEsubset !subsetI !subsetIl /= -!astabCin ?quotientS //.
+have cfixH F: H \subset 'C(S :&: F | to).
+ by rewrite (subset_trans cSH) // astabS ?subsetIl.
+rewrite andbC astab_mod ?quotientS //=; last by rewrite astabCin ?subsetIr.
+by rewrite -(quotientSGK nHG) //= -astab_mod // astabCin ?quotientS ?subsetIr.
+Qed.
+
+End GenericMod.
+
+Lemma modact_faithful G S :
+ [faithful G / 'C_G(S | to), on S | mod_action 'C_G(S | to)].
+Proof.
+rewrite /faithful astab_mod ?subsetIr //=.
+by rewrite -quotientIG ?subsetIr ?trivg_quotient.
+Qed.
+
+End ModAction.
+
+Notation "to %% H" := (mod_action to H) : action_scope.
+
+Section ActPerm.
+(* Morphism to permutations induced by an action. *)
+
+Variables (aT : finGroupType) (D : {set aT}) (rT : finType).
+Variable to : action D rT.
+
+Definition actperm a := perm (act_inj to a).
+
+Lemma actpermM : {in D &, {morph actperm : a b / a * b}}.
+Proof. by move=> a b Da Db; apply/permP=> x; rewrite permM !permE actMin. Qed.
+
+Canonical actperm_morphism := Morphism actpermM.
+
+Lemma actpermE a x : actperm a x = to x a.
+Proof. by rewrite permE. Qed.
+
+Lemma actpermK x a : aperm x (actperm a) = to x a.
+Proof. exact: actpermE. Qed.
+
+Lemma ker_actperm : 'ker actperm = 'C(setT | to).
+Proof.
+congr (_ :&: _); apply/setP=> a; rewrite !inE /=.
+apply/eqP/subsetP=> [a1 x _ | a1]; first by rewrite inE -actpermE a1 perm1.
+by apply/permP=> x; apply/eqP; have:= a1 x; rewrite !inE actpermE perm1 => ->.
+Qed.
+
+End ActPerm.
+
+Section RestrictActionTheory.
+
+Variables (aT : finGroupType) (D : {set aT}) (rT : finType).
+Variables (to : action D rT).
+
+Lemma faithful_isom (A : {group aT}) S (nSA : actby_cond A S to) :
+ [faithful A, on S | to] -> isom A (actperm <[nSA]> @* A) (actperm <[nSA]>).
+Proof.
+by move=> ffulAS; apply/isomP; rewrite ker_actperm astab_actby setIT.
+Qed.
+
+Variables (A : {set aT}) (sAD : A \subset D).
+
+Lemma ractpermE : actperm (to \ sAD) =1 actperm to.
+Proof. by move=> a; apply/permP=> x; rewrite !permE. Qed.
+
+Lemma afix_ract B : 'Fix_(to \ sAD)(B) = 'Fix_to(B). Proof. by []. Qed.
+
+Lemma astab_ract S : 'C(S | to \ sAD) = 'C_A(S | to).
+Proof. by rewrite setIA (setIidPl sAD). Qed.
+
+Lemma astabs_ract S : 'N(S | to \ sAD) = 'N_A(S | to).
+Proof. by rewrite setIA (setIidPl sAD). Qed.
+
+Lemma acts_ract (B : {set aT}) S :
+ [acts B, on S | to \ sAD] = (B \subset A) && [acts B, on S | to].
+Proof. by rewrite astabs_ract subsetI. Qed.
+
+End RestrictActionTheory.
+
+Section MorphAct.
+(* Action induced by a morphism to permutations. *)
+
+Variables (aT : finGroupType) (D : {group aT}) (rT : finType).
+Variable phi : {morphism D >-> {perm rT}}.
+
+Definition mact x a := phi a x.
+
+Lemma mact_is_action : is_action D mact.
+Proof.
+split=> [a x y | x a b Da Db]; first exact: perm_inj.
+by rewrite /mact morphM //= permM.
+Qed.
+
+Canonical morph_action := Action mact_is_action.
+
+Lemma mactE x a : morph_action x a = phi a x. Proof. by []. Qed.
+
+Lemma injm_faithful : 'injm phi -> [faithful D, on setT | morph_action].
+Proof.
+move/injmP=> phi_inj; apply/subsetP=> a /setIP[Da /astab_act a1].
+apply/set1P/phi_inj => //; apply/permP=> x.
+by rewrite morph1 perm1 -mactE a1 ?inE.
+Qed.
+
+Lemma perm_mact a : actperm morph_action a = phi a.
+Proof. by apply/permP=> x; rewrite permE. Qed.
+
+End MorphAct.
+
+Notation "<< phi >>" := (morph_action phi) : action_scope.
+
+Section CompAct.
+
+Variables (gT aT : finGroupType) (rT : finType).
+Variables (D : {set aT}) (to : action D rT).
+Variables (B : {set gT}) (f : {morphism B >-> aT}).
+
+Definition comp_act x e := to x (f e).
+Lemma comp_is_action : is_action (f @*^-1 D) comp_act.
+Proof.
+split=> [e | x e1 e2]; first exact: act_inj.
+case/morphpreP=> Be1 Dfe1; case/morphpreP=> Be2 Dfe2.
+by rewrite /comp_act morphM ?actMin.
+Qed.
+Canonical comp_action := Action comp_is_action.
+
+Lemma comp_actE x e : comp_action x e = to x (f e). Proof. by []. Qed.
+
+Lemma afix_comp (A : {set gT}) :
+ A \subset B -> 'Fix_comp_action(A) = 'Fix_to(f @* A).
+Proof.
+move=> sAB; apply/setP=> x; rewrite !inE /morphim (setIidPr sAB).
+apply/subsetP/subsetP=> [cAx _ /imsetP[a Aa ->] | cfAx a Aa].
+ by move/cAx: Aa; rewrite !inE.
+by rewrite inE; move/(_ (f a)): cfAx; rewrite inE mem_imset // => ->.
+Qed.
+
+Lemma astab_comp S : 'C(S | comp_action) = f @*^-1 'C(S | to).
+Proof. by apply/setP=> x; rewrite !inE -andbA. Qed.
+
+Lemma astabs_comp S : 'N(S | comp_action) = f @*^-1 'N(S | to).
+Proof. by apply/setP=> x; rewrite !inE -andbA. Qed.
+
+End CompAct.
+
+Notation "to \o f" := (comp_action to f) : action_scope.
+
+Section PermAction.
+(* Natural action of permutation groups. *)
+
+Variable rT : finType.
+Local Notation gT := {perm rT}.
+Implicit Types a b c : gT.
+
+Lemma aperm_is_action : is_action setT (@aperm rT).
+Proof.
+by apply: is_total_action => [x|x a b]; rewrite apermE (perm1, permM).
+Qed.
+
+Canonical perm_action := Action aperm_is_action.
+
+Lemma pcycleE a : pcycle a = orbit perm_action <[a]>%g.
+Proof. by []. Qed.
+
+Lemma perm_act1P a : reflect (forall x, aperm x a = x) (a == 1).
+Proof.
+apply: (iffP eqP) => [-> x | a1]; first exact: act1.
+by apply/permP=> x; rewrite -apermE a1 perm1.
+Qed.
+
+Lemma perm_faithful A : [faithful A, on setT | perm_action].
+Proof.
+apply/subsetP=> a /setIP[Da crTa].
+by apply/set1P; apply/permP=> x; rewrite -apermE perm1 (astabP crTa) ?inE.
+Qed.
+
+Lemma actperm_id p : actperm perm_action p = p.
+Proof. by apply/permP=> x; rewrite permE. Qed.
+
+End PermAction.
+
+Implicit Arguments perm_act1P [rT a].
+Prenex Implicits perm_act1P.
+
+Notation "'P" := (perm_action _) (at level 8) : action_scope.
+
+Section ActpermOrbits.
+
+Variables (aT : finGroupType) (D : {group aT}) (rT : finType).
+Variable to : action D rT.
+
+Lemma orbit_morphim_actperm (A : {set aT}) :
+ A \subset D -> orbit 'P (actperm to @* A) =1 orbit to A.
+Proof.
+move=> sAD x; rewrite morphimEsub // /orbit -imset_comp.
+by apply: eq_imset => a //=; rewrite actpermK.
+Qed.
+
+Lemma pcycle_actperm (a : aT) :
+ a \in D -> pcycle (actperm to a) =1 orbit to <[a]>.
+Proof.
+move=> Da x.
+by rewrite pcycleE -orbit_morphim_actperm ?cycle_subG ?morphim_cycle.
+Qed.
+
+End ActpermOrbits.
+
+Section RestrictPerm.
+
+Variables (T : finType) (S : {set T}).
+
+Definition restr_perm := actperm (<[subxx 'N(S | 'P)]>).
+Canonical restr_perm_morphism := [morphism of restr_perm].
+
+Lemma restr_perm_on p : perm_on S (restr_perm p).
+Proof.
+apply/subsetP=> x; apply: contraR => notSx.
+by rewrite permE /= /actby (negPf notSx).
+Qed.
+
+Lemma triv_restr_perm p : p \notin 'N(S | 'P) -> restr_perm p = 1.
+Proof.
+move=> not_nSp; apply/permP=> x.
+by rewrite !permE /= /actby (negPf not_nSp) andbF.
+Qed.
+
+Lemma restr_permE : {in 'N(S | 'P) & S, forall p, restr_perm p =1 p}.
+Proof. by move=> y x nSp Sx; rewrite /= actpermE actbyE. Qed.
+
+Lemma ker_restr_perm : 'ker restr_perm = 'C(S | 'P).
+Proof. by rewrite ker_actperm astab_actby setIT (setIidPr (astab_sub _ _)). Qed.
+
+Lemma im_restr_perm p : restr_perm p @: S = S.
+Proof. exact: im_perm_on (restr_perm_on p). Qed.
+
+End RestrictPerm.
+
+Section AutIn.
+
+Variable gT : finGroupType.
+
+Definition Aut_in A (B : {set gT}) := 'N_A(B | 'P) / 'C_A(B | 'P).
+
+Variables G H : {group gT}.
+Hypothesis sHG: H \subset G.
+
+Lemma Aut_restr_perm a : a \in Aut G -> restr_perm H a \in Aut H.
+Proof.
+move=> AutGa.
+case nHa: (a \in 'N(H | 'P)); last by rewrite triv_restr_perm ?nHa ?group1.
+rewrite inE restr_perm_on; apply/morphicP=> x y Hx Hy /=.
+by rewrite !restr_permE ?groupM // -(autmE AutGa) morphM ?(subsetP sHG).
+Qed.
+
+Lemma restr_perm_Aut : restr_perm H @* Aut G \subset Aut H.
+Proof.
+by apply/subsetP=> a'; case/morphimP=> a _ AutGa ->{a'}; exact: Aut_restr_perm.
+Qed.
+
+Lemma Aut_in_isog : Aut_in (Aut G) H \isog restr_perm H @* Aut G.
+Proof.
+rewrite /Aut_in -ker_restr_perm kerE -morphpreIdom -morphimIdom -kerE /=.
+by rewrite setIA (setIC _ (Aut G)) first_isog_loc ?subsetIr.
+Qed.
+
+Lemma Aut_sub_fullP :
+ reflect (forall h : {morphism H >-> gT}, 'injm h -> h @* H = H ->
+ exists g : {morphism G >-> gT},
+ [/\ 'injm g, g @* G = G & {in H, g =1 h}])
+ (Aut_in (Aut G) H \isog Aut H).
+Proof.
+rewrite (isog_transl _ Aut_in_isog) /=; set rG := _ @* _.
+apply: (iffP idP) => [iso_rG h injh hH| AutHinG].
+ have: aut injh hH \in rG; last case/morphimP=> g nHg AutGg def_g.
+ suffices ->: rG = Aut H by exact: Aut_aut.
+ by apply/eqP; rewrite eqEcard restr_perm_Aut /= (card_isog iso_rG).
+ exists (autm_morphism AutGg); rewrite injm_autm im_autm; split=> // x Hx.
+ by rewrite -(autE injh hH Hx) def_g actpermE actbyE.
+suffices ->: rG = Aut H by exact: isog_refl.
+apply/eqP; rewrite eqEsubset restr_perm_Aut /=.
+apply/subsetP=> h AutHh; have hH := im_autm AutHh.
+have [g [injg gG eq_gh]] := AutHinG _ (injm_autm AutHh) hH.
+have [Ng AutGg]: aut injg gG \in 'N(H | 'P) /\ aut injg gG \in Aut G.
+ rewrite Aut_aut !inE; split=> //; apply/subsetP=> x Hx.
+ by rewrite inE /= /aperm autE ?(subsetP sHG) // -hH eq_gh ?mem_morphim.
+apply/morphimP; exists (aut injg gG) => //; apply: (eq_Aut AutHh) => [|x Hx].
+ by rewrite (subsetP restr_perm_Aut) // mem_morphim.
+by rewrite restr_permE //= /aperm autE ?eq_gh ?(subsetP sHG).
+Qed.
+
+End AutIn.
+
+Arguments Scope Aut_in [_ group_scope group_scope].
+
+Section InjmAutIn.
+
+Variables (gT rT : finGroupType) (D G H : {group gT}) (f : {morphism D >-> rT}).
+Hypotheses (injf : 'injm f) (sGD : G \subset D) (sHG : H \subset G).
+Let sHD := subset_trans sHG sGD.
+Local Notation fGisom := (Aut_isom injf sGD).
+Local Notation fHisom := (Aut_isom injf sHD).
+Local Notation inH := (restr_perm H).
+Local Notation infH := (restr_perm (f @* H)).
+
+Lemma astabs_Aut_isom a :
+ a \in Aut G -> (fGisom a \in 'N(f @* H | 'P)) = (a \in 'N(H | 'P)).
+Proof.
+move=> AutGa; rewrite !inE sub_morphim_pre // subsetI sHD /= /aperm.
+rewrite !(sameP setIidPl eqP) !eqEsubset !subsetIl; apply: eq_subset_r => x.
+rewrite !inE; apply: andb_id2l => Hx; have Gx: x \in G := subsetP sHG x Hx.
+have Dax: a x \in D by rewrite (subsetP sGD) // Aut_closed.
+by rewrite Aut_isomE // -!sub1set -morphim_set1 // injmSK ?sub1set.
+Qed.
+
+Lemma isom_restr_perm a : a \in Aut G -> fHisom (inH a) = infH (fGisom a).
+Proof.
+move=> AutGa; case nHa: (a \in 'N(H | 'P)); last first.
+ by rewrite !triv_restr_perm ?astabs_Aut_isom ?nHa ?morph1.
+apply: (eq_Aut (Aut_Aut_isom injf sHD _)) => [|fx Hfx /=].
+ by rewrite (Aut_restr_perm (morphimS f sHG)) ?Aut_Aut_isom.
+have [x Dx Hx def_fx] := morphimP Hfx; have Gx := subsetP sHG x Hx.
+rewrite {1}def_fx Aut_isomE ?(Aut_restr_perm sHG) //.
+by rewrite !restr_permE ?astabs_Aut_isom // def_fx Aut_isomE.
+Qed.
+
+Lemma restr_perm_isom : isom (inH @* Aut G) (infH @* Aut (f @* G)) fHisom.
+Proof.
+apply: sub_isom; rewrite ?restr_perm_Aut ?injm_Aut_isom //=.
+rewrite -(im_Aut_isom injf sGD) -!morphim_comp.
+apply: eq_in_morphim; last exact: isom_restr_perm.
+apply/setP=> a; rewrite 2!in_setI; apply: andb_id2r => AutGa.
+rewrite /= inE andbC inE (Aut_restr_perm sHG) //=.
+by symmetry; rewrite inE AutGa inE astabs_Aut_isom.
+Qed.
+
+Lemma injm_Aut_sub : Aut_in (Aut (f @* G)) (f @* H) \isog Aut_in (Aut G) H.
+Proof.
+do 2!rewrite isog_sym (isog_transl _ (Aut_in_isog _ _)).
+by rewrite isog_sym (isom_isog _ _ restr_perm_isom) // restr_perm_Aut.
+Qed.
+
+Lemma injm_Aut_full :
+ (Aut_in (Aut (f @* G)) (f @* H) \isog Aut (f @* H))
+ = (Aut_in (Aut G) H \isog Aut H).
+Proof.
+by rewrite (isog_transl _ injm_Aut_sub) (isog_transr _ (injm_Aut injf sHD)).
+Qed.
+
+End InjmAutIn.
+
+Section GroupAction.
+
+Variables (aT rT : finGroupType) (D : {set aT}) (R : {set rT}).
+Local Notation actT := (action D rT).
+
+Definition is_groupAction (to : actT) :=
+ {in D, forall a, actperm to a \in Aut R}.
+
+Structure groupAction := GroupAction {gact :> actT; _ : is_groupAction gact}.
+
+Definition clone_groupAction to :=
+ let: GroupAction _ toA := to return {type of GroupAction for to} -> _ in
+ fun k => k toA : groupAction.
+
+End GroupAction.
+
+Delimit Scope groupAction_scope with gact.
+Bind Scope groupAction_scope with groupAction.
+
+Arguments Scope is_groupAction [_ _ group_scope group_scope action_scope].
+Arguments Scope groupAction [_ _ group_scope group_scope].
+Arguments Scope gact [_ _ group_scope group_scope groupAction_scope].
+
+Notation "[ 'groupAction' 'of' to ]" :=
+ (clone_groupAction (@GroupAction _ _ _ _ to))
+ (at level 0, format "[ 'groupAction' 'of' to ]") : form_scope.
+
+Section GroupActionDefs.
+
+Variables (aT rT : finGroupType) (D : {set aT}) (R : {set rT}).
+Implicit Type A : {set aT}.
+Implicit Type S : {set rT}.
+Implicit Type to : groupAction D R.
+
+Definition gact_range of groupAction D R := R.
+
+Definition gacent to A := 'Fix_(R | to)(D :&: A).
+
+Definition acts_on_group A S to := [acts A, on S | to] /\ S \subset R.
+
+Coercion actby_cond_group A S to : acts_on_group A S to -> actby_cond A S to :=
+ @proj1 _ _.
+
+Definition acts_irreducibly A S to :=
+ [min S of G | G :!=: 1 & [acts A, on G | to]].
+
+End GroupActionDefs.
+
+Arguments Scope gacent
+ [_ _ group_scope group_scope groupAction_scope group_scope].
+Arguments Scope acts_on_group
+ [_ _ group_scope group_scope group_scope group_scope groupAction_scope].
+Arguments Scope acts_irreducibly
+ [_ _ group_scope group_scope group_scope group_scope groupAction_scope].
+
+Notation "''C_' ( | to ) ( A )" := (gacent to A)
+ (at level 8, format "''C_' ( | to ) ( A )") : group_scope.
+Notation "''C_' ( G | to ) ( A )" := (G :&: 'C_(|to)(A))
+ (at level 8, format "''C_' ( G | to ) ( A )") : group_scope.
+Notation "''C_' ( | to ) [ a ]" := 'C_(|to)([set a])
+ (at level 8, format "''C_' ( | to ) [ a ]") : group_scope.
+Notation "''C_' ( G | to ) [ a ]" := 'C_(G | to)([set a])
+ (at level 8, format "''C_' ( G | to ) [ a ]") : group_scope.
+
+Notation "{ 'acts' A , 'on' 'group' G | to }" := (acts_on_group A G to)
+ (at level 0, format "{ 'acts' A , 'on' 'group' G | to }") : form_scope.
+
+Section RawGroupAction.
+
+Variables (aT rT : finGroupType) (D : {set aT}) (R : {set rT}).
+Variable to : groupAction D R.
+
+Lemma actperm_Aut : is_groupAction R to. Proof. by case: to. Qed.
+
+Lemma im_actperm_Aut : actperm to @* D \subset Aut R.
+Proof. by apply/subsetP=> _ /morphimP[a _ Da ->]; exact: actperm_Aut. Qed.
+
+Lemma gact_out x a : a \in D -> x \notin R -> to x a = x.
+Proof. by move=> Da Rx; rewrite -actpermE (out_Aut _ Rx) ?actperm_Aut. Qed.
+
+Lemma gactM : {in D, forall a, {in R &, {morph to^~ a : x y / x * y}}}.
+Proof.
+move=> a Da /= x y; rewrite -!(actpermE to); apply: morphicP x y.
+by rewrite Aut_morphic ?actperm_Aut.
+Qed.
+
+Lemma actmM a : {in R &, {morph actm to a : x y / x * y}}.
+Proof. rewrite /actm; case: ifP => //; exact: gactM. Qed.
+
+Canonical act_morphism a := Morphism (actmM a).
+
+Lemma morphim_actm :
+ {in D, forall a (S : {set rT}), S \subset R -> actm to a @* S = to^* S a}.
+Proof. by move=> a Da /= S sSR; rewrite /morphim /= actmEfun ?(setIidPr _). Qed.
+
+Variables (a : aT) (A B : {set aT}) (S : {set rT}).
+
+Lemma gacentIdom : 'C_(|to)(D :&: A) = 'C_(|to)(A).
+Proof. by rewrite /gacent setIA setIid. Qed.
+
+Lemma gacentIim : 'C_(R | to)(A) = 'C_(|to)(A).
+Proof. by rewrite setIA setIid. Qed.
+
+Lemma gacentS : A \subset B -> 'C_(|to)(B) \subset 'C_(|to)(A).
+Proof. by move=> sAB; rewrite !(setIS, afixS). Qed.
+
+Lemma gacentU : 'C_(|to)(A :|: B) = 'C_(|to)(A) :&: 'C_(|to)(B).
+Proof. by rewrite -setIIr -afixU -setIUr. Qed.
+
+Hypotheses (Da : a \in D) (sAD : A \subset D) (sSR : S \subset R).
+
+Lemma gacentE : 'C_(|to)(A) = 'Fix_(R | to)(A).
+Proof. by rewrite -{2}(setIidPr sAD). Qed.
+
+Lemma gacent1E : 'C_(|to)[a] = 'Fix_(R | to)[a].
+Proof. by rewrite /gacent [D :&: _](setIidPr _) ?sub1set. Qed.
+
+Lemma subgacentE : 'C_(S | to)(A) = 'Fix_(S | to)(A).
+Proof. by rewrite gacentE setIA (setIidPl sSR). Qed.
+
+Lemma subgacent1E : 'C_(S | to)[a] = 'Fix_(S | to)[a].
+Proof. by rewrite gacent1E setIA (setIidPl sSR). Qed.
+
+End RawGroupAction.
+
+Section GroupActionTheory.
+
+Variables aT rT : finGroupType.
+Variables (D : {group aT}) (R : {group rT}) (to : groupAction D R).
+Implicit Type A B : {set aT}.
+Implicit Types G H : {group aT}.
+Implicit Type S : {set rT}.
+Implicit Types M N : {group rT}.
+
+Lemma gact1 : {in D, forall a, to 1 a = 1}.
+Proof. by move=> a Da; rewrite /= -actmE ?morph1. Qed.
+
+Lemma gactV : {in D, forall a, {in R, {morph to^~ a : x / x^-1}}}.
+Proof. by move=> a Da /= x Rx; move; rewrite -!actmE ?morphV. Qed.
+
+Lemma gactX : {in D, forall a n, {in R, {morph to^~ a : x / x ^+ n}}}.
+Proof. by move=> a Da /= n x Rx; rewrite -!actmE // morphX. Qed.
+
+Lemma gactJ : {in D, forall a, {in R &, {morph to^~ a : x y / x ^ y}}}.
+Proof. by move=> a Da /= x Rx y Ry; rewrite -!actmE // morphJ. Qed.
+
+Lemma gactR : {in D, forall a, {in R &, {morph to^~ a : x y / [~ x, y]}}}.
+Proof. by move=> a Da /= x Rx y Ry; rewrite -!actmE // morphR. Qed.
+
+Lemma gact_stable : {acts D, on R | to}.
+Proof.
+apply: acts_act; apply/subsetP=> a Da; rewrite !inE Da.
+apply/subsetP=> x; rewrite inE; apply: contraLR => R'xa.
+by rewrite -(actKin to Da x) gact_out ?groupV.
+Qed.
+
+Lemma group_set_gacent A : group_set 'C_(|to)(A).
+Proof.
+apply/group_setP; split=> [|x y].
+ by rewrite !inE group1; apply/subsetP=> a /setIP[Da _]; rewrite inE gact1.
+case/setIP=> Rx /afixP cAx /setIP[Ry /afixP cAy].
+rewrite inE groupM //; apply/afixP=> a Aa.
+by rewrite gactM ?cAx ?cAy //; case/setIP: Aa.
+Qed.
+
+Canonical gacent_group A := Group (group_set_gacent A).
+
+Lemma gacent1 : 'C_(|to)(1) = R.
+Proof. by rewrite /gacent (setIidPr (sub1G _)) afix1 setIT. Qed.
+
+Lemma gacent_gen A : A \subset D -> 'C_(|to)(<<A>>) = 'C_(|to)(A).
+Proof.
+by move=> sAD; rewrite /gacent ![D :&: _](setIidPr _) ?gen_subG ?afix_gen_in.
+Qed.
+
+Lemma gacentD1 A : 'C_(|to)(A^#) = 'C_(|to)(A).
+Proof.
+rewrite -gacentIdom -gacent_gen ?subsetIl // setIDA genD1 ?group1 //.
+by rewrite gacent_gen ?subsetIl // gacentIdom.
+Qed.
+
+Lemma gacent_cycle a : a \in D -> 'C_(|to)(<[a]>) = 'C_(|to)[a].
+Proof. by move=> Da; rewrite gacent_gen ?sub1set. Qed.
+
+Lemma gacentY A B :
+ A \subset D -> B \subset D -> 'C_(|to)(A <*> B) = 'C_(|to)(A) :&: 'C_(|to)(B).
+Proof. by move=> sAD sBD; rewrite gacent_gen ?gacentU // subUset sAD. Qed.
+
+Lemma gacentM G H :
+ G \subset D -> H \subset D -> 'C_(|to)(G * H) = 'C_(|to)(G) :&: 'C_(|to)(H).
+Proof.
+by move=> sGD sHB; rewrite -gacent_gen ?mul_subG // genM_join gacentY.
+Qed.
+
+Lemma astab1 : 'C(1 | to) = D.
+Proof.
+by apply/setP=> x; rewrite ?(inE, sub1set) andb_idr //; move/gact1=> ->.
+Qed.
+
+Lemma astab_range : 'C(R | to) = 'C(setT | to).
+Proof.
+apply/eqP; rewrite eqEsubset andbC astabS ?subsetT //=.
+apply/subsetP=> a cRa; have Da := astab_dom cRa; rewrite !inE Da.
+apply/subsetP=> x; rewrite -(setUCr R) !inE.
+by case/orP=> ?; [rewrite (astab_act cRa) | rewrite gact_out].
+Qed.
+
+Lemma gacentC A S :
+ A \subset D -> S \subset R ->
+ (S \subset 'C_(|to)(A)) = (A \subset 'C(S | to)).
+Proof. by move=> sAD sSR; rewrite subsetI sSR astabCin // (setIidPr sAD). Qed.
+
+Lemma astab_gen S : S \subset R -> 'C(<<S>> | to) = 'C(S | to).
+Proof.
+move=> sSR; apply/setP=> a; case Da: (a \in D); last by rewrite !inE Da.
+by rewrite -!sub1set -!gacentC ?sub1set ?gen_subG.
+Qed.
+
+Lemma astabM M N :
+ M \subset R -> N \subset R -> 'C(M * N | to) = 'C(M | to) :&: 'C(N | to).
+Proof.
+move=> sMR sNR; rewrite -astabU -astab_gen ?mul_subG // genM_join.
+by rewrite astab_gen // subUset sMR.
+Qed.
+
+Lemma astabs1 : 'N(1 | to) = D.
+Proof. by rewrite astabs_set1 astab1. Qed.
+
+Lemma astabs_range : 'N(R | to) = D.
+Proof.
+apply/setIidPl; apply/subsetP=> a Da; rewrite inE.
+by apply/subsetP=> x Rx; rewrite inE gact_stable.
+Qed.
+
+Lemma astabsD1 S : 'N(S^# | to) = 'N(S | to).
+Proof.
+case S1: (1 \in S); last first.
+ by rewrite (setDidPl _) // disjoint_sym disjoints_subset sub1set inE S1.
+apply/eqP; rewrite eqEsubset andbC -{1}astabsIdom -{1}astabs1 setIC astabsD /=.
+by rewrite -{2}(setD1K S1) -astabsIdom -{1}astabs1 astabsU.
+Qed.
+
+Lemma gacts_range A : A \subset D -> {acts A, on group R | to}.
+Proof. by move=> sAD; split; rewrite ?astabs_range. Qed.
+
+Lemma acts_subnorm_gacent A : A \subset D ->
+ [acts 'N_D(A), on 'C_(| to)(A) | to].
+Proof.
+move=> sAD; rewrite gacentE // actsI ?astabs_range ?subsetIl //.
+by rewrite -{2}(setIidPr sAD) acts_subnorm_fix.
+Qed.
+
+Lemma acts_subnorm_subgacent A B S :
+ A \subset D -> [acts B, on S | to] -> [acts 'N_B(A), on 'C_(S | to)(A) | to].
+Proof.
+move=> sAD actsB; rewrite actsI //; first by rewrite subIset ?actsB.
+by rewrite (subset_trans _ (acts_subnorm_gacent sAD)) ?setSI ?(acts_dom actsB).
+Qed.
+
+Lemma acts_gen A S :
+ S \subset R -> [acts A, on S | to] -> [acts A, on <<S>> | to].
+Proof.
+move=> sSR actsA; apply: {A}subset_trans actsA _.
+apply/subsetP=> a nSa; have Da := astabs_dom nSa; rewrite !inE Da.
+apply: subset_trans (_ : <<S>> \subset actm to a @*^-1 <<S>>) _.
+ rewrite gen_subG subsetI sSR; apply/subsetP=> x Sx.
+ by rewrite inE /= actmE ?mem_gen // astabs_act.
+by apply/subsetP=> x; rewrite !inE; case/andP=> Rx; rewrite /= actmE.
+Qed.
+
+Lemma acts_joing A M N :
+ M \subset R -> N \subset R -> [acts A, on M | to] -> [acts A, on N | to] ->
+ [acts A, on M <*> N | to].
+Proof. by move=> sMR sNR nMA nNA; rewrite acts_gen ?actsU // subUset sMR. Qed.
+
+Lemma injm_actm a : 'injm (actm to a).
+Proof.
+apply/injmP=> x y Rx Ry; rewrite /= /actm; case: ifP => Da //.
+exact: act_inj.
+Qed.
+
+Lemma im_actm a : actm to a @* R = R.
+Proof.
+apply/eqP; rewrite eqEcard (card_injm (injm_actm a)) // leqnn andbT.
+apply/subsetP=> _ /morphimP[x Rx _ ->] /=.
+by rewrite /actm; case: ifP => // Da; rewrite gact_stable.
+Qed.
+
+Lemma acts_char G M : G \subset D -> M \char R -> [acts G, on M | to].
+Proof.
+move=> sGD /charP[sMR charM].
+apply/subsetP=> a Ga; have Da := subsetP sGD a Ga; rewrite !inE Da.
+apply/subsetP=> x Mx; have Rx := subsetP sMR x Mx.
+by rewrite inE -(charM _ (injm_actm a) (im_actm a)) -actmE // mem_morphim.
+Qed.
+
+Lemma gacts_char G M :
+ G \subset D -> M \char R -> {acts G, on group M | to}.
+Proof. by move=> sGD charM; split; rewrite (acts_char, char_sub). Qed.
+
+Section Restrict.
+
+Variables (A : {group aT}) (sAD : A \subset D).
+
+Lemma ract_is_groupAction : is_groupAction R (to \ sAD).
+Proof. by move=> a Aa /=; rewrite ractpermE actperm_Aut ?(subsetP sAD). Qed.
+
+Canonical ract_groupAction := GroupAction ract_is_groupAction.
+
+Lemma gacent_ract B : 'C_(|ract_groupAction)(B) = 'C_(|to)(A :&: B).
+Proof. by rewrite /gacent afix_ract setIA (setIidPr sAD). Qed.
+
+End Restrict.
+
+Section ActBy.
+
+Variables (A : {group aT}) (G : {group rT}) (nGAg : {acts A, on group G | to}).
+
+Lemma actby_is_groupAction : is_groupAction G <[nGAg]>.
+Proof.
+move=> a Aa; rewrite /= inE; apply/andP; split.
+ apply/subsetP=> x; apply: contraR => Gx.
+ by rewrite actpermE /= /actby (negbTE Gx).
+apply/morphicP=> x y Gx Gy; rewrite !actpermE /= /actby Aa groupM ?Gx ?Gy //=.
+by case nGAg; move/acts_dom; do 2!move/subsetP=> ?; rewrite gactM; auto.
+Qed.
+
+Canonical actby_groupAction := GroupAction actby_is_groupAction.
+
+Lemma gacent_actby B :
+ 'C_(|actby_groupAction)(B) = 'C_(G | to)(A :&: B).
+Proof.
+rewrite /gacent afix_actby !setIA setIid setIUr setICr set0U.
+by have [nAG sGR] := nGAg; rewrite (setIidPr (acts_dom nAG)) (setIidPl sGR).
+Qed.
+
+End ActBy.
+
+Section Quotient.
+
+Variable H : {group rT}.
+
+Lemma acts_qact_dom_norm : {acts qact_dom to H, on 'N(H) | to}.
+Proof.
+move=> a HDa /= x; rewrite {2}(('N(H) =P to^~ a @^-1: 'N(H)) _) ?inE {x}//.
+rewrite eqEcard (card_preimset _ (act_inj _ _)) leqnn andbT.
+apply/subsetP=> x Nx; rewrite inE; move/(astabs_act (H :* x)): HDa.
+rewrite mem_rcosets mulSGid ?normG // Nx => /rcosetsP[y Ny defHy].
+suffices: to x a \in H :* y by apply: subsetP; rewrite mul_subG ?sub1set ?normG.
+by rewrite -defHy; apply: mem_imset; exact: rcoset_refl.
+Qed.
+
+Lemma qact_is_groupAction : is_groupAction (R / H) (to / H).
+Proof.
+move=> a HDa /=; have Da := astabs_dom HDa.
+rewrite inE; apply/andP; split.
+ apply/subsetP=> Hx /=; case: (cosetP Hx) => x Nx ->{Hx}.
+ apply: contraR => R'Hx; rewrite actpermE qactE // gact_out //.
+ by apply: contra R'Hx; apply: mem_morphim.
+apply/morphicP=> Hx Hy; rewrite !actpermE.
+case/morphimP=> x Nx Gx ->{Hx}; case/morphimP=> y Ny Gy ->{Hy}.
+by rewrite -morphM ?qactE ?groupM ?gactM // morphM ?acts_qact_dom_norm.
+Qed.
+
+Canonical quotient_groupAction := GroupAction qact_is_groupAction.
+
+Lemma qact_domE : H \subset R -> qact_dom to H = 'N(H | to).
+Proof.
+move=> sHR; apply/setP=> a; apply/idP/idP=> nHa; have Da := astabs_dom nHa.
+ rewrite !inE Da; apply/subsetP=> x Hx; rewrite inE -(rcoset1 H).
+ have /rcosetsP[y Ny defHy]: to^~ a @: H \in rcosets H 'N(H).
+ by rewrite (astabs_act _ nHa) -{1}(mulg1 H) -rcosetE mem_imset ?group1.
+ by rewrite (@rcoset_transl _ H 1 y) -defHy -1?(gact1 Da) mem_setact.
+rewrite !inE Da; apply/subsetP=> Hx; rewrite inE => /rcosetsP[x Nx ->{Hx}].
+apply/imsetP; exists (to x a).
+ case Rx: (x \in R); last by rewrite gact_out ?Rx.
+ rewrite inE; apply/subsetP=> _ /imsetP[y Hy ->].
+ rewrite -(actKVin to Da y) -gactJ // ?(subsetP sHR, astabs_act, groupV) //.
+ by rewrite memJ_norm // astabs_act ?groupV.
+apply/eqP; rewrite rcosetE eqEcard.
+rewrite (card_imset _ (act_inj _ _)) !card_rcoset leqnn andbT.
+apply/subsetP=> _ /imsetP[y Hxy ->]; rewrite !mem_rcoset in Hxy *.
+have Rxy := subsetP sHR _ Hxy; rewrite -(mulgKV x y).
+case Rx: (x \in R); last by rewrite !gact_out ?mulgK // 1?groupMl ?Rx.
+by rewrite -gactV // -gactM 1?groupMr ?groupV // mulgK astabs_act.
+Qed.
+
+End Quotient.
+
+Section Mod.
+
+Variable H : {group aT}.
+
+Lemma modact_is_groupAction : is_groupAction 'C_(|to)(H) (to %% H).
+Proof.
+move=> Ha /morphimP[a Na Da ->]; have NDa: a \in 'N_D(H) by exact/setIP.
+rewrite inE; apply/andP; split.
+ apply/subsetP=> x; rewrite 2!inE andbC actpermE /= modactEcond //.
+ by apply: contraR; case: ifP => // E Rx; rewrite gact_out.
+apply/morphicP=> x y /setIP[Rx cHx] /setIP[Ry cHy].
+rewrite /= !actpermE /= !modactE ?gactM //.
+suffices: x * y \in 'C_(|to)(H) by case/setIP.
+rewrite groupM //; exact/setIP.
+Qed.
+
+Canonical mod_groupAction := GroupAction modact_is_groupAction.
+
+Lemma modgactE x a :
+ H \subset 'C(R | to) -> a \in 'N_D(H) -> (to %% H)%act x (coset H a) = to x a.
+Proof.
+move=> cRH NDa /=; have [Da Na] := setIP NDa.
+have [Rx | notRx] := boolP (x \in R).
+ by rewrite modactE //; apply/afixP=> b /setIP[_ /(subsetP cRH)/astab_act->].
+rewrite gact_out //= /modact; case: ifP => // _; rewrite gact_out //.
+suffices: a \in D :&: coset H a by case/mem_repr/setIP.
+by rewrite inE Da val_coset // rcoset_refl.
+Qed.
+
+Lemma gacent_mod G M :
+ H \subset 'C(M | to) -> G \subset 'N(H) ->
+ 'C_(M | mod_groupAction)(G / H) = 'C_(M | to)(G).
+Proof.
+move=> cMH nHG; rewrite -gacentIdom gacentE ?subsetIl // setICA.
+have sHD: H \subset D by rewrite (subset_trans cMH) ?subsetIl.
+rewrite -quotientGI // afix_mod ?setIS // setICA -gacentIim (setIC R) -setIA.
+rewrite -gacentE ?subsetIl // gacentIdom setICA (setIidPr _) //.
+by rewrite gacentC // ?(subset_trans cMH) ?astabS ?subsetIl // setICA subsetIl.
+Qed.
+
+Lemma acts_irr_mod G M :
+ H \subset 'C(M | to) -> G \subset 'N(H) -> acts_irreducibly G M to ->
+ acts_irreducibly (G / H) M mod_groupAction.
+Proof.
+move=> cMH nHG /mingroupP[/andP[ntM nMG] minM].
+apply/mingroupP; rewrite ntM astabs_mod ?quotientS //; split=> // L modL ntL.
+have cLH: H \subset 'C(L | to) by rewrite (subset_trans cMH) ?astabS //.
+apply: minM => //; case/andP: modL => ->; rewrite astabs_mod ?quotientSGK //.
+by rewrite (subset_trans cLH) ?astab_sub.
+Qed.
+
+End Mod.
+
+Lemma modact_coset_astab x a :
+ a \in D -> (to %% 'C(R | to))%act x (coset _ a) = to x a.
+Proof.
+move=> Da; apply: modgactE => {x}//.
+rewrite !inE Da; apply/subsetP=> _ /imsetP[c Cc ->].
+have Dc := astab_dom Cc; rewrite !inE groupJ //.
+apply/subsetP=> x Rx; rewrite inE conjgE !actMin ?groupM ?groupV //.
+by rewrite (astab_act Cc) ?actKVin // gact_stable ?groupV.
+Qed.
+
+Lemma acts_irr_mod_astab G M :
+ acts_irreducibly G M to ->
+ acts_irreducibly (G / 'C_G(M | to)) M (mod_groupAction _).
+Proof.
+move=> irrG; have /andP[_ nMG] := mingroupp irrG.
+apply: acts_irr_mod irrG; first exact: subsetIr.
+by rewrite normsI ?normG // (subset_trans nMG) // astab_norm.
+Qed.
+
+Section CompAct.
+
+Variables (gT : finGroupType) (G : {group gT}) (f : {morphism G >-> aT}).
+
+Lemma comp_is_groupAction : is_groupAction R (comp_action to f).
+Proof.
+move=> a /morphpreP[Ba Dfa]; apply: etrans (actperm_Aut to Dfa).
+by congr (_ \in Aut R); apply/permP=> x; rewrite !actpermE.
+Qed.
+Canonical comp_groupAction := GroupAction comp_is_groupAction.
+
+Lemma gacent_comp U : 'C_(|comp_groupAction)(U) = 'C_(|to)(f @* U).
+Proof.
+rewrite /gacent afix_comp ?subIset ?subxx //.
+by rewrite -(setIC U) (setIC D) morphim_setIpre.
+Qed.
+
+End CompAct.
+
+End GroupActionTheory.
+
+Notation "''C_' ( | to ) ( A )" := (gacent_group to A) : Group_scope.
+Notation "''C_' ( G | to ) ( A )" :=
+ (setI_group G 'C_(|to)(A)) : Group_scope.
+Notation "''C_' ( | to ) [ a ]" := (gacent_group to [set a%g]) : Group_scope.
+Notation "''C_' ( G | to ) [ a ]" :=
+ (setI_group G 'C_(|to)[a]) : Group_scope.
+
+Notation "to \ sAD" := (ract_groupAction to sAD) : groupAction_scope.
+Notation "<[ nGA ] >" := (actby_groupAction nGA) : groupAction_scope.
+Notation "to / H" := (quotient_groupAction to H) : groupAction_scope.
+Notation "to %% H" := (mod_groupAction to H) : groupAction_scope.
+Notation "to \o f" := (comp_groupAction to f) : groupAction_scope.
+
+(* Operator group isomorphism. *)
+Section MorphAction.
+
+Variables (aT1 aT2 : finGroupType) (rT1 rT2 : finType).
+Variables (D1 : {group aT1}) (D2 : {group aT2}).
+Variables (to1 : action D1 rT1) (to2 : action D2 rT2).
+Variables (A : {set aT1}) (R S : {set rT1}).
+Variables (h : rT1 -> rT2) (f : {morphism D1 >-> aT2}).
+Hypotheses (actsDR : {acts D1, on R | to1}) (injh : {in R &, injective h}).
+Hypothesis defD2 : f @* D1 = D2.
+Hypotheses (sSR : S \subset R) (sAD1 : A \subset D1).
+Hypothesis hfJ : {in S & D1, morph_act to1 to2 h f}.
+
+Lemma morph_astabs : f @* 'N(S | to1) = 'N(h @: S | to2).
+Proof.
+apply/setP=> fx; apply/morphimP/idP=> [[x D1x nSx ->] | nSx].
+ rewrite 2!inE -{1}defD2 mem_morphim //=; apply/subsetP=> _ /imsetP[u Su ->].
+ by rewrite inE -hfJ ?mem_imset // (astabs_act _ nSx).
+have [|x D1x _ def_fx] := morphimP (_ : fx \in f @* D1).
+ by rewrite defD2 (astabs_dom nSx).
+exists x => //; rewrite !inE D1x; apply/subsetP=> u Su.
+have /imsetP[u' Su' /injh def_u']: h (to1 u x) \in h @: S.
+ by rewrite hfJ // -def_fx (astabs_act _ nSx) mem_imset.
+by rewrite inE def_u' ?actsDR ?(subsetP sSR).
+Qed.
+
+Lemma morph_astab : f @* 'C(S | to1) = 'C(h @: S | to2).
+Proof.
+apply/setP=> fx; apply/morphimP/idP=> [[x D1x cSx ->] | cSx].
+ rewrite 2!inE -{1}defD2 mem_morphim //=; apply/subsetP=> _ /imsetP[u Su ->].
+ by rewrite inE -hfJ // (astab_act cSx).
+have [|x D1x _ def_fx] := morphimP (_ : fx \in f @* D1).
+ by rewrite defD2 (astab_dom cSx).
+exists x => //; rewrite !inE D1x; apply/subsetP=> u Su.
+rewrite inE -(inj_in_eq injh) ?actsDR ?(subsetP sSR) ?hfJ //.
+by rewrite -def_fx (astab_act cSx) ?mem_imset.
+Qed.
+
+Lemma morph_afix : h @: 'Fix_(S | to1)(A) = 'Fix_(h @: S | to2)(f @* A).
+Proof.
+apply/setP=> hu; apply/imsetP/setIP=> [[u /setIP[Su cAu] ->]|].
+ split; first by rewrite mem_imset.
+ by apply/afixP=> _ /morphimP[x D1x Ax ->]; rewrite -hfJ ?(afixP cAu).
+case=> /imsetP[u Su ->] /afixP c_hu_fA; exists u; rewrite // inE Su.
+apply/afixP=> x Ax; have Dx := subsetP sAD1 x Ax.
+by apply: injh; rewrite ?actsDR ?(subsetP sSR) ?hfJ // c_hu_fA ?mem_morphim.
+Qed.
+
+End MorphAction.
+
+Section MorphGroupAction.
+
+Variables (aT1 aT2 rT1 rT2 : finGroupType).
+Variables (D1 : {group aT1}) (D2 : {group aT2}).
+Variables (R1 : {group rT1}) (R2 : {group rT2}).
+Variables (to1 : groupAction D1 R1) (to2 : groupAction D2 R2).
+Variables (h : {morphism R1 >-> rT2}) (f : {morphism D1 >-> aT2}).
+Hypotheses (iso_h : isom R1 R2 h) (iso_f : isom D1 D2 f).
+Hypothesis hfJ : {in R1 & D1, morph_act to1 to2 h f}.
+Implicit Types (A : {set aT1}) (S : {set rT1}) (M : {group rT1}).
+
+Lemma morph_gastabs S : S \subset R1 -> f @* 'N(S | to1) = 'N(h @* S | to2).
+Proof.
+have [[_ defD2] [injh _]] := (isomP iso_f, isomP iso_h).
+move=> sSR1; rewrite (morphimEsub _ sSR1).
+apply: (morph_astabs (gact_stable to1) (injmP injh)) => // u x.
+by move/(subsetP sSR1); exact: hfJ.
+Qed.
+
+Lemma morph_gastab S : S \subset R1 -> f @* 'C(S | to1) = 'C(h @* S | to2).
+Proof.
+have [[_ defD2] [injh _]] := (isomP iso_f, isomP iso_h).
+move=> sSR1; rewrite (morphimEsub _ sSR1).
+apply: (morph_astab (gact_stable to1) (injmP injh)) => // u x.
+by move/(subsetP sSR1); exact: hfJ.
+Qed.
+
+Lemma morph_gacent A : A \subset D1 -> h @* 'C_(|to1)(A) = 'C_(|to2)(f @* A).
+Proof.
+have [[_ defD2] [injh defR2]] := (isomP iso_f, isomP iso_h).
+move=> sAD1; rewrite !gacentE //; last by rewrite -defD2 morphimS.
+rewrite morphimEsub ?subsetIl // -{1}defR2 morphimEdom.
+exact: (morph_afix (gact_stable to1) (injmP injh)).
+Qed.
+
+Lemma morph_gact_irr A M :
+ A \subset D1 -> M \subset R1 ->
+ acts_irreducibly (f @* A) (h @* M) to2 = acts_irreducibly A M to1.
+Proof.
+move=> sAD1 sMR1.
+have [[injf defD2] [injh defR2]] := (isomP iso_f, isomP iso_h).
+have h_eq1 := morphim_injm_eq1 injh.
+apply/mingroupP/mingroupP=> [] [/andP[ntM actAM] minM].
+ split=> [|U]; first by rewrite -h_eq1 // ntM -(injmSK injf) ?morph_gastabs.
+ case/andP=> ntU acts_fAU sUM; have sUR1 := subset_trans sUM sMR1.
+ apply: (injm_morphim_inj injh) => //; apply: minM; last exact: morphimS.
+ by rewrite h_eq1 // ntU -morph_gastabs ?morphimS.
+split=> [|U]; first by rewrite h_eq1 // ntM -morph_gastabs ?morphimS.
+case/andP=> ntU acts_fAU sUhM.
+have sUhR1 := subset_trans sUhM (morphimS h sMR1).
+have sU'M: h @*^-1 U \subset M by rewrite sub_morphpre_injm.
+rewrite /= -(minM _ _ sU'M) ?morphpreK // -h_eq1 ?subsetIl // -(injmSK injf) //.
+by rewrite morph_gastabs ?(subset_trans sU'M) // morphpreK ?ntU.
+Qed.
+
+End MorphGroupAction.
+
+(* Conjugation and right translation actions. *)
+Section InternalActionDefs.
+
+Variable gT : finGroupType.
+Implicit Type A : {set gT}.
+Implicit Type G : {group gT}.
+
+(* This is not a Canonical action because it is seldom used, and it would *)
+(* cause too many spurious matches (any group product would be viewed as an *)
+(* action!). *)
+Definition mulgr_action := TotalAction (@mulg1 gT) (@mulgA gT).
+
+Canonical conjg_action := TotalAction (@conjg1 gT) (@conjgM gT).
+
+Lemma conjg_is_groupAction : is_groupAction setT conjg_action.
+Proof.
+move=> a _; rewrite /= inE; apply/andP; split.
+ by apply/subsetP=> x _; rewrite inE.
+by apply/morphicP=> x y _ _; rewrite !actpermE /= conjMg.
+Qed.
+
+Canonical conjg_groupAction := GroupAction conjg_is_groupAction.
+
+Lemma rcoset_is_action : is_action setT (@rcoset gT).
+Proof.
+by apply: is_total_action => [A|A x y]; rewrite !rcosetE (mulg1, rcosetM).
+Qed.
+
+Canonical rcoset_action := Action rcoset_is_action.
+
+Canonical conjsg_action := TotalAction (@conjsg1 gT) (@conjsgM gT).
+
+Lemma conjG_is_action : is_action setT (@conjG_group gT).
+Proof.
+apply: is_total_action => [G | G x y]; apply: val_inj; rewrite /= ?act1 //.
+exact: actM.
+Qed.
+
+Definition conjG_action := Action conjG_is_action.
+
+End InternalActionDefs.
+
+Notation "'R" := (@mulgr_action _) (at level 8) : action_scope.
+Notation "'Rs" := (@rcoset_action _) (at level 8) : action_scope.
+Notation "'J" := (@conjg_action _) (at level 8) : action_scope.
+Notation "'J" := (@conjg_groupAction _) (at level 8) : groupAction_scope.
+Notation "'Js" := (@conjsg_action _) (at level 8) : action_scope.
+Notation "'JG" := (@conjG_action _) (at level 8) : action_scope.
+Notation "'Q" := ('J / _)%act (at level 8) : action_scope.
+Notation "'Q" := ('J / _)%gact (at level 8) : groupAction_scope.
+
+Section InternalGroupAction.
+
+Variable gT : finGroupType.
+Implicit Types A B : {set gT}.
+Implicit Types G H : {group gT}.
+Implicit Type x : gT.
+
+(* Various identities for actions on groups. *)
+
+Lemma orbitR G x : orbit 'R G x = x *: G.
+Proof. by rewrite -lcosetE. Qed.
+
+Lemma astab1R x : 'C[x | 'R] = 1.
+Proof.
+apply/trivgP/subsetP=> y cxy.
+by rewrite -(mulKg x y) [x * y](astab1P cxy) mulVg set11.
+Qed.
+
+Lemma astabR G : 'C(G | 'R) = 1.
+Proof.
+apply/trivgP/subsetP=> x cGx.
+by rewrite -(mul1g x) [1 * x](astabP cGx) group1.
+Qed.
+
+Lemma astabsR G : 'N(G | 'R) = G.
+Proof.
+apply/setP=> x; rewrite !inE -setactVin ?inE //=.
+by rewrite -groupV -{1 3}(mulg1 G) rcoset_sym -sub1set -mulGS -!rcosetE.
+Qed.
+
+Lemma atransR G : [transitive G, on G | 'R].
+Proof. by rewrite /atrans -{1}(mul1g G) -orbitR mem_imset. Qed.
+
+Lemma faithfulR G : [faithful G, on G | 'R].
+Proof. by rewrite /faithful astabR subsetIr. Qed.
+
+Definition Cayley_repr G := actperm <[atrans_acts (atransR G)]>.
+
+Theorem Cayley_isom G : isom G (Cayley_repr G @* G) (Cayley_repr G).
+Proof. exact: faithful_isom (faithfulR G). Qed.
+
+Theorem Cayley_isog G : G \isog Cayley_repr G @* G.
+Proof. exact: isom_isog (Cayley_isom G). Qed.
+
+Lemma orbitJ G x : orbit 'J G x = x ^: G. Proof. by []. Qed.
+
+Lemma afixJ A : 'Fix_('J)(A) = 'C(A).
+Proof.
+apply/setP=> x; apply/afixP/centP=> cAx y Ay /=.
+ by rewrite /commute conjgC cAx.
+by rewrite conjgE cAx ?mulKg.
+Qed.
+
+Lemma astabJ A : 'C(A |'J) = 'C(A).
+Proof.
+apply/setP=> x; apply/astabP/centP=> cAx y Ay /=.
+ by apply: esym; rewrite conjgC cAx.
+by rewrite conjgE -cAx ?mulKg.
+Qed.
+
+Lemma astab1J x : 'C[x |'J] = 'C[x].
+Proof. by rewrite astabJ cent_set1. Qed.
+
+Lemma astabsJ A : 'N(A | 'J) = 'N(A).
+Proof. by apply/setP=> x; rewrite -2!groupV !inE -conjg_preim -sub_conjg. Qed.
+
+Lemma setactJ A x : 'J^*%act A x = A :^ x. Proof. by []. Qed.
+
+Lemma gacentJ A : 'C_(|'J)(A) = 'C(A).
+Proof. by rewrite gacentE ?setTI ?subsetT ?afixJ. Qed.
+
+Lemma orbitRs G A : orbit 'Rs G A = rcosets A G. Proof. by []. Qed.
+
+Lemma sub_afixRs_norms G x A : (G :* x \in 'Fix_('Rs)(A)) = (A \subset G :^ x).
+Proof.
+rewrite inE /=; apply: eq_subset_r => a.
+rewrite inE rcosetE -(can2_eq (rcosetKV x) (rcosetK x)) -!rcosetM.
+rewrite eqEcard card_rcoset leqnn andbT mulgA (conjgCV x) mulgK.
+by rewrite -{2 3}(mulGid G) mulGS sub1set -mem_conjg.
+Qed.
+
+Lemma sub_afixRs_norm G x : (G :* x \in 'Fix_('Rs)(G)) = (x \in 'N(G)).
+Proof. by rewrite sub_afixRs_norms -groupV inE sub_conjgV. Qed.
+
+Lemma afixRs_rcosets A G : 'Fix_(rcosets G A | 'Rs)(G) = rcosets G 'N_A(G).
+Proof.
+apply/setP=> Gx; apply/setIP/rcosetsP=> [[/rcosetsP[x Ax ->]]|[x]].
+ by rewrite sub_afixRs_norm => Nx; exists x; rewrite // inE Ax.
+by case/setIP=> Ax Nx ->; rewrite -{1}rcosetE mem_imset // sub_afixRs_norm.
+Qed.
+
+Lemma astab1Rs G : 'C[G : {set gT} | 'Rs] = G.
+Proof.
+apply/setP=> x.
+by apply/astab1P/idP=> /= [<- | Gx]; rewrite rcosetE ?rcoset_refl ?rcoset_id.
+Qed.
+
+Lemma actsRs_rcosets H G : [acts G, on rcosets H G | 'Rs].
+Proof. by rewrite -orbitRs acts_orbit ?subsetT. Qed.
+
+Lemma transRs_rcosets H G : [transitive G, on rcosets H G | 'Rs].
+Proof. by rewrite -orbitRs atrans_orbit. Qed.
+
+(* This is the second part of Aschbacher (5.7) *)
+Lemma astabRs_rcosets H G : 'C(rcosets H G | 'Rs) = gcore H G.
+Proof.
+have transGH := transRs_rcosets H G.
+by rewrite (astab_trans_gcore transGH (orbit_refl _ G _)) astab1Rs.
+Qed.
+
+Lemma orbitJs G A : orbit 'Js G A = A :^: G. Proof. by []. Qed.
+
+Lemma astab1Js A : 'C[A | 'Js] = 'N(A).
+Proof. by apply/setP=> x; apply/astab1P/normP. Qed.
+
+Lemma card_conjugates A G : #|A :^: G| = #|G : 'N_G(A)|.
+Proof. by rewrite card_orbit astab1Js. Qed.
+
+Lemma afixJG G A : (G \in 'Fix_('JG)(A)) = (A \subset 'N(G)).
+Proof. by apply/afixP/normsP=> nG x Ax; apply/eqP; move/eqP: (nG x Ax). Qed.
+
+Lemma astab1JG G : 'C[G | 'JG] = 'N(G).
+Proof.
+by apply/setP=> x; apply/astab1P/normP=> [/congr_group | /group_inj].
+Qed.
+
+Lemma dom_qactJ H : qact_dom 'J H = 'N(H).
+Proof. by rewrite qact_domE ?subsetT ?astabsJ. Qed.
+
+Lemma qactJ H (Hy : coset_of H) x :
+ 'Q%act Hy x = if x \in 'N(H) then Hy ^ coset H x else Hy.
+Proof.
+case: (cosetP Hy) => y Ny ->{Hy}.
+by rewrite qactEcond // dom_qactJ; case Nx: (x \in 'N(H)); rewrite ?morphJ.
+Qed.
+
+Lemma actsQ A B H :
+ A \subset 'N(H) -> A \subset 'N(B) -> [acts A, on B / H | 'Q].
+Proof.
+by move=> nHA nBA; rewrite acts_quotient // subsetI dom_qactJ nHA astabsJ.
+Qed.
+
+Lemma astabsQ G H : H <| G -> 'N(G / H | 'Q) = 'N(H) :&: 'N(G).
+Proof. by move=> nsHG; rewrite astabs_quotient // dom_qactJ astabsJ. Qed.
+
+Lemma astabQ H Abar : 'C(Abar |'Q) = coset H @*^-1 'C(Abar).
+Proof.
+apply/setP=> x; rewrite inE /= dom_qactJ morphpreE in_setI /=.
+apply: andb_id2l => Nx; rewrite !inE -sub1set centsC cent_set1.
+apply: eq_subset_r => {Abar} Hy; rewrite inE qactJ Nx (sameP eqP conjg_fixP).
+by rewrite (sameP cent1P eqP) (sameP commgP eqP).
+Qed.
+
+Lemma sub_astabQ A H Bbar :
+ (A \subset 'C(Bbar | 'Q)) = (A \subset 'N(H)) && (A / H \subset 'C(Bbar)).
+Proof.
+rewrite astabQ -morphpreIdom subsetI; apply: andb_id2l => nHA.
+by rewrite -sub_quotient_pre.
+Qed.
+
+Lemma sub_astabQR A B H :
+ A \subset 'N(H) -> B \subset 'N(H) ->
+ (A \subset 'C(B / H | 'Q)) = ([~: A, B] \subset H).
+Proof.
+move=> nHA nHB; rewrite sub_astabQ nHA /= (sameP commG1P eqP).
+by rewrite eqEsubset sub1G andbT -quotientR // quotient_sub1 // comm_subG.
+Qed.
+
+Lemma astabQR A H : A \subset 'N(H) ->
+ 'C(A / H | 'Q) = [set x in 'N(H) | [~: [set x], A] \subset H].
+Proof.
+move=> nHA; apply/setP=> x; rewrite astabQ -morphpreIdom 2!inE -astabQ.
+by case nHx: (x \in _); rewrite //= -sub1set sub_astabQR ?sub1set.
+Qed.
+
+Lemma quotient_astabQ H Abar : 'C(Abar | 'Q) / H = 'C(Abar).
+Proof. by rewrite astabQ cosetpreK. Qed.
+
+Lemma conj_astabQ A H x :
+ x \in 'N(H) -> 'C(A / H | 'Q) :^ x = 'C(A :^ x / H | 'Q).
+Proof.
+move=> nHx; apply/setP=> y; rewrite !astabQ mem_conjg !in_setI -mem_conjg.
+rewrite -normJ (normP nHx) quotientJ //; apply/andb_id2l => nHy.
+by rewrite !inE centJ morphJ ?groupV ?morphV // -mem_conjg.
+Qed.
+
+Section CardClass.
+
+Variable G : {group gT}.
+
+Lemma index_cent1 x : #|G : 'C_G[x]| = #|x ^: G|.
+Proof. by rewrite -astab1J -card_orbit. Qed.
+
+Lemma classes_partition : partition (classes G) G.
+Proof. by apply: orbit_partition; apply/actsP=> x Gx y; exact: groupJr. Qed.
+
+Lemma sum_card_class : \sum_(C in classes G) #|C| = #|G|.
+Proof. by apply: acts_sum_card_orbit; apply/actsP=> x Gx y; exact: groupJr. Qed.
+
+Lemma class_formula : \sum_(C in classes G) #|G : 'C_G[repr C]| = #|G|.
+Proof.
+rewrite -sum_card_class; apply: eq_bigr => _ /imsetP[x Gx ->].
+have: x \in x ^: G by rewrite -{1}(conjg1 x) mem_imset.
+by case/mem_repr/imsetP=> y Gy ->; rewrite index_cent1 classGidl.
+Qed.
+
+Lemma abelian_classP : reflect {in G, forall x, x ^: G = [set x]} (abelian G).
+Proof.
+rewrite /abelian -astabJ astabC.
+by apply: (iffP subsetP) => cGG x Gx; apply/orbit1P; exact: cGG.
+Qed.
+
+Lemma card_classes_abelian : abelian G = (#|classes G| == #|G|).
+Proof.
+have cGgt0 C: C \in classes G -> 1 <= #|C| ?= iff (#|C| == 1)%N.
+ by case/imsetP=> x _ ->; rewrite eq_sym -index_cent1.
+rewrite -sum_card_class -sum1_card (leqif_sum cGgt0).
+apply/abelian_classP/forall_inP=> [cGG _ /imsetP[x Gx ->]| cGG x Gx].
+ by rewrite cGG ?cards1.
+apply/esym/eqP; rewrite eqEcard sub1set cards1 class_refl leq_eqVlt cGG //.
+exact: mem_imset.
+Qed.
+
+End CardClass.
+
+End InternalGroupAction.
+
+Lemma gacentQ (gT : finGroupType) (H : {group gT}) (A : {set gT}) :
+ 'C_(|'Q)(A) = 'C(A / H).
+Proof.
+apply/setP=> Hx; case: (cosetP Hx) => x Nx ->{Hx}.
+rewrite -sub_cent1 -astab1J astabC sub1set -(quotientInorm H A).
+have defD: qact_dom 'J H = 'N(H) by rewrite qact_domE ?subsetT ?astabsJ.
+rewrite !(inE, mem_quotient) //= defD setIC.
+apply/subsetP/subsetP=> [cAx _ /morphimP[a Na Aa ->] | cAx a Aa].
+ by move/cAx: Aa; rewrite !inE qactE ?defD ?morphJ.
+have [_ Na] := setIP Aa; move/implyP: (cAx (coset H a)); rewrite mem_morphim //.
+by rewrite !inE qactE ?defD ?morphJ.
+Qed.
+
+Section AutAct.
+
+Variable (gT : finGroupType) (G : {set gT}).
+
+Definition autact := act ('P \ subsetT (Aut G)).
+Canonical aut_action := [action of autact].
+
+Lemma autactK a : actperm aut_action a = a.
+Proof. by apply/permP=> x; rewrite permE. Qed.
+
+Lemma autact_is_groupAction : is_groupAction G aut_action.
+Proof. by move=> a Aa /=; rewrite autactK. Qed.
+Canonical aut_groupAction := GroupAction autact_is_groupAction.
+
+End AutAct.
+
+Arguments Scope aut_action [_ group_scope].
+Arguments Scope aut_groupAction [_ group_scope].
+Notation "[ 'Aut' G ]" := (aut_action G) : action_scope.
+Notation "[ 'Aut' G ]" := (aut_groupAction G) : groupAction_scope.
+
+
diff --git a/mathcomp/fingroup/all.v b/mathcomp/fingroup/all.v
new file mode 100644
index 0000000..0cf8995
--- /dev/null
+++ b/mathcomp/fingroup/all.v
@@ -0,0 +1,10 @@
+Require Export action.
+Require Export automorphism.
+Require Export cyclic.
+Require Export fingroup.
+Require Export gproduct.
+Require Export morphism.
+Require Export perm.
+Require Export presentation.
+Require Export quotient.
+Require Export zmodp.
diff --git a/mathcomp/fingroup/automorphism.v b/mathcomp/fingroup/automorphism.v
new file mode 100644
index 0000000..1644ebf
--- /dev/null
+++ b/mathcomp/fingroup/automorphism.v
@@ -0,0 +1,489 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat fintype finset.
+Require Import fingroup perm morphism.
+
+(******************************************************************************)
+(* Group automorphisms and characteristic subgroups. *)
+(* Unlike morphisms on a group G, which are functions of type gT -> rT, with *)
+(* a canonical structure of dependent type {morphim G >-> rT}, automorphisms *)
+(* are permutations of type {perm gT} contained in Aut G : {set {perm gT}}. *)
+(* This lets us use the finGroupType of {perm gT}. Note also that while *)
+(* morphisms on G are undefined outside G, automorphisms have their support *)
+(* in G, i.e., they are the identity ouside G. *)
+(* Definitions: *)
+(* Aut G (or [Aut G]) == the automorphism group of G. *)
+(* [Aut G]%G == the group structure for Aut G. *)
+(* autm AutGa == the morphism on G induced by a, given *)
+(* AutGa : a \in Aut G. *)
+(* perm_in injf fA == the permutation with support B in induced by f, *)
+(* given injf : {in A &, injective f} and *)
+(* fA : f @: A \subset A. *)
+(* aut injf fG == the automorphism of G induced by the morphism f, *)
+(* given injf : 'injm f and fG : f @* G \subset G. *)
+(* Aut_isom injf sDom == the injective homomorphism that maps Aut G to *)
+(* Aut (f @* G), with f : {morphism D >-> rT} and *)
+(* given injf: 'injm f and sDom : G \subset D. *)
+(* conjgm G == the conjugation automorphism on G. *)
+(* H \char G == H is a characteristic subgroup of G. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+(***********************************************************************)
+(* A group automorphism, defined as a permutation on a subset of a *)
+(* finGroupType that respects the morphism law. *)
+(* Here perm_on is used as a closure rule for the set A. *)
+(***********************************************************************)
+
+Section Automorphism.
+
+Variable gT : finGroupType.
+Implicit Type A : {set gT}.
+Implicit Types a b : {perm gT}.
+
+Definition Aut A := [set a | perm_on A a & morphic A a].
+
+Lemma Aut_morphic A a : a \in Aut A -> morphic A a.
+Proof. by case/setIdP. Qed.
+
+Lemma out_Aut A a x : a \in Aut A -> x \notin A -> a x = x.
+Proof. by case/setIdP=> Aa _; exact: out_perm. Qed.
+
+Lemma eq_Aut A : {in Aut A &, forall a b, {in A, a =1 b} -> a = b}.
+Proof.
+move=> a g Aa Ag /= eqag; apply/permP=> x.
+by have [/eqag // | /out_Aut out] := boolP (x \in A); rewrite !out.
+Qed.
+
+(* The morphism that is represented by a given element of Aut A. *)
+
+Definition autm A a (AutAa : a \in Aut A) := morphm (Aut_morphic AutAa).
+Lemma autmE A a (AutAa : a \in Aut A) : autm AutAa = a.
+Proof. by []. Qed.
+
+Canonical autm_morphism A a aM := Eval hnf in [morphism of @autm A a aM].
+
+Section AutGroup.
+
+Variable G : {group gT}.
+
+Lemma Aut_group_set : group_set (Aut G).
+Proof.
+apply/group_setP; split=> [|a b].
+ by rewrite inE perm_on1; apply/morphicP=> ? *; rewrite !permE.
+rewrite !inE => /andP[Ga aM] /andP[Gb bM]; rewrite perm_onM //=.
+apply/morphicP=> x y Gx Gy; rewrite !permM (morphicP aM) //.
+by rewrite (morphicP bM) ?perm_closed.
+Qed.
+
+Canonical Aut_group := group Aut_group_set.
+
+Variable (a : {perm gT}) (AutGa : a \in Aut G).
+Notation f := (autm AutGa).
+Notation fE := (autmE AutGa).
+
+Lemma injm_autm : 'injm f.
+Proof. apply/injmP; apply: in2W; exact: perm_inj. Qed.
+
+Lemma ker_autm : 'ker f = 1. Proof. by move/trivgP: injm_autm. Qed.
+
+Lemma im_autm : f @* G = G.
+Proof.
+apply/setP=> x; rewrite morphimEdom (can_imset_pre _ (permK a)) inE.
+by have:= AutGa; rewrite inE => /andP[/perm_closed <-]; rewrite permKV.
+Qed.
+
+Lemma Aut_closed x : x \in G -> a x \in G.
+Proof. by move=> Gx; rewrite -im_autm; exact: mem_morphim. Qed.
+
+End AutGroup.
+
+Lemma Aut1 : Aut 1 = 1.
+Proof.
+apply/trivgP/subsetP=> a /= AutGa; apply/set1P.
+apply: eq_Aut (AutGa) (group1 _) _ => _ /set1P->.
+by rewrite -(autmE AutGa) morph1 perm1.
+Qed.
+
+End Automorphism.
+
+Arguments Scope Aut [_ group_scope].
+Notation "[ 'Aut' G ]" := (Aut_group G)
+ (at level 0, format "[ 'Aut' G ]") : Group_scope.
+Notation "[ 'Aut' G ]" := (Aut G)
+ (at level 0, only parsing) : group_scope.
+
+Prenex Implicits Aut autm.
+
+(* The permutation function (total on the underlying groupType) that is the *)
+(* representant of a given morphism f with domain A in (Aut A). *)
+
+Section PermIn.
+
+Variables (T : finType) (A : {set T}) (f : T -> T).
+
+Hypotheses (injf : {in A &, injective f}) (sBf : f @: A \subset A).
+
+Lemma perm_in_inj : injective (fun x => if x \in A then f x else x).
+Proof.
+move=> x y /=; wlog Ay: x y / y \in A.
+ by move=> IH eqfxy; case: ifP (eqfxy); [symmetry | case: ifP => //]; auto.
+rewrite Ay; case: ifP => [Ax | nAx def_x]; first exact: injf.
+by case/negP: nAx; rewrite def_x (subsetP sBf) ?mem_imset.
+Qed.
+
+Definition perm_in := perm perm_in_inj.
+
+Lemma perm_in_on : perm_on A perm_in.
+Proof.
+by apply/subsetP=> x; rewrite inE /= permE; case: ifP => // _; case/eqP.
+Qed.
+
+Lemma perm_inE : {in A, perm_in =1 f}.
+Proof. by move=> x Ax; rewrite /= permE Ax. Qed.
+
+End PermIn.
+
+(* properties of injective endomorphisms *)
+
+Section MakeAut.
+
+Variables (gT : finGroupType) (G : {group gT}) (f : {morphism G >-> gT}).
+Implicit Type A : {set gT}.
+
+Hypothesis injf : 'injm f.
+
+Lemma morphim_fixP A : A \subset G -> reflect (f @* A = A) (f @* A \subset A).
+Proof.
+rewrite /morphim => sAG; have:= eqEcard (f @: A) A.
+rewrite (setIidPr sAG) card_in_imset ?leqnn ?andbT => [<-|]; first exact: eqP.
+move/injmP: injf; apply: sub_in2; exact/subsetP.
+Qed.
+
+Hypothesis Gf : f @* G = G.
+
+Lemma aut_closed : f @: G \subset G.
+Proof. by rewrite -morphimEdom; exact/morphim_fixP. Qed.
+
+Definition aut := perm_in (injmP injf) aut_closed.
+
+Lemma autE : {in G, aut =1 f}.
+Proof. exact: perm_inE. Qed.
+
+Lemma morphic_aut : morphic G aut.
+Proof. by apply/morphicP=> x y Gx Gy /=; rewrite !autE ?groupM // morphM. Qed.
+
+Lemma Aut_aut : aut \in Aut G.
+Proof. by rewrite inE morphic_aut perm_in_on. Qed.
+
+Lemma imset_autE A : A \subset G -> aut @: A = f @* A.
+Proof.
+move=> sAG; rewrite /morphim (setIidPr sAG).
+apply: eq_in_imset; apply: sub_in1 autE; exact/subsetP.
+Qed.
+
+Lemma preim_autE A : A \subset G -> aut @^-1: A = f @*^-1 A.
+Proof.
+move=> sAG; apply/setP=> x; rewrite !inE permE /=.
+by case Gx: (x \in G) => //; apply/negP=> Ax; rewrite (subsetP sAG) in Gx.
+Qed.
+
+End MakeAut.
+
+Implicit Arguments morphim_fixP [gT G f].
+Prenex Implicits aut morphim_fixP.
+
+Section AutIsom.
+
+Variables (gT rT : finGroupType) (G D : {group gT}) (f : {morphism D >-> rT}).
+
+Hypotheses (injf : 'injm f) (sGD : G \subset D).
+Let domG := subsetP sGD.
+
+Lemma Aut_isom_subproof a :
+ {a' | a' \in Aut (f @* G) & a \in Aut G -> {in G, a' \o f =1 f \o a}}.
+Proof.
+set Aut_a := autm (subgP (subg [Aut G] a)).
+have aDom: 'dom (f \o Aut_a \o invm injf) = f @* G.
+ rewrite /dom /= morphpre_invm -morphpreIim; congr (f @* _).
+ by rewrite [_ :&: D](setIidPl _) ?injmK ?injm_autm ?im_autm.
+have [af [def_af ker_af _ im_af]] := domP _ aDom.
+have inj_a': 'injm af by rewrite ker_af !injm_comp ?injm_autm ?injm_invm.
+have im_a': af @* (f @* G) = f @* G.
+ by rewrite im_af !morphim_comp morphim_invm // im_autm.
+pose a' := aut inj_a' im_a'; exists a' => [|AutGa x Gx]; first exact: Aut_aut.
+have Dx := domG Gx; rewrite /= [a' _]autE ?mem_morphim //.
+by rewrite def_af /= invmE // autmE subgK.
+Qed.
+
+Definition Aut_isom a := s2val (Aut_isom_subproof a).
+
+Lemma Aut_Aut_isom a : Aut_isom a \in Aut (f @* G).
+Proof. by rewrite /Aut_isom; case: (Aut_isom_subproof a). Qed.
+
+Lemma Aut_isomE a : a \in Aut G -> {in G, forall x, Aut_isom a (f x) = f (a x)}.
+Proof. by rewrite /Aut_isom; case: (Aut_isom_subproof a). Qed.
+
+Lemma Aut_isomM : {in Aut G &, {morph Aut_isom: x y / x * y}}.
+Proof.
+move=> a b AutGa AutGb.
+apply: (eq_Aut (Aut_Aut_isom _)); rewrite ?groupM ?Aut_Aut_isom // => fx.
+case/morphimP=> x Dx Gx ->{fx}.
+by rewrite permM !Aut_isomE ?groupM /= ?permM ?Aut_closed.
+Qed.
+Canonical Aut_isom_morphism := Morphism Aut_isomM.
+
+Lemma injm_Aut_isom : 'injm Aut_isom.
+Proof.
+apply/injmP=> a b AutGa AutGb eq_ab'; apply: (eq_Aut AutGa AutGb) => x Gx.
+by apply: (injmP injf); rewrite ?domG ?Aut_closed // -!Aut_isomE //= eq_ab'.
+Qed.
+
+End AutIsom.
+
+Section InjmAut.
+
+Variables (gT rT : finGroupType) (G D : {group gT}) (f : {morphism D >-> rT}).
+
+Hypotheses (injf : 'injm f) (sGD : G \subset D).
+Let domG := subsetP sGD.
+
+Lemma im_Aut_isom : Aut_isom injf sGD @* Aut G = Aut (f @* G).
+Proof.
+apply/eqP; rewrite eqEcard; apply/andP; split.
+ by apply/subsetP=> _ /morphimP[a _ AutGa ->]; exact: Aut_Aut_isom.
+have inj_isom' := injm_Aut_isom (injm_invm injf) (morphimS _ sGD).
+rewrite card_injm ?injm_Aut_isom // -(card_injm inj_isom') ?subset_leq_card //.
+apply/subsetP=> a /morphimP[a' _ AutfGa' def_a].
+by rewrite -(morphim_invm injf sGD) def_a Aut_Aut_isom.
+Qed.
+
+Lemma Aut_isomP : isom (Aut G) (Aut (f @* G)) (Aut_isom injf sGD).
+Proof. by apply/isomP; split; [exact: injm_Aut_isom | exact: im_Aut_isom]. Qed.
+
+Lemma injm_Aut : Aut (f @* G) \isog Aut G.
+Proof. by rewrite isog_sym (isom_isog _ _ Aut_isomP). Qed.
+
+End InjmAut.
+
+(* conjugation automorphism *)
+
+Section ConjugationMorphism.
+
+Variable gT : finGroupType.
+Implicit Type A : {set gT}.
+
+Definition conjgm of {set gT} := fun x y : gT => y ^ x.
+
+Lemma conjgmE A x y : conjgm A x y = y ^ x. Proof. by []. Qed.
+
+Canonical conjgm_morphism A x :=
+ @Morphism _ _ A (conjgm A x) (in2W (fun y z => conjMg y z x)).
+
+Lemma morphim_conj A x B : conjgm A x @* B = (A :&: B) :^ x.
+Proof. by []. Qed.
+
+Variable G : {group gT}.
+
+Lemma injm_conj x : 'injm (conjgm G x).
+Proof. by apply/injmP; apply: in2W; exact: conjg_inj. Qed.
+
+Lemma conj_isom x : isom G (G :^ x) (conjgm G x).
+Proof. by apply/isomP; rewrite morphim_conj setIid injm_conj. Qed.
+
+Lemma conj_isog x : G \isog G :^ x.
+Proof. exact: isom_isog (conj_isom x). Qed.
+
+Lemma norm_conjg_im x : x \in 'N(G) -> conjgm G x @* G = G.
+Proof. by rewrite morphimEdom; exact: normP. Qed.
+
+Lemma norm_conj_isom x : x \in 'N(G) -> isom G G (conjgm G x).
+Proof. by move/norm_conjg_im/restr_isom_to/(_ (conj_isom x))->. Qed.
+
+Definition conj_aut x := aut (injm_conj _) (norm_conjg_im (subgP (subg _ x))).
+
+Lemma norm_conj_autE : {in 'N(G) & G, forall x y, conj_aut x y = y ^ x}.
+Proof. by move=> x y nGx Gy; rewrite /= autE //= subgK. Qed.
+
+Lemma conj_autE : {in G &, forall x y, conj_aut x y = y ^ x}.
+Proof. by apply: sub_in11 norm_conj_autE => //; exact: subsetP (normG G). Qed.
+
+Lemma conj_aut_morphM : {in 'N(G) &, {morph conj_aut : x y / x * y}}.
+Proof.
+move=> x y nGx nGy; apply/permP=> z /=; rewrite permM.
+case Gz: (z \in G); last by rewrite !permE /= !Gz.
+by rewrite !norm_conj_autE // (conjgM, memJ_norm, groupM).
+Qed.
+
+Canonical conj_aut_morphism := Morphism conj_aut_morphM.
+
+Lemma ker_conj_aut : 'ker conj_aut = 'C(G).
+Proof.
+apply/setP=> x; rewrite inE; case nGx: (x \in 'N(G)); last first.
+ by symmetry; apply/idP=> cGx; rewrite (subsetP (cent_sub G)) in nGx.
+rewrite 2!inE /=; apply/eqP/centP=> [cx1 y Gy | cGx].
+ by rewrite /commute (conjgC y) -norm_conj_autE // cx1 perm1.
+apply/permP=> y; case Gy: (y \in G); last by rewrite !permE Gy.
+by rewrite perm1 norm_conj_autE // conjgE -cGx ?mulKg.
+Qed.
+
+Lemma Aut_conj_aut A : conj_aut @* A \subset Aut G.
+Proof. by apply/subsetP=> _ /imsetP[x _ ->]; exact: Aut_aut. Qed.
+
+End ConjugationMorphism.
+
+Arguments Scope conjgm [_ group_scope].
+Prenex Implicits conjgm conj_aut.
+
+Reserved Notation "G \char H" (at level 70).
+
+(* Characteristic subgroup *)
+
+Section Characteristicity.
+
+Variable gT : finGroupType.
+Implicit Types A B : {set gT}.
+Implicit Types G H K L : {group gT}.
+
+Definition characteristic A B :=
+ (A \subset B) && [forall f in Aut B, f @: A \subset A].
+
+Infix "\char" := characteristic.
+
+Lemma charP H G :
+ reflect [/\ H \subset G
+ & forall f : {morphism G >-> gT},
+ 'injm f -> f @* G = G -> f @* H = H]
+ (H \char G).
+Proof.
+apply: (iffP andP) => [] [sHG chHG]; split=> //.
+ move=> f injf Gf; apply/morphim_fixP=> //.
+ by have:= forallP chHG (aut injf Gf); rewrite Aut_aut imset_autE.
+apply/forall_inP=> f Af; have injf := injm_autm Af.
+move/(morphim_fixP injf _ sHG): (chHG _ injf (im_autm Af)).
+by rewrite /morphim (setIidPr _).
+Qed.
+
+(* Characteristic subgroup properties : composition, relational properties *)
+
+Lemma char1 G : 1 \char G.
+Proof. by apply/charP; split=> [|f _ _]; rewrite (sub1G, morphim1). Qed.
+
+Lemma char_refl G : G \char G.
+Proof. exact/charP. Qed.
+
+Lemma char_trans H G K : K \char H -> H \char G -> K \char G.
+Proof.
+case/charP=> sKH chKH; case/charP=> sHG chHG.
+apply/charP; split=> [|f injf Gf]; first exact: subset_trans sHG.
+rewrite -{1}(setIidPr sKH) -(morphim_restrm sHG) chKH //.
+ rewrite ker_restrm; move/trivgP: injf => ->; exact: subsetIr.
+by rewrite morphim_restrm setIid chHG.
+Qed.
+
+Lemma char_norms H G : H \char G -> 'N(G) \subset 'N(H).
+Proof.
+case/charP=> sHG chHG; apply/normsP=> x /normP Nx.
+have:= chHG [morphism of conjgm G x] => /=.
+by rewrite !morphimEsub //=; apply; rewrite // injm_conj.
+Qed.
+
+Lemma char_sub A B : A \char B -> A \subset B.
+Proof. by case/andP. Qed.
+
+Lemma char_norm_trans H G A : H \char G -> A \subset 'N(G) -> A \subset 'N(H).
+Proof. by move/char_norms=> nHnG nGA; exact: subset_trans nHnG. Qed.
+
+Lemma char_normal_trans H G K : K \char H -> H <| G -> K <| G.
+Proof.
+move=> chKH /andP[sHG nHG].
+by rewrite /normal (subset_trans (char_sub chKH)) // (char_norm_trans chKH).
+Qed.
+
+Lemma char_normal H G : H \char G -> H <| G.
+Proof. by move/char_normal_trans; apply; apply/andP; rewrite normG. Qed.
+
+Lemma char_norm H G : H \char G -> G \subset 'N(H).
+Proof. by case/char_normal/andP. Qed.
+
+Lemma charI G H K : H \char G -> K \char G -> H :&: K \char G.
+Proof.
+case/charP=> sHG chHG; case/charP=> _ chKG.
+apply/charP; split=> [|f injf Gf]; first by rewrite subIset // sHG.
+rewrite morphimGI ?(chHG, chKG) //; exact: subset_trans (sub1G H).
+Qed.
+
+Lemma charMgen G H K : H \char G -> K \char G -> H <*> K \char G.
+Proof.
+case/charP=> sHG chHG; case/charP=> sKG chKG.
+apply/charP; split=> [|f injf Gf]; first by rewrite gen_subG subUset sHG.
+by rewrite morphim_gen ?(morphimU, subUset, sHG, chHG, chKG).
+Qed.
+
+Lemma charM G H K : H \char G -> K \char G -> H * K \char G.
+Proof.
+move=> chHG chKG; rewrite -norm_joinEl ?charMgen //.
+exact: subset_trans (char_sub chHG) (char_norm chKG).
+Qed.
+
+Lemma lone_subgroup_char G H :
+ H \subset G -> (forall K, K \subset G -> K \isog H -> K \subset H) ->
+ H \char G.
+Proof.
+move=> sHG Huniq; apply/charP; split=> // f injf Gf; apply/eqP.
+have{injf} injf: {in H &, injective f}.
+ by move/injmP: injf; apply: sub_in2; exact/subsetP.
+have fH: f @* H = f @: H by rewrite /morphim (setIidPr sHG).
+rewrite eqEcard {2}fH card_in_imset ?{}Huniq //=.
+ by rewrite -{3}Gf morphimS.
+rewrite isog_sym; apply/isogP.
+exists [morphism of restrm sHG f] => //=; first exact/injmP.
+by rewrite morphimEdom fH.
+Qed.
+
+End Characteristicity.
+
+Arguments Scope characteristic [_ group_scope group_scope].
+Notation "H \char G" := (characteristic H G) : group_scope.
+
+Section InjmChar.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+
+Hypothesis injf : 'injm f.
+
+Lemma injm_char (G H : {group aT}) :
+ G \subset D -> H \char G -> f @* H \char f @* G.
+Proof.
+move=> sGD /charP[sHG charH].
+apply/charP; split=> [|g injg gfG]; first exact: morphimS.
+have /domP[h [_ ker_h _ im_h]]: 'dom (invm injf \o g \o f) = G.
+ by rewrite /dom /= -(morphpreIim g) (setIidPl _) ?injmK // gfG morphimS.
+have hH: h @* H = H.
+ apply: charH; first by rewrite ker_h !injm_comp ?injm_invm.
+ by rewrite im_h !morphim_comp gfG morphim_invm.
+rewrite /= -{2}hH im_h !morphim_comp morphim_invmE morphpreK //.
+by rewrite (subset_trans _ (morphimS f sGD)) //= -{3}gfG !morphimS.
+Qed.
+
+End InjmChar.
+
+Section CharInjm.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+Hypothesis injf : 'injm f.
+
+Lemma char_injm (G H : {group aT}) :
+ G \subset D -> H \subset D -> (f @* H \char f @* G) = (H \char G).
+Proof.
+move=> sGD sHD; apply/idP/idP; last exact: injm_char.
+by move/(injm_char (injm_invm injf)); rewrite !morphim_invm ?morphimS // => ->.
+Qed.
+
+End CharInjm.
+
+Unset Implicit Arguments.
diff --git a/mathcomp/fingroup/cyclic.v b/mathcomp/fingroup/cyclic.v
new file mode 100644
index 0000000..ad10492
--- /dev/null
+++ b/mathcomp/fingroup/cyclic.v
@@ -0,0 +1,865 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype bigop.
+Require Import prime finset fingroup morphism perm automorphism quotient.
+Require Import gproduct ssralg finalg zmodp poly.
+
+(******************************************************************************)
+(* Properties of cyclic groups. *)
+(* Definitions: *)
+(* Defined in fingroup.v: *)
+(* <[x]> == the cycle (cyclic group) generated by x. *)
+(* #[x] == the order of x, i.e., the cardinal of <[x]>. *)
+(* Defined in prime.v: *)
+(* totient n == Euler's totient function *)
+(* Definitions in this file: *)
+(* cyclic G <=> G is a cyclic group. *)
+(* metacyclic G <=> G is a metacyclic group (i.e., a cyclic extension of a *)
+(* cyclic group). *)
+(* generator G x <=> x is a generator of the (cyclic) group G. *)
+(* Zpm x == the isomorphism mapping the additive group of integers *)
+(* mod #[x] to the cyclic group <[x]>. *)
+(* cyclem x n == the endomorphism y |-> y ^+ n of <[x]>. *)
+(* Zp_unitm x == the isomorphism mapping the multiplicative group of the *)
+(* units of the ring of integers mod #[x] to the group of *)
+(* automorphisms of <[x]> (i.e., Aut <[x]>). *)
+(* Zp_unitm x maps u to cyclem x u. *)
+(* eltm dvd_y_x == the smallest morphism (with domain <[x]>) mapping x to *)
+(* y, given a proof dvd_y_x : #[y] %| #[x]. *)
+(* expg_invn G k == if coprime #|G| k, the inverse of exponent k in G. *)
+(* Basic results for these notions, plus the classical result that any finite *)
+(* group isomorphic to a subgroup of a field is cyclic, hence that Aut G is *)
+(* cyclic when G is of prime order. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope GRing.Theory.
+
+(***********************************************************************)
+(* Cyclic groups. *)
+(***********************************************************************)
+
+Section Cyclic.
+
+Variable gT : finGroupType.
+Implicit Types (a x y : gT) (A B : {set gT}) (G K H : {group gT}).
+
+Definition cyclic A := [exists x, A == <[x]>].
+
+Lemma cyclicP A : reflect (exists x, A = <[x]>) (cyclic A).
+Proof. exact: exists_eqP. Qed.
+
+Lemma cycle_cyclic x : cyclic <[x]>.
+Proof. by apply/cyclicP; exists x. Qed.
+
+Lemma cyclic1 : cyclic [1 gT].
+Proof. by rewrite -cycle1 cycle_cyclic. Qed.
+
+(***********************************************************************)
+(* Isomorphism with the additive group *)
+(***********************************************************************)
+
+Section Zpm.
+
+Variable a : gT.
+
+Definition Zpm (i : 'Z_#[a]) := a ^+ i.
+
+Lemma ZpmM : {in Zp #[a] &, {morph Zpm : x y / x * y}}.
+Proof.
+rewrite /Zpm; case: (eqVneq a 1) => [-> | nta] i j _ _.
+ by rewrite !expg1n ?mulg1.
+by rewrite /= {3}Zp_cast ?order_gt1 // expg_mod_order expgD.
+Qed.
+
+Canonical Zpm_morphism := Morphism ZpmM.
+
+Lemma im_Zpm : Zpm @* Zp #[a] = <[a]>.
+Proof.
+apply/eqP; rewrite eq_sym eqEcard cycle_subG /= andbC morphimEdom.
+rewrite (leq_trans (leq_imset_card _ _)) ?card_Zp //= /Zp order_gt1.
+case: eqP => /= [a1 | _]; first by rewrite imset_set1 morph1 a1 set11.
+by apply/imsetP; exists 1%R; rewrite ?expg1 ?inE.
+Qed.
+
+Lemma injm_Zpm : 'injm Zpm.
+Proof.
+apply/injmP/dinjectiveP/card_uniqP.
+rewrite size_map -cardE card_Zp //= {7}/order -im_Zpm morphimEdom /=.
+by apply: eq_card => x; apply/imageP/imsetP=> [] [i Zp_i ->]; exists i.
+Qed.
+
+Lemma eq_expg_mod_order m n : (a ^+ m == a ^+ n) = (m == n %[mod #[a]]).
+Proof.
+have [->|] := eqVneq a 1; first by rewrite order1 !modn1 !expg1n eqxx.
+rewrite -order_gt1 => lt1a; have ZpT: Zp #[a] = setT by rewrite /Zp lt1a.
+have: injective Zpm by move=> i j; apply (injmP injm_Zpm); rewrite /= ZpT inE.
+move/inj_eq=> eqZ; symmetry; rewrite -(Zp_cast lt1a).
+by rewrite -[_ == _](eqZ (inZp m) (inZp n)) /Zpm /= Zp_cast ?expg_mod_order.
+Qed.
+
+Lemma Zp_isom : isom (Zp #[a]) <[a]> Zpm.
+Proof. by apply/isomP; rewrite injm_Zpm im_Zpm. Qed.
+
+Lemma Zp_isog : isog (Zp #[a]) <[a]>.
+Proof. exact: isom_isog Zp_isom. Qed.
+
+End Zpm.
+
+(***********************************************************************)
+(* Central and direct product of cycles *)
+(***********************************************************************)
+
+Lemma cyclic_abelian A : cyclic A -> abelian A.
+Proof. by case/cyclicP=> a ->; exact: cycle_abelian. Qed.
+
+Lemma cycleMsub a b :
+ commute a b -> coprime #[a] #[b] -> <[a]> \subset <[a * b]>.
+Proof.
+move=> cab co_ab; apply/subsetP=> _ /cycleP[k ->].
+apply/cycleP; exists (chinese #[a] #[b] k 0); symmetry.
+rewrite expgMn // -expg_mod_order chinese_modl // expg_mod_order.
+by rewrite /chinese addn0 -mulnA mulnCA expgM expg_order expg1n mulg1.
+Qed.
+
+Lemma cycleM a b :
+ commute a b -> coprime #[a] #[b] -> <[a * b]> = <[a]> * <[b]>.
+Proof.
+move=> cab co_ab; apply/eqP; rewrite eqEsubset -(cent_joinEl (cents_cycle cab)).
+rewrite join_subG {3}cab !cycleMsub // 1?coprime_sym //.
+by rewrite -genM_join cycle_subG mem_gen // mem_imset2 ?cycle_id.
+Qed.
+
+Lemma cyclicM A B :
+ cyclic A -> cyclic B -> B \subset 'C(A) -> coprime #|A| #|B| ->
+ cyclic (A * B).
+Proof.
+move=> /cyclicP[a ->] /cyclicP[b ->]; rewrite cent_cycle cycle_subG => cab coab.
+by rewrite -cycleM ?cycle_cyclic //; exact/esym/cent1P.
+Qed.
+
+Lemma cyclicY K H :
+ cyclic K -> cyclic H -> H \subset 'C(K) -> coprime #|K| #|H| ->
+ cyclic (K <*> H).
+Proof. by move=> cycK cycH cKH coKH; rewrite cent_joinEr // cyclicM. Qed.
+
+(***********************************************************************)
+(* Order properties *)
+(***********************************************************************)
+
+Lemma order_dvdn a n : #[a] %| n = (a ^+ n == 1).
+Proof. by rewrite (eq_expg_mod_order a n 0) mod0n. Qed.
+
+Lemma order_inf a n : a ^+ n.+1 == 1 -> #[a] <= n.+1.
+Proof. by rewrite -order_dvdn; exact: dvdn_leq. Qed.
+
+Lemma order_dvdG G a : a \in G -> #[a] %| #|G|.
+Proof. by move=> Ga; apply: cardSg; rewrite cycle_subG. Qed.
+
+Lemma expg_cardG G a : a \in G -> a ^+ #|G| = 1.
+Proof. by move=> Ga; apply/eqP; rewrite -order_dvdn order_dvdG. Qed.
+
+Lemma expg_znat G x k : x \in G -> x ^+ (k%:R : 'Z_(#|G|))%R = x ^+ k.
+Proof.
+case: (eqsVneq G 1) => [-> /set1P-> | ntG Gx]; first by rewrite !expg1n.
+apply/eqP; rewrite val_Zp_nat ?cardG_gt1 // eq_expg_mod_order.
+by rewrite modn_dvdm ?order_dvdG.
+Qed.
+
+Lemma expg_zneg G x (k : 'Z_(#|G|)) : x \in G -> x ^+ (- k)%R = x ^- k.
+Proof.
+move=> Gx; apply/eqP; rewrite eq_sym eq_invg_mul -expgD.
+by rewrite -(expg_znat _ Gx) natrD natr_Zp natr_negZp subrr.
+Qed.
+
+Lemma nt_gen_prime G x : prime #|G| -> x \in G^# -> G :=: <[x]>.
+Proof.
+move=> Gpr /setD1P[]; rewrite -cycle_subG -cycle_eq1 => ntX sXG.
+apply/eqP; rewrite eqEsubset sXG andbT.
+by apply: contraR ntX => /(prime_TIg Gpr); rewrite (setIidPr sXG) => ->.
+Qed.
+
+Lemma nt_prime_order p x : prime p -> x ^+ p = 1 -> x != 1 -> #[x] = p.
+Proof.
+move=> p_pr xp ntx; apply/prime_nt_dvdP; rewrite ?order_eq1 //.
+by rewrite order_dvdn xp.
+Qed.
+
+Lemma orderXdvd a n : #[a ^+ n] %| #[a].
+Proof. by apply: order_dvdG; exact: mem_cycle. Qed.
+
+Lemma orderXgcd a n : #[a ^+ n] = #[a] %/ gcdn #[a] n.
+Proof.
+apply/eqP; rewrite eqn_dvd; apply/andP; split.
+ rewrite order_dvdn -expgM -muln_divCA_gcd //.
+ by rewrite expgM expg_order expg1n.
+have [-> | n_gt0] := posnP n; first by rewrite gcdn0 divnn order_gt0 dvd1n.
+rewrite -(dvdn_pmul2r n_gt0) divn_mulAC ?dvdn_gcdl // dvdn_lcm.
+by rewrite order_dvdn mulnC expgM expg_order eqxx dvdn_mulr.
+Qed.
+
+Lemma orderXdiv a n : n %| #[a] -> #[a ^+ n] = #[a] %/ n.
+Proof. by case/dvdnP=> q defq; rewrite orderXgcd {2}defq gcdnC gcdnMl. Qed.
+
+Lemma orderXexp p m n x : #[x] = (p ^ n)%N -> #[x ^+ (p ^ m)] = (p ^ (n - m))%N.
+Proof.
+move=> ox; have [n_le_m | m_lt_n] := leqP n m.
+ rewrite -(subnKC n_le_m) subnDA subnn expnD expgM -ox.
+ by rewrite expg_order expg1n order1.
+rewrite orderXdiv ox ?dvdn_exp2l ?expnB ?(ltnW m_lt_n) //.
+by have:= order_gt0 x; rewrite ox expn_gt0 orbC -(ltn_predK m_lt_n).
+Qed.
+
+Lemma orderXpfactor p k n x :
+ #[x ^+ (p ^ k)] = n -> prime p -> p %| n -> #[x] = (p ^ k * n)%N.
+Proof.
+move=> oxp p_pr dv_p_n.
+suffices pk_x: p ^ k %| #[x] by rewrite -oxp orderXdiv // mulnC divnK.
+rewrite pfactor_dvdn // leqNgt; apply: contraL dv_p_n => lt_x_k.
+rewrite -oxp -p'natE // -(subnKC (ltnW lt_x_k)) expnD expgM.
+rewrite (pnat_dvd (orderXdvd _ _)) // -p_part // orderXdiv ?dvdn_part //.
+by rewrite -{1}[#[x]](partnC p) // mulKn // part_pnat.
+Qed.
+
+Lemma orderXprime p n x :
+ #[x ^+ p] = n -> prime p -> p %| n -> #[x] = (p * n)%N.
+Proof. exact: (@orderXpfactor p 1). Qed.
+
+Lemma orderXpnat m n x : #[x ^+ m] = n -> \pi(n).-nat m -> #[x] = (m * n)%N.
+Proof.
+move=> oxm n_m; have [m_gt0 _] := andP n_m.
+suffices m_x: m %| #[x] by rewrite -oxm orderXdiv // mulnC divnK.
+apply/dvdn_partP=> // p; rewrite mem_primes => /and3P[p_pr _ p_m].
+have n_p: p \in \pi(n) by apply: (pnatP _ _ n_m).
+have p_oxm: p %| #[x ^+ (p ^ logn p m)].
+ apply: dvdn_trans (orderXdvd _ m`_p^'); rewrite -expgM -p_part ?partnC //.
+ by rewrite oxm; rewrite mem_primes in n_p; case/and3P: n_p.
+by rewrite (orderXpfactor (erefl _) p_pr p_oxm) p_part // dvdn_mulr.
+Qed.
+
+Lemma orderM a b :
+ commute a b -> coprime #[a] #[b] -> #[a * b] = (#[a] * #[b])%N.
+Proof. by move=> cab co_ab; rewrite -coprime_cardMg -?cycleM. Qed.
+
+Definition expg_invn A k := (egcdn k #|A|).1.
+
+Lemma expgK G k :
+ coprime #|G| k -> {in G, cancel (expgn^~ k) (expgn^~ (expg_invn G k))}.
+Proof.
+move=> coGk x /order_dvdG Gx; apply/eqP.
+rewrite -expgM (eq_expg_mod_order _ _ 1) -(modn_dvdm 1 Gx).
+by rewrite -(chinese_modl coGk 1 0) /chinese mul1n addn0 modn_dvdm.
+Qed.
+
+Lemma cyclic_dprod K H G :
+ K \x H = G -> cyclic K -> cyclic H -> cyclic G = coprime #|K| #|H| .
+Proof.
+case/dprodP=> _ defKH cKH tiKH cycK cycH; pose m := lcmn #|K| #|H|.
+apply/idP/idP=> [/cyclicP[x defG] | coKH]; last by rewrite -defKH cyclicM.
+rewrite /coprime -dvdn1 -(@dvdn_pmul2l m) ?lcmn_gt0 ?cardG_gt0 //.
+rewrite muln_lcm_gcd muln1 -TI_cardMg // defKH defG order_dvdn.
+have /mulsgP[y z Ky Hz ->]: x \in K * H by rewrite defKH defG cycle_id.
+rewrite -[1]mulg1 expgMn; last exact/commute_sym/(centsP cKH).
+apply/eqP; congr (_ * _); apply/eqP; rewrite -order_dvdn.
+ exact: dvdn_trans (order_dvdG Ky) (dvdn_lcml _ _).
+exact: dvdn_trans (order_dvdG Hz) (dvdn_lcmr _ _).
+Qed.
+
+(***********************************************************************)
+(* Generator *)
+(***********************************************************************)
+
+Definition generator (A : {set gT}) a := A == <[a]>.
+
+Lemma generator_cycle a : generator <[a]> a.
+Proof. exact: eqxx. Qed.
+
+Lemma cycle_generator a x : generator <[a]> x -> x \in <[a]>.
+Proof. by move/(<[a]> =P _)->; exact: cycle_id. Qed.
+
+Lemma generator_order a b : generator <[a]> b -> #[a] = #[b].
+Proof. by rewrite /order => /(<[a]> =P _)->. Qed.
+
+End Cyclic.
+
+Arguments Scope cyclic [_ group_scope].
+Arguments Scope generator [_ group_scope group_scope].
+Arguments Scope expg_invn [_ group_scope nat_scope].
+Implicit Arguments cyclicP [gT A].
+Prenex Implicits cyclic Zpm generator expg_invn.
+
+(* Euler's theorem *)
+Theorem Euler_exp_totient a n : coprime a n -> a ^ totient n = 1 %[mod n].
+Proof.
+case: n => [|[|n']] //; [by rewrite !modn1 | set n := n'.+2] => co_a_n.
+have{co_a_n} Ua: coprime n (inZp a : 'I_n) by rewrite coprime_sym coprime_modl.
+have: FinRing.unit 'Z_n Ua ^+ totient n == 1.
+ by rewrite -card_units_Zp // -order_dvdn order_dvdG ?inE.
+by rewrite -2!val_eqE unit_Zp_expg /= -/n modnXm => /eqP.
+Qed.
+
+Section Eltm.
+
+Variables (aT rT : finGroupType) (x : aT) (y : rT).
+
+Definition eltm of #[y] %| #[x] := fun x_i => y ^+ invm (injm_Zpm x) x_i.
+
+Hypothesis dvd_y_x : #[y] %| #[x].
+
+Lemma eltmE i : eltm dvd_y_x (x ^+ i) = y ^+ i.
+Proof.
+apply/eqP; rewrite eq_expg_mod_order.
+have [x_le1 | x_gt1] := leqP #[x] 1.
+ suffices: #[y] %| 1 by rewrite dvdn1 => /eqP->; rewrite !modn1.
+ by rewrite (dvdn_trans dvd_y_x) // dvdn1 order_eq1 -cycle_eq1 trivg_card_le1.
+rewrite -(expg_znat i (cycle_id x)) invmE /=; last by rewrite /Zp x_gt1 inE.
+by rewrite val_Zp_nat // modn_dvdm.
+Qed.
+
+Lemma eltm_id : eltm dvd_y_x x = y. Proof. exact: (eltmE 1). Qed.
+
+Lemma eltmM : {in <[x]> &, {morph eltm dvd_y_x : x_i x_j / x_i * x_j}}.
+Proof.
+move=> _ _ /cycleP[i ->] /cycleP[j ->].
+by apply/eqP; rewrite -expgD !eltmE expgD.
+Qed.
+Canonical eltm_morphism := Morphism eltmM.
+
+Lemma im_eltm : eltm dvd_y_x @* <[x]> = <[y]>.
+Proof. by rewrite morphim_cycle ?cycle_id //= eltm_id. Qed.
+
+Lemma ker_eltm : 'ker (eltm dvd_y_x) = <[x ^+ #[y]]>.
+Proof.
+apply/eqP; rewrite eq_sym eqEcard cycle_subG 3!inE mem_cycle /= eltmE.
+rewrite expg_order eqxx (orderE y) -im_eltm card_morphim setIid -orderE.
+by rewrite orderXdiv ?dvdn_indexg //= leq_divRL ?indexg_gt0 ?Lagrange ?subsetIl.
+Qed.
+
+Lemma injm_eltm : 'injm (eltm dvd_y_x) = (#[x] %| #[y]).
+Proof. by rewrite ker_eltm subG1 cycle_eq1 -order_dvdn. Qed.
+
+End Eltm.
+
+Section CycleSubGroup.
+
+Variable gT : finGroupType.
+
+(* Gorenstein, 1.3.1 (i) *)
+Lemma cycle_sub_group (a : gT) m :
+ m %| #[a] ->
+ [set H : {group gT} | H \subset <[a]> & #|H| == m]
+ = [set <[a ^+ (#[a] %/ m)]>%G].
+Proof.
+move=> m_dv_a; have m_gt0: 0 < m by apply: dvdn_gt0 m_dv_a.
+have oam: #|<[a ^+ (#[a] %/ m)]>| = m.
+ apply/eqP; rewrite [#|_|]orderXgcd -(divnMr m_gt0) muln_gcdl divnK //.
+ by rewrite gcdnC gcdnMr mulKn.
+apply/eqP; rewrite eqEsubset sub1set inE /= cycleX oam eqxx !andbT.
+apply/subsetP=> X; rewrite in_set1 inE -val_eqE /= eqEcard oam.
+case/andP=> sXa /eqP oX; rewrite oX leqnn andbT.
+apply/subsetP=> x Xx; case/cycleP: (subsetP sXa _ Xx) => k def_x.
+have: (x ^+ m == 1)%g by rewrite -oX -order_dvdn cardSg // gen_subG sub1set.
+rewrite {x Xx}def_x -expgM -order_dvdn -[#[a]](Lagrange sXa) -oX mulnC.
+rewrite dvdn_pmul2r // mulnK // => /dvdnP[i ->].
+by rewrite mulnC expgM groupX // cycle_id.
+Qed.
+
+Lemma cycle_subgroup_char a (H : {group gT}) : H \subset <[a]> -> H \char <[a]>.
+Proof.
+move=> sHa; apply: lone_subgroup_char => // J sJa isoJH.
+have dvHa: #|H| %| #[a] by exact: cardSg.
+have{dvHa} /setP Huniq := esym (cycle_sub_group dvHa).
+move: (Huniq H) (Huniq J); rewrite !inE /=.
+by rewrite sHa sJa (card_isog isoJH) eqxx => /eqP<- /eqP<-.
+Qed.
+
+End CycleSubGroup.
+
+(***********************************************************************)
+(* Reflected boolean property and morphic image, injection, bijection *)
+(***********************************************************************)
+
+Section MorphicImage.
+
+Variables aT rT : finGroupType.
+Variables (D : {group aT}) (f : {morphism D >-> rT}) (x : aT).
+Hypothesis Dx : x \in D.
+
+Lemma morph_order : #[f x] %| #[x].
+Proof. by rewrite order_dvdn -morphX // expg_order morph1. Qed.
+
+Lemma morph_generator A : generator A x -> generator (f @* A) (f x).
+Proof. by move/(A =P _)->; rewrite /generator morphim_cycle. Qed.
+
+End MorphicImage.
+
+Section CyclicProps.
+
+Variables gT : finGroupType.
+Implicit Types (aT rT : finGroupType) (G H K : {group gT}).
+
+Lemma cyclicS G H : H \subset G -> cyclic G -> cyclic H.
+Proof.
+move=> sHG /cyclicP[x defG]; apply/cyclicP.
+exists (x ^+ (#[x] %/ #|H|)); apply/congr_group/set1P.
+by rewrite -cycle_sub_group /order -defG ?cardSg // inE sHG eqxx.
+Qed.
+
+Lemma cyclicJ G x : cyclic (G :^ x) = cyclic G.
+Proof.
+apply/cyclicP/cyclicP=> [[y /(canRL (conjsgK x))] | [y ->]].
+ by rewrite -cycleJ; exists (y ^ x^-1).
+by exists (y ^ x); rewrite cycleJ.
+Qed.
+
+Lemma eq_subG_cyclic G H K :
+ cyclic G -> H \subset G -> K \subset G -> (H :==: K) = (#|H| == #|K|).
+Proof.
+case/cyclicP=> x -> sHx sKx; apply/eqP/eqP=> [-> //| eqHK].
+have def_GHx := cycle_sub_group (cardSg sHx); set GHx := [set _] in def_GHx.
+have []: H \in GHx /\ K \in GHx by rewrite -def_GHx !inE sHx sKx eqHK /=.
+by do 2!move/set1P->.
+Qed.
+
+Lemma cardSg_cyclic G H K :
+ cyclic G -> H \subset G -> K \subset G -> (#|H| %| #|K|) = (H \subset K).
+Proof.
+move=> cycG sHG sKG; apply/idP/idP; last exact: cardSg.
+case/cyclicP: (cyclicS sKG cycG) => x defK; rewrite {K}defK in sKG *.
+case/dvdnP=> k ox; suffices ->: H :=: <[x ^+ k]> by exact: cycleX.
+apply/eqP; rewrite (eq_subG_cyclic cycG) ?(subset_trans (cycleX _ _)) //.
+rewrite -orderE orderXdiv orderE ox ?dvdn_mulr ?mulKn //.
+by have:= order_gt0 x; rewrite orderE ox; case k.
+Qed.
+
+Lemma sub_cyclic_char G H : cyclic G -> (H \char G) = (H \subset G).
+Proof.
+case/cyclicP=> x ->; apply/idP/idP => [/andP[] //|].
+exact: cycle_subgroup_char.
+Qed.
+
+Lemma morphim_cyclic rT G H (f : {morphism G >-> rT}) :
+ cyclic H -> cyclic (f @* H).
+Proof.
+move=> cycH; wlog sHG: H cycH / H \subset G.
+ by rewrite -morphimIdom; apply; rewrite (cyclicS _ cycH, subsetIl) ?subsetIr.
+case/cyclicP: cycH sHG => x ->; rewrite gen_subG sub1set => Gx.
+by apply/cyclicP; exists (f x); rewrite morphim_cycle.
+Qed.
+
+Lemma quotient_cycle x H : x \in 'N(H) -> <[x]> / H = <[coset H x]>.
+Proof. exact: morphim_cycle. Qed.
+
+Lemma quotient_cyclic G H : cyclic G -> cyclic (G / H).
+Proof. exact: morphim_cyclic. Qed.
+
+Lemma quotient_generator x G H :
+ x \in 'N(H) -> generator G x -> generator (G / H) (coset H x).
+Proof. by move=> Nx; apply: morph_generator. Qed.
+
+Lemma prime_cyclic G : prime #|G| -> cyclic G.
+Proof.
+case/primeP; rewrite ltnNge -trivg_card_le1.
+case/trivgPn=> x Gx ntx /(_ _ (order_dvdG Gx)).
+rewrite order_eq1 (negbTE ntx) => /eqnP oxG; apply/cyclicP.
+by exists x; apply/eqP; rewrite eq_sym eqEcard -oxG cycle_subG Gx leqnn.
+Qed.
+
+Lemma dvdn_prime_cyclic G p : prime p -> #|G| %| p -> cyclic G.
+Proof.
+move=> p_pr pG; case: (eqsVneq G 1) => [-> | ntG]; first exact: cyclic1.
+by rewrite prime_cyclic // (prime_nt_dvdP p_pr _ pG) -?trivg_card1.
+Qed.
+
+Lemma cyclic_small G : #|G| <= 3 -> cyclic G.
+Proof.
+rewrite 4!(ltnS, leq_eqVlt) -trivg_card_le1 orbA orbC.
+case/predU1P=> [-> | oG]; first exact: cyclic1.
+by apply: prime_cyclic; case/pred2P: oG => ->.
+Qed.
+
+End CyclicProps.
+
+Section IsoCyclic.
+
+Variables gT rT : finGroupType.
+Implicit Types (G H : {group gT}) (M : {group rT}).
+
+Lemma injm_cyclic G H (f : {morphism G >-> rT}) :
+ 'injm f -> H \subset G -> cyclic (f @* H) = cyclic H.
+Proof.
+move=> injf sHG; apply/idP/idP; last exact: morphim_cyclic.
+rewrite -{2}(morphim_invm injf sHG); exact: morphim_cyclic.
+Qed.
+
+Lemma isog_cyclic G M : G \isog M -> cyclic G = cyclic M.
+Proof. by case/isogP=> f injf <-; rewrite injm_cyclic. Qed.
+
+Lemma isog_cyclic_card G M : cyclic G -> isog G M = cyclic M && (#|M| == #|G|).
+Proof.
+move=> cycG; apply/idP/idP=> [isoGM | ].
+ by rewrite (card_isog isoGM) -(isog_cyclic isoGM) cycG /=.
+case/cyclicP: cycG => x ->{G} /andP[/cyclicP[y ->] /eqP oy].
+by apply: isog_trans (isog_symr _) (Zp_isog y); rewrite /order oy Zp_isog.
+Qed.
+
+Lemma injm_generator G H (f : {morphism G >-> rT}) x :
+ 'injm f -> x \in G -> H \subset G ->
+ generator (f @* H) (f x) = generator H x.
+Proof.
+move=> injf Gx sHG; apply/idP/idP; last exact: morph_generator.
+rewrite -{2}(morphim_invm injf sHG) -{2}(invmE injf Gx).
+by apply: morph_generator; exact: mem_morphim.
+Qed.
+
+End IsoCyclic.
+
+(* Metacyclic groups. *)
+Section Metacyclic.
+
+Variable gT : finGroupType.
+Implicit Types (A : {set gT}) (G H : {group gT}).
+
+Definition metacyclic A :=
+ [exists H : {group gT}, [&& cyclic H, H <| A & cyclic (A / H)]].
+
+Lemma metacyclicP A :
+ reflect (exists H : {group gT}, [/\ cyclic H, H <| A & cyclic (A / H)])
+ (metacyclic A).
+Proof. exact: 'exists_and3P. Qed.
+
+Lemma metacyclic1 : metacyclic 1.
+Proof.
+by apply/existsP; exists 1%G; rewrite normal1 trivg_quotient !cyclic1.
+Qed.
+
+Lemma cyclic_metacyclic A : cyclic A -> metacyclic A.
+Proof.
+case/cyclicP=> x ->; apply/existsP; exists (<[x]>)%G.
+by rewrite normal_refl cycle_cyclic trivg_quotient cyclic1.
+Qed.
+
+Lemma metacyclicS G H : H \subset G -> metacyclic G -> metacyclic H.
+Proof.
+move=> sHG /metacyclicP[K [cycK nsKG cycGq]]; apply/metacyclicP.
+exists (H :&: K)%G; rewrite (cyclicS (subsetIr H K)) ?(normalGI sHG) //=.
+rewrite setIC (isog_cyclic (second_isog _)) ?(cyclicS _ cycGq) ?quotientS //.
+by rewrite (subset_trans sHG) ?normal_norm.
+Qed.
+
+End Metacyclic.
+
+Arguments Scope metacyclic [_ group_scope].
+Prenex Implicits metacyclic.
+Implicit Arguments metacyclicP [gT A].
+
+(* Automorphisms of cyclic groups. *)
+Section CyclicAutomorphism.
+
+Variable gT : finGroupType.
+
+Section CycleAutomorphism.
+
+Variable a : gT.
+
+Section CycleMorphism.
+
+Variable n : nat.
+
+Definition cyclem of gT := fun x : gT => x ^+ n.
+
+Lemma cyclemM : {in <[a]> & , {morph cyclem a : x y / x * y}}.
+Proof.
+by move=> x y ax ay; apply: expgMn; exact: (centsP (cycle_abelian a)).
+Qed.
+
+Canonical cyclem_morphism := Morphism cyclemM.
+
+End CycleMorphism.
+
+Section ZpUnitMorphism.
+
+Variable u : {unit 'Z_#[a]}.
+
+Lemma injm_cyclem : 'injm (cyclem (val u) a).
+Proof.
+apply/subsetP=> x /setIdP[ax]; rewrite !inE -order_dvdn.
+case: (eqVneq a 1) => [a1 | nta]; first by rewrite a1 cycle1 inE in ax.
+rewrite -order_eq1 -dvdn1; move/eqnP: (valP u) => /= <-.
+by rewrite dvdn_gcd {2}Zp_cast ?order_gt1 // order_dvdG.
+Qed.
+
+Lemma im_cyclem : cyclem (val u) a @* <[a]> = <[a]>.
+Proof.
+apply/morphim_fixP=> //; first exact: injm_cyclem.
+by rewrite morphim_cycle ?cycle_id ?cycleX.
+Qed.
+
+Definition Zp_unitm := aut injm_cyclem im_cyclem.
+
+End ZpUnitMorphism.
+
+Lemma Zp_unitmM : {in units_Zp #[a] &, {morph Zp_unitm : u v / u * v}}.
+Proof.
+move=> u v _ _; apply: (eq_Aut (Aut_aut _ _)) => [|x a_x].
+ by rewrite groupM ?Aut_aut.
+rewrite permM !autE ?groupX //= /cyclem -expgM.
+rewrite -expg_mod_order modn_dvdm ?expg_mod_order //.
+case: (leqP #[a] 1) => [lea1 | lt1a]; last by rewrite Zp_cast ?order_dvdG.
+by rewrite card_le1_trivg // in a_x; rewrite (set1P a_x) order1 dvd1n.
+Qed.
+
+Canonical Zp_unit_morphism := Morphism Zp_unitmM.
+
+Lemma injm_Zp_unitm : 'injm Zp_unitm.
+Proof.
+case: (eqVneq a 1) => [a1 | nta].
+ by rewrite subIset //= card_le1_trivg ?subxx // card_units_Zp a1 order1.
+apply/subsetP=> /= u /morphpreP[_ /set1P/= um1].
+have{um1}: Zp_unitm u a == Zp_unitm 1 a by rewrite um1 morph1.
+rewrite !autE ?cycle_id // eq_expg_mod_order.
+by rewrite -[n in _ == _ %[mod n]]Zp_cast ?order_gt1 // !modZp inE.
+Qed.
+
+Lemma generator_coprime m : generator <[a]> (a ^+ m) = coprime #[a] m.
+Proof.
+rewrite /generator eq_sym eqEcard cycleX -/#[a] [#|_|]orderXgcd /=.
+apply/idP/idP=> [le_a_am|co_am]; last by rewrite (eqnP co_am) divn1.
+have am_gt0: 0 < gcdn #[a] m by rewrite gcdn_gt0 order_gt0.
+by rewrite /coprime eqn_leq am_gt0 andbT -(@leq_pmul2l #[a]) ?muln1 -?leq_divRL.
+Qed.
+
+Lemma im_Zp_unitm : Zp_unitm @* units_Zp #[a] = Aut <[a]>.
+Proof.
+rewrite morphimEdom; apply/setP=> f; pose n := invm (injm_Zpm a) (f a).
+apply/imsetP/idP=> [[u _ ->] | Af]; first exact: Aut_aut.
+have [a1 | nta] := eqVneq a 1.
+ by rewrite a1 cycle1 Aut1 in Af; exists 1; rewrite // morph1 (set1P Af).
+have a_fa: <[a]> = <[f a]>.
+ by rewrite -(autmE Af) -morphim_cycle ?im_autm ?cycle_id.
+have def_n: a ^+ n = f a.
+ by rewrite -/(Zpm n) invmK // im_Zpm a_fa cycle_id.
+have co_a_n: coprime #[a].-2.+2 n.
+ by rewrite {1}Zp_cast ?order_gt1 // -generator_coprime def_n; exact/eqP.
+exists (FinRing.unit 'Z_#[a] co_a_n); rewrite ?inE //.
+apply: eq_Aut (Af) (Aut_aut _ _) _ => x ax.
+rewrite autE //= /cyclem; case/cycleP: ax => k ->{x}.
+by rewrite -(autmE Af) morphX ?cycle_id //= autmE -def_n -!expgM mulnC.
+Qed.
+
+Lemma Zp_unit_isom : isom (units_Zp #[a]) (Aut <[a]>) Zp_unitm.
+Proof. by apply/isomP; rewrite ?injm_Zp_unitm ?im_Zp_unitm. Qed.
+
+Lemma Zp_unit_isog : isog (units_Zp #[a]) (Aut <[a]>).
+Proof. exact: isom_isog Zp_unit_isom. Qed.
+
+Lemma card_Aut_cycle : #|Aut <[a]>| = totient #[a].
+Proof. by rewrite -(card_isog Zp_unit_isog) card_units_Zp. Qed.
+
+Lemma totient_gen : totient #[a] = #|[set x | generator <[a]> x]|.
+Proof.
+have [lea1 | lt1a] := leqP #[a] 1.
+ rewrite /order card_le1_trivg // cards1 (@eq_card1 _ 1) // => x.
+ by rewrite !inE -cycle_eq1 eq_sym.
+rewrite -(card_injm (injm_invm (injm_Zpm a))) /= ?im_Zpm; last first.
+ by apply/subsetP=> x; rewrite inE; exact: cycle_generator.
+rewrite -card_units_Zp // cardsE card_sub morphim_invmE; apply: eq_card => /= d.
+by rewrite !inE /= qualifE /= /Zp lt1a inE /= generator_coprime {1}Zp_cast.
+Qed.
+
+Lemma Aut_cycle_abelian : abelian (Aut <[a]>).
+Proof. by rewrite -im_Zp_unitm morphim_abelian ?units_Zp_abelian. Qed.
+
+End CycleAutomorphism.
+
+Variable G : {group gT}.
+
+Lemma Aut_cyclic_abelian : cyclic G -> abelian (Aut G).
+Proof. by case/cyclicP=> x ->; exact: Aut_cycle_abelian. Qed.
+
+Lemma card_Aut_cyclic : cyclic G -> #|Aut G| = totient #|G|.
+Proof. by case/cyclicP=> x ->; exact: card_Aut_cycle. Qed.
+
+Lemma sum_ncycle_totient :
+ \sum_(d < #|G|.+1) #|[set <[x]> | x in G & #[x] == d]| * totient d = #|G|.
+Proof.
+pose h (x : gT) : 'I_#|G|.+1 := inord #[x].
+symmetry; rewrite -{1}sum1_card (partition_big h xpredT) //=.
+apply: eq_bigr => d _; set Gd := finset _.
+rewrite -sum_nat_const sum1dep_card -sum1_card (_ : finset _ = Gd); last first.
+ apply/setP=> x; rewrite !inE; apply: andb_id2l => Gx.
+ by rewrite /eq_op /= inordK // ltnS subset_leq_card ?cycle_subG.
+rewrite (partition_big_imset cycle) {}/Gd; apply: eq_bigr => C /=.
+case/imsetP=> x /setIdP[Gx /eqP <-] -> {C d}.
+rewrite sum1dep_card totient_gen; apply: eq_card => y; rewrite !inE /generator.
+move: Gx; rewrite andbC eq_sym -!cycle_subG /order.
+by case: eqP => // -> ->; rewrite eqxx.
+Qed.
+
+End CyclicAutomorphism.
+
+Lemma sum_totient_dvd n : \sum_(d < n.+1 | d %| n) totient d = n.
+Proof.
+case: n => [|[|n']]; try by rewrite big_mkcond !big_ord_recl big_ord0.
+set n := n'.+2; pose x1 : 'Z_n := 1%R.
+have ox1: #[x1] = n by rewrite /order -Zp_cycle card_Zp.
+rewrite -[rhs in _ = rhs]ox1 -[#[_]]sum_ncycle_totient [#|_|]ox1 big_mkcond /=.
+apply: eq_bigr => d _; rewrite -{2}ox1; case: ifP => [|ndv_dG]; last first.
+ rewrite eq_card0 // => C; apply/imsetP=> [[x /setIdP[Gx oxd] _{C}]].
+ by rewrite -(eqP oxd) order_dvdG in ndv_dG.
+move/cycle_sub_group; set Gd := [set _] => def_Gd.
+rewrite (_ : _ @: _ = @gval _ @: Gd); first by rewrite imset_set1 cards1 mul1n.
+apply/setP=> C; apply/idP/imsetP=> [| [gC GdC ->{C}]].
+ case/imsetP=> x /setIdP[_ oxd] ->; exists <[x]>%G => //.
+ by rewrite -def_Gd inE -Zp_cycle subsetT.
+have:= GdC; rewrite -def_Gd => /setIdP[_ /eqP <-].
+by rewrite (set1P GdC) /= mem_imset // inE eqxx (mem_cycle x1).
+Qed.
+
+Section FieldMulCyclic.
+
+(***********************************************************************)
+(* A classic application to finite multiplicative subgroups of fields. *)
+(***********************************************************************)
+
+Import GRing.Theory.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Lemma order_inj_cyclic :
+ {in G &, forall x y, #[x] = #[y] -> <[x]> = <[y]>} -> cyclic G.
+Proof.
+move=> ucG; apply: negbNE (contra _ (negbT (ltnn #|G|))) => ncG.
+rewrite -{2}[#|G|]sum_totient_dvd big_mkcond (bigD1 ord_max) ?dvdnn //=.
+rewrite -{1}[#|G|]sum_ncycle_totient (bigD1 ord_max) //= -addSn leq_add //.
+ rewrite eq_card0 ?totient_gt0 ?cardG_gt0 // => C.
+ apply/imsetP=> [[x /setIdP[Gx /eqP oxG]]]; case/cyclicP: ncG.
+ by exists x; apply/eqP; rewrite eq_sym eqEcard cycle_subG Gx -oxG /=.
+elim/big_ind2: _ => // [m1 n1 m2 n2 | d _]; first exact: leq_add.
+set Gd := _ @: _; case: (set_0Vmem Gd) => [-> | [C]]; first by rewrite cards0.
+rewrite {}/Gd => /imsetP[x /setIdP[Gx /eqP <-] _ {C d}].
+rewrite order_dvdG // (@eq_card1 _ <[x]>) ?mul1n // => C.
+apply/idP/eqP=> [|-> {C}]; last by rewrite mem_imset // inE Gx eqxx.
+by case/imsetP=> y /setIdP[Gy /eqP/ucG->].
+Qed.
+
+Lemma div_ring_mul_group_cyclic (R : unitRingType) (f : gT -> R) :
+ f 1 = 1%R -> {in G &, {morph f : u v / u * v >-> (u * v)%R}} ->
+ {in G^#, forall x, f x - 1 \in GRing.unit}%R ->
+ abelian G -> cyclic G.
+Proof.
+move=> f1 fM f1P abelG.
+have fX n: {in G, {morph f : u / u ^+ n >-> (u ^+ n)%R}}.
+ by case: n => // n x Gx; elim: n => //= n IHn; rewrite expgS fM ?groupX ?IHn.
+have fU x: x \in G -> f x \in GRing.unit.
+ by move=> Gx; apply/unitrP; exists (f x^-1); rewrite -!fM ?groupV ?gsimp.
+apply: order_inj_cyclic => x y Gx Gy; set n := #[x] => yn.
+apply/eqP; rewrite eq_sym eqEcard -[#|_|]/n yn leqnn andbT cycle_subG /=.
+suff{y Gy yn} ->: <[x]> = G :&: [set z | #[z] %| n] by rewrite !inE Gy yn /=.
+apply/eqP; rewrite eqEcard subsetI cycle_subG {}Gx /= cardE; set rs := enum _.
+apply/andP; split; first by apply/subsetP=> y xy; rewrite inE order_dvdG.
+pose P : {poly R} := ('X^n - 1)%R; have n_gt0: n > 0 by exact: order_gt0.
+have szP: size P = n.+1 by rewrite size_addl size_polyXn ?size_opp ?size_poly1.
+rewrite -ltnS -szP -(size_map f) max_ring_poly_roots -?size_poly_eq0 ?{}szP //.
+ apply/allP=> fy /mapP[y]; rewrite mem_enum !inE order_dvdn => /andP[Gy].
+ move/eqP=> yn1 ->{fy}; apply/eqP.
+ by rewrite !(hornerE, hornerXn) -fX // yn1 f1 subrr.
+have: uniq rs by exact: enum_uniq.
+have: all (mem G) rs by apply/allP=> y; rewrite mem_enum; case/setIP.
+elim: rs => //= y rs IHrs /andP[Gy Grs] /andP[y_rs]; rewrite andbC.
+move/IHrs=> -> {IHrs}//; apply/allP=> _ /mapP[z rs_z ->].
+have{Grs} Gz := allP Grs z rs_z; rewrite /diff_roots -!fM // (centsP abelG) //.
+rewrite eqxx -[f y]mul1r -(mulgKV y z) fM ?groupM ?groupV //=.
+rewrite -mulNr -mulrDl unitrMl ?fU ?f1P // !inE.
+by rewrite groupM ?groupV // andbT -eq_mulgV1; apply: contra y_rs; move/eqP <-.
+Qed.
+
+Lemma field_mul_group_cyclic (F : fieldType) (f : gT -> F) :
+ {in G &, {morph f : u v / u * v >-> (u * v)%R}} ->
+ {in G, forall x, f x = 1%R <-> x = 1} ->
+ cyclic G.
+Proof.
+move=> fM f1P; have f1 : f 1 = 1%R by exact/f1P.
+apply: (div_ring_mul_group_cyclic f1 fM) => [x|].
+ case/setD1P=> x1 Gx; rewrite unitfE; apply: contra x1.
+ by rewrite subr_eq0 => /eqP/f1P->.
+apply/centsP=> x Gx y Gy; apply/commgP/eqP.
+apply/f1P; rewrite ?fM ?groupM ?groupV //.
+by rewrite mulrCA -!fM ?groupM ?groupV // mulKg mulVg.
+Qed.
+
+End FieldMulCyclic.
+
+Lemma field_unit_group_cyclic (F : finFieldType) (G : {group {unit F}}) :
+ cyclic G.
+Proof.
+apply: field_mul_group_cyclic FinRing.uval _ _ => // u _.
+by split=> /eqP ?; exact/eqP.
+Qed.
+
+Section PrimitiveRoots.
+
+Open Scope ring_scope.
+Import GRing.Theory.
+
+Lemma has_prim_root (F : fieldType) (n : nat) (rs : seq F) :
+ n > 0 -> all n.-unity_root rs -> uniq rs -> size rs >= n ->
+ has n.-primitive_root rs.
+Proof.
+move=> n_gt0 rsn1 Urs; rewrite leq_eqVlt ltnNge max_unity_roots // orbF eq_sym.
+move/eqP=> sz_rs; pose r := val (_ : seq_sub rs).
+have rn1 x: r x ^+ n = 1.
+ by apply/eqP; rewrite -unity_rootE (allP rsn1) ?(valP x).
+have prim_r z: z ^+ n = 1 -> z \in rs.
+ by move/eqP; rewrite -unity_rootE -(mem_unity_roots n_gt0).
+pose r' := SeqSub (prim_r _ _); pose sG_1 := r' _ (expr1n _ _).
+have sG_VP: r _ ^+ n.-1 ^+ n = 1.
+ by move=> x; rewrite -exprM mulnC exprM rn1 expr1n.
+have sG_MP: (r _ * r _) ^+ n = 1 by move=> x y; rewrite exprMn !rn1 mul1r.
+pose sG_V := r' _ (sG_VP _); pose sG_M := r' _ (sG_MP _ _).
+have sG_Ag: associative sG_M by move=> x y z; apply: val_inj; rewrite /= mulrA.
+have sG_1g: left_id sG_1 sG_M by move=> x; apply: val_inj; rewrite /= mul1r.
+have sG_Vg: left_inverse sG_1 sG_V sG_M.
+ by move=> x; apply: val_inj; rewrite /= -exprSr prednK ?rn1.
+pose sgT := BaseFinGroupType _ (FinGroup.Mixin sG_Ag sG_1g sG_Vg).
+pose gT := @FinGroupType sgT sG_Vg.
+have /cyclicP[x gen_x]: @cyclic gT setT.
+ apply: (@field_mul_group_cyclic gT [set: _] F r) => // x _.
+ by split=> [ri1 | ->]; first exact: val_inj.
+apply/hasP; exists (r x); first exact: (valP x).
+have [m prim_x dvdmn] := prim_order_exists n_gt0 (rn1 x).
+rewrite -((m =P n) _) // eqn_dvd {}dvdmn -sz_rs -(card_seq_sub Urs) -cardsT.
+rewrite gen_x (@order_dvdn gT) /(_ == _) /= -{prim_x}(prim_expr_order prim_x).
+by apply/eqP; elim: m => //= m IHm; rewrite exprS expgS /= -IHm.
+Qed.
+
+End PrimitiveRoots.
+
+(***********************************************************************)
+(* Cycles of prime order *)
+(***********************************************************************)
+
+Section AutPrime.
+
+Variable gT : finGroupType.
+
+Lemma Aut_prime_cycle_cyclic (a : gT) : prime #[a] -> cyclic (Aut <[a]>).
+Proof.
+move=> pr_a; have inj_um := injm_Zp_unitm a; have eq_a := Fp_Zcast pr_a.
+pose fm := cast_ord (esym eq_a) \o val \o invm inj_um.
+apply: (@field_mul_group_cyclic _ _ _ fm) => [f g Af Ag | f Af] /=.
+ by apply: val_inj; rewrite /= morphM ?im_Zp_unitm //= eq_a.
+split=> [/= fm1 |->]; last by apply: val_inj; rewrite /= morph1.
+apply: (injm1 (injm_invm inj_um)); first by rewrite /= im_Zp_unitm.
+by do 2!apply: val_inj; move/(congr1 val): fm1.
+Qed.
+
+Lemma Aut_prime_cyclic (G : {group gT}) : prime #|G| -> cyclic (Aut G).
+Proof.
+move=> pr_G; case/cyclicP: (prime_cyclic pr_G) (pr_G) => x ->.
+exact: Aut_prime_cycle_cyclic.
+Qed.
+
+End AutPrime.
diff --git a/mathcomp/fingroup/fingroup.v b/mathcomp/fingroup/fingroup.v
new file mode 100644
index 0000000..fe78559
--- /dev/null
+++ b/mathcomp/fingroup/fingroup.v
@@ -0,0 +1,3096 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq choice fintype.
+Require Import div path bigop prime finset.
+
+(******************************************************************************)
+(* This file defines the main interface for finite groups : *)
+(* finGroupType == the structure for finite types with a group law. *)
+(* {group gT} == type of groups with elements of type gT. *)
+(* baseFinGroupType == the structure for finite types with a monoid law *)
+(* and an involutive antimorphism; finGroupType is *)
+(* derived from baseFinGroupType (via a telescope). *)
+(* FinGroupType mulVg == the finGroupType structure for an existing *)
+(* baseFinGroupType structure, built from a proof of *)
+(* the left inverse group axiom for that structure's *)
+(* operations. *)
+(* BaseFinGroupType bgm == the baseFingroupType structure built by packaging *)
+(* bgm : FinGroup.mixin_of T for a type T with an *)
+(* existing finType structure. *)
+(* FinGroup.BaseMixin mulA mul1x invK invM == *)
+(* the mixin for a baseFinGroupType structure, built *)
+(* from proofs of the baseFinGroupType axioms. *)
+(* FinGroup.Mixin mulA mul1x mulVg == *)
+(* the mixin for a baseFinGroupType structure, built *)
+(* from proofs of the group axioms. *)
+(* [baseFinGroupType of T] == a clone of an existing baseFinGroupType *)
+(* structure on T, for T (the existing structure *)
+(* might be for som delta-expansion of T). *)
+(* [finGroupType of T] == a clone of an existing finGroupType structure on *)
+(* T, for the canonical baseFinGroupType structure *)
+(* of T (the existing structure might be for the *)
+(* baseFinGroupType of some delta-expansion of T). *)
+(* [group of G] == a clone for an existing {group gT} structure on *)
+(* G : {set gT} (the existing structure might be for *)
+(* some delta-expansion of G). *)
+(* If gT implements finGroupType, then we can form {set gT}, the type of *)
+(* finite sets with elements of type gT (as finGroupType extends finType). *)
+(* The group law extends pointwise to {set gT}, which thus implements a sub- *)
+(* interface baseFinGroupType of finGroupType. To be consistent with the *)
+(* predType interface, this is done by coercion to FinGroup.arg_sort, an *)
+(* alias for FinGroup.sort. Accordingly, all pointwise group operations below *)
+(* have arguments of type (FinGroup.arg_sort) gT and return results of type *)
+(* FinGroup.sort gT. *)
+(* The notations below are declared in two scopes: *)
+(* group_scope (delimiter %g) for point operations and set constructs. *)
+(* Group_scope (delimiter %G) for explicit {group gT} structures. *)
+(* These scopes should not be opened globally, although group_scope is often *)
+(* opened locally in group-theory files (via Import GroupScope). *)
+(* As {group gT} is both a subtype and an interface structure for {set gT}, *)
+(* the fact that a given G : {set gT} is a group can (and usually should) be *)
+(* inferred by type inference with canonical structures. This means that all *)
+(* `group' constructions (e.g., the normaliser 'N_G(H)) actually define sets *)
+(* with a canonical {group gT} structure; the %G delimiter can be used to *)
+(* specify the actual {group gT} structure (e.g., 'N_G(H)%G). *)
+(* Operations on elements of a group: *)
+(* x * y == the group product of x and y. *)
+(* x ^+ n == the nth power of x, i.e., x * ... * x (n times). *)
+(* x^-1 == the group inverse of x. *)
+(* x ^- n == the inverse of x ^+ n (notation for (x ^+ n)^-1). *)
+(* 1 == the unit element. *)
+(* x ^ y == the conjugate of x by y. *)
+(* \prod_(i ...) x i == the product of the x i (order-sensitive). *)
+(* commute x y <-> x and y commute. *)
+(* centralises x A <-> x centralises A. *)
+(* 'C[x] == the set of elements that commute with x. *)
+(* 'C_G[x] == the set of elements of G that commute with x. *)
+(* <[x]> == the cyclic subgroup generated by the element x. *)
+(* #[x] == the order of the element x, i.e., #|<[x]>|. *)
+(* [~ x1, ..., xn] == the commutator of x1, ..., xn. *)
+(* Operations on subsets/subgroups of a finite group: *)
+(* H * G == {xy | x \in H, y \in G}. *)
+(* 1 or [1] or [1 gT] == the unit group. *)
+(* [set: gT]%G == the group of all x : gT (in Group_scope). *)
+(* [subg G] == the subtype, set, or group of all x \in G: this *)
+(* notation is defined simultaneously in %type, %g *)
+(* and %G scopes, and G must denote a {group gT} *)
+(* structure (G is in the %G scope). *)
+(* subg, sgval == the projection into and injection from [subg G]. *)
+(* H^# == the set H minus the unit element *)
+(* repr H == some element of H if 1 \notin H != set0, else 1. *)
+(* (repr is defined over sets of a baseFinGroupType, *)
+(* so it can be used, e.g., to pick right cosets.) *)
+(* x *: H == left coset of H by x. *)
+(* lcosets H G == the set of the left cosets of H by elements of G. *)
+(* H :* x == right coset of H by x. *)
+(* rcosets H G == the set of the right cosets of H by elements of G. *)
+(* #|G : H| == the index of H in G, i.e., #|rcosets G H|. *)
+(* H :^ x == the conjugate of H by x. *)
+(* x ^: H == the conjugate class of x in H. *)
+(* classes G == the set of all conjugate classes of G. *)
+(* G :^: H == {G :^ x | x \in H}. *)
+(* class_support G H == {x ^ y | x \in G, y \in H}. *)
+(* [~: H1, ..., Hn] == commutator subgroup of H1, ..., Hn. *)
+(*{in G, centralised H} <-> G centralises H. *)
+(* {in G, normalised H} <-> G normalises H. *)
+(* <-> forall x, x \in G -> H :^ x = H. *)
+(* 'N(H) == the normaliser of H. *)
+(* 'N_G(H) == the normaliser of H in G. *)
+(* H <| G <=> H is normal in G. *)
+(* 'C(H) == the centraliser of H. *)
+(* 'C_G(H) == the centraliser of H in G. *)
+(* <<H>> == the subgroup generated by the set H. *)
+(* H <*> G == the subgroup generated by sets H and G (H join G). *)
+(* (H * G)%G == the join of G H : {group gT} (convertible, but not *)
+(* identical to (G <*> H)%G). *)
+(* (\prod_(i ...) H i)%G == the group generated by the H i. *)
+(* gcore H G == the largest subgroup of H normalised by G. *)
+(* If H is a subgroup of G, this is the largest *)
+(* normal subgroup of G contained in H). *)
+(* abelian H <=> H is abelian. *)
+(* subgroups G == the set of subgroups of G, i.e., the set of all *)
+(* H : {group gT} such that H \subset G. *)
+(* In the notation below G is a variable that is bound in P. *)
+(* [max G | P] <=> G is the largest group such that P holds. *)
+(* [max H of G | P] <=> H is the largest group G such that P holds. *)
+(* [max G | P & Q] := [max G | P && Q], likewise [max H of G | P & Q]. *)
+(* [min G | P] <=> G is the smallest group such that P holds. *)
+(* [min G | P & Q] := [min G | P && Q], likewise [min H of G | P & Q]. *)
+(* [min H of G | P] <=> H is the smallest group G such that P holds. *)
+(* In addition to the generic suffixes described in ssrbool.v and finset.v, *)
+(* we associate the following suffixes to group operations: *)
+(* 1 - identity element, as in group1 : 1 \in G. *)
+(* M - multiplication, as is invMg : (x * y)^-1 = x^-1 * y^-1. *)
+(* Also nat multiplication, for expgM : x ^+ (m * n) = x ^+ m ^+ n. *)
+(* D - (nat) addition, for expgD : x ^+ (m + n) = x ^+ m * x ^+ n. *)
+(* V - inverse, as in mulgV : x * x^-1 = 1. *)
+(* X - exponentiation, as in conjXg : (x ^+ n) ^ y = (x ^ y) ^+ n. *)
+(* J - conjugation, as in orderJ : #[x ^ y] = #[x]. *)
+(* R - commutator, as in conjRg : [~ x, y] ^ z = [~ x ^ z, y ^ z]. *)
+(* Y - join, as in centY : 'C(G <*> H) = 'C(G) :&: 'C(H). *)
+(* We sometimes prefix these with an `s' to indicate a set-lifted operation, *)
+(* e.g., conjsMg : (A * B) :^ x = A :^ x * B :^ x. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Delimit Scope group_scope with g.
+Delimit Scope Group_scope with G.
+
+(* This module can be imported to open the scope for group element *)
+(* operations locally to a file, without exporing the Open to *)
+(* clients of that file (as Open would do). *)
+Module GroupScope.
+Open Scope group_scope.
+End GroupScope.
+Import GroupScope.
+
+(* These are the operation notations introduced by this file. *)
+Reserved Notation "[ ~ x1 , x2 , .. , xn ]" (at level 0,
+ format "'[ ' [ ~ x1 , '/' x2 , '/' .. , '/' xn ] ']'").
+Reserved Notation "[ 1 gT ]" (at level 0, format "[ 1 gT ]").
+Reserved Notation "[ 1 ]" (at level 0, format "[ 1 ]").
+Reserved Notation "[ 'subg' G ]" (at level 0, format "[ 'subg' G ]").
+Reserved Notation "A ^#" (at level 2, format "A ^#").
+Reserved Notation "A :^ x" (at level 35, right associativity).
+Reserved Notation "x ^: B" (at level 35, right associativity).
+Reserved Notation "A :^: B" (at level 35, right associativity).
+Reserved Notation "#| B : A |" (at level 0, B, A at level 99,
+ format "#| B : A |").
+Reserved Notation "''N' ( A )" (at level 8, format "''N' ( A )").
+Reserved Notation "''N_' G ( A )" (at level 8, G at level 2,
+ format "''N_' G ( A )").
+Reserved Notation "A <| B" (at level 70, no associativity).
+Reserved Notation "#[ x ]" (at level 0, format "#[ x ]").
+Reserved Notation "A <*> B" (at level 40, left associativity).
+Reserved Notation "[ ~: A1 , A2 , .. , An ]" (at level 0,
+ format "[ ~: '[' A1 , '/' A2 , '/' .. , '/' An ']' ]").
+Reserved Notation "[ 'max' A 'of' G | gP ]" (at level 0,
+ format "[ '[hv' 'max' A 'of' G '/ ' | gP ']' ]").
+Reserved Notation "[ 'max' G | gP ]" (at level 0,
+ format "[ '[hv' 'max' G '/ ' | gP ']' ]").
+Reserved Notation "[ 'max' A 'of' G | gP & gQ ]" (at level 0,
+ format "[ '[hv' 'max' A 'of' G '/ ' | gP '/ ' & gQ ']' ]").
+Reserved Notation "[ 'max' G | gP & gQ ]" (at level 0,
+ format "[ '[hv' 'max' G '/ ' | gP '/ ' & gQ ']' ]").
+Reserved Notation "[ 'min' A 'of' G | gP ]" (at level 0,
+ format "[ '[hv' 'min' A 'of' G '/ ' | gP ']' ]").
+Reserved Notation "[ 'min' G | gP ]" (at level 0,
+ format "[ '[hv' 'min' G '/ ' | gP ']' ]").
+Reserved Notation "[ 'min' A 'of' G | gP & gQ ]" (at level 0,
+ format "[ '[hv' 'min' A 'of' G '/ ' | gP '/ ' & gQ ']' ]").
+Reserved Notation "[ 'min' G | gP & gQ ]" (at level 0,
+ format "[ '[hv' 'min' G '/ ' | gP '/ ' & gQ ']' ]").
+
+Module FinGroup.
+
+(* We split the group axiomatisation in two. We define a *)
+(* class of "base groups", which are basically monoids *)
+(* with an involutive antimorphism, from which we derive *)
+(* the class of groups proper. This allows use to reuse *)
+(* much of the group notation and algebraic axioms for *)
+(* group subsets, by defining a base group class on them. *)
+(* We use class/mixins here rather than telescopes to *)
+(* be able to interoperate with the type coercions. *)
+(* Another potential benefit (not exploited here) would *)
+(* be to define a class for infinite groups, which could *)
+(* share all of the algebraic laws. *)
+Record mixin_of (T : Type) : Type := BaseMixin {
+ mul : T -> T -> T;
+ one : T;
+ inv : T -> T;
+ _ : associative mul;
+ _ : left_id one mul;
+ _ : involutive inv;
+ _ : {morph inv : x y / mul x y >-> mul y x}
+}.
+
+Structure base_type : Type := PackBase {
+ sort : Type;
+ _ : mixin_of sort;
+ _ : Finite.class_of sort
+}.
+
+(* We want to use sort as a coercion class, both to infer *)
+(* argument scopes properly, and to allow groups and cosets to *)
+(* coerce to the base group of group subsets. *)
+(* However, the return type of group operations should NOT be a *)
+(* coercion class, since this would trump the real (head-normal) *)
+(* coercion class for concrete group types, thus spoiling the *)
+(* coercion of A * B to pred_sort in x \in A * B, or rho * tau to *)
+(* ffun and Funclass in (rho * tau) x, when rho tau : perm T. *)
+(* Therefore we define an alias of sort for argument types, and *)
+(* make it the default coercion FinGroup.base_class >-> Sortclass *)
+(* so that arguments of a functions whose parameters are of type, *)
+(* say, gT : finGroupType, can be coerced to the coercion class *)
+(* of arg_sort. Care should be taken, however, to declare the *)
+(* return type of functions and operators as FinGroup.sort gT *)
+(* rather than gT, e.g., mulg : gT -> gT -> FinGroup.sort gT. *)
+(* Note that since we do this here and in quotient.v for all the *)
+(* basic functions, the inferred return type should generally be *)
+(* correct. *)
+Definition arg_sort := sort.
+
+Definition mixin T :=
+ let: PackBase _ m _ := T return mixin_of (sort T) in m.
+
+Definition finClass T :=
+ let: PackBase _ _ m := T return Finite.class_of (sort T) in m.
+
+Structure type : Type := Pack {
+ base : base_type;
+ _ : left_inverse (one (mixin base)) (inv (mixin base)) (mul (mixin base))
+}.
+
+(* We only need three axioms to make a true group. *)
+
+Section Mixin.
+
+Variables (T : Type) (one : T) (mul : T -> T -> T) (inv : T -> T).
+
+Hypothesis mulA : associative mul.
+Hypothesis mul1 : left_id one mul.
+Hypothesis mulV : left_inverse one inv mul.
+Notation "1" := one.
+Infix "*" := mul.
+Notation "x ^-1" := (inv x).
+
+Lemma mk_invgK : involutive inv.
+Proof.
+have mulV21 x: x^-1^-1 * 1 = x by rewrite -(mulV x) mulA mulV mul1.
+by move=> x; rewrite -[_ ^-1]mulV21 -(mul1 1) mulA !mulV21.
+Qed.
+
+Lemma mk_invMg : {morph inv : x y / x * y >-> y * x}.
+Proof.
+have mulxV x: x * x^-1 = 1 by rewrite -{1}[x]mk_invgK mulV.
+move=> x y /=; rewrite -[y^-1 * _]mul1 -(mulV (x * y)) -2!mulA (mulA y).
+by rewrite mulxV mul1 mulxV -(mulxV (x * y)) mulA mulV mul1.
+Qed.
+
+Definition Mixin := BaseMixin mulA mul1 mk_invgK mk_invMg.
+
+End Mixin.
+
+Definition pack_base T m :=
+ fun c cT & phant_id (Finite.class cT) c => @PackBase T m c.
+
+Definition clone_base T :=
+ fun bT & sort bT -> T =>
+ fun m c (bT' := @PackBase T m c) & phant_id bT' bT => bT'.
+
+Definition clone T :=
+ fun bT gT & sort bT * sort (base gT) -> T * T =>
+ fun m (gT' := @Pack bT m) & phant_id gT' gT => gT'.
+
+Section InheritedClasses.
+
+Variable bT : base_type.
+Local Notation T := (arg_sort bT).
+Local Notation rT := (sort bT).
+Local Notation class := (finClass bT).
+
+Canonical eqType := Equality.Pack class rT.
+Canonical choiceType := Choice.Pack class rT.
+Canonical countType := Countable.Pack class rT.
+Canonical finType := Finite.Pack class rT.
+Definition arg_eqType := Eval hnf in [eqType of T].
+Definition arg_choiceType := Eval hnf in [choiceType of T].
+Definition arg_countType := Eval hnf in [countType of T].
+Definition arg_finType := Eval hnf in [finType of T].
+
+End InheritedClasses.
+
+Module Import Exports.
+(* Declaring sort as a Coercion is clearly redundant; it only *)
+(* serves the purpose of eliding FinGroup.sort in the display of *)
+(* return types. The warning could be eliminated by using the *)
+(* functor trick to replace Sortclass by a dummy target. *)
+Coercion arg_sort : base_type >-> Sortclass.
+Coercion sort : base_type >-> Sortclass.
+Coercion mixin : base_type >-> mixin_of.
+Coercion base : type >-> base_type.
+Canonical eqType.
+Canonical choiceType.
+Canonical countType.
+Canonical finType.
+Coercion arg_eqType : base_type >-> Equality.type.
+Canonical arg_eqType.
+Coercion arg_choiceType : base_type >-> Choice.type.
+Canonical arg_choiceType.
+Coercion arg_countType : base_type >-> Countable.type.
+Canonical arg_countType.
+Coercion arg_finType : base_type >-> Finite.type.
+Canonical arg_finType.
+Bind Scope group_scope with sort.
+Bind Scope group_scope with arg_sort.
+Notation baseFinGroupType := base_type.
+Notation finGroupType := type.
+Notation BaseFinGroupType T m := (@pack_base T m _ _ id).
+Notation FinGroupType := Pack.
+Notation "[ 'baseFinGroupType' 'of' T ]" := (@clone_base T _ id _ _ id)
+ (at level 0, format "[ 'baseFinGroupType' 'of' T ]") : form_scope.
+Notation "[ 'finGroupType' 'of' T ]" := (@clone T _ _ id _ id)
+ (at level 0, format "[ 'finGroupType' 'of' T ]") : form_scope.
+End Exports.
+
+End FinGroup.
+Export FinGroup.Exports.
+
+Section ElementOps.
+
+Variable T : baseFinGroupType.
+Notation rT := (FinGroup.sort T).
+
+Definition oneg : rT := FinGroup.one T.
+Definition mulg : T -> T -> rT := FinGroup.mul T.
+Definition invg : T -> rT := FinGroup.inv T.
+Definition expgn_rec (x : T) n : rT := iterop n mulg x oneg.
+
+End ElementOps.
+
+Definition expgn := nosimpl expgn_rec.
+
+Notation "1" := (oneg _) : group_scope.
+Notation "x1 * x2" := (mulg x1 x2) : group_scope.
+Notation "x ^-1" := (invg x) : group_scope.
+Notation "x ^+ n" := (expgn x n) : group_scope.
+Notation "x ^- n" := (x ^+ n)^-1 : group_scope.
+
+(* Arguments of conjg are restricted to true groups to avoid an *)
+(* improper interpretation of A ^ B with A and B sets, namely: *)
+(* {x^-1 * (y * z) | y \in A, x, z \in B} *)
+Definition conjg (T : finGroupType) (x y : T) := y^-1 * (x * y).
+Notation "x1 ^ x2" := (conjg x1 x2) : group_scope.
+
+Definition commg (T : finGroupType) (x y : T) := x^-1 * x ^ y.
+Notation "[ ~ x1 , x2 , .. , xn ]" := (commg .. (commg x1 x2) .. xn)
+ : group_scope.
+
+Prenex Implicits mulg invg expgn conjg commg.
+
+Notation "\prod_ ( i <- r | P ) F" :=
+ (\big[mulg/1]_(i <- r | P%B) F%g) : group_scope.
+Notation "\prod_ ( i <- r ) F" :=
+ (\big[mulg/1]_(i <- r) F%g) : group_scope.
+Notation "\prod_ ( m <= i < n | P ) F" :=
+ (\big[mulg/1]_(m <= i < n | P%B) F%g) : group_scope.
+Notation "\prod_ ( m <= i < n ) F" :=
+ (\big[mulg/1]_(m <= i < n) F%g) : group_scope.
+Notation "\prod_ ( i | P ) F" :=
+ (\big[mulg/1]_(i | P%B) F%g) : group_scope.
+Notation "\prod_ i F" :=
+ (\big[mulg/1]_i F%g) : group_scope.
+Notation "\prod_ ( i : t | P ) F" :=
+ (\big[mulg/1]_(i : t | P%B) F%g) (only parsing) : group_scope.
+Notation "\prod_ ( i : t ) F" :=
+ (\big[mulg/1]_(i : t) F%g) (only parsing) : group_scope.
+Notation "\prod_ ( i < n | P ) F" :=
+ (\big[mulg/1]_(i < n | P%B) F%g) : group_scope.
+Notation "\prod_ ( i < n ) F" :=
+ (\big[mulg/1]_(i < n) F%g) : group_scope.
+Notation "\prod_ ( i 'in' A | P ) F" :=
+ (\big[mulg/1]_(i in A | P%B) F%g) : group_scope.
+Notation "\prod_ ( i 'in' A ) F" :=
+ (\big[mulg/1]_(i in A) F%g) : group_scope.
+
+Section PreGroupIdentities.
+
+Variable T : baseFinGroupType.
+Implicit Types x y z : T.
+Local Notation mulgT := (@mulg T).
+
+Lemma mulgA : associative mulgT. Proof. by case: T => ? []. Qed.
+Lemma mul1g : left_id 1 mulgT. Proof. by case: T => ? []. Qed.
+Lemma invgK : @involutive T invg. Proof. by case: T => ? []. Qed.
+Lemma invMg x y : (x * y)^-1 = y^-1 * x^-1. Proof. by case: T x y => ? []. Qed.
+
+Lemma invg_inj : @injective T T invg. Proof. exact: can_inj invgK. Qed.
+
+Lemma eq_invg_sym x y : (x^-1 == y) = (x == y^-1).
+Proof. by exact: (inv_eq invgK). Qed.
+
+Lemma invg1 : 1^-1 = 1 :> T.
+Proof. by apply: invg_inj; rewrite -{1}[1^-1]mul1g invMg invgK mul1g. Qed.
+
+Lemma eq_invg1 x : (x^-1 == 1) = (x == 1).
+Proof. by rewrite eq_invg_sym invg1. Qed.
+
+Lemma mulg1 : right_id 1 mulgT.
+Proof. by move=> x; apply: invg_inj; rewrite invMg invg1 mul1g. Qed.
+
+Canonical finGroup_law := Monoid.Law mulgA mul1g mulg1.
+
+Lemma expgnE x n : x ^+ n = expgn_rec x n. Proof. by []. Qed.
+
+Lemma expg0 x : x ^+ 0 = 1. Proof. by []. Qed.
+Lemma expg1 x : x ^+ 1 = x. Proof. by []. Qed.
+
+Lemma expgS x n : x ^+ n.+1 = x * x ^+ n.
+Proof. by case: n => //; rewrite mulg1. Qed.
+
+Lemma expg1n n : 1 ^+ n = 1 :> T.
+Proof. by elim: n => // n IHn; rewrite expgS mul1g. Qed.
+
+Lemma expgD x n m : x ^+ (n + m) = x ^+ n * x ^+ m.
+Proof. by elim: n => [|n IHn]; rewrite ?mul1g // !expgS IHn mulgA. Qed.
+
+Lemma expgSr x n : x ^+ n.+1 = x ^+ n * x.
+Proof. by rewrite -addn1 expgD expg1. Qed.
+
+Lemma expgM x n m : x ^+ (n * m) = x ^+ n ^+ m.
+Proof.
+elim: m => [|m IHm]; first by rewrite muln0 expg0.
+by rewrite mulnS expgD IHm expgS.
+Qed.
+
+Lemma expgAC x m n : x ^+ m ^+ n = x ^+ n ^+ m.
+Proof. by rewrite -!expgM mulnC. Qed.
+
+Definition commute x y := x * y = y * x.
+
+Lemma commute_refl x : commute x x.
+Proof. by []. Qed.
+
+Lemma commute_sym x y : commute x y -> commute y x.
+Proof. by []. Qed.
+
+Lemma commute1 x : commute x 1.
+Proof. by rewrite /commute mulg1 mul1g. Qed.
+
+Lemma commuteM x y z : commute x y -> commute x z -> commute x (y * z).
+Proof. by move=> cxy cxz; rewrite /commute -mulgA -cxz !mulgA cxy. Qed.
+
+Lemma commuteX x y n : commute x y -> commute x (y ^+ n).
+Proof.
+move=> cxy; elim: n => [|n]; [exact: commute1 | rewrite expgS; exact: commuteM].
+Qed.
+
+Lemma commuteX2 x y m n : commute x y -> commute (x ^+ m) (y ^+ n).
+Proof. by move=> cxy; exact/commuteX/commute_sym/commuteX. Qed.
+
+Lemma expgVn x n : x^-1 ^+ n = x ^- n.
+Proof. by elim: n => [|n IHn]; rewrite ?invg1 // expgSr expgS invMg IHn. Qed.
+
+Lemma expgMn x y n : commute x y -> (x * y) ^+ n = x ^+ n * y ^+ n.
+Proof.
+move=> cxy; elim: n => [|n IHn]; first by rewrite mulg1.
+by rewrite !expgS IHn -mulgA (mulgA y) (commuteX _ (commute_sym cxy)) !mulgA.
+Qed.
+
+End PreGroupIdentities.
+
+Hint Resolve commute1.
+Implicit Arguments invg_inj [T].
+Prenex Implicits commute invgK invg_inj.
+
+Section GroupIdentities.
+
+Variable T : finGroupType.
+Implicit Types x y z : T.
+Local Notation mulgT := (@mulg T).
+
+Lemma mulVg : left_inverse 1 invg mulgT.
+Proof. by case T. Qed.
+
+Lemma mulgV : right_inverse 1 invg mulgT.
+Proof. by move=> x; rewrite -{1}(invgK x) mulVg. Qed.
+
+Lemma mulKg : left_loop invg mulgT.
+Proof. by move=> x y; rewrite mulgA mulVg mul1g. Qed.
+
+Lemma mulKVg : rev_left_loop invg mulgT.
+Proof. by move=> x y; rewrite mulgA mulgV mul1g. Qed.
+
+Lemma mulgI : right_injective mulgT.
+Proof. move=> x; exact: can_inj (mulKg x). Qed.
+
+Lemma mulgK : right_loop invg mulgT.
+Proof. by move=> x y; rewrite -mulgA mulgV mulg1. Qed.
+
+Lemma mulgKV : rev_right_loop invg mulgT.
+Proof. by move=> x y; rewrite -mulgA mulVg mulg1. Qed.
+
+Lemma mulIg : left_injective mulgT.
+Proof. move=> x; exact: can_inj (mulgK x). Qed.
+
+Lemma eq_invg_mul x y : (x^-1 == y :> T) = (x * y == 1 :> T).
+Proof. by rewrite -(inj_eq (@mulgI x)) mulgV eq_sym. Qed.
+
+Lemma eq_mulgV1 x y : (x == y) = (x * y^-1 == 1 :> T).
+Proof. by rewrite -(inj_eq invg_inj) eq_invg_mul. Qed.
+
+Lemma eq_mulVg1 x y : (x == y) = (x^-1 * y == 1 :> T).
+Proof. by rewrite -eq_invg_mul invgK. Qed.
+
+Lemma commuteV x y : commute x y -> commute x y^-1.
+Proof. by move=> cxy; apply: (@mulIg y); rewrite mulgKV -mulgA cxy mulKg. Qed.
+
+Lemma conjgE x y : x ^ y = y^-1 * (x * y). Proof. by []. Qed.
+
+Lemma conjgC x y : x * y = y * x ^ y.
+Proof. by rewrite mulKVg. Qed.
+
+Lemma conjgCV x y : x * y = y ^ x^-1 * x.
+Proof. by rewrite -mulgA mulgKV invgK. Qed.
+
+Lemma conjg1 x : x ^ 1 = x.
+Proof. by rewrite conjgE commute1 mulKg. Qed.
+
+Lemma conj1g x : 1 ^ x = 1.
+Proof. by rewrite conjgE mul1g mulVg. Qed.
+
+Lemma conjMg x y z : (x * y) ^ z = x ^ z * y ^ z.
+Proof. by rewrite !conjgE !mulgA mulgK. Qed.
+
+Lemma conjgM x y z : x ^ (y * z) = (x ^ y) ^ z.
+Proof. by rewrite !conjgE invMg !mulgA. Qed.
+
+Lemma conjVg x y : x^-1 ^ y = (x ^ y)^-1.
+Proof. by rewrite !conjgE !invMg invgK mulgA. Qed.
+
+Lemma conjJg x y z : (x ^ y) ^ z = (x ^ z) ^ y ^ z.
+Proof. by rewrite 2!conjMg conjVg. Qed.
+
+Lemma conjXg x y n : (x ^+ n) ^ y = (x ^ y) ^+ n.
+Proof. by elim: n => [|n IHn]; rewrite ?conj1g // !expgS conjMg IHn. Qed.
+
+Lemma conjgK : @right_loop T T invg conjg.
+Proof. by move=> y x; rewrite -conjgM mulgV conjg1. Qed.
+
+Lemma conjgKV : @rev_right_loop T T invg conjg.
+Proof. by move=> y x; rewrite -conjgM mulVg conjg1. Qed.
+
+Lemma conjg_inj : @left_injective T T T conjg.
+Proof. move=> y; exact: can_inj (conjgK y). Qed.
+
+Lemma conjg_eq1 x y : (x ^ y == 1) = (x == 1).
+Proof. by rewrite -(inj_eq (@conjg_inj y) x) conj1g. Qed.
+
+Lemma conjg_prod I r (P : pred I) F z :
+ (\prod_(i <- r | P i) F i) ^ z = \prod_(i <- r | P i) (F i ^ z).
+Proof.
+by apply: (big_morph (conjg^~ z)) => [x y|]; rewrite ?conj1g ?conjMg.
+Qed.
+
+Lemma commgEl x y : [~ x, y] = x^-1 * x ^ y. Proof. by []. Qed.
+
+Lemma commgEr x y : [~ x, y] = y^-1 ^ x * y.
+Proof. by rewrite -!mulgA. Qed.
+
+Lemma commgC x y : x * y = y * x * [~ x, y].
+Proof. by rewrite -mulgA !mulKVg. Qed.
+
+Lemma commgCV x y : x * y = [~ x^-1, y^-1] * (y * x).
+Proof. by rewrite commgEl !mulgA !invgK !mulgKV. Qed.
+
+Lemma conjRg x y z : [~ x, y] ^ z = [~ x ^ z, y ^ z].
+Proof. by rewrite !conjMg !conjVg. Qed.
+
+Lemma invg_comm x y : [~ x, y]^-1 = [~ y, x].
+Proof. by rewrite commgEr conjVg invMg invgK. Qed.
+
+Lemma commgP x y : reflect (commute x y) ([~ x, y] == 1 :> T).
+Proof. by rewrite [[~ x, y]]mulgA -invMg -eq_mulVg1 eq_sym; exact: eqP. Qed.
+
+Lemma conjg_fixP x y : reflect (x ^ y = x) ([~ x, y] == 1 :> T).
+Proof. by rewrite -eq_mulVg1 eq_sym; exact: eqP. Qed.
+
+Lemma commg1_sym x y : ([~ x, y] == 1 :> T) = ([~ y, x] == 1 :> T).
+Proof. by rewrite -invg_comm (inv_eq invgK) invg1. Qed.
+
+Lemma commg1 x : [~ x, 1] = 1.
+Proof. exact/eqP/commgP. Qed.
+
+Lemma comm1g x : [~ 1, x] = 1.
+Proof. by rewrite -invg_comm commg1 invg1. Qed.
+
+Lemma commgg x : [~ x, x] = 1.
+Proof. by exact/eqP/commgP. Qed.
+
+Lemma commgXg x n : [~ x, x ^+ n] = 1.
+Proof. exact/eqP/commgP/commuteX. Qed.
+
+Lemma commgVg x : [~ x, x^-1] = 1.
+Proof. by exact/eqP/commgP/commuteV. Qed.
+
+Lemma commgXVg x n : [~ x, x ^- n] = 1.
+Proof. exact/eqP/commgP/commuteV/commuteX. Qed.
+
+(* Other commg identities should slot in here. *)
+
+End GroupIdentities.
+
+Hint Rewrite mulg1 mul1g invg1 mulVg mulgV (@invgK) mulgK mulgKV
+ invMg mulgA : gsimpl.
+
+Ltac gsimpl := autorewrite with gsimpl; try done.
+
+Definition gsimp := (mulg1 , mul1g, (invg1, @invgK), (mulgV, mulVg)).
+Definition gnorm := (gsimp, (mulgK, mulgKV, (mulgA, invMg))).
+
+Implicit Arguments mulgI [T].
+Implicit Arguments mulIg [T].
+Implicit Arguments conjg_inj [T].
+Implicit Arguments commgP [T x y].
+Implicit Arguments conjg_fixP [T x y].
+Prenex Implicits conjg_fixP commgP.
+
+Section Repr.
+(* Plucking a set representative. *)
+
+Variable gT : baseFinGroupType.
+Implicit Type A : {set gT}.
+
+Definition repr A := if 1 \in A then 1 else odflt 1 [pick x in A].
+
+Lemma mem_repr A x : x \in A -> repr A \in A.
+Proof.
+by rewrite /repr; case: ifP => // _; case: pickP => // A0; rewrite [x \in A]A0.
+Qed.
+
+Lemma card_mem_repr A : #|A| > 0 -> repr A \in A.
+Proof. by rewrite lt0n => /existsP[x]; exact: mem_repr. Qed.
+
+Lemma repr_set1 x : repr [set x] = x.
+Proof. by apply/set1P/card_mem_repr; rewrite cards1. Qed.
+
+Lemma repr_set0 : repr set0 = 1.
+Proof. by rewrite /repr; case: pickP => [x|_]; rewrite !inE. Qed.
+
+End Repr.
+
+Implicit Arguments mem_repr [gT A].
+
+Section BaseSetMulDef.
+(* We only assume a baseFinGroupType to allow this construct to be iterated. *)
+Variable gT : baseFinGroupType.
+Implicit Types A B : {set gT}.
+
+(* Set-lifted group operations. *)
+
+Definition set_mulg A B := mulg @2: (A, B).
+Definition set_invg A := invg @^-1: A.
+
+(* The pre-group structure of group subsets. *)
+
+Lemma set_mul1g : left_id [set 1] set_mulg.
+Proof.
+move=> A; apply/setP=> y; apply/imset2P/idP=> [[_ x /set1P-> Ax ->] | Ay].
+ by rewrite mul1g.
+by exists (1 : gT) y; rewrite ?(set11, mul1g).
+Qed.
+
+Lemma set_mulgA : associative set_mulg.
+Proof.
+move=> A B C; apply/setP=> y.
+apply/imset2P/imset2P=> [[x1 z Ax1 /imset2P[x2 x3 Bx2 Cx3 ->] ->]| [z x3]].
+ by exists (x1 * x2) x3; rewrite ?mulgA //; apply/imset2P; exists x1 x2.
+case/imset2P=> x1 x2 Ax1 Bx2 -> Cx3 ->.
+by exists x1 (x2 * x3); rewrite ?mulgA //; apply/imset2P; exists x2 x3.
+Qed.
+
+Lemma set_invgK : involutive set_invg.
+Proof. by move=> A; apply/setP=> x; rewrite !inE invgK. Qed.
+
+Lemma set_invgM : {morph set_invg : A B / set_mulg A B >-> set_mulg B A}.
+Proof.
+move=> A B; apply/setP=> z; rewrite inE.
+apply/imset2P/imset2P=> [[x y Ax By /(canRL invgK)->] | [y x]].
+ by exists y^-1 x^-1; rewrite ?invMg // inE invgK.
+by rewrite !inE => By1 Ax1 ->; exists x^-1 y^-1; rewrite ?invMg.
+Qed.
+
+Definition group_set_baseGroupMixin : FinGroup.mixin_of (set_type gT) :=
+ FinGroup.BaseMixin set_mulgA set_mul1g set_invgK set_invgM.
+
+Canonical group_set_baseGroupType :=
+ Eval hnf in BaseFinGroupType (set_type gT) group_set_baseGroupMixin.
+
+Canonical group_set_of_baseGroupType :=
+ Eval hnf in [baseFinGroupType of {set gT}].
+
+End BaseSetMulDef.
+
+(* Time to open the bag of dirty tricks. When we define groups down below *)
+(* as a subtype of {set gT}, we need them to be able to coerce to sets in *)
+(* both set-style contexts (x \in G) and monoid-style contexts (G * H), *)
+(* and we need the coercion function to be EXACTLY the structure *)
+(* projection in BOTH cases -- otherwise the canonical unification breaks.*)
+(* Alas, Coq doesn't let us use the same coercion function twice, even *)
+(* when the targets are convertible. Our workaround (ab)uses the module *)
+(* system to declare two different identity coercions on an alias class. *)
+
+Module GroupSet.
+Definition sort (gT : baseFinGroupType) := {set gT}.
+End GroupSet.
+Identity Coercion GroupSet_of_sort : GroupSet.sort >-> set_of.
+
+Module Type GroupSetBaseGroupSig.
+Definition sort gT := group_set_of_baseGroupType gT : Type.
+End GroupSetBaseGroupSig.
+
+Module MakeGroupSetBaseGroup (Gset_base : GroupSetBaseGroupSig).
+Identity Coercion of_sort : Gset_base.sort >-> FinGroup.arg_sort.
+End MakeGroupSetBaseGroup.
+
+Module Export GroupSetBaseGroup := MakeGroupSetBaseGroup GroupSet.
+
+Canonical group_set_eqType gT := Eval hnf in [eqType of GroupSet.sort gT].
+Canonical group_set_choiceType gT :=
+ Eval hnf in [choiceType of GroupSet.sort gT].
+Canonical group_set_countType gT := Eval hnf in [countType of GroupSet.sort gT].
+Canonical group_set_finType gT := Eval hnf in [finType of GroupSet.sort gT].
+
+Section GroupSetMulDef.
+(* Some of these constructs could be defined on a baseFinGroupType. *)
+(* We restrict them to proper finGroupType because we only develop *)
+(* the theory for that case. *)
+Variable gT : finGroupType.
+Implicit Types A B : {set gT}.
+Implicit Type x y : gT.
+
+Definition lcoset A x := mulg x @: A.
+Definition rcoset A x := mulg^~ x @: A.
+Definition lcosets A B := lcoset A @: B.
+Definition rcosets A B := rcoset A @: B.
+Definition indexg B A := #|rcosets A B|.
+
+Definition conjugate A x := conjg^~ x @: A.
+Definition conjugates A B := conjugate A @: B.
+Definition class x B := conjg x @: B.
+Definition classes A := class^~ A @: A.
+Definition class_support A B := conjg @2: (A, B).
+
+Definition commg_set A B := commg @2: (A, B).
+
+(* These will only be used later, but are defined here so that we can *)
+(* keep all the Notation together. *)
+Definition normaliser A := [set x | conjugate A x \subset A].
+Definition centraliser A := \bigcap_(x in A) normaliser [set x].
+Definition abelian A := A \subset centraliser A.
+Definition normal A B := (A \subset B) && (B \subset normaliser A).
+
+(* "normalised" and "centralise[s|d]" are intended to be used with *)
+(* the {in ...} form, as in abelian below. *)
+Definition normalised A := forall x, conjugate A x = A.
+Definition centralises x A := forall y, y \in A -> commute x y.
+Definition centralised A := forall x, centralises x A.
+
+End GroupSetMulDef.
+
+Arguments Scope lcoset [_ group_scope group_scope].
+Arguments Scope rcoset [_ group_scope group_scope].
+Arguments Scope rcosets [_ group_scope group_scope].
+Arguments Scope lcosets [_ group_scope group_scope].
+Arguments Scope indexg [_ group_scope group_scope].
+Arguments Scope conjugate [_ group_scope group_scope].
+Arguments Scope conjugates [_ group_scope group_scope].
+Arguments Scope class [_ group_scope group_scope].
+Arguments Scope classes [_ group_scope].
+Arguments Scope class_support [_ group_scope group_scope].
+Arguments Scope commg_set [_ group_scope group_scope].
+Arguments Scope normaliser [_ group_scope].
+Arguments Scope centraliser [_ group_scope].
+Arguments Scope abelian [_ group_scope].
+Arguments Scope normal [_ group_scope group_scope].
+Arguments Scope centralised [_ group_scope].
+Arguments Scope normalised [_ group_scope].
+Arguments Scope centralises [_ group_scope group_scope].
+Arguments Scope centralised [_ group_scope].
+
+Notation "[ 1 gT ]" := (1 : {set gT}) : group_scope.
+Notation "[ 1 ]" := [1 FinGroup.sort _] : group_scope.
+
+Notation "A ^#" := (A :\ 1) : group_scope.
+
+Notation "x *: A" := ([set x%g] * A) : group_scope.
+Notation "A :* x" := (A * [set x%g]) : group_scope.
+Notation "A :^ x" := (conjugate A x) : group_scope.
+Notation "x ^: B" := (class x B) : group_scope.
+Notation "A :^: B" := (conjugates A B) : group_scope.
+
+Notation "#| B : A |" := (indexg B A) : group_scope.
+
+(* No notation for lcoset and rcoset, which are to be used mostly *)
+(* in curried form; x *: B and A :* 1 denote singleton products, *)
+(* so thus we can use mulgA, mulg1, etc, on, say, A :* 1 * B :* x. *)
+(* No notation for the set commutator generator set set_commg. *)
+
+Notation "''N' ( A )" := (normaliser A) : group_scope.
+Notation "''N_' G ( A )" := (G%g :&: 'N(A)) : group_scope.
+Notation "A <| B" := (normal A B) : group_scope.
+Notation "''C' ( A )" := (centraliser A) : group_scope.
+Notation "''C_' G ( A )" := (G%g :&: 'C(A)) : group_scope.
+Notation "''C_' ( G ) ( A )" := 'C_G(A) (only parsing) : group_scope.
+Notation "''C' [ x ]" := 'N([set x%g]) : group_scope.
+Notation "''C_' G [ x ]" := 'N_G([set x%g]) : group_scope.
+Notation "''C_' ( G ) [ x ]" := 'C_G[x] (only parsing) : group_scope.
+
+Prenex Implicits repr lcoset rcoset lcosets rcosets normal.
+Prenex Implicits conjugate conjugates class classes class_support.
+Prenex Implicits commg_set normalised centralised abelian.
+
+Section BaseSetMulProp.
+(* Properties of the purely multiplicative structure. *)
+Variable gT : baseFinGroupType.
+Implicit Types A B C D : {set gT}.
+Implicit Type x y z : gT.
+
+(* Set product. We already have all the pregroup identities, so we *)
+(* only need to add the monotonicity rules. *)
+
+Lemma mulsgP A B x :
+ reflect (imset2_spec mulg (mem A) (fun _ => mem B) x) (x \in A * B).
+Proof. exact: imset2P. Qed.
+
+Lemma mem_mulg A B x y : x \in A -> y \in B -> x * y \in A * B.
+Proof. by move=> Ax By; apply/mulsgP; exists x y. Qed.
+
+Lemma prodsgP (I : finType) (P : pred I) (A : I -> {set gT}) x :
+ reflect (exists2 c, forall i, P i -> c i \in A i & x = \prod_(i | P i) c i)
+ (x \in \prod_(i | P i) A i).
+Proof.
+rewrite -big_filter filter_index_enum; set r := enum P.
+pose inA c i := c i \in A i; set RHS := x \in _.
+suffices IHr: reflect (exists2 c, all (inA c) r & x = \prod_(i <- r) c i) RHS.
+ apply: (iffP IHr) => [][c inAc ->].
+ rewrite -[r]filter_index_enum big_filter; exists c => // i Pi.
+ by apply: (allP inAc); rewrite mem_enum.
+ rewrite -big_filter filter_index_enum; exists c => //; apply/allP=> i.
+ rewrite mem_enum; exact: inAc.
+have: uniq r by [rewrite enum_uniq]; rewrite {}/RHS.
+elim: {P}r x => /= [x _|i r IHr x /andP[r'i Ur]].
+ by rewrite unlock; apply: (iffP set1P) => [-> | [] //]; exists (fun _ => 1).
+rewrite big_cons; apply: (iffP idP) => [|[c /andP[Aci Ac] ->]]; last first.
+ by rewrite big_cons mem_mulg //; apply/IHr=> //; exists c.
+case/mulsgP=> y _ Ai_y /IHr[//| c Ac ->] ->{x}.
+exists [eta c with i |-> y] => /=.
+ rewrite /inA /= eqxx Ai_y; apply/allP=> j rj.
+ by case: eqP rj r'i => [-> -> // | _ rj _]; exact: (allP Ac).
+rewrite big_cons eqxx !big_seq; congr (_ * _).
+by apply: eq_bigr => j rj; case: eqP rj r'i => // -> ->.
+Qed.
+
+Lemma mem_prodg (I : finType) (P : pred I) (A : I -> {set gT}) c :
+ (forall i, P i -> c i \in A i) -> \prod_(i | P i) c i \in \prod_(i | P i) A i.
+Proof. by move=> Ac; apply/prodsgP; exists c. Qed.
+
+Lemma mulSg A B C : A \subset B -> A * C \subset B * C.
+Proof. exact: imset2Sl. Qed.
+
+Lemma mulgS A B C : B \subset C -> A * B \subset A * C.
+Proof. exact: imset2Sr. Qed.
+
+Lemma mulgSS A B C D : A \subset B -> C \subset D -> A * C \subset B * D.
+Proof. exact: imset2S. Qed.
+
+Lemma mulg_subl A B : 1 \in B -> A \subset A * B.
+Proof. by move=> B1; rewrite -{1}(mulg1 A) mulgS ?sub1set. Qed.
+
+Lemma mulg_subr A B : 1 \in A -> B \subset A * B.
+Proof. by move=> A1; rewrite -{1}(mul1g B) mulSg ?sub1set. Qed.
+
+Lemma mulUg A B C : (A :|: B) * C = (A * C) :|: (B * C).
+Proof. exact: imset2Ul. Qed.
+
+Lemma mulgU A B C : A * (B :|: C) = (A * B) :|: (A * C).
+Proof. exact: imset2Ur. Qed.
+
+(* Set (pointwise) inverse. *)
+
+Lemma invUg A B : (A :|: B)^-1 = A^-1 :|: B^-1.
+Proof. exact: preimsetU. Qed.
+
+Lemma invIg A B : (A :&: B)^-1 = A^-1 :&: B^-1.
+Proof. exact: preimsetI. Qed.
+
+Lemma invDg A B : (A :\: B)^-1 = A^-1 :\: B^-1.
+Proof. exact: preimsetD. Qed.
+
+Lemma invCg A : (~: A)^-1 = ~: A^-1.
+Proof. exact: preimsetC. Qed.
+
+Lemma invSg A B : (A^-1 \subset B^-1) = (A \subset B).
+Proof. by rewrite !(sameP setIidPl eqP) -invIg (inj_eq invg_inj). Qed.
+
+Lemma mem_invg x A : (x \in A^-1) = (x^-1 \in A).
+Proof. by rewrite inE. Qed.
+
+Lemma memV_invg x A : (x^-1 \in A^-1) = (x \in A).
+Proof. by rewrite inE invgK. Qed.
+
+Lemma card_invg A : #|A^-1| = #|A|.
+Proof. by apply: card_preimset; exact: invg_inj. Qed.
+
+(* Product with singletons. *)
+
+Lemma set1gE : 1 = [set 1] :> {set gT}. Proof. by []. Qed.
+
+Lemma set1gP x : reflect (x = 1) (x \in [1]).
+Proof. exact: set1P. Qed.
+
+Lemma mulg_set1 x y : [set x] :* y = [set x * y].
+Proof. by rewrite [_ * _]imset2_set1l imset_set1. Qed.
+
+Lemma invg_set1 x : [set x]^-1 = [set x^-1].
+Proof. by apply/setP=> y; rewrite !inE inv_eq //; exact: invgK. Qed.
+
+End BaseSetMulProp.
+
+Implicit Arguments set1gP [gT x].
+Implicit Arguments mulsgP [gT A B x].
+Implicit Arguments prodsgP [gT I P A x].
+
+Section GroupSetMulProp.
+(* Constructs that need a finGroupType *)
+Variable gT : finGroupType.
+Implicit Types A B C D : {set gT}.
+Implicit Type x y z : gT.
+
+(* Left cosets. *)
+
+Lemma lcosetE A x : lcoset A x = x *: A.
+Proof. by rewrite [_ * _]imset2_set1l. Qed.
+
+Lemma card_lcoset A x : #|x *: A| = #|A|.
+Proof. by rewrite -lcosetE (card_imset _ (mulgI _)). Qed.
+
+Lemma mem_lcoset A x y : (y \in x *: A) = (x^-1 * y \in A).
+Proof. by rewrite -lcosetE [_ x](can_imset_pre _ (mulKg _)) inE. Qed.
+
+Lemma lcosetP A x y : reflect (exists2 a, a \in A & y = x * a) (y \in x *: A).
+Proof. rewrite -lcosetE; exact: imsetP. Qed.
+
+Lemma lcosetsP A B C :
+ reflect (exists2 x, x \in B & C = x *: A) (C \in lcosets A B).
+Proof. by apply: (iffP imsetP) => [] [x Bx ->]; exists x; rewrite ?lcosetE. Qed.
+
+Lemma lcosetM A x y : (x * y) *: A = x *: (y *: A).
+Proof. by rewrite -mulg_set1 mulgA. Qed.
+
+Lemma lcoset1 A : 1 *: A = A.
+Proof. exact: mul1g. Qed.
+
+Lemma lcosetK : left_loop invg (fun x A => x *: A).
+Proof. by move=> x A; rewrite -lcosetM mulVg mul1g. Qed.
+
+Lemma lcosetKV : rev_left_loop invg (fun x A => x *: A).
+Proof. by move=> x A; rewrite -lcosetM mulgV mul1g. Qed.
+
+Lemma lcoset_inj : right_injective (fun x A => x *: A).
+Proof. by move=> x; exact: can_inj (lcosetK x). Qed.
+
+Lemma lcosetS x A B : (x *: A \subset x *: B) = (A \subset B).
+Proof.
+apply/idP/idP=> sAB; last exact: mulgS.
+by rewrite -(lcosetK x A) -(lcosetK x B) mulgS.
+Qed.
+
+Lemma sub_lcoset x A B : (A \subset x *: B) = (x^-1 *: A \subset B).
+Proof. by rewrite -(lcosetS x^-1) lcosetK. Qed.
+
+Lemma sub_lcosetV x A B : (A \subset x^-1 *: B) = (x *: A \subset B).
+Proof. by rewrite sub_lcoset invgK. Qed.
+
+(* Right cosets. *)
+
+Lemma rcosetE A x : rcoset A x = A :* x.
+Proof. by rewrite [_ * _]imset2_set1r. Qed.
+
+Lemma card_rcoset A x : #|A :* x| = #|A|.
+Proof. by rewrite -rcosetE (card_imset _ (mulIg _)). Qed.
+
+Lemma mem_rcoset A x y : (y \in A :* x) = (y * x^-1 \in A).
+Proof. by rewrite -rcosetE [_ x](can_imset_pre A (mulgK _)) inE. Qed.
+
+Lemma rcosetP A x y : reflect (exists2 a, a \in A & y = a * x) (y \in A :* x).
+Proof. by rewrite -rcosetE; exact: imsetP. Qed.
+
+Lemma rcosetsP A B C :
+ reflect (exists2 x, x \in B & C = A :* x) (C \in rcosets A B).
+Proof. by apply: (iffP imsetP) => [] [x Bx ->]; exists x; rewrite ?rcosetE. Qed.
+
+Lemma rcosetM A x y : A :* (x * y) = A :* x :* y.
+Proof. by rewrite -mulg_set1 mulgA. Qed.
+
+Lemma rcoset1 A : A :* 1 = A.
+Proof. exact: mulg1. Qed.
+
+Lemma rcosetK : right_loop invg (fun A x => A :* x).
+Proof. by move=> x A; rewrite -rcosetM mulgV mulg1. Qed.
+
+Lemma rcosetKV : rev_right_loop invg (fun A x => A :* x).
+Proof. by move=> x A; rewrite -rcosetM mulVg mulg1. Qed.
+
+Lemma rcoset_inj : left_injective (fun A x => A :* x).
+Proof. by move=> x; exact: can_inj (rcosetK x). Qed.
+
+Lemma rcosetS x A B : (A :* x \subset B :* x) = (A \subset B).
+Proof.
+apply/idP/idP=> sAB; last exact: mulSg.
+by rewrite -(rcosetK x A) -(rcosetK x B) mulSg.
+Qed.
+
+Lemma sub_rcoset x A B : (A \subset B :* x) = (A :* x ^-1 \subset B).
+Proof. by rewrite -(rcosetS x^-1) rcosetK. Qed.
+
+Lemma sub_rcosetV x A B : (A \subset B :* x^-1) = (A :* x \subset B).
+Proof. by rewrite sub_rcoset invgK. Qed.
+
+(* Inverse map lcosets to rcosets *)
+
+Lemma lcosets_invg A B : lcosets A^-1 B^-1 = invg @^-1: rcosets A B.
+Proof.
+apply/setP=> C; rewrite inE.
+apply/imsetP/imsetP=> [] [a]; rewrite -memV_invg ?invgK => Aa;
+ try move/(canRL invgK); move->; exists a^-1;
+ by rewrite // lcosetE rcosetE invMg invg_set1 ?invgK.
+Qed.
+
+(* Conjugates. *)
+
+Lemma conjg_preim A x : A :^ x = (conjg^~ x^-1) @^-1: A.
+Proof. exact: can_imset_pre (conjgK _). Qed.
+
+Lemma mem_conjg A x y : (y \in A :^ x) = (y ^ x^-1 \in A).
+Proof. by rewrite conjg_preim inE. Qed.
+
+Lemma mem_conjgV A x y : (y \in A :^ x^-1) = (y ^ x \in A).
+Proof. by rewrite mem_conjg invgK. Qed.
+
+Lemma memJ_conjg A x y : (y ^ x \in A :^ x) = (y \in A).
+Proof. by rewrite mem_conjg conjgK. Qed.
+
+Lemma conjsgE A x : A :^ x = x^-1 *: (A :* x).
+Proof. by apply/setP=> y; rewrite mem_lcoset mem_rcoset -mulgA mem_conjg. Qed.
+
+Lemma conjsg1 A : A :^ 1 = A.
+Proof. by rewrite conjsgE invg1 mul1g mulg1. Qed.
+
+Lemma conjsgM A x y : A :^ (x * y) = (A :^ x) :^ y.
+Proof. by rewrite !conjsgE invMg -!mulg_set1 !mulgA. Qed.
+
+Lemma conjsgK : @right_loop _ gT invg conjugate.
+Proof. by move=> x A; rewrite -conjsgM mulgV conjsg1. Qed.
+
+Lemma conjsgKV : @rev_right_loop _ gT invg conjugate.
+Proof. by move=> x A; rewrite -conjsgM mulVg conjsg1. Qed.
+
+Lemma conjsg_inj : @left_injective _ gT _ conjugate.
+Proof. by move=> x; exact: can_inj (conjsgK x). Qed.
+
+Lemma cardJg A x : #|A :^ x| = #|A|.
+Proof. by rewrite (card_imset _ (conjg_inj x)). Qed.
+
+Lemma conjSg A B x : (A :^ x \subset B :^ x) = (A \subset B).
+Proof. by rewrite !conjsgE lcosetS rcosetS. Qed.
+
+Lemma properJ A B x : (A :^ x \proper B :^ x) = (A \proper B).
+Proof. by rewrite /proper !conjSg. Qed.
+
+Lemma sub_conjg A B x : (A :^ x \subset B) = (A \subset B :^ x^-1).
+Proof. by rewrite -(conjSg A _ x) conjsgKV. Qed.
+
+Lemma sub_conjgV A B x : (A :^ x^-1 \subset B) = (A \subset B :^ x).
+Proof. by rewrite -(conjSg _ B x) conjsgKV. Qed.
+
+Lemma conjg_set1 x y : [set x] :^ y = [set x ^ y].
+Proof. by rewrite [_ :^ _]imset_set1. Qed.
+
+Lemma conjs1g x : 1 :^ x = 1.
+Proof. by rewrite conjg_set1 conj1g. Qed.
+
+Lemma conjsg_eq1 A x : (A :^ x == 1%g) = (A == 1%g).
+Proof. by rewrite (canF_eq (conjsgK x)) conjs1g. Qed.
+
+Lemma conjsMg A B x : (A * B) :^ x = A :^ x * B :^ x.
+Proof. by rewrite !conjsgE !mulgA rcosetK. Qed.
+
+Lemma conjIg A B x : (A :&: B) :^ x = A :^ x :&: B :^ x.
+Proof. by rewrite !conjg_preim preimsetI. Qed.
+
+Lemma conj0g x : set0 :^ x = set0.
+Proof. exact: imset0. Qed.
+
+Lemma conjTg x : [set: gT] :^ x = [set: gT].
+Proof. by rewrite conjg_preim preimsetT. Qed.
+
+Lemma bigcapJ I r (P : pred I) (B : I -> {set gT}) x :
+ \bigcap_(i <- r | P i) (B i :^ x) = (\bigcap_(i <- r | P i) B i) :^ x.
+Proof.
+by rewrite (big_endo (conjugate^~ x)) => // [B1 B2|]; rewrite (conjTg, conjIg).
+Qed.
+
+Lemma conjUg A B x : (A :|: B) :^ x = A :^ x :|: B :^ x.
+Proof. by rewrite !conjg_preim preimsetU. Qed.
+
+Lemma bigcupJ I r (P : pred I) (B : I -> {set gT}) x :
+ \bigcup_(i <- r | P i) (B i :^ x) = (\bigcup_(i <- r | P i) B i) :^ x.
+Proof.
+rewrite (big_endo (conjugate^~ x)) => // [B1 B2|]; first by rewrite conjUg.
+exact: imset0.
+Qed.
+
+Lemma conjCg A x : (~: A) :^ x = ~: A :^ x.
+Proof. by rewrite !conjg_preim preimsetC. Qed.
+
+Lemma conjDg A B x : (A :\: B) :^ x = A :^ x :\: B :^ x.
+Proof. by rewrite !setDE !(conjCg, conjIg). Qed.
+
+Lemma conjD1g A x : A^# :^ x = (A :^ x)^#.
+Proof. by rewrite conjDg conjs1g. Qed.
+
+(* Classes; not much for now. *)
+
+Lemma memJ_class x y A : y \in A -> x ^ y \in x ^: A.
+Proof. exact: mem_imset. Qed.
+
+Lemma classS x A B : A \subset B -> x ^: A \subset x ^: B.
+Proof. exact: imsetS. Qed.
+
+Lemma class_set1 x y : x ^: [set y] = [set x ^ y].
+Proof. exact: imset_set1. Qed.
+
+Lemma class1g x A : x \in A -> 1 ^: A = 1.
+Proof.
+move=> Ax; apply/setP=> y.
+by apply/imsetP/set1P=> [[a Aa]|] ->; last exists x; rewrite ?conj1g.
+Qed.
+
+Lemma classVg x A : x^-1 ^: A = (x ^: A)^-1.
+Proof.
+apply/setP=> xy; rewrite inE; apply/imsetP/imsetP=> [] [y Ay def_xy].
+ by rewrite def_xy conjVg invgK; exists y.
+by rewrite -[xy]invgK def_xy -conjVg; exists y.
+Qed.
+
+Lemma mem_classes x A : x \in A -> x ^: A \in classes A.
+Proof. exact: mem_imset. Qed.
+
+Lemma memJ_class_support A B x y :
+ x \in A -> y \in B -> x ^ y \in class_support A B.
+Proof. by move=> Ax By; apply: mem_imset2. Qed.
+
+Lemma class_supportM A B C :
+ class_support A (B * C) = class_support (class_support A B) C.
+Proof.
+apply/setP=> x; apply/imset2P/imset2P=> [[a y Aa] | [y c]].
+ case/mulsgP=> b c Bb Cc -> ->{x y}.
+ by exists (a ^ b) c; rewrite ?(mem_imset2, conjgM).
+case/imset2P=> a b Aa Bb -> Cc ->{x y}.
+by exists a (b * c); rewrite ?(mem_mulg, conjgM).
+Qed.
+
+Lemma class_support_set1l A x : class_support [set x] A = x ^: A.
+Proof. exact: imset2_set1l. Qed.
+
+Lemma class_support_set1r A x : class_support A [set x] = A :^ x.
+Proof. exact: imset2_set1r. Qed.
+
+Lemma classM x A B : x ^: (A * B) = class_support (x ^: A) B.
+Proof. by rewrite -!class_support_set1l class_supportM. Qed.
+
+Lemma class_lcoset x y A : x ^: (y *: A) = (x ^ y) ^: A.
+Proof. by rewrite classM class_set1 class_support_set1l. Qed.
+
+Lemma class_rcoset x A y : x ^: (A :* y) = (x ^: A) :^ y.
+Proof. by rewrite -class_support_set1r classM. Qed.
+
+(* Conjugate set. *)
+
+Lemma conjugatesS A B C : B \subset C -> A :^: B \subset A :^: C.
+Proof. exact: imsetS. Qed.
+
+Lemma conjugates_set1 A x : A :^: [set x] = [set A :^ x].
+Proof. exact: imset_set1. Qed.
+
+Lemma conjugates_conj A x B : (A :^ x) :^: B = A :^: (x *: B).
+Proof.
+rewrite /conjugates [x *: B]imset2_set1l -imset_comp.
+by apply: eq_imset => y /=; rewrite conjsgM.
+Qed.
+
+(* Class support. *)
+
+Lemma class_supportEl A B : class_support A B = \bigcup_(x in A) x ^: B.
+Proof. exact: curry_imset2l. Qed.
+
+Lemma class_supportEr A B : class_support A B = \bigcup_(x in B) A :^ x.
+Proof. exact: curry_imset2r. Qed.
+
+(* Groups (at last!) *)
+
+Definition group_set A := (1 \in A) && (A * A \subset A).
+
+Lemma group_setP A :
+ reflect (1 \in A /\ {in A & A, forall x y, x * y \in A}) (group_set A).
+Proof.
+apply: (iffP andP) => [] [A1 AM]; split=> {A1}//.
+ by move=> x y Ax Ay; apply: (subsetP AM); rewrite mem_mulg.
+apply/subsetP=> _ /mulsgP[x y Ax Ay ->]; exact: AM.
+Qed.
+
+Structure group_type : Type := Group {
+ gval :> GroupSet.sort gT;
+ _ : group_set gval
+}.
+
+Definition group_of of phant gT : predArgType := group_type.
+Notation Local groupT := (group_of (Phant gT)).
+Identity Coercion type_of_group : group_of >-> group_type.
+
+Canonical group_subType := Eval hnf in [subType for gval].
+Definition group_eqMixin := Eval hnf in [eqMixin of group_type by <:].
+Canonical group_eqType := Eval hnf in EqType group_type group_eqMixin.
+Definition group_choiceMixin := [choiceMixin of group_type by <:].
+Canonical group_choiceType :=
+ Eval hnf in ChoiceType group_type group_choiceMixin.
+Definition group_countMixin := [countMixin of group_type by <:].
+Canonical group_countType := Eval hnf in CountType group_type group_countMixin.
+Canonical group_subCountType := Eval hnf in [subCountType of group_type].
+Definition group_finMixin := [finMixin of group_type by <:].
+Canonical group_finType := Eval hnf in FinType group_type group_finMixin.
+Canonical group_subFinType := Eval hnf in [subFinType of group_type].
+
+(* No predType or baseFinGroupType structures, as these would hide the *)
+(* group-to-set coercion and thus spoil unification. *)
+
+Canonical group_of_subType := Eval hnf in [subType of groupT].
+Canonical group_of_eqType := Eval hnf in [eqType of groupT].
+Canonical group_of_choiceType := Eval hnf in [choiceType of groupT].
+Canonical group_of_countType := Eval hnf in [countType of groupT].
+Canonical group_of_subCountType := Eval hnf in [subCountType of groupT].
+Canonical group_of_finType := Eval hnf in [finType of groupT].
+Canonical group_of_subFinType := Eval hnf in [subFinType of groupT].
+
+Definition group (A : {set gT}) gA : groupT := @Group A gA.
+
+Definition clone_group G :=
+ let: Group _ gP := G return {type of Group for G} -> groupT in fun k => k gP.
+
+Lemma group_inj : injective gval. Proof. exact: val_inj. Qed.
+Lemma groupP (G : groupT) : group_set G. Proof. by case: G. Qed.
+
+Lemma congr_group (H K : groupT) : H = K -> H :=: K.
+Proof. exact: congr1. Qed.
+
+Lemma isgroupP A : reflect (exists G : groupT, A = G) (group_set A).
+Proof. by apply: (iffP idP) => [gA | [[B gB] -> //]]; exists (Group gA). Qed.
+
+Lemma group_set_one : group_set 1.
+Proof. by rewrite /group_set set11 mulg1 subxx. Qed.
+
+Canonical one_group := group group_set_one.
+Canonical set1_group := @group [set 1] group_set_one.
+
+Lemma group_setT (phT : phant gT) : group_set (setTfor phT).
+Proof. by apply/group_setP; split=> [|x y _ _]; rewrite inE. Qed.
+
+Canonical setT_group phT := group (group_setT phT).
+
+(* These definitions come early so we can establish the Notation. *)
+Definition generated A := \bigcap_(G : groupT | A \subset G) G.
+Definition gcore A B := \bigcap_(x in B) A :^ x.
+Definition joing A B := generated (A :|: B).
+Definition commutator A B := generated (commg_set A B).
+Definition cycle x := generated [set x].
+Definition order x := #|cycle x|.
+
+End GroupSetMulProp.
+
+Implicit Arguments lcosetP [gT A x y].
+Implicit Arguments lcosetsP [gT A B C].
+Implicit Arguments rcosetP [gT A x y].
+Implicit Arguments rcosetsP [gT A B C].
+Implicit Arguments group_setP [gT A].
+Prenex Implicits group_set mulsgP set1gP.
+Prenex Implicits lcosetP lcosetsP rcosetP rcosetsP group_setP.
+
+Arguments Scope commutator [_ group_scope group_scope].
+Arguments Scope joing [_ group_scope group_scope].
+Arguments Scope generated [_ group_scope].
+
+Notation "{ 'group' gT }" := (group_of (Phant gT))
+ (at level 0, format "{ 'group' gT }") : type_scope.
+
+Notation "[ 'group' 'of' G ]" := (clone_group (@group _ G))
+ (at level 0, format "[ 'group' 'of' G ]") : form_scope.
+
+Bind Scope Group_scope with group_type.
+Bind Scope Group_scope with group_of.
+Notation "1" := (one_group _) : Group_scope.
+Notation "[ 1 gT ]" := (1%G : {group gT}) : Group_scope.
+Notation "[ 'set' : gT ]" := (setT_group (Phant gT)) : Group_scope.
+
+(* Helper notation for defining new groups that need a bespoke finGroupType. *)
+(* The actual group for such a type (say, my_gT) will be the full group, *)
+(* i.e., [set: my_gT] or [set: my_gT]%G, but Coq will not recognize *)
+(* specific notation for these because of the coercions inserted during type *)
+(* inference, unless they are defined as [set: gsort my_gT] using the *)
+(* Notation below. *)
+Notation gsort gT := (FinGroup.arg_sort (FinGroup.base gT%type)) (only parsing).
+Notation "<< A >>" := (generated A) : group_scope.
+Notation "<[ x ] >" := (cycle x) : group_scope.
+Notation "#[ x ]" := (order x) : group_scope.
+Notation "A <*> B" := (joing A B) : group_scope.
+Notation "[ ~: A1 , A2 , .. , An ]" :=
+ (commutator .. (commutator A1 A2) .. An) : group_scope.
+
+Prenex Implicits order cycle gcore.
+
+Section GroupProp.
+
+Variable gT : finGroupType.
+Notation sT := {set gT}.
+Implicit Types A B C D : sT.
+Implicit Types x y z : gT.
+Implicit Types G H K : {group gT}.
+
+Section OneGroup.
+
+Variable G : {group gT}.
+
+Lemma valG : val G = G. Proof. by []. Qed.
+
+(* Non-triviality. *)
+
+Lemma group1 : 1 \in G. Proof. by case/group_setP: (valP G). Qed.
+Hint Resolve group1.
+
+(* Loads of silly variants to placate the incompleteness of trivial. *)
+(* An alternative would be to upgrade done, pending better support *)
+(* in the ssreflect ML code. *)
+Notation gTr := (FinGroup.sort gT).
+Notation Gcl := (pred_of_set G : pred gTr).
+Lemma group1_class1 : (1 : gTr) \in G. Proof. by []. Qed.
+Lemma group1_class2 : 1 \in Gcl. Proof. by []. Qed.
+Lemma group1_class12 : (1 : gTr) \in Gcl. Proof. by []. Qed.
+Lemma group1_eqType : (1 : gT : eqType) \in G. Proof. by []. Qed.
+Lemma group1_finType : (1 : gT : finType) \in G. Proof. by []. Qed.
+
+Lemma group1_contra x : x \notin G -> x != 1.
+Proof. by apply: contraNneq => ->. Qed.
+
+Lemma sub1G : [1 gT] \subset G. Proof. by rewrite sub1set. Qed.
+Lemma subG1 : (G \subset [1]) = (G :==: 1).
+Proof. by rewrite eqEsubset sub1G andbT. Qed.
+
+Lemma setI1g : 1 :&: G = 1. Proof. exact: (setIidPl sub1G). Qed.
+Lemma setIg1 : G :&: 1 = 1. Proof. exact: (setIidPr sub1G). Qed.
+
+Lemma subG1_contra H : G \subset H -> G :!=: 1 -> H :!=: 1.
+Proof. by move=> sGH; rewrite -subG1; apply: contraNneq => <-. Qed.
+
+Lemma repr_group : repr G = 1. Proof. by rewrite /repr group1. Qed.
+
+Lemma cardG_gt0 : 0 < #|G|.
+Proof. by rewrite lt0n; apply/existsP; exists (1 : gT). Qed.
+(* Workaround for the fact that the simple matching used by Trivial does not *)
+(* always allow conversion. In particular cardG_gt0 always fails to apply to *)
+(* subgoals that have been simplified (by /=) because type inference in the *)
+(* notation #|G| introduces redexes of the form *)
+(* Finite.sort (arg_finGroupType (FinGroup.base gT)) *)
+(* which get collapsed to Fingroup.arg_sort (FinGroup.base gT). *)
+Definition cardG_gt0_reduced : 0 < card (@mem gT (predPredType gT) G)
+ := cardG_gt0.
+
+Lemma indexg_gt0 A : 0 < #|G : A|.
+Proof.
+rewrite lt0n; apply/existsP; exists A.
+rewrite -{2}[A]mulg1 -rcosetE; exact: mem_imset.
+Qed.
+
+Lemma trivgP : reflect (G :=: 1) (G \subset [1]).
+Proof. by rewrite subG1; exact: eqP. Qed.
+
+Lemma trivGP : reflect (G = 1%G) (G \subset [1]).
+Proof. by rewrite subG1; exact: eqP. Qed.
+
+Lemma proper1G : ([1] \proper G) = (G :!=: 1).
+Proof. by rewrite properEneq sub1G andbT eq_sym. Qed.
+
+Lemma trivgPn : reflect (exists2 x, x \in G & x != 1) (G :!=: 1).
+Proof.
+rewrite -subG1.
+by apply: (iffP subsetPn) => [] [x Gx x1]; exists x; rewrite ?inE in x1 *.
+Qed.
+
+Lemma trivg_card_le1 : (G :==: 1) = (#|G| <= 1).
+Proof. by rewrite eq_sym eqEcard cards1 sub1G. Qed.
+
+Lemma trivg_card1 : (G :==: 1) = (#|G| == 1%N).
+Proof. by rewrite trivg_card_le1 eqn_leq cardG_gt0 andbT. Qed.
+
+Lemma cardG_gt1 : (#|G| > 1) = (G :!=: 1).
+Proof. by rewrite trivg_card_le1 ltnNge. Qed.
+
+Lemma card_le1_trivg : #|G| <= 1 -> G :=: 1.
+Proof. by rewrite -trivg_card_le1; move/eqP. Qed.
+
+Lemma card1_trivg : #|G| = 1%N -> G :=: 1.
+Proof. by move=> G1; rewrite card_le1_trivg ?G1. Qed.
+
+(* Inclusion and product. *)
+
+Lemma mulG_subl A : A \subset A * G.
+Proof. exact: mulg_subl group1. Qed.
+
+Lemma mulG_subr A : A \subset G * A.
+Proof. exact: mulg_subr group1. Qed.
+
+Lemma mulGid : G * G = G.
+Proof.
+by apply/eqP; rewrite eqEsubset mulG_subr andbT; case/andP: (valP G).
+Qed.
+
+Lemma mulGS A B : (G * A \subset G * B) = (A \subset G * B).
+Proof.
+apply/idP/idP; first exact: subset_trans (mulG_subr A).
+by move/(mulgS G); rewrite mulgA mulGid.
+Qed.
+
+Lemma mulSG A B : (A * G \subset B * G) = (A \subset B * G).
+Proof.
+apply/idP/idP; first exact: subset_trans (mulG_subl A).
+by move/(mulSg G); rewrite -mulgA mulGid.
+Qed.
+
+Lemma mul_subG A B : A \subset G -> B \subset G -> A * B \subset G.
+Proof. by move=> sAG sBG; rewrite -mulGid mulgSS. Qed.
+
+(* Membership lemmas *)
+
+Lemma groupM x y : x \in G -> y \in G -> x * y \in G.
+Proof. by case/group_setP: (valP G) x y. Qed.
+
+Lemma groupX x n : x \in G -> x ^+ n \in G.
+Proof. by move=> Gx; elim: n => [|n IHn]; rewrite ?group1 // expgS groupM. Qed.
+
+Lemma groupVr x : x \in G -> x^-1 \in G.
+Proof.
+move=> Gx; rewrite -(mul1g x^-1) -mem_rcoset ((G :* x =P G) _) //.
+by rewrite eqEcard card_rcoset leqnn mul_subG ?sub1set.
+Qed.
+
+Lemma groupVl x : x^-1 \in G -> x \in G.
+Proof. by move/groupVr; rewrite invgK. Qed.
+
+Lemma groupV x : (x^-1 \in G) = (x \in G).
+Proof. by apply/idP/idP; [exact: groupVl | exact: groupVr]. Qed.
+
+Lemma groupMl x y : x \in G -> (x * y \in G) = (y \in G).
+Proof.
+move=> Gx; apply/idP/idP=> Gy; last exact: groupM.
+rewrite -(mulKg x y); exact: groupM (groupVr _) _.
+Qed.
+
+Lemma groupMr x y : x \in G -> (y * x \in G) = (y \in G).
+Proof. by move=> Gx; rewrite -[_ \in G]groupV invMg groupMl groupV. Qed.
+
+Definition in_group := (group1, groupV, (groupMl, groupX)).
+
+Lemma groupJ x y : x \in G -> y \in G -> x ^ y \in G.
+Proof. by move=> Gx Gy; rewrite !in_group. Qed.
+
+Lemma groupJr x y : y \in G -> (x ^ y \in G) = (x \in G).
+Proof. by move=> Gy; rewrite groupMl (groupMr, groupV). Qed.
+
+Lemma groupR x y : x \in G -> y \in G -> [~ x, y] \in G.
+Proof. by move=> Gx Gy; rewrite !in_group. Qed.
+
+Lemma group_prod I r (P : pred I) F :
+ (forall i, P i -> F i \in G) -> \prod_(i <- r | P i) F i \in G.
+Proof. by move=> G_P; elim/big_ind: _ => //; exact: groupM. Qed.
+
+(* Inverse is an anti-morphism. *)
+
+Lemma invGid : G^-1 = G. Proof. by apply/setP=> x; rewrite inE groupV. Qed.
+
+Lemma inv_subG A : (A^-1 \subset G) = (A \subset G).
+Proof. by rewrite -{1}invGid invSg. Qed.
+
+Lemma invg_lcoset x : (x *: G)^-1 = G :* x^-1.
+Proof. by rewrite invMg invGid invg_set1. Qed.
+
+Lemma invg_rcoset x : (G :* x)^-1 = x^-1 *: G.
+Proof. by rewrite invMg invGid invg_set1. Qed.
+
+Lemma memV_lcosetV x y : (y^-1 \in x^-1 *: G) = (y \in G :* x).
+Proof. by rewrite -invg_rcoset memV_invg. Qed.
+
+Lemma memV_rcosetV x y : (y^-1 \in G :* x^-1) = (y \in x *: G).
+Proof. by rewrite -invg_lcoset memV_invg. Qed.
+
+(* Product idempotence *)
+
+Lemma mulSgGid A x : x \in A -> A \subset G -> A * G = G.
+Proof.
+move=> Ax sAG; apply/eqP; rewrite eqEsubset -{2}mulGid mulSg //=.
+apply/subsetP=> y Gy; rewrite -(mulKVg x y) mem_mulg // groupMr // groupV.
+exact: (subsetP sAG).
+Qed.
+
+Lemma mulGSgid A x : x \in A -> A \subset G -> G * A = G.
+Proof.
+rewrite -memV_invg -invSg invGid => Ax sAG.
+by apply: invg_inj; rewrite invMg invGid (mulSgGid Ax).
+Qed.
+
+(* Left cosets *)
+
+Lemma lcoset_refl x : x \in x *: G.
+Proof. by rewrite mem_lcoset mulVg group1. Qed.
+
+Lemma lcoset_sym x y : (x \in y *: G) = (y \in x *: G).
+Proof. by rewrite !mem_lcoset -groupV invMg invgK. Qed.
+
+Lemma lcoset_transl x y : x \in y *: G -> x *: G = y *: G.
+Proof.
+move=> Gyx; apply/setP=> u; rewrite !mem_lcoset in Gyx *.
+by rewrite -{2}(mulKVg x u) mulgA (groupMl _ Gyx).
+Qed.
+
+Lemma lcoset_transr x y z : x \in y *: G -> (x \in z *: G) = (y \in z *: G).
+Proof. by move=> Gyx; rewrite -2!(lcoset_sym z) (lcoset_transl Gyx). Qed.
+
+Lemma lcoset_trans x y z : x \in y *: G -> y \in z *: G -> x \in z *: G.
+Proof. by move/lcoset_transr->. Qed.
+
+Lemma lcoset_id x : x \in G -> x *: G = G.
+Proof. rewrite -{-2}(mul1g G); exact: lcoset_transl. Qed.
+
+(* Right cosets, with an elimination form for repr. *)
+
+Lemma rcoset_refl x : x \in G :* x.
+Proof. by rewrite mem_rcoset mulgV group1. Qed.
+
+Lemma rcoset_sym x y : (x \in G :* y) = (y \in G :* x).
+Proof. by rewrite -!memV_lcosetV lcoset_sym. Qed.
+
+Lemma rcoset_transl x y : x \in G :* y -> G :* x = G :* y.
+Proof.
+move=> Gyx; apply: invg_inj; rewrite !invg_rcoset.
+by apply: lcoset_transl; rewrite memV_lcosetV.
+Qed.
+
+Lemma rcoset_transr x y z : x \in G :* y -> (x \in G :* z) = (y \in G :* z).
+Proof. by move=> Gyx; rewrite -2!(rcoset_sym z) (rcoset_transl Gyx). Qed.
+
+Lemma rcoset_trans x y z : y \in G :* x -> z \in G :* y -> z \in G :* x.
+Proof. by move/rcoset_transl->. Qed.
+
+Lemma rcoset_id x : x \in G -> G :* x = G.
+Proof. by rewrite -{-2}(mulg1 G); exact: rcoset_transl. Qed.
+
+(* Elimination form. *)
+
+CoInductive rcoset_repr_spec x : gT -> Type :=
+ RcosetReprSpec g : g \in G -> rcoset_repr_spec x (g * x).
+
+Lemma mem_repr_rcoset x : repr (G :* x) \in G :* x.
+Proof. exact: mem_repr (rcoset_refl x). Qed.
+
+(* This form sometimes fails because ssreflect 1.1 delegates matching to the *)
+(* (weaker) primitive Coq algorithm for general (co)inductive type families. *)
+Lemma repr_rcosetP x : rcoset_repr_spec x (repr (G :* x)).
+Proof.
+by rewrite -[repr _](mulgKV x); split; rewrite -mem_rcoset mem_repr_rcoset.
+Qed.
+
+Lemma rcoset_repr x : G :* (repr (G :* x)) = G :* x.
+Proof. by apply: rcoset_transl; exact: mem_repr (rcoset_refl x). Qed.
+
+(* Coset spaces. *)
+
+Lemma mem_lcosets A x : (x *: G \in lcosets G A) = (x \in A * G).
+Proof.
+apply/imsetP/mulsgP=> [[a Aa eqxaG] | [a g Aa Gg ->{x}]].
+ exists a (a^-1 * x); rewrite ?mulKVg //.
+ by rewrite -mem_lcoset -lcosetE -eqxaG lcoset_refl.
+by exists a; rewrite // lcosetM lcosetE lcoset_id.
+Qed.
+
+Lemma mem_rcosets A x : (G :* x \in rcosets G A) = (x \in G * A).
+Proof.
+rewrite -memV_invg invMg invGid -mem_lcosets.
+by rewrite -{4}invGid lcosets_invg inE invg_lcoset invgK.
+Qed.
+
+(* Conjugates. *)
+
+Lemma group_setJ A x : group_set (A :^ x) = group_set A.
+Proof. by rewrite /group_set mem_conjg conj1g -conjsMg conjSg. Qed.
+
+Lemma group_set_conjG x : group_set (G :^ x).
+Proof. by rewrite group_setJ groupP. Qed.
+
+Canonical conjG_group x := group (group_set_conjG x).
+
+Lemma conjGid : {in G, normalised G}.
+Proof. by move=> x Gx; apply/setP=> y; rewrite mem_conjg groupJr ?groupV. Qed.
+
+Lemma conj_subG x A : x \in G -> A \subset G -> A :^ x \subset G.
+Proof. by move=> Gx sAG; rewrite -(conjGid Gx) conjSg. Qed.
+
+(* Classes *)
+
+Lemma class1G : 1 ^: G = 1. Proof. exact: class1g group1. Qed.
+
+Lemma classes1 : [1] \in classes G. Proof. by rewrite -class1G mem_classes. Qed.
+
+Lemma classGidl x y : y \in G -> (x ^ y) ^: G = x ^: G.
+Proof. by move=> Gy; rewrite -class_lcoset lcoset_id. Qed.
+
+Lemma classGidr x : {in G, normalised (x ^: G)}.
+Proof. by move=> y Gy /=; rewrite -class_rcoset rcoset_id. Qed.
+
+Lemma class_refl x : x \in x ^: G.
+Proof. by apply/imsetP; exists (1 : gT); rewrite ?conjg1. Qed.
+Hint Resolve class_refl.
+
+Lemma class_transr x y : x \in y ^: G -> x ^: G = y ^: G.
+Proof. by case/imsetP=> z Gz ->; rewrite classGidl. Qed.
+
+Lemma class_sym x y : (x \in y ^: G) = (y \in x ^: G).
+Proof. by apply/idP/idP=> /class_transr->. Qed.
+
+Lemma class_transl x y z : x \in y ^: G -> (x \in z ^: G) = (y \in z ^: G).
+Proof. by rewrite -!(class_sym z) => /class_transr->. Qed.
+
+Lemma class_trans x y z : x \in y ^: G -> y \in z ^: G -> x \in z ^: G.
+Proof. by move/class_transl->. Qed.
+
+Lemma repr_class x : {y | y \in G & repr (x ^: G) = x ^ y}.
+Proof.
+set z := repr _; have: #|[set y in G | z == x ^ y]| > 0.
+ have: z \in x ^: G by exact: (mem_repr x).
+ by case/imsetP=> y Gy ->; rewrite (cardD1 y) inE Gy eqxx.
+by move/card_mem_repr; move: (repr _) => y /setIdP[Gy /eqP]; exists y.
+Qed.
+
+Lemma classG_eq1 x : (x ^: G == 1) = (x == 1).
+Proof.
+apply/eqP/eqP=> [xG1 | ->]; last exact: class1G.
+by have:= class_refl x; rewrite xG1 => /set1P.
+Qed.
+
+Lemma class_subG x A : x \in G -> A \subset G -> x ^: A \subset G.
+Proof.
+move=> Gx sAG; apply/subsetP=> _ /imsetP[y Ay ->].
+by rewrite groupJ // (subsetP sAG).
+Qed.
+
+Lemma repr_classesP xG :
+ reflect (repr xG \in G /\ xG = repr xG ^: G) (xG \in classes G).
+Proof.
+apply: (iffP imsetP) => [[x Gx ->] | []]; last by exists (repr xG).
+by have [y Gy ->] := repr_class x; rewrite classGidl ?groupJ.
+Qed.
+
+Lemma mem_repr_classes xG : xG \in classes G -> repr xG \in xG.
+Proof. by case/repr_classesP=> _ {2}->; exact: class_refl. Qed.
+
+Lemma classes_gt0 : 0 < #|classes G|.
+Proof. by rewrite (cardsD1 1) classes1. Qed.
+
+Lemma classes_gt1 : (#|classes G| > 1) = (G :!=: 1).
+Proof.
+rewrite (cardsD1 1) classes1 ltnS lt0n cards_eq0.
+apply/set0Pn/trivgPn=> [[xG /setD1P[nt_xG]] | [x Gx ntx]].
+ by case/imsetP=> x Gx def_xG; rewrite def_xG classG_eq1 in nt_xG; exists x.
+by exists (x ^: G); rewrite !inE classG_eq1 ntx; exact: mem_imset.
+Qed.
+
+Lemma mem_class_support A x : x \in A -> x \in class_support A G.
+Proof. by move=> Ax; rewrite -[x]conjg1 memJ_class_support. Qed.
+
+Lemma class_supportGidl A x :
+ x \in G -> class_support (A :^ x) G = class_support A G.
+Proof.
+by move=> Gx; rewrite -class_support_set1r -class_supportM lcoset_id.
+Qed.
+
+Lemma class_supportGidr A : {in G, normalised (class_support A G)}.
+Proof.
+by move=> x Gx /=; rewrite -class_support_set1r -class_supportM rcoset_id.
+Qed.
+
+Lemma class_support_subG A : A \subset G -> class_support A G \subset G.
+Proof.
+by move=> sAG; rewrite class_supportEr; apply/bigcupsP=> x Gx; exact: conj_subG.
+Qed.
+
+Lemma sub_class_support A : A \subset class_support A G.
+Proof. by rewrite class_supportEr (bigcup_max 1) ?conjsg1. Qed.
+
+Lemma class_support_id : class_support G G = G.
+Proof.
+by apply/eqP; rewrite eqEsubset sub_class_support class_support_subG.
+Qed.
+
+Lemma class_supportD1 A : (class_support A G)^# = cover (A^# :^: G).
+Proof.
+rewrite cover_imset class_supportEr setDE big_distrl /=.
+by apply: eq_bigr => x _; rewrite -setDE conjD1g.
+Qed.
+
+(* Subgroup Type construction. *)
+(* We only expect to use this for abstract groups, so we don't project *)
+(* the argument to a set. *)
+
+Inductive subg_of : predArgType := Subg x & x \in G.
+Definition sgval u := let: Subg x _ := u in x.
+Canonical subg_subType := Eval hnf in [subType for sgval].
+Definition subg_eqMixin := Eval hnf in [eqMixin of subg_of by <:].
+Canonical subg_eqType := Eval hnf in EqType subg_of subg_eqMixin.
+Definition subg_choiceMixin := [choiceMixin of subg_of by <:].
+Canonical subg_choiceType := Eval hnf in ChoiceType subg_of subg_choiceMixin.
+Definition subg_countMixin := [countMixin of subg_of by <:].
+Canonical subg_countType := Eval hnf in CountType subg_of subg_countMixin.
+Canonical subg_subCountType := Eval hnf in [subCountType of subg_of].
+Definition subg_finMixin := [finMixin of subg_of by <:].
+Canonical subg_finType := Eval hnf in FinType subg_of subg_finMixin.
+Canonical subg_subFinType := Eval hnf in [subFinType of subg_of].
+
+Lemma subgP u : sgval u \in G.
+Proof. exact: valP. Qed.
+Lemma subg_inj : injective sgval.
+Proof. exact: val_inj. Qed.
+Lemma congr_subg u v : u = v -> sgval u = sgval v.
+Proof. exact: congr1. Qed.
+
+Definition subg_one := Subg group1.
+Definition subg_inv u := Subg (groupVr (subgP u)).
+Definition subg_mul u v := Subg (groupM (subgP u) (subgP v)).
+Lemma subg_oneP : left_id subg_one subg_mul.
+Proof. move=> u; apply: val_inj; exact: mul1g. Qed.
+
+Lemma subg_invP : left_inverse subg_one subg_inv subg_mul.
+Proof. move=> u; apply: val_inj; exact: mulVg. Qed.
+Lemma subg_mulP : associative subg_mul.
+Proof. move=> u v w; apply: val_inj; exact: mulgA. Qed.
+
+Definition subFinGroupMixin := FinGroup.Mixin subg_mulP subg_oneP subg_invP.
+Canonical subBaseFinGroupType :=
+ Eval hnf in BaseFinGroupType subg_of subFinGroupMixin.
+Canonical subFinGroupType := FinGroupType subg_invP.
+
+Lemma sgvalM : {in setT &, {morph sgval : x y / x * y}}. Proof. by []. Qed.
+Lemma valgM : {in setT &, {morph val : x y / (x : subg_of) * y >-> x * y}}.
+Proof. by []. Qed.
+
+Definition subg : gT -> subg_of := insubd (1 : subg_of).
+Lemma subgK x : x \in G -> val (subg x) = x.
+Proof. by move=> Gx; rewrite insubdK. Qed.
+Lemma sgvalK : cancel sgval subg.
+Proof. case=> x Gx; apply: val_inj; exact: subgK. Qed.
+Lemma subg_default x : (x \in G) = false -> val (subg x) = 1.
+Proof. by move=> Gx; rewrite val_insubd Gx. Qed.
+Lemma subgM : {in G &, {morph subg : x y / x * y}}.
+Proof. by move=> x y Gx Gy; apply: val_inj; rewrite /= !subgK ?groupM. Qed.
+
+End OneGroup.
+
+Hint Resolve group1.
+
+Lemma groupD1_inj G H : G^# = H^# -> G :=: H.
+Proof. by move/(congr1 (setU 1)); rewrite !setD1K. Qed.
+
+Lemma invMG G H : (G * H)^-1 = H * G.
+Proof. by rewrite invMg !invGid. Qed.
+
+Lemma mulSGid G H : H \subset G -> H * G = G.
+Proof. exact: mulSgGid (group1 H). Qed.
+
+Lemma mulGSid G H : H \subset G -> G * H = G.
+Proof. exact: mulGSgid (group1 H). Qed.
+
+Lemma mulGidPl G H : reflect (G * H = G) (H \subset G).
+Proof. by apply: (iffP idP) => [|<-]; [exact: mulGSid | exact: mulG_subr]. Qed.
+
+Lemma mulGidPr G H : reflect (G * H = H) (G \subset H).
+Proof. by apply: (iffP idP) => [|<-]; [exact: mulSGid | exact: mulG_subl]. Qed.
+
+Lemma comm_group_setP G H : reflect (commute G H) (group_set (G * H)).
+Proof.
+rewrite /group_set (subsetP (mulG_subl _ _)) ?group1 // andbC.
+have <-: #|G * H| <= #|H * G| by rewrite -invMG card_invg.
+rewrite -mulgA mulGS mulgA mulSG -eqEcard eq_sym; exact: eqP.
+Qed.
+
+Lemma card_lcosets G H : #|lcosets H G| = #|G : H|.
+Proof.
+by rewrite -[#|G : H|](card_preimset _ invg_inj) -lcosets_invg !invGid.
+Qed.
+
+(* Group Modularity equations *)
+
+Lemma group_modl A B G : A \subset G -> A * (B :&: G) = A * B :&: G.
+Proof.
+move=> sAG; apply/eqP; rewrite eqEsubset subsetI mulgS ?subsetIl //.
+rewrite -{2}mulGid mulgSS ?subsetIr //.
+apply/subsetP => _ /setIP[/mulsgP[a b Aa Bb ->] Gab].
+by rewrite mem_mulg // inE Bb -(groupMl _ (subsetP sAG _ Aa)).
+Qed.
+
+Lemma group_modr A B G : B \subset G -> (G :&: A) * B = G :&: A * B.
+Proof.
+move=> sBG; apply: invg_inj; rewrite !(invMg, invIg) invGid !(setIC G).
+by rewrite group_modl // -invGid invSg.
+Qed.
+
+End GroupProp.
+
+Hint Resolve group1 group1_class1 group1_class12 group1_class12.
+Hint Resolve group1_eqType group1_finType.
+Hint Resolve cardG_gt0 cardG_gt0_reduced indexg_gt0.
+
+Notation "G :^ x" := (conjG_group G x) : Group_scope.
+
+Notation "[ 'subg' G ]" := (subg_of G) : type_scope.
+Notation "[ 'subg' G ]" := [set: subg_of G] : group_scope.
+Notation "[ 'subg' G ]" := [set: subg_of G]%G : Group_scope.
+
+Prenex Implicits subg sgval subg_of.
+Bind Scope group_scope with subg_of.
+
+Implicit Arguments trivgP [gT G].
+Implicit Arguments trivGP [gT G].
+Implicit Arguments mulGidPl [gT G H].
+Implicit Arguments mulGidPr [gT G H].
+Implicit Arguments comm_group_setP [gT G H].
+Implicit Arguments repr_classesP [gT G xG].
+Prenex Implicits trivgP trivGP comm_group_setP.
+
+Section GroupInter.
+
+Variable gT : finGroupType.
+Implicit Types A B : {set gT}.
+Implicit Types G H : {group gT}.
+
+Lemma group_setI G H : group_set (G :&: H).
+Proof.
+apply/group_setP; split=> [|x y]; rewrite !inE ?group1 //.
+by case/andP=> Gx Hx; rewrite !groupMl.
+Qed.
+
+Canonical setI_group G H := group (group_setI G H).
+
+Section Nary.
+
+Variables (I : finType) (P : pred I) (F : I -> {group gT}).
+
+Lemma group_set_bigcap : group_set (\bigcap_(i | P i) F i).
+Proof.
+elim/big_rec: _ => [|i G _ gG]; first exact: groupP.
+exact: group_setI (Group gG).
+Qed.
+
+Canonical bigcap_group := group group_set_bigcap.
+
+End Nary.
+
+Canonical generated_group A : {group _} := Eval hnf in [group of <<A>>].
+Canonical gcore_group G A : {group _} := Eval hnf in [group of gcore G A].
+Canonical commutator_group A B : {group _} := Eval hnf in [group of [~: A, B]].
+Canonical joing_group A B : {group _} := Eval hnf in [group of A <*> B].
+Canonical cycle_group x : {group _} := Eval hnf in [group of <[x]>].
+
+Lemma order_gt0 (x : gT) : 0 < #[x].
+Proof. exact: cardG_gt0. Qed.
+
+End GroupInter.
+
+Hint Resolve order_gt0.
+
+Definition joinG (gT : finGroupType) (G H : {group gT}) := joing_group G H.
+
+Definition subgroups (gT : finGroupType) (G : {set gT}) :=
+ [set H : {group gT} | H \subset G].
+
+Arguments Scope generated_group [_ group_scope].
+Arguments Scope joing_group [_ group_scope group_scope].
+
+Notation "G :&: H" := (setI_group G H) : Group_scope.
+Notation "<< A >>" := (generated_group A) : Group_scope.
+Notation "<[ x ] >" := (cycle_group x) : Group_scope.
+Notation "[ ~: A1 , A2 , .. , An ]" :=
+ (commutator_group .. (commutator_group A1 A2) .. An) : Group_scope.
+Notation "A <*> B" := (joing_group A B) : Group_scope.
+Notation "G * H" := (joinG G H) : Group_scope.
+Prenex Implicits joinG.
+
+Notation "\prod_ ( i <- r | P ) F" :=
+ (\big[joinG/1%G]_(i <- r | P%B) F%G) : Group_scope.
+Notation "\prod_ ( i <- r ) F" :=
+ (\big[joinG/1%G]_(i <- r) F%G) : Group_scope.
+Notation "\prod_ ( m <= i < n | P ) F" :=
+ (\big[joinG/1%G]_(m <= i < n | P%B) F%G) : Group_scope.
+Notation "\prod_ ( m <= i < n ) F" :=
+ (\big[joinG/1%G]_(m <= i < n) F%G) : Group_scope.
+Notation "\prod_ ( i | P ) F" :=
+ (\big[joinG/1%G]_(i | P%B) F%G) : Group_scope.
+Notation "\prod_ i F" :=
+ (\big[joinG/1%G]_i F%G) : Group_scope.
+Notation "\prod_ ( i : t | P ) F" :=
+ (\big[joinG/1%G]_(i : t | P%B) F%G) (only parsing) : Group_scope.
+Notation "\prod_ ( i : t ) F" :=
+ (\big[joinG/1%G]_(i : t) F%G) (only parsing) : Group_scope.
+Notation "\prod_ ( i < n | P ) F" :=
+ (\big[joinG/1%G]_(i < n | P%B) F%G) : Group_scope.
+Notation "\prod_ ( i < n ) F" :=
+ (\big[joinG/1%G]_(i < n) F%G) : Group_scope.
+Notation "\prod_ ( i 'in' A | P ) F" :=
+ (\big[joinG/1%G]_(i in A | P%B) F%G) : Group_scope.
+Notation "\prod_ ( i 'in' A ) F" :=
+ (\big[joinG/1%G]_(i in A) F%G) : Group_scope.
+
+Section Lagrange.
+
+Variable gT : finGroupType.
+Implicit Types G H K : {group gT}.
+
+Lemma LagrangeI G H : (#|G :&: H| * #|G : H|)%N = #|G|.
+Proof.
+rewrite -[#|G|]sum1_card (partition_big_imset (rcoset H)) /=.
+rewrite mulnC -sum_nat_const; apply: eq_bigr => _ /rcosetsP[x Gx ->].
+rewrite -(card_rcoset _ x) -sum1_card; apply: eq_bigl => y.
+rewrite rcosetE eqEcard mulGS !card_rcoset leqnn andbT.
+by rewrite group_modr sub1set // inE.
+Qed.
+
+Lemma divgI G H : #|G| %/ #|G :&: H| = #|G : H|.
+Proof. by rewrite -(LagrangeI G H) mulKn ?cardG_gt0. Qed.
+
+Lemma divg_index G H : #|G| %/ #|G : H| = #|G :&: H|.
+Proof. by rewrite -(LagrangeI G H) mulnK. Qed.
+
+Lemma dvdn_indexg G H : #|G : H| %| #|G|.
+Proof. by rewrite -(LagrangeI G H) dvdn_mull. Qed.
+
+Theorem Lagrange G H : H \subset G -> (#|H| * #|G : H|)%N = #|G|.
+Proof. by move/setIidPr=> sHG; rewrite -{1}sHG LagrangeI. Qed.
+
+Lemma cardSg G H : H \subset G -> #|H| %| #|G|.
+Proof. by move/Lagrange <-; rewrite dvdn_mulr. Qed.
+
+Lemma lognSg p G H : G \subset H -> logn p #|G| <= logn p #|H|.
+Proof. by move=> sGH; rewrite dvdn_leq_log ?cardSg. Qed.
+
+Lemma piSg G H : G \subset H -> {subset \pi(gval G) <= \pi(gval H)}.
+Proof.
+move=> sGH p; rewrite !mem_primes !cardG_gt0 => /and3P[-> _ pG].
+exact: dvdn_trans (cardSg sGH).
+Qed.
+
+Lemma divgS G H : H \subset G -> #|G| %/ #|H| = #|G : H|.
+Proof. by move/Lagrange <-; rewrite mulKn. Qed.
+
+Lemma divg_indexS G H : H \subset G -> #|G| %/ #|G : H| = #|H|.
+Proof. by move/Lagrange <-; rewrite mulnK. Qed.
+
+Lemma coprimeSg G H p : H \subset G -> coprime #|G| p -> coprime #|H| p.
+Proof. by move=> sHG; exact: coprime_dvdl (cardSg sHG). Qed.
+
+Lemma coprimegS G H p : H \subset G -> coprime p #|G| -> coprime p #|H|.
+Proof. by move=> sHG; exact: coprime_dvdr (cardSg sHG). Qed.
+
+Lemma indexJg G H x : #|G :^ x : H :^ x| = #|G : H|.
+Proof. by rewrite -!divgI -conjIg !cardJg. Qed.
+
+Lemma indexgg G : #|G : G| = 1%N.
+Proof. by rewrite -divgS // divnn cardG_gt0. Qed.
+
+Lemma rcosets_id G : rcosets G G = [set G : {set gT}].
+Proof.
+apply/esym/eqP; rewrite eqEcard sub1set [#|_|]indexgg cards1 andbT.
+by apply/rcosetsP; exists 1; rewrite ?mulg1.
+Qed.
+
+Lemma Lagrange_index G H K :
+ H \subset G -> K \subset H -> (#|G : H| * #|H : K|)%N = #|G : K|.
+Proof.
+move=> sHG sKH; apply/eqP; rewrite mulnC -(eqn_pmul2l (cardG_gt0 K)).
+by rewrite mulnA !Lagrange // (subset_trans sKH).
+Qed.
+
+Lemma indexgI G H : #|G : G :&: H| = #|G : H|.
+Proof. by rewrite -divgI divgS ?subsetIl. Qed.
+
+Lemma indexgS G H K : H \subset K -> #|G : K| %| #|G : H|.
+Proof.
+move=> sHK; rewrite -(@dvdn_pmul2l #|G :&: K|) ?cardG_gt0 // LagrangeI.
+by rewrite -(Lagrange (setIS G sHK)) mulnAC LagrangeI dvdn_mulr.
+Qed.
+
+Lemma indexSg G H K : H \subset K -> K \subset G -> #|K : H| %| #|G : H|.
+Proof.
+move=> sHK sKG; rewrite -(@dvdn_pmul2l #|H|) ?cardG_gt0 //.
+by rewrite !Lagrange ?(cardSg, subset_trans sHK).
+Qed.
+
+Lemma indexg_eq1 G H : (#|G : H| == 1%N) = (G \subset H).
+Proof.
+rewrite eqn_leq -(leq_pmul2l (cardG_gt0 (G :&: H))) LagrangeI muln1.
+by rewrite indexg_gt0 andbT (sameP setIidPl eqP) eqEcard subsetIl.
+Qed.
+
+Lemma indexg_gt1 G H : (#|G : H| > 1) = ~~ (G \subset H).
+Proof. by rewrite -indexg_eq1 eqn_leq indexg_gt0 andbT -ltnNge. Qed.
+
+Lemma index1g G H : H \subset G -> #|G : H| = 1%N -> H :=: G.
+Proof. by move=> sHG iHG; apply/eqP; rewrite eqEsubset sHG -indexg_eq1 iHG. Qed.
+
+Lemma indexg1 G : #|G : 1| = #|G|.
+Proof. by rewrite -divgS ?sub1G // cards1 divn1. Qed.
+
+Lemma indexMg G A : #|G * A : G| = #|A : G|.
+Proof.
+congr #|(_ : {set _})|; apply/eqP; rewrite eqEsubset andbC imsetS ?mulG_subr //.
+by apply/subsetP=> _ /imsetP[x GAx ->]; rewrite rcosetE mem_rcosets.
+Qed.
+
+Lemma rcosets_partition_mul G H : partition (rcosets H G) (H * G).
+Proof.
+have eqiR: {in H * G & &, equivalence_rel [rel x y | y \in rcoset H x]}.
+ by move=> *; rewrite /= !rcosetE rcoset_refl; split=> // /rcoset_transl->.
+congr (partition _ _): (equivalence_partitionP eqiR); apply/setP=> Hx.
+apply/imsetP/idP=> [[x HGx defHx] | /rcosetsP[x Gx ->]].
+ suffices ->: Hx = H :* x by rewrite mem_rcosets.
+ apply/setP=> y; rewrite defHx inE /= rcosetE andb_idl //.
+ by apply: subsetP y; rewrite mulGS sub1set.
+exists (1 * x); rewrite ?mem_mulg // mul1g.
+apply/setP=> y; rewrite inE /= rcosetE andb_idl //.
+by apply: subsetP y; rewrite mulgS ?sub1set.
+Qed.
+
+Lemma rcosets_partition G H : H \subset G -> partition (rcosets H G) G.
+Proof. by move/mulSGid=> {2}<-; exact: rcosets_partition_mul. Qed.
+
+Lemma LagrangeMl G H : (#|G| * #|H : G|)%N = #|G * H|.
+Proof.
+rewrite mulnC -(card_uniform_partition _ (rcosets_partition_mul H G)) //.
+by move=> _ /rcosetsP[x Hx ->]; rewrite card_rcoset.
+Qed.
+
+Lemma LagrangeMr G H : (#|G : H| * #|H|)%N = #|G * H|.
+Proof. by rewrite mulnC LagrangeMl -card_invg invMg !invGid. Qed.
+
+Lemma mul_cardG G H : (#|G| * #|H| = #|G * H|%g * #|G :&: H|)%N.
+Proof. by rewrite -LagrangeMr -(LagrangeI G H) -mulnA mulnC. Qed.
+
+Lemma dvdn_cardMg G H : #|G * H| %| #|G| * #|H|.
+Proof. by rewrite mul_cardG dvdn_mulr. Qed.
+
+Lemma cardMg_divn G H : #|G * H| = (#|G| * #|H|) %/ #|G :&: H|.
+Proof. by rewrite mul_cardG mulnK ?cardG_gt0. Qed.
+
+Lemma cardIg_divn G H : #|G :&: H| = (#|G| * #|H|) %/ #|G * H|.
+Proof. by rewrite mul_cardG mulKn // (cardD1 (1 * 1)) mem_mulg. Qed.
+
+Lemma TI_cardMg G H : G :&: H = 1 -> #|G * H| = (#|G| * #|H|)%N.
+Proof. by move=> tiGH; rewrite mul_cardG tiGH cards1 muln1. Qed.
+
+Lemma cardMg_TI G H : #|G| * #|H| <= #|G * H| -> G :&: H = 1.
+Proof.
+move=> leGH; apply: card_le1_trivg.
+rewrite -(@leq_pmul2l #|G * H|); first by rewrite -mul_cardG muln1.
+by apply: leq_trans leGH; rewrite muln_gt0 !cardG_gt0.
+Qed.
+
+Lemma coprime_TIg G H : coprime #|G| #|H| -> G :&: H = 1.
+Proof.
+move=> coGH; apply/eqP; rewrite trivg_card1 -dvdn1 -{}(eqnP coGH).
+by rewrite dvdn_gcd /= {2}setIC !cardSg ?subsetIl.
+Qed.
+
+Lemma prime_TIg G H : prime #|G| -> ~~ (G \subset H) -> G :&: H = 1.
+Proof.
+case/primeP=> _; move/(_ _ (cardSg (subsetIl G H))).
+rewrite (sameP setIidPl eqP) eqEcard subsetIl -ltnNge ltn_neqAle -trivg_card1.
+by case/predU1P=> ->.
+Qed.
+
+Lemma prime_meetG G H : prime #|G| -> G :&: H != 1 -> G \subset H.
+Proof. by move=> prG; apply: contraR; move/prime_TIg->. Qed.
+
+Lemma coprime_cardMg G H : coprime #|G| #|H| -> #|G * H| = (#|G| * #|H|)%N.
+Proof. by move=> coGH; rewrite TI_cardMg ?coprime_TIg. Qed.
+
+Lemma coprime_index_mulG G H K :
+ H \subset G -> K \subset G -> coprime #|G : H| #|G : K| -> H * K = G.
+Proof.
+move=> sHG sKG co_iG_HK; apply/eqP; rewrite eqEcard mul_subG //=.
+rewrite -(@leq_pmul2r #|H :&: K|) ?cardG_gt0 // -mul_cardG.
+rewrite -(Lagrange sHG) -(LagrangeI K H) mulnAC setIC -mulnA.
+rewrite !leq_pmul2l ?cardG_gt0 // dvdn_leq // -(Gauss_dvdr _ co_iG_HK).
+by rewrite -(indexgI K) Lagrange_index ?indexgS ?subsetIl ?subsetIr.
+Qed.
+
+End Lagrange.
+
+Section GeneratedGroup.
+
+Variable gT : finGroupType.
+Implicit Types x y z : gT.
+Implicit Types A B C D : {set gT}.
+Implicit Types G H K : {group gT}.
+
+Lemma subset_gen A : A \subset <<A>>.
+Proof. exact/bigcapsP. Qed.
+
+Lemma sub_gen A B : A \subset B -> A \subset <<B>>.
+Proof. by move/subset_trans=> -> //; exact: subset_gen. Qed.
+
+Lemma mem_gen x A : x \in A -> x \in <<A>>.
+Proof. exact: subsetP (subset_gen A) x. Qed.
+
+Lemma generatedP x A : reflect (forall G, A \subset G -> x \in G) (x \in <<A>>).
+Proof. exact: bigcapP. Qed.
+
+Lemma gen_subG A G : (<<A>> \subset G) = (A \subset G).
+Proof.
+apply/idP/idP=> [|sAG]; first exact: subset_trans (subset_gen A).
+by apply/subsetP=> x /generatedP; apply.
+Qed.
+
+Lemma genGid G : <<G>> = G.
+Proof. by apply/eqP; rewrite eqEsubset gen_subG subset_gen andbT. Qed.
+
+Lemma genGidG G : <<G>>%G = G.
+Proof. by apply: val_inj; exact: genGid. Qed.
+
+Lemma gen_set_id A : group_set A -> <<A>> = A.
+Proof. by move=> gA; exact: (genGid (group gA)). Qed.
+
+Lemma genS A B : A \subset B -> <<A>> \subset <<B>>.
+Proof. by move=> sAB; rewrite gen_subG sub_gen. Qed.
+
+Lemma gen0 : <<set0>> = 1 :> {set gT}.
+Proof. by apply/eqP; rewrite eqEsubset sub1G gen_subG sub0set. Qed.
+
+Lemma gen_expgs A : {n | <<A>> = (1 |: A) ^+ n}.
+Proof.
+set B := (1 |: A); pose N := #|gT|.
+have BsubG n : B ^+ n \subset <<A>>.
+ by elim: n => [|n IHn]; rewrite ?expgS ?mul_subG ?subUset ?sub1G ?subset_gen.
+have B_1 n : 1 \in B ^+ n.
+ by elim: n => [|n IHn]; rewrite ?set11 // expgS mulUg mul1g inE IHn.
+case: (pickP (fun i : 'I_N => B ^+ i.+1 \subset B ^+ i)) => [n fixBn | no_fix].
+ exists n; apply/eqP; rewrite eqEsubset BsubG andbT.
+ rewrite -[B ^+ n]gen_set_id ?genS ?subsetUr //.
+ by apply: subset_trans fixBn; rewrite expgS mulUg subsetU ?mulg_subl ?orbT.
+ rewrite /group_set B_1 /=.
+ elim: {2}(n : nat) => [|m IHm]; first by rewrite mulg1.
+ by apply: subset_trans fixBn; rewrite !expgSr mulgA mulSg.
+suffices: N < #|B ^+ N| by rewrite ltnNge max_card.
+elim: {-2}N (leqnn N) => [|n IHn] lt_nN; first by rewrite cards1.
+apply: leq_ltn_trans (IHn (ltnW lt_nN)) (proper_card _).
+by rewrite /proper (no_fix (Ordinal lt_nN)) expgS mulUg mul1g subsetUl.
+Qed.
+
+Lemma gen_prodgP A x :
+ reflect (exists n, exists2 c, forall i : 'I_n, c i \in A & x = \prod_i c i)
+ (x \in <<A>>).
+Proof.
+apply: (iffP idP) => [|[n [c Ac ->]]]; last first.
+ by apply: group_prod => i _; rewrite mem_gen ?Ac.
+have [n ->] := gen_expgs A; rewrite /expgn /expgn_rec Monoid.iteropE.
+have ->: n = count 'I_n (index_enum _).
+ by rewrite -size_filter filter_index_enum -cardT card_ord.
+rewrite -big_const_seq; case/prodsgP=> /= c Ac def_x.
+have{Ac def_x} ->: x = \prod_(i | c i \in A) c i.
+ rewrite big_mkcond {x}def_x; apply: eq_bigr => i _.
+ by case/setU1P: (Ac i isT) => -> //; rewrite if_same.
+rewrite -big_filter; set e := filter _ _; case def_e: e => [|i e'].
+ by exists 0; exists (fun _ => 1) => [[] // |]; rewrite big_nil big_ord0.
+rewrite -{e'}def_e (big_nth i) big_mkord.
+exists (size e); exists (c \o nth i e \o val) => // j /=.
+have: nth i e j \in e by rewrite mem_nth.
+by rewrite mem_filter; case/andP.
+Qed.
+
+Lemma genD A B : A \subset <<A :\: B>> -> <<A :\: B>> = <<A>>.
+Proof.
+by move=> sAB; apply/eqP; rewrite eqEsubset genS (subsetDl, gen_subG).
+Qed.
+
+Lemma genV A : <<A^-1>> = <<A>>.
+Proof.
+apply/eqP; rewrite eqEsubset !gen_subG -!(invSg _ <<_>>) invgK.
+by rewrite !invGid !subset_gen.
+Qed.
+
+Lemma genJ A z : <<A :^z>> = <<A>> :^ z.
+Proof.
+by apply/eqP; rewrite eqEsubset sub_conjg !gen_subG conjSg -?sub_conjg !sub_gen.
+Qed.
+
+Lemma conjYg A B z : (A <*> B) :^z = A :^ z <*> B :^ z.
+Proof. by rewrite -genJ conjUg. Qed.
+
+Lemma genD1 A x : x \in <<A :\ x>> -> <<A :\ x>> = <<A>>.
+Proof.
+move=> gA'x; apply/eqP; rewrite eqEsubset genS; last by rewrite subsetDl.
+rewrite gen_subG; apply/subsetP=> y Ay.
+by case: (y =P x) => [-> //|]; move/eqP=> nyx; rewrite mem_gen // !inE nyx.
+Qed.
+
+Lemma genD1id A : <<A^#>> = <<A>>.
+Proof. by rewrite genD1 ?group1. Qed.
+
+Notation joingT := (@joing gT) (only parsing).
+Notation joinGT := (@joinG gT) (only parsing).
+
+Lemma joingE A B : A <*> B = <<A :|: B>>. Proof. by []. Qed.
+
+Lemma joinGE G H : (G * H)%G = (G <*> H)%G. Proof. by []. Qed.
+
+Lemma joingC : commutative joingT.
+Proof. by move=> A B; rewrite /joing setUC. Qed.
+
+Lemma joing_idr A B : A <*> <<B>> = A <*> B.
+Proof.
+apply/eqP; rewrite eqEsubset gen_subG subUset gen_subG /=.
+by rewrite -subUset subset_gen genS // setUS // subset_gen.
+Qed.
+
+Lemma joing_idl A B : <<A>> <*> B = A <*> B.
+Proof. by rewrite -!(joingC B) joing_idr. Qed.
+
+Lemma joing_subl A B : A \subset A <*> B.
+Proof. by rewrite sub_gen ?subsetUl. Qed.
+
+Lemma joing_subr A B : B \subset A <*> B.
+Proof. by rewrite sub_gen ?subsetUr. Qed.
+
+Lemma join_subG A B G : (A <*> B \subset G) = (A \subset G) && (B \subset G).
+Proof. by rewrite gen_subG subUset. Qed.
+
+Lemma joing_idPl G A : reflect (G <*> A = G) (A \subset G).
+Proof.
+apply: (iffP idP) => [sHG | <-]; last by rewrite joing_subr.
+by rewrite joingE (setUidPl sHG) genGid.
+Qed.
+
+Lemma joing_idPr A G : reflect (A <*> G = G) (A \subset G).
+Proof. by rewrite joingC; exact: joing_idPl. Qed.
+
+Lemma joing_subP A B G :
+ reflect (A \subset G /\ B \subset G) (A <*> B \subset G).
+Proof. by rewrite join_subG; exact: andP. Qed.
+
+Lemma joing_sub A B C : A <*> B = C -> A \subset C /\ B \subset C.
+Proof. by move <-; exact/joing_subP. Qed.
+
+Lemma genDU A B C : A \subset C -> <<C :\: A>> = <<B>> -> <<A :|: B>> = <<C>>.
+Proof.
+move=> sAC; rewrite -joingE -joing_idr => <- {B}; rewrite joing_idr.
+by congr <<_>>; rewrite setDE setUIr setUCr setIT; exact/setUidPr.
+Qed.
+
+Lemma joingA : associative joingT.
+Proof. by move=> A B C; rewrite joing_idl joing_idr /joing setUA. Qed.
+
+Lemma joing1G G : 1 <*> G = G.
+Proof. by rewrite -gen0 joing_idl /joing set0U genGid. Qed.
+
+Lemma joingG1 G : G <*> 1 = G.
+Proof. by rewrite joingC joing1G. Qed.
+
+Lemma genM_join G H : <<G * H>> = G <*> H.
+Proof.
+apply/eqP; rewrite eqEsubset gen_subG /= -{1}[G <*> H]mulGid.
+rewrite genS; last by rewrite subUset mulG_subl mulG_subr.
+by rewrite mulgSS ?(sub_gen, subsetUl, subsetUr).
+Qed.
+
+Lemma mulG_subG G H K : (G * H \subset K) = (G \subset K) && (H \subset K).
+Proof. by rewrite -gen_subG genM_join join_subG. Qed.
+
+Lemma mulGsubP K H G : reflect (K \subset G /\ H \subset G) (K * H \subset G).
+Proof. by rewrite mulG_subG; exact: andP. Qed.
+
+Lemma mulG_sub K H A : K * H = A -> K \subset A /\ H \subset A.
+Proof. by move <-; rewrite mulG_subl mulG_subr. Qed.
+
+Lemma trivMg G H : (G * H == 1) = (G :==: 1) && (H :==: 1).
+Proof.
+by rewrite !eqEsubset -{2}[1]mulGid mulgSS ?sub1G // !andbT mulG_subG.
+Qed.
+
+Lemma comm_joingE G H : commute G H -> G <*> H = G * H.
+Proof.
+by move/comm_group_setP=> gGH; rewrite -genM_join; exact: (genGid (group gGH)).
+Qed.
+
+Lemma joinGC : commutative joinGT.
+Proof. by move=> G H; apply: val_inj; exact: joingC. Qed.
+
+Lemma joinGA : associative joinGT.
+Proof. by move=> G H K; apply: val_inj; exact: joingA. Qed.
+
+Lemma join1G : left_id 1%G joinGT.
+Proof. by move=> G; apply: val_inj; exact: joing1G. Qed.
+
+Lemma joinG1 : right_id 1%G joinGT.
+Proof. by move=> G; apply: val_inj; exact: joingG1. Qed.
+
+Canonical joinG_law := Monoid.Law joinGA join1G joinG1.
+Canonical joinG_abelaw := Monoid.ComLaw joinGC.
+
+Lemma bigprodGEgen I r (P : pred I) (F : I -> {set gT}) :
+ (\prod_(i <- r | P i) <<F i>>)%G :=: << \bigcup_(i <- r | P i) F i >>.
+Proof.
+elim/big_rec2: _ => /= [|i A _ _ ->]; first by rewrite gen0.
+by rewrite joing_idl joing_idr.
+Qed.
+
+Lemma bigprodGE I r (P : pred I) (F : I -> {group gT}) :
+ (\prod_(i <- r | P i) F i)%G :=: << \bigcup_(i <- r | P i) F i >>.
+Proof.
+rewrite -bigprodGEgen /=; apply: congr_group.
+by apply: eq_bigr => i _; rewrite genGidG.
+Qed.
+
+Lemma mem_commg A B x y : x \in A -> y \in B -> [~ x, y] \in [~: A, B].
+Proof. by move=> Ax By; rewrite mem_gen ?mem_imset2. Qed.
+
+Lemma commSg A B C : A \subset B -> [~: A, C] \subset [~: B, C].
+Proof. by move=> sAC; rewrite genS ?imset2S. Qed.
+
+Lemma commgS A B C : B \subset C -> [~: A, B] \subset [~: A, C].
+Proof. by move=> sBC; rewrite genS ?imset2S. Qed.
+
+Lemma commgSS A B C D :
+ A \subset B -> C \subset D -> [~: A, C] \subset [~: B, D].
+Proof. by move=> sAB sCD; rewrite genS ?imset2S. Qed.
+
+Lemma der1_subG G : [~: G, G] \subset G.
+Proof.
+by rewrite gen_subG; apply/subsetP=> _ /imset2P[x y Gx Gy ->]; exact: groupR.
+Qed.
+
+Lemma comm_subG A B G : A \subset G -> B \subset G -> [~: A, B] \subset G.
+Proof.
+by move=> sAG sBG; apply: subset_trans (der1_subG G); exact: commgSS.
+Qed.
+
+Lemma commGC A B : [~: A, B] = [~: B, A].
+Proof.
+rewrite -[[~: A, B]]genV; congr <<_>>; apply/setP=> z; rewrite inE.
+by apply/imset2P/imset2P=> [] [x y Ax Ay]; last rewrite -{1}(invgK z);
+ rewrite -invg_comm => /invg_inj->; exists y x.
+Qed.
+
+Lemma conjsRg A B x : [~: A, B] :^ x = [~: A :^ x, B :^ x].
+Proof.
+wlog suffices: A B x / [~: A, B] :^ x \subset [~: A :^ x, B :^ x].
+ move=> subJ; apply/eqP; rewrite eqEsubset subJ /= -sub_conjgV.
+ by rewrite -{2}(conjsgK x A) -{2}(conjsgK x B).
+rewrite -genJ gen_subG; apply/subsetP=> _ /imsetP[_ /imset2P[y z Ay Bz ->] ->].
+by rewrite conjRg mem_commg ?memJ_conjg.
+Qed.
+
+End GeneratedGroup.
+
+Implicit Arguments gen_prodgP [gT A x].
+Implicit Arguments joing_idPl [gT G A].
+Implicit Arguments joing_idPr [gT A G].
+Implicit Arguments mulGsubP [gT K H G].
+Implicit Arguments joing_subP [gT A B G].
+
+Section Cycles.
+
+(* Elementary properties of cycles and order, needed in perm.v. *)
+(* More advanced results on the structure of cyclic groups will *)
+(* be given in cyclic.v. *)
+
+Variable gT : finGroupType.
+Implicit Types x y : gT.
+Implicit Types G : {group gT}.
+
+Import Monoid.Theory.
+
+Lemma cycle1 : <[1]> = [1 gT].
+Proof. exact: genGid. Qed.
+
+Lemma order1 : #[1 : gT] = 1%N.
+Proof. by rewrite /order cycle1 cards1. Qed.
+
+Lemma cycle_id x : x \in <[x]>.
+Proof. by rewrite mem_gen // set11. Qed.
+
+Lemma mem_cycle x i : x ^+ i \in <[x]>.
+Proof. by rewrite groupX // cycle_id. Qed.
+
+Lemma cycle_subG x G : (<[x]> \subset G) = (x \in G).
+Proof. by rewrite gen_subG sub1set. Qed.
+
+Lemma cycle_eq1 x : (<[x]> == 1) = (x == 1).
+Proof. by rewrite eqEsubset sub1G andbT cycle_subG inE. Qed.
+
+Lemma orderE x : #[x] = #|<[x]>|. Proof. by []. Qed.
+
+Lemma order_eq1 x : (#[x] == 1%N) = (x == 1).
+Proof. by rewrite -trivg_card1 cycle_eq1. Qed.
+
+Lemma order_gt1 x : (#[x] > 1) = (x != 1).
+Proof. by rewrite ltnNge -trivg_card_le1 cycle_eq1. Qed.
+
+Lemma cycle_traject x : <[x]> =i traject (mulg x) 1 #[x].
+Proof.
+set t := _ 1; apply: fsym; apply/subset_cardP; last first.
+ by apply/subsetP=> _ /trajectP[i _ ->]; rewrite -iteropE mem_cycle.
+rewrite (card_uniqP _) ?size_traject //; case def_n: #[_] => // [n].
+rewrite looping_uniq; apply: contraL (card_size (t n)) => /loopingP t_xi.
+rewrite -ltnNge size_traject -def_n ?subset_leq_card //.
+rewrite -(eq_subset_r (in_set _)) {}/t; set G := finset _.
+rewrite -[x]mulg1 -[G]gen_set_id ?genS ?sub1set ?inE ?(t_xi 1%N)//.
+apply/group_setP; split=> [|y z]; rewrite !inE ?(t_xi 0) //.
+by do 2!case/trajectP=> ? _ ->; rewrite -!iteropE -expgD [x ^+ _]iteropE.
+Qed.
+
+Lemma cycle2g x : #[x] = 2 -> <[x]> = [set 1; x].
+Proof. by move=> ox; apply/setP=> y; rewrite cycle_traject ox !inE mulg1. Qed.
+
+Lemma cyclePmin x y : y \in <[x]> -> {i | i < #[x] & y = x ^+ i}.
+Proof.
+rewrite cycle_traject; set tx := traject _ _ #[x] => tx_y; pose i := index y tx.
+have lt_i_x : i < #[x] by rewrite -index_mem size_traject in tx_y.
+by exists i; rewrite // [x ^+ i]iteropE /= -(nth_traject _ lt_i_x) nth_index.
+Qed.
+
+Lemma cycleP x y : reflect (exists i, y = x ^+ i) (y \in <[x]>).
+Proof.
+by apply: (iffP idP) => [/cyclePmin[i _]|[i ->]]; [exists i | exact: mem_cycle].
+Qed.
+
+Lemma expg_order x : x ^+ #[x] = 1.
+Proof.
+have: uniq (traject (mulg x) 1 #[x]).
+ by apply/card_uniqP; rewrite size_traject -(eq_card (cycle_traject x)).
+case/cyclePmin: (mem_cycle x #[x]) => [] [//|i] ltix.
+rewrite -(subnKC ltix) addSnnS /= expgD; move: (_ - _) => j x_j1.
+case/andP=> /trajectP[]; exists j; first exact: leq_addl.
+by apply: (mulgI (x ^+ i.+1)); rewrite -iterSr iterS -iteropE -expgS mulg1.
+Qed.
+
+Lemma expg_mod p k x : x ^+ p = 1 -> x ^+ (k %% p) = x ^+ k.
+Proof.
+move=> xp.
+by rewrite {2}(divn_eq k p) expgD mulnC expgM xp expg1n mul1g.
+Qed.
+
+Lemma expg_mod_order x i : x ^+ (i %% #[x]) = x ^+ i.
+Proof. by rewrite expg_mod // expg_order. Qed.
+
+Lemma invg_expg x : x^-1 = x ^+ #[x].-1.
+Proof. by apply/eqP; rewrite eq_invg_mul -expgS prednK ?expg_order. Qed.
+
+Lemma invg2id x : #[x] = 2 -> x^-1 = x.
+Proof. by move=> ox; rewrite invg_expg ox. Qed.
+
+Lemma cycleX x i : <[x ^+ i]> \subset <[x]>.
+Proof. rewrite cycle_subG; exact: mem_cycle. Qed.
+
+Lemma cycleV x : <[x^-1]> = <[x]>.
+Proof.
+by apply/eqP; rewrite eq_sym eqEsubset !cycle_subG groupV -groupV !cycle_id.
+Qed.
+
+Lemma orderV x : #[x^-1] = #[x].
+Proof. by rewrite /order cycleV. Qed.
+
+Lemma cycleJ x y : <[x ^ y]> = <[x]> :^ y.
+Proof. by rewrite -genJ conjg_set1. Qed.
+
+Lemma orderJ x y : #[x ^ y] = #[x].
+Proof. by rewrite /order cycleJ cardJg. Qed.
+
+End Cycles.
+
+Section Normaliser.
+
+Variable gT : finGroupType.
+Implicit Types x y z : gT.
+Implicit Types A B C D : {set gT}.
+Implicit Type G H K : {group gT}.
+
+Lemma normP x A : reflect (A :^ x = A) (x \in 'N(A)).
+Proof.
+suffices ->: (x \in 'N(A)) = (A :^ x == A) by exact: eqP.
+by rewrite eqEcard cardJg leqnn andbT inE.
+Qed.
+Implicit Arguments normP [x A].
+
+Lemma group_set_normaliser A : group_set 'N(A).
+Proof.
+apply/group_setP; split=> [|x y Nx Ny]; rewrite inE ?conjsg1 //.
+by rewrite conjsgM !(normP _).
+Qed.
+
+Canonical normaliser_group A := group (group_set_normaliser A).
+
+Lemma normsP A B : reflect {in A, normalised B} (A \subset 'N(B)).
+Proof.
+apply: (iffP subsetP) => nBA x Ax; last by rewrite inE nBA //.
+by apply/normP; exact: nBA.
+Qed.
+Implicit Arguments normsP [A B].
+
+Lemma memJ_norm x y A : x \in 'N(A) -> (y ^ x \in A) = (y \in A).
+Proof. by move=> Nx; rewrite -{1}(normP Nx) memJ_conjg. Qed.
+
+Lemma norms_cycle x y : (<[y]> \subset 'N(<[x]>)) = (x ^ y \in <[x]>).
+Proof. by rewrite cycle_subG inE -cycleJ cycle_subG. Qed.
+
+Lemma norm1 : 'N(1) = setT :> {set gT}.
+Proof. by apply/setP=> x; rewrite !inE conjs1g subxx. Qed.
+
+Lemma norms1 A : A \subset 'N(1).
+Proof. by rewrite norm1 subsetT. Qed.
+
+Lemma normCs A : 'N(~: A) = 'N(A).
+Proof. by apply/setP=> x; rewrite -groupV !inE conjCg setCS sub_conjg. Qed.
+
+Lemma normG G : G \subset 'N(G).
+Proof. by apply/normsP; exact: conjGid. Qed.
+
+Lemma normT : 'N([set: gT]) = [set: gT].
+Proof. by apply/eqP; rewrite -subTset normG. Qed.
+
+Lemma normsG A G : A \subset G -> A \subset 'N(G).
+Proof. move=> sAG; exact: subset_trans (normG G). Qed.
+
+Lemma normC A B : A \subset 'N(B) -> commute A B.
+Proof.
+move/subsetP=> nBA; apply/setP=> u.
+apply/mulsgP/mulsgP=> [[x y Ax By] | [y x By Ax]] -> {u}.
+ by exists (y ^ x^-1) x; rewrite -?conjgCV // memJ_norm // groupV nBA.
+by exists x (y ^ x); rewrite -?conjgC // memJ_norm // nBA.
+Qed.
+
+Lemma norm_joinEl G H : G \subset 'N(H) -> G <*> H = G * H.
+Proof. by move/normC/comm_joingE. Qed.
+
+Lemma norm_joinEr G H : H \subset 'N(G) -> G <*> H = G * H.
+Proof. by move/normC=> cHG; exact: comm_joingE. Qed.
+
+Lemma norm_rlcoset G x : x \in 'N(G) -> G :* x = x *: G.
+Proof. by rewrite -sub1set => /normC. Qed.
+
+Lemma rcoset_mul G x y : x \in 'N(G) -> (G :* x) * (G :* y) = G :* (x * y).
+Proof.
+move/norm_rlcoset=> GxxG.
+by rewrite mulgA -(mulgA _ _ G) -GxxG mulgA mulGid -mulgA mulg_set1.
+Qed.
+
+Lemma normJ A x : 'N(A :^ x) = 'N(A) :^ x.
+Proof.
+by apply/setP=> y; rewrite mem_conjg !inE -conjsgM conjgCV conjsgM conjSg.
+Qed.
+
+Lemma norm_conj_norm x A B :
+ x \in 'N(A) -> (A \subset 'N(B :^ x)) = (A \subset 'N(B)).
+Proof. by move=> Nx; rewrite normJ -sub_conjgV (normP _) ?groupV. Qed.
+
+Lemma norm_gen A : 'N(A) \subset 'N(<<A>>).
+Proof. by apply/normsP=> x Nx; rewrite -genJ (normP Nx). Qed.
+
+Lemma class_norm x G : G \subset 'N(x ^: G).
+Proof. by apply/normsP=> y; exact: classGidr. Qed.
+
+Lemma class_normal x G : x \in G -> x ^: G <| G.
+Proof. by move=> Gx; rewrite /normal class_norm class_subG. Qed.
+
+Lemma class_sub_norm G A x : G \subset 'N(A) -> (x ^: G \subset A) = (x \in A).
+Proof.
+move=> nAG; apply/subsetP/idP=> [-> // | Ax xy]; first exact: class_refl.
+by case/imsetP=> y Gy ->; rewrite memJ_norm ?(subsetP nAG).
+Qed.
+
+Lemma class_support_norm A G : G \subset 'N(class_support A G).
+Proof. by apply/normsP; exact: class_supportGidr. Qed.
+
+Lemma class_support_sub_norm A B G :
+ A \subset G -> B \subset 'N(G) -> class_support A B \subset G.
+Proof.
+move=> sAG nGB; rewrite class_supportEr.
+by apply/bigcupsP=> x Bx; rewrite -(normsP nGB x Bx) conjSg.
+Qed.
+
+Section norm_trans.
+
+Variables (A B C D : {set gT}).
+Hypotheses (nBA : A \subset 'N(B)) (nCA : A \subset 'N(C)).
+
+Lemma norms_gen : A \subset 'N(<<B>>).
+Proof. exact: subset_trans nBA (norm_gen B). Qed.
+
+Lemma norms_norm : A \subset 'N('N(B)).
+Proof. by apply/normsP=> x Ax; rewrite -normJ (normsP nBA). Qed.
+
+Lemma normsI : A \subset 'N(B :&: C).
+Proof. by apply/normsP=> x Ax; rewrite conjIg !(normsP _ x Ax). Qed.
+
+Lemma normsU : A \subset 'N(B :|: C).
+Proof. by apply/normsP=> x Ax; rewrite conjUg !(normsP _ x Ax). Qed.
+
+Lemma normsIs : B \subset 'N(D) -> A :&: B \subset 'N(C :&: D).
+Proof.
+move/normsP=> nDB; apply/normsP=> x; case/setIP=> Ax Bx.
+by rewrite conjIg (normsP nCA) ?nDB.
+Qed.
+
+Lemma normsD : A \subset 'N(B :\: C).
+Proof. by apply/normsP=> x Ax; rewrite conjDg !(normsP _ x Ax). Qed.
+
+Lemma normsM : A \subset 'N(B * C).
+Proof. by apply/normsP=> x Ax; rewrite conjsMg !(normsP _ x Ax). Qed.
+
+Lemma normsY : A \subset 'N(B <*> C).
+Proof. by apply/normsP=> x Ax; rewrite -genJ conjUg !(normsP _ x Ax). Qed.
+
+Lemma normsR : A \subset 'N([~: B, C]).
+Proof. by apply/normsP=> x Ax; rewrite conjsRg !(normsP _ x Ax). Qed.
+
+Lemma norms_class_support : A \subset 'N(class_support B C).
+Proof.
+apply/subsetP=> x Ax; rewrite inE sub_conjg class_supportEr.
+apply/bigcupsP=> y Cy; rewrite -sub_conjg -conjsgM conjgC conjsgM.
+by rewrite (normsP nBA) // bigcup_sup ?memJ_norm ?(subsetP nCA).
+Qed.
+
+End norm_trans.
+
+Lemma normsIG A B G : A \subset 'N(B) -> A :&: G \subset 'N(B :&: G).
+Proof. by move/normsIs->; rewrite ?normG. Qed.
+
+Lemma normsGI A B G : A \subset 'N(B) -> G :&: A \subset 'N(G :&: B).
+Proof. by move=> nBA; rewrite !(setIC G) normsIG. Qed.
+
+Lemma norms_bigcap I r (P : pred I) A (B_ : I -> {set gT}) :
+ A \subset \bigcap_(i <- r | P i) 'N(B_ i) ->
+ A \subset 'N(\bigcap_(i <- r | P i) B_ i).
+Proof.
+elim/big_rec2: _ => [|i B N _ IH /subsetIP[nBiA /IH]]; last exact: normsI.
+by rewrite normT.
+Qed.
+
+Lemma norms_bigcup I r (P : pred I) A (B_ : I -> {set gT}) :
+ A \subset \bigcap_(i <- r | P i) 'N(B_ i) ->
+ A \subset 'N(\bigcup_(i <- r | P i) B_ i).
+Proof.
+move=> nBA; rewrite -normCs setC_bigcup norms_bigcap //.
+by rewrite (eq_bigr _ (fun _ _ => normCs _)).
+Qed.
+
+Lemma normsD1 A B : A \subset 'N(B) -> A \subset 'N(B^#).
+Proof. by move/normsD->; rewrite ?norms1. Qed.
+
+Lemma normD1 A : 'N(A^#) = 'N(A).
+Proof.
+apply/eqP; rewrite eqEsubset normsD1 //.
+rewrite -{2}(setID A 1) setIC normsU //; apply/normsP=> x _; apply/setP=> y.
+by rewrite conjIg conjs1g !inE mem_conjg; case: eqP => // ->; rewrite conj1g.
+Qed.
+
+Lemma normalP A B : reflect (A \subset B /\ {in B, normalised A}) (A <| B).
+Proof. by apply: (iffP andP)=> [] [sAB]; move/normsP. Qed.
+
+Lemma normal_sub A B : A <| B -> A \subset B.
+Proof. by case/andP. Qed.
+
+Lemma normal_norm A B : A <| B -> B \subset 'N(A).
+Proof. by case/andP. Qed.
+
+Lemma normalS G H K : K \subset H -> H \subset G -> K <| G -> K <| H.
+Proof.
+by move=> sKH sHG /andP[_ nKG]; rewrite /(K <| _) sKH (subset_trans sHG).
+Qed.
+
+Lemma normal1 G : 1 <| G.
+Proof. by rewrite /normal sub1set group1 norms1. Qed.
+
+Lemma normal_refl G : G <| G.
+Proof. by rewrite /(G <| _) normG subxx. Qed.
+
+Lemma normalG G : G <| 'N(G).
+Proof. by rewrite /(G <| _) normG subxx. Qed.
+
+Lemma normalSG G H : H \subset G -> H <| 'N_G(H).
+Proof. by move=> sHG; rewrite /normal subsetI sHG normG subsetIr. Qed.
+
+Lemma normalJ A B x : (A :^ x <| B :^ x) = (A <| B).
+Proof. by rewrite /normal normJ !conjSg. Qed.
+
+Lemma normalM G A B : A <| G -> B <| G -> A * B <| G.
+Proof.
+by case/andP=> sAG nAG /andP[sBG nBG]; rewrite /normal mul_subG ?normsM.
+Qed.
+
+Lemma normalY G A B : A <| G -> B <| G -> A <*> B <| G.
+Proof.
+by case/andP=> sAG ? /andP[sBG ?]; rewrite /normal join_subG sAG sBG ?normsY.
+Qed.
+
+Lemma normalYl G H : (H <| H <*> G) = (G \subset 'N(H)).
+Proof. by rewrite /normal joing_subl join_subG normG. Qed.
+
+Lemma normalYr G H : (H <| G <*> H) = (G \subset 'N(H)).
+Proof. by rewrite joingC normalYl. Qed.
+
+Lemma normalI G A B : A <| G -> B <| G -> A :&: B <| G.
+Proof.
+by case/andP=> sAG nAG /andP[_ nBG]; rewrite /normal subIset ?sAG // normsI.
+Qed.
+
+Lemma norm_normalI G A : G \subset 'N(A) -> G :&: A <| G.
+Proof. by move=> nAG; rewrite /normal subsetIl normsI ?normG. Qed.
+
+Lemma normalGI G H A : H \subset G -> A <| G -> H :&: A <| H.
+Proof.
+by move=> sHG /andP[_ nAG]; exact: norm_normalI (subset_trans sHG nAG).
+Qed.
+
+Lemma normal_subnorm G H : (H <| 'N_G(H)) = (H \subset G).
+Proof. by rewrite /normal subsetIr subsetI normG !andbT. Qed.
+
+Lemma normalD1 A G : (A^# <| G) = (A <| G).
+Proof. by rewrite /normal normD1 subDset (setUidPr (sub1G G)). Qed.
+
+Lemma gcore_sub A G : gcore A G \subset A.
+Proof. by rewrite (bigcap_min 1) ?conjsg1. Qed.
+
+Lemma gcore_norm A G : G \subset 'N(gcore A G).
+Proof.
+apply/subsetP=> x Gx; rewrite inE; apply/bigcapsP=> y Gy.
+by rewrite sub_conjg -conjsgM bigcap_inf ?groupM ?groupV.
+Qed.
+
+Lemma gcore_normal A G : A \subset G -> gcore A G <| G.
+Proof.
+by move=> sAG; rewrite /normal gcore_norm (subset_trans (gcore_sub A G)).
+Qed.
+
+Lemma gcore_max A B G : B \subset A -> G \subset 'N(B) -> B \subset gcore A G.
+Proof.
+move=> sBA nBG; apply/bigcapsP=> y Gy.
+by rewrite -sub_conjgV (normsP nBG) ?groupV.
+Qed.
+
+Lemma sub_gcore A B G :
+ G \subset 'N(B) -> (B \subset gcore A G) = (B \subset A).
+Proof.
+move=> nBG; apply/idP/idP=> [sBAG | sBA]; last exact: gcore_max.
+exact: subset_trans (gcore_sub A G).
+Qed.
+
+(* An elementary proof that subgroups of index 2 are normal; it is almost as *)
+(* short as the "advanced" proof using group actions; besides, the fact that *)
+(* the coset is equal to the complement is used in extremal.v. *)
+Lemma rcoset_index2 G H x :
+ H \subset G -> #|G : H| = 2 -> x \in G :\: H -> H :* x = G :\: H.
+Proof.
+move=> sHG indexHG => /setDP[Gx notHx]; apply/eqP.
+rewrite eqEcard -(leq_add2l #|G :&: H|) cardsID -(LagrangeI G H) indexHG muln2.
+rewrite (setIidPr sHG) card_rcoset addnn leqnn andbT.
+apply/subsetP=> _ /rcosetP[y Hy ->]; apply/setDP.
+by rewrite !groupMl // (subsetP sHG).
+Qed.
+
+Lemma index2_normal G H : H \subset G -> #|G : H| = 2 -> H <| G.
+Proof.
+move=> sHG indexHG; rewrite /normal sHG; apply/subsetP=> x Gx.
+case Hx: (x \in H); first by rewrite inE conjGid.
+rewrite inE conjsgE mulgA -sub_rcosetV -invg_rcoset.
+by rewrite !(rcoset_index2 sHG) ?inE ?groupV ?Hx // invDg !invGid.
+Qed.
+
+Lemma cent1P x y : reflect (commute x y) (x \in 'C[y]).
+Proof.
+rewrite inE conjg_set1 sub1set inE (sameP eqP conjg_fixP)commg1_sym.
+exact: commgP.
+Qed.
+
+Lemma cent1id x : x \in 'C[x]. Proof. exact/cent1P. Qed.
+
+Lemma cent1E x y : (x \in 'C[y]) = (x * y == y * x).
+Proof. by rewrite (sameP (cent1P x y) eqP). Qed.
+
+Lemma cent1C x y : (x \in 'C[y]) = (y \in 'C[x]).
+Proof. by rewrite !cent1E eq_sym. Qed.
+
+Canonical centraliser_group A : {group _} := Eval hnf in [group of 'C(A)].
+
+Lemma cent_set1 x : 'C([set x]) = 'C[x].
+Proof. by apply: big_pred1 => y /=; rewrite inE. Qed.
+
+Lemma cent1J x y : 'C[x ^ y] = 'C[x] :^ y.
+Proof. by rewrite -conjg_set1 normJ. Qed.
+
+Lemma centP A x : reflect (centralises x A) (x \in 'C(A)).
+Proof. by apply: (iffP bigcapP) => cxA y /cxA/cent1P. Qed.
+
+Lemma centsP A B : reflect {in A, centralised B} (A \subset 'C(B)).
+Proof. by apply: (iffP subsetP) => cAB x /cAB/centP. Qed.
+
+Lemma centsC A B : (A \subset 'C(B)) = (B \subset 'C(A)).
+Proof. by apply/centsP/centsP=> cAB x ? y ?; rewrite /commute -cAB. Qed.
+
+Lemma cents1 A : A \subset 'C(1).
+Proof. by rewrite centsC sub1G. Qed.
+
+Lemma cent1T : 'C(1) = setT :> {set gT}.
+Proof. by apply/eqP; rewrite -subTset cents1. Qed.
+
+Lemma cent11T : 'C[1] = setT :> {set gT}.
+Proof. by rewrite -cent_set1 cent1T. Qed.
+
+Lemma cent_sub A : 'C(A) \subset 'N(A).
+Proof.
+apply/subsetP=> x /centP cAx; rewrite inE.
+by apply/subsetP=> _ /imsetP[y Ay ->]; rewrite /conjg -cAx ?mulKg.
+Qed.
+
+Lemma cents_norm A B : A \subset 'C(B) -> A \subset 'N(B).
+Proof. by move=> cAB; exact: subset_trans (cent_sub B). Qed.
+
+Lemma centC A B : A \subset 'C(B) -> commute A B.
+Proof. by move=> cAB; exact: normC (cents_norm cAB). Qed.
+
+Lemma cent_joinEl G H : G \subset 'C(H) -> G <*> H = G * H.
+Proof. by move=> cGH; exact: norm_joinEl (cents_norm cGH). Qed.
+
+Lemma cent_joinEr G H : H \subset 'C(G) -> G <*> H = G * H.
+Proof. by move=> cGH; exact: norm_joinEr (cents_norm cGH). Qed.
+
+Lemma centJ A x : 'C(A :^ x) = 'C(A) :^ x.
+Proof.
+apply/setP=> y; rewrite mem_conjg; apply/centP/centP=> cAy z Az.
+ by apply: (conjg_inj x); rewrite 2!conjMg conjgKV cAy ?memJ_conjg.
+by apply: (conjg_inj x^-1); rewrite 2!conjMg cAy -?mem_conjg.
+Qed.
+
+Lemma cent_norm A : 'N(A) \subset 'N('C(A)).
+Proof. by apply/normsP=> x nCx; rewrite -centJ (normP nCx). Qed.
+
+Lemma norms_cent A B : A \subset 'N(B) -> A \subset 'N('C(B)).
+Proof. move=> nBA; exact: subset_trans nBA (cent_norm B). Qed.
+
+Lemma cent_normal A : 'C(A) <| 'N(A).
+Proof. by rewrite /(_ <| _) cent_sub cent_norm. Qed.
+
+Lemma centS A B : B \subset A -> 'C(A) \subset 'C(B).
+Proof. by move=> sAB; rewrite centsC (subset_trans sAB) 1?centsC. Qed.
+
+Lemma centsS A B C : A \subset B -> C \subset 'C(B) -> C \subset 'C(A).
+Proof. by move=> sAB cCB; exact: subset_trans cCB (centS sAB). Qed.
+
+Lemma centSS A B C D :
+ A \subset C -> B \subset D -> C \subset 'C(D) -> A \subset 'C(B).
+Proof. move=> sAC sBD cCD; exact: subset_trans (centsS sBD cCD). Qed.
+
+Lemma centI A B : 'C(A) <*> 'C(B) \subset 'C(A :&: B).
+Proof. by rewrite gen_subG subUset !centS ?(subsetIl, subsetIr). Qed.
+
+Lemma centU A B : 'C(A :|: B) = 'C(A) :&: 'C(B).
+Proof.
+apply/eqP; rewrite eqEsubset subsetI 2?centS ?(subsetUl, subsetUr) //=.
+by rewrite centsC subUset -centsC subsetIl -centsC subsetIr.
+Qed.
+
+Lemma cent_gen A : 'C(<<A>>) = 'C(A).
+Proof. by apply/setP=> x; rewrite -!sub1set centsC gen_subG centsC. Qed.
+
+Lemma cent_cycle x : 'C(<[x]>) = 'C[x].
+Proof. by rewrite cent_gen cent_set1. Qed.
+
+Lemma sub_cent1 A x : (A \subset 'C[x]) = (x \in 'C(A)).
+Proof. by rewrite -cent_cycle centsC cycle_subG. Qed.
+
+Lemma cents_cycle x y : commute x y -> <[x]> \subset 'C(<[y]>).
+Proof. move=> cxy; rewrite cent_cycle cycle_subG; exact/cent1P. Qed.
+
+Lemma cycle_abelian x : abelian <[x]>.
+Proof. exact: cents_cycle. Qed.
+
+Lemma centY A B : 'C(A <*> B) = 'C(A) :&: 'C(B).
+Proof. by rewrite cent_gen centU. Qed.
+
+Lemma centM G H : 'C(G * H) = 'C(G) :&: 'C(H).
+Proof. by rewrite -cent_gen genM_join centY. Qed.
+
+Lemma cent_classP x G : reflect (x ^: G = [set x]) (x \in 'C(G)).
+Proof.
+apply: (iffP (centP _ _)) => [Cx | Cx1 y Gy].
+ apply/eqP; rewrite eqEsubset sub1set class_refl andbT.
+ by apply/subsetP=> _ /imsetP[y Gy ->]; rewrite inE conjgE Cx ?mulKg.
+by apply/commgP/conjg_fixP/set1P; rewrite -Cx1; apply/imsetP; exists y.
+Qed.
+
+Lemma commG1P A B : reflect ([~: A, B] = 1) (A \subset 'C(B)).
+Proof.
+apply: (iffP (centsP A B)) => [cAB | cAB1 x Ax y By].
+ apply/trivgP; rewrite gen_subG; apply/subsetP=> _ /imset2P[x y Ax Ay ->].
+ by rewrite inE; apply/commgP; exact: cAB.
+by apply/commgP; rewrite -in_set1 -[[set 1]]cAB1 mem_commg.
+Qed.
+
+Lemma abelianE A : abelian A = (A \subset 'C(A)). Proof. by []. Qed.
+
+Lemma abelian1 : abelian [1 gT]. Proof. exact: sub1G. Qed.
+
+Lemma abelianS A B : A \subset B -> abelian B -> abelian A.
+Proof. by move=> sAB; exact: centSS. Qed.
+
+Lemma abelianJ A x : abelian (A :^ x) = abelian A.
+Proof. by rewrite /abelian centJ conjSg. Qed.
+
+Lemma abelian_gen A : abelian <<A>> = abelian A.
+Proof. by rewrite /abelian cent_gen gen_subG. Qed.
+
+Lemma abelianY A B :
+ abelian (A <*> B) = [&& abelian A, abelian B & B \subset 'C(A)].
+Proof.
+rewrite /abelian join_subG /= centY !subsetI -!andbA; congr (_ && _).
+by rewrite centsC andbA andbb andbC.
+Qed.
+
+Lemma abelianM G H :
+ abelian (G * H) = [&& abelian G, abelian H & H \subset 'C(G)].
+Proof. by rewrite -abelian_gen genM_join abelianY. Qed.
+
+Section SubAbelian.
+
+Variable A B C : {set gT}.
+Hypothesis cAA : abelian A.
+
+Lemma sub_abelian_cent : C \subset A -> A \subset 'C(C).
+Proof. by move=> sCA; rewrite centsC (subset_trans sCA). Qed.
+
+Lemma sub_abelian_cent2 : B \subset A -> C \subset A -> B \subset 'C(C).
+Proof. by move=> sBA; move/sub_abelian_cent; exact: subset_trans. Qed.
+
+Lemma sub_abelian_norm : C \subset A -> A \subset 'N(C).
+Proof. by move=> sCA; rewrite cents_norm ?sub_abelian_cent. Qed.
+
+Lemma sub_abelian_normal : (C \subset A) = (C <| A).
+Proof.
+by rewrite /normal; case sHG: (C \subset A); rewrite // sub_abelian_norm.
+Qed.
+
+End SubAbelian.
+
+End Normaliser.
+
+Implicit Arguments normP [gT x A].
+Implicit Arguments centP [gT x A].
+Implicit Arguments normsP [gT A B].
+Implicit Arguments cent1P [gT x y].
+Implicit Arguments normalP [gT A B].
+Implicit Arguments centsP [gT A B].
+Implicit Arguments commG1P [gT A B].
+
+Prenex Implicits normP normsP cent1P normalP centP centsP commG1P.
+
+Arguments Scope normaliser_group [_ group_scope].
+Arguments Scope centraliser_group [_ group_scope].
+
+Notation "''N' ( A )" := (normaliser_group A) : Group_scope.
+Notation "''C' ( A )" := (centraliser_group A) : Group_scope.
+Notation "''C' [ x ]" := (normaliser_group [set x%g]) : Group_scope.
+Notation "''N_' G ( A )" := (setI_group G 'N(A)) : Group_scope.
+Notation "''C_' G ( A )" := (setI_group G 'C(A)) : Group_scope.
+Notation "''C_' ( G ) ( A )" := (setI_group G 'C(A))
+ (only parsing) : Group_scope.
+Notation "''C_' G [ x ]" := (setI_group G 'C[x]) : Group_scope.
+Notation "''C_' ( G ) [ x ]" := (setI_group G 'C[x])
+ (only parsing) : Group_scope.
+
+Hint Resolve normG normal_refl.
+
+Section MinMaxGroup.
+
+Variable gT : finGroupType.
+Variable gP : pred {group gT}.
+Arguments Scope gP [Group_scope].
+
+Definition maxgroup := maxset (fun A => group_set A && gP <<A>>).
+Definition mingroup := minset (fun A => group_set A && gP <<A>>).
+
+Lemma ex_maxgroup : (exists G, gP G) -> {G : {group gT} | maxgroup G}.
+Proof.
+move=> exP; have [A maxA]: {A | maxgroup A}.
+ apply: ex_maxset; case: exP => G gPG.
+ by exists (G : {set gT}); rewrite groupP genGidG.
+by exists <<A>>%G; rewrite /= gen_set_id; case/andP: (maxsetp maxA).
+Qed.
+
+Lemma ex_mingroup : (exists G, gP G) -> {G : {group gT} | mingroup G}.
+Proof.
+move=> exP; have [A minA]: {A | mingroup A}.
+ apply: ex_minset; case: exP => G gPG.
+ by exists (G : {set gT}); rewrite groupP genGidG.
+by exists <<A>>%G; rewrite /= gen_set_id; case/andP: (minsetp minA).
+Qed.
+
+Variable G : {group gT}.
+
+Lemma mingroupP :
+ reflect (gP G /\ forall H, gP H -> H \subset G -> H :=: G) (mingroup G).
+Proof.
+apply: (iffP minsetP); rewrite /= groupP genGidG /= => [] [-> minG].
+ by split=> // H gPH sGH; apply: minG; rewrite // groupP genGidG.
+split=> // A; case/andP=> gA gPA; rewrite -(gen_set_id gA); exact: minG.
+Qed.
+
+Lemma maxgroupP :
+ reflect (gP G /\ forall H, gP H -> G \subset H -> H :=: G) (maxgroup G).
+Proof.
+apply: (iffP maxsetP); rewrite /= groupP genGidG /= => [] [-> maxG].
+ by split=> // H gPH sGH; apply: maxG; rewrite // groupP genGidG.
+split=> // A; case/andP=> gA gPA; rewrite -(gen_set_id gA); exact: maxG.
+Qed.
+
+Lemma maxgroupp : maxgroup G -> gP G. Proof. by case/maxgroupP. Qed.
+
+Lemma mingroupp : mingroup G -> gP G. Proof. by case/mingroupP. Qed.
+
+Hypothesis gPG : gP G.
+
+Lemma maxgroup_exists : {H : {group gT} | maxgroup H & G \subset H}.
+Proof.
+have [A maxA sGA]: {A | maxgroup A & G \subset A}.
+ by apply: maxset_exists; rewrite groupP genGidG.
+by exists <<A>>%G; rewrite /= gen_set_id; case/andP: (maxsetp maxA).
+Qed.
+
+Lemma mingroup_exists : {H : {group gT} | mingroup H & H \subset G}.
+Proof.
+have [A maxA sGA]: {A | mingroup A & A \subset G}.
+ by apply: minset_exists; rewrite groupP genGidG.
+by exists <<A>>%G; rewrite /= gen_set_id; case/andP: (minsetp maxA).
+Qed.
+
+End MinMaxGroup.
+
+Notation "[ 'max' A 'of' G | gP ]" :=
+ (maxgroup (fun G : {group _} => gP) A) : group_scope.
+Notation "[ 'max' G | gP ]" := [max gval G of G | gP] : group_scope.
+Notation "[ 'max' A 'of' G | gP & gQ ]" :=
+ [max A of G | gP && gQ] : group_scope.
+Notation "[ 'max' G | gP & gQ ]" := [max G | gP && gQ] : group_scope.
+Notation "[ 'min' A 'of' G | gP ]" :=
+ (mingroup (fun G : {group _} => gP) A) : group_scope.
+Notation "[ 'min' G | gP ]" := [min gval G of G | gP] : group_scope.
+Notation "[ 'min' A 'of' G | gP & gQ ]" :=
+ [min A of G | gP && gQ] : group_scope.
+Notation "[ 'min' G | gP & gQ ]" := [min G | gP && gQ] : group_scope.
+
+Implicit Arguments mingroupP [gT gP G].
+Implicit Arguments maxgroupP [gT gP G].
+Prenex Implicits mingroupP maxgroupP.
diff --git a/mathcomp/fingroup/gproduct.v b/mathcomp/fingroup/gproduct.v
new file mode 100644
index 0000000..924779f
--- /dev/null
+++ b/mathcomp/fingroup/gproduct.v
@@ -0,0 +1,1703 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div choice fintype.
+Require Import bigop finset fingroup morphism quotient action.
+
+(******************************************************************************)
+(* Partial, semidirect, central, and direct products. *)
+(* ++ Internal products, with A, B : {set gT}, are partial operations : *)
+(* partial_product A B == A * B if A is a group normalised by the group B, *)
+(* and the empty set otherwise. *)
+(* A ><| B == A * B if this is a semi-direct product (i.e., if A *)
+(* is normalised by B and intersects it trivially). *)
+(* A \* B == A * B if this is a central product ([A, B] = 1). *)
+(* A \x B == A * B if this is a direct product. *)
+(* [complements to K in G] == set of groups H s.t. K * H = G and K :&: H = 1. *)
+(* [splits G, over K] == [complements to K in G] is not empty. *)
+(* remgr A B x == the right remainder in B of x mod A, i.e., *)
+(* some element of (A :* x) :&: B. *)
+(* divgr A B x == the "quotient" in B of x by A: for all x, *)
+(* x = divgr A B x * remgr A B x. *)
+(* ++ External products : *)
+(* pairg1, pair1g == the isomorphisms aT1 -> aT1 * aT2, aT2 -> aT1 * aT2. *)
+(* (aT1 * aT2 has a direct product group structure.) *)
+(* sdprod_by to == the semidirect product defined by to : groupAction H K. *)
+(* This is a finGroupType; the actual semidirect product is *)
+(* the total set [set: sdprod_by to] on that type. *)
+(* sdpair[12] to == the isomorphisms injecting K and H into *)
+(* sdprod_by to = sdpair1 to @* K ><| sdpair2 to @* H. *)
+(* External central products (with identified centers) will be defined later *)
+(* in file center.v. *)
+(* ++ Morphisms on product groups: *)
+(* pprodm nAB fJ fAB == the morphism extending fA and fB on A <*> B when *)
+(* nAB : B \subset 'N(A), *)
+(* fJ : {in A & B, morph_actj fA fB}, and *)
+(* fAB : {in A :&: B, fA =1 fB}. *)
+(* sdprodm defG fJ == the morphism extending fA and fB on G, given *)
+(* defG : A ><| B = G and *)
+(* fJ : {in A & B, morph_act 'J 'J fA fB}. *)
+(* xsdprodm fHKact == the total morphism on sdprod_by to induced by *)
+(* fH : {morphism H >-> rT}, fK : {morphism K >-> rT}, *)
+(* with to : groupAction K H, *)
+(* given fHKact : morph_act to 'J fH fK. *)
+(* cprodm defG cAB fAB == the morphism extending fA and fB on G, when *)
+(* defG : A \* B = G, *)
+(* cAB : fB @* B \subset 'C(fB @* A), *)
+(* and fAB : {in A :&: B, fA =1 fB}. *)
+(* dprodm defG cAB == the morphism extending fA and fB on G, when *)
+(* defG : A \x B = G and *)
+(* cAB : fA @* B \subset 'C(fA @* A) *)
+(* mulgm (x, y) == x * y; mulgm is an isomorphism from setX A B to G *)
+(* iff A \x B = G . *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section Defs.
+
+Variables gT : finGroupType.
+Implicit Types A B C : {set gT}.
+
+Definition partial_product A B :=
+ if A == 1 then B else if B == 1 then A else
+ if [&& group_set A, group_set B & B \subset 'N(A)] then A * B else set0.
+
+Definition semidirect_product A B :=
+ if A :&: B \subset 1%G then partial_product A B else set0.
+
+Definition central_product A B :=
+ if B \subset 'C(A) then partial_product A B else set0.
+
+Definition direct_product A B :=
+ if A :&: B \subset 1%G then central_product A B else set0.
+
+Definition complements_to_in A B :=
+ [set K : {group gT} | A :&: K == 1 & A * K == B].
+
+Definition splits_over B A := complements_to_in A B != set0.
+
+(* Product remainder functions -- right variant only. *)
+Definition remgr A B x := repr (A :* x :&: B).
+Definition divgr A B x := x * (remgr A B x)^-1.
+
+End Defs.
+
+Arguments Scope partial_product [_ group_scope group_scope].
+Arguments Scope semidirect_product [_ group_scope group_scope].
+Arguments Scope central_product [_ group_scope group_scope].
+Arguments Scope complements_to_in [_ group_scope group_scope].
+Arguments Scope splits_over [_ group_scope group_scope].
+Arguments Scope remgr [_ group_scope group_scope group_scope].
+Arguments Scope divgr [_ group_scope group_scope group_scope].
+Implicit Arguments partial_product [].
+Implicit Arguments semidirect_product [].
+Implicit Arguments central_product [].
+Implicit Arguments direct_product [].
+Notation pprod := (partial_product _).
+Notation sdprod := (semidirect_product _).
+Notation cprod := (central_product _).
+Notation dprod := (direct_product _).
+
+Notation "G ><| H" := (sdprod G H)%g (at level 40, left associativity).
+Notation "G \* H" := (cprod G H)%g (at level 40, left associativity).
+Notation "G \x H" := (dprod G H)%g (at level 40, left associativity).
+
+Notation "[ 'complements' 'to' A 'in' B ]" := (complements_to_in A B)
+ (at level 0, format "[ 'complements' 'to' A 'in' B ]") : group_scope.
+
+Notation "[ 'splits' B , 'over' A ]" := (splits_over B A)
+ (at level 0, format "[ 'splits' B , 'over' A ]") : group_scope.
+
+(* Prenex Implicits remgl divgl. *)
+Prenex Implicits remgr divgr.
+
+Section InternalProd.
+
+Variable gT : finGroupType.
+Implicit Types A B C : {set gT}.
+Implicit Types G H K L M : {group gT}.
+
+Local Notation pprod := (partial_product gT).
+Local Notation sdprod := (semidirect_product gT) (only parsing).
+Local Notation cprod := (central_product gT) (only parsing).
+Local Notation dprod := (direct_product gT) (only parsing).
+
+Lemma pprod1g : left_id 1 pprod.
+Proof. by move=> A; rewrite /pprod eqxx. Qed.
+
+Lemma pprodg1 : right_id 1 pprod.
+Proof. by move=> A; rewrite /pprod eqxx; case: eqP. Qed.
+
+CoInductive are_groups A B : Prop := AreGroups K H of A = K & B = H.
+
+Lemma group_not0 G : set0 <> G.
+Proof. by move/setP/(_ 1); rewrite inE group1. Qed.
+
+Lemma mulg0 : right_zero (@set0 gT) mulg.
+Proof.
+by move=> A; apply/setP=> x; rewrite inE; apply/imset2P=> [[y z]]; rewrite inE.
+Qed.
+
+Lemma mul0g : left_zero (@set0 gT) mulg.
+Proof.
+by move=> A; apply/setP=> x; rewrite inE; apply/imset2P=> [[y z]]; rewrite inE.
+Qed.
+
+Lemma pprodP A B G :
+ pprod A B = G -> [/\ are_groups A B, A * B = G & B \subset 'N(A)].
+Proof.
+have Gnot0 := @group_not0 G; rewrite /pprod; do 2?case: eqP => [-> ->| _].
+- by rewrite mul1g norms1; split; first exists 1%G G.
+- by rewrite mulg1 sub1G; split; first exists G 1%G.
+by case: and3P => // [[gA gB ->]]; split; first exists (Group gA) (Group gB).
+Qed.
+
+Lemma pprodE K H : H \subset 'N(K) -> pprod K H = K * H.
+Proof.
+move=> nKH; rewrite /pprod nKH !groupP /=.
+by do 2?case: eqP => [-> | _]; rewrite ?mulg1 ?mul1g.
+Qed.
+
+Lemma pprodEY K H : H \subset 'N(K) -> pprod K H = K <*> H.
+Proof. by move=> nKH; rewrite pprodE ?norm_joinEr. Qed.
+
+Lemma pprodW A B G : pprod A B = G -> A * B = G. Proof. by case/pprodP. Qed.
+
+Lemma pprodWC A B G : pprod A B = G -> B * A = G.
+Proof. by case/pprodP=> _ <- /normC. Qed.
+
+Lemma pprodWY A B G : pprod A B = G -> A <*> B = G.
+Proof. by case/pprodP=> [[K H -> ->] <- /norm_joinEr]. Qed.
+
+Lemma pprodJ A B x : pprod A B :^ x = pprod (A :^ x) (B :^ x).
+Proof.
+rewrite /pprod !conjsg_eq1 !group_setJ normJ conjSg -conjsMg.
+by do 3?case: ifP => // _; exact: conj0g.
+Qed.
+
+(* Properties of the remainders *)
+
+Lemma remgrMl K B x y : y \in K -> remgr K B (y * x) = remgr K B x.
+Proof. by move=> Ky; rewrite {1}/remgr rcosetM rcoset_id. Qed.
+
+Lemma remgrP K B x : (remgr K B x \in K :* x :&: B) = (x \in K * B).
+Proof.
+set y := _ x; apply/idP/mulsgP=> [|[g b Kg Bb x_gb]].
+ rewrite inE rcoset_sym mem_rcoset => /andP[Kxy' By].
+ by exists (x * y^-1) y; rewrite ?mulgKV.
+by apply: (mem_repr b); rewrite inE rcoset_sym mem_rcoset x_gb mulgK Kg.
+Qed.
+
+Lemma remgr1 K H x : x \in K -> remgr K H x = 1.
+Proof. by move=> Kx; rewrite /remgr rcoset_id ?repr_group. Qed.
+
+Lemma divgr_eq A B x : x = divgr A B x * remgr A B x.
+Proof. by rewrite mulgKV. Qed.
+
+Lemma divgrMl K B x y : x \in K -> divgr K B (x * y) = x * divgr K B y.
+Proof. by move=> Hx; rewrite /divgr remgrMl ?mulgA. Qed.
+
+Lemma divgr_id K H x : x \in K -> divgr K H x = x.
+Proof. by move=> Kx; rewrite /divgr remgr1 // invg1 mulg1. Qed.
+
+Lemma mem_remgr K B x : x \in K * B -> remgr K B x \in B.
+Proof. by rewrite -remgrP => /setIP[]. Qed.
+
+Lemma mem_divgr K B x : x \in K * B -> divgr K B x \in K.
+Proof. by rewrite -remgrP inE rcoset_sym mem_rcoset => /andP[]. Qed.
+
+Section DisjointRem.
+
+Variables K H : {group gT}.
+
+Hypothesis tiKH : K :&: H = 1.
+
+Lemma remgr_id x : x \in H -> remgr K H x = x.
+Proof.
+move=> Hx; apply/eqP; rewrite eq_mulgV1 (sameP eqP set1gP) -tiKH inE.
+rewrite -mem_rcoset groupMr ?groupV // -in_setI remgrP.
+by apply: subsetP Hx; exact: mulG_subr.
+Qed.
+
+Lemma remgrMid x y : x \in K -> y \in H -> remgr K H (x * y) = y.
+Proof. by move=> Kx Hy; rewrite remgrMl ?remgr_id. Qed.
+
+Lemma divgrMid x y : x \in K -> y \in H -> divgr K H (x * y) = x.
+Proof. by move=> Kx Hy; rewrite /divgr remgrMid ?mulgK. Qed.
+
+End DisjointRem.
+
+(* Intersection of a centraliser with a disjoint product. *)
+
+Lemma subcent_TImulg K H A :
+ K :&: H = 1 -> A \subset 'N(K) :&: 'N(H) -> 'C_K(A) * 'C_H(A) = 'C_(K * H)(A).
+Proof.
+move=> tiKH /subsetIP[nKA nHA]; apply/eqP.
+rewrite group_modl ?subsetIr // eqEsubset setSI ?mulSg ?subsetIl //=.
+apply/subsetP=> _ /setIP[/mulsgP[x y Kx Hy ->] cAxy].
+rewrite inE cAxy mem_mulg // inE Kx /=.
+apply/centP=> z Az; apply/commgP/conjg_fixP.
+move/commgP/conjg_fixP/(congr1 (divgr K H)): (centP cAxy z Az).
+by rewrite conjMg !divgrMid ?memJ_norm // (subsetP nKA, subsetP nHA).
+Qed.
+
+(* Complements, and splitting. *)
+
+Lemma complP H A B :
+ reflect (A :&: H = 1 /\ A * H = B) (H \in [complements to A in B]).
+Proof. by apply: (iffP setIdP); case; split; apply/eqP. Qed.
+
+Lemma splitsP B A :
+ reflect (exists H, H \in [complements to A in B]) [splits B, over A].
+Proof. exact: set0Pn. Qed.
+
+Lemma complgC H K G :
+ (H \in [complements to K in G]) = (K \in [complements to H in G]).
+Proof.
+rewrite !inE setIC; congr (_ && _).
+by apply/eqP/eqP=> defG; rewrite -(comm_group_setP _) // defG groupP.
+Qed.
+
+Section NormalComplement.
+
+Variables K H G : {group gT}.
+
+Hypothesis complH_K : H \in [complements to K in G].
+
+Lemma remgrM : K <| G -> {in G &, {morph remgr K H : x y / x * y}}.
+Proof.
+case/normalP=> _; case/complP: complH_K => tiKH <- nK_KH x y KHx KHy.
+rewrite {1}(divgr_eq K H y) mulgA (conjgCV x) {2}(divgr_eq K H x) -2!mulgA.
+rewrite mulgA remgrMid //; last by rewrite groupMl mem_remgr.
+by rewrite groupMl !(=^~ mem_conjg, nK_KH, mem_divgr).
+Qed.
+
+Lemma divgrM : H \subset 'C(K) -> {in G &, {morph divgr K H : x y / x * y}}.
+Proof.
+move=> cKH; have /complP[_ defG] := complH_K.
+have nsKG: K <| G by rewrite -defG -cent_joinEr // normalYl cents_norm.
+move=> x y Gx Gy; rewrite {1}/divgr remgrM // invMg -!mulgA (mulgA y).
+by congr (_ * _); rewrite -(centsP cKH) ?groupV ?(mem_remgr, mem_divgr, defG).
+Qed.
+
+End NormalComplement.
+
+(* Semi-direct product *)
+
+Lemma sdprod1g : left_id 1 sdprod.
+Proof. by move=> A; rewrite /sdprod subsetIl pprod1g. Qed.
+
+Lemma sdprodg1 : right_id 1 sdprod.
+Proof. by move=> A; rewrite /sdprod subsetIr pprodg1. Qed.
+
+Lemma sdprodP A B G :
+ A ><| B = G -> [/\ are_groups A B, A * B = G, B \subset 'N(A) & A :&: B = 1].
+Proof.
+rewrite /sdprod; case: ifP => [trAB | _ /group_not0[] //].
+case/pprodP=> gAB defG nBA; split=> {defG nBA}//.
+by case: gAB trAB => H K -> -> /trivgP.
+Qed.
+
+Lemma sdprodE K H : H \subset 'N(K) -> K :&: H = 1 -> K ><| H = K * H.
+Proof. by move=> nKH tiKH; rewrite /sdprod tiKH subxx pprodE. Qed.
+
+Lemma sdprodEY K H : H \subset 'N(K) -> K :&: H = 1 -> K ><| H = K <*> H.
+Proof. by move=> nKH tiKH; rewrite sdprodE ?norm_joinEr. Qed.
+
+Lemma sdprodWpp A B G : A ><| B = G -> pprod A B = G.
+Proof. by case/sdprodP=> [[K H -> ->] <- /pprodE]. Qed.
+
+Lemma sdprodW A B G : A ><| B = G -> A * B = G.
+Proof. by move/sdprodWpp/pprodW. Qed.
+
+Lemma sdprodWC A B G : A ><| B = G -> B * A = G.
+Proof. by move/sdprodWpp/pprodWC. Qed.
+
+Lemma sdprodWY A B G : A ><| B = G -> A <*> B = G.
+Proof. by move/sdprodWpp/pprodWY. Qed.
+
+Lemma sdprodJ A B x : (A ><| B) :^ x = A :^ x ><| B :^ x.
+Proof.
+rewrite /sdprod -conjIg sub_conjg conjs1g -pprodJ.
+by case: ifP => _ //; exact: imset0.
+Qed.
+
+Lemma sdprod_context G K H : K ><| H = G ->
+ [/\ K <| G, H \subset G, K * H = G, H \subset 'N(K) & K :&: H = 1].
+Proof.
+case/sdprodP=> _ <- nKH tiKH.
+by rewrite /normal mulG_subl mulG_subr mulG_subG normG.
+Qed.
+
+Lemma sdprod_compl G K H : K ><| H = G -> H \in [complements to K in G].
+Proof. by case/sdprodP=> _ mulKH _ tiKH; exact/complP. Qed.
+
+Lemma sdprod_normal_complP G K H :
+ K <| G -> reflect (K ><| H = G) (K \in [complements to H in G]).
+Proof.
+case/andP=> _ nKG; rewrite complgC.
+apply: (iffP idP); [case/complP=> tiKH mulKH | exact: sdprod_compl].
+by rewrite sdprodE ?(subset_trans _ nKG) // -mulKH mulG_subr.
+Qed.
+
+Lemma sdprod_card G A B : A ><| B = G -> (#|A| * #|B|)%N = #|G|.
+Proof. by case/sdprodP=> [[H K -> ->] <- _ /TI_cardMg]. Qed.
+
+Lemma sdprod_isom G A B :
+ A ><| B = G ->
+ {nAB : B \subset 'N(A) | isom B (G / A) (restrm nAB (coset A))}.
+Proof.
+case/sdprodP=> [[K H -> ->] <- nKH tiKH].
+by exists nKH; rewrite quotientMidl quotient_isom.
+Qed.
+
+Lemma sdprod_isog G A B : A ><| B = G -> B \isog G / A.
+Proof. by case/sdprod_isom=> nAB; apply: isom_isog. Qed.
+
+Lemma sdprod_subr G A B M : A ><| B = G -> M \subset B -> A ><| M = A <*> M.
+Proof.
+case/sdprodP=> [[K H -> ->] _ nKH tiKH] sMH.
+by rewrite sdprodEY ?(subset_trans sMH) //; apply/trivgP; rewrite -tiKH setIS.
+Qed.
+
+Lemma index_sdprod G A B : A ><| B = G -> #|B| = #|G : A|.
+Proof.
+case/sdprodP=> [[K H -> ->] <- _ tiHK].
+by rewrite indexMg -indexgI setIC tiHK indexg1.
+Qed.
+
+Lemma index_sdprodr G A B M :
+ A ><| B = G -> M \subset B -> #|B : M| = #|G : A <*> M|.
+Proof.
+move=> defG; case/sdprodP: defG (defG) => [[K H -> ->] mulKH nKH _] defG sMH.
+rewrite -!divgS //=; last by rewrite -genM_join gen_subG -mulKH mulgS.
+by rewrite -(sdprod_card defG) -(sdprod_card (sdprod_subr defG sMH)) divnMl.
+Qed.
+
+Lemma quotient_sdprodr_isom G A B M :
+ A ><| B = G -> M <| B ->
+ {f : {morphism B / M >-> coset_of (A <*> M)} |
+ isom (B / M) (G / (A <*> M)) f
+ & forall L, L \subset B -> f @* (L / M) = A <*> L / (A <*> M)}.
+Proof.
+move=> defG nsMH; have [defA defB]: A = <<A>>%G /\ B = <<B>>%G.
+ by have [[K1 H1 -> ->] _ _ _] := sdprodP defG; rewrite /= !genGid.
+do [rewrite {}defA {}defB; move: {A}<<A>>%G {B}<<B>>%G => K H] in defG nsMH *.
+have [[nKH /isomP[injKH imKH]] sMH] := (sdprod_isom defG, normal_sub nsMH).
+have [[nsKG sHG mulKH _ _] nKM] := (sdprod_context defG, subset_trans sMH nKH).
+have nsKMG: K <*> M <| G.
+ by rewrite -quotientYK // -mulKH -quotientK ?cosetpre_normal ?quotient_normal.
+have [/= f inj_f im_f] := third_isom (joing_subl K M) nsKG nsKMG.
+rewrite quotientYidl //= -imKH -(restrm_quotientE nKH sMH) in f inj_f im_f.
+have /domP[h [_ ker_h _ im_h]]: 'dom (f \o quotm _ nsMH) = H / M.
+ by rewrite ['dom _]morphpre_quotm injmK.
+have{im_h} im_h L: L \subset H -> h @* (L / M) = K <*> L / (K <*> M).
+ move=> sLH; have [sLG sKKM] := (subset_trans sLH sHG, joing_subl K M).
+ rewrite im_h morphim_comp morphim_quotm [_ @* L]restrm_quotientE ?im_f //.
+ rewrite quotientY ?(normsG sKKM) ?(subset_trans sLG) ?normal_norm //.
+ by rewrite (quotientS1 sKKM) joing1G.
+exists h => //; apply/isomP; split; last by rewrite im_h //= (sdprodWY defG).
+by rewrite ker_h injm_comp ?injm_quotm.
+Qed.
+
+Lemma quotient_sdprodr_isog G A B M :
+ A ><| B = G -> M <| B -> B / M \isog G / (A <*> M).
+Proof.
+move=> defG; case/sdprodP: defG (defG) => [[K H -> ->] _ _ _] => defG nsMH.
+by have [h /isom_isog->] := quotient_sdprodr_isom defG nsMH.
+Qed.
+
+Lemma sdprod_modl A B G H :
+ A ><| B = G -> A \subset H -> A ><| (B :&: H) = G :&: H.
+Proof.
+case/sdprodP=> {A B} [[A B -> ->]] <- nAB tiAB sAH.
+rewrite -group_modl ?sdprodE ?subIset ?nAB //.
+by rewrite setIA tiAB (setIidPl _) ?sub1G.
+Qed.
+
+Lemma sdprod_modr A B G H :
+ A ><| B = G -> B \subset H -> (H :&: A) ><| B = H :&: G.
+Proof.
+case/sdprodP=> {A B}[[A B -> ->]] <- nAB tiAB sAH.
+rewrite -group_modr ?sdprodE ?normsI // ?normsG //.
+by rewrite -setIA tiAB (setIidPr _) ?sub1G.
+Qed.
+
+Lemma subcent_sdprod B C G A :
+ B ><| C = G -> A \subset 'N(B) :&: 'N(C) -> 'C_B(A) ><| 'C_C(A) = 'C_G(A).
+Proof.
+case/sdprodP=> [[H K -> ->] <- nHK tiHK] nHKA {B C G}.
+rewrite sdprodE ?subcent_TImulg ?normsIG //.
+by rewrite -setIIl tiHK (setIidPl (sub1G _)).
+Qed.
+
+Lemma sdprod_recl n G K H K1 :
+ #|G| <= n -> K ><| H = G -> K1 \proper K -> H \subset 'N(K1) ->
+ exists G1 : {group gT}, [/\ #|G1| < n, G1 \subset G & K1 ><| H = G1].
+Proof.
+move=> leGn; case/sdprodP=> _ defG nKH tiKH ltK1K nK1H.
+have tiK1H: K1 :&: H = 1 by apply/trivgP; rewrite -tiKH setSI ?proper_sub.
+exists (K1 <*> H)%G; rewrite /= -defG sdprodE // norm_joinEr //.
+rewrite ?mulSg ?proper_sub ?(leq_trans _ leGn) //=.
+by rewrite -defG ?TI_cardMg // ltn_pmul2r ?proper_card.
+Qed.
+
+Lemma sdprod_recr n G K H H1 :
+ #|G| <= n -> K ><| H = G -> H1 \proper H ->
+ exists G1 : {group gT}, [/\ #|G1| < n, G1 \subset G & K ><| H1 = G1].
+Proof.
+move=> leGn; case/sdprodP=> _ defG nKH tiKH ltH1H.
+have [sH1H _] := andP ltH1H; have nKH1 := subset_trans sH1H nKH.
+have tiKH1: K :&: H1 = 1 by apply/trivgP; rewrite -tiKH setIS.
+exists (K <*> H1)%G; rewrite /= -defG sdprodE // norm_joinEr //.
+rewrite ?mulgS // ?(leq_trans _ leGn) //=.
+by rewrite -defG ?TI_cardMg // ltn_pmul2l ?proper_card.
+Qed.
+
+Lemma mem_sdprod G A B x : A ><| B = G -> x \in G ->
+ exists y, exists z,
+ [/\ y \in A, z \in B, x = y * z &
+ {in A & B, forall u t, x = u * t -> u = y /\ t = z}].
+Proof.
+case/sdprodP=> [[K H -> ->{A B}] <- _ tiKH] /mulsgP[y z Ky Hz ->{x}].
+exists y; exists z; split=> // u t Ku Ht eqyzut.
+move: (congr1 (divgr K H) eqyzut) (congr1 (remgr K H) eqyzut).
+by rewrite !remgrMid // !divgrMid.
+Qed.
+
+(* Central product *)
+
+Lemma cprod1g : left_id 1 cprod.
+Proof. by move=> A; rewrite /cprod cents1 pprod1g. Qed.
+
+Lemma cprodg1 : right_id 1 cprod.
+Proof. by move=> A; rewrite /cprod sub1G pprodg1. Qed.
+
+Lemma cprodP A B G :
+ A \* B = G -> [/\ are_groups A B, A * B = G & B \subset 'C(A)].
+Proof. by rewrite /cprod; case: ifP => [cAB /pprodP[] | _ /group_not0[]]. Qed.
+
+Lemma cprodE G H : H \subset 'C(G) -> G \* H = G * H.
+Proof. by move=> cGH; rewrite /cprod cGH pprodE ?cents_norm. Qed.
+
+Lemma cprodEY G H : H \subset 'C(G) -> G \* H = G <*> H.
+Proof. by move=> cGH; rewrite cprodE ?cent_joinEr. Qed.
+
+Lemma cprodWpp A B G : A \* B = G -> pprod A B = G.
+Proof. by case/cprodP=> [[K H -> ->] <- /cents_norm/pprodE]. Qed.
+
+Lemma cprodW A B G : A \* B = G -> A * B = G.
+Proof. by move/cprodWpp/pprodW. Qed.
+
+Lemma cprodWC A B G : A \* B = G -> B * A = G.
+Proof. by move/cprodWpp/pprodWC. Qed.
+
+Lemma cprodWY A B G : A \* B = G -> A <*> B = G.
+Proof. by move/cprodWpp/pprodWY. Qed.
+
+Lemma cprodJ A B x : (A \* B) :^ x = A :^ x \* B :^ x.
+Proof.
+by rewrite /cprod centJ conjSg -pprodJ; case: ifP => _ //; exact: imset0.
+Qed.
+
+Lemma cprod_normal2 A B G : A \* B = G -> A <| G /\ B <| G.
+Proof.
+case/cprodP=> [[K H -> ->] <- cKH]; rewrite -cent_joinEr //.
+by rewrite normalYl normalYr !cents_norm // centsC.
+Qed.
+
+Lemma bigcprodW I (r : seq I) P F G :
+ \big[cprod/1]_(i <- r | P i) F i = G -> \prod_(i <- r | P i) F i = G.
+Proof.
+elim/big_rec2: _ G => // i A B _ IH G /cprodP[[_ H _ defB] <- _].
+by rewrite (IH H) defB.
+Qed.
+
+Lemma bigcprodWY I (r : seq I) P F G :
+ \big[cprod/1]_(i <- r | P i) F i = G -> << \bigcup_(i <- r | P i) F i >> = G.
+Proof.
+elim/big_rec2: _ G => [|i A B _ IH G]; first by rewrite gen0.
+case /cprodP => [[K H -> defB] <- cKH].
+by rewrite -[<<_>>]joing_idr (IH H) ?cent_joinEr -?defB.
+Qed.
+
+Lemma triv_cprod A B : (A \* B == 1) = (A == 1) && (B == 1).
+Proof.
+case A1: (A == 1); first by rewrite (eqP A1) cprod1g.
+apply/eqP=> /cprodP[[G H defA ->]] /eqP.
+by rewrite defA trivMg -defA A1.
+Qed.
+
+Lemma cprod_ntriv A B : A != 1 -> B != 1 ->
+ A \* B =
+ if [&& group_set A, group_set B & B \subset 'C(A)] then A * B else set0.
+Proof.
+move=> A1 B1; rewrite /cprod; case: ifP => cAB; rewrite ?cAB ?andbF //=.
+by rewrite /pprod -if_neg A1 -if_neg B1 cents_norm.
+Qed.
+
+Lemma trivg0 : (@set0 gT == 1) = false.
+Proof. by rewrite eqEcard cards0 cards1 andbF. Qed.
+
+Lemma group0 : group_set (@set0 gT) = false.
+Proof. by rewrite /group_set inE. Qed.
+
+Lemma cprod0g A : set0 \* A = set0.
+Proof. by rewrite /cprod centsC sub0set /pprod group0 trivg0 !if_same. Qed.
+
+Lemma cprodC : commutative cprod.
+Proof.
+rewrite /cprod => A B; case: ifP => cAB; rewrite centsC cAB // /pprod.
+by rewrite andbCA normC !cents_norm // 1?centsC //; do 2!case: eqP => // ->.
+Qed.
+
+Lemma cprodA : associative cprod.
+Proof.
+move=> A B C; case A1: (A == 1); first by rewrite (eqP A1) !cprod1g.
+case B1: (B == 1); first by rewrite (eqP B1) cprod1g cprodg1.
+case C1: (C == 1); first by rewrite (eqP C1) !cprodg1.
+rewrite !(triv_cprod, cprod_ntriv) ?{}A1 ?{}B1 ?{}C1 //.
+case: isgroupP => [[G ->{A}] | _]; last by rewrite group0.
+case: (isgroupP B) => [[H ->{B}] | _]; last by rewrite group0.
+case: (isgroupP C) => [[K ->{C}] | _]; last by rewrite group0 !andbF.
+case cGH: (H \subset 'C(G)); case cHK: (K \subset 'C(H)); last first.
+- by rewrite group0.
+- by rewrite group0 /= mulG_subG cGH andbF.
+- by rewrite group0 /= centM subsetI cHK !andbF.
+rewrite /= mulgA mulG_subG centM subsetI cGH cHK andbT -(cent_joinEr cHK).
+by rewrite -(cent_joinEr cGH) !groupP.
+Qed.
+
+Canonical cprod_law := Monoid.Law cprodA cprod1g cprodg1.
+Canonical cprod_abelaw := Monoid.ComLaw cprodC.
+
+Lemma cprod_modl A B G H :
+ A \* B = G -> A \subset H -> A \* (B :&: H) = G :&: H.
+Proof.
+case/cprodP=> [[U V -> -> {A B}]] defG cUV sUH.
+by rewrite cprodE; [rewrite group_modl ?defG | rewrite subIset ?cUV].
+Qed.
+
+Lemma cprod_modr A B G H :
+ A \* B = G -> B \subset H -> (H :&: A) \* B = H :&: G.
+Proof. by rewrite -!(cprodC B) !(setIC H); exact: cprod_modl. Qed.
+
+Lemma bigcprodYP (I : finType) (P : pred I) (H : I -> {group gT}) :
+ reflect (forall i j, P i -> P j -> i != j -> H i \subset 'C(H j))
+ (\big[cprod/1]_(i | P i) H i == (\prod_(i | P i) H i)%G).
+Proof.
+apply: (iffP eqP) => [defG i j Pi Pj neq_ij | cHH].
+ rewrite (bigD1 j) // (bigD1 i) /= ?cprodA in defG; last exact/andP.
+ by case/cprodP: defG => [[K _ /cprodP[//]]].
+set Q := P; have: subpred Q P by [].
+elim: {Q}_.+1 {-2}Q (ltnSn #|Q|) => // n IHn Q leQn sQP.
+have [i Qi | Q0] := pickP Q; last by rewrite !big_pred0.
+rewrite (cardD1x Qi) add1n ltnS !(bigD1 i Qi) /= in leQn *.
+rewrite {}IHn {n leQn}// => [|j /andP[/sQP //]].
+rewrite bigprodGE cprodEY // gen_subG; apply/bigcupsP=> j /andP[neq_ji Qj].
+by rewrite cHH ?sQP.
+Qed.
+
+Lemma bigcprodEY I r (P : pred I) (H : I -> {group gT}) G :
+ abelian G -> (forall i, P i -> H i \subset G) ->
+ \big[cprod/1]_(i <- r | P i) H i = (\prod_(i <- r | P i) H i)%G.
+Proof.
+move=> cGG sHG; apply/eqP; rewrite !(big_tnth _ _ r).
+by apply/bigcprodYP=> i j Pi Pj _; rewrite (sub_abelian_cent2 cGG) ?sHG.
+Qed.
+
+Lemma perm_bigcprod (I : eqType) r1 r2 (A : I -> {set gT}) G x :
+ \big[cprod/1]_(i <- r1) A i = G -> {in r1, forall i, x i \in A i} ->
+ perm_eq r1 r2 ->
+ \prod_(i <- r1) x i = \prod_(i <- r2) x i.
+Proof.
+elim: r1 r2 G => [|i r1 IHr] r2 G defG Ax eq_r12.
+ by rewrite perm_eq_sym in eq_r12; rewrite (perm_eq_small _ eq_r12) ?big_nil.
+have /rot_to[n r3 Dr2]: i \in r2 by rewrite -(perm_eq_mem eq_r12) mem_head.
+transitivity (\prod_(j <- rot n r2) x j).
+ rewrite Dr2 !big_cons in defG Ax *; have [[_ G1 _ defG1] _ _] := cprodP defG.
+ rewrite (IHr r3 G1) //; first by case/allP/andP: Ax => _ /allP.
+ by rewrite -(perm_cons i) -Dr2 perm_eq_sym perm_rot perm_eq_sym.
+rewrite -{-2}(cat_take_drop n r2) in eq_r12 *.
+rewrite (eq_big_perm _ eq_r12) !big_cat /= !(big_nth i) !big_mkord in defG *.
+have /cprodP[[G1 G2 defG1 defG2] _ /centsP-> //] := defG.
+ rewrite defG2 -(bigcprodW defG2) mem_prodg // => k _; apply: Ax.
+ by rewrite (perm_eq_mem eq_r12) mem_cat orbC mem_nth.
+rewrite defG1 -(bigcprodW defG1) mem_prodg // => k _; apply: Ax.
+by rewrite (perm_eq_mem eq_r12) mem_cat mem_nth.
+Qed.
+
+Lemma reindex_bigcprod (I J : finType) (h : J -> I) P (A : I -> {set gT}) G x :
+ {on SimplPred P, bijective h} -> \big[cprod/1]_(i | P i) A i = G ->
+ {in SimplPred P, forall i, x i \in A i} ->
+ \prod_(i | P i) x i = \prod_(j | P (h j)) x (h j).
+Proof.
+case=> h1 hK h1K; rewrite -!(big_filter _ P) filter_index_enum => defG Ax.
+rewrite -(big_map h P x) -(big_filter _ P) filter_map filter_index_enum.
+apply: perm_bigcprod defG _ _ => [i|]; first by rewrite mem_enum => /Ax.
+apply: uniq_perm_eq (enum_uniq P) _ _ => [|i].
+ by apply/dinjectiveP; apply: (can_in_inj hK).
+rewrite mem_enum; apply/idP/imageP=> [Pi | [j Phj ->//]].
+by exists (h1 i); rewrite ?inE h1K.
+Qed.
+
+(* Direct product *)
+
+Lemma dprod1g : left_id 1 dprod.
+Proof. by move=> A; rewrite /dprod subsetIl cprod1g. Qed.
+
+Lemma dprodg1 : right_id 1 dprod.
+Proof. by move=> A; rewrite /dprod subsetIr cprodg1. Qed.
+
+Lemma dprodP A B G :
+ A \x B = G -> [/\ are_groups A B, A * B = G, B \subset 'C(A) & A :&: B = 1].
+Proof.
+rewrite /dprod; case: ifP => trAB; last by case/group_not0.
+by case/cprodP=> gAB; split=> //; case: gAB trAB => ? ? -> -> /trivgP.
+Qed.
+
+Lemma dprodE G H : H \subset 'C(G) -> G :&: H = 1 -> G \x H = G * H.
+Proof. by move=> cGH trGH; rewrite /dprod trGH sub1G cprodE. Qed.
+
+Lemma dprodEY G H : H \subset 'C(G) -> G :&: H = 1 -> G \x H = G <*> H.
+Proof. by move=> cGH trGH; rewrite /dprod trGH subxx cprodEY. Qed.
+
+Lemma dprodEcp A B : A :&: B = 1 -> A \x B = A \* B.
+Proof. by move=> trAB; rewrite /dprod trAB subxx. Qed.
+
+Lemma dprodEsd A B : B \subset 'C(A) -> A \x B = A ><| B.
+Proof. by rewrite /dprod /cprod => ->. Qed.
+
+Lemma dprodWcp A B G : A \x B = G -> A \* B = G.
+Proof. by move=> defG; have [_ _ _ /dprodEcp <-] := dprodP defG. Qed.
+
+Lemma dprodWsd A B G : A \x B = G -> A ><| B = G.
+Proof. by move=> defG; have [_ _ /dprodEsd <-] := dprodP defG. Qed.
+
+Lemma dprodW A B G : A \x B = G -> A * B = G.
+Proof. by move/dprodWsd/sdprodW. Qed.
+
+Lemma dprodWC A B G : A \x B = G -> B * A = G.
+Proof. by move/dprodWsd/sdprodWC. Qed.
+
+Lemma dprodWY A B G : A \x B = G -> A <*> B = G.
+Proof. by move/dprodWsd/sdprodWY. Qed.
+
+Lemma cprod_card_dprod G A B :
+ A \* B = G -> #|A| * #|B| <= #|G| -> A \x B = G.
+Proof. by case/cprodP=> [[K H -> ->] <- cKH] /cardMg_TI; exact: dprodE. Qed.
+
+Lemma dprodJ A B x : (A \x B) :^ x = A :^ x \x B :^ x.
+Proof.
+rewrite /dprod -conjIg sub_conjg conjs1g -cprodJ.
+by case: ifP => _ //; exact: imset0.
+Qed.
+
+Lemma dprod_normal2 A B G : A \x B = G -> A <| G /\ B <| G.
+Proof. by move/dprodWcp/cprod_normal2. Qed.
+
+Lemma dprodYP K H : reflect (K \x H = K <*> H) (H \subset 'C(K) :\: K^#).
+Proof.
+rewrite subsetD -setI_eq0 setIDA setD_eq0 setIC subG1 /=.
+by apply: (iffP andP) => [[cKH /eqP/dprodEY->] | /dprodP[_ _ -> ->]].
+Qed.
+
+Lemma dprodC : commutative dprod.
+Proof. by move=> A B; rewrite /dprod setIC cprodC. Qed.
+
+Lemma dprodWsdC A B G : A \x B = G -> B ><| A = G.
+Proof. by rewrite dprodC => /dprodWsd. Qed.
+
+Lemma dprodA : associative dprod.
+Proof.
+move=> A B C; case A1: (A == 1); first by rewrite (eqP A1) !dprod1g.
+case B1: (B == 1); first by rewrite (eqP B1) dprod1g dprodg1.
+case C1: (C == 1); first by rewrite (eqP C1) !dprodg1.
+rewrite /dprod (fun_if (cprod A)) (fun_if (cprod^~ C)) -cprodA.
+rewrite -(cprodC set0) !cprod0g cprod_ntriv ?B1 ?{}C1 //.
+case: and3P B1 => [[] | _ _]; last by rewrite cprodC cprod0g !if_same.
+case/isgroupP=> H ->; case/isgroupP=> K -> {B C}; move/cent_joinEr=> eHK H1.
+rewrite cprod_ntriv ?trivMg ?{}A1 ?{}H1 // mulG_subG.
+case: and4P => [[] | _]; last by rewrite !if_same.
+case/isgroupP=> G ->{A} _ cGH _; rewrite cprodEY // -eHK.
+case trGH: (G :&: H \subset _); case trHK: (H :&: K \subset _); last first.
+- by rewrite !if_same.
+- rewrite if_same; case: ifP => // trG_HK; case/negP: trGH.
+ by apply: subset_trans trG_HK; rewrite setIS ?joing_subl.
+- rewrite if_same; case: ifP => // trGH_K; case/negP: trHK.
+ by apply: subset_trans trGH_K; rewrite setSI ?joing_subr.
+do 2![case: ifP] => // trGH_K trG_HK; [case/negP: trGH_K | case/negP: trG_HK].
+ apply: subset_trans trHK; rewrite subsetI subsetIr -{2}(mulg1 H) -mulGS.
+ rewrite setIC group_modl ?joing_subr //= cent_joinEr // -eHK.
+ by rewrite -group_modr ?joing_subl //= setIC -(normC (sub1G _)) mulSg.
+apply: subset_trans trGH; rewrite subsetI subsetIl -{2}(mul1g H) -mulSG.
+rewrite setIC group_modr ?joing_subl //= eHK -(cent_joinEr cGH).
+by rewrite -group_modl ?joing_subr //= setIC (normC (sub1G _)) mulgS.
+Qed.
+
+Canonical dprod_law := Monoid.Law dprodA dprod1g dprodg1.
+Canonical dprod_abelaw := Monoid.ComLaw dprodC.
+
+Lemma bigdprodWcp I (r : seq I) P F G :
+ \big[dprod/1]_(i <- r | P i) F i = G -> \big[cprod/1]_(i <- r | P i) F i = G.
+Proof.
+elim/big_rec2: _ G => // i A B _ IH G /dprodP[[K H -> defB] <- cKH _].
+by rewrite (IH H) // cprodE -defB.
+Qed.
+
+Lemma bigdprodW I (r : seq I) P F G :
+ \big[dprod/1]_(i <- r | P i) F i = G -> \prod_(i <- r | P i) F i = G.
+Proof. by move/bigdprodWcp; exact: bigcprodW. Qed.
+
+Lemma bigdprodWY I (r : seq I) P F G :
+ \big[dprod/1]_(i <- r | P i) F i = G -> << \bigcup_(i <- r | P i) F i >> = G.
+Proof. by move/bigdprodWcp; exact: bigcprodWY. Qed.
+
+Lemma bigdprodYP (I : finType) (P : pred I) (F : I -> {group gT}) :
+ reflect (forall i, P i ->
+ (\prod_(j | P j && (j != i)) F j)%G \subset 'C(F i) :\: (F i)^#)
+ (\big[dprod/1]_(i | P i) F i == (\prod_(i | P i) F i)%G).
+Proof.
+apply: (iffP eqP) => [defG i Pi | dxG].
+ rewrite !(bigD1 i Pi) /= in defG; have [[_ G' _ defG'] _ _ _] := dprodP defG.
+ by apply/dprodYP; rewrite -defG defG' bigprodGE (bigdprodWY defG').
+set Q := P; have: subpred Q P by [].
+elim: {Q}_.+1 {-2}Q (ltnSn #|Q|) => // n IHn Q leQn sQP.
+have [i Qi | Q0] := pickP Q; last by rewrite !big_pred0.
+rewrite (cardD1x Qi) add1n ltnS !(bigD1 i Qi) /= in leQn *.
+rewrite {}IHn {n leQn}// => [|j /andP[/sQP //]].
+apply/dprodYP; apply: subset_trans (dxG i (sQP i Qi)); rewrite !bigprodGE.
+by apply: genS; apply/bigcupsP=> j /andP[Qj ne_ji]; rewrite (bigcup_max j) ?sQP.
+Qed.
+
+Lemma dprod_modl A B G H :
+ A \x B = G -> A \subset H -> A \x (B :&: H) = G :&: H.
+Proof.
+case/dprodP=> [[U V -> -> {A B}]] defG cUV trUV sUH.
+rewrite dprodEcp; first by apply: cprod_modl; rewrite ?cprodE.
+by rewrite setIA trUV (setIidPl _) ?sub1G.
+Qed.
+
+Lemma dprod_modr A B G H :
+ A \x B = G -> B \subset H -> (H :&: A) \x B = H :&: G.
+Proof. by rewrite -!(dprodC B) !(setIC H); exact: dprod_modl. Qed.
+
+Lemma subcent_dprod B C G A :
+ B \x C = G -> A \subset 'N(B) :&: 'N(C) -> 'C_B(A) \x 'C_C(A) = 'C_G(A).
+Proof.
+move=> defG; have [_ _ cBC _] := dprodP defG; move: defG.
+by rewrite !dprodEsd 1?(centSS _ _ cBC) ?subsetIl //; exact: subcent_sdprod.
+Qed.
+
+Lemma dprod_card A B G : A \x B = G -> (#|A| * #|B|)%N = #|G|.
+Proof. by case/dprodP=> [[H K -> ->] <- _]; move/TI_cardMg. Qed.
+
+Lemma bigdprod_card I r (P : pred I) E G :
+ \big[dprod/1]_(i <- r | P i) E i = G ->
+ (\prod_(i <- r | P i) #|E i|)%N = #|G|.
+Proof.
+elim/big_rec2: _ G => [G <- | i A B _ IH G defG]; first by rewrite cards1.
+have [[_ H _ defH] _ _ _] := dprodP defG.
+by rewrite -(dprod_card defG) (IH H) defH.
+Qed.
+
+Lemma bigcprod_card_dprod I r (P : pred I) (A : I -> {set gT}) G :
+ \big[cprod/1]_(i <- r | P i) A i = G ->
+ \prod_(i <- r | P i) #|A i| <= #|G| ->
+ \big[dprod/1]_(i <- r | P i) A i = G.
+Proof.
+elim: r G => [|i r IHr]; rewrite !(big_nil, big_cons) //; case: ifP => _ // G.
+case/cprodP=> [[K H -> defH]]; rewrite defH => <- cKH leKH_G.
+have /implyP := leq_trans leKH_G (dvdn_leq _ (dvdn_cardMg K H)).
+rewrite muln_gt0 leq_pmul2l !cardG_gt0 //= => /(IHr H defH){defH}defH.
+by rewrite defH dprodE // cardMg_TI // -(bigdprod_card defH).
+Qed.
+
+Lemma bigcprod_coprime_dprod (I : finType) (P : pred I) (A : I -> {set gT}) G :
+ \big[cprod/1]_(i | P i) A i = G ->
+ (forall i j, P i -> P j -> i != j -> coprime #|A i| #|A j|) ->
+ \big[dprod/1]_(i | P i) A i = G.
+Proof.
+move=> defG coA; set Q := P in defG *; have: subpred Q P by [].
+elim: {Q}_.+1 {-2}Q (ltnSn #|Q|) => // m IHm Q leQm in G defG * => sQP.
+have [i Qi | Q0] := pickP Q; last by rewrite !big_pred0 in defG *.
+move: defG; rewrite !(bigD1 i Qi) /= => /cprodP[[Hi Gi defAi defGi] <-].
+rewrite defAi defGi => cHGi.
+have{defGi} defGi: \big[dprod/1]_(j | Q j && (j != i)) A j = Gi.
+ by apply: IHm => [||j /andP[/sQP]] //; rewrite (cardD1x Qi) in leQm.
+rewrite defGi dprodE // coprime_TIg // -defAi -(bigdprod_card defGi).
+elim/big_rec: _ => [|j n /andP[neq_ji Qj] IHn]; first exact: coprimen1.
+by rewrite coprime_mulr coprime_sym coA ?sQP.
+Qed.
+
+Lemma mem_dprod G A B x : A \x B = G -> x \in G ->
+ exists y, exists z,
+ [/\ y \in A, z \in B, x = y * z &
+ {in A & B, forall u t, x = u * t -> u = y /\ t = z}].
+Proof.
+move=> defG; have [_ _ cBA _] := dprodP defG.
+by apply: mem_sdprod; rewrite -dprodEsd.
+Qed.
+
+Lemma mem_bigdprod (I : finType) (P : pred I) F G x :
+ \big[dprod/1]_(i | P i) F i = G -> x \in G ->
+ exists c, [/\ forall i, P i -> c i \in F i, x = \prod_(i | P i) c i
+ & forall e, (forall i, P i -> e i \in F i) ->
+ x = \prod_(i | P i) e i ->
+ forall i, P i -> e i = c i].
+Proof.
+move=> defG; rewrite -(bigdprodW defG) => /prodsgP[c Fc ->].
+exists c; split=> // e Fe eq_ce i Pi.
+set r := index_enum _ in defG eq_ce.
+have: i \in r by rewrite -[r]enumT mem_enum.
+elim: r G defG eq_ce => // j r IHr G; rewrite !big_cons inE.
+case Pj: (P j); last by case: eqP (IHr G) => // eq_ij; rewrite eq_ij Pj in Pi.
+case/dprodP=> [[K H defK defH] _ _]; rewrite defK defH => tiFjH eq_ce.
+suffices{i Pi IHr} eq_cej: c j = e j.
+ case/predU1P=> [-> //|]; apply: IHr defH _.
+ by apply: (mulgI (c j)); rewrite eq_ce eq_cej.
+rewrite !(big_nth j) !big_mkord in defH eq_ce.
+move/(congr1 (divgr K H)) : eq_ce; move/bigdprodW: defH => defH.
+by rewrite !divgrMid // -?defK -?defH ?mem_prodg // => *; rewrite ?Fc ?Fe.
+Qed.
+
+End InternalProd.
+
+Implicit Arguments complP [gT H A B].
+Implicit Arguments splitsP [gT A B].
+Implicit Arguments sdprod_normal_complP [gT K H G].
+Implicit Arguments dprodYP [gT K H].
+Implicit Arguments bigdprodYP [gT I P F].
+
+Section MorphimInternalProd.
+
+Variables (gT rT : finGroupType) (D : {group gT}) (f : {morphism D >-> rT}).
+
+Section OneProd.
+
+Variables G H K : {group gT}.
+Hypothesis sGD : G \subset D.
+
+Lemma morphim_pprod : pprod K H = G -> pprod (f @* K) (f @* H) = f @* G.
+Proof.
+case/pprodP=> _ defG mKH; rewrite pprodE ?morphim_norms //.
+by rewrite -morphimMl ?(subset_trans _ sGD) -?defG // mulG_subl.
+Qed.
+
+Lemma morphim_coprime_sdprod :
+ K ><| H = G -> coprime #|K| #|H| -> f @* K ><| f @* H = f @* G.
+Proof.
+rewrite /sdprod => defG coHK; move: defG.
+by rewrite !coprime_TIg ?coprime_morph // !subxx; exact: morphim_pprod.
+Qed.
+
+Lemma injm_sdprod : 'injm f -> K ><| H = G -> f @* K ><| f @* H = f @* G.
+Proof.
+move=> inj_f; case/sdprodP=> _ defG nKH tiKH.
+by rewrite /sdprod -injmI // tiKH morphim1 subxx morphim_pprod // pprodE.
+Qed.
+
+Lemma morphim_cprod : K \* H = G -> f @* K \* f @* H = f @* G.
+Proof.
+case/cprodP=> _ defG cKH; rewrite /cprod morphim_cents // morphim_pprod //.
+by rewrite pprodE // cents_norm // centsC.
+Qed.
+
+Lemma injm_dprod : 'injm f -> K \x H = G -> f @* K \x f @* H = f @* G.
+Proof.
+move=> inj_f; case/dprodP=> _ defG cHK tiKH.
+by rewrite /dprod -injmI // tiKH morphim1 subxx morphim_cprod // cprodE.
+Qed.
+
+Lemma morphim_coprime_dprod :
+ K \x H = G -> coprime #|K| #|H| -> f @* K \x f @* H = f @* G.
+Proof.
+rewrite /dprod => defG coHK; move: defG.
+by rewrite !coprime_TIg ?coprime_morph // !subxx; exact: morphim_cprod.
+Qed.
+
+End OneProd.
+
+Implicit Type G : {group gT}.
+
+Lemma morphim_bigcprod I r (P : pred I) (H : I -> {group gT}) G :
+ G \subset D -> \big[cprod/1]_(i <- r | P i) H i = G ->
+ \big[cprod/1]_(i <- r | P i) f @* H i = f @* G.
+Proof.
+elim/big_rec2: _ G => [|i fB B Pi def_fB] G sGD defG.
+ by rewrite -defG morphim1.
+case/cprodP: defG (defG) => [[Hi Gi -> defB] _ _]; rewrite defB => defG.
+rewrite (def_fB Gi) //; first exact: morphim_cprod.
+by apply: subset_trans sGD; case/cprod_normal2: defG => _ /andP[].
+Qed.
+
+Lemma injm_bigdprod I r (P : pred I) (H : I -> {group gT}) G :
+ G \subset D -> 'injm f -> \big[dprod/1]_(i <- r | P i) H i = G ->
+ \big[dprod/1]_(i <- r | P i) f @* H i = f @* G.
+Proof.
+move=> sGD injf; elim/big_rec2: _ G sGD => [|i fB B Pi def_fB] G sGD defG.
+ by rewrite -defG morphim1.
+case/dprodP: defG (defG) => [[Hi Gi -> defB] _ _ _]; rewrite defB => defG.
+rewrite (def_fB Gi) //; first exact: injm_dprod.
+by apply: subset_trans sGD; case/dprod_normal2: defG => _ /andP[].
+Qed.
+
+Lemma morphim_coprime_bigdprod (I : finType) P (H : I -> {group gT}) G :
+ G \subset D -> \big[dprod/1]_(i | P i) H i = G ->
+ (forall i j, P i -> P j -> i != j -> coprime #|H i| #|H j|) ->
+ \big[dprod/1]_(i | P i) f @* H i = f @* G.
+Proof.
+move=> sGD /bigdprodWcp defG coH; have def_fG := morphim_bigcprod sGD defG.
+by apply: bigcprod_coprime_dprod => // i j *; rewrite coprime_morph ?coH.
+Qed.
+
+End MorphimInternalProd.
+
+Section QuotientInternalProd.
+
+Variables (gT : finGroupType) (G K H M : {group gT}).
+
+Hypothesis nMG: G \subset 'N(M).
+
+Lemma quotient_pprod : pprod K H = G -> pprod (K / M) (H / M) = G / M.
+Proof. exact: morphim_pprod. Qed.
+
+Lemma quotient_coprime_sdprod :
+ K ><| H = G -> coprime #|K| #|H| -> (K / M) ><| (H / M) = G / M.
+Proof. exact: morphim_coprime_sdprod. Qed.
+
+Lemma quotient_cprod : K \* H = G -> (K / M) \* (H / M) = G / M.
+Proof. exact: morphim_cprod. Qed.
+
+Lemma quotient_coprime_dprod :
+ K \x H = G -> coprime #|K| #|H| -> (K / M) \x (H / M) = G / M.
+Proof. exact: morphim_coprime_dprod. Qed.
+
+End QuotientInternalProd.
+
+Section ExternalDirProd.
+
+Variables gT1 gT2 : finGroupType.
+
+Definition extprod_mulg (x y : gT1 * gT2) := (x.1 * y.1, x.2 * y.2).
+Definition extprod_invg (x : gT1 * gT2) := (x.1^-1, x.2^-1).
+
+Lemma extprod_mul1g : left_id (1, 1) extprod_mulg.
+Proof. case=> x1 x2; congr (_, _); exact: mul1g. Qed.
+
+Lemma extprod_mulVg : left_inverse (1, 1) extprod_invg extprod_mulg.
+Proof. by move=> x; congr (_, _); exact: mulVg. Qed.
+
+Lemma extprod_mulgA : associative extprod_mulg.
+Proof. by move=> x y z; congr (_, _); exact: mulgA. Qed.
+
+Definition extprod_groupMixin :=
+ Eval hnf in FinGroup.Mixin extprod_mulgA extprod_mul1g extprod_mulVg.
+Canonical extprod_baseFinGroupType :=
+ Eval hnf in BaseFinGroupType (gT1 * gT2) extprod_groupMixin.
+Canonical prod_group := FinGroupType extprod_mulVg.
+
+Lemma group_setX (H1 : {group gT1}) (H2 : {group gT2}) : group_set (setX H1 H2).
+Proof.
+apply/group_setP; split; first by rewrite inE !group1.
+case=> [x1 x2] [y1 y2]; rewrite !inE; case/andP=> Hx1 Hx2; case/andP=> Hy1 Hy2.
+by rewrite /= !groupM.
+Qed.
+
+Canonical setX_group H1 H2 := Group (group_setX H1 H2).
+
+Definition pairg1 x : gT1 * gT2 := (x, 1).
+Definition pair1g x : gT1 * gT2 := (1, x).
+
+Lemma pairg1_morphM : {morph pairg1 : x y / x * y}.
+Proof. by move=> x y /=; rewrite {2}/mulg /= /extprod_mulg /= mul1g. Qed.
+
+Canonical pairg1_morphism := @Morphism _ _ setT _ (in2W pairg1_morphM).
+
+Lemma pair1g_morphM : {morph pair1g : x y / x * y}.
+Proof. by move=> x y /=; rewrite {2}/mulg /= /extprod_mulg /= mul1g. Qed.
+
+Canonical pair1g_morphism := @Morphism _ _ setT _ (in2W pair1g_morphM).
+
+Lemma fst_morphM : {morph (@fst gT1 gT2) : x y / x * y}.
+Proof. by move=> x y. Qed.
+
+Lemma snd_morphM : {morph (@snd gT1 gT2) : x y / x * y}.
+Proof. by move=> x y. Qed.
+
+Canonical fst_morphism := @Morphism _ _ setT _ (in2W fst_morphM).
+
+Canonical snd_morphism := @Morphism _ _ setT _ (in2W snd_morphM).
+
+Lemma injm_pair1g : 'injm pair1g.
+Proof. by apply/subsetP=> x /morphpreP[_ /set1P[->]]; exact: set11. Qed.
+
+Lemma injm_pairg1 : 'injm pairg1.
+Proof. by apply/subsetP=> x /morphpreP[_ /set1P[->]]; exact: set11. Qed.
+
+Lemma morphim_pairg1 (H1 : {set gT1}) : pairg1 @* H1 = setX H1 1.
+Proof. by rewrite -imset2_pair imset2_set1r morphimEsub ?subsetT. Qed.
+
+Lemma morphim_pair1g (H2 : {set gT2}) : pair1g @* H2 = setX 1 H2.
+Proof. by rewrite -imset2_pair imset2_set1l morphimEsub ?subsetT. Qed.
+
+Lemma morphim_fstX (H1: {set gT1}) (H2 : {group gT2}) :
+ [morphism of fun x => x.1] @* setX H1 H2 = H1.
+Proof.
+apply/eqP; rewrite eqEsubset morphimE setTI /=.
+apply/andP; split; apply/subsetP=> x.
+ by case/imsetP=> x0; rewrite inE; move/andP=> [Hx1 _] ->.
+move=> Hx1; apply/imsetP; exists (x, 1); last by trivial.
+by rewrite in_setX Hx1 /=.
+Qed.
+
+Lemma morphim_sndX (H1: {group gT1}) (H2 : {set gT2}) :
+ [morphism of fun x => x.2] @* setX H1 H2 = H2.
+Proof.
+apply/eqP; rewrite eqEsubset morphimE setTI /=.
+apply/andP; split; apply/subsetP=> x.
+ by case/imsetP=> x0; rewrite inE; move/andP=> [_ Hx2] ->.
+move=>Hx2; apply/imsetP; exists (1, x); last by [].
+by rewrite in_setX Hx2 andbT.
+Qed.
+
+Lemma setX_prod (H1 : {set gT1}) (H2 : {set gT2}) :
+ setX H1 1 * setX 1 H2 = setX H1 H2.
+Proof.
+apply/setP=> [[x y]]; rewrite !inE /=.
+apply/imset2P/andP=> [[[x1 u1] [v1 y1]] | [Hx Hy]].
+ rewrite !inE /= => /andP[Hx1 /eqP->] /andP[/eqP-> Hx] [-> ->].
+ by rewrite mulg1 mul1g.
+exists (x, 1 : gT2) (1 : gT1, y); rewrite ?inE ?Hx ?eqxx //.
+by rewrite /mulg /= /extprod_mulg /= mulg1 mul1g.
+Qed.
+
+Lemma setX_dprod (H1 : {group gT1}) (H2 : {group gT2}) :
+ setX H1 1 \x setX 1 H2 = setX H1 H2.
+Proof.
+rewrite dprodE ?setX_prod //.
+ apply/centsP=> [[x u]]; rewrite !inE /= => /andP[/eqP-> _] [v y].
+ by rewrite !inE /= => /andP[_ /eqP->]; congr (_, _); rewrite ?mul1g ?mulg1.
+apply/trivgP; apply/subsetP=> [[x y]]; rewrite !inE /= -!andbA.
+by case/and4P=> _ /eqP-> /eqP->; rewrite eqxx.
+Qed.
+
+Lemma isog_setX1 (H1 : {group gT1}) : isog H1 (setX H1 1).
+Proof.
+apply/isogP; exists [morphism of restrm (subsetT H1) pairg1].
+ by rewrite injm_restrm ?injm_pairg1.
+by rewrite morphim_restrm morphim_pairg1 setIid.
+Qed.
+
+Lemma isog_set1X (H2 : {group gT2}) : isog H2 (setX 1 H2).
+Proof.
+apply/isogP; exists [morphism of restrm (subsetT H2) pair1g].
+ by rewrite injm_restrm ?injm_pair1g.
+by rewrite morphim_restrm morphim_pair1g setIid.
+Qed.
+
+Lemma setX_gen (H1 : {set gT1}) (H2 : {set gT2}) :
+ 1 \in H1 -> 1 \in H2 -> <<setX H1 H2>> = setX <<H1>> <<H2>>.
+Proof.
+move=> H1_1 H2_1; apply/eqP.
+rewrite eqEsubset gen_subG setXS ?subset_gen //.
+rewrite -setX_prod -morphim_pair1g -morphim_pairg1 !morphim_gen ?subsetT //.
+by rewrite morphim_pair1g morphim_pairg1 mul_subG // genS // setXS ?sub1set.
+Qed.
+
+End ExternalDirProd.
+
+Section ExternalSDirProd.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (R : {group rT}).
+
+(* The pair (a, x) denotes the product sdpair2 a * sdpair1 x *)
+
+Inductive sdprod_by (to : groupAction D R) : predArgType :=
+ SdPair (ax : aT * rT) of ax \in setX D R.
+
+Coercion pair_of_sd to (u : sdprod_by to) := let: SdPair ax _ := u in ax.
+
+Variable to : groupAction D R.
+
+Notation sdT := (sdprod_by to).
+Notation sdval := (@pair_of_sd to).
+
+Canonical sdprod_subType := Eval hnf in [subType for sdval].
+Definition sdprod_eqMixin := Eval hnf in [eqMixin of sdT by <:].
+Canonical sdprod_eqType := Eval hnf in EqType sdT sdprod_eqMixin.
+Definition sdprod_choiceMixin := [choiceMixin of sdT by <:].
+Canonical sdprod_choiceType := ChoiceType sdT sdprod_choiceMixin.
+Definition sdprod_countMixin := [countMixin of sdT by <:].
+Canonical sdprod_countType := CountType sdT sdprod_countMixin.
+Canonical sdprod_subCountType := Eval hnf in [subCountType of sdT].
+Definition sdprod_finMixin := [finMixin of sdT by <:].
+Canonical sdprod_finType := FinType sdT sdprod_finMixin.
+Canonical sdprod_subFinType := Eval hnf in [subFinType of sdT].
+
+Definition sdprod_one := SdPair to (group1 _).
+
+Lemma sdprod_inv_proof (u : sdT) : (u.1^-1, to u.2^-1 u.1^-1) \in setX D R.
+Proof.
+by case: u => [[a x]] /= /setXP[Da Rx]; rewrite inE gact_stable !groupV ?Da.
+Qed.
+
+Definition sdprod_inv u := SdPair to (sdprod_inv_proof u).
+
+Lemma sdprod_mul_proof (u v : sdT) :
+ (u.1 * v.1, to u.2 v.1 * v.2) \in setX D R.
+Proof.
+case: u v => [[a x] /= /setXP[Da Rx]] [[b y] /= /setXP[Db Ry]].
+by rewrite inE !groupM //= gact_stable.
+Qed.
+
+Definition sdprod_mul u v := SdPair to (sdprod_mul_proof u v).
+
+Lemma sdprod_mul1g : left_id sdprod_one sdprod_mul.
+Proof.
+move=> u; apply: val_inj; case: u => [[a x] /=]; case/setXP=> Da _.
+by rewrite gact1 // !mul1g.
+Qed.
+
+Lemma sdprod_mulVg : left_inverse sdprod_one sdprod_inv sdprod_mul.
+Proof.
+move=> u; apply: val_inj; case: u => [[a x] /=]; case/setXP=> Da _.
+by rewrite actKVin ?mulVg.
+Qed.
+
+Lemma sdprod_mulgA : associative sdprod_mul.
+Proof.
+move=> u v w; apply: val_inj; case: u => [[a x]] /=; case/setXP=> Da Rx.
+case: v w => [[b y]] /=; case/setXP=> Db Ry [[c z]] /=; case/setXP=> Dc Rz.
+by rewrite !(actMin to) // gactM ?gact_stable // !mulgA.
+Qed.
+
+Canonical sdprod_groupMixin :=
+ FinGroup.Mixin sdprod_mulgA sdprod_mul1g sdprod_mulVg.
+
+Canonical sdprod_baseFinGroupType :=
+ Eval hnf in BaseFinGroupType sdT sdprod_groupMixin.
+
+Canonical sdprod_groupType := FinGroupType sdprod_mulVg.
+
+Definition sdpair1 x := insubd sdprod_one (1, x) : sdT.
+Definition sdpair2 a := insubd sdprod_one (a, 1) : sdT.
+
+Lemma sdpair1_morphM : {in R &, {morph sdpair1 : x y / x * y}}.
+Proof.
+move=> x y Rx Ry; apply: val_inj.
+by rewrite /= !val_insubd !inE !group1 !groupM ?Rx ?Ry //= mulg1 act1.
+Qed.
+
+Lemma sdpair2_morphM : {in D &, {morph sdpair2 : a b / a * b}}.
+Proof.
+move=> a b Da Db; apply: val_inj.
+by rewrite /= !val_insubd !inE !group1 !groupM ?Da ?Db //= mulg1 gact1.
+Qed.
+
+Canonical sdpair1_morphism := Morphism sdpair1_morphM.
+
+Canonical sdpair2_morphism := Morphism sdpair2_morphM.
+
+Lemma injm_sdpair1 : 'injm sdpair1.
+Proof.
+apply/subsetP=> x /setIP[Rx].
+by rewrite !inE -val_eqE val_insubd inE Rx group1 /=; case/andP.
+Qed.
+
+Lemma injm_sdpair2 : 'injm sdpair2.
+Proof.
+apply/subsetP=> a /setIP[Da].
+by rewrite !inE -val_eqE val_insubd inE Da group1 /=; case/andP.
+Qed.
+
+Lemma sdpairE (u : sdT) : u = sdpair2 u.1 * sdpair1 u.2.
+Proof.
+apply: val_inj; case: u => [[a x] /= /setXP[Da Rx]].
+by rewrite !val_insubd !inE Da Rx !(group1, gact1) // mulg1 mul1g.
+Qed.
+
+Lemma sdpair_act : {in R & D,
+ forall x a, sdpair1 (to x a) = sdpair1 x ^ sdpair2 a}.
+Proof.
+move=> x a Rx Da; apply: val_inj.
+rewrite /= !val_insubd !inE !group1 gact_stable ?Da ?Rx //=.
+by rewrite !mul1g mulVg invg1 mulg1 actKVin ?mul1g.
+Qed.
+
+Lemma sdpair_setact (G : {set rT}) a : G \subset R -> a \in D ->
+ sdpair1 @* (to^~ a @: G) = (sdpair1 @* G) :^ sdpair2 a.
+Proof.
+move=> sGR Da; have GtoR := subsetP sGR; apply/eqP.
+rewrite eqEcard cardJg !(card_injm injm_sdpair1) //; last first.
+ by apply/subsetP=> _ /imsetP[x Gx ->]; rewrite gact_stable ?GtoR.
+rewrite (card_imset _ (act_inj _ _)) leqnn andbT.
+apply/subsetP=> _ /morphimP[xa Rxa /imsetP[x Gx def_xa ->]].
+rewrite mem_conjg -morphV // -sdpair_act ?groupV // def_xa actKin //.
+by rewrite mem_morphim ?GtoR.
+Qed.
+
+Lemma im_sdpair_norm : sdpair2 @* D \subset 'N(sdpair1 @* R).
+Proof.
+apply/subsetP=> _ /morphimP[a _ Da ->].
+rewrite inE -sdpair_setact // morphimS //.
+by apply/subsetP=> _ /imsetP[x Rx ->]; rewrite gact_stable.
+Qed.
+
+Lemma im_sdpair_TI : (sdpair1 @* R) :&: (sdpair2 @* D) = 1.
+Proof.
+apply/trivgP; apply/subsetP=> _ /setIP[/morphimP[x _ Rx ->]].
+case/morphimP=> a _ Da /eqP; rewrite inE -!val_eqE.
+by rewrite !val_insubd !inE Da Rx !group1 /eq_op /= eqxx; case/andP.
+Qed.
+
+Lemma im_sdpair : (sdpair1 @* R) * (sdpair2 @* D) = setT.
+Proof.
+apply/eqP; rewrite -subTset -(normC im_sdpair_norm).
+apply/subsetP=> /= u _; rewrite [u]sdpairE.
+by case: u => [[a x] /= /setXP[Da Rx]]; rewrite mem_mulg ?mem_morphim.
+Qed.
+
+Lemma sdprod_sdpair : sdpair1 @* R ><| sdpair2 @* D = setT.
+Proof. by rewrite sdprodE ?(im_sdpair_norm, im_sdpair, im_sdpair_TI). Qed.
+
+Variables (A : {set aT}) (G : {set rT}).
+
+Lemma gacentEsd : 'C_(|to)(A) = sdpair1 @*^-1 'C(sdpair2 @* A).
+Proof.
+apply/setP=> x; apply/idP/idP.
+ case/setIP=> Rx /afixP cDAx; rewrite mem_morphpre //.
+ apply/centP=> _ /morphimP[a Da Aa ->]; red.
+ by rewrite conjgC -sdpair_act // cDAx // inE Da.
+case/morphpreP=> Rx cAx; rewrite inE Rx; apply/afixP=> a /setIP[Da Aa].
+apply: (injmP injm_sdpair1); rewrite ?gact_stable /= ?sdpair_act //=.
+by rewrite /conjg (centP cAx) ?mulKg ?mem_morphim.
+Qed.
+
+Hypotheses (sAD : A \subset D) (sGR : G \subset R).
+
+Lemma astabEsd : 'C(G | to) = sdpair2 @*^-1 'C(sdpair1 @* G).
+Proof.
+have ssGR := subsetP sGR; apply/setP=> a; apply/idP/idP=> [cGa|].
+ rewrite mem_morphpre ?(astab_dom cGa) //.
+ apply/centP=> _ /morphimP[x Rx Gx ->]; symmetry.
+ by rewrite conjgC -sdpair_act ?(astab_act cGa) ?(astab_dom cGa).
+case/morphpreP=> Da cGa; rewrite !inE Da; apply/subsetP=> x Gx; rewrite inE.
+apply/eqP; apply: (injmP injm_sdpair1); rewrite ?gact_stable ?ssGR //=.
+by rewrite sdpair_act ?ssGR // /conjg -(centP cGa) ?mulKg ?mem_morphim ?ssGR.
+Qed.
+
+Lemma astabsEsd : 'N(G | to) = sdpair2 @*^-1 'N(sdpair1 @* G).
+Proof.
+apply/setP=> a; apply/idP/idP=> [nGa|].
+ have Da := astabs_dom nGa; rewrite mem_morphpre // inE sub_conjg.
+ apply/subsetP=> _ /morphimP[x Rx Gx ->].
+ by rewrite mem_conjgV -sdpair_act // mem_morphim ?gact_stable ?astabs_act.
+case/morphpreP=> Da nGa; rewrite !inE Da; apply/subsetP=> x Gx.
+have Rx := subsetP sGR _ Gx; have Rxa: to x a \in R by rewrite gact_stable.
+rewrite inE -sub1set -(injmSK injm_sdpair1) ?morphim_set1 ?sub1set //=.
+by rewrite sdpair_act ?memJ_norm ?mem_morphim.
+Qed.
+
+Lemma actsEsd : [acts A, on G | to] = (sdpair2 @* A \subset 'N(sdpair1 @* G)).
+Proof. by rewrite sub_morphim_pre -?astabsEsd. Qed.
+
+End ExternalSDirProd.
+
+Section ProdMorph.
+
+Variables gT rT : finGroupType.
+Implicit Types A B : {set gT}.
+Implicit Types G H K : {group gT}.
+Implicit Types C D : {set rT}.
+Implicit Type L : {group rT}.
+
+Section defs.
+
+Variables (A B : {set gT}) (fA fB : gT -> FinGroup.sort rT).
+
+Definition pprodm of B \subset 'N(A) & {in A & B, morph_act 'J 'J fA fB}
+ & {in A :&: B, fA =1 fB} :=
+ fun x => fA (divgr A B x) * fB (remgr A B x).
+
+End defs.
+
+Section Props.
+
+Variables H K : {group gT}.
+Variables (fH : {morphism H >-> rT}) (fK : {morphism K >-> rT}).
+Hypothesis nHK : K \subset 'N(H).
+Hypothesis actf : {in H & K, morph_act 'J 'J fH fK}.
+Hypothesis eqfHK : {in H :&: K, fH =1 fK}.
+
+Notation Local f := (pprodm nHK actf eqfHK).
+
+Lemma pprodmE x a : x \in H -> a \in K -> f (x * a) = fH x * fK a.
+Proof.
+move=> Hx Ka; have: x * a \in H * K by rewrite mem_mulg.
+rewrite -remgrP inE /f rcoset_sym mem_rcoset /divgr -mulgA groupMl //.
+case/andP; move: (remgr H K _) => b Hab Kb; rewrite morphM // -mulgA.
+have Kab: a * b^-1 \in K by rewrite groupM ?groupV.
+by congr (_ * _); rewrite eqfHK 1?inE ?Hab // -morphM // mulgKV.
+Qed.
+
+Lemma pprodmEl : {in H, f =1 fH}.
+Proof. by move=> x Hx; rewrite -(mulg1 x) pprodmE // morph1 !mulg1. Qed.
+
+Lemma pprodmEr : {in K, f =1 fK}.
+Proof. by move=> a Ka; rewrite -(mul1g a) pprodmE // morph1 !mul1g. Qed.
+
+Lemma pprodmM : {in H <*> K &, {morph f: x y / x * y}}.
+Proof.
+move=> xa yb; rewrite norm_joinEr //.
+move=> /imset2P[x a Ha Ka ->{xa}] /imset2P[y b Hy Kb ->{yb}].
+have Hya: y ^ a^-1 \in H by rewrite -mem_conjg (normsP nHK).
+rewrite mulgA -(mulgA x) (conjgCV a y) (mulgA x) -mulgA !pprodmE 1?groupMl //.
+by rewrite morphM // actf ?groupV ?morphV // morphM // !mulgA mulgKV invgK.
+Qed.
+
+Canonical pprodm_morphism := Morphism pprodmM.
+
+Lemma morphim_pprodm A B :
+ A \subset H -> B \subset K -> f @* (A * B) = fH @* A * fK @* B.
+Proof.
+move=> sAH sBK; rewrite [f @* _]morphimEsub /=; last first.
+ by rewrite norm_joinEr // mulgSS.
+apply/setP=> y; apply/imsetP/idP=> [[_ /mulsgP[x a Ax Ba ->] ->{y}] |].
+ have Hx := subsetP sAH x Ax; have Ka := subsetP sBK a Ba.
+ by rewrite pprodmE // mem_imset2 ?mem_morphim.
+case/mulsgP=> _ _ /morphimP[x Hx Ax ->] /morphimP[a Ka Ba ->] ->{y}.
+by exists (x * a); rewrite ?mem_mulg ?pprodmE.
+Qed.
+
+Lemma morphim_pprodml A : A \subset H -> f @* A = fH @* A.
+Proof.
+by move=> sAH; rewrite -{1}(mulg1 A) morphim_pprodm ?sub1G // morphim1 mulg1.
+Qed.
+
+Lemma morphim_pprodmr B : B \subset K -> f @* B = fK @* B.
+Proof.
+by move=> sBK; rewrite -{1}(mul1g B) morphim_pprodm ?sub1G // morphim1 mul1g.
+Qed.
+
+Lemma ker_pprodm : 'ker f = [set x * a^-1 | x in H, a in K & fH x == fK a].
+Proof.
+apply/setP=> y; rewrite 3!inE {1}norm_joinEr //=.
+apply/andP/imset2P=> [[/mulsgP[x a Hx Ka ->{y}]]|[x a Hx]].
+ rewrite pprodmE // => fxa1.
+ by exists x a^-1; rewrite ?invgK // inE groupVr ?morphV // eq_mulgV1 invgK.
+case/setIdP=> Kx /eqP fx ->{y}.
+by rewrite mem_imset2 ?pprodmE ?groupV ?morphV // fx mulgV.
+Qed.
+
+Lemma injm_pprodm :
+ 'injm f = [&& 'injm fH, 'injm fK & fH @* H :&: fK @* K == fH @* K].
+Proof.
+apply/idP/and3P=> [injf | [injfH injfK]].
+ rewrite eq_sym -{1}morphimIdom -(morphim_pprodml (subsetIl _ _)) injmI //.
+ rewrite morphim_pprodml // morphim_pprodmr //=; split=> //.
+ apply/injmP=> x y Hx Hy /=; rewrite -!pprodmEl //.
+ by apply: (injmP injf); rewrite ?mem_gen ?inE ?Hx ?Hy.
+ apply/injmP=> a b Ka Kb /=; rewrite -!pprodmEr //.
+ by apply: (injmP injf); rewrite ?mem_gen //; apply/setUP; right.
+move/eqP=> fHK; rewrite ker_pprodm; apply/subsetP=> y.
+case/imset2P=> x a Hx /setIdP[Ka /eqP fxa] ->.
+have: fH x \in fH @* K by rewrite -fHK inE {2}fxa !mem_morphim.
+case/morphimP=> z Hz Kz /(injmP injfH) def_x.
+rewrite def_x // eqfHK ?inE ?Hz // in fxa.
+by rewrite def_x // (injmP injfK _ _ Kz Ka fxa) mulgV set11.
+Qed.
+
+End Props.
+
+Section Sdprodm.
+
+Variables H K G : {group gT}.
+Variables (fH : {morphism H >-> rT}) (fK : {morphism K >-> rT}).
+Hypothesis eqHK_G : H ><| K = G.
+Hypothesis actf : {in H & K, morph_act 'J 'J fH fK}.
+
+Lemma sdprodm_norm : K \subset 'N(H).
+Proof. by case/sdprodP: eqHK_G. Qed.
+
+Lemma sdprodm_sub : G \subset H <*> K.
+Proof. by case/sdprodP: eqHK_G => _ <- nHK _; rewrite norm_joinEr. Qed.
+
+Lemma sdprodm_eqf : {in H :&: K, fH =1 fK}.
+Proof.
+by case/sdprodP: eqHK_G => _ _ _ -> _ /set1P->; rewrite !morph1.
+Qed.
+
+Definition sdprodm :=
+ restrm sdprodm_sub (pprodm sdprodm_norm actf sdprodm_eqf).
+
+Canonical sdprodm_morphism := Eval hnf in [morphism of sdprodm].
+
+Lemma sdprodmE a b : a \in H -> b \in K -> sdprodm (a * b) = fH a * fK b.
+Proof. exact: pprodmE. Qed.
+
+Lemma sdprodmEl a : a \in H -> sdprodm a = fH a.
+Proof. exact: pprodmEl. Qed.
+
+Lemma sdprodmEr b : b \in K -> sdprodm b = fK b.
+Proof. exact: pprodmEr. Qed.
+
+Lemma morphim_sdprodm A B :
+ A \subset H -> B \subset K -> sdprodm @* (A * B) = fH @* A * fK @* B.
+Proof.
+move=> sAH sBK; rewrite morphim_restrm /= (setIidPr _) ?morphim_pprodm //.
+case/sdprodP: eqHK_G => _ <- _ _; exact: mulgSS.
+Qed.
+
+Lemma im_sdprodm : sdprodm @* G = fH @* H * fK @* K.
+Proof. by rewrite -morphim_sdprodm //; case/sdprodP: eqHK_G => _ ->. Qed.
+
+Lemma morphim_sdprodml A : A \subset H -> sdprodm @* A = fH @* A.
+Proof.
+by move=> sHA; rewrite -{1}(mulg1 A) morphim_sdprodm ?sub1G // morphim1 mulg1.
+Qed.
+
+Lemma morphim_sdprodmr B : B \subset K -> sdprodm @* B = fK @* B.
+Proof.
+by move=> sBK; rewrite -{1}(mul1g B) morphim_sdprodm ?sub1G // morphim1 mul1g.
+Qed.
+
+Lemma ker_sdprodm :
+ 'ker sdprodm = [set a * b^-1 | a in H, b in K & fH a == fK b].
+Proof.
+rewrite ker_restrm (setIidPr _) ?subIset ?ker_pprodm //; apply/orP; left.
+by case/sdprodP: eqHK_G => _ <- nHK _; rewrite norm_joinEr.
+Qed.
+
+Lemma injm_sdprodm :
+ 'injm sdprodm = [&& 'injm fH, 'injm fK & fH @* H :&: fK @* K == 1].
+Proof.
+rewrite ker_sdprodm -(ker_pprodm sdprodm_norm actf sdprodm_eqf) injm_pprodm.
+congr [&& _, _ & _ == _]; have [_ _ _ tiHK] := sdprodP eqHK_G.
+by rewrite -morphimIdom tiHK morphim1.
+Qed.
+
+End Sdprodm.
+
+Section Cprodm.
+
+Variables H K G : {group gT}.
+Variables (fH : {morphism H >-> rT}) (fK : {morphism K >-> rT}).
+Hypothesis eqHK_G : H \* K = G.
+Hypothesis cfHK : fK @* K \subset 'C(fH @* H).
+Hypothesis eqfHK : {in H :&: K, fH =1 fK}.
+
+Lemma cprodm_norm : K \subset 'N(H).
+Proof. by rewrite cents_norm //; case/cprodP: eqHK_G. Qed.
+
+Lemma cprodm_sub : G \subset H <*> K.
+Proof. by case/cprodP: eqHK_G => _ <- cHK; rewrite cent_joinEr. Qed.
+
+Lemma cprodm_actf : {in H & K, morph_act 'J 'J fH fK}.
+Proof.
+case/cprodP: eqHK_G => _ _ cHK a b Ha Kb /=.
+by rewrite /conjg -(centsP cHK b) // -(centsP cfHK (fK b)) ?mulKg ?mem_morphim.
+Qed.
+
+Definition cprodm := restrm cprodm_sub (pprodm cprodm_norm cprodm_actf eqfHK).
+
+Canonical cprodm_morphism := Eval hnf in [morphism of cprodm].
+
+Lemma cprodmE a b : a \in H -> b \in K -> cprodm (a * b) = fH a * fK b.
+Proof. exact: pprodmE. Qed.
+
+Lemma cprodmEl a : a \in H -> cprodm a = fH a.
+Proof. exact: pprodmEl. Qed.
+
+Lemma cprodmEr b : b \in K -> cprodm b = fK b.
+Proof. exact: pprodmEr. Qed.
+
+Lemma morphim_cprodm A B :
+ A \subset H -> B \subset K -> cprodm @* (A * B) = fH @* A * fK @* B.
+Proof.
+move=> sAH sBK; rewrite morphim_restrm /= (setIidPr _) ?morphim_pprodm //.
+case/cprodP: eqHK_G => _ <- _; exact: mulgSS.
+Qed.
+
+Lemma im_cprodm : cprodm @* G = fH @* H * fK @* K.
+Proof.
+by have [_ defHK _] := cprodP eqHK_G; rewrite -{2}defHK morphim_cprodm.
+Qed.
+
+Lemma morphim_cprodml A : A \subset H -> cprodm @* A = fH @* A.
+Proof.
+by move=> sHA; rewrite -{1}(mulg1 A) morphim_cprodm ?sub1G // morphim1 mulg1.
+Qed.
+
+Lemma morphim_cprodmr B : B \subset K -> cprodm @* B = fK @* B.
+Proof.
+by move=> sBK; rewrite -{1}(mul1g B) morphim_cprodm ?sub1G // morphim1 mul1g.
+Qed.
+
+Lemma ker_cprodm : 'ker cprodm = [set a * b^-1 | a in H, b in K & fH a == fK b].
+Proof.
+rewrite ker_restrm (setIidPr _) ?subIset ?ker_pprodm //; apply/orP; left.
+by case/cprodP: eqHK_G => _ <- cHK; rewrite cent_joinEr.
+Qed.
+
+Lemma injm_cprodm :
+ 'injm cprodm = [&& 'injm fH, 'injm fK & fH @* H :&: fK @* K == fH @* K].
+Proof.
+by rewrite ker_cprodm -(ker_pprodm cprodm_norm cprodm_actf eqfHK) injm_pprodm.
+Qed.
+
+End Cprodm.
+
+Section Dprodm.
+
+Variables G H K : {group gT}.
+Variables (fH : {morphism H >-> rT}) (fK : {morphism K >-> rT}).
+Hypothesis eqHK_G : H \x K = G.
+Hypothesis cfHK : fK @* K \subset 'C(fH @* H).
+
+Lemma dprodm_cprod : H \* K = G.
+Proof.
+by rewrite -eqHK_G /dprod; case/dprodP: eqHK_G => _ _ _ ->; rewrite subxx.
+Qed.
+
+Lemma dprodm_eqf : {in H :&: K, fH =1 fK}.
+Proof. by case/dprodP: eqHK_G => _ _ _ -> _ /set1P->; rewrite !morph1. Qed.
+
+Definition dprodm := cprodm dprodm_cprod cfHK dprodm_eqf.
+
+Canonical dprodm_morphism := Eval hnf in [morphism of dprodm].
+
+Lemma dprodmE a b : a \in H -> b \in K -> dprodm (a * b) = fH a * fK b.
+Proof. exact: pprodmE. Qed.
+
+Lemma dprodmEl a : a \in H -> dprodm a = fH a.
+Proof. exact: pprodmEl. Qed.
+
+Lemma dprodmEr b : b \in K -> dprodm b = fK b.
+Proof. exact: pprodmEr. Qed.
+
+Lemma morphim_dprodm A B :
+ A \subset H -> B \subset K -> dprodm @* (A * B) = fH @* A * fK @* B.
+Proof. exact: morphim_cprodm. Qed.
+
+Lemma im_dprodm : dprodm @* G = fH @* H * fK @* K.
+Proof. exact: im_cprodm. Qed.
+
+Lemma morphim_dprodml A : A \subset H -> dprodm @* A = fH @* A.
+Proof. exact: morphim_cprodml. Qed.
+
+Lemma morphim_dprodmr B : B \subset K -> dprodm @* B = fK @* B.
+Proof. exact: morphim_cprodmr. Qed.
+
+Lemma ker_dprodm : 'ker dprodm = [set a * b^-1 | a in H, b in K & fH a == fK b].
+Proof. exact: ker_cprodm. Qed.
+
+Lemma injm_dprodm :
+ 'injm dprodm = [&& 'injm fH, 'injm fK & fH @* H :&: fK @* K == 1].
+Proof.
+rewrite injm_cprodm -(morphimIdom fH K).
+by case/dprodP: eqHK_G => _ _ _ ->; rewrite morphim1.
+Qed.
+
+End Dprodm.
+
+Lemma isog_dprod A B G C D L :
+ A \x B = G -> C \x D = L -> isog A C -> isog B D -> isog G L.
+Proof.
+move=> defG {C D} /dprodP[[C D -> ->] defL cCD trCD].
+case/dprodP: defG (defG) => {A B} [[A B -> ->] defG _ _] dG defC defD.
+case/isogP: defC defL cCD trCD => fA injfA <-{C}.
+case/isogP: defD => fB injfB <-{D} defL cCD trCD.
+apply/isogP; exists (dprodm_morphism dG cCD).
+ by rewrite injm_dprodm injfA injfB trCD eqxx.
+by rewrite /= -{2}defG morphim_dprodm.
+Qed.
+
+End ProdMorph.
+
+Section ExtSdprodm.
+
+Variables gT aT rT : finGroupType.
+Variables (H : {group gT}) (K : {group aT}) (to : groupAction K H).
+Variables (fH : {morphism H >-> rT}) (fK : {morphism K >-> rT}).
+
+Hypothesis actf : {in H & K, morph_act to 'J fH fK}.
+
+Local Notation fsH := (fH \o invm (injm_sdpair1 to)).
+Local Notation fsK := (fK \o invm (injm_sdpair2 to)).
+Let DgH := sdpair1 to @* H.
+Let DgK := sdpair2 to @* K.
+
+Lemma xsdprodm_dom1 : DgH \subset 'dom fsH.
+Proof. by rewrite ['dom _]morphpre_invm. Qed.
+Local Notation gH := (restrm xsdprodm_dom1 fsH).
+
+Lemma xsdprodm_dom2 : DgK \subset 'dom fsK.
+Proof. by rewrite ['dom _]morphpre_invm. Qed.
+Local Notation gK := (restrm xsdprodm_dom2 fsK).
+
+Lemma im_sdprodm1 : gH @* DgH = fH @* H.
+Proof. by rewrite morphim_restrm setIid morphim_comp im_invm. Qed.
+
+Lemma im_sdprodm2 : gK @* DgK = fK @* K.
+Proof. by rewrite morphim_restrm setIid morphim_comp im_invm. Qed.
+
+Lemma xsdprodm_act : {in DgH & DgK, morph_act 'J 'J gH gK}.
+Proof.
+move=> fh fk; case/morphimP=> h _ Hh ->{fh}; case/morphimP=> k _ Kk ->{fk}.
+by rewrite /= -sdpair_act // /restrm /= !invmE ?actf ?gact_stable.
+Qed.
+
+Definition xsdprodm := sdprodm (sdprod_sdpair to) xsdprodm_act.
+Canonical xsdprod_morphism := [morphism of xsdprodm].
+
+Lemma im_xsdprodm : xsdprodm @* setT = fH @* H * fK @* K.
+Proof. by rewrite -im_sdpair morphim_sdprodm // im_sdprodm1 im_sdprodm2. Qed.
+
+Lemma injm_xsdprodm :
+ 'injm xsdprodm = [&& 'injm fH, 'injm fK & fH @* H :&: fK @* K == 1].
+Proof.
+rewrite injm_sdprodm im_sdprodm1 im_sdprodm2 !subG1 /= !ker_restrm !ker_comp.
+rewrite !morphpre_invm !morphimIim.
+by rewrite !morphim_injm_eq1 ?subsetIl ?injm_sdpair1 ?injm_sdpair2.
+Qed.
+
+End ExtSdprodm.
+
+Section DirprodIsom.
+
+Variable gT : finGroupType.
+Implicit Types G H : {group gT}.
+
+Definition mulgm : gT * gT -> _ := prod_curry mulg.
+
+Lemma imset_mulgm (A B : {set gT}) : mulgm @: setX A B = A * B.
+Proof. by rewrite -curry_imset2X. Qed.
+
+Lemma mulgmP H1 H2 G : reflect (H1 \x H2 = G) (misom (setX H1 H2) G mulgm).
+Proof.
+apply: (iffP misomP) => [[pM /isomP[injf /= <-]] | ].
+ have /dprodP[_ /= defX cH12] := setX_dprod H1 H2.
+ rewrite -{4}defX {}defX => /(congr1 (fun A => morphm pM @* A)).
+ move/(morphimS (morphm_morphism pM)): cH12 => /=.
+ have sH1H: setX H1 1 \subset setX H1 H2 by rewrite setXS ?sub1G.
+ have sH2H: setX 1 H2 \subset setX H1 H2 by rewrite setXS ?sub1G.
+ rewrite morphim1 injm_cent ?injmI //= subsetI => /andP[_].
+ by rewrite !morphimEsub //= !imset_mulgm mulg1 mul1g; exact: dprodE.
+case/dprodP=> _ defG cH12 trH12.
+have fM: morphic (setX H1 H2) mulgm.
+ apply/morphicP=> [[x1 x2] [y1 y2] /setXP[_ Hx2] /setXP[Hy1 _]].
+ by rewrite /= mulgA -(mulgA x1) -(centsP cH12 x2) ?mulgA.
+exists fM; apply/isomP; split; last by rewrite morphimEsub //= imset_mulgm.
+apply/subsetP=> [[x1 x2]]; rewrite !inE /= andbC -eq_invg_mul.
+case: eqP => //= <-; rewrite groupV -in_setI trH12 => /set1P->.
+by rewrite invg1 eqxx.
+Qed.
+
+End DirprodIsom.
+
+Implicit Arguments mulgmP [gT H1 H2 G].
+Prenex Implicits mulgm mulgmP.
diff --git a/mathcomp/fingroup/morphism.v b/mathcomp/fingroup/morphism.v
new file mode 100644
index 0000000..7013264
--- /dev/null
+++ b/mathcomp/fingroup/morphism.v
@@ -0,0 +1,1539 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq choice fintype finfun.
+Require Import bigop finset fingroup.
+
+(******************************************************************************)
+(* This file contains the definitions of: *)
+(* *)
+(* {morphism D >-> rT} == *)
+(* the structure type of functions that are group morphisms mapping a *)
+(* domain set D : {set aT} to a type rT; rT must have a finGroupType *)
+(* structure, and D is usually a group (most of the theory expects this). *)
+(* mfun == the coercion projecting {morphism D >-> rT} to aT -> rT *)
+(* *)
+(* Basic examples: *)
+(* idm D == the identity morphism with domain D, or more precisely *)
+(* the identity function, but with a canonical *)
+(* {morphism G -> gT} structure. *)
+(* trivm D == the trivial morphism with domain D *)
+(* If f has a {morphism D >-> rT} structure *)
+(* 'dom f == D *)
+(* f @* A == the image of A by f, where f is defined *)
+(* := f @: (D :&: A) *)
+(* f @*^-1 R == the pre-image of R by f, where f is defined *)
+(* := D :&: f @^-1: R *)
+(* 'ker f == the kernel of f *)
+(* := f @^-1: 1 *)
+(* 'ker_G f == the kernel of f restricted to G *)
+(* := G :&: 'ker f (this is a pure notation) *)
+(* 'injm f <=> f injective on D *)
+(* <-> ker f \subset 1 (this is a pure notation) *)
+(* invm injf == the inverse morphism of f, with domain f @* D, when f *)
+(* is injective (injf : 'injm f) *)
+(* restrm f sDom == the restriction of f to a subset A of D, given *)
+(* (sDom : A \subset D); restrm f sDom is transparently *)
+(* identical to f; the restrmP and domP lemmas provide *)
+(* opaque restrictions. *)
+(* invm f infj == the inverse morphism for an injective f, with domain *)
+(* f @* D, given (injf : 'injm f) *)
+(* *)
+(* G \isog H <=> G and H are isomorphic as groups *)
+(* H \homg G <=> H is a homomorphic image of G *)
+(* isom G H f <=> f maps G isomorphically to H, provided D contains G *)
+(* <-> f @: G^# == H^# *)
+(* *)
+(* If, moreover, g : {morphism G >-> gT} with G : {group aT}, *)
+(* factm sKer sDom == the (natural) factor morphism mapping f @* G to g @* G *)
+(* given sDom : G \subset D, sKer : 'ker f \subset 'ker g *)
+(* ifactm injf g == the (natural) factor morphism mapping f @* G to g @* G *)
+(* when f is injective (injf : 'injm f); here g must *)
+(* be an actual morphism structure, not its function *)
+(* projection. *)
+(* *)
+(* If g has a {morphism G >-> aT} structure for any G : {group gT}, then *)
+(* f \o g has a canonical {morphism g @*^-1 D >-> rT} structure *)
+(* *)
+(* Finally, for an arbitrary function f : aT -> rT *)
+(* morphic D f <=> f preserves group multiplication in D, i.e., *)
+(* f (x * y) = (f x) * (f y) for all x, y in D *)
+(* morphm fM == a function identical to f, but with a canonical *)
+(* {morphism D >-> rT} structure, given fM : morphic D f *)
+(* misom D C f <=> f maps D isomorphically to C *)
+(* := morphic D f && isom D C f *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Reserved Notation "x \isog y" (at level 70).
+
+Section MorphismStructure.
+
+Variables aT rT : finGroupType.
+
+Structure morphism (D : {set aT}) : Type := Morphism {
+ mfun :> aT -> FinGroup.sort rT;
+ _ : {in D &, {morph mfun : x y / x * y}}
+}.
+
+(* We give the most 'lightweight' possible specification to define morphisms:*)
+(* local congruence with the group law of aT. We then provide the properties *)
+(* for the 'textbook' notion of morphism, when the required structures are *)
+(* available (e.g. its domain is a group). *)
+
+Definition morphism_for D of phant rT := morphism D.
+
+Definition clone_morphism D f :=
+ let: Morphism _ fM := f
+ return {type of @Morphism D for f} -> morphism_for D (Phant rT)
+ in fun k => k fM.
+
+Variables (D A : {set aT}) (R : {set rT}) (x : aT) (y : rT) (f : aT -> rT).
+
+CoInductive morphim_spec : Prop := MorphimSpec z & z \in D & z \in A & y = f z.
+
+Lemma morphimP : reflect morphim_spec (y \in f @: (D :&: A)).
+Proof.
+apply: (iffP imsetP) => [] [z]; first by case/setIP; exists z.
+by exists z; first apply/setIP.
+Qed.
+
+Lemma morphpreP : reflect (x \in D /\ f x \in R) (x \in D :&: f @^-1: R).
+Proof. rewrite !inE; exact: andP. Qed.
+
+End MorphismStructure.
+
+Notation "{ 'morphism' D >-> T }" := (morphism_for D (Phant T))
+ (at level 0, format "{ 'morphism' D >-> T }") : group_scope.
+Notation "[ 'morphism' D 'of' f ]" :=
+ (@clone_morphism _ _ D _ (fun fM => @Morphism _ _ D f fM))
+ (at level 0, format "[ 'morphism' D 'of' f ]") : form_scope.
+Notation "[ 'morphism' 'of' f ]" := (clone_morphism (@Morphism _ _ _ f))
+ (at level 0, format "[ 'morphism' 'of' f ]") : form_scope.
+
+Implicit Arguments morphimP [aT rT D A f y].
+Implicit Arguments morphpreP [aT rT D R f x].
+Prenex Implicits morphimP morphpreP.
+
+(* domain, image, preimage, kernel, using phantom types to infer the domain *)
+
+Section MorphismOps1.
+
+Variables (aT rT : finGroupType) (D : {set aT}) (f : {morphism D >-> rT}).
+
+Lemma morphM : {in D &, {morph f : x y / x * y}}.
+Proof. by case f. Qed.
+
+Notation morPhantom := (phantom (aT -> rT)).
+Definition MorPhantom := Phantom (aT -> rT).
+
+Definition dom of morPhantom f := D.
+
+Definition morphim of morPhantom f := fun A => f @: (D :&: A).
+
+Definition morphpre of morPhantom f := fun R : {set rT} => D :&: f @^-1: R.
+
+Definition ker mph := morphpre mph 1.
+
+End MorphismOps1.
+
+Arguments Scope morphim [_ _ group_scope _ _ group_scope].
+Arguments Scope morphpre [_ _ group_scope _ _ group_scope].
+
+Notation "''dom' f" := (dom (MorPhantom f))
+ (at level 10, f at level 8, format "''dom' f") : group_scope.
+
+Notation "''ker' f" := (ker (MorPhantom f))
+ (at level 10, f at level 8, format "''ker' f") : group_scope.
+
+Notation "''ker_' H f" := (H :&: 'ker f)
+ (at level 10, H at level 2, f at level 8, format "''ker_' H f")
+ : group_scope.
+
+Notation "f @* A" := (morphim (MorPhantom f) A)
+ (at level 24, format "f @* A") : group_scope.
+
+Notation "f @*^-1 R" := (morphpre (MorPhantom f) R)
+ (at level 24, format "f @*^-1 R") : group_scope.
+
+Notation "''injm' f" := (pred_of_set ('ker f) \subset pred_of_set 1)
+ (at level 10, f at level 8, format "''injm' f") : group_scope.
+
+Section MorphismTheory.
+
+Variables aT rT : finGroupType.
+Implicit Types A B : {set aT}.
+Implicit Types G H : {group aT}.
+Implicit Types R S : {set rT}.
+Implicit Types M : {group rT}.
+
+(* Most properties of morphims hold only when the domain is a group. *)
+Variables (D : {group aT}) (f : {morphism D >-> rT}).
+
+Lemma morph1 : f 1 = 1.
+Proof. by apply: (mulgI (f 1)); rewrite -morphM ?mulg1. Qed.
+
+Lemma morph_prod I r (P : pred I) F :
+ (forall i, P i -> F i \in D) ->
+ f (\prod_(i <- r | P i) F i) = \prod_( i <- r | P i) f (F i).
+Proof.
+move=> D_F; elim/(big_load (fun x => x \in D)): _.
+elim/big_rec2: _ => [|i _ x Pi [Dx <-]]; first by rewrite morph1.
+by rewrite groupM ?morphM // D_F.
+Qed.
+
+Lemma morphV : {in D, {morph f : x / x^-1}}.
+Proof.
+move=> x Dx; apply: (mulgI (f x)).
+by rewrite -morphM ?groupV // !mulgV morph1.
+Qed.
+
+Lemma morphJ : {in D &, {morph f : x y / x ^ y}}.
+Proof. by move=> * /=; rewrite !morphM ?morphV // ?groupM ?groupV. Qed.
+
+Lemma morphX n : {in D, {morph f : x / x ^+ n}}.
+Proof.
+by elim: n => [|n IHn] x Dx; rewrite ?morph1 // !expgS morphM ?(groupX, IHn).
+Qed.
+
+Lemma morphR : {in D &, {morph f : x y / [~ x, y]}}.
+Proof. by move=> * /=; rewrite morphM ?(groupV, groupJ) // morphJ ?morphV. Qed.
+
+(* morphic image,preimage properties w.r.t. set-theoretic operations *)
+
+Lemma morphimE A : f @* A = f @: (D :&: A). Proof. by []. Qed.
+Lemma morphpreE R : f @*^-1 R = D :&: f @^-1: R. Proof. by []. Qed.
+Lemma kerE : 'ker f = f @*^-1 1. Proof. by []. Qed.
+
+Lemma morphimEsub A : A \subset D -> f @* A = f @: A.
+Proof. by move=> sAD; rewrite /morphim (setIidPr sAD). Qed.
+
+Lemma morphimEdom : f @* D = f @: D.
+Proof. exact: morphimEsub. Qed.
+
+Lemma morphimIdom A : f @* (D :&: A) = f @* A.
+Proof. by rewrite /morphim setIA setIid. Qed.
+
+Lemma morphpreIdom R : D :&: f @*^-1 R = f @*^-1 R.
+Proof. by rewrite /morphim setIA setIid. Qed.
+
+Lemma morphpreIim R : f @*^-1 (f @* D :&: R) = f @*^-1 R.
+Proof.
+apply/setP=> x; rewrite morphimEdom !inE.
+by case Dx: (x \in D); rewrite // mem_imset.
+Qed.
+
+Lemma morphimIim A : f @* D :&: f @* A = f @* A.
+Proof. by apply/setIidPr; rewrite imsetS // setIid subsetIl. Qed.
+
+Lemma mem_morphim A x : x \in D -> x \in A -> f x \in f @* A.
+Proof. by move=> Dx Ax; apply/morphimP; exists x. Qed.
+
+Lemma mem_morphpre R x : x \in D -> f x \in R -> x \in f @*^-1 R.
+Proof. by move=> Dx Rfx; exact/morphpreP. Qed.
+
+Lemma morphimS A B : A \subset B -> f @* A \subset f @* B.
+Proof. by move=> sAB; rewrite imsetS ?setIS. Qed.
+
+Lemma morphim_sub A : f @* A \subset f @* D.
+Proof. by rewrite imsetS // setIid subsetIl. Qed.
+
+Lemma leq_morphim A : #|f @* A| <= #|A|.
+Proof.
+by apply: (leq_trans (leq_imset_card _ _)); rewrite subset_leq_card ?subsetIr.
+Qed.
+
+Lemma morphpreS R S : R \subset S -> f @*^-1 R \subset f @*^-1 S.
+Proof. by move=> sRS; rewrite setIS ?preimsetS. Qed.
+
+Lemma morphpre_sub R : f @*^-1 R \subset D.
+Proof. exact: subsetIl. Qed.
+
+Lemma morphim_setIpre A R : f @* (A :&: f @*^-1 R) = f @* A :&: R.
+Proof.
+apply/setP=> fa; apply/morphimP/setIP=> [[a Da] | [/morphimP[a Da Aa ->] Rfa]].
+ by rewrite !inE Da /= => /andP[Aa Rfa] ->; rewrite mem_morphim.
+by exists a; rewrite // !inE Aa Da.
+Qed.
+
+Lemma morphim0 : f @* set0 = set0.
+Proof. by rewrite morphimE setI0 imset0. Qed.
+
+Lemma morphim_eq0 A : A \subset D -> (f @* A == set0) = (A == set0).
+Proof. by rewrite imset_eq0 => /setIidPr->. Qed.
+
+Lemma morphim_set1 x : x \in D -> f @* [set x] = [set f x].
+Proof. by rewrite /morphim -sub1set => /setIidPr->; exact: imset_set1. Qed.
+
+Lemma morphim1 : f @* 1 = 1.
+Proof. by rewrite morphim_set1 ?morph1. Qed.
+
+Lemma morphimV A : f @* A^-1 = (f @* A)^-1.
+Proof.
+wlog suffices: A / f @* A^-1 \subset (f @* A)^-1.
+ by move=> IH; apply/eqP; rewrite eqEsubset IH -invSg invgK -{1}(invgK A) IH.
+apply/subsetP=> _ /morphimP[x Dx Ax' ->]; rewrite !inE in Ax' *.
+by rewrite -morphV // mem_imset // inE groupV Dx.
+Qed.
+
+Lemma morphpreV R : f @*^-1 R^-1 = (f @*^-1 R)^-1.
+Proof.
+apply/setP=> x; rewrite !inE groupV; case Dx: (x \in D) => //=.
+by rewrite morphV.
+Qed.
+
+Lemma morphimMl A B : A \subset D -> f @* (A * B) = f @* A * f @* B.
+Proof.
+move=> sAD; rewrite /morphim setIC -group_modl // (setIidPr sAD).
+apply/setP=> fxy; apply/idP/idP.
+ case/imsetP=> _ /imset2P[x y Ax /setIP[Dy By] ->] ->{fxy}.
+ by rewrite morphM // (subsetP sAD, mem_imset2) // mem_imset // inE By.
+case/imset2P=> _ _ /imsetP[x Ax ->] /morphimP[y Dy By ->] ->{fxy}.
+by rewrite -morphM // (subsetP sAD, mem_imset) // mem_mulg // inE By.
+Qed.
+
+Lemma morphimMr A B : B \subset D -> f @* (A * B) = f @* A * f @* B.
+Proof.
+move=> sBD; apply: invg_inj.
+by rewrite invMg -!morphimV invMg morphimMl // -invGid invSg.
+Qed.
+
+Lemma morphpreMl R S :
+ R \subset f @* D -> f @*^-1 (R * S) = f @*^-1 R * f @*^-1 S.
+Proof.
+move=> sRfD; apply/setP=> x; rewrite !inE.
+apply/andP/imset2P=> [[Dx] | [y z]]; last first.
+ rewrite !inE => /andP[Dy Rfy] /andP[Dz Rfz] ->.
+ by rewrite ?(groupM, morphM, mem_imset2).
+case/imset2P=> fy fz Rfy Rfz def_fx.
+have /morphimP[y Dy _ def_fy]: fy \in f @* D := subsetP sRfD fy Rfy.
+exists y (y^-1 * x); last by rewrite mulKVg.
+ by rewrite !inE Dy -def_fy.
+by rewrite !inE groupM ?(morphM, morphV, groupV) // def_fx -def_fy mulKg.
+Qed.
+
+Lemma morphimJ A x : x \in D -> f @* (A :^ x) = f @* A :^ f x.
+Proof.
+move=> Dx; rewrite !conjsgE morphimMl ?(morphimMr, sub1set, groupV) //.
+by rewrite !(morphim_set1, groupV, morphV).
+Qed.
+
+Lemma morphpreJ R x : x \in D -> f @*^-1 (R :^ f x) = f @*^-1 R :^ x.
+Proof.
+move=> Dx; apply/setP=> y; rewrite conjIg !inE conjGid // !mem_conjg inE.
+by case Dy: (y \in D); rewrite // morphJ ?(morphV, groupV).
+Qed.
+
+Lemma morphim_class x A :
+ x \in D -> A \subset D -> f @* (x ^: A) = f x ^: f @* A.
+Proof.
+move=> Dx sAD; rewrite !morphimEsub ?class_subG // /class -!imset_comp.
+by apply: eq_in_imset => y Ay /=; rewrite morphJ // (subsetP sAD).
+Qed.
+
+Lemma classes_morphim A :
+ A \subset D -> classes (f @* A) = [set f @* xA | xA in classes A].
+Proof.
+move=> sAD; rewrite morphimEsub // /classes -!imset_comp.
+apply: eq_in_imset => x /(subsetP sAD) Dx /=.
+by rewrite morphim_class ?morphimEsub.
+Qed.
+
+Lemma morphimT : f @* setT = f @* D.
+Proof. by rewrite -morphimIdom setIT. Qed.
+
+Lemma morphimU A B : f @* (A :|: B) = f @* A :|: f @* B.
+Proof. by rewrite -imsetU -setIUr. Qed.
+
+Lemma morphimI A B : f @* (A :&: B) \subset f @* A :&: f @* B.
+Proof. by rewrite subsetI // ?morphimS ?(subsetIl, subsetIr). Qed.
+
+Lemma morphpre0 : f @*^-1 set0 = set0.
+Proof. by rewrite morphpreE preimset0 setI0. Qed.
+
+Lemma morphpreT : f @*^-1 setT = D.
+Proof. by rewrite morphpreE preimsetT setIT. Qed.
+
+Lemma morphpreU R S : f @*^-1 (R :|: S) = f @*^-1 R :|: f @*^-1 S.
+Proof. by rewrite -setIUr -preimsetU. Qed.
+
+Lemma morphpreI R S : f @*^-1 (R :&: S) = f @*^-1 R :&: f @*^-1 S.
+Proof. by rewrite -setIIr -preimsetI. Qed.
+
+Lemma morphpreD R S : f @*^-1 (R :\: S) = f @*^-1 R :\: f @*^-1 S.
+Proof. by apply/setP=> x; rewrite !inE; case: (x \in D). Qed.
+
+(* kernel, domain properties *)
+
+Lemma kerP x : x \in D -> reflect (f x = 1) (x \in 'ker f).
+Proof. move=> Dx; rewrite 2!inE Dx; exact: set1P. Qed.
+
+Lemma dom_ker : {subset 'ker f <= D}.
+Proof. by move=> x /morphpreP[]. Qed.
+
+Lemma mker x : x \in 'ker f -> f x = 1.
+Proof. by move=> Kx; apply/kerP=> //; exact: dom_ker. Qed.
+
+Lemma mkerl x y : x \in 'ker f -> y \in D -> f (x * y) = f y.
+Proof. by move=> Kx Dy; rewrite morphM // ?(dom_ker, mker Kx, mul1g). Qed.
+
+Lemma mkerr x y : x \in D -> y \in 'ker f -> f (x * y) = f x.
+Proof. by move=> Dx Ky; rewrite morphM // ?(dom_ker, mker Ky, mulg1). Qed.
+
+Lemma rcoset_kerP x y :
+ x \in D -> y \in D -> reflect (f x = f y) (x \in 'ker f :* y).
+Proof.
+move=> Dx Dy; rewrite mem_rcoset !inE groupM ?morphM ?groupV //=.
+rewrite morphV // -eq_mulgV1; exact: eqP.
+Qed.
+
+Lemma ker_rcoset x y :
+ x \in D -> y \in D -> f x = f y -> exists2 z, z \in 'ker f & x = z * y.
+Proof. move=> Dx Dy eqfxy; apply/rcosetP; exact/rcoset_kerP. Qed.
+
+Lemma ker_norm : D \subset 'N('ker f).
+Proof.
+apply/subsetP=> x Dx; rewrite inE; apply/subsetP=> _ /imsetP[y Ky ->].
+by rewrite !inE groupJ ?morphJ // ?dom_ker //= mker ?conj1g.
+Qed.
+
+Lemma ker_normal : 'ker f <| D.
+Proof. by rewrite /(_ <| D) subsetIl ker_norm. Qed.
+
+Lemma morphimGI G A : 'ker f \subset G -> f @* (G :&: A) = f @* G :&: f @* A.
+Proof.
+move=> sKG; apply/eqP; rewrite eqEsubset morphimI setIC.
+apply/subsetP=> _ /setIP[/morphimP[x Dx Ax ->] /morphimP[z Dz Gz]].
+case/ker_rcoset=> {Dz}// y Ky def_x.
+have{z Gz y Ky def_x} Gx: x \in G by rewrite def_x groupMl // (subsetP sKG).
+by rewrite mem_imset ?inE // Dx Gx Ax.
+Qed.
+
+Lemma morphimIG A G : 'ker f \subset G -> f @* (A :&: G) = f @* A :&: f @* G.
+Proof. by move=> sKG; rewrite setIC morphimGI // setIC. Qed.
+
+Lemma morphimD A B : f @* A :\: f @* B \subset f @* (A :\: B).
+Proof.
+rewrite subDset -morphimU morphimS //.
+by rewrite setDE setUIr setUCr setIT subsetUr.
+Qed.
+
+Lemma morphimDG A G : 'ker f \subset G -> f @* (A :\: G) = f @* A :\: f @* G.
+Proof.
+move=> sKG; apply/eqP; rewrite eqEsubset morphimD andbT !setDE subsetI.
+rewrite morphimS ?subsetIl // -[~: f @* G]setU0 -subDset setDE setCK.
+by rewrite -morphimIG //= setIAC -setIA setICr setI0 morphim0.
+Qed.
+
+Lemma morphimD1 A : (f @* A)^# \subset f @* A^#.
+Proof. by rewrite -!set1gE -morphim1 morphimD. Qed.
+
+(* group structure preservation *)
+
+Lemma morphpre_groupset M : group_set (f @*^-1 M).
+Proof.
+apply/group_setP; split=> [|x y]; rewrite !inE ?(morph1, group1) //.
+by case/andP=> Dx Mfx /andP[Dy Mfy]; rewrite morphM ?groupM.
+Qed.
+
+Lemma morphim_groupset G : group_set (f @* G).
+Proof.
+apply/group_setP; split=> [|_ _ /morphimP[x Dx Gx ->] /morphimP[y Dy Gy ->]].
+ by rewrite -morph1 mem_imset ?group1.
+by rewrite -morphM ?mem_imset ?inE ?groupM.
+Qed.
+
+Canonical morphpre_group fPh M :=
+ @group _ (morphpre fPh M) (morphpre_groupset M).
+Canonical morphim_group fPh G := @group _ (morphim fPh G) (morphim_groupset G).
+Canonical ker_group fPh : {group aT} := Eval hnf in [group of ker fPh].
+
+Lemma morph_dom_groupset : group_set (f @: D).
+Proof. by rewrite -morphimEdom groupP. Qed.
+
+Canonical morph_dom_group := group morph_dom_groupset.
+
+Lemma morphpreMr R S :
+ S \subset f @* D -> f @*^-1 (R * S) = f @*^-1 R * f @*^-1 S.
+Proof.
+move=> sSfD; apply: invg_inj.
+by rewrite invMg -!morphpreV invMg morphpreMl // -invSg invgK invGid.
+Qed.
+
+Lemma morphimK A : A \subset D -> f @*^-1 (f @* A) = 'ker f * A.
+Proof.
+move=> sAD; apply/setP=> x; rewrite !inE.
+apply/idP/idP=> [/andP[Dx /morphimP[y Dy Ay eqxy]] | /imset2P[z y Kz Ay ->{x}]].
+ rewrite -(mulgKV y x) mem_mulg // !inE !(groupM, morphM, groupV) //.
+ by rewrite morphV //= eqxy mulgV.
+have [Dy Dz]: y \in D /\ z \in D by rewrite (subsetP sAD) // dom_ker.
+by rewrite groupM // morphM // mker // mul1g mem_imset // inE Dy.
+Qed.
+
+Lemma morphimGK G : 'ker f \subset G -> G \subset D -> f @*^-1 (f @* G) = G.
+Proof. by move=> sKG sGD; rewrite morphimK // mulSGid. Qed.
+
+Lemma morphpre_set1 x : x \in D -> f @*^-1 [set f x] = 'ker f :* x.
+Proof. by move=> Dx; rewrite -morphim_set1 // morphimK ?sub1set. Qed.
+
+Lemma morphpreK R : R \subset f @* D -> f @* (f @*^-1 R) = R.
+Proof.
+move=> sRfD; apply/setP=> y; apply/morphimP/idP=> [[x _] | Ry].
+ by rewrite !inE; case/andP=> _ Rfx ->.
+have /morphimP[x Dx _ defy]: y \in f @* D := subsetP sRfD y Ry.
+by exists x; rewrite // !inE Dx -defy.
+Qed.
+
+Lemma morphim_ker : f @* 'ker f = 1.
+Proof. by rewrite morphpreK ?sub1G. Qed.
+
+Lemma ker_sub_pre M : 'ker f \subset f @*^-1 M.
+Proof. by rewrite morphpreS ?sub1G. Qed.
+
+Lemma ker_normal_pre M : 'ker f <| f @*^-1 M.
+Proof. by rewrite /normal ker_sub_pre subIset ?ker_norm. Qed.
+
+Lemma morphpreSK R S :
+ R \subset f @* D -> (f @*^-1 R \subset f @*^-1 S) = (R \subset S).
+Proof.
+move=> sRfD; apply/idP/idP=> [sf'RS|]; last exact: morphpreS.
+suffices: R \subset f @* D :&: S by rewrite subsetI sRfD.
+rewrite -(morphpreK sRfD) -[_ :&: S]morphpreK (morphimS, subsetIl) //.
+by rewrite morphpreI morphimGK ?subsetIl // setIA setIid.
+Qed.
+
+Lemma sub_morphim_pre A R :
+ A \subset D -> (f @* A \subset R) = (A \subset f @*^-1 R).
+Proof.
+move=> sAD; rewrite -morphpreSK (morphimS, morphimK) //.
+apply/idP/idP; first by apply: subset_trans; exact: mulG_subr.
+by move/(mulgS ('ker f)); rewrite -morphpreMl ?(sub1G, mul1g).
+Qed.
+
+Lemma morphpre_proper R S :
+ R \subset f @* D -> S \subset f @* D ->
+ (f @*^-1 R \proper f @*^-1 S) = (R \proper S).
+Proof. by move=> dQ dR; rewrite /proper !morphpreSK. Qed.
+
+Lemma sub_morphpre_im R G :
+ 'ker f \subset G -> G \subset D -> R \subset f @* D ->
+ (f @*^-1 R \subset G) = (R \subset f @* G).
+Proof. by symmetry; rewrite -morphpreSK ?morphimGK. Qed.
+
+Lemma ker_trivg_morphim A :
+ (A \subset 'ker f) = (A \subset D) && (f @* A \subset [1]).
+Proof.
+case sAD: (A \subset D); first by rewrite sub_morphim_pre.
+by rewrite subsetI sAD.
+Qed.
+
+Lemma morphimSK A B :
+ A \subset D -> (f @* A \subset f @* B) = (A \subset 'ker f * B).
+Proof.
+move=> sAD; transitivity (A \subset 'ker f * (D :&: B)).
+ by rewrite -morphimK ?subsetIl // -sub_morphim_pre // /morphim setIA setIid.
+by rewrite setIC group_modl (subsetIl, subsetI) // andbC sAD.
+Qed.
+
+Lemma morphimSGK A G :
+ A \subset D -> 'ker f \subset G -> (f @* A \subset f @* G) = (A \subset G).
+Proof. by move=> sGD skfK; rewrite morphimSK // mulSGid. Qed.
+
+Lemma ltn_morphim A : [1] \proper 'ker_A f -> #|f @* A| < #|A|.
+Proof.
+case/properP; rewrite sub1set => /setIP[A1 _] [x /setIP[Ax kx] x1].
+rewrite (cardsD1 1 A) A1 ltnS -{1}(setD1K A1) morphimU morphim1.
+rewrite (setUidPr _) ?sub1set; last first.
+ by rewrite -(mker kx) mem_morphim ?(dom_ker kx) // inE x1.
+by rewrite (leq_trans (leq_imset_card _ _)) ?subset_leq_card ?subsetIr.
+Qed.
+
+(* injectivity of image and preimage *)
+
+Lemma morphpre_inj :
+ {in [pred R : {set rT} | R \subset f @* D] &, injective (fun R => f @*^-1 R)}.
+Proof. exact: can_in_inj morphpreK. Qed.
+
+Lemma morphim_injG :
+ {in [pred G : {group aT} | 'ker f \subset G & G \subset D] &,
+ injective (fun G => f @* G)}.
+Proof.
+move=> G H /andP[sKG sGD] /andP[sKH sHD] eqfGH.
+by apply: val_inj; rewrite /= -(morphimGK sKG sGD) eqfGH morphimGK.
+Qed.
+
+Lemma morphim_inj G H :
+ ('ker f \subset G) && (G \subset D) ->
+ ('ker f \subset H) && (H \subset D) ->
+ f @* G = f @* H -> G :=: H.
+Proof. by move=> nsGf nsHf /morphim_injG->. Qed.
+
+(* commutation with generated groups and cycles *)
+
+Lemma morphim_gen A : A \subset D -> f @* <<A>> = <<f @* A>>.
+Proof.
+move=> sAD; apply/eqP.
+rewrite eqEsubset andbC gen_subG morphimS; last exact: subset_gen.
+by rewrite sub_morphim_pre gen_subG // -sub_morphim_pre // subset_gen.
+Qed.
+
+Lemma morphim_cycle x : x \in D -> f @* <[x]> = <[f x]>.
+Proof. by move=> Dx; rewrite morphim_gen (sub1set, morphim_set1). Qed.
+
+Lemma morphimY A B :
+ A \subset D -> B \subset D -> f @* (A <*> B) = f @* A <*> f @* B.
+Proof. by move=> sAD sBD; rewrite morphim_gen ?morphimU // subUset sAD. Qed.
+
+Lemma morphpre_gen R :
+ 1 \in R -> R \subset f @* D -> f @*^-1 <<R>> = <<f @*^-1 R>>.
+Proof.
+move=> R1 sRfD; apply/eqP.
+rewrite eqEsubset andbC gen_subG morphpreS; last exact: subset_gen.
+rewrite -{1}(morphpreK sRfD) -morphim_gen ?subsetIl // morphimGK //=.
+ by rewrite sub_gen // setIS // preimsetS ?sub1set.
+by rewrite gen_subG subsetIl.
+Qed.
+
+(* commutator, normaliser, normal, center properties*)
+
+Lemma morphimR A B :
+ A \subset D -> B \subset D -> f @* [~: A, B] = [~: f @* A, f @* B].
+Proof.
+move/subsetP=> sAD /subsetP sBD.
+rewrite morphim_gen; last first; last congr <<_>>.
+ by apply/subsetP=> _ /imset2P[x y Ax By ->]; rewrite groupR; auto.
+apply/setP=> fz; apply/morphimP/imset2P=> [[z _] | [fx fy]].
+ case/imset2P=> x y Ax By -> -> {z fz}.
+ have Dx := sAD x Ax; have Dy := sBD y By.
+ by exists (f x) (f y); rewrite ?(mem_imset, morphR) // ?(inE, Dx, Dy).
+case/morphimP=> x Dx Ax ->{fx}; case/morphimP=> y Dy By ->{fy} -> {fz}.
+by exists [~ x, y]; rewrite ?(inE, morphR, groupR, mem_imset2).
+Qed.
+
+Lemma morphim_norm A : f @* 'N(A) \subset 'N(f @* A).
+Proof.
+apply/subsetP=> fx; case/morphimP=> x Dx Nx -> {fx}.
+by rewrite inE -morphimJ ?(normP Nx).
+Qed.
+
+Lemma morphim_norms A B : A \subset 'N(B) -> f @* A \subset 'N(f @* B).
+Proof.
+by move=> nBA; apply: subset_trans (morphim_norm B); exact: morphimS.
+Qed.
+
+Lemma morphim_subnorm A B : f @* 'N_A(B) \subset 'N_(f @* A)(f @* B).
+Proof. exact: subset_trans (morphimI A _) (setIS _ (morphim_norm B)). Qed.
+
+Lemma morphim_normal A B : A <| B -> f @* A <| f @* B.
+Proof. by case/andP=> sAB nAB; rewrite /(_ <| _) morphimS // morphim_norms. Qed.
+
+Lemma morphim_cent1 x : x \in D -> f @* 'C[x] \subset 'C[f x].
+Proof. by move=> Dx; rewrite -(morphim_set1 Dx) morphim_norm. Qed.
+
+Lemma morphim_cent1s A x : x \in D -> A \subset 'C[x] -> f @* A \subset 'C[f x].
+Proof.
+by move=> Dx cAx; apply: subset_trans (morphim_cent1 Dx); exact: morphimS.
+Qed.
+
+Lemma morphim_subcent1 A x : x \in D -> f @* 'C_A[x] \subset 'C_(f @* A)[f x].
+Proof. by move=> Dx; rewrite -(morphim_set1 Dx) morphim_subnorm. Qed.
+
+Lemma morphim_cent A : f @* 'C(A) \subset 'C(f @* A).
+Proof.
+apply/bigcapsP=> fx; case/morphimP=> x Dx Ax ->{fx}.
+by apply: subset_trans (morphim_cent1 Dx); apply: morphimS; exact: bigcap_inf.
+Qed.
+
+Lemma morphim_cents A B : A \subset 'C(B) -> f @* A \subset 'C(f @* B).
+Proof.
+by move=> cBA; apply: subset_trans (morphim_cent B); exact: morphimS.
+Qed.
+
+Lemma morphim_subcent A B : f @* 'C_A(B) \subset 'C_(f @* A)(f @* B).
+Proof. exact: subset_trans (morphimI A _) (setIS _ (morphim_cent B)). Qed.
+
+Lemma morphim_abelian A : abelian A -> abelian (f @* A).
+Proof. exact: morphim_cents. Qed.
+
+Lemma morphpre_norm R : f @*^-1 'N(R) \subset 'N(f @*^-1 R).
+Proof.
+apply/subsetP=> x; rewrite !inE => /andP[Dx Nfx].
+by rewrite -morphpreJ ?morphpreS.
+Qed.
+
+Lemma morphpre_norms R S : R \subset 'N(S) -> f @*^-1 R \subset 'N(f @*^-1 S).
+Proof.
+by move=> nSR; apply: subset_trans (morphpre_norm S); exact: morphpreS.
+Qed.
+
+Lemma morphpre_normal R S :
+ R \subset f @* D -> S \subset f @* D -> (f @*^-1 R <| f @*^-1 S) = (R <| S).
+Proof.
+move=> sRfD sSfD; apply/idP/andP=> [|[sRS nSR]].
+ by move/morphim_normal; rewrite !morphpreK //; case/andP.
+by rewrite /(_ <| _) (subset_trans _ (morphpre_norm _)) morphpreS.
+Qed.
+
+Lemma morphpre_subnorm R S : f @*^-1 'N_R(S) \subset 'N_(f @*^-1 R)(f @*^-1 S).
+Proof. by rewrite morphpreI setIS ?morphpre_norm. Qed.
+
+Lemma morphim_normG G :
+ 'ker f \subset G -> G \subset D -> f @* 'N(G) = 'N_(f @* D)(f @* G).
+Proof.
+move=> sKG sGD; apply/eqP; rewrite eqEsubset -{1}morphimIdom morphim_subnorm.
+rewrite -(morphpreK (subsetIl _ _)) morphimS //= morphpreI subIset // orbC.
+by rewrite -{2}(morphimGK sKG sGD) morphpre_norm.
+Qed.
+
+Lemma morphim_subnormG A G :
+ 'ker f \subset G -> G \subset D -> f @* 'N_A(G) = 'N_(f @* A)(f @* G).
+Proof.
+move=> sKB sBD; rewrite morphimIG ?normsG // morphim_normG //.
+by rewrite setICA setIA morphimIim.
+Qed.
+
+Lemma morphpre_cent1 x : x \in D -> 'C_D[x] \subset f @*^-1 'C[f x].
+Proof.
+move=> Dx; rewrite -sub_morphim_pre ?subsetIl //.
+by apply: subset_trans (morphim_cent1 Dx); rewrite morphimS ?subsetIr.
+Qed.
+
+Lemma morphpre_cent1s R x :
+ x \in D -> R \subset f @* D -> f @*^-1 R \subset 'C[x] -> R \subset 'C[f x].
+Proof. by move=> Dx sRfD; move/(morphim_cent1s Dx); rewrite morphpreK. Qed.
+
+Lemma morphpre_subcent1 R x :
+ x \in D -> 'C_(f @*^-1 R)[x] \subset f @*^-1 'C_R[f x].
+Proof.
+move=> Dx; rewrite -morphpreIdom -setIA setICA morphpreI setIS //.
+exact: morphpre_cent1.
+Qed.
+
+Lemma morphpre_cent A : 'C_D(A) \subset f @*^-1 'C(f @* A).
+Proof.
+rewrite -sub_morphim_pre ?subsetIl // morphimGI ?(subsetIl, subIset) // orbC.
+by rewrite (subset_trans (morphim_cent _)).
+Qed.
+
+Lemma morphpre_cents A R :
+ R \subset f @* D -> f @*^-1 R \subset 'C(A) -> R \subset 'C(f @* A).
+Proof. by move=> sRfD; move/morphim_cents; rewrite morphpreK. Qed.
+
+Lemma morphpre_subcent R A : 'C_(f @*^-1 R)(A) \subset f @*^-1 'C_R(f @* A).
+Proof.
+by rewrite -morphpreIdom -setIA setICA morphpreI setIS //; exact: morphpre_cent.
+Qed.
+
+(* local injectivity properties *)
+
+Lemma injmP : reflect {in D &, injective f} ('injm f).
+Proof.
+apply: (iffP subsetP) => [injf x y Dx Dy | injf x /= Kx].
+ by case/ker_rcoset=> // z /injf/set1P->; rewrite mul1g.
+have Dx := dom_ker Kx; apply/set1P/injf => //.
+by apply/rcoset_kerP; rewrite // mulg1.
+Qed.
+
+Lemma card_im_injm : (#|f @* D| == #|D|) = 'injm f.
+Proof. by rewrite morphimEdom (sameP imset_injP injmP). Qed.
+
+Section Injective.
+
+Hypothesis injf : 'injm f.
+
+Lemma ker_injm : 'ker f = 1.
+Proof. exact/trivgP. Qed.
+
+Lemma injmK A : A \subset D -> f @*^-1 (f @* A) = A.
+Proof. by move=> sAD; rewrite morphimK // ker_injm // mul1g. Qed.
+
+Lemma injm_morphim_inj A B :
+ A \subset D -> B \subset D -> f @* A = f @* B -> A = B.
+Proof. by move=> sAD sBD eqAB; rewrite -(injmK sAD) eqAB injmK. Qed.
+
+Lemma card_injm A : A \subset D -> #|f @* A| = #|A|.
+Proof.
+move=> sAD; rewrite morphimEsub // card_in_imset //.
+exact: (sub_in2 (subsetP sAD) (injmP injf)).
+Qed.
+
+Lemma order_injm x : x \in D -> #[f x] = #[x].
+Proof.
+by move=> Dx; rewrite orderE -morphim_cycle // card_injm ?cycle_subG.
+Qed.
+
+Lemma injm1 x : x \in D -> f x = 1 -> x = 1.
+Proof. by move=> Dx; move/(kerP Dx); rewrite ker_injm; move/set1P. Qed.
+
+Lemma morph_injm_eq1 x : x \in D -> (f x == 1) = (x == 1).
+Proof. by move=> Dx; rewrite -morph1 (inj_in_eq (injmP injf)) ?group1. Qed.
+
+Lemma injmSK A B :
+ A \subset D -> (f @* A \subset f @* B) = (A \subset B).
+Proof. by move=> sAD; rewrite morphimSK // ker_injm mul1g. Qed.
+
+Lemma sub_morphpre_injm R A :
+ A \subset D -> R \subset f @* D ->
+ (f @*^-1 R \subset A) = (R \subset f @* A).
+Proof. by move=> sAD sRfD; rewrite -morphpreSK ?injmK. Qed.
+
+Lemma injm_eq A B : A \subset D -> B \subset D -> (f @* A == f @* B) = (A == B).
+Proof. by move=> sAD sBD; rewrite !eqEsubset !injmSK. Qed.
+
+Lemma morphim_injm_eq1 A : A \subset D -> (f @* A == 1) = (A == 1).
+Proof. by move=> sAD; rewrite -morphim1 injm_eq ?sub1G. Qed.
+
+Lemma injmI A B : f @* (A :&: B) = f @* A :&: f @* B.
+Proof.
+rewrite -morphimIdom setIIr -4!(injmK (subsetIl D _), =^~ morphimIdom).
+by rewrite -morphpreI morphpreK // subIset ?morphim_sub.
+Qed.
+
+Lemma injmD1 A : f @* A^# = (f @* A)^#.
+Proof. by have:= morphimDG A injf; rewrite morphim1. Qed.
+
+Lemma nclasses_injm A : A \subset D -> #|classes (f @* A)| = #|classes A|.
+Proof.
+move=> sAD; rewrite classes_morphim // card_in_imset //.
+move=> _ _ /imsetP[x Ax ->] /imsetP[y Ay ->].
+by apply: injm_morphim_inj; rewrite // class_subG ?(subsetP sAD).
+Qed.
+
+Lemma injm_norm A : A \subset D -> f @* 'N(A) = 'N_(f @* D)(f @* A).
+Proof.
+move=> sAD; apply/eqP; rewrite -morphimIdom eqEsubset morphim_subnorm.
+rewrite -sub_morphpre_injm ?subsetIl // morphpreI injmK // setIS //.
+by rewrite -{2}(injmK sAD) morphpre_norm.
+Qed.
+
+Lemma injm_norms A B :
+ A \subset D -> B \subset D -> (f @* A \subset 'N(f @* B)) = (A \subset 'N(B)).
+Proof. by move=> sAD sBD; rewrite -injmSK // injm_norm // subsetI morphimS. Qed.
+
+Lemma injm_normal A B :
+ A \subset D -> B \subset D -> (f @* A <| f @* B) = (A <| B).
+Proof. by move=> sAD sBD; rewrite /normal injmSK ?injm_norms. Qed.
+
+Lemma injm_subnorm A B : B \subset D -> f @* 'N_A(B) = 'N_(f @* A)(f @* B).
+Proof. by move=> sBD; rewrite injmI injm_norm // setICA setIA morphimIim. Qed.
+
+Lemma injm_cent1 x : x \in D -> f @* 'C[x] = 'C_(f @* D)[f x].
+Proof. by move=> Dx; rewrite injm_norm ?morphim_set1 ?sub1set. Qed.
+
+Lemma injm_subcent1 A x : x \in D -> f @* 'C_A[x] = 'C_(f @* A)[f x].
+Proof. by move=> Dx; rewrite injm_subnorm ?morphim_set1 ?sub1set. Qed.
+
+Lemma injm_cent A : A \subset D -> f @* 'C(A) = 'C_(f @* D)(f @* A).
+Proof.
+move=> sAD; apply/eqP; rewrite -morphimIdom eqEsubset morphim_subcent.
+apply/subsetP=> fx; case/setIP; case/morphimP=> x Dx _ ->{fx} cAfx.
+rewrite mem_morphim // inE Dx -sub1set centsC cent_set1 -injmSK //.
+by rewrite injm_cent1 // subsetI morphimS // -cent_set1 centsC sub1set.
+Qed.
+
+Lemma injm_cents A B :
+ A \subset D -> B \subset D -> (f @* A \subset 'C(f @* B)) = (A \subset 'C(B)).
+Proof. by move=> sAD sBD; rewrite -injmSK // injm_cent // subsetI morphimS. Qed.
+
+Lemma injm_subcent A B : B \subset D -> f @* 'C_A(B) = 'C_(f @* A)(f @* B).
+Proof. by move=> sBD; rewrite injmI injm_cent // setICA setIA morphimIim. Qed.
+
+Lemma injm_abelian A : A \subset D -> abelian (f @* A) = abelian A.
+Proof.
+by move=> sAD; rewrite /abelian -subsetIidl -injm_subcent // injmSK ?subsetIidl.
+Qed.
+
+End Injective.
+
+Lemma eq_morphim (g : {morphism D >-> rT}):
+ {in D, f =1 g} -> forall A, f @* A = g @* A.
+Proof.
+by move=> efg A; apply: eq_in_imset; apply: sub_in1 efg => x /setIP[].
+Qed.
+
+Lemma eq_in_morphim B A (g : {morphism B >-> rT}) :
+ D :&: A = B :&: A -> {in A, f =1 g} -> f @* A = g @* A.
+Proof.
+move=> eqDBA eqAfg; rewrite /morphim /= eqDBA.
+by apply: eq_in_imset => x /setIP[_]/eqAfg.
+Qed.
+
+End MorphismTheory.
+
+Notation "''ker' f" := (ker_group (MorPhantom f)) : Group_scope.
+Notation "''ker_' G f" := (G :&: 'ker f)%G : Group_scope.
+Notation "f @* G" := (morphim_group (MorPhantom f) G) : Group_scope.
+Notation "f @*^-1 M" := (morphpre_group (MorPhantom f) M) : Group_scope.
+Notation "f @: D" := (morph_dom_group f D) : Group_scope.
+
+Implicit Arguments injmP [aT rT D f].
+
+Section IdentityMorphism.
+
+Variable gT : finGroupType.
+Implicit Types A B : {set gT}.
+Implicit Type G : {group gT}.
+
+Definition idm of {set gT} := fun x : gT => x : FinGroup.sort gT.
+
+Lemma idm_morphM A : {in A & , {morph idm A : x y / x * y}}.
+Proof. by []. Qed.
+
+Canonical idm_morphism A := Morphism (@idm_morphM A).
+
+Lemma injm_idm G : 'injm (idm G).
+Proof. by apply/injmP=> x y _ _. Qed.
+
+Lemma ker_idm G : 'ker (idm G) = 1.
+Proof. by apply/trivgP; exact: injm_idm. Qed.
+
+Lemma morphim_idm A B : B \subset A -> idm A @* B = B.
+Proof.
+rewrite /morphim /= /idm => /setIidPr->.
+by apply/setP=> x; apply/imsetP/idP=> [[y By ->]|Bx]; last exists x.
+Qed.
+
+Lemma morphpre_idm A B : idm A @*^-1 B = A :&: B.
+Proof. by apply/setP=> x; rewrite !inE. Qed.
+
+Lemma im_idm A : idm A @* A = A.
+Proof. exact: morphim_idm. Qed.
+
+End IdentityMorphism.
+
+Arguments Scope idm [_ group_scope group_scope].
+Prenex Implicits idm.
+
+Section RestrictedMorphism.
+
+Variables aT rT : finGroupType.
+Variables A D : {set aT}.
+Implicit Type B : {set aT}.
+Implicit Type R : {set rT}.
+
+Definition restrm of A \subset D := @id (aT -> FinGroup.sort rT).
+
+Section Props.
+
+Hypothesis sAD : A \subset D.
+Variable f : {morphism D >-> rT}.
+Local Notation fA := (restrm sAD (mfun f)).
+
+Canonical restrm_morphism :=
+ @Morphism aT rT A fA (sub_in2 (subsetP sAD) (morphM f)).
+
+Lemma morphim_restrm B : fA @* B = f @* (A :&: B).
+Proof. by rewrite {2}/morphim setIA (setIidPr sAD). Qed.
+
+Lemma restrmEsub B : B \subset A -> fA @* B = f @* B.
+Proof. by rewrite morphim_restrm => /setIidPr->. Qed.
+
+Lemma im_restrm : fA @* A = f @* A.
+Proof. exact: restrmEsub. Qed.
+
+Lemma morphpre_restrm R : fA @*^-1 R = A :&: f @*^-1 R.
+Proof. by rewrite setIA (setIidPl sAD). Qed.
+
+Lemma ker_restrm : 'ker fA = 'ker_A f.
+Proof. exact: morphpre_restrm. Qed.
+
+Lemma injm_restrm : 'injm f -> 'injm fA.
+Proof. by apply: subset_trans; rewrite ker_restrm subsetIr. Qed.
+
+End Props.
+
+Lemma restrmP (f : {morphism D >-> rT}) : A \subset 'dom f ->
+ {g : {morphism A >-> rT} | [/\ g = f :> (aT -> rT), 'ker g = 'ker_A f,
+ forall R, g @*^-1 R = A :&: f @*^-1 R
+ & forall B, B \subset A -> g @* B = f @* B]}.
+Proof.
+move=> sAD; exists (restrm_morphism sAD f).
+split=> // [|R|B sBA]; first 1 [exact: ker_restrm | exact: morphpre_restrm].
+by rewrite morphim_restrm (setIidPr sBA).
+Qed.
+
+Lemma domP (f : {morphism D >-> rT}) : 'dom f = A ->
+ {g : {morphism A >-> rT} | [/\ g = f :> (aT -> rT), 'ker g = 'ker f,
+ forall R, g @*^-1 R = f @*^-1 R
+ & forall B, g @* B = f @* B]}.
+Proof. by move <-; exists f. Qed.
+
+End RestrictedMorphism.
+
+Arguments Scope restrm [_ _ group_scope group_scope _ group_scope].
+Prenex Implicits restrm.
+Implicit Arguments restrmP [aT rT D A].
+Implicit Arguments domP [aT rT D A].
+
+Section TrivMorphism.
+
+Variables aT rT : finGroupType.
+
+Definition trivm of {set aT} & aT := 1 : FinGroup.sort rT.
+
+Lemma trivm_morphM (A : {set aT}) : {in A &, {morph trivm A : x y / x * y}}.
+Proof. by move=> x y /=; rewrite mulg1. Qed.
+
+Canonical triv_morph A := Morphism (@trivm_morphM A).
+
+Lemma morphim_trivm (G H : {group aT}) : trivm G @* H = 1.
+Proof.
+apply/setP=> /= y; rewrite inE; apply/idP/eqP=> [|->]; first by case/morphimP.
+by apply/morphimP; exists (1 : aT); rewrite /= ?group1.
+Qed.
+
+Lemma ker_trivm (G : {group aT}) : 'ker (trivm G) = G.
+Proof. by apply/setIidPl/subsetP=> x _; rewrite !inE /=. Qed.
+
+End TrivMorphism.
+
+Arguments Scope trivm [_ _ group_scope group_scope].
+Implicit Arguments trivm [[aT] [rT]].
+
+(* The composition of two morphisms is a Canonical morphism instance. *)
+Section MorphismComposition.
+
+Variables gT hT rT : finGroupType.
+Variables (G : {group gT}) (H : {group hT}).
+
+Variable f : {morphism G >-> hT}.
+Variable g : {morphism H >-> rT}.
+
+Notation Local gof := (mfun g \o mfun f).
+
+Lemma comp_morphM : {in f @*^-1 H &, {morph gof: x y / x * y}}.
+Proof.
+by move=> x y; rewrite /= !inE => /andP[? ?] /andP[? ?]; rewrite !morphM.
+Qed.
+
+Canonical comp_morphism := Morphism comp_morphM.
+
+Lemma ker_comp : 'ker gof = f @*^-1 'ker g.
+Proof. by apply/setP=> x; rewrite !inE andbA. Qed.
+
+Lemma injm_comp : 'injm f -> 'injm g -> 'injm gof.
+Proof. by move=> injf; rewrite ker_comp; move/trivgP=> ->. Qed.
+
+Lemma morphim_comp (A : {set gT}) : gof @* A = g @* (f @* A).
+Proof.
+apply/setP=> z; apply/morphimP/morphimP=> [[x]|[y Hy fAy ->{z}]].
+ rewrite !inE => /andP[Gx Hfx]; exists (f x) => //.
+ by apply/morphimP; exists x.
+by case/morphimP: fAy Hy => x Gx Ax ->{y} Hfx; exists x; rewrite ?inE ?Gx.
+Qed.
+
+Lemma morphpre_comp (C : {set rT}) : gof @*^-1 C = f @*^-1 (g @*^-1 C).
+Proof. by apply/setP=> z; rewrite !inE andbA. Qed.
+
+End MorphismComposition.
+
+(* The factor morphism *)
+Section FactorMorphism.
+
+Variables aT qT rT : finGroupType.
+
+Variables G H : {group aT}.
+Variable f : {morphism G >-> rT}.
+Variable q : {morphism H >-> qT}.
+
+Definition factm of 'ker q \subset 'ker f & G \subset H :=
+ fun x => f (repr (q @*^-1 [set x])).
+
+Hypothesis sKqKf : 'ker q \subset 'ker f.
+Hypothesis sGH : G \subset H.
+
+Notation ff := (factm sKqKf sGH).
+
+Lemma factmE x : x \in G -> ff (q x) = f x.
+Proof.
+rewrite /ff => Gx; have Hx := subsetP sGH x Gx.
+have /mem_repr: x \in q @*^-1 [set q x] by rewrite !inE Hx /=.
+case/morphpreP; move: (repr _) => y Hy /set1P.
+by case/ker_rcoset=> // z Kz ->; rewrite mkerl ?(subsetP sKqKf).
+Qed.
+
+Lemma factm_morphM : {in q @* G &, {morph ff : x y / x * y}}.
+Proof.
+move=> _ _ /morphimP[x Hx Gx ->] /morphimP[y Hy Gy ->].
+by rewrite -morphM ?factmE ?groupM // morphM.
+Qed.
+
+Canonical factm_morphism := Morphism factm_morphM.
+
+Lemma morphim_factm (A : {set aT}) : ff @* (q @* A) = f @* A.
+Proof.
+rewrite -morphim_comp /= {1}/morphim /= morphimGK //; last first.
+ by rewrite (subset_trans sKqKf) ?subsetIl.
+apply/setP=> y; apply/morphimP/morphimP;
+ by case=> x Gx Ax ->{y}; exists x; rewrite //= factmE.
+Qed.
+
+Lemma morphpre_factm (C : {set rT}) : ff @*^-1 C = q @* (f @*^-1 C).
+Proof.
+apply/setP=> y; rewrite !inE /=; apply/andP/morphimP=> [[]|[x Hx]]; last first.
+ by case/morphpreP=> Gx Cfx ->; rewrite factmE ?mem_imset ?inE ?Hx.
+case/morphimP=> x Hx Gx ->; rewrite factmE //.
+by exists x; rewrite // !inE Gx.
+Qed.
+
+Lemma ker_factm : 'ker ff = q @* 'ker f.
+Proof. exact: morphpre_factm. Qed.
+
+Lemma injm_factm : 'injm f -> 'injm ff.
+Proof. by rewrite ker_factm => /trivgP->; rewrite morphim1. Qed.
+
+Lemma injm_factmP : reflect ('ker f = 'ker q) ('injm ff).
+Proof.
+rewrite ker_factm -morphimIdom sub_morphim_pre ?subsetIl //.
+rewrite setIA (setIidPr sGH) (sameP setIidPr eqP) (setIidPl _) // eq_sym.
+exact: eqP.
+Qed.
+
+Lemma ker_factm_loc (K : {group aT}) : 'ker_(q @* K) ff = q @* 'ker_K f.
+Proof. by rewrite ker_factm -morphimIG. Qed.
+
+End FactorMorphism.
+
+Prenex Implicits factm.
+
+Section InverseMorphism.
+
+Variables aT rT : finGroupType.
+Implicit Types A B : {set aT}.
+Implicit Types C D : {set rT}.
+Variables (G : {group aT}) (f : {morphism G >-> rT}).
+Hypothesis injf : 'injm f.
+
+Lemma invm_subker : 'ker f \subset 'ker (idm G).
+Proof. by rewrite ker_idm. Qed.
+
+Definition invm := factm invm_subker (subxx _).
+
+Canonical invm_morphism := Eval hnf in [morphism of invm].
+
+Lemma invmE : {in G, cancel f invm}.
+Proof. exact: factmE. Qed.
+
+Lemma invmK : {in f @* G, cancel invm f}.
+Proof. by move=> fx; case/morphimP=> x _ Gx ->; rewrite invmE. Qed.
+
+Lemma morphpre_invm A : invm @*^-1 A = f @* A.
+Proof. by rewrite morphpre_factm morphpre_idm morphimIdom. Qed.
+
+Lemma morphim_invm A : A \subset G -> invm @* (f @* A) = A.
+Proof. by move=> sAG; rewrite morphim_factm morphim_idm. Qed.
+
+Lemma morphim_invmE C : invm @* C = f @*^-1 C.
+Proof.
+rewrite -morphpreIdom -(morphim_invm (subsetIl _ _)).
+by rewrite morphimIdom -morphpreIim morphpreK (subsetIl, morphimIdom).
+Qed.
+
+Lemma injm_proper A B :
+ A \subset G -> B \subset G -> (f @* A \proper f @* B) = (A \proper B).
+Proof.
+move=> dA dB; rewrite -morphpre_invm -(morphpre_invm B).
+by rewrite morphpre_proper ?morphim_invm.
+Qed.
+
+Lemma injm_invm : 'injm invm.
+Proof. by move/can_in_inj/injmP: invmK. Qed.
+
+Lemma ker_invm : 'ker invm = 1.
+Proof. by move/trivgP: injm_invm. Qed.
+
+Lemma im_invm : invm @* (f @* G) = G.
+Proof. exact: morphim_invm. Qed.
+
+End InverseMorphism.
+
+Prenex Implicits invm.
+
+Section InjFactm.
+
+Variables (gT aT rT : finGroupType) (D G : {group gT}).
+Variables (g : {morphism G >-> rT}) (f : {morphism D >-> aT}) (injf : 'injm f).
+
+Definition ifactm :=
+ tag (domP [morphism of g \o invm injf] (morphpre_invm injf G)).
+
+Lemma ifactmE : {in D, forall x, ifactm (f x) = g x}.
+Proof.
+rewrite /ifactm => x Dx; case: domP => f' /= [def_f' _ _ _].
+by rewrite {f'}def_f' //= invmE.
+Qed.
+
+Lemma morphim_ifactm (A : {set gT}) :
+ A \subset D -> ifactm @* (f @* A) = g @* A.
+Proof.
+rewrite /ifactm => sAD; case: domP => _ /= [_ _ _ ->].
+by rewrite morphim_comp morphim_invm.
+Qed.
+
+Lemma im_ifactm : G \subset D -> ifactm @* (f @* G) = g @* G.
+Proof. exact: morphim_ifactm. Qed.
+
+Lemma morphpre_ifactm C : ifactm @*^-1 C = f @* (g @*^-1 C).
+Proof.
+rewrite /ifactm; case: domP => _ /= [_ _ -> _].
+by rewrite morphpre_comp morphpre_invm.
+Qed.
+
+Lemma ker_ifactm : 'ker ifactm = f @* 'ker g.
+Proof. exact: morphpre_ifactm. Qed.
+
+Lemma injm_ifactm : 'injm g -> 'injm ifactm.
+Proof. by rewrite ker_ifactm => /trivgP->; rewrite morphim1. Qed.
+
+End InjFactm.
+
+(* Reflected (boolean) form of morphism and isomorphism properties *)
+
+Section ReflectProp.
+
+Variables aT rT : finGroupType.
+
+Section Defs.
+
+Variables (A : {set aT}) (B : {set rT}).
+
+(* morphic is the morphM property of morphisms seen through morphicP *)
+Definition morphic (f : aT -> rT) :=
+ [forall u in [predX A & A], f (u.1 * u.2) == f u.1 * f u.2].
+
+Definition isom f := f @: A^# == B^#.
+
+Definition misom f := morphic f && isom f.
+
+Definition isog := [exists f : {ffun aT -> rT}, misom f].
+
+Section MorphicProps.
+
+Variable f : aT -> rT.
+
+Lemma morphicP : reflect {in A &, {morph f : x y / x * y}} (morphic f).
+Proof.
+apply: (iffP forallP) => [fM x y Ax Ay | fM [x y] /=].
+ by apply/eqP; have:= fM (x, y); rewrite inE /= Ax Ay.
+by apply/implyP=> /andP[Ax Ay]; rewrite fM.
+Qed.
+
+Definition morphm of morphic f := f : aT -> FinGroup.sort rT.
+
+Lemma morphmE fM : morphm fM = f. Proof. by []. Qed.
+
+Canonical morphm_morphism fM := @Morphism _ _ A (morphm fM) (morphicP fM).
+
+End MorphicProps.
+
+Lemma misomP f : reflect {fM : morphic f & isom (morphm fM)} (misom f).
+Proof. by apply: (iffP andP) => [] [fM fiso] //; exists fM. Qed.
+
+Lemma misom_isog f : misom f -> isog.
+Proof.
+case/andP=> fM iso_f; apply/existsP; exists (finfun f).
+apply/andP; split; last by rewrite /misom /isom !(eq_imset _ (ffunE f)).
+apply/forallP=> u; rewrite !ffunE; exact: forallP fM u.
+Qed.
+
+Lemma isom_isog (D : {group aT}) (f : {morphism D >-> rT}) :
+ A \subset D -> isom f -> isog.
+Proof.
+move=> sAD isof; apply: (@misom_isog f); rewrite /misom isof andbT.
+apply/morphicP; exact: (sub_in2 (subsetP sAD) (morphM f)).
+Qed.
+
+Lemma isog_isom : isog -> {f : {morphism A >-> rT} | isom f}.
+Proof.
+by case/existsP/sigW=> f /misomP[fM isom_f]; exists (morphm_morphism fM).
+Qed.
+
+End Defs.
+
+Infix "\isog" := isog.
+
+Implicit Arguments isom_isog [A B D].
+
+(* The real reflection properties only hold for true groups and morphisms. *)
+
+Section Main.
+
+Variables (G : {group aT}) (H : {group rT}).
+
+Lemma isomP (f : {morphism G >-> rT}) :
+ reflect ('injm f /\ f @* G = H) (isom G H f).
+Proof.
+apply: (iffP eqP) => [eqfGH | [injf <-]]; last first.
+ by rewrite -injmD1 // morphimEsub ?subsetDl.
+split.
+ apply/subsetP=> x /morphpreP[Gx fx1]; have: f x \notin H^# by rewrite inE fx1.
+ by apply: contraR => ntx; rewrite -eqfGH mem_imset // inE ntx.
+rewrite morphimEdom -{2}(setD1K (group1 G)) imsetU eqfGH.
+by rewrite imset_set1 morph1 setD1K.
+Qed.
+
+Lemma isogP :
+ reflect (exists2 f : {morphism G >-> rT}, 'injm f & f @* G = H) (G \isog H).
+Proof.
+apply: (iffP idP) => [/isog_isom[f /isomP[]] | [f injf fG]]; first by exists f.
+by apply: (isom_isog f) => //; apply/isomP.
+Qed.
+
+Variable f : {morphism G >-> rT}.
+Hypothesis isoGH : isom G H f.
+
+Lemma isom_inj : 'injm f. Proof. by have /isomP[] := isoGH. Qed.
+Lemma isom_im : f @* G = H. Proof. by have /isomP[] := isoGH. Qed.
+Lemma isom_card : #|G| = #|H|.
+Proof. by rewrite -isom_im card_injm ?isom_inj. Qed.
+Lemma isom_sub_im : H \subset f @* G. Proof. by rewrite isom_im. Qed.
+Definition isom_inv := restrm isom_sub_im (invm isom_inj).
+
+End Main.
+
+Variables (G : {group aT}) (f : {morphism G >-> rT}).
+
+Lemma morphim_isom (H : {group aT}) (K : {group rT}) :
+ H \subset G -> isom H K f -> f @* H = K.
+Proof. by case/(restrmP f)=> g [gf _ _ <- //]; rewrite -gf; case/isomP. Qed.
+
+Lemma sub_isom (A : {set aT}) (C : {set rT}) :
+ A \subset G -> f @* A = C -> 'injm f -> isom A C f.
+Proof.
+move=> sAG; case: (restrmP f sAG) => g [_ _ _ img] <-{C} injf.
+rewrite /isom -morphimEsub ?morphimDG ?morphim1 //.
+by rewrite subDset setUC subsetU ?sAG.
+Qed.
+
+Lemma sub_isog (A : {set aT}) : A \subset G -> 'injm f -> isog A (f @* A).
+Proof. by move=> sAG injf; apply: (isom_isog f sAG); exact: sub_isom. Qed.
+
+Lemma restr_isom_to (A : {set aT}) (C R : {group rT}) (sAG : A \subset G) :
+ f @* A = C -> isom G R f -> isom A C (restrm sAG f).
+Proof. by move=> defC /isomP[inj_f _]; apply: sub_isom. Qed.
+
+Lemma restr_isom (A : {group aT}) (R : {group rT}) (sAG : A \subset G) :
+ isom G R f -> isom A (f @* A) (restrm sAG f).
+Proof. exact: restr_isom_to. Qed.
+
+End ReflectProp.
+
+Arguments Scope isom [_ _ group_scope group_scope _].
+Arguments Scope morphic [_ _ group_scope _].
+Arguments Scope misom [_ _ group_scope group_scope _].
+Arguments Scope isog [_ _ group_scope group_scope].
+
+Implicit Arguments morphicP [aT rT A f].
+Implicit Arguments misomP [aT rT A B f].
+Implicit Arguments isom_isog [aT rT A B D].
+Implicit Arguments isomP [aT rT G H f].
+Implicit Arguments isogP [aT rT G H].
+Prenex Implicits morphic morphicP morphm isom isog isomP misomP isogP.
+Notation "x \isog y":= (isog x y).
+
+Section Isomorphisms.
+
+Variables gT hT kT : finGroupType.
+Variables (G : {group gT}) (H : {group hT}) (K : {group kT}).
+
+Lemma idm_isom : isom G G (idm G).
+Proof. exact: sub_isom (im_idm G) (injm_idm G). Qed.
+
+Lemma isog_refl : G \isog G. Proof. exact: isom_isog idm_isom. Qed.
+
+Lemma card_isog : G \isog H -> #|G| = #|H|.
+Proof. case/isogP=> f injf <-; apply: isom_card (f) _; exact/isomP. Qed.
+
+Lemma isog_abelian : G \isog H -> abelian G = abelian H.
+Proof. by case/isogP=> f injf <-; rewrite injm_abelian. Qed.
+
+Lemma trivial_isog : G :=: 1 -> H :=: 1 -> G \isog H.
+Proof.
+move=> -> ->; apply/isogP.
+exists [morphism of @trivm gT hT 1]; rewrite /= ?morphim1 //.
+rewrite ker_trivm; exact: subxx.
+Qed.
+
+Lemma isog_eq1 : G \isog H -> (G :==: 1) = (H :==: 1).
+Proof. by move=> isoGH; rewrite !trivg_card1 card_isog. Qed.
+
+Lemma isom_sym (f : {morphism G >-> hT}) (isoGH : isom G H f) :
+ isom H G (isom_inv isoGH).
+Proof.
+rewrite sub_isom 1?injm_restrm ?injm_invm // im_restrm.
+by rewrite -(isom_im isoGH) im_invm.
+Qed.
+
+Lemma isog_symr : G \isog H -> H \isog G.
+Proof. by case/isog_isom=> f /isom_sym/isom_isog->. Qed.
+
+Lemma isog_trans : G \isog H -> H \isog K -> G \isog K.
+Proof.
+case/isogP=> f injf <-; case/isogP=> g injg <-.
+have defG: f @*^-1 (f @* G) = G by rewrite morphimGK ?subsetIl.
+rewrite -morphim_comp -{1 8}defG.
+by apply/isogP; exists [morphism of g \o f]; rewrite ?injm_comp.
+Qed.
+
+Lemma nclasses_isog : G \isog H -> #|classes G| = #|classes H|.
+Proof. by case/isogP=> f injf <-; rewrite nclasses_injm. Qed.
+
+End Isomorphisms.
+
+Section IsoBoolEquiv.
+
+Variables gT hT kT : finGroupType.
+Variables (G : {group gT}) (H : {group hT}) (K : {group kT}).
+
+Lemma isog_sym : (G \isog H) = (H \isog G).
+Proof. apply/idP/idP; exact: isog_symr. Qed.
+
+Lemma isog_transl : G \isog H -> (G \isog K) = (H \isog K).
+Proof.
+by move=> iso; apply/idP/idP; apply: isog_trans; rewrite // -isog_sym.
+Qed.
+
+Lemma isog_transr : G \isog H -> (K \isog G) = (K \isog H).
+Proof.
+by move=> iso; apply/idP/idP; move/isog_trans; apply; rewrite // -isog_sym.
+Qed.
+
+End IsoBoolEquiv.
+
+Section Homg.
+
+Implicit Types rT gT aT : finGroupType.
+
+Definition homg rT aT (C : {set rT}) (D : {set aT}) :=
+ [exists (f : {ffun aT -> rT} | morphic D f), f @: D == C].
+
+Lemma homgP rT aT (C : {set rT}) (D : {set aT}) :
+ reflect (exists f : {morphism D >-> rT}, f @* D = C) (homg C D).
+Proof.
+apply: (iffP exists_eq_inP) => [[f fM <-] | [f <-]].
+ by exists (morphm_morphism fM); rewrite /morphim /= setIid.
+exists (finfun f); first by apply/morphicP=> x y Dx Dy; rewrite !ffunE morphM.
+by rewrite /morphim setIid; apply: eq_imset => x; rewrite ffunE.
+Qed.
+
+Lemma morphim_homg aT rT (A D : {set aT}) (f : {morphism D >-> rT}) :
+ A \subset D -> homg (f @* A) A.
+Proof.
+move=> sAD; apply/homgP; exists (restrm_morphism sAD f).
+by rewrite morphim_restrm setIid.
+Qed.
+
+Lemma leq_homg rT aT (C : {set rT}) (G : {group aT}) :
+ homg C G -> #|C| <= #|G|.
+Proof. by case/homgP=> f <-; apply: leq_morphim. Qed.
+
+Lemma homg_refl aT (A : {set aT}) : homg A A.
+Proof. by apply/homgP; exists (idm_morphism A); rewrite im_idm. Qed.
+
+Lemma homg_trans aT (B : {set aT}) rT (C : {set rT}) gT (G : {group gT}) :
+ homg C B -> homg B G -> homg C G.
+Proof.
+move=> homCB homBG; case/homgP: homBG homCB => fG <- /homgP[fK <-].
+by rewrite -morphim_comp morphim_homg // -sub_morphim_pre.
+Qed.
+
+Lemma isogEcard rT aT (G : {group rT}) (H : {group aT}) :
+ (G \isog H) = (homg G H) && (#|H| <= #|G|).
+Proof.
+rewrite isog_sym; apply/isogP/andP=> [[f injf <-] | []].
+ by rewrite leq_eqVlt eq_sym card_im_injm injf morphim_homg.
+case/homgP=> f <-; rewrite leq_eqVlt eq_sym card_im_injm.
+by rewrite ltnNge leq_morphim orbF; exists f.
+Qed.
+
+Lemma isog_hom rT aT (G : {group rT}) (H : {group aT}) : G \isog H -> homg G H.
+Proof. by rewrite isogEcard; case/andP. Qed.
+
+Lemma isogEhom rT aT (G : {group rT}) (H : {group aT}) :
+ (G \isog H) = homg G H && homg H G.
+Proof.
+apply/idP/andP=> [isoGH | [homGH homHG]].
+ by rewrite !isog_hom // isog_sym.
+by rewrite isogEcard homGH leq_homg.
+Qed.
+
+Lemma eq_homgl gT aT rT (G : {group gT}) (H : {group aT}) (K : {group rT}) :
+ G \isog H -> homg G K = homg H K.
+Proof.
+by rewrite isogEhom => /andP[homGH homHG]; apply/idP/idP; exact: homg_trans.
+Qed.
+
+Lemma eq_homgr gT rT aT (G : {group gT}) (H : {group rT}) (K : {group aT}) :
+ G \isog H -> homg K G = homg K H.
+Proof.
+rewrite isogEhom => /andP[homGH homHG].
+by apply/idP/idP=> homK; exact: homg_trans homK _.
+Qed.
+
+End Homg.
+
+Arguments Scope homg [_ _ group_scope group_scope].
+Notation "G \homg H" := (homg G H)
+ (at level 70, no associativity) : group_scope.
+
+Implicit Arguments homgP [rT aT C D].
+
+(* Isomorphism between a group and its subtype. *)
+
+Section SubMorphism.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Canonical sgval_morphism := Morphism (@sgvalM _ G).
+Canonical subg_morphism := Morphism (@subgM _ G).
+
+Lemma injm_sgval : 'injm sgval.
+Proof. apply/injmP; apply: in2W; exact: subg_inj. Qed.
+
+Lemma injm_subg : 'injm (subg G).
+Proof. apply/injmP; exact: can_in_inj (@subgK _ _). Qed.
+Hint Resolve injm_sgval injm_subg.
+
+Lemma ker_sgval : 'ker sgval = 1. Proof. exact/trivgP. Qed.
+Lemma ker_subg : 'ker (subg G) = 1. Proof. exact/trivgP. Qed.
+
+Lemma im_subg : subg G @* G = [subg G].
+Proof.
+apply/eqP; rewrite -subTset morphimEdom.
+by apply/subsetP=> u _; rewrite -(sgvalK u) mem_imset ?subgP.
+Qed.
+
+Lemma sgval_sub A : sgval @* A \subset G.
+Proof. apply/subsetP=> x; case/imsetP=> u _ ->; exact: subgP. Qed.
+
+Lemma sgvalmK A : subg G @* (sgval @* A) = A.
+Proof.
+apply/eqP; rewrite eqEcard !card_injm ?subsetT ?sgval_sub // leqnn andbT.
+rewrite -morphim_comp; apply/subsetP=> _ /morphimP[v _ Av ->] /=.
+by rewrite sgvalK.
+Qed.
+
+Lemma subgmK (A : {set gT}) : A \subset G -> sgval @* (subg G @* A) = A.
+Proof.
+move=> sAG; apply/eqP; rewrite eqEcard !card_injm ?subsetT //.
+rewrite leqnn andbT -morphim_comp morphimE /= morphpreT.
+by apply/subsetP=> _ /morphimP[v Gv Av ->] /=; rewrite subgK.
+Qed.
+
+Lemma im_sgval : sgval @* [subg G] = G.
+Proof. by rewrite -{2}im_subg subgmK. Qed.
+
+Lemma isom_subg : isom G [subg G] (subg G).
+Proof. by apply/isomP; rewrite im_subg. Qed.
+
+Lemma isom_sgval : isom [subg G] G sgval.
+Proof. by apply/isomP; rewrite im_sgval. Qed.
+
+Lemma isog_subg : isog G [subg G].
+Proof. exact: isom_isog isom_subg. Qed.
+
+End SubMorphism.
+
diff --git a/mathcomp/fingroup/perm.v b/mathcomp/fingroup/perm.v
new file mode 100644
index 0000000..9c31176
--- /dev/null
+++ b/mathcomp/fingroup/perm.v
@@ -0,0 +1,576 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq path choice fintype.
+Require Import tuple finfun bigop finset binomial fingroup.
+
+(******************************************************************************)
+(* This file contains the definition and properties associated to the group *)
+(* of permutations of an arbitrary finite type. *)
+(* {perm T} == the type of permutations of a finite type T, i.e., *)
+(* injective (finite) functions from T to T. Permutations *)
+(* coerce to CiC functions. *)
+(* 'S_n == the set of all permutations of 'I_n, i.e., of {0,.., n-1} *)
+(* perm_on A u == u is a permutation with support A, i.e., u only displaces *)
+(* elements of A (u x != x implies x \in A). *)
+(* tperm x y == the transposition of x, y *)
+(* aperm x s == the image of x under the action of the permutation s *)
+(* := s x *)
+(* pcycle s x == the set of all elements that are in the same cycle of the *)
+(* permutation s as x, i.e., {x, s x, (s ^+ 2) x, ...} *)
+(* pcycles s == the set of all the cycles of the permutation s *)
+(* (s : bool) == s is an odd permutation (the coercion is called odd_perm) *)
+(* dpair u == u is a pair (x, y) of distinct objects (i.e., x != y) *)
+(* lift_perm i j s == the permutation obtained by lifting s : 'S_n.-1 over *)
+(* (i |-> j), that maps i to j and lift i k to lift j (s k). *)
+(* Canonical structures are defined allowing permutations to be an eqType, *)
+(* choiceType, countType, finType, subType, finGroupType; permutations with *)
+(* composition form a group, therefore inherit all generic group notations: *)
+(* 1 == identity permutation, * == composition, ^-1 == inverse permutation. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section PermDefSection.
+
+Variable T : finType.
+
+Inductive perm_type : predArgType :=
+ Perm (pval : {ffun T -> T}) & injectiveb pval.
+Definition pval p := let: Perm f _ := p in f.
+Definition perm_of of phant T := perm_type.
+Identity Coercion type_of_perm : perm_of >-> perm_type.
+
+Notation pT := (perm_of (Phant T)).
+
+Canonical perm_subType := Eval hnf in [subType for pval].
+Definition perm_eqMixin := Eval hnf in [eqMixin of perm_type by <:].
+Canonical perm_eqType := Eval hnf in EqType perm_type perm_eqMixin.
+Definition perm_choiceMixin := [choiceMixin of perm_type by <:].
+Canonical perm_choiceType := Eval hnf in ChoiceType perm_type perm_choiceMixin.
+Definition perm_countMixin := [countMixin of perm_type by <:].
+Canonical perm_countType := Eval hnf in CountType perm_type perm_countMixin.
+Canonical perm_subCountType := Eval hnf in [subCountType of perm_type].
+Definition perm_finMixin := [finMixin of perm_type by <:].
+Canonical perm_finType := Eval hnf in FinType perm_type perm_finMixin.
+Canonical perm_subFinType := Eval hnf in [subFinType of perm_type].
+
+Canonical perm_for_subType := Eval hnf in [subType of pT].
+Canonical perm_for_eqType := Eval hnf in [eqType of pT].
+Canonical perm_for_choiceType := Eval hnf in [choiceType of pT].
+Canonical perm_for_countType := Eval hnf in [countType of pT].
+Canonical perm_for_subCountType := Eval hnf in [subCountType of pT].
+Canonical perm_for_finType := Eval hnf in [finType of pT].
+Canonical perm_for_subFinType := Eval hnf in [subFinType of pT].
+
+Lemma perm_proof (f : T -> T) : injective f -> injectiveb (finfun f).
+Proof.
+by move=> f_inj; apply/injectiveP; apply: eq_inj f_inj _ => x; rewrite ffunE.
+Qed.
+
+End PermDefSection.
+
+Notation "{ 'perm' T }" := (perm_of (Phant T))
+ (at level 0, format "{ 'perm' T }") : type_scope.
+
+Arguments Scope pval [_ group_scope].
+
+Bind Scope group_scope with perm_type.
+Bind Scope group_scope with perm_of.
+
+Notation "''S_' n" := {perm 'I_n}
+ (at level 8, n at level 2, format "''S_' n").
+
+Notation Local fun_of_perm_def := (fun T (u : perm_type T) => val u : T -> T).
+Notation Local perm_def := (fun T f injf => Perm (@perm_proof T f injf)).
+
+Module Type PermDefSig.
+Parameter fun_of_perm : forall T, perm_type T -> T -> T.
+Parameter perm : forall (T : finType) (f : T -> T), injective f -> {perm T}.
+Axiom fun_of_permE : fun_of_perm = fun_of_perm_def.
+Axiom permE : perm = perm_def.
+End PermDefSig.
+
+Module PermDef : PermDefSig.
+Definition fun_of_perm := fun_of_perm_def.
+Definition perm := perm_def.
+Lemma fun_of_permE : fun_of_perm = fun_of_perm_def. Proof. by []. Qed.
+Lemma permE : perm = perm_def. Proof. by []. Qed.
+End PermDef.
+
+Notation fun_of_perm := PermDef.fun_of_perm.
+Notation "@ 'perm'" := (@PermDef.perm) (at level 10, format "@ 'perm'").
+Notation perm := (@PermDef.perm _ _).
+Canonical fun_of_perm_unlock := Unlockable PermDef.fun_of_permE.
+Canonical perm_unlock := Unlockable PermDef.permE.
+Coercion fun_of_perm : perm_type >-> Funclass.
+
+Section Theory.
+
+Variable T : finType.
+Implicit Types (x y : T) (s t : {perm T}) (S : {set T}).
+
+Lemma permP s t : s =1 t <-> s = t.
+Proof. by split=> [| -> //]; rewrite unlock => eq_sv; exact/val_inj/ffunP. Qed.
+
+Lemma pvalE s : pval s = s :> (T -> T).
+Proof. by rewrite [@fun_of_perm]unlock. Qed.
+
+Lemma permE f f_inj : @perm T f f_inj =1 f.
+Proof. by move=> x; rewrite -pvalE [@perm]unlock ffunE. Qed.
+
+Lemma perm_inj s : injective s.
+Proof. by rewrite -!pvalE; exact: (injectiveP _ (valP s)). Qed.
+
+Implicit Arguments perm_inj [].
+Hint Resolve perm_inj.
+
+Lemma perm_onto s : codom s =i predT.
+Proof. by apply/subset_cardP; rewrite ?card_codom ?subset_predT. Qed.
+
+Definition perm_one := perm (@inj_id T).
+
+Lemma perm_invK s : cancel (fun x => iinv (perm_onto s x)) s.
+Proof. by move=> x /=; rewrite f_iinv. Qed.
+
+Definition perm_inv s := perm (can_inj (perm_invK s)).
+
+Definition perm_mul s t := perm (inj_comp (perm_inj t) (perm_inj s)).
+
+Lemma perm_oneP : left_id perm_one perm_mul.
+Proof. by move=> s; apply/permP => x; rewrite permE /= permE. Qed.
+
+Lemma perm_invP : left_inverse perm_one perm_inv perm_mul.
+Proof. by move=> s; apply/permP=> x; rewrite !permE /= permE f_iinv. Qed.
+
+Lemma perm_mulP : associative perm_mul.
+Proof. by move=> s t u; apply/permP=> x; do !rewrite permE /=. Qed.
+
+Definition perm_of_baseFinGroupMixin : FinGroup.mixin_of (perm_type T) :=
+ FinGroup.Mixin perm_mulP perm_oneP perm_invP.
+Canonical perm_baseFinGroupType :=
+ Eval hnf in BaseFinGroupType (perm_type T) perm_of_baseFinGroupMixin.
+Canonical perm_finGroupType := @FinGroupType perm_baseFinGroupType perm_invP.
+
+Canonical perm_of_baseFinGroupType :=
+ Eval hnf in [baseFinGroupType of {perm T}].
+Canonical perm_of_finGroupType := Eval hnf in [finGroupType of {perm T} ].
+
+Lemma perm1 x : (1 : {perm T}) x = x.
+Proof. by rewrite permE. Qed.
+
+Lemma permM s t x : (s * t) x = t (s x).
+Proof. by rewrite permE. Qed.
+
+Lemma permK s : cancel s s^-1.
+Proof. by move=> x; rewrite -permM mulgV perm1. Qed.
+
+Lemma permKV s : cancel s^-1 s.
+Proof. by have:= permK s^-1; rewrite invgK. Qed.
+
+Lemma permJ s t x : (s ^ t) (t x) = t (s x).
+Proof. by rewrite !permM permK. Qed.
+
+Lemma permX s x n : (s ^+ n) x = iter n s x.
+Proof. by elim: n => [|n /= <-]; rewrite ?perm1 // -permM expgSr. Qed.
+
+Lemma im_permV s S : s^-1 @: S = s @^-1: S.
+Proof. exact: can2_imset_pre (permKV s) (permK s). Qed.
+
+Lemma preim_permV s S : s^-1 @^-1: S = s @: S.
+Proof. by rewrite -im_permV invgK. Qed.
+
+Definition perm_on S : pred {perm T} := fun s => [pred x | s x != x] \subset S.
+
+Lemma perm_closed S s x : perm_on S s -> (s x \in S) = (x \in S).
+Proof.
+move/subsetP=> s_on_S; have [-> // | nfix_s_x] := eqVneq (s x) x.
+by rewrite !s_on_S // inE /= ?(inj_eq (perm_inj s)).
+Qed.
+
+Lemma perm_on1 H : perm_on H 1.
+Proof. by apply/subsetP=> x; rewrite inE /= perm1 eqxx. Qed.
+
+Lemma perm_onM H s t : perm_on H s -> perm_on H t -> perm_on H (s * t).
+Proof.
+move/subsetP=> sH /subsetP tH; apply/subsetP => x; rewrite inE /= permM.
+by have [-> /tH | /sH] := eqVneq (s x) x.
+Qed.
+
+Lemma out_perm S u x : perm_on S u -> x \notin S -> u x = x.
+Proof. by move=> uS; exact: contraNeq (subsetP uS x). Qed.
+
+Lemma im_perm_on u S : perm_on S u -> u @: S = S.
+Proof.
+move=> Su; rewrite -preim_permV; apply/setP=> x.
+by rewrite !inE -(perm_closed _ Su) permKV.
+Qed.
+
+Lemma tperm_proof x y : involutive [fun z => z with x |-> y, y |-> x].
+Proof.
+move=> z /=; case: (z =P x) => [-> | ne_zx]; first by rewrite eqxx; case: eqP.
+by case: (z =P y) => [->| ne_zy]; [rewrite eqxx | do 2?case: eqP].
+Qed.
+
+Definition tperm x y := perm (can_inj (tperm_proof x y)).
+
+CoInductive tperm_spec x y z : T -> Type :=
+ | TpermFirst of z = x : tperm_spec x y z y
+ | TpermSecond of z = y : tperm_spec x y z x
+ | TpermNone of z <> x & z <> y : tperm_spec x y z z.
+
+Lemma tpermP x y z : tperm_spec x y z (tperm x y z).
+Proof. by rewrite permE /=; do 2?[case: eqP => /=]; constructor; auto. Qed.
+
+Lemma tpermL x y : tperm x y x = y.
+Proof. by case: tpermP. Qed.
+
+Lemma tpermR x y : tperm x y y = x.
+Proof. by case: tpermP. Qed.
+
+Lemma tpermD x y z : x != z -> y != z -> tperm x y z = z.
+Proof. by case: tpermP => // ->; rewrite eqxx. Qed.
+
+Lemma tpermC x y : tperm x y = tperm y x.
+Proof. by apply/permP => z; do 2![case: tpermP => //] => ->. Qed.
+
+Lemma tperm1 x : tperm x x = 1.
+Proof. by apply/permP => z; rewrite perm1; case: tpermP. Qed.
+
+Lemma tpermK x y : involutive (tperm x y).
+Proof. by move=> z; rewrite !permE tperm_proof. Qed.
+
+Lemma tpermKg x y : involutive (mulg (tperm x y)).
+Proof. by move=> s; apply/permP=> z; rewrite !permM tpermK. Qed.
+
+Lemma tpermV x y : (tperm x y)^-1 = tperm x y.
+Proof. by set t := tperm x y; rewrite -{2}(mulgK t t) -mulgA tpermKg. Qed.
+
+Lemma tperm2 x y : tperm x y * tperm x y = 1.
+Proof. by rewrite -{1}tpermV mulVg. Qed.
+
+Lemma card_perm A : #|perm_on A| = (#|A|)`!.
+Proof.
+pose ffA := {ffun {x | x \in A} -> T}.
+rewrite -ffactnn -{2}(card_sig (mem A)) /= -card_inj_ffuns_on.
+pose fT (f : ffA) := [ffun x => oapp f x (insub x)].
+pose pfT f := insubd (1 : {perm T}) (fT f).
+pose fA s : ffA := [ffun u => s (val u)].
+rewrite -!sum1dep_card -sum1_card (reindex_onto fA pfT) => [|f].
+ apply: eq_bigl => p; rewrite andbC; apply/idP/and3P=> [onA | []]; first split.
+ - apply/eqP; suffices fTAp: fT (fA p) = pval p.
+ by apply/permP=> x; rewrite -!pvalE insubdK fTAp //; exact: (valP p).
+ apply/ffunP=> x; rewrite ffunE pvalE.
+ by case: insubP => [u _ <- | /out_perm->] //=; rewrite ffunE.
+ - by apply/forallP=> [[x Ax]]; rewrite ffunE /= perm_closed.
+ - by apply/injectiveP=> u v; rewrite !ffunE => /perm_inj; exact: val_inj.
+ move/eqP=> <- _ _; apply/subsetP=> x; rewrite !inE -pvalE val_insubd fun_if.
+ by rewrite if_arg ffunE; case: insubP; rewrite // pvalE perm1 if_same eqxx.
+case/andP=> /forallP-onA /injectiveP-f_inj.
+apply/ffunP=> u; rewrite ffunE -pvalE insubdK; first by rewrite ffunE valK.
+apply/injectiveP=> {u} x y; rewrite !ffunE.
+case: insubP => [u _ <-|]; case: insubP => [v _ <-|] //=; first by move/f_inj->.
+ by move=> Ay' def_y; rewrite -def_y [_ \in A]onA in Ay'.
+by move=> Ax' def_x; rewrite def_x [_ \in A]onA in Ax'.
+Qed.
+
+End Theory.
+
+Prenex Implicits tperm.
+
+Lemma inj_tperm (T T' : finType) (f : T -> T') x y z :
+ injective f -> f (tperm x y z) = tperm (f x) (f y) (f z).
+Proof. by move=> injf; rewrite !permE /= !(inj_eq injf) !(fun_if f). Qed.
+
+Lemma tpermJ (T : finType) x y (s : {perm T}) :
+ (tperm x y) ^ s = tperm (s x) (s y).
+Proof.
+apply/permP => z; rewrite -(permKV s z) permJ; apply: inj_tperm.
+exact: perm_inj.
+Qed.
+
+Lemma tuple_perm_eqP {T : eqType} {n} {s : seq T} {t : n.-tuple T} :
+ reflect (exists p : 'S_n, s = [tuple tnth t (p i) | i < n]) (perm_eq s t).
+Proof.
+apply: (iffP idP) => [|[p ->]]; last first.
+ rewrite /= (map_comp (tnth t)) -{1}(map_tnth_enum t) perm_map //.
+ apply: uniq_perm_eq => [||i]; rewrite ?enum_uniq //.
+ by apply/injectiveP; apply: perm_inj.
+ by rewrite mem_enum -[i](permKV p) image_f.
+case: n => [|n] in t *; last have x0 := tnth t ord0.
+ rewrite tuple0 => /perm_eq_small-> //.
+ by exists 1; rewrite [mktuple _]tuple0.
+case/(perm_eq_iotaP x0); rewrite size_tuple => Is eqIst ->{s}.
+have uniqIs: uniq Is by rewrite (perm_eq_uniq eqIst) iota_uniq.
+have szIs: size Is == n.+1 by rewrite (perm_eq_size eqIst) !size_tuple.
+have pP i : tnth (Tuple szIs) i < n.+1.
+ by rewrite -[_ < _](mem_iota 0) -(perm_eq_mem eqIst) mem_tnth.
+have inj_p: injective (fun i => Ordinal (pP i)).
+ by apply/injectiveP/(@map_uniq _ _ val); rewrite -map_comp map_tnth_enum.
+exists (perm inj_p); rewrite -[Is]/(tval (Tuple szIs)); congr (tval _).
+by apply: eq_from_tnth => i; rewrite tnth_map tnth_mktuple permE (tnth_nth x0).
+Qed.
+
+Section PermutationParity.
+
+Variable T : finType.
+
+Implicit Types (s t u v : {perm T}) (x y z a b : T).
+
+(* Note that pcycle s x is the orbit of x by <[s]> under the action aperm. *)
+(* Hence, the pcycle lemmas below are special cases of more general lemmas *)
+(* on orbits that will be stated in action.v. *)
+(* Defining pcycle directly here avoids a dependency of matrix.v on *)
+(* action.v and hence morphism.v. *)
+
+Definition aperm x s := s x.
+Definition pcycle s x := aperm x @: <[s]>.
+Definition pcycles s := pcycle s @: T.
+Definition odd_perm (s : perm_type T) := odd #|T| (+) odd #|pcycles s|.
+
+Lemma apermE x s : aperm x s = s x. Proof. by []. Qed.
+
+Lemma mem_pcycle s i x : (s ^+ i) x \in pcycle s x.
+Proof. by rewrite (mem_imset (aperm x)) ?mem_cycle. Qed.
+
+Lemma pcycle_id s x : x \in pcycle s x.
+Proof. by rewrite -{1}[x]perm1 (mem_pcycle s 0). Qed.
+
+Lemma uniq_traject_pcycle s x : uniq (traject s x #|pcycle s x|).
+Proof.
+case def_n: #|_| => // [n]; rewrite looping_uniq.
+apply: contraL (card_size (traject s x n)) => /loopingP t_sx.
+rewrite -ltnNge size_traject -def_n ?subset_leq_card //.
+by apply/subsetP=> _ /imsetP[_ /cycleP[i ->] ->]; rewrite /aperm permX t_sx.
+Qed.
+
+Lemma pcycle_traject s x : pcycle s x =i traject s x #|pcycle s x|.
+Proof.
+apply: fsym; apply/subset_cardP.
+ by rewrite (card_uniqP _) ?size_traject ?uniq_traject_pcycle.
+by apply/subsetP=> _ /trajectP[i _ ->]; rewrite -permX mem_pcycle.
+Qed.
+
+Lemma iter_pcycle s x : iter #|pcycle s x| s x = x.
+Proof.
+case def_n: #|_| (uniq_traject_pcycle s x) => [//|n] Ut.
+have: looping s x n.+1.
+ by rewrite -def_n -[looping _ _ _]pcycle_traject -permX mem_pcycle.
+rewrite /looping => /trajectP[[|i] //= lt_i_n /perm_inj eq_i_n_sx].
+move: lt_i_n; rewrite ltnS ltn_neqAle andbC => /andP[le_i_n /negP[]].
+by rewrite -(nth_uniq x _ _ Ut) ?size_traject ?nth_traject // eq_i_n_sx.
+Qed.
+
+Lemma eq_pcycle_mem s x y : (pcycle s x == pcycle s y) = (x \in pcycle s y).
+Proof.
+apply/eqP/idP=> [<- | /imsetP[si s_si ->]]; first exact: pcycle_id.
+apply/setP => z; apply/imsetP/imsetP=> [] [sj s_sj ->].
+ by exists (si * sj); rewrite ?groupM /aperm ?permM.
+exists (si^-1 * sj); first by rewrite groupM ?groupV.
+by rewrite /aperm permM permK.
+Qed.
+
+Lemma pcycle_sym s x y : (x \in pcycle s y) = (y \in pcycle s x).
+Proof. by rewrite -!eq_pcycle_mem eq_sym. Qed.
+
+Lemma pcycle_perm s i x : pcycle s ((s ^+ i) x) = pcycle s x.
+Proof. by apply/eqP; rewrite eq_pcycle_mem mem_pcycle. Qed.
+
+Lemma ncycles_mul_tperm s x y : let t := tperm x y in
+ #|pcycles (t * s)| + (x \notin pcycle s y).*2 = #|pcycles s| + (x != y).
+Proof.
+pose xf a b u := find (pred2 a b) (traject u (u a) #|pcycle u a|).
+have xf_size a b u: xf a b u <= #|pcycle u a|.
+ by rewrite (leq_trans (find_size _ _)) ?size_traject.
+have lt_xf a b u n : n < xf a b u -> ~~ pred2 a b ((u ^+ n.+1) a).
+ move=> lt_n; apply: contraFN (before_find (u a) lt_n).
+ by rewrite permX iterSr nth_traject // (leq_trans lt_n).
+pose t a b u := tperm a b * u.
+have tC a b u : t a b u = t b a u by rewrite /t tpermC.
+have tK a b: involutive (t a b) by move=> u; exact: tpermKg.
+have tXC a b u n: n <= xf a b u -> (t a b u ^+ n.+1) b = (u ^+ n.+1) a.
+ elim: n => [|n IHn] lt_n_f; first by rewrite permM tpermR.
+ rewrite !(expgSr _ n.+1) !permM {}IHn 1?ltnW //; congr (u _).
+ by case/lt_xf/norP: lt_n_f => ne_a ne_b; rewrite tpermD // eq_sym.
+have eq_xf a b u: pred2 a b ((u ^+ (xf a b u).+1) a).
+ have ua_a: a \in pcycle u (u a) by rewrite pcycle_sym (mem_pcycle _ 1).
+ have has_f: has (pred2 a b) (traject u (u a) #|pcycle u (u a)|).
+ by apply/hasP; exists a; rewrite /= ?eqxx -?pcycle_traject.
+ have:= nth_find (u a) has_f; rewrite has_find size_traject in has_f.
+ rewrite -eq_pcycle_mem in ua_a.
+ by rewrite nth_traject // -iterSr -permX -(eqP ua_a).
+have xfC a b u: xf b a (t a b u) = xf a b u.
+ without loss lt_a: a b u / xf b a (t a b u) < xf a b u.
+ move=> IHab; set m := xf b a _; set n := xf a b u.
+ by case: (ltngtP m n) => // ltx; [exact: IHab | rewrite -[m]IHab tC tK].
+ by move/lt_xf: (lt_a); rewrite -(tXC a b) 1?ltnW //= orbC [_ || _]eq_xf.
+pose ts := t x y s; rewrite /= -[_ * s]/ts.
+pose dp u := #|pcycles u :\ pcycle u y :\ pcycle u x|.
+rewrite !(addnC #|_|) (cardsD1 (pcycle ts y)) mem_imset ?inE //.
+rewrite (cardsD1 (pcycle ts x)) inE mem_imset ?inE //= -/(dp ts) {}/ts.
+rewrite (cardsD1 (pcycle s y)) (cardsD1 (pcycle s x)) !(mem_imset, inE) //.
+rewrite -/(dp s) !addnA !eq_pcycle_mem andbT; congr (_ + _); last first.
+ wlog suffices: s / dp s <= dp (t x y s).
+ by move=> IHs; apply/eqP; rewrite eqn_leq -{2}(tK x y s) !IHs.
+ apply/subset_leq_card/subsetP=> {dp} C.
+ rewrite !inE andbA andbC !(eq_sym C) => /and3P[/imsetP[z _ ->{C}]].
+ rewrite 2!eq_pcycle_mem => sxz syz.
+ suffices ts_z: pcycle (t x y s) z = pcycle s z.
+ by rewrite -ts_z !eq_pcycle_mem {1 2}ts_z sxz syz mem_imset ?inE.
+ suffices exp_id n: ((t x y s) ^+ n) z = (s ^+ n) z.
+ apply/setP=> u; apply/idP/idP=> /imsetP[_ /cycleP[i ->] ->].
+ by rewrite /aperm exp_id mem_pcycle.
+ by rewrite /aperm -exp_id mem_pcycle.
+ elim: n => // n IHn; rewrite !expgSr !permM {}IHn tpermD //.
+ apply: contraNneq sxz => ->; exact: mem_pcycle.
+ apply: contraNneq syz => ->; exact: mem_pcycle.
+case: eqP {dp} => [<- | ne_xy]; first by rewrite /t tperm1 mul1g pcycle_id.
+suff ->: (x \in pcycle (t x y s) y) = (x \notin pcycle s y) by case: (x \in _).
+without loss xf_x: s x y ne_xy / (s ^+ (xf x y s).+1) x = x.
+ move=> IHs; have ne_yx := nesym ne_xy; have:= eq_xf x y s; set n := xf x y s.
+ case/pred2P=> [|snx]; first exact: IHs.
+ by rewrite -[x \in _]negbK ![x \in _]pcycle_sym -{}IHs ?xfC ?tXC // tC tK.
+rewrite -{1}xf_x -(tXC _ _ _ _ (leqnn _)) mem_pcycle; symmetry.
+rewrite -eq_pcycle_mem eq_sym eq_pcycle_mem pcycle_traject.
+apply/trajectP=> [[n _ snx]].
+have: looping s x (xf x y s).+1 by rewrite /looping -permX xf_x inE eqxx.
+move/loopingP/(_ n); rewrite -{n}snx.
+case/trajectP=> [[_|i]]; first exact: nesym; rewrite ltnS -permX => lt_i def_y.
+by move/lt_xf: lt_i; rewrite def_y /= eqxx orbT.
+Qed.
+
+Lemma odd_perm1 : odd_perm 1 = false.
+Proof.
+rewrite /odd_perm card_imset ?addbb // => x y; move/eqP.
+by rewrite eq_pcycle_mem /pcycle cycle1 imset_set1 /aperm perm1; move/set1P.
+Qed.
+
+Lemma odd_mul_tperm x y s : odd_perm (tperm x y * s) = (x != y) (+) odd_perm s.
+Proof.
+rewrite addbC -addbA -[~~ _]oddb -odd_add -ncycles_mul_tperm.
+by rewrite odd_add odd_double addbF.
+Qed.
+
+Lemma odd_tperm x y : odd_perm (tperm x y) = (x != y).
+Proof. by rewrite -[_ y]mulg1 odd_mul_tperm odd_perm1 addbF. Qed.
+
+Definition dpair (eT : eqType) := [pred t | t.1 != t.2 :> eT].
+Implicit Arguments dpair [eT].
+
+Lemma prod_tpermP s :
+ {ts : seq (T * T) | s = \prod_(t <- ts) tperm t.1 t.2 & all dpair ts}.
+Proof.
+elim: {s}_.+1 {-2}s (ltnSn #|[pred x | s x != x]|) => // n IHn s.
+rewrite ltnS => le_s_n; case: (pickP (fun x => s x != x)) => [x s_x | s_id].
+ have [|ts def_s ne_ts] := IHn (tperm x (s^-1 x) * s).
+ rewrite (cardD1 x) !inE s_x in le_s_n; apply: leq_ltn_trans le_s_n.
+ apply: subset_leq_card; apply/subsetP=> y.
+ rewrite !inE permM permE /= -(canF_eq (permK _)).
+ have [-> | ne_yx] := altP (y =P x); first by rewrite permKV eqxx.
+ by case: (s y =P x) => // -> _; rewrite eq_sym.
+ exists ((x, s^-1 x) :: ts); last by rewrite /= -(canF_eq (permK _)) s_x.
+ by rewrite big_cons -def_s mulgA tperm2 mul1g.
+exists nil; rewrite // big_nil; apply/permP=> x.
+by apply/eqP/idPn; rewrite perm1 s_id.
+Qed.
+
+Lemma odd_perm_prod ts :
+ all dpair ts -> odd_perm (\prod_(t <- ts) tperm t.1 t.2) = odd (size ts).
+Proof.
+elim: ts => [_|t ts IHts] /=; first by rewrite big_nil odd_perm1.
+by case/andP=> dt12 dts; rewrite big_cons odd_mul_tperm dt12 IHts.
+Qed.
+
+Lemma odd_permM : {morph odd_perm : s1 s2 / s1 * s2 >-> s1 (+) s2}.
+Proof.
+move=> s1 s2; case: (prod_tpermP s1) => ts1 ->{s1} dts1.
+case: (prod_tpermP s2) => ts2 ->{s2} dts2.
+by rewrite -big_cat !odd_perm_prod ?all_cat ?dts1 // size_cat odd_add.
+Qed.
+
+Lemma odd_permV s : odd_perm s^-1 = odd_perm s.
+Proof. by rewrite -{2}(mulgK s s) !odd_permM -addbA addKb. Qed.
+
+Lemma odd_permJ s1 s2 : odd_perm (s1 ^ s2) = odd_perm s1.
+Proof. by rewrite !odd_permM odd_permV addbC addbK. Qed.
+
+End PermutationParity.
+
+Coercion odd_perm : perm_type >-> bool.
+Implicit Arguments dpair [eT].
+Prenex Implicits pcycle dpair pcycles aperm.
+
+Section LiftPerm.
+(* Somewhat more specialised constructs for permutations on ordinals. *)
+
+Variable n : nat.
+Implicit Types i j : 'I_n.+1.
+Implicit Types s t : 'S_n.
+
+Definition lift_perm_fun i j s k :=
+ if unlift i k is Some k' then lift j (s k') else j.
+
+Lemma lift_permK i j s :
+ cancel (lift_perm_fun i j s) (lift_perm_fun j i s^-1).
+Proof.
+rewrite /lift_perm_fun => k.
+by case: (unliftP i k) => [j'|] ->; rewrite (liftK, unlift_none) ?permK.
+Qed.
+
+Definition lift_perm i j s := perm (can_inj (lift_permK i j s)).
+
+Lemma lift_perm_id i j s : lift_perm i j s i = j.
+Proof. by rewrite permE /lift_perm_fun unlift_none. Qed.
+
+Lemma lift_perm_lift i j s k' :
+ lift_perm i j s (lift i k') = lift j (s k') :> 'I_n.+1.
+Proof. by rewrite permE /lift_perm_fun liftK. Qed.
+
+Lemma lift_permM i j k s t :
+ lift_perm i j s * lift_perm j k t = lift_perm i k (s * t).
+Proof.
+apply/permP=> i1; case: (unliftP i i1) => [i2|] ->{i1}.
+ by rewrite !(permM, lift_perm_lift).
+by rewrite permM !lift_perm_id.
+Qed.
+
+Lemma lift_perm1 i : lift_perm i i 1 = 1.
+Proof. by apply: (mulgI (lift_perm i i 1)); rewrite lift_permM !mulg1. Qed.
+
+Lemma lift_permV i j s : (lift_perm i j s)^-1 = lift_perm j i s^-1.
+Proof. by apply/eqP; rewrite eq_invg_mul lift_permM mulgV lift_perm1. Qed.
+
+Lemma odd_lift_perm i j s : lift_perm i j s = odd i (+) odd j (+) s :> bool.
+Proof.
+rewrite -{1}(mul1g s) -(lift_permM _ j) odd_permM.
+congr (_ (+) _); last first.
+ case: (prod_tpermP s) => ts ->{s} _.
+ elim: ts => [|t ts IHts] /=; first by rewrite big_nil lift_perm1 !odd_perm1.
+ rewrite big_cons odd_mul_tperm -(lift_permM _ j) odd_permM {}IHts //.
+ congr (_ (+) _); transitivity (tperm (lift j t.1) (lift j t.2)); last first.
+ by rewrite odd_tperm (inj_eq (@lift_inj _ _)).
+ congr odd_perm; apply/permP=> k; case: (unliftP j k) => [k'|] ->.
+ rewrite lift_perm_lift inj_tperm //; exact: lift_inj.
+ by rewrite lift_perm_id tpermD // eq_sym neq_lift.
+suff{i j s} odd_lift0 (k : 'I_n.+1): lift_perm ord0 k 1 = odd k :> bool.
+ rewrite -!odd_lift0 -{2}invg1 -lift_permV odd_permV -odd_permM.
+ by rewrite lift_permM mulg1.
+elim: {k}(k : nat) {1 3}k (erefl (k : nat)) => [|m IHm] k def_k.
+ rewrite (_ : k = ord0) ?lift_perm1 ?odd_perm1 //; exact: val_inj.
+have le_mn: m < n.+1 by [rewrite -def_k ltnW]; pose j := Ordinal le_mn.
+rewrite -(mulg1 1)%g -(lift_permM _ j) odd_permM {}IHm // addbC.
+rewrite (_ : _ 1 = tperm j k); first by rewrite odd_tperm neq_ltn def_k leqnn.
+apply/permP=> i; case: (unliftP j i) => [i'|] ->; last first.
+ by rewrite lift_perm_id tpermL.
+apply: ord_inj; rewrite lift_perm_lift !permE /= eq_sym -if_neg neq_lift.
+rewrite fun_if -val_eqE /= def_k /bump ltn_neqAle andbC.
+case: leqP => [_ | lt_i'm] /=; last by rewrite -if_neg neq_ltn leqW.
+by rewrite add1n eqSS eq_sym; case: eqP.
+Qed.
+
+End LiftPerm.
+
+
+
diff --git a/mathcomp/fingroup/presentation.v b/mathcomp/fingroup/presentation.v
new file mode 100644
index 0000000..46658d5
--- /dev/null
+++ b/mathcomp/fingroup/presentation.v
@@ -0,0 +1,254 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq fintype finset.
+Require Import fingroup morphism.
+
+(******************************************************************************)
+(* Support for generator-and-relation presentations of groups. We provide the *)
+(* syntax: *)
+(* G \homg Grp (x_1 : ... x_n : (s_1 = t_1, ..., s_m = t_m)) *)
+(* <=> G is generated by elements x_1, ..., x_m satisfying the relations *)
+(* s_1 = t_1, ..., s_m = t_m, i.e., G is a homomorphic image of the *)
+(* group generated by the x_i, subject to the relations s_j = t_j. *)
+(* G \isog Grp (x_1 : ... x_n : (s_1 = t_1, ..., s_m = t_m)) *)
+(* <=> G is isomorphic to the group generated by the x_i, subject to the *)
+(* relations s_j = t_j. This is an intensional predicate (in Prop), as *)
+(* even the non-triviality of a generated group is undecidable. *)
+(* Syntax details: *)
+(* - Grp is a litteral constant. *)
+(* - There must be at least one generator and one relation. *)
+(* - A relation s_j = 1 can be abbreviated as simply s_j (a.k.a. a relator). *)
+(* - Two consecutive relations s_j = t, s_j+1 = t can be abbreviated *)
+(* s_j = s_j+1 = t. *)
+(* - The s_j and t_j are terms built from the x_i and the standard group *)
+(* operators *, 1, ^-1, ^+, ^-, ^, [~ u_1, ..., u_k]; no other operator or *)
+(* abbreviation may be used, as the notation is implemented using static *)
+(* overloading. *)
+(* - This is the closest we could get to the notation used in Aschbacher, *)
+(* Grp (x_1, ... x_n : t_1,1 = ... = t_1,k1, ..., t_m,1 = ... = t_m,km) *)
+(* under the current limitations of the Coq Notation facility. *)
+(* Semantics details: *)
+(* - G \isog Grp (...) : Prop expands to the statement *)
+(* forall rT (H : {group rT}), (H \homg G) = (H \homg Grp (...)) *)
+(* (with rT : finGroupType). *)
+(* - G \homg Grp (x_1 : ... x_n : (s_1 = t_1, ..., s_m = t_m)) : bool, with *)
+(* G : {set gT}, is convertible to the boolean expression *)
+(* [exists t : gT * ... gT, let: (x_1, ..., x_n) := t in *)
+(* (<[x_1]> <*> ... <*> <[x_n]>, (s_1, ... (s_m-1, s_m) ...)) *)
+(* == (G, (t_1, ... (t_m-1, t_m) ...))] *)
+(* where the tuple comparison above is convertible to the conjunction *)
+(* [&& <[x_1]> <*> ... <*> <[x_n]> == G, s_1 == t_1, ... & s_m == t_m] *)
+(* Thus G \homg Grp (...) can be easily exploited by destructing the tuple *)
+(* created case/existsP, then destructing the tuple equality with case/eqP. *)
+(* Conversely it can be proved by using apply/existsP, providing the tuple *)
+(* with a single exists (u_1, ..., u_n), then using rewrite !xpair_eqE /= *)
+(* to expose the conjunction, and optionally using an apply/and{m+1}P view *)
+(* to split it into subgoals (in that case, the rewrite is in principle *)
+(* redundant, but necessary in practice because of the poor performance of *)
+(* conversion in the Coq unifier). *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Module Presentation.
+
+Section Presentation.
+
+Implicit Types gT rT : finGroupType.
+Implicit Type vT : finType. (* tuple value type *)
+
+Inductive term :=
+ | Cst of nat
+ | Idx
+ | Inv of term
+ | Exp of term & nat
+ | Mul of term & term
+ | Conj of term & term
+ | Comm of term & term.
+
+Fixpoint eval {gT} e t : gT :=
+ match t with
+ | Cst i => nth 1 e i
+ | Idx => 1
+ | Inv t1 => (eval e t1)^-1
+ | Exp t1 n => eval e t1 ^+ n
+ | Mul t1 t2 => eval e t1 * eval e t2
+ | Conj t1 t2 => eval e t1 ^ eval e t2
+ | Comm t1 t2 => [~ eval e t1, eval e t2]
+ end.
+
+Inductive formula := Eq2 of term & term | And of formula & formula.
+Definition Eq1 s := Eq2 s Idx.
+Definition Eq3 s1 s2 t := And (Eq2 s1 t) (Eq2 s2 t).
+
+Inductive rel_type := NoRel | Rel vT of vT & vT.
+
+Definition bool_of_rel r := if r is Rel vT v1 v2 then v1 == v2 else true.
+Local Coercion bool_of_rel : rel_type >-> bool.
+
+Definition and_rel vT (v1 v2 : vT) r :=
+ if r is Rel wT w1 w2 then Rel (v1, w1) (v2, w2) else Rel v1 v2.
+
+Fixpoint rel {gT} (e : seq gT) f r :=
+ match f with
+ | Eq2 s t => and_rel (eval e s) (eval e t) r
+ | And f1 f2 => rel e f1 (rel e f2 r)
+ end.
+
+Inductive type := Generator of term -> type | Formula of formula.
+Definition Cast p : type := p. (* syntactic scope cast *)
+Local Coercion Formula : formula >-> type.
+
+Inductive env gT := Env of {set gT} & seq gT.
+Definition env1 {gT} (x : gT : finType) := Env <[x]> [:: x].
+
+Fixpoint sat gT vT B n (s : vT -> env gT) p :=
+ match p with
+ | Formula f =>
+ [exists v, let: Env A e := s v in and_rel A B (rel (rev e) f NoRel)]
+ | Generator p' =>
+ let s' v := let: Env A e := s v.1 in Env (A <*> <[v.2]>) (v.2 :: e) in
+ sat B n.+1 s' (p' (Cst n))
+ end.
+
+Definition hom gT (B : {set gT}) p := sat B 1 env1 (p (Cst 0)).
+Definition iso gT (B : {set gT}) p :=
+ forall rT (H : {group rT}), (H \homg B) = hom H p.
+
+End Presentation.
+
+End Presentation.
+
+Import Presentation.
+
+Coercion bool_of_rel : rel_type >-> bool.
+Coercion Eq1 : term >-> formula.
+Coercion Formula : formula >-> type.
+
+(* Declare (implicitly) the argument scope tags. *)
+Notation "1" := Idx : group_presentation.
+Arguments Scope Inv [group_presentation].
+Arguments Scope Exp [group_presentation nat_scope].
+Arguments Scope Mul [group_presentation group_presentation].
+Arguments Scope Conj [group_presentation group_presentation].
+Arguments Scope Comm [group_presentation group_presentation].
+Arguments Scope Eq1 [group_presentation].
+Arguments Scope Eq2 [group_presentation group_presentation].
+Arguments Scope Eq3 [group_presentation group_presentation group_presentation].
+Arguments Scope And [group_presentation group_presentation].
+Arguments Scope Formula [group_presentation].
+Arguments Scope Cast [group_presentation].
+
+Infix "*" := Mul : group_presentation.
+Infix "^+" := Exp : group_presentation.
+Infix "^" := Conj : group_presentation.
+Notation "x ^-1" := (Inv x) : group_presentation.
+Notation "x ^- n" := (Inv (x ^+ n)) : group_presentation.
+Notation "[ ~ x1 , x2 , .. , xn ]" :=
+ (Comm .. (Comm x1 x2) .. xn) : group_presentation.
+Notation "x = y" := (Eq2 x y) : group_presentation.
+Notation "x = y = z" := (Eq3 x y z) : group_presentation.
+Notation "( r1 , r2 , .. , rn )" :=
+ (And .. (And r1 r2) .. rn) : group_presentation.
+
+(* Declare (implicitly) the argument scope tags. *)
+Notation "x : p" := (fun x => Cast p) : nt_group_presentation.
+Arguments Scope Generator [nt_group_presentation].
+Arguments Scope hom [_ group_scope nt_group_presentation].
+Arguments Scope iso [_ group_scope nt_group_presentation].
+
+Notation "x : p" := (Generator (x : p)) : group_presentation.
+
+Notation "H \homg 'Grp' p" := (hom H p)
+ (at level 70, p at level 0, format "H \homg 'Grp' p") : group_scope.
+
+Notation "H \isog 'Grp' p" := (iso H p)
+ (at level 70, p at level 0, format "H \isog 'Grp' p") : group_scope.
+
+Notation "H \homg 'Grp' ( x : p )" := (hom H (x : p))
+ (at level 70, x at level 0,
+ format "'[hv' H '/ ' \homg 'Grp' ( x : p ) ']'") : group_scope.
+
+Notation "H \isog 'Grp' ( x : p )" := (iso H (x : p))
+ (at level 70, x at level 0,
+ format "'[hv' H '/ ' \isog 'Grp' ( x : p ) ']'") : group_scope.
+
+Section PresentationTheory.
+
+Implicit Types gT rT : finGroupType.
+
+Import Presentation.
+
+Lemma isoGrp_hom gT (G : {group gT}) p : G \isog Grp p -> G \homg Grp p.
+Proof. by move <-; exact: homg_refl. Qed.
+
+Lemma isoGrpP gT (G : {group gT}) p rT (H : {group rT}) :
+ G \isog Grp p -> reflect (#|H| = #|G| /\ H \homg Grp p) (H \isog G).
+Proof.
+move=> isoGp; apply: (iffP idP) => [isoGH | [oH homHp]].
+ by rewrite (card_isog isoGH) -isoGp isog_hom.
+by rewrite isogEcard isoGp homHp /= oH.
+Qed.
+
+Lemma homGrp_trans rT gT (H : {set rT}) (G : {group gT}) p :
+ H \homg G -> G \homg Grp p -> H \homg Grp p.
+Proof.
+case/homgP=> h <-{H}; rewrite /hom; move: {p}(p _) => p.
+have evalG e t: all (mem G) e -> eval (map h e) t = h (eval e t).
+ move=> Ge; apply: (@proj2 (eval e t \in G)); elim: t => /=.
+ - move=> i; case: (leqP (size e) i) => [le_e_i | lt_i_e].
+ by rewrite !nth_default ?size_map ?morph1.
+ by rewrite (nth_map 1) // [_ \in G](allP Ge) ?mem_nth.
+ - by rewrite morph1.
+ - by move=> t [Gt ->]; rewrite groupV morphV.
+ - by move=> t [Gt ->] n; rewrite groupX ?morphX.
+ - by move=> t1 [Gt1 ->] t2 [Gt2 ->]; rewrite groupM ?morphM.
+ - by move=> t1 [Gt1 ->] t2 [Gt2 ->]; rewrite groupJ ?morphJ.
+ by move=> t1 [Gt1 ->] t2 [Gt2 ->]; rewrite groupR ?morphR.
+have and_relE xT x1 x2 r: @and_rel xT x1 x2 r = (x1 == x2) && r :> bool.
+ by case: r => //=; rewrite andbT.
+have rsatG e f: all (mem G) e -> rel e f NoRel -> rel (map h e) f NoRel.
+ move=> Ge; have: NoRel -> NoRel by []; move: NoRel {2 4}NoRel.
+ elim: f => [x1 x2 | f1 IH1 f2 IH2] r hr IHr; last by apply: IH1; exact: IH2.
+ by rewrite !and_relE !evalG //; case/andP; move/eqP->; rewrite eqxx.
+set s := env1; set vT := gT : finType in s *.
+set s' := env1; set vT' := rT : finType in s' *.
+have (v): let: Env A e := s v in
+ A \subset G -> all (mem G) e /\ exists v', s' v' = Env (h @* A) (map h e).
+- rewrite /= cycle_subG andbT => Gv; rewrite morphim_cycle //.
+ by split; last exists (h v).
+elim: p 1%N vT vT' s s' => /= [p IHp | f] n vT vT' s s' Gs.
+ apply: IHp => [[v x]] /=; case: (s v) {Gs}(Gs v) => A e /= Gs.
+ rewrite join_subG cycle_subG; case/andP=> sAG Gx; rewrite Gx.
+ have [//|-> [v' def_v']] := Gs; split=> //; exists (v', h x); rewrite def_v'.
+ by congr (Env _ _); rewrite morphimY ?cycle_subG // morphim_cycle.
+case/existsP=> v; case: (s v) {Gs}(Gs v) => /= A e Gs.
+rewrite and_relE => /andP[/eqP defA rel_f].
+have{Gs} [|Ge [v' def_v']] := Gs; first by rewrite defA.
+apply/existsP; exists v'; rewrite def_v' and_relE defA eqxx /=.
+by rewrite -map_rev rsatG ?(eq_all_r (mem_rev e)).
+Qed.
+
+Lemma eq_homGrp gT rT (G : {group gT}) (H : {group rT}) p :
+ G \isog H -> (G \homg Grp p) = (H \homg Grp p).
+Proof.
+by rewrite isogEhom => /andP[homGH homHG]; apply/idP/idP; exact: homGrp_trans.
+Qed.
+
+Lemma isoGrp_trans gT rT (G : {group gT}) (H : {group rT}) p :
+ G \isog H -> H \isog Grp p -> G \isog Grp p.
+Proof. by move=> isoGH isoHp kT K; rewrite -isoHp; exact: eq_homgr. Qed.
+
+Lemma intro_isoGrp gT (G : {group gT}) p :
+ G \homg Grp p -> (forall rT (H : {group rT}), H \homg Grp p -> H \homg G) ->
+ G \isog Grp p.
+Proof.
+move=> homGp freeG rT H.
+by apply/idP/idP=> [homHp|]; [exact: homGrp_trans homGp | exact: freeG].
+Qed.
+
+End PresentationTheory.
+
diff --git a/mathcomp/fingroup/quotient.v b/mathcomp/fingroup/quotient.v
new file mode 100644
index 0000000..6688509
--- /dev/null
+++ b/mathcomp/fingroup/quotient.v
@@ -0,0 +1,972 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div choice.
+Require Import fintype prime finset fingroup morphism automorphism.
+
+(******************************************************************************)
+(* This file contains the definitions of: *)
+(* coset_of H == the (sub)type of bilateral cosets of H (see below) *)
+(* coset H == the canonical projection into coset_of H *)
+(* A / H == the quotient of A by H, that is, the morphic image *)
+(* of A by coset H. We do not require H <| A, so in a *)
+(* textbook A / H would be written 'N_A(H) * H / H. *)
+(* quotm f (nHG : H <| G) == the quotient morphism induced by f, *)
+(* mapping G / H onto f @* G / f @* H *)
+(* qisom f (eqHG : H = G) == the identity isomorphism between *)
+(* [set: coset_of G] and [set: coset_of H]. *)
+(* We also prove the three isomorphism theorems, and counting lemmas for *)
+(* morphisms. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section Cosets.
+
+Variables (gT : finGroupType) (Q A : {set gT}).
+
+(******************************************************************************)
+(* Cosets are right cosets of elements in the normaliser *)
+(* We let cosets coerce to GroupSet.sort, so they inherit the group subset *)
+(* base group structure. Later we will define a proper group structure on *)
+(* cosets, which will then hide the inherited structure once coset_of unifies *)
+(* with FinGroup.sort; the coercion to GroupSet.sort will no longer be used. *)
+(* Note that for Hx Hy : coset_of H, Hx * Hy : {set gT} can mean either *)
+(* set_of_coset (mulg Hx Hy) OR mulg (set_of_coset Hx) (set_of_coset Hy) *)
+(* However, since the two terms are actually convertible, we can live with *)
+(* this ambiguity. *)
+(* We take great care that neither the type coset_of H, nor its Canonical *)
+(* finGroupType structure, nor the coset H morphism depend on the actual *)
+(* group structure of H. Otherwise, rewriting would be extremely awkward *)
+(* because all our equalities are stated at the set level. *)
+(* The trick we use is to interpret coset_of A, when A is any set, as the *)
+(* type of cosets of the group <A> generated by A, in the group A <*> N(A) *)
+(* generated by A and its normaliser. This coincides with the type of *)
+(* bilateral cosets of A when A is a group. We restrict the domain of coset A *)
+(* to 'N(A), so that we get almost all the same conversion equalities as if *)
+(* we had forced A to be a group in the first place; the only exception, that *)
+(* 1 : coset_of A : set _ = <<A>> rather than A, is covered by genGid. *)
+(******************************************************************************)
+
+Notation H := <<A>>.
+Definition coset_range := [pred B in rcosets H 'N(A)].
+
+Record coset_of : Type :=
+ Coset { set_of_coset :> GroupSet.sort gT; _ : coset_range set_of_coset }.
+
+Canonical coset_subType := Eval hnf in [subType for set_of_coset].
+Definition coset_eqMixin := Eval hnf in [eqMixin of coset_of by <:].
+Canonical coset_eqType := Eval hnf in EqType coset_of coset_eqMixin.
+Definition coset_choiceMixin := [choiceMixin of coset_of by <:].
+Canonical coset_choiceType := Eval hnf in ChoiceType coset_of coset_choiceMixin.
+Definition coset_countMixin := [countMixin of coset_of by <:].
+Canonical coset_countType := Eval hnf in CountType coset_of coset_countMixin.
+Canonical coset_subCountType := Eval hnf in [subCountType of coset_of].
+Definition coset_finMixin := [finMixin of coset_of by <:].
+Canonical coset_finType := Eval hnf in FinType coset_of coset_finMixin.
+Canonical coset_subFinType := Eval hnf in [subFinType of coset_of].
+
+(* We build a new (canonical) structure of groupType for cosets. *)
+(* When A is a group, this is the largest possible quotient 'N(A) / A. *)
+
+Lemma coset_one_proof : coset_range H.
+Proof. by apply/rcosetsP; exists (1 : gT); rewrite (group1, mulg1). Qed.
+Definition coset_one := Coset coset_one_proof.
+
+Let nNH := subsetP (norm_gen A).
+
+Lemma coset_range_mul (B C : coset_of) : coset_range (B * C).
+Proof.
+case: B C => _ /= /rcosetsP[x Nx ->] [_ /= /rcosetsP[y Ny ->]].
+by apply/rcosetsP; exists (x * y); rewrite !(groupM, rcoset_mul, nNH).
+Qed.
+
+Definition coset_mul B C := Coset (coset_range_mul B C).
+
+Lemma coset_range_inv (B : coset_of) : coset_range B^-1.
+Proof.
+case: B => _ /= /rcosetsP[x Nx ->]; rewrite norm_rlcoset ?nNH // invg_lcoset.
+by apply/rcosetsP; exists x^-1; rewrite ?groupV.
+Qed.
+
+Definition coset_inv B := Coset (coset_range_inv B).
+
+Lemma coset_mulP : associative coset_mul.
+Proof. by move=> B C D; apply: val_inj; rewrite /= mulgA. Qed.
+
+Lemma coset_oneP : left_id coset_one coset_mul.
+Proof.
+case=> B coB; apply: val_inj => /=; case/rcosetsP: coB => x Hx ->{B}.
+by rewrite mulgA mulGid.
+Qed.
+
+Lemma coset_invP : left_inverse coset_one coset_inv coset_mul.
+Proof.
+case=> B coB; apply: val_inj => /=; case/rcosetsP: coB => x Hx ->{B}.
+rewrite invg_rcoset -mulgA (mulgA H) mulGid.
+by rewrite norm_rlcoset ?nNH // -lcosetM mulVg mul1g.
+Qed.
+
+Definition coset_of_groupMixin :=
+ FinGroup.Mixin coset_mulP coset_oneP coset_invP.
+
+Canonical coset_baseGroupType :=
+ Eval hnf in BaseFinGroupType coset_of coset_of_groupMixin.
+Canonical coset_groupType := FinGroupType coset_invP.
+
+(* Projection of the initial group type over the cosets groupType *)
+
+Definition coset x : coset_of := insubd (1 : coset_of) (H :* x).
+
+(* This is a primitive lemma -- we'll need to restate it for *)
+(* the case where A is a group. *)
+Lemma val_coset_prim x : x \in 'N(A) -> coset x :=: H :* x.
+Proof.
+by move=> Nx; rewrite val_insubd /= mem_rcosets -{1}(mul1g x) mem_mulg.
+Qed.
+
+Lemma coset_morphM : {in 'N(A) &, {morph coset : x y / x * y}}.
+Proof.
+move=> x y Nx Ny; apply: val_inj.
+by rewrite /= !val_coset_prim ?groupM //= rcoset_mul ?nNH.
+Qed.
+
+Canonical coset_morphism := Morphism coset_morphM.
+
+Lemma ker_coset_prim : 'ker coset = 'N_H(A).
+Proof.
+apply/setP=> z; rewrite !in_setI andbC 2!inE -val_eqE /=.
+case Nz: (z \in 'N(A)); rewrite ?andbF ?val_coset_prim // !andbT.
+by apply/eqP/idP=> [<-| Az]; rewrite (rcoset_refl, rcoset_id).
+Qed.
+
+Implicit Type xbar : coset_of.
+
+Lemma coset_mem y xbar : y \in xbar -> coset y = xbar.
+Proof.
+case: xbar => /= Hx NHx Hxy; apply: val_inj=> /=.
+case/rcosetsP: NHx (NHx) Hxy => x Nx -> NHx Hxy.
+by rewrite val_insubd /= (rcoset_transl Hxy) NHx.
+Qed.
+
+(* coset is an inverse to repr *)
+
+Lemma mem_repr_coset xbar : repr xbar \in xbar.
+Proof. case: xbar => /= _ /rcosetsP[x _ ->]; exact: mem_repr_rcoset. Qed.
+
+Lemma repr_coset1 : repr (1 : coset_of) = 1.
+Proof. exact: repr_group. Qed.
+
+Lemma coset_reprK : cancel (fun xbar => repr xbar) coset.
+Proof. by move=> xbar; exact: coset_mem (mem_repr_coset xbar). Qed.
+
+(* cosetP is slightly stronger than using repr because we only *)
+(* guarantee repr xbar \in 'N(A) when A is a group. *)
+Lemma cosetP xbar : {x | x \in 'N(A) & xbar = coset x}.
+Proof.
+pose x := repr 'N_xbar(A).
+have [xbar_x Nx]: x \in xbar /\ x \in 'N(A).
+ apply/setIP; rewrite {}/x; case: xbar => /= _ /rcosetsP[y Ny ->].
+ by apply: (mem_repr y); rewrite inE rcoset_refl.
+by exists x; last rewrite (coset_mem xbar_x).
+Qed.
+
+Lemma coset_id x : x \in A -> coset x = 1.
+Proof. by move=> Ax; apply: coset_mem; exact: mem_gen. Qed.
+
+Lemma im_coset : coset @* 'N(A) = setT.
+Proof.
+by apply/setP=> xbar; case: (cosetP xbar) => x Nx ->; rewrite inE mem_morphim.
+Qed.
+
+Lemma sub_im_coset (C : {set coset_of}) : C \subset coset @* 'N(A).
+Proof. by rewrite im_coset subsetT. Qed.
+
+Lemma cosetpre_proper C D :
+ (coset @*^-1 C \proper coset @*^-1 D) = (C \proper D).
+Proof. by rewrite morphpre_proper ?sub_im_coset. Qed.
+
+Definition quotient : {set coset_of} := coset @* Q.
+
+Lemma quotientE : quotient = coset @* Q. Proof. by []. Qed.
+
+End Cosets.
+
+Arguments Scope coset_of [_ group_scope].
+Arguments Scope coset [_ group_scope group_scope].
+Arguments Scope quotient [_ group_scope group_scope].
+Prenex Implicits coset_of coset.
+
+Bind Scope group_scope with coset_of.
+
+Notation "A / B" := (quotient A B) : group_scope.
+
+Section CosetOfGroupTheory.
+
+Variables (gT : finGroupType) (H : {group gT}).
+Implicit Types (A B : {set gT}) (G K : {group gT}) (xbar yb : coset_of H).
+Implicit Types (C D : {set coset_of H}) (L M : {group coset_of H}).
+
+Canonical quotient_group G A : {group coset_of A} :=
+ Eval hnf in [group of G / A].
+
+Infix "/" := quotient_group : Group_scope.
+
+Lemma val_coset x : x \in 'N(H) -> coset H x :=: H :* x.
+Proof. by move=> Nx; rewrite val_coset_prim // genGid. Qed.
+
+Lemma coset_default x : (x \in 'N(H)) = false -> coset H x = 1.
+Proof.
+move=> Nx; apply: val_inj.
+by rewrite val_insubd /= mem_rcosets /= genGid mulSGid ?normG ?Nx.
+Qed.
+
+Lemma coset_norm xbar : xbar \subset 'N(H).
+Proof.
+case: xbar => /= _ /rcosetsP[x Nx ->].
+by rewrite genGid mul_subG ?sub1set ?normG.
+Qed.
+
+Lemma ker_coset : 'ker (coset H) = H.
+Proof. by rewrite ker_coset_prim genGid (setIidPl _) ?normG. Qed.
+
+Lemma coset_idr x : x \in 'N(H) -> coset H x = 1 -> x \in H.
+Proof. by move=> Nx Hx1; rewrite -ker_coset mem_morphpre //= Hx1 set11. Qed.
+
+Lemma repr_coset_norm xbar : repr xbar \in 'N(H).
+Proof. exact: subsetP (coset_norm _) _ (mem_repr_coset _). Qed.
+
+Lemma imset_coset G : coset H @: G = G / H.
+Proof.
+apply/eqP; rewrite eqEsubset andbC imsetS ?subsetIr //=.
+apply/subsetP=> _ /imsetP[x Gx ->].
+by case Nx: (x \in 'N(H)); rewrite ?(coset_default Nx) ?mem_morphim ?group1.
+Qed.
+
+Lemma val_quotient A : val @: (A / H) = rcosets H 'N_A(H).
+Proof.
+apply/setP=> B; apply/imsetP/rcosetsP=> [[xbar Axbar]|[x /setIP[Ax Nx]]] ->{B}.
+ case/morphimP: Axbar => x Nx Ax ->{xbar}.
+ by exists x; [rewrite inE Ax | rewrite /= val_coset].
+by exists (coset H x); [apply/morphimP; exists x | rewrite /= val_coset].
+Qed.
+
+Lemma card_quotient_subnorm A : #|A / H| = #|'N_A(H) : H|.
+Proof. by rewrite -(card_imset _ val_inj) val_quotient. Qed.
+
+Lemma leq_quotient A : #|A / H| <= #|A|.
+Proof. exact: leq_morphim. Qed.
+
+Lemma ltn_quotient A : H :!=: 1 -> H \subset A -> #|A / H| < #|A|.
+Proof.
+by move=> ntH sHA; rewrite ltn_morphim // ker_coset (setIidPr sHA) proper1G.
+Qed.
+
+Lemma card_quotient A : A \subset 'N(H) -> #|A / H| = #|A : H|.
+Proof. by move=> nHA; rewrite card_quotient_subnorm (setIidPl nHA). Qed.
+
+Lemma divg_normal G : H <| G -> #|G| %/ #|H| = #|G / H|.
+Proof. by case/andP=> sHG nHG; rewrite divgS ?card_quotient. Qed.
+
+(* Specializing all the morphisms lemmas that have different assumptions *)
+(* (e.g., because 'ker (coset H) = H), or conclusions (e.g., because we use *)
+(* A / H rather than coset H @* A). We may want to reevaluate later, and *)
+(* eliminate variants that aren't used . *)
+
+(* Variant of morph1; no specialization for other morph lemmas. *)
+Lemma coset1 : coset H 1 :=: H.
+Proof. by rewrite morph1 /= genGid. Qed.
+
+(* Variant of kerE. *)
+Lemma cosetpre1 : coset H @*^-1 1 = H.
+Proof. by rewrite -kerE ker_coset. Qed.
+
+(* Variant of morphimEdom; mophimE[sub] covered by imset_coset. *)
+(* morph[im|pre]Iim are also covered by im_quotient. *)
+Lemma im_quotient : 'N(H) / H = setT.
+Proof. exact: im_coset. Qed.
+
+Lemma quotientT : setT / H = setT.
+Proof. by rewrite -im_quotient; exact: morphimT. Qed.
+
+(* Variant of morphimIdom. *)
+Lemma quotientInorm A : 'N_A(H) / H = A / H.
+Proof. by rewrite /quotient setIC morphimIdom. Qed.
+
+Lemma quotient_setIpre A D : (A :&: coset H @*^-1 D) / H = A / H :&: D.
+Proof. exact: morphim_setIpre. Qed.
+
+Lemma mem_quotient x G : x \in G -> coset H x \in G / H.
+Proof. by move=> Gx; rewrite -imset_coset mem_imset. Qed.
+
+Lemma quotientS A B : A \subset B -> A / H \subset B / H.
+Proof. exact: morphimS. Qed.
+
+Lemma quotient0 : set0 / H = set0.
+Proof. exact: morphim0. Qed.
+
+Lemma quotient_set1 x : x \in 'N(H) -> [set x] / H = [set coset H x].
+Proof. exact: morphim_set1. Qed.
+
+Lemma quotient1 : 1 / H = 1.
+Proof. exact: morphim1. Qed.
+
+Lemma quotientV A : A^-1 / H = (A / H)^-1.
+Proof. exact: morphimV. Qed.
+
+Lemma quotientMl A B : A \subset 'N(H) -> A * B / H = (A / H) * (B / H).
+Proof. exact: morphimMl. Qed.
+
+Lemma quotientMr A B : B \subset 'N(H) -> A * B / H = (A / H) * (B / H).
+Proof. exact: morphimMr. Qed.
+
+Lemma cosetpreM C D : coset H @*^-1 (C * D) = coset H @*^-1 C * coset H @*^-1 D.
+Proof. by rewrite morphpreMl ?sub_im_coset. Qed.
+
+Lemma quotientJ A x : x \in 'N(H) -> A :^ x / H = (A / H) :^ coset H x.
+Proof. exact: morphimJ. Qed.
+
+Lemma quotientU A B : (A :|: B) / H = A / H :|: B / H.
+Proof. exact: morphimU. Qed.
+
+Lemma quotientI A B : (A :&: B) / H \subset A / H :&: B / H.
+Proof. exact: morphimI. Qed.
+
+Lemma quotientY A B :
+ A \subset 'N(H) -> B \subset 'N(H) -> (A <*> B) / H = (A / H) <*> (B / H).
+Proof. exact: morphimY. Qed.
+
+Lemma quotient_homg A : A \subset 'N(H) -> homg (A / H) A.
+Proof. exact: morphim_homg. Qed.
+
+Lemma coset_kerl x y : x \in H -> coset H (x * y) = coset H y.
+Proof.
+move=> Hx; case Ny: (y \in 'N(H)); first by rewrite mkerl ?ker_coset.
+by rewrite !coset_default ?groupMl // (subsetP (normG H)).
+Qed.
+
+Lemma coset_kerr x y : y \in H -> coset H (x * y) = coset H x.
+Proof.
+move=> Hy; case Nx: (x \in 'N(H)); first by rewrite mkerr ?ker_coset.
+by rewrite !coset_default ?groupMr // (subsetP (normG H)).
+Qed.
+
+Lemma rcoset_kercosetP x y :
+ x \in 'N(H) -> y \in 'N(H) -> reflect (coset H x = coset H y) (x \in H :* y).
+Proof. rewrite -{6}ker_coset; exact: rcoset_kerP. Qed.
+
+Lemma kercoset_rcoset x y :
+ x \in 'N(H) -> y \in 'N(H) ->
+ coset H x = coset H y -> exists2 z, z \in H & x = z * y.
+Proof. by move=> Nx Ny eqfxy; rewrite -ker_coset; exact: ker_rcoset. Qed.
+
+Lemma quotientGI G A : H \subset G -> (G :&: A) / H = G / H :&: A / H.
+Proof. by rewrite -{1}ker_coset; exact: morphimGI. Qed.
+
+Lemma quotientIG A G : H \subset G -> (A :&: G) / H = A / H :&: G / H.
+Proof. by rewrite -{1}ker_coset; exact: morphimIG. Qed.
+
+Lemma quotientD A B : A / H :\: B / H \subset (A :\: B) / H.
+Proof. exact: morphimD. Qed.
+
+Lemma quotientD1 A : (A / H)^# \subset A^# / H.
+Proof. exact: morphimD1. Qed.
+
+Lemma quotientDG A G : H \subset G -> (A :\: G) / H = A / H :\: G / H.
+Proof. by rewrite -{1}ker_coset; exact: morphimDG. Qed.
+
+Lemma quotientK A : A \subset 'N(H) -> coset H @*^-1 (A / H) = H * A.
+Proof. by rewrite -{8}ker_coset; exact: morphimK. Qed.
+
+Lemma quotientYK G : G \subset 'N(H) -> coset H @*^-1 (G / H) = H <*> G.
+Proof. by move=> nHG; rewrite quotientK ?norm_joinEr. Qed.
+
+Lemma quotientGK G : H <| G -> coset H @*^-1 (G / H) = G.
+Proof. by case/andP; rewrite -{1}ker_coset; exact: morphimGK. Qed.
+
+Lemma quotient_class x A :
+ x \in 'N(H) -> A \subset 'N(H) -> x ^: A / H = coset H x ^: (A / H).
+Proof. exact: morphim_class. Qed.
+
+Lemma classes_quotient A :
+ A \subset 'N(H) -> classes (A / H) = [set xA / H | xA in classes A].
+Proof. exact: classes_morphim. Qed.
+
+Lemma cosetpre_set1 x :
+ x \in 'N(H) -> coset H @*^-1 [set coset H x] = H :* x.
+Proof. by rewrite -{9}ker_coset; exact: morphpre_set1. Qed.
+
+Lemma cosetpre_set1_coset xbar : coset H @*^-1 [set xbar] = xbar.
+Proof. by case: (cosetP xbar) => x Nx ->; rewrite cosetpre_set1 ?val_coset. Qed.
+
+Lemma cosetpreK C : coset H @*^-1 C / H = C.
+Proof. by rewrite /quotient morphpreK ?sub_im_coset. Qed.
+
+(* Variant of morhphim_ker *)
+Lemma trivg_quotient : H / H = 1.
+Proof. by rewrite -{3}ker_coset /quotient morphim_ker. Qed.
+
+Lemma quotientS1 G : G \subset H -> G / H = 1.
+Proof. by move=> sGH; apply/trivgP; rewrite -trivg_quotient quotientS. Qed.
+
+Lemma sub_cosetpre M : H \subset coset H @*^-1 M.
+Proof. by rewrite -{1}ker_coset; exact: ker_sub_pre. Qed.
+
+Lemma quotient_proper G K :
+ H <| G -> H <| K -> (G / H \proper K / H) = (G \proper K).
+Proof. by move=> nHG nHK; rewrite -cosetpre_proper ?quotientGK. Qed.
+
+Lemma normal_cosetpre M : H <| coset H @*^-1 M.
+Proof. rewrite -{1}ker_coset; exact: ker_normal_pre. Qed.
+
+Lemma cosetpreSK C D :
+ (coset H @*^-1 C \subset coset H @*^-1 D) = (C \subset D).
+Proof. by rewrite morphpreSK ?sub_im_coset. Qed.
+
+Lemma sub_quotient_pre A C :
+ A \subset 'N(H) -> (A / H \subset C) = (A \subset coset H @*^-1 C).
+Proof. exact: sub_morphim_pre. Qed.
+
+Lemma sub_cosetpre_quo C G :
+ H <| G -> (coset H @*^-1 C \subset G) = (C \subset G / H).
+Proof. by move=> nHG; rewrite -cosetpreSK quotientGK. Qed.
+
+(* Variant of ker_trivg_morphim. *)
+Lemma quotient_sub1 A : A \subset 'N(H) -> (A / H \subset [1]) = (A \subset H).
+Proof. by move=> nHA /=; rewrite -{10}ker_coset ker_trivg_morphim nHA. Qed.
+
+Lemma quotientSK A B :
+ A \subset 'N(H) -> (A / H \subset B / H) = (A \subset H * B).
+Proof. by move=> nHA; rewrite morphimSK ?ker_coset. Qed.
+
+Lemma quotientSGK A G :
+ A \subset 'N(H) -> H \subset G -> (A / H \subset G / H) = (A \subset G).
+Proof. by rewrite -{2}ker_coset; exact: morphimSGK. Qed.
+
+Lemma quotient_injG :
+ {in [pred G : {group gT} | H <| G] &, injective (fun G => G / H)}.
+Proof. by rewrite /normal -{1}ker_coset; exact: morphim_injG. Qed.
+
+Lemma quotient_inj G1 G2 :
+ H <| G1 -> H <| G2 -> G1 / H = G2 / H -> G1 :=: G2.
+Proof. by rewrite /normal -{1 3}ker_coset; exact: morphim_inj. Qed.
+
+Lemma quotient_neq1 A : H <| A -> (A / H != 1) = (H \proper A).
+Proof.
+case/andP=> sHA nHA; rewrite /proper sHA -trivg_quotient eqEsubset andbC.
+by rewrite quotientS //= quotientSGK.
+Qed.
+
+Lemma quotient_gen A : A \subset 'N(H) -> <<A>> / H = <<A / H>>.
+Proof. exact: morphim_gen. Qed.
+
+Lemma cosetpre_gen C :
+ 1 \in C -> coset H @*^-1 <<C>> = <<coset H @*^-1 C>>.
+Proof. by move=> C1; rewrite morphpre_gen ?sub_im_coset. Qed.
+
+Lemma quotientR A B :
+ A \subset 'N(H) -> B \subset 'N(H) -> [~: A, B] / H = [~: A / H, B / H].
+Proof. exact: morphimR. Qed.
+
+Lemma quotient_norm A : 'N(A) / H \subset 'N(A / H).
+Proof. exact: morphim_norm. Qed.
+
+Lemma quotient_norms A B : A \subset 'N(B) -> A / H \subset 'N(B / H).
+Proof. exact: morphim_norms. Qed.
+
+Lemma quotient_subnorm A B : 'N_A(B) / H \subset 'N_(A / H)(B / H).
+Proof. exact: morphim_subnorm. Qed.
+
+Lemma quotient_normal A B : A <| B -> A / H <| B / H.
+Proof. exact: morphim_normal. Qed.
+
+Lemma quotient_cent1 x : 'C[x] / H \subset 'C[coset H x].
+Proof.
+case Nx: (x \in 'N(H)); first exact: morphim_cent1.
+by rewrite coset_default // cent11T subsetT.
+Qed.
+
+Lemma quotient_cent1s A x : A \subset 'C[x] -> A / H \subset 'C[coset H x].
+Proof.
+by move=> sAC; exact: subset_trans (quotientS sAC) (quotient_cent1 x).
+Qed.
+
+Lemma quotient_subcent1 A x : 'C_A[x] / H \subset 'C_(A / H)[coset H x].
+Proof. exact: subset_trans (quotientI _ _) (setIS _ (quotient_cent1 x)). Qed.
+
+Lemma quotient_cent A : 'C(A) / H \subset 'C(A / H).
+Proof. exact: morphim_cent. Qed.
+
+Lemma quotient_cents A B : A \subset 'C(B) -> A / H \subset 'C(B / H).
+Proof. exact: morphim_cents. Qed.
+
+Lemma quotient_abelian A : abelian A -> abelian (A / H).
+Proof. exact: morphim_abelian. Qed.
+
+Lemma quotient_subcent A B : 'C_A(B) / H \subset 'C_(A / H)(B / H).
+Proof. exact: morphim_subcent. Qed.
+
+Lemma norm_quotient_pre A C :
+ A \subset 'N(H) -> A / H \subset 'N(C) -> A \subset 'N(coset H @*^-1 C).
+Proof.
+by move/sub_quotient_pre=> -> /subset_trans-> //; exact: morphpre_norm.
+Qed.
+
+Lemma cosetpre_normal C D : (coset H @*^-1 C <| coset H @*^-1 D) = (C <| D).
+Proof. by rewrite morphpre_normal ?sub_im_coset. Qed.
+
+Lemma quotient_normG G : H <| G -> 'N(G) / H = 'N(G / H).
+Proof.
+case/andP=> sHG nHG.
+by rewrite [_ / _]morphim_normG ?ker_coset // im_coset setTI.
+Qed.
+
+Lemma quotient_subnormG A G : H <| G -> 'N_A(G) / H = 'N_(A / H)(G / H).
+Proof. by case/andP=> sHG nHG; rewrite -morphim_subnormG ?ker_coset. Qed.
+
+Lemma cosetpre_cent1 x : 'C_('N(H))[x] \subset coset H @*^-1 'C[coset H x].
+Proof.
+case Nx: (x \in 'N(H)); first by rewrite morphpre_cent1.
+by rewrite coset_default // cent11T morphpreT subsetIl.
+Qed.
+
+Lemma cosetpre_cent1s C x :
+ coset H @*^-1 C \subset 'C[x] -> C \subset 'C[coset H x].
+Proof.
+move=> sC; rewrite -cosetpreSK; apply: subset_trans (cosetpre_cent1 x).
+by rewrite subsetI subsetIl.
+Qed.
+
+Lemma cosetpre_subcent1 C x :
+ 'C_(coset H @*^-1 C)[x] \subset coset H @*^-1 'C_C[coset H x].
+Proof.
+by rewrite -morphpreIdom -setIA setICA morphpreI setIS // cosetpre_cent1.
+Qed.
+
+Lemma cosetpre_cent A : 'C_('N(H))(A) \subset coset H @*^-1 'C(A / H).
+Proof. exact: morphpre_cent. Qed.
+
+Lemma cosetpre_cents A C : coset H @*^-1 C \subset 'C(A) -> C \subset 'C(A / H).
+Proof. by apply: morphpre_cents; rewrite ?sub_im_coset. Qed.
+
+Lemma cosetpre_subcent C A :
+ 'C_(coset H @*^-1 C)(A) \subset coset H @*^-1 'C_C(A / H).
+Proof. exact: morphpre_subcent. Qed.
+
+Lemma restrm_quotientE G A (nHG : G \subset 'N(H)) :
+ A \subset G -> restrm nHG (coset H) @* A = A / H.
+Proof. exact: restrmEsub. Qed.
+
+Section InverseImage.
+
+Variables (G : {group gT}) (Kbar : {group coset_of H}).
+
+Hypothesis nHG : H <| G.
+
+CoInductive inv_quotient_spec (P : pred {group gT}) : Prop :=
+ InvQuotientSpec K of Kbar :=: K / H & H \subset K & P K.
+
+Lemma inv_quotientS :
+ Kbar \subset G / H -> inv_quotient_spec (fun K => K \subset G).
+Proof.
+case/andP: nHG => sHG nHG' sKbarG.
+have sKdH: Kbar \subset 'N(H) / H by rewrite (subset_trans sKbarG) ?morphimS.
+exists (coset H @*^-1 Kbar)%G; first by rewrite cosetpreK.
+ by rewrite -{1}ker_coset morphpreS ?sub1G.
+by rewrite sub_cosetpre_quo.
+Qed.
+
+Lemma inv_quotientN : Kbar <| G / H -> inv_quotient_spec (fun K => K <| G).
+Proof.
+move=> nKbar; case/inv_quotientS: (normal_sub nKbar) => K defKbar sHK sKG.
+exists K => //; rewrite defKbar -cosetpre_normal !quotientGK // in nKbar.
+exact: normalS nHG.
+Qed.
+
+End InverseImage.
+
+Lemma quotientMidr A : A * H / H = A / H.
+Proof.
+by rewrite [_ /_]morphimMr ?normG //= -!quotientE trivg_quotient mulg1.
+Qed.
+
+Lemma quotientMidl A : H * A / H = A / H.
+Proof.
+by rewrite [_ /_]morphimMl ?normG //= -!quotientE trivg_quotient mul1g.
+Qed.
+
+Lemma quotientYidr G : G \subset 'N(H) -> G <*> H / H = G / H.
+Proof.
+move=> nHG; rewrite -genM_join quotient_gen ?mul_subG ?normG //.
+by rewrite quotientMidr genGid.
+Qed.
+
+Lemma quotientYidl G : G \subset 'N(H) -> H <*> G / H = G / H.
+Proof. by move=> nHG; rewrite joingC quotientYidr. Qed.
+
+Section Injective.
+
+Variables (G : {group gT}).
+Hypotheses (nHG : G \subset 'N(H)) (tiHG : H :&: G = 1).
+
+Lemma quotient_isom : isom G (G / H) (restrm nHG (coset H)).
+Proof. by apply/isomP; rewrite ker_restrm setIC ker_coset tiHG im_restrm. Qed.
+
+Lemma quotient_isog : isog G (G / H).
+Proof. exact: isom_isog quotient_isom. Qed.
+
+End Injective.
+
+End CosetOfGroupTheory.
+
+Notation "A / H" := (quotient_group A H) : Group_scope.
+
+Section Quotient1.
+
+Variables (gT : finGroupType) (A : {set gT}).
+
+Lemma coset1_injm : 'injm (@coset gT 1).
+Proof. by rewrite ker_coset /=. Qed.
+
+Lemma quotient1_isom : isom A (A / 1) (coset 1).
+Proof. by apply: sub_isom coset1_injm; rewrite ?norms1. Qed.
+
+Lemma quotient1_isog : isog A (A / 1).
+Proof. apply: isom_isog quotient1_isom; exact: norms1. Qed.
+
+End Quotient1.
+
+Section QuotientMorphism.
+
+Variable (gT rT : finGroupType) (G H : {group gT}) (f : {morphism G >-> rT}).
+
+Implicit Types A : {set gT}.
+Implicit Types B : {set (coset_of H)}.
+Hypotheses (nsHG : H <| G).
+Let sHG : H \subset G := normal_sub nsHG.
+Let nHG : G \subset 'N(H) := normal_norm nsHG.
+Let nfHfG : f @* G \subset 'N(f @* H) := morphim_norms f nHG.
+
+Notation fH := (coset (f @* H) \o f).
+
+Lemma quotm_dom_proof : G \subset 'dom fH.
+Proof. by rewrite -sub_morphim_pre. Qed.
+
+Notation fH_G := (restrm quotm_dom_proof fH).
+
+Lemma quotm_ker_proof : 'ker (coset H) \subset 'ker fH_G.
+Proof.
+by rewrite ker_restrm ker_comp !ker_coset morphpreIdom morphimK ?mulG_subr.
+Qed.
+
+Definition quotm := factm quotm_ker_proof nHG.
+
+Canonical quotm_morphism := [morphism G / H of quotm].
+
+Lemma quotmE x : x \in G -> quotm (coset H x) = coset (f @* H) (f x).
+Proof. exact: factmE. Qed.
+
+Lemma morphim_quotm A : quotm @* (A / H) = f @* A / f @* H.
+Proof. by rewrite morphim_factm morphim_restrm morphim_comp morphimIdom. Qed.
+
+Lemma morphpre_quotm Abar : quotm @*^-1 (Abar / f @* H) = f @*^-1 Abar / H.
+Proof.
+rewrite morphpre_factm morphpre_restrm morphpre_comp /=.
+rewrite morphpreIdom -[Abar / _]quotientInorm quotientK ?subsetIr //=.
+rewrite morphpreMl ?morphimS // morphimK // [_ * H]normC ?subIset ?nHG //.
+rewrite -quotientE -mulgA quotientMidl /= setIC -morphpreIim setIA.
+by rewrite (setIidPl nfHfG) morphpreIim -morphpreMl ?sub1G ?mul1g.
+Qed.
+
+Lemma ker_quotm : 'ker quotm = 'ker f / H.
+Proof. by rewrite -morphpre_quotm /quotient morphim1. Qed.
+
+Lemma injm_quotm : 'injm f -> 'injm quotm.
+Proof. by move/trivgP=> /= kf1; rewrite ker_quotm kf1 quotientE morphim1. Qed.
+
+End QuotientMorphism.
+
+Section EqIso.
+
+Variables (gT : finGroupType) (G H : {group gT}).
+
+Hypothesis (eqGH : G :=: H).
+
+Lemma im_qisom_proof : 'N(H) \subset 'N(G). Proof. by rewrite eqGH. Qed.
+Lemma qisom_ker_proof : 'ker (coset G) \subset 'ker (coset H).
+Proof. by rewrite eqGH. Qed.
+Lemma qisom_restr_proof : setT \subset 'N(H) / G.
+Proof. by rewrite eqGH im_quotient. Qed.
+
+Definition qisom :=
+ restrm qisom_restr_proof (factm qisom_ker_proof im_qisom_proof).
+
+Canonical qisom_morphism := Eval hnf in [morphism of qisom].
+
+Lemma qisomE x : qisom (coset G x) = coset H x.
+Proof.
+case Nx: (x \in 'N(H)); first exact: factmE.
+by rewrite !coset_default ?eqGH ?morph1.
+Qed.
+
+Lemma val_qisom Gx : val (qisom Gx) = val Gx.
+Proof.
+by case: (cosetP Gx) => x Nx ->{Gx}; rewrite qisomE /= !val_coset -?eqGH.
+Qed.
+
+Lemma morphim_qisom A : qisom @* (A / G) = A / H.
+Proof. by rewrite morphim_restrm setTI morphim_factm. Qed.
+
+Lemma morphpre_qisom A : qisom @*^-1 (A / H) = A / G.
+Proof.
+rewrite morphpre_restrm setTI morphpre_factm eqGH.
+by rewrite morphpreK // im_coset subsetT.
+Qed.
+
+Lemma injm_qisom : 'injm qisom.
+Proof. by rewrite -quotient1 -morphpre_qisom morphpreS ?sub1G. Qed.
+
+Lemma im_qisom : qisom @* setT = setT.
+Proof. by rewrite -{2}im_quotient morphim_qisom eqGH im_quotient. Qed.
+
+Lemma qisom_isom : isom setT setT qisom.
+Proof. by apply/isomP; rewrite injm_qisom im_qisom. Qed.
+
+Lemma qisom_isog : [set: coset_of G] \isog [set: coset_of H].
+Proof. exact: isom_isog qisom_isom. Qed.
+
+Lemma qisom_inj : injective qisom.
+Proof. by move=> x y; apply: (injmP injm_qisom); rewrite inE. Qed.
+
+Lemma morphim_qisom_inj : injective (fun Gx => qisom @* Gx).
+Proof.
+by move=> Gx Gy; apply: injm_morphim_inj; rewrite (injm_qisom, subsetT).
+Qed.
+
+End EqIso.
+
+Implicit Arguments qisom_inj [gT G H].
+Implicit Arguments morphim_qisom_inj [gT G H].
+
+Section FirstIsomorphism.
+
+Variables aT rT : finGroupType.
+
+Lemma first_isom (G : {group aT}) (f : {morphism G >-> rT}) :
+ {g : {morphism G / 'ker f >-> rT} | 'injm g &
+ forall A : {set aT}, g @* (A / 'ker f) = f @* A}.
+Proof.
+have nkG := ker_norm f.
+have skk: 'ker (coset ('ker f)) \subset 'ker f by rewrite ker_coset.
+exists (factm_morphism skk nkG) => /=; last exact: morphim_factm.
+by rewrite ker_factm -quotientE trivg_quotient.
+Qed.
+
+Variables (G H : {group aT}) (f : {morphism G >-> rT}).
+Hypothesis sHG : H \subset G.
+
+Lemma first_isog : (G / 'ker f) \isog (f @* G).
+Proof.
+by case: (first_isom f) => g injg im_g; apply/isogP; exists g; rewrite ?im_g.
+Qed.
+
+Lemma first_isom_loc : {g : {morphism H / 'ker_H f >-> rT} |
+ 'injm g & forall A : {set aT}, A \subset H -> g @* (A / 'ker_H f) = f @* A}.
+Proof.
+case: (first_isom (restrm_morphism sHG f)).
+rewrite ker_restrm => g injg im_g; exists g => // A sAH.
+by rewrite im_g morphim_restrm (setIidPr sAH).
+Qed.
+
+Lemma first_isog_loc : (H / 'ker_H f) \isog (f @* H).
+Proof.
+by case: first_isom_loc => g injg im_g; apply/isogP; exists g; rewrite ?im_g.
+Qed.
+
+End FirstIsomorphism.
+
+Section SecondIsomorphism.
+
+Variables (gT : finGroupType) (H K : {group gT}).
+
+Hypothesis nKH : H \subset 'N(K).
+
+Lemma second_isom : {f : {morphism H / (K :&: H) >-> coset_of K} |
+ 'injm f & forall A : {set gT}, A \subset H -> f @* (A / (K :&: H)) = A / K}.
+Proof.
+have ->: K :&: H = 'ker_H (coset K) by rewrite ker_coset setIC.
+exact: first_isom_loc.
+Qed.
+
+Lemma second_isog : H / (K :&: H) \isog H / K.
+Proof. by rewrite setIC -{1 3}(ker_coset K); exact: first_isog_loc. Qed.
+
+Lemma weak_second_isog : H / (K :&: H) \isog H * K / K.
+Proof. by rewrite quotientMidr; exact: second_isog. Qed.
+
+End SecondIsomorphism.
+
+Section ThirdIsomorphism.
+
+Variables (gT : finGroupType) (G H K : {group gT}).
+
+Lemma homg_quotientS (A : {set gT}) :
+ A \subset 'N(H) -> A \subset 'N(K) -> H \subset K -> A / K \homg A / H.
+Proof.
+rewrite -!(gen_subG A) /=; set L := <<A>> => nHL nKL sKH.
+have sub_ker: 'ker (restrm nHL (coset H)) \subset 'ker (restrm nKL (coset K)).
+ by rewrite !ker_restrm !ker_coset setIS.
+have sAL: A \subset L := subset_gen A; rewrite -(setIidPr sAL).
+rewrite -[_ / H](morphim_restrm nHL) -[_ / K](morphim_restrm nKL) /=.
+by rewrite -(morphim_factm sub_ker (subxx L)) morphim_homg ?morphimS.
+Qed.
+
+Hypothesis sHK : H \subset K.
+Hypothesis snHG : H <| G.
+Hypothesis snKG : K <| G.
+
+Theorem third_isom : {f : {morphism (G / H) / (K / H) >-> coset_of K} | 'injm f
+ & forall A : {set gT}, A \subset G -> f @* (A / H / (K / H)) = A / K}.
+Proof.
+have [[sKG nKG] [sHG nHG]] := (andP snKG, andP snHG).
+have sHker: 'ker (coset H) \subset 'ker (restrm nKG (coset K)).
+ by rewrite ker_restrm !ker_coset subsetI sHG.
+have:= first_isom_loc (factm_morphism sHker nHG) (subxx _) => /=.
+rewrite ker_factm_loc ker_restrm ker_coset !(setIidPr sKG) /= -!quotientE.
+case=> f injf im_f; exists f => // A sAG; rewrite im_f ?morphimS //.
+by rewrite morphim_factm morphim_restrm (setIidPr sAG).
+Qed.
+
+Theorem third_isog : (G / H / (K / H)) \isog (G / K).
+Proof.
+by case: third_isom => f inj_f im_f; apply/isogP; exists f; rewrite ?im_f.
+Qed.
+
+End ThirdIsomorphism.
+
+Lemma char_from_quotient (gT : finGroupType) (G H K : {group gT}) :
+ H <| K -> H \char G -> K / H \char G / H -> K \char G.
+Proof.
+case/andP=> sHK nHK chHG.
+have nsHG := char_normal chHG; have [sHG nHG] := andP nsHG.
+case/charP; rewrite quotientSGK // => sKG /= chKG.
+apply/charP; split=> // f injf Gf; apply/morphim_fixP => //.
+rewrite -(quotientSGK _ sHK); last by rewrite -morphimIim Gf subIset ?nHG.
+have{chHG} Hf: f @* H = H by case/charP: chHG => _; apply.
+set q := quotm_morphism f nsHG; have{injf}: 'injm q by exact: injm_quotm.
+have: q @* _ = _ := morphim_quotm _ _ _; move: q; rewrite Hf => q im_q injq.
+by rewrite -im_q chKG // im_q Gf.
+Qed.
+
+(* Counting lemmas for morphisms. *)
+
+Section CardMorphism.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+Implicit Types G H : {group aT}.
+Implicit Types L M : {group rT}.
+
+Lemma card_morphim G : #|f @* G| = #|D :&: G : 'ker f|.
+Proof.
+rewrite -morphimIdom -indexgI -card_quotient; last first.
+ by rewrite normsI ?normG ?subIset ?ker_norm.
+by apply: esym (card_isog _); rewrite first_isog_loc ?subsetIl.
+Qed.
+
+Lemma dvdn_morphim G : #|f @* G| %| #|G|.
+Proof.
+rewrite card_morphim (dvdn_trans (dvdn_indexg _ _)) //.
+by rewrite cardSg ?subsetIr.
+Qed.
+
+Lemma logn_morphim p G : logn p #|f @* G| <= logn p #|G|.
+Proof. by rewrite dvdn_leq_log ?dvdn_morphim. Qed.
+
+Lemma coprime_morphl G p : coprime #|G| p -> coprime #|f @* G| p.
+Proof. exact: coprime_dvdl (dvdn_morphim G). Qed.
+
+Lemma coprime_morphr G p : coprime p #|G| -> coprime p #|f @* G|.
+Proof. exact: coprime_dvdr (dvdn_morphim G). Qed.
+
+Lemma coprime_morph G H : coprime #|G| #|H| -> coprime #|f @* G| #|f @* H|.
+Proof. by move=> coGH; rewrite coprime_morphl // coprime_morphr. Qed.
+
+Lemma index_morphim_ker G H :
+ H \subset G -> G \subset D ->
+ (#|f @* G : f @* H| * #|'ker_G f : H|)%N = #|G : H|.
+Proof.
+move=> sHG sGD; apply/eqP.
+rewrite -(eqn_pmul2l (cardG_gt0 (f @* H))) mulnA Lagrange ?morphimS //.
+rewrite !card_morphim (setIidPr sGD) (setIidPr (subset_trans sHG sGD)).
+rewrite -(eqn_pmul2l (cardG_gt0 ('ker_H f))) /=.
+by rewrite -{1}(setIidPr sHG) setIAC mulnCA mulnC mulnA !LagrangeI Lagrange.
+Qed.
+
+Lemma index_morphim G H : G :&: H \subset D -> #|f @* G : f @* H| %| #|G : H|.
+Proof.
+move=> dGH; rewrite -(indexgI G) -(setIidPr dGH) setIA.
+apply: dvdn_trans (indexSg (subsetIl _ H) (subsetIr D G)).
+rewrite -index_morphim_ker ?subsetIl ?subsetIr ?dvdn_mulr //= morphimIdom.
+by rewrite indexgS ?morphimS ?subsetIr.
+Qed.
+
+Lemma index_injm G H : 'injm f -> G \subset D -> #|f @* G : f @* H| = #|G : H|.
+Proof.
+move=> injf dG; rewrite -{2}(setIidPr dG) -(indexgI _ H) /=.
+rewrite -index_morphim_ker ?subsetIl ?subsetIr //= setIAC morphimIdom setIC.
+rewrite injmI ?subsetIr // indexgI /= morphimIdom setIC ker_injm //.
+by rewrite -(indexgI (1 :&: _)) /= -setIA !(setIidPl (sub1G _)) indexgg muln1.
+Qed.
+
+Lemma card_morphpre L : L \subset f @* D -> #|f @*^-1 L| = (#|'ker f| * #|L|)%N.
+Proof.
+move/morphpreK=> {2} <-; rewrite card_morphim morphpreIdom.
+by rewrite Lagrange // morphpreS ?sub1G.
+Qed.
+
+Lemma index_morphpre L M :
+ L \subset f @* D -> #|f @*^-1 L : f @*^-1 M| = #|L : M|.
+Proof.
+move=> dL; rewrite -!divgI -morphpreI card_morphpre //.
+have: L :&: M \subset f @* D by rewrite subIset ?dL.
+by move/card_morphpre->; rewrite divnMl ?cardG_gt0.
+Qed.
+
+End CardMorphism.
+
+Lemma card_homg (aT rT : finGroupType) (G : {group aT}) (R : {group rT}) :
+ G \homg R -> #|G| %| #|R|.
+Proof. by case/homgP=> f <-; rewrite card_morphim setIid dvdn_indexg. Qed.
+
+Section CardCosetpre.
+
+Variables (gT : finGroupType) (G H K : {group gT}) (L M : {group coset_of H}).
+
+Lemma dvdn_quotient : #|G / H| %| #|G|.
+Proof. exact: dvdn_morphim. Qed.
+
+Lemma index_quotient_ker :
+ K \subset G -> G \subset 'N(H) ->
+ (#|G / H : K / H| * #|G :&: H : K|)%N = #|G : K|.
+Proof. by rewrite -{5}(ker_coset H); exact: index_morphim_ker. Qed.
+
+Lemma index_quotient : G :&: K \subset 'N(H) -> #|G / H : K / H| %| #|G : K|.
+Proof. exact: index_morphim. Qed.
+
+Lemma index_quotient_eq :
+ G :&: H \subset K -> K \subset G -> G \subset 'N(H) ->
+ #|G / H : K / H| = #|G : K|.
+Proof.
+move=> sGH_K sKG sGN; rewrite -index_quotient_ker {sKG sGN}//.
+by rewrite -(indexgI _ K) (setIidPl sGH_K) indexgg muln1.
+Qed.
+
+Lemma card_cosetpre : #|coset H @*^-1 L| = (#|H| * #|L|)%N.
+Proof. by rewrite card_morphpre ?ker_coset ?sub_im_coset. Qed.
+
+Lemma index_cosetpre : #|coset H @*^-1 L : coset H @*^-1 M| = #|L : M|.
+Proof. by rewrite index_morphpre ?sub_im_coset. Qed.
+
+End CardCosetpre.
diff --git a/mathcomp/fingroup/zmodp.v b/mathcomp/fingroup/zmodp.v
new file mode 100644
index 0000000..df5b378
--- /dev/null
+++ b/mathcomp/fingroup/zmodp.v
@@ -0,0 +1,362 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div.
+Require Import fintype bigop finset prime fingroup ssralg finalg.
+
+(******************************************************************************)
+(* Definition of the additive group and ring Zp, represented as 'I_p *)
+(******************************************************************************)
+(* Definitions: *)
+(* From fintype.v: *)
+(* 'I_p == the subtype of integers less than p, taken here as the type of *)
+(* the integers mod p. *)
+(* This file: *)
+(* inZp == the natural projection from nat into the integers mod p, *)
+(* represented as 'I_p. Here p is implicit, but MUST be of the *)
+(* form n.+1. *)
+(* The operations: *)
+(* Zp0 == the identity element for addition *)
+(* Zp1 == the identity element for multiplication, and a generator of *)
+(* additive group *)
+(* Zp_opp == inverse function for addition *)
+(* Zp_add == addition *)
+(* Zp_mul == multiplication *)
+(* Zp_inv == inverse function for multiplication *)
+(* Note that while 'I_n.+1 has canonical finZmodType and finGroupType *)
+(* structures, only 'I_n.+2 has a canonical ring structure (it has, in fact, *)
+(* a canonical finComUnitRing structure), and hence an associated *)
+(* multiplicative unit finGroupType. To mitigate the issues caused by the *)
+(* trivial "ring" (which is, indeed is NOT a ring in the ssralg/finalg *)
+(* formalization), we define additional notation: *)
+(* 'Z_p == the type of integers mod (max p 2); this is always a proper *)
+(* ring, by constructions. Note that 'Z_p is provably equal to *)
+(* 'I_p if p > 1, and convertible to 'I_p if p is of the form *)
+(* n.+2. *)
+(* Zp p == the subgroup of integers mod (max p 1) in 'Z_p; this is thus *)
+(* is thus all of 'Z_p if p > 1, and else the trivial group. *)
+(* units_Zp p == the group of all units of 'Z_p -- i.e., the group of *)
+(* (multiplicative) automorphisms of Zp p. *)
+(* We show that Zp and units_Zp are abelian, and compute their orders. *)
+(* We use a similar technique to represent the prime fields: *)
+(* 'F_p == the finite field of integers mod the first prime divisor of *)
+(* maxn p 2. This is provably equal to 'Z_p and 'I_p if p is *)
+(* provably prime, and indeed convertible to the above if p is *)
+(* a concrete prime such as 2, 5 or 23. *)
+(* Note finally that due to the canonical structures it is possible to use *)
+(* 0%R instead of Zp0, and 1%R instead of Zp1 (for the latter, p must be of *)
+(* the form n.+2, and 1%R : nat will simplify to 1%N). *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Local Open Scope ring_scope.
+
+Section ZpDef.
+
+(***********************************************************************)
+(* *)
+(* Mod p arithmetic on the finite set {0, 1, 2, ..., p - 1} *)
+(* *)
+(***********************************************************************)
+
+Variable p' : nat.
+Local Notation p := p'.+1.
+
+Implicit Types x y z : 'I_p.
+
+(* Standard injection; val (inZp i) = i %% p *)
+Definition inZp i := Ordinal (ltn_pmod i (ltn0Sn p')).
+Lemma modZp x : x %% p = x.
+Proof. by rewrite modn_small ?ltn_ord. Qed.
+Lemma valZpK x : inZp x = x.
+Proof. by apply: val_inj; rewrite /= modZp. Qed.
+
+(* Operations *)
+Definition Zp0 : 'I_p := ord0.
+Definition Zp1 := inZp 1.
+Definition Zp_opp x := inZp (p - x).
+Definition Zp_add x y := inZp (x + y).
+Definition Zp_mul x y := inZp (x * y).
+Definition Zp_inv x := if coprime p x then inZp (egcdn x p).1 else x.
+
+(* Additive group structure. *)
+
+Lemma Zp_add0z : left_id Zp0 Zp_add.
+Proof. exact: valZpK. Qed.
+
+Lemma Zp_addNz : left_inverse Zp0 Zp_opp Zp_add.
+Proof.
+by move=> x; apply: val_inj; rewrite /= modnDml subnK ?modnn // ltnW.
+Qed.
+
+Lemma Zp_addA : associative Zp_add.
+Proof.
+by move=> x y z; apply: val_inj; rewrite /= modnDml modnDmr addnA.
+Qed.
+
+Lemma Zp_addC : commutative Zp_add.
+Proof. by move=> x y; apply: val_inj; rewrite /= addnC. Qed.
+
+Definition Zp_zmodMixin := ZmodMixin Zp_addA Zp_addC Zp_add0z Zp_addNz.
+Canonical Zp_zmodType := Eval hnf in ZmodType 'I_p Zp_zmodMixin.
+Canonical Zp_finZmodType := Eval hnf in [finZmodType of 'I_p].
+Canonical Zp_baseFinGroupType := Eval hnf in [baseFinGroupType of 'I_p for +%R].
+Canonical Zp_finGroupType := Eval hnf in [finGroupType of 'I_p for +%R].
+
+(* Ring operations *)
+
+Lemma Zp_mul1z : left_id Zp1 Zp_mul.
+Proof. by move=> x; apply: val_inj; rewrite /= modnMml mul1n modZp. Qed.
+
+Lemma Zp_mulC : commutative Zp_mul.
+Proof. by move=> x y; apply: val_inj; rewrite /= mulnC. Qed.
+
+Lemma Zp_mulz1 : right_id Zp1 Zp_mul.
+Proof. by move=> x; rewrite Zp_mulC Zp_mul1z. Qed.
+
+Lemma Zp_mulA : associative Zp_mul.
+Proof.
+by move=> x y z; apply: val_inj; rewrite /= modnMml modnMmr mulnA.
+Qed.
+
+Lemma Zp_mul_addr : right_distributive Zp_mul Zp_add.
+Proof.
+by move=> x y z; apply: val_inj; rewrite /= modnMmr modnDm mulnDr.
+Qed.
+
+Lemma Zp_mul_addl : left_distributive Zp_mul Zp_add.
+Proof. by move=> x y z; rewrite -!(Zp_mulC z) Zp_mul_addr. Qed.
+
+Lemma Zp_mulVz x : coprime p x -> Zp_mul (Zp_inv x) x = Zp1.
+Proof.
+move=> co_p_x; apply: val_inj; rewrite /Zp_inv co_p_x /= modnMml.
+by rewrite -(chinese_modl co_p_x 1 0) /chinese addn0 mul1n mulnC.
+Qed.
+
+Lemma Zp_mulzV x : coprime p x -> Zp_mul x (Zp_inv x) = Zp1.
+Proof. by move=> Ux; rewrite /= Zp_mulC Zp_mulVz. Qed.
+
+Lemma Zp_intro_unit x y : Zp_mul y x = Zp1 -> coprime p x.
+Proof.
+case=> yx1; have:= coprimen1 p.
+by rewrite -coprime_modr -yx1 coprime_modr coprime_mulr; case/andP.
+Qed.
+
+Lemma Zp_inv_out x : ~~ coprime p x -> Zp_inv x = x.
+Proof. by rewrite /Zp_inv => /negPf->. Qed.
+
+Lemma Zp_mulrn x n : x *+ n = inZp (x * n).
+Proof.
+apply: val_inj => /=; elim: n => [|n IHn]; first by rewrite muln0 modn_small.
+by rewrite !GRing.mulrS /= IHn modnDmr mulnS.
+Qed.
+
+Import GroupScope.
+
+Lemma Zp_mulgC : @commutative 'I_p _ mulg.
+Proof. exact: Zp_addC. Qed.
+
+Lemma Zp_abelian : abelian [set: 'I_p].
+Proof. exact: FinRing.zmod_abelian. Qed.
+
+Lemma Zp_expg x n : x ^+ n = inZp (x * n).
+Proof. exact: Zp_mulrn. Qed.
+
+Lemma Zp1_expgz x : Zp1 ^+ x = x.
+Proof. by rewrite Zp_expg; exact: Zp_mul1z. Qed.
+
+Lemma Zp_cycle : setT = <[Zp1]>.
+Proof. by apply/setP=> x; rewrite -[x]Zp1_expgz inE groupX ?mem_gen ?set11. Qed.
+
+Lemma order_Zp1 : #[Zp1] = p.
+Proof. by rewrite orderE -Zp_cycle cardsT card_ord. Qed.
+
+End ZpDef.
+
+Implicit Arguments Zp0 [[p']].
+Implicit Arguments Zp1 [[p']].
+Implicit Arguments inZp [[p']].
+
+Lemma ord1 : all_equal_to (0 : 'I_1).
+Proof. by case=> [[] // ?]; exact: val_inj. Qed.
+
+Lemma lshift0 m n : lshift m (0 : 'I_n.+1) = (0 : 'I_(n + m).+1).
+Proof. exact: val_inj. Qed.
+
+Lemma rshift1 n : @rshift 1 n =1 lift (0 : 'I_n.+1).
+Proof. by move=> i; exact: val_inj. Qed.
+
+Lemma split1 n i :
+ split (i : 'I_(1 + n)) = oapp (@inr _ _) (inl _ 0) (unlift 0 i).
+Proof.
+case: unliftP => [i'|] -> /=.
+ by rewrite -rshift1 (unsplitK (inr _ _)).
+by rewrite -(lshift0 n 0) (unsplitK (inl _ _)).
+Qed.
+
+Lemma big_ord1 R idx (op : @Monoid.law R idx) F :
+ \big[op/idx]_(i < 1) F i = F 0.
+Proof. by rewrite big_ord_recl big_ord0 Monoid.mulm1. Qed.
+
+Lemma big_ord1_cond R idx (op : @Monoid.law R idx) P F :
+ \big[op/idx]_(i < 1 | P i) F i = if P 0 then F 0 else idx.
+Proof. by rewrite big_mkcond big_ord1. Qed.
+
+Section ZpRing.
+
+Variable p' : nat.
+Local Notation p := p'.+2.
+
+Lemma Zp_nontrivial : Zp1 != 0 :> 'I_p. Proof. by []. Qed.
+
+Definition Zp_ringMixin :=
+ ComRingMixin (@Zp_mulA _) (@Zp_mulC _) (@Zp_mul1z _) (@Zp_mul_addl _)
+ Zp_nontrivial.
+Canonical Zp_ringType := Eval hnf in RingType 'I_p Zp_ringMixin.
+Canonical Zp_finRingType := Eval hnf in [finRingType of 'I_p].
+Canonical Zp_comRingType := Eval hnf in ComRingType 'I_p (@Zp_mulC _).
+Canonical Zp_finComRingType := Eval hnf in [finComRingType of 'I_p].
+
+Definition Zp_unitRingMixin :=
+ ComUnitRingMixin (@Zp_mulVz _) (@Zp_intro_unit _) (@Zp_inv_out _).
+Canonical Zp_unitRingType := Eval hnf in UnitRingType 'I_p Zp_unitRingMixin.
+Canonical Zp_finUnitRingType := Eval hnf in [finUnitRingType of 'I_p].
+Canonical Zp_comUnitRingType := Eval hnf in [comUnitRingType of 'I_p].
+Canonical Zp_finComUnitRingType := Eval hnf in [finComUnitRingType of 'I_p].
+
+Lemma Zp_nat n : n%:R = inZp n :> 'I_p.
+Proof. by apply: val_inj; rewrite [n%:R]Zp_mulrn /= modnMml mul1n. Qed.
+
+Lemma natr_Zp (x : 'I_p) : x%:R = x.
+Proof. by rewrite Zp_nat valZpK. Qed.
+
+Lemma natr_negZp (x : 'I_p) : (- x)%:R = - x.
+Proof. by apply: val_inj; rewrite /= Zp_nat /= modn_mod. Qed.
+
+Import GroupScope.
+
+Lemma unit_Zp_mulgC : @commutative {unit 'I_p} _ mulg.
+Proof. by move=> u v; apply: val_inj; rewrite /= GRing.mulrC. Qed.
+
+Lemma unit_Zp_expg (u : {unit 'I_p}) n :
+ val (u ^+ n) = inZp (val u ^ n) :> 'I_p.
+Proof.
+apply: val_inj => /=; elim: n => [|n IHn] //.
+by rewrite expgS /= IHn expnS modnMmr.
+Qed.
+
+End ZpRing.
+
+Definition Zp_trunc p := p.-2.
+
+Notation "''Z_' p" := 'I_(Zp_trunc p).+2
+ (at level 8, p at level 2, format "''Z_' p") : type_scope.
+Notation "''F_' p" := 'Z_(pdiv p)
+ (at level 8, p at level 2, format "''F_' p") : type_scope.
+
+Section Groups.
+
+Variable p : nat.
+
+Definition Zp := if p > 1 then [set: 'Z_p] else 1%g.
+Definition units_Zp := [set: {unit 'Z_p}].
+
+Lemma Zp_cast : p > 1 -> (Zp_trunc p).+2 = p.
+Proof. by case: p => [|[]]. Qed.
+
+Lemma val_Zp_nat (p_gt1 : p > 1) n : (n%:R : 'Z_p) = (n %% p)%N :> nat.
+Proof. by rewrite Zp_nat /= Zp_cast. Qed.
+
+Lemma Zp_nat_mod (p_gt1 : p > 1)m : (m %% p)%:R = m%:R :> 'Z_p.
+Proof. by apply: ord_inj; rewrite !val_Zp_nat // modn_mod. Qed.
+
+Lemma char_Zp : p > 1 -> p%:R = 0 :> 'Z_p.
+Proof. by move=> p_gt1; rewrite -Zp_nat_mod ?modnn. Qed.
+
+Lemma unitZpE x : p > 1 -> ((x%:R : 'Z_p) \is a GRing.unit) = coprime p x.
+Proof.
+by move=> p_gt1; rewrite qualifE /= val_Zp_nat ?Zp_cast ?coprime_modr.
+Qed.
+
+Lemma Zp_group_set : group_set Zp.
+Proof. rewrite /Zp; case: (p > 1); exact: groupP. Qed.
+Canonical Zp_group := Group Zp_group_set.
+
+Lemma card_Zp : p > 0 -> #|Zp| = p.
+Proof.
+rewrite /Zp; case: p => [|[|p']] //= _; first by rewrite cards1.
+by rewrite cardsT card_ord.
+Qed.
+
+Lemma mem_Zp x : p > 1 -> x \in Zp. Proof. by rewrite /Zp => ->. Qed.
+
+Canonical units_Zp_group := [group of units_Zp].
+
+Lemma card_units_Zp : p > 0 -> #|units_Zp| = totient p.
+Proof.
+move=> p_gt0; transitivity (totient p.-2.+2); last by case: p p_gt0 => [|[|p']].
+rewrite cardsT card_sub -sum1_card big_mkcond /=.
+by rewrite totient_count_coprime big_mkord.
+Qed.
+
+Lemma units_Zp_abelian : abelian units_Zp.
+Proof. apply/centsP=> u _ v _; exact: unit_Zp_mulgC. Qed.
+
+End Groups.
+
+(* Field structure for primes. *)
+
+Section PrimeField.
+
+Open Scope ring_scope.
+
+Variable p : nat.
+
+Section F_prime.
+
+Hypothesis p_pr : prime p.
+
+Lemma Fp_Zcast : (Zp_trunc (pdiv p)).+2 = (Zp_trunc p).+2.
+Proof. by rewrite /pdiv primes_prime. Qed.
+
+Lemma Fp_cast : (Zp_trunc (pdiv p)).+2 = p.
+Proof. by rewrite Fp_Zcast ?Zp_cast ?prime_gt1. Qed.
+
+Lemma card_Fp : #|'F_p| = p.
+Proof. by rewrite card_ord Fp_cast. Qed.
+
+Lemma val_Fp_nat n : (n%:R : 'F_p) = (n %% p)%N :> nat.
+Proof. by rewrite Zp_nat /= Fp_cast. Qed.
+
+Lemma Fp_nat_mod m : (m %% p)%:R = m%:R :> 'F_p.
+Proof. by apply: ord_inj; rewrite !val_Fp_nat // modn_mod. Qed.
+
+Lemma char_Fp : p \in [char 'F_p].
+Proof. by rewrite !inE -Fp_nat_mod p_pr ?modnn. Qed.
+
+Lemma char_Fp_0 : p%:R = 0 :> 'F_p.
+Proof. exact: GRing.charf0 char_Fp. Qed.
+
+Lemma unitFpE x : ((x%:R : 'F_p) \is a GRing.unit) = coprime p x.
+Proof. by rewrite pdiv_id // unitZpE // prime_gt1. Qed.
+
+End F_prime.
+
+Lemma Fp_fieldMixin : GRing.Field.mixin_of [the unitRingType of 'F_p].
+Proof.
+move=> x nzx; rewrite qualifE /= prime_coprime ?gtnNdvd ?lt0n //.
+case: (ltnP 1 p) => [lt1p | ]; last by case: p => [|[|p']].
+by rewrite Zp_cast ?prime_gt1 ?pdiv_prime.
+Qed.
+
+Definition Fp_idomainMixin := FieldIdomainMixin Fp_fieldMixin.
+
+Canonical Fp_idomainType := Eval hnf in IdomainType 'F_p Fp_idomainMixin.
+Canonical Fp_finIdomainType := Eval hnf in [finIdomainType of 'F_p].
+Canonical Fp_fieldType := Eval hnf in FieldType 'F_p Fp_fieldMixin.
+Canonical Fp_finFieldType := Eval hnf in [finFieldType of 'F_p].
+Canonical Fp_decFieldType :=
+ Eval hnf in [decFieldType of 'F_p for Fp_finFieldType].
+
+End PrimeField.