aboutsummaryrefslogtreecommitdiff
path: root/mathcomp/algebra
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/algebra
Initial commit
Diffstat (limited to 'mathcomp/algebra')
-rw-r--r--mathcomp/algebra/all.v16
-rw-r--r--mathcomp/algebra/finalg.v1297
-rw-r--r--mathcomp/algebra/fraction.v384
-rw-r--r--mathcomp/algebra/intdiv.v1076
-rw-r--r--mathcomp/algebra/interval.v403
-rw-r--r--mathcomp/algebra/matrix.v2872
-rw-r--r--mathcomp/algebra/mxalgebra.v2764
-rw-r--r--mathcomp/algebra/mxpoly.v1109
-rw-r--r--mathcomp/algebra/poly.v2591
-rw-r--r--mathcomp/algebra/polyXY.v405
-rw-r--r--mathcomp/algebra/polydiv.v3418
-rw-r--r--mathcomp/algebra/rat.v808
-rw-r--r--mathcomp/algebra/ring_quotient.v650
-rw-r--r--mathcomp/algebra/ssralg.v6230
-rw-r--r--mathcomp/algebra/ssrint.v1782
-rw-r--r--mathcomp/algebra/ssrnum.v4219
-rw-r--r--mathcomp/algebra/vector.v2040
17 files changed, 32064 insertions, 0 deletions
diff --git a/mathcomp/algebra/all.v b/mathcomp/algebra/all.v
new file mode 100644
index 0000000..f4406a9
--- /dev/null
+++ b/mathcomp/algebra/all.v
@@ -0,0 +1,16 @@
+Require Export finalg.
+Require Export fraction.
+Require Export intdiv.
+Require Export interval.
+Require Export matrix.
+Require Export mxalgebra.
+Require Export mxpoly.
+Require Export polydiv.
+Require Export poly.
+Require Export polyXY.
+Require Export rat.
+Require Export ring_quotient.
+Require Export ssralg.
+Require Export ssrint.
+Require Export ssrnum.
+Require Export vector.
diff --git a/mathcomp/algebra/finalg.v b/mathcomp/algebra/finalg.v
new file mode 100644
index 0000000..b903577
--- /dev/null
+++ b/mathcomp/algebra/finalg.v
@@ -0,0 +1,1297 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype.
+Require Import ssralg finset fingroup morphism perm action.
+
+(*****************************************************************************)
+(* This file clones the entire ssralg hierachy for finite types; this allows *)
+(* type inference to function properly on expressions that mix combinatorial *)
+(* and algebraic operators (e.g., [set x + y | x in A, y in A]). *)
+(* finZmodType, finRingType, finComRingType, finUnitRingType, *)
+(* finComUnitRingType, finIdomType, finFieldType finLmodType, *)
+(* finLalgType finAlgType finUnitAlgType *)
+(* == the finite counterparts of zmodType, etc. *)
+(* Note that a finFieldType is canonically decidable. All these structures *)
+(* can be derived using [xxxType of T] forms, e.g., if R has both canonical *)
+(* finType and ringType structures, then *)
+(* Canonical R_finRingType := Eval hnf in [finRingType of R]. *)
+(* declares the derived finRingType structure for R. As the implementation *)
+(* of the derivation is somewhat involved, the Eval hnf normalization is *)
+(* strongly recommended. *)
+(* This file also provides direct tie-ins with finite group theory: *)
+(* [baseFinGroupType of R for +%R] == the (canonical) additive group *)
+(* [finGroupType of R for +%R] structures for R *)
+(* {unit R} == the type of units of R, which has a *)
+(* canonical group structure. *)
+(* FinRing.unit R Ux == the element of {unit R} corresponding *)
+(* to x, where Ux : x \in GRing.unit. *)
+(* 'U%act == the action by right multiplication of *)
+(* {unit R} on R, via FinRing.unit_act. *)
+(* (This is also a group action.) *)
+(*****************************************************************************)
+
+Local Open Scope ring_scope.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Module FinRing.
+
+Local Notation mixin_of T b := (Finite.mixin_of (EqType T b)).
+
+Section Generic.
+
+(* Implicits *)
+Variables (type base_type : Type) (class_of base_of : Type -> Type).
+Variable to_choice : forall T, base_of T -> Choice.class_of T.
+Variable base_sort : base_type -> Type.
+
+(* Explicits *)
+Variable Pack : forall T, class_of T -> Type -> type.
+Variable Class : forall T b, mixin_of T (to_choice b) -> class_of T.
+Variable base_class : forall bT, base_of (base_sort bT).
+
+Definition gen_pack T :=
+ fun bT b & phant_id (base_class bT) b =>
+ fun fT m & phant_id (Finite.class fT) (Finite.Class m) =>
+ Pack (@Class T b m) T.
+
+End Generic.
+
+Implicit Arguments
+ gen_pack [type base_type class_of base_of to_choice base_sort].
+Local Notation fin_ c := (@Finite.Class _ c c).
+Local Notation do_pack pack T := (pack T _ _ id _ _ id).
+Import GRing.Theory.
+
+Definition groupMixin V := FinGroup.Mixin (@addrA V) (@add0r V) (@addNr V).
+Local Notation base_group T vT fT :=
+ (@FinGroup.PackBase T (groupMixin vT) (Finite.class fT)).
+Local Notation fin_group B V := (@FinGroup.Pack B (@addNr V)).
+
+Module Zmodule.
+
+Section ClassDef.
+
+Record class_of M :=
+ Class { base : GRing.Zmodule.class_of M; mixin : mixin_of M base }.
+Local Coercion base : class_of >-> GRing.Zmodule.class_of.
+Local Coercion mixin : class_of >-> mixin_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Definition pack := gen_pack Pack Class GRing.Zmodule.class.
+Variable cT : type.
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+
+Definition join_finType := @Finite.Pack zmodType (fin_ xclass) xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.Zmodule.class_of.
+Coercion mixin : class_of >-> mixin_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Canonical join_finType.
+Notation finZmodType := type.
+Notation "[ 'finZmodType' 'of' T ]" := (do_pack pack T)
+ (at level 0, format "[ 'finZmodType' 'of' T ]") : form_scope.
+Coercion baseFinGroupType : type >-> FinGroup.base_type.
+Canonical baseFinGroupType.
+Coercion finGroupType : type >-> FinGroup.type.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+Notation "[ 'baseFinGroupType' 'of' R 'for' +%R ]" :=
+ (BaseFinGroupType R (groupMixin _))
+ (at level 0, format "[ 'baseFinGroupType' 'of' R 'for' +%R ]")
+ : form_scope.
+Notation "[ 'finGroupType' 'of' R 'for' +%R ]" :=
+ (@FinGroup.clone R _ (finGroupType _) id _ id)
+ (at level 0, format "[ 'finGroupType' 'of' R 'for' +%R ]") : form_scope.
+End Exports.
+
+End Zmodule.
+Import Zmodule.Exports.
+
+Section AdditiveGroup.
+
+Variable U : finZmodType.
+Implicit Types x y : U.
+
+Lemma zmod1gE : 1%g = 0 :> U. Proof. by []. Qed.
+Lemma zmodVgE x : x^-1%g = - x. Proof. by []. Qed.
+Lemma zmodMgE x y : (x * y)%g = x + y. Proof. by []. Qed.
+Lemma zmodXgE n x : (x ^+ n)%g = x *+ n. Proof. by []. Qed.
+Lemma zmod_mulgC x y : commute x y. Proof. exact: GRing.addrC. Qed.
+Lemma zmod_abelian (A : {set U}) : abelian A.
+Proof. by apply/centsP=> x _ y _; exact: zmod_mulgC. Qed.
+
+End AdditiveGroup.
+
+Module Ring.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : GRing.Ring.class_of R; mixin : mixin_of R base }.
+Definition base2 R (c : class_of R) := Zmodule.Class (mixin c).
+Local Coercion base : class_of >-> GRing.Ring.class_of.
+Local Coercion base2 : class_of >-> Zmodule.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Definition pack := gen_pack Pack Class GRing.Ring.class.
+Variable cT : type.
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) cT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass cT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition join_finType := @Finite.Pack ringType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack ringType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Import Exports.
+Coercion base : class_of >-> GRing.Ring.class_of.
+Coercion base2 : class_of >-> Zmodule.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Notation finRingType := type.
+Notation "[ 'finRingType' 'of' T ]" := (do_pack pack T)
+ (at level 0, format "[ 'finRingType' 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+Section Unit.
+
+Variable R : finRingType.
+
+Definition is_inv (x y : R) := (x * y == 1) && (y * x == 1).
+Definition unit := [qualify a x : R | [exists y, is_inv x y]].
+Definition inv x := odflt x (pick (is_inv x)).
+
+Lemma mulVr : {in unit, left_inverse 1 inv *%R}.
+Proof.
+rewrite /inv => x Ux; case: pickP => [y | no_y]; last by case/pred0P: Ux.
+by case/andP=> _; move/eqP.
+Qed.
+
+Lemma mulrV : {in unit, right_inverse 1 inv *%R}.
+Proof.
+rewrite /inv => x Ux; case: pickP => [y | no_y]; last by case/pred0P: Ux.
+by case/andP; move/eqP.
+Qed.
+
+Lemma intro_unit x y : y * x = 1 /\ x * y = 1 -> x \is a unit.
+Proof.
+by case=> yx1 xy1; apply/existsP; exists y; rewrite /is_inv xy1 yx1 !eqxx.
+Qed.
+
+Lemma invr_out : {in [predC unit], inv =1 id}.
+Proof.
+rewrite /inv => x nUx; case: pickP => // y invxy.
+by case/existsP: nUx; exists y.
+Qed.
+
+Definition UnitMixin := GRing.UnitRing.Mixin mulVr mulrV intro_unit invr_out.
+
+End Unit.
+
+End Ring.
+Import Ring.Exports.
+
+Module ComRing.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : GRing.ComRing.class_of R; mixin : mixin_of R base }.
+Definition base2 R (c : class_of R) := Ring.Class (mixin c).
+Local Coercion base : class_of >-> GRing.ComRing.class_of.
+Local Coercion base2 : class_of >-> Ring.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Definition pack := gen_pack Pack Class GRing.ComRing.class.
+Variable cT : type.
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition finRingType := @Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition join_finType := @Finite.Pack comRingType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack comRingType xclass xT.
+Definition join_finRingType := @Ring.Pack comRingType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.ComRing.class_of.
+Coercion base2 : class_of >-> Ring.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion finRingType : type >-> Ring.type.
+Canonical finRingType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Canonical join_finRingType.
+Notation finComRingType := FinRing.ComRing.type.
+Notation "[ 'finComRingType' 'of' T ]" := (do_pack pack T)
+ (at level 0, format "[ 'finComRingType' 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+End ComRing.
+Import ComRing.Exports.
+
+Module UnitRing.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : GRing.UnitRing.class_of R; mixin : mixin_of R base }.
+Definition base2 R (c : class_of R) := Ring.Class (mixin c).
+Local Coercion base : class_of >-> GRing.UnitRing.class_of.
+Local Coercion base2 : class_of >-> Ring.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Definition pack := gen_pack Pack Class GRing.UnitRing.class.
+Variable cT : type.
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition finRingType := @Ring.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+
+Definition join_finType := @Finite.Pack unitRingType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack unitRingType xclass xT.
+Definition join_finRingType := @Ring.Pack unitRingType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.UnitRing.class_of.
+Coercion base2 : class_of >-> Ring.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion finRingType : type >-> Ring.type.
+Canonical finRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Canonical join_finRingType.
+Notation finUnitRingType := FinRing.UnitRing.type.
+Notation "[ 'finUnitRingType' 'of' T ]" := (do_pack pack T)
+ (at level 0, format "[ 'finUnitRingType' 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+End UnitRing.
+Import UnitRing.Exports.
+
+Section UnitsGroup.
+
+Variable R : finUnitRingType.
+
+Inductive unit_of (phR : phant R) := Unit (x : R) of x \is a GRing.unit.
+Bind Scope group_scope with unit_of.
+
+Let phR := Phant R.
+Local Notation uT := (unit_of phR).
+Implicit Types u v : uT.
+Definition uval u := let: Unit x _ := u in x.
+
+Canonical unit_subType := [subType for uval].
+Definition unit_eqMixin := Eval hnf in [eqMixin of uT by <:].
+Canonical unit_eqType := Eval hnf in EqType uT unit_eqMixin.
+Definition unit_choiceMixin := [choiceMixin of uT by <:].
+Canonical unit_choiceType := Eval hnf in ChoiceType uT unit_choiceMixin.
+Definition unit_countMixin := [countMixin of uT by <:].
+Canonical unit_countType := Eval hnf in CountType uT unit_countMixin.
+Canonical unit_subCountType := Eval hnf in [subCountType of uT].
+Definition unit_finMixin := [finMixin of uT by <:].
+Canonical unit_finType := Eval hnf in FinType uT unit_finMixin.
+Canonical unit_subFinType := Eval hnf in [subFinType of uT].
+
+Definition unit1 := Unit phR (@GRing.unitr1 _).
+Lemma unit_inv_proof u : (val u)^-1 \is a GRing.unit.
+Proof. by rewrite GRing.unitrV ?(valP u). Qed.
+Definition unit_inv u := Unit phR (unit_inv_proof u).
+Lemma unit_mul_proof u v : val u * val v \is a GRing.unit.
+Proof. by rewrite (GRing.unitrMr _ (valP u)) ?(valP v). Qed.
+Definition unit_mul u v := Unit phR (unit_mul_proof u v).
+Lemma unit_muluA : associative unit_mul.
+Proof. move=> u v w; apply: val_inj; exact: GRing.mulrA. Qed.
+Lemma unit_mul1u : left_id unit1 unit_mul.
+Proof. move=> u; apply: val_inj; exact: GRing.mul1r. Qed.
+Lemma unit_mulVu : left_inverse unit1 unit_inv unit_mul.
+Proof. move=> u; apply: val_inj; exact: GRing.mulVr (valP u). Qed.
+
+Definition unit_GroupMixin := FinGroup.Mixin unit_muluA unit_mul1u unit_mulVu.
+Canonical unit_baseFinGroupType :=
+ Eval hnf in BaseFinGroupType uT unit_GroupMixin.
+Canonical unit_finGroupType := Eval hnf in FinGroupType unit_mulVu.
+
+Lemma val_unit1 : val (1%g : uT) = 1. Proof. by []. Qed.
+Lemma val_unitM x y : val (x * y : uT)%g = val x * val y. Proof. by []. Qed.
+Lemma val_unitV x : val (x^-1 : uT)%g = (val x)^-1. Proof. by []. Qed.
+Lemma val_unitX n x : val (x ^+ n : uT)%g = val x ^+ n.
+Proof. by case: n; last by elim=> //= n ->. Qed.
+
+Definition unit_act x u := x * val u.
+Lemma unit_actE x u : unit_act x u = x * val u. Proof. by []. Qed.
+
+Canonical unit_action :=
+ @TotalAction _ _ unit_act (@GRing.mulr1 _) (fun _ _ _ => GRing.mulrA _ _ _).
+Lemma unit_is_groupAction : @is_groupAction _ R setT setT unit_action.
+Proof.
+move=> u _ /=; rewrite inE; apply/andP; split.
+ by apply/subsetP=> x _; rewrite inE.
+by apply/morphicP=> x y _ _; rewrite !actpermE /= [_ u]GRing.mulrDl.
+Qed.
+Canonical unit_groupAction := GroupAction unit_is_groupAction.
+
+End UnitsGroup.
+
+Module Import UnitsGroupExports.
+Bind Scope group_scope with unit_of.
+Canonical unit_subType.
+Canonical unit_eqType.
+Canonical unit_choiceType.
+Canonical unit_countType.
+Canonical unit_subCountType.
+Canonical unit_finType.
+Canonical unit_subFinType.
+Canonical unit_baseFinGroupType.
+Canonical unit_finGroupType.
+Canonical unit_action.
+Canonical unit_groupAction.
+End UnitsGroupExports.
+
+Notation unit R Ux := (Unit (Phant R) Ux).
+
+Module ComUnitRing.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : GRing.ComUnitRing.class_of R; mixin : mixin_of R base }.
+Definition base2 R (c : class_of R) := ComRing.Class (mixin c).
+Definition base3 R (c : class_of R) := @UnitRing.Class R (base c) (mixin c).
+Local Coercion base : class_of >-> GRing.ComUnitRing.class_of.
+Local Coercion base2 : class_of >-> ComRing.class_of.
+Local Coercion base3 : class_of >-> UnitRing.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Definition pack := gen_pack Pack Class GRing.ComUnitRing.class.
+Variable cT : type.
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition finRingType := @Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition finComRingType := @ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition finUnitRingType := @UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+
+Definition join_finType := @Finite.Pack comUnitRingType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack comUnitRingType xclass xT.
+Definition join_finRingType := @Ring.Pack comUnitRingType xclass xT.
+Definition join_finComRingType := @ComRing.Pack comUnitRingType xclass xT.
+Definition join_finUnitRingType := @UnitRing.Pack comUnitRingType xclass xT.
+Definition ujoin_finComRingType := @ComRing.Pack unitRingType xclass xT.
+Definition cjoin_finUnitRingType := @UnitRing.Pack comRingType xclass xT.
+Definition fcjoin_finUnitRingType := @UnitRing.Pack finComRingType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.ComUnitRing.class_of.
+Coercion base2 : class_of >-> ComRing.class_of.
+Coercion base3 : class_of >-> UnitRing.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion finRingType : type >-> Ring.type.
+Canonical finRingType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion finComRingType : type >-> ComRing.type.
+Canonical finComRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion finUnitRingType : type >-> UnitRing.type.
+Canonical finUnitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Canonical join_finRingType.
+Canonical join_finComRingType.
+Canonical join_finUnitRingType.
+Canonical ujoin_finComRingType.
+Canonical cjoin_finUnitRingType.
+Canonical fcjoin_finUnitRingType.
+Notation finComUnitRingType := FinRing.ComUnitRing.type.
+Notation "[ 'finComUnitRingType' 'of' T ]" := (do_pack pack T)
+ (at level 0, format "[ 'finComUnitRingType' 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+End ComUnitRing.
+Import ComUnitRing.Exports.
+
+Module IntegralDomain.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : GRing.IntegralDomain.class_of R; mixin : mixin_of R base }.
+Definition base2 R (c : class_of R) := ComUnitRing.Class (mixin c).
+Local Coercion base : class_of >-> GRing.IntegralDomain.class_of.
+Local Coercion base2 : class_of >-> ComUnitRing.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Definition pack := gen_pack Pack Class GRing.IntegralDomain.class.
+Variable cT : type.
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition finRingType := @Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition finComRingType := @ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition finUnitRingType := @UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+Definition finComUnitRingType := @ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT.
+
+Definition join_finType := @Finite.Pack idomainType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack idomainType xclass xT.
+Definition join_finRingType := @Ring.Pack idomainType xclass xT.
+Definition join_finUnitRingType := @UnitRing.Pack idomainType xclass xT.
+Definition join_finComRingType := @ComRing.Pack idomainType xclass xT.
+Definition join_finComUnitRingType := @ComUnitRing.Pack idomainType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.IntegralDomain.class_of.
+Coercion base2 : class_of >-> ComUnitRing.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion finRingType : type >-> Ring.type.
+Canonical finRingType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion finComRingType : type >-> ComRing.type.
+Canonical finComRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion finUnitRingType : type >-> UnitRing.type.
+Canonical finUnitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion finComUnitRingType : type >-> ComUnitRing.type.
+Canonical finComUnitRingType.
+Coercion idomainType : type >-> GRing.IntegralDomain.type.
+Canonical idomainType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Canonical join_finRingType.
+Canonical join_finComRingType.
+Canonical join_finUnitRingType.
+Canonical join_finComUnitRingType.
+Notation finIdomainType := FinRing.IntegralDomain.type.
+Notation "[ 'finIdomainType' 'of' T ]" := (do_pack pack T)
+ (at level 0, format "[ 'finIdomainType' 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+End IntegralDomain.
+Import IntegralDomain.Exports.
+
+Module Field.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : GRing.Field.class_of R; mixin : mixin_of R base }.
+Definition base2 R (c : class_of R) := IntegralDomain.Class (mixin c).
+Local Coercion base : class_of >-> GRing.Field.class_of.
+Local Coercion base2 : class_of >-> IntegralDomain.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Definition pack := gen_pack Pack Class GRing.Field.class.
+Variable cT : type.
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition finRingType := @Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition finComRingType := @ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition finUnitRingType := @UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+Definition finComUnitRingType := @ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT.
+Definition finIdomainType := @IntegralDomain.Pack cT xclass xT.
+Definition fieldType := @GRing.Field.Pack cT xclass xT.
+
+Definition join_finType := @Finite.Pack fieldType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack fieldType xclass xT.
+Definition join_finRingType := @Ring.Pack fieldType xclass xT.
+Definition join_finUnitRingType := @UnitRing.Pack fieldType xclass xT.
+Definition join_finComRingType := @ComRing.Pack fieldType xclass xT.
+Definition join_finComUnitRingType := @ComUnitRing.Pack fieldType xclass xT.
+Definition join_finIdomainType := @IntegralDomain.Pack fieldType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.Field.class_of.
+Coercion base2 : class_of >-> IntegralDomain.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion finRingType : type >-> Ring.type.
+Canonical finRingType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion finComRingType : type >-> ComRing.type.
+Canonical finComRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion finUnitRingType : type >-> UnitRing.type.
+Canonical finUnitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion finComUnitRingType : type >-> ComUnitRing.type.
+Canonical finComUnitRingType.
+Coercion idomainType : type >-> GRing.IntegralDomain.type.
+Canonical idomainType.
+Coercion finIdomainType : type >-> IntegralDomain.type.
+Canonical finIdomainType.
+Coercion fieldType : type >-> GRing.Field.type.
+Canonical fieldType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Canonical join_finRingType.
+Canonical join_finComRingType.
+Canonical join_finUnitRingType.
+Canonical join_finComUnitRingType.
+Canonical join_finIdomainType.
+Notation finFieldType := FinRing.Field.type.
+Notation "[ 'finFieldType' 'of' T ]" := (do_pack pack T)
+ (at level 0, format "[ 'finFieldType' 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+End Field.
+Import Field.Exports.
+
+Section DecideField.
+
+Variable F : Field.type.
+
+Fixpoint sat e f :=
+ match f with
+ | GRing.Bool b => b
+ | t1 == t2 => (GRing.eval e t1 == GRing.eval e t2)%bool
+ | GRing.Unit t => GRing.eval e t \is a GRing.unit
+ | f1 /\ f2 => sat e f1 && sat e f2
+ | f1 \/ f2 => sat e f1 || sat e f2
+ | f1 ==> f2 => (sat e f1 ==> sat e f2)%bool
+ | ~ f1 => ~~ sat e f1
+ | ('exists 'X_k, f1) => [exists x : F, sat (set_nth 0%R e k x) f1]
+ | ('forall 'X_k, f1) => [forall x : F, sat (set_nth 0%R e k x) f1]
+ end%T.
+
+Lemma decidable : GRing.DecidableField.axiom sat.
+Proof.
+move=> e f; elim: f e;
+ try by move=> f1 IH1 f2 IH2 e /=; case IH1; case IH2; constructor; tauto.
+- by move=> b e; exact: idP.
+- by move=> t1 t2 e; exact: eqP.
+- by move=> t e; exact: idP.
+- by move=> f IH e /=; case: IH; constructor.
+- by move=> i f IH e; apply: (iffP existsP) => [] [x fx]; exists x; exact/IH.
+by move=> i f IH e; apply: (iffP forallP) => f_ x; exact/IH.
+Qed.
+
+Definition DecidableFieldMixin := DecFieldMixin decidable.
+
+End DecideField.
+
+Module DecField.
+
+Section Joins.
+
+Variable cT : Field.type.
+Let xT := let: Field.Pack T _ _ := cT in T.
+Let xclass : Field.class_of xT := Field.class cT.
+
+Definition type := Eval hnf in DecFieldType cT (DecidableFieldMixin cT).
+Definition finType := @Finite.Pack type (fin_ xclass) xT.
+Definition finZmodType := @Zmodule.Pack type xclass xT.
+Definition finRingType := @Ring.Pack type xclass xT.
+Definition finUnitRingType := @UnitRing.Pack type xclass xT.
+Definition finComRingType := @ComRing.Pack type xclass xT.
+Definition finComUnitRingType := @ComUnitRing.Pack type xclass xT.
+Definition finIdomainType := @IntegralDomain.Pack type xclass xT.
+Definition baseFinGroupType := base_group type finZmodType finZmodType.
+Definition finGroupType := fin_group baseFinGroupType cT.
+
+End Joins.
+
+Module Exports.
+Coercion type : Field.type >-> GRing.DecidableField.type.
+Canonical type.
+Canonical finType.
+Canonical finZmodType.
+Canonical finRingType.
+Canonical finUnitRingType.
+Canonical finComRingType.
+Canonical finComUnitRingType.
+Canonical finIdomainType.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+End Exports.
+
+End DecField.
+
+Module Lmodule.
+
+Section ClassDef.
+
+Variable R : ringType.
+
+Record class_of M :=
+ Class { base : GRing.Lmodule.class_of R M ; mixin : mixin_of M base }.
+Definition base2 R (c : class_of R) := Zmodule.Class (mixin c).
+Local Coercion base : class_of >-> GRing.Lmodule.class_of.
+Local Coercion base2 : class_of >-> Zmodule.class_of.
+
+Structure type (phR : phant R) := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (phR : phant R) (cT : type phR).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition pack := gen_pack (Pack phR) Class (@GRing.Lmodule.class R phR).
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition lmodType := @GRing.Lmodule.Pack R phR cT xclass xT.
+Definition join_finType := @Finite.Pack lmodType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack lmodType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Import Exports.
+Coercion base : class_of >-> GRing.Lmodule.class_of.
+Coercion base2 : class_of >-> Zmodule.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion lmodType : type >-> GRing.Lmodule.type.
+Canonical lmodType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Notation finLmodType R := (FinRing.Lmodule.type (Phant R)).
+Notation "[ 'finLmodType' R 'of' T ]" := (do_pack (@pack _ (Phant R)) T)
+ (at level 0, format "[ 'finLmodType' R 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+End Lmodule.
+Import Lmodule.Exports.
+
+Module Lalgebra.
+
+Section ClassDef.
+
+Variable R : ringType.
+
+Record class_of M :=
+ Class { base : GRing.Lalgebra.class_of R M; mixin : mixin_of M base }.
+Definition base2 M (c : class_of M) := Ring.Class (mixin c).
+Definition base3 M (c : class_of M) := @Lmodule.Class _ _ (base c) (mixin c).
+Local Coercion base : class_of >-> GRing.Lalgebra.class_of.
+Local Coercion base2 : class_of >-> Ring.class_of.
+Local Coercion base3 : class_of >-> Lmodule.class_of.
+
+Structure type (phR : phant R) := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (phR : phant R) (cT : type phR).
+Definition pack := gen_pack (Pack phR) Class (@GRing.Lalgebra.class R phR).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition finRingType := @Ring.Pack cT xclass xT.
+Definition lmodType := @GRing.Lmodule.Pack R phR cT xclass xT.
+Definition finLmodType := @Lmodule.Pack R phR cT xclass xT.
+Definition lalgType := @GRing.Lalgebra.Pack R phR cT xclass xT.
+
+Definition join_finType := @Finite.Pack lalgType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack lalgType xclass xT.
+Definition join_finLmodType := @Lmodule.Pack R phR lalgType xclass xT.
+Definition join_finRingType := @Ring.Pack lalgType xclass xT.
+Definition rjoin_finLmodType := @Lmodule.Pack R phR ringType xclass xT.
+Definition ljoin_finRingType := @Ring.Pack lmodType xclass xT.
+Definition fljoin_finRingType := @Ring.Pack finLmodType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.Lalgebra.class_of.
+Coercion base2 : class_of >-> Ring.class_of.
+Coercion base3 : class_of >-> Lmodule.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion finRingType : type >-> Ring.type.
+Canonical finRingType.
+Coercion lmodType : type >-> GRing.Lmodule.type.
+Canonical lmodType.
+Coercion finLmodType : type >-> Lmodule.type.
+Canonical finLmodType.
+Coercion lalgType : type >-> GRing.Lalgebra.type.
+Canonical lalgType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Canonical join_finLmodType.
+Canonical join_finRingType.
+Canonical rjoin_finLmodType.
+Canonical ljoin_finRingType.
+Canonical fljoin_finRingType.
+Notation finLalgType R := (FinRing.Lalgebra.type (Phant R)).
+Notation "[ 'finLalgType' R 'of' T ]" := (do_pack (@pack _ (Phant R)) T)
+ (at level 0, format "[ 'finLalgType' R 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+End Lalgebra.
+Import Lalgebra.Exports.
+
+Module Algebra.
+
+Section ClassDef.
+
+Variable R : ringType.
+
+Record class_of M :=
+ Class { base : GRing.Algebra.class_of R M; mixin : mixin_of M base }.
+Definition base2 M (c : class_of M) := Lalgebra.Class (mixin c).
+Local Coercion base : class_of >-> GRing.Algebra.class_of.
+Local Coercion base2 : class_of >->Lalgebra.class_of.
+
+Structure type (phR : phant R) := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (phR : phant R) (cT : type phR).
+Definition pack := gen_pack (Pack phR) Class (@GRing.Algebra.class R phR).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition finRingType := @Ring.Pack cT xclass xT.
+Definition lmodType := @GRing.Lmodule.Pack R phR cT xclass xT.
+Definition finLmodType := @Lmodule.Pack R phR cT xclass xT.
+Definition lalgType := @GRing.Lalgebra.Pack R phR cT xclass xT.
+Definition finLalgType := @Lalgebra.Pack R phR cT xclass xT.
+Definition algType := @GRing.Algebra.Pack R phR cT xclass xT.
+
+Definition join_finType := @Finite.Pack algType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack algType xclass xT.
+Definition join_finRingType := @Ring.Pack algType xclass xT.
+Definition join_finLmodType := @Lmodule.Pack R phR algType xclass xT.
+Definition join_finLalgType := @Lalgebra.Pack R phR algType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.Algebra.class_of.
+Coercion base2 : class_of >-> Lalgebra.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion finRingType : type >-> Ring.type.
+Canonical finRingType.
+Coercion lmodType : type >-> GRing.Lmodule.type.
+Canonical lmodType.
+Coercion finLmodType : type >-> Lmodule.type.
+Canonical finLmodType.
+Coercion lalgType : type >-> GRing.Lalgebra.type.
+Canonical lalgType.
+Coercion finLalgType : type >-> Lalgebra.type.
+Canonical finLalgType.
+Coercion algType : type >-> GRing.Algebra.type.
+Canonical algType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Canonical join_finLmodType.
+Canonical join_finRingType.
+Canonical join_finLalgType.
+Notation finAlgType R := (type (Phant R)).
+Notation "[ 'finAlgType' R 'of' T ]" := (do_pack (@pack _ (Phant R)) T)
+ (at level 0, format "[ 'finAlgType' R 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+End Algebra.
+Import Algebra.Exports.
+
+Module UnitAlgebra.
+
+Section ClassDef.
+
+Variable R : unitRingType.
+
+Record class_of M :=
+ Class { base : GRing.UnitAlgebra.class_of R M ; mixin : mixin_of M base }.
+Definition base2 M (c : class_of M) := Algebra.Class (mixin c).
+Definition base3 M (c : class_of M) := @UnitRing.Class _ (base c) (mixin c).
+
+Local Coercion base : class_of >-> GRing.UnitAlgebra.class_of.
+Local Coercion base2 : class_of >-> Algebra.class_of.
+Local Coercion base3 : class_of >-> UnitRing.class_of.
+
+Structure type (phR : phant R) := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (phR : phant R) (cT : type phR).
+Definition pack := gen_pack (Pack phR) Class (@GRing.UnitAlgebra.class R phR).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition countType := @Countable.Pack cT (fin_ xclass) xT.
+Definition finType := @Finite.Pack cT (fin_ xclass) xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition finZmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition finRingType := @Ring.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition finUnitRingType := @UnitRing.Pack cT xclass xT.
+Definition lmodType := @GRing.Lmodule.Pack R phR cT xclass xT.
+Definition finLmodType := @Lmodule.Pack R phR cT xclass xT.
+Definition lalgType := @GRing.Lalgebra.Pack R phR cT xclass xT.
+Definition finLalgType := @Lalgebra.Pack R phR cT xclass xT.
+Definition algType := @GRing.Algebra.Pack R phR cT xclass xT.
+Definition finAlgType := @Algebra.Pack R phR cT xclass xT.
+Definition unitAlgType := @GRing.UnitAlgebra.Pack R phR cT xclass xT.
+
+Definition join_finType := @Finite.Pack unitAlgType (fin_ xclass) xT.
+Definition join_finZmodType := @Zmodule.Pack unitAlgType xclass xT.
+Definition join_finRingType := @Ring.Pack unitAlgType xclass xT.
+Definition join_finUnitRingType := @UnitRing.Pack unitAlgType xclass xT.
+Definition join_finLmodType := @Lmodule.Pack R phR unitAlgType xclass xT.
+Definition join_finLalgType := @Lalgebra.Pack R phR unitAlgType xclass xT.
+Definition join_finAlgType := @Algebra.Pack R phR unitAlgType xclass xT.
+Definition ljoin_finUnitRingType := @UnitRing.Pack lmodType xclass xT.
+Definition fljoin_finUnitRingType := @UnitRing.Pack finLmodType xclass xT.
+Definition njoin_finUnitRingType := @UnitRing.Pack lalgType xclass xT.
+Definition fnjoin_finUnitRingType := @UnitRing.Pack finLalgType xclass xT.
+Definition ajoin_finUnitRingType := @UnitRing.Pack algType xclass xT.
+Definition fajoin_finUnitRingType := @UnitRing.Pack finAlgType xclass xT.
+Definition ujoin_finLmodType := @Lmodule.Pack R phR unitRingType xclass xT.
+Definition ujoin_finLalgType := @Lalgebra.Pack R phR unitRingType xclass xT.
+Definition ujoin_finAlgType := @Algebra.Pack R phR unitRingType xclass xT.
+
+Definition baseFinGroupType := base_group cT zmodType finType.
+Definition finGroupType := fin_group baseFinGroupType zmodType.
+Definition join_baseFinGroupType := base_group zmodType zmodType finType.
+Definition join_finGroupType := fin_group join_baseFinGroupType zmodType.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.UnitAlgebra.class_of.
+Coercion base2 : class_of >-> Algebra.class_of.
+Coercion base3 : class_of >-> UnitRing.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion countType : type >-> Countable.type.
+Canonical countType.
+Coercion finType : type >-> Finite.type.
+Canonical finType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion finZmodType : type >-> Zmodule.type.
+Canonical finZmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion finRingType : type >-> Ring.type.
+Canonical finRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion finUnitRingType : type >-> UnitRing.type.
+Canonical finUnitRingType.
+Coercion lmodType : type >-> GRing.Lmodule.type.
+Canonical lmodType.
+Coercion finLmodType : type >-> Lmodule.type.
+Canonical finLmodType.
+Coercion lalgType : type >-> GRing.Lalgebra.type.
+Canonical lalgType.
+Coercion finLalgType : type >-> Lalgebra.type.
+Canonical finLalgType.
+Coercion algType : type >-> GRing.Algebra.type.
+Canonical algType.
+Coercion finAlgType : type >-> Algebra.type.
+Canonical finAlgType.
+Coercion unitAlgType : type >-> GRing.UnitAlgebra.type.
+Canonical unitAlgType.
+Canonical join_finType.
+Canonical join_finZmodType.
+Canonical join_finLmodType.
+Canonical join_finRingType.
+Canonical join_finLalgType.
+Canonical join_finAlgType.
+Canonical ljoin_finUnitRingType.
+Canonical fljoin_finUnitRingType.
+Canonical njoin_finUnitRingType.
+Canonical fnjoin_finUnitRingType.
+Canonical ajoin_finUnitRingType.
+Canonical fajoin_finUnitRingType.
+Canonical ujoin_finLmodType.
+Canonical ujoin_finLalgType.
+Canonical ujoin_finAlgType.
+Notation finUnitAlgType R := (type (Phant R)).
+Notation "[ 'finUnitAlgType' R 'of' T ]" := (do_pack (@pack _ (Phant R)) T)
+ (at level 0, format "[ 'finUnitAlgType' R 'of' T ]") : form_scope.
+Canonical baseFinGroupType.
+Canonical finGroupType.
+Canonical join_baseFinGroupType.
+Canonical join_finGroupType.
+End Exports.
+
+End UnitAlgebra.
+Import UnitAlgebra.Exports.
+
+Module Theory.
+
+Definition zmod1gE := zmod1gE.
+Definition zmodVgE := zmodVgE.
+Definition zmodMgE := zmodMgE.
+Definition zmodXgE := zmodXgE.
+Definition zmod_mulgC := zmod_mulgC.
+Definition zmod_abelian := zmod_abelian.
+Definition val_unit1 := val_unit1.
+Definition val_unitM := val_unitM.
+Definition val_unitX := val_unitX.
+Definition val_unitV := val_unitV.
+Definition unit_actE := unit_actE.
+
+End Theory.
+
+End FinRing.
+
+Import FinRing.
+Export Zmodule.Exports Ring.Exports ComRing.Exports.
+Export UnitRing.Exports UnitsGroupExports ComUnitRing.Exports.
+Export IntegralDomain.Exports Field.Exports DecField.Exports.
+Export Lmodule.Exports Lalgebra.Exports Algebra.Exports UnitAlgebra.Exports.
+
+Notation "{ 'unit' R }" := (unit_of (Phant R))
+ (at level 0, format "{ 'unit' R }") : type_scope.
+Prenex Implicits FinRing.uval.
+Notation "''U'" := (unit_action _) (at level 8) : action_scope.
+Notation "''U'" := (unit_groupAction _) (at level 8) : groupAction_scope.
+
diff --git a/mathcomp/algebra/fraction.v b/mathcomp/algebra/fraction.v
new file mode 100644
index 0000000..732cc19
--- /dev/null
+++ b/mathcomp/algebra/fraction.v
@@ -0,0 +1,384 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat div seq choice tuple.
+Require Import bigop ssralg poly polydiv generic_quotient.
+
+(* This file builds the field of fraction of any integral domain. *)
+(* The main result of this file is the existence of the field *)
+(* and of the tofrac function which is a injective ring morphism from R *)
+(* to its fraction field {fraction R} *)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GRing.Theory.
+Open Local Scope ring_scope.
+Open Local Scope quotient_scope.
+
+Reserved Notation "{ 'ratio' T }" (at level 0, format "{ 'ratio' T }").
+Reserved Notation "{ 'fraction' T }" (at level 0, format "{ 'fraction' T }").
+Reserved Notation "x %:F" (at level 2, format "x %:F").
+
+Section FracDomain.
+Variable R : ringType.
+
+(* ratios are pairs of R, such that the second member is nonzero *)
+Inductive ratio := mkRatio { frac :> R * R; _ : frac.2 != 0 }.
+Definition ratio_of of phant R := ratio.
+Local Notation "{ 'ratio' T }" := (ratio_of (Phant T)).
+
+Canonical ratio_subType := Eval hnf in [subType for frac].
+Canonical ratio_of_subType := Eval hnf in [subType of {ratio R}].
+Definition ratio_EqMixin := [eqMixin of ratio by <:].
+Canonical ratio_eqType := EqType ratio ratio_EqMixin.
+Canonical ratio_of_eqType := Eval hnf in [eqType of {ratio R}].
+Definition ratio_ChoiceMixin := [choiceMixin of ratio by <:].
+Canonical ratio_choiceType := ChoiceType ratio ratio_ChoiceMixin.
+Canonical ratio_of_choiceType := Eval hnf in [choiceType of {ratio R}].
+
+Lemma denom_ratioP : forall f : ratio, f.2 != 0. Proof. by case. Qed.
+
+Definition ratio0 := (@mkRatio (0, 1) (oner_neq0 _)).
+Definition Ratio x y : {ratio R} := insubd ratio0 (x, y).
+
+Lemma numer_Ratio x y : y != 0 -> (Ratio x y).1 = x.
+Proof. by move=> ny0; rewrite /Ratio /insubd insubT. Qed.
+
+Lemma denom_Ratio x y : y != 0 -> (Ratio x y).2 = y.
+Proof. by move=> ny0; rewrite /Ratio /insubd insubT. Qed.
+
+Definition numden_Ratio := (numer_Ratio, denom_Ratio).
+
+CoInductive Ratio_spec (n d : R) : {ratio R} -> R -> R -> Type :=
+ | RatioNull of d = 0 : Ratio_spec n d ratio0 n 0
+ | RatioNonNull (d_neq0 : d != 0) :
+ Ratio_spec n d (@mkRatio (n, d) d_neq0) n d.
+
+Lemma RatioP n d : Ratio_spec n d (Ratio n d) n d.
+Proof.
+rewrite /Ratio /insubd; case: insubP=> /= [x /= d_neq0 hx|].
+ have ->: x = @mkRatio (n, d) d_neq0 by apply: val_inj.
+ by constructor.
+by rewrite negbK=> /eqP hx; rewrite {2}hx; constructor.
+Qed.
+
+Lemma Ratio0 x : Ratio x 0 = ratio0.
+Proof. by rewrite /Ratio /insubd; case: insubP; rewrite //= eqxx. Qed.
+
+End FracDomain.
+
+Notation "{ 'ratio' T }" := (ratio_of (Phant T)).
+Identity Coercion type_fracdomain_of : ratio_of >-> ratio.
+
+Notation "'\n_' x" := (frac x).1
+ (at level 8, x at level 2, format "'\n_' x").
+Notation "'\d_' x" := (frac x).2
+ (at level 8, x at level 2, format "'\d_' x").
+
+Module FracField.
+Section FracField.
+
+Variable R : idomainType.
+
+Local Notation frac := (R * R).
+Local Notation dom := (ratio R).
+Local Notation domP := denom_ratioP.
+
+Implicit Types x y z : dom.
+
+(* We define a relation in ratios *)
+Local Notation equivf_notation x y := (\n_x * \d_y == \d_x * \n_y).
+Definition equivf x y := equivf_notation x y.
+
+Lemma equivfE x y : equivf x y = equivf_notation x y.
+Proof. by []. Qed.
+
+Lemma equivf_refl : reflexive equivf.
+Proof. by move=> x; rewrite /equivf mulrC. Qed.
+
+Lemma equivf_sym : symmetric equivf.
+Proof. by move=> x y; rewrite /equivf eq_sym; congr (_==_); rewrite mulrC. Qed.
+
+Lemma equivf_trans : transitive equivf.
+Proof.
+move=> [x Px] [y Py] [z Pz]; rewrite /equivf /= mulrC => /eqP xy /eqP yz.
+by rewrite -(inj_eq (mulfI Px)) mulrA xy -mulrA yz mulrCA.
+Qed.
+
+(* we show that equivf is an equivalence *)
+Canonical equivf_equiv := EquivRel equivf equivf_refl equivf_sym equivf_trans.
+
+Definition type := {eq_quot equivf}.
+Definition type_of of phant R := type.
+Notation "{ 'fraction' T }" := (type_of (Phant T)).
+
+(* we recover some structure for the quotient *)
+Canonical frac_quotType := [quotType of type].
+Canonical frac_eqType := [eqType of type].
+Canonical frac_choiceType := [choiceType of type].
+Canonical frac_eqQuotType := [eqQuotType equivf of type].
+
+Canonical frac_of_quotType := [quotType of {fraction R}].
+Canonical frac_of_eqType := [eqType of {fraction R}].
+Canonical frac_of_choiceType := [choiceType of {fraction R}].
+Canonical frac_of_eqQuotType := [eqQuotType equivf of {fraction R}].
+
+(* we explain what was the equivalence on the quotient *)
+Lemma equivf_def (x y : ratio R) : x == y %[mod type]
+ = (\n_x * \d_y == \d_x * \n_y).
+Proof. by rewrite eqmodE. Qed.
+
+Lemma equivf_r x : \n_x * \d_(repr (\pi_type x)) = \d_x * \n_(repr (\pi_type x)).
+Proof. by apply/eqP; rewrite -equivf_def reprK. Qed.
+
+Lemma equivf_l x : \n_(repr (\pi_type x)) * \d_x = \d_(repr (\pi_type x)) * \n_x.
+Proof. by apply/eqP; rewrite -equivf_def reprK. Qed.
+
+Lemma numer0 x : (\n_x == 0) = (x == (ratio0 R) %[mod_eq equivf]).
+Proof. by rewrite eqmodE /= !equivfE // mulr1 mulr0. Qed.
+
+Lemma Ratio_numden : forall x, Ratio \n_x \d_x = x.
+Proof.
+case=> [[n d] /= nd]; rewrite /Ratio /insubd; apply: val_inj=> /=.
+by case: insubP=> //=; rewrite nd.
+Qed.
+
+Definition tofrac := lift_embed {fraction R} (fun x : R => Ratio x 1).
+Canonical tofrac_pi_morph := PiEmbed tofrac.
+
+Notation "x %:F" := (@tofrac x).
+
+Implicit Types a b c : type.
+
+Definition addf x y : dom := Ratio (\n_x * \d_y + \n_y * \d_x) (\d_x * \d_y).
+Definition add := lift_op2 {fraction R} addf.
+
+Lemma pi_add : {morph \pi : x y / addf x y >-> add x y}.
+Proof.
+move=> x y; unlock add; apply/eqmodP; rewrite /= equivfE.
+rewrite /addf /= !numden_Ratio ?mulf_neq0 ?domP //.
+rewrite mulrDr mulrDl eq_sym; apply/eqP.
+rewrite !mulrA ![_ * \n__]mulrC !mulrA equivf_l.
+congr (_ + _); first by rewrite -mulrA mulrCA !mulrA.
+rewrite -!mulrA [X in _ * X]mulrCA !mulrA equivf_l.
+by rewrite mulrC !mulrA -mulrA mulrC mulrA.
+Qed.
+Canonical pi_add_morph := PiMorph2 pi_add.
+
+Definition oppf x : dom := Ratio (- \n_x) \d_x.
+Definition opp := lift_op1 {fraction R} oppf.
+Lemma pi_opp : {morph \pi : x / oppf x >-> opp x}.
+Proof.
+move=> x; unlock opp; apply/eqmodP; rewrite /= /equivf /oppf /=.
+by rewrite !numden_Ratio ?(domP,mulf_neq0) // mulNr mulrN -equivf_r.
+Qed.
+Canonical pi_opp_morph := PiMorph1 pi_opp.
+
+Definition mulf x y : dom := Ratio (\n_x * \n_y) (\d_x * \d_y).
+Definition mul := lift_op2 {fraction R} mulf.
+
+Lemma pi_mul : {morph \pi : x y / mulf x y >-> mul x y}.
+Proof.
+move=> x y; unlock mul; apply/eqmodP=> /=.
+rewrite equivfE /= /addf /= !numden_Ratio ?mulf_neq0 ?domP //.
+rewrite mulrAC !mulrA -mulrA equivf_r -equivf_l.
+by rewrite mulrA ![_ * \d_y]mulrC !mulrA.
+Qed.
+Canonical pi_mul_morph := PiMorph2 pi_mul.
+
+Definition invf x : dom := Ratio \d_x \n_x.
+Definition inv := lift_op1 {fraction R} invf.
+
+Lemma pi_inv : {morph \pi : x / invf x >-> inv x}.
+Proof.
+move=> x; unlock inv; apply/eqmodP=> /=; rewrite equivfE /invf eq_sym.
+do 2?case: RatioP=> /= [/eqP|];
+ rewrite ?mul0r ?mul1r -?equivf_def ?numer0 ?reprK //.
+ by move=> hx /eqP hx'; rewrite hx' eqxx in hx.
+by move=> /eqP ->; rewrite eqxx.
+Qed.
+Canonical pi_inv_morph := PiMorph1 pi_inv.
+
+Lemma addA : associative add.
+Proof.
+elim/quotW=> x; elim/quotW=> y; elim/quotW=> z; rewrite !piE.
+rewrite /addf /= !numden_Ratio ?mulf_neq0 ?domP // !mulrDl !mulrA !addrA.
+by congr (\pi (Ratio (_ + _ + _) _)); rewrite mulrAC.
+Qed.
+
+Lemma addC : commutative add.
+Proof.
+by elim/quotW=> x; elim/quotW=> y; rewrite !piE /addf addrC [\d__ * _]mulrC.
+Qed.
+
+Lemma add0_l : left_id 0%:F add.
+Proof.
+elim/quotW=> x; rewrite !piE /addf !numden_Ratio ?oner_eq0 //.
+by rewrite mul0r mul1r mulr1 add0r Ratio_numden.
+Qed.
+
+Lemma addN_l : left_inverse 0%:F opp add.
+Proof.
+elim/quotW=> x; apply/eqP; rewrite piE /equivf.
+rewrite /addf /oppf !numden_Ratio ?(oner_eq0, mulf_neq0, domP) //.
+by rewrite mulr1 mulr0 mulNr addNr.
+Qed.
+
+(* fracions form an abelian group *)
+Definition frac_zmodMixin := ZmodMixin addA addC add0_l addN_l.
+Canonical frac_zmodType := Eval hnf in ZmodType type frac_zmodMixin.
+
+Lemma mulA : associative mul.
+Proof.
+elim/quotW=> x; elim/quotW=> y; elim/quotW=> z; rewrite !piE.
+by rewrite /mulf !numden_Ratio ?mulf_neq0 ?domP // !mulrA.
+Qed.
+
+Lemma mulC : commutative mul.
+Proof.
+elim/quotW=> x; elim/quotW=> y; rewrite !piE /mulf.
+by rewrite [_ * (\d_x)]mulrC [_ * (\n_x)]mulrC.
+Qed.
+
+Lemma mul1_l : left_id 1%:F mul.
+Proof.
+elim/quotW=> x; rewrite !piE /mulf.
+by rewrite !numden_Ratio ?oner_eq0 // !mul1r Ratio_numden.
+Qed.
+
+Lemma mul_addl : left_distributive mul add.
+Proof.
+elim/quotW=> x; elim/quotW=> y; elim/quotW=> z; apply/eqP.
+rewrite !piE /equivf /mulf /addf !numden_Ratio ?mulf_neq0 ?domP //; apply/eqP.
+rewrite !(mulrDr, mulrDl) !mulrA; congr (_ * _ + _ * _).
+ rewrite ![_ * \n_z]mulrC -!mulrA; congr (_ * _).
+ rewrite ![\d_y * _]mulrC !mulrA; congr (_ * _ * _).
+ by rewrite [X in _ = X]mulrC mulrA.
+rewrite ![_ * \n_z]mulrC -!mulrA; congr (_ * _).
+rewrite ![\d_x * _]mulrC !mulrA; congr (_ * _ * _).
+by rewrite -mulrA mulrC [X in X * _] mulrC.
+Qed.
+
+Lemma nonzero1 : 1%:F != 0%:F :> type.
+Proof. by rewrite piE equivfE !numden_Ratio ?mul1r ?oner_eq0. Qed.
+
+(* fracions form a commutative ring *)
+Definition frac_comRingMixin := ComRingMixin mulA mulC mul1_l mul_addl nonzero1.
+Canonical frac_ringType := Eval hnf in RingType type frac_comRingMixin.
+Canonical frac_comRingType := Eval hnf in ComRingType type mulC.
+
+Lemma mulV_l : forall a, a != 0%:F -> mul (inv a) a = 1%:F.
+Proof.
+elim/quotW=> x /=; rewrite !piE.
+rewrite /equivf !numden_Ratio ?oner_eq0 // mulr1 mulr0=> nx0.
+apply/eqmodP; rewrite /= equivfE.
+by rewrite !numden_Ratio ?(oner_eq0, mulf_neq0, domP) // !mulr1 mulrC.
+Qed.
+
+Lemma inv0 : inv 0%:F = 0%:F.
+Proof.
+rewrite !piE /invf !numden_Ratio ?oner_eq0 // /Ratio /insubd.
+do 2?case: insubP; rewrite //= ?eqxx ?oner_eq0 // => u _ hu _.
+by congr \pi; apply: val_inj; rewrite /= hu.
+Qed.
+
+(* fractions form a ring with explicit unit *)
+Definition RatFieldUnitMixin := FieldUnitMixin mulV_l inv0.
+Canonical frac_unitRingType := Eval hnf in UnitRingType type RatFieldUnitMixin.
+Canonical frac_comUnitRingType := [comUnitRingType of type].
+
+Lemma field_axiom : GRing.Field.mixin_of frac_unitRingType.
+Proof. exact. Qed.
+
+(* fractions form a field *)
+Definition RatFieldIdomainMixin := (FieldIdomainMixin field_axiom).
+Canonical frac_idomainType :=
+ Eval hnf in IdomainType type (FieldIdomainMixin field_axiom).
+Canonical frac_fieldType := FieldType type field_axiom.
+
+End FracField.
+End FracField.
+
+Notation "{ 'fraction' T }" := (FracField.type_of (Phant T)).
+Notation equivf := (@FracField.equivf _).
+Hint Resolve denom_ratioP.
+
+Section Canonicals.
+
+Variable R : idomainType.
+
+(* reexporting the structures *)
+Canonical FracField.frac_quotType.
+Canonical FracField.frac_eqType.
+Canonical FracField.frac_choiceType.
+Canonical FracField.frac_zmodType.
+Canonical FracField.frac_ringType.
+Canonical FracField.frac_comRingType.
+Canonical FracField.frac_unitRingType.
+Canonical FracField.frac_comUnitRingType.
+Canonical FracField.frac_idomainType.
+Canonical FracField.frac_fieldType.
+Canonical FracField.tofrac_pi_morph.
+Canonical frac_of_quotType := Eval hnf in [quotType of {fraction R}].
+Canonical frac_of_eqType := Eval hnf in [eqType of {fraction R}].
+Canonical frac_of_choiceType := Eval hnf in [choiceType of {fraction R}].
+Canonical frac_of_zmodType := Eval hnf in [zmodType of {fraction R}].
+Canonical frac_of_ringType := Eval hnf in [ringType of {fraction R}].
+Canonical frac_of_comRingType := Eval hnf in [comRingType of {fraction R}].
+Canonical frac_of_unitRingType := Eval hnf in [unitRingType of {fraction R}].
+Canonical frac_of_comUnitRingType := Eval hnf in [comUnitRingType of {fraction R}].
+Canonical frac_of_idomainType := Eval hnf in [idomainType of {fraction R}].
+Canonical frac_of_fieldType := Eval hnf in [fieldType of {fraction R}].
+
+End Canonicals.
+
+Section FracFieldTheory.
+
+Import FracField.
+
+Variable R : idomainType.
+
+Lemma Ratio_numden (x : {ratio R}) : Ratio \n_x \d_x = x.
+Proof. exact: FracField.Ratio_numden. Qed.
+
+(* exporting the embeding from R to {fraction R} *)
+Local Notation tofrac := (@FracField.tofrac R).
+Local Notation "x %:F" := (tofrac x).
+
+Lemma tofrac_is_additive: additive tofrac.
+Proof.
+move=> p q /=; unlock tofrac.
+rewrite -[X in _ = _ + X]pi_opp -[X in _ = X]pi_add.
+by rewrite /addf /oppf /= !numden_Ratio ?(oner_neq0, mul1r, mulr1).
+Qed.
+
+Canonical tofrac_additive := Additive tofrac_is_additive.
+
+Lemma tofrac_is_multiplicative: multiplicative tofrac.
+Proof.
+split=> [p q|//]; unlock tofrac; rewrite -[X in _ = X]pi_mul.
+by rewrite /mulf /= !numden_Ratio ?(oner_neq0, mul1r, mulr1).
+Qed.
+
+Canonical tofrac_rmorphism := AddRMorphism tofrac_is_multiplicative.
+
+(* tests *)
+Lemma tofrac0 : 0%:F = 0. Proof. exact: rmorph0. Qed.
+Lemma tofracN : {morph tofrac: x / - x}. Proof. exact: rmorphN. Qed.
+Lemma tofracD : {morph tofrac: x y / x + y}. Proof. exact: rmorphD. Qed.
+Lemma tofracB : {morph tofrac: x y / x - y}. Proof. exact: rmorphB. Qed.
+Lemma tofracMn n : {morph tofrac: x / x *+ n}. Proof. exact: rmorphMn. Qed.
+Lemma tofracMNn n : {morph tofrac: x / x *- n}. Proof. exact: rmorphMNn. Qed.
+Lemma tofrac1 : 1%:F = 1. Proof. exact: rmorph1. Qed.
+Lemma tofracM : {morph tofrac: x y / x * y}. Proof. exact: rmorphM. Qed.
+Lemma tofracX n : {morph tofrac: x / x ^+ n}. Proof. exact: rmorphX. Qed.
+
+Lemma tofrac_eq (p q : R): (p%:F == q%:F) = (p == q).
+Proof.
+apply/eqP/eqP=> [|->//]; unlock tofrac=> /eqmodP /eqP /=.
+by rewrite !numden_Ratio ?(oner_eq0, mul1r, mulr1).
+Qed.
+
+Lemma tofrac_eq0 (p : R): (p%:F == 0) = (p == 0).
+Proof. by rewrite tofrac_eq. Qed.
+End FracFieldTheory.
diff --git a/mathcomp/algebra/intdiv.v b/mathcomp/algebra/intdiv.v
new file mode 100644
index 0000000..4f1ce95
--- /dev/null
+++ b/mathcomp/algebra/intdiv.v
@@ -0,0 +1,1076 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice.
+Require Import fintype tuple finfun bigop prime ssralg poly ssrnum ssrint rat.
+Require Import polydiv finalg perm zmodp matrix mxalgebra vector.
+
+(******************************************************************************)
+(* This file provides various results on divisibility of integers. *)
+(* It defines, for m, n, d : int, *)
+(* (m %% d)%Z == the remainder of the Euclidean division of m by d; this is *)
+(* the least non-negative element of the coset m + dZ when *)
+(* d != 0, and m if d = 0. *)
+(* (m %/ d)%Z == the quotient of the Euclidean division of m by d, such *)
+(* that m = (m %/ d)%Z * d + (m %% d)%Z. Since for d != 0 the *)
+(* remainder is non-negative, (m %/ d)%Z is non-zero for *)
+(* (d %| m)%Z <=> m is divisible by d; dvdz d is the (collective) predicate *)
+(* for integers divisible by d, and (d %| m)%Z is actually *)
+(* (transposing) notation for m \in dvdz d. *)
+(* (m = n %[mod d])%Z, (m == n %[mod d])%Z, (m != n %[mod d])%Z *)
+(* m and n are (resp. compare, don't compare) equal mod d. *)
+(* gcdz m n == the (non-negative) greatest common divisor of m and n, *)
+(* with gcdz 0 0 = 0. *)
+(* coprimez m n <=> m and n are coprime. *)
+(* egcdz m n == the Bezout coefficients of the gcd of m and n: a pair *)
+(* (u, v) of coprime integers such that u*m + v*n = gcdz m n. *)
+(* Alternatively, a Bezoutz lemma states such u and v exist. *)
+(* zchinese m1 m2 n1 n2 == for coprime m1 and m2, a solution to the Chinese *)
+(* remainder problem for n1 and n2, i.e., and integer n such *)
+(* that n = n1 %[mod m1] and n = n2 %[mod m2]. *)
+(* zcontents p == the contents of p : {poly int}, that is, the gcd of the *)
+(* coefficients of p, with the lead coefficient of p, *)
+(* zprimitive p == the primitive part of p : {poly int}, i.e., p divided by *)
+(* its contents. *)
+(* inIntSpan X v <-> v is an integral linear combination of elements of *)
+(* X : seq V, where V is a zmodType. We prove that this is a *)
+(* decidable property for Q-vector spaces. *)
+(* int_Smith_normal_form :: a theorem asserting the existence of the Smith *)
+(* normal form for integer matrices. *)
+(* Note that many of the concepts and results in this file could and perhaps *)
+(* sould be generalized to the more general setting of integral, unique *)
+(* factorization, principal ideal, or Euclidean domains. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GRing.Theory Num.Theory.
+Local Open Scope ring_scope.
+
+Definition divz (m d : int) :=
+ let: (K, n) := match m with Posz n => (Posz, n) | Negz n => (Negz, n) end in
+ sgz d * K (n %/ `|d|)%N.
+
+Definition modz (m d : int) : int := m - divz m d * d.
+
+Definition dvdz d m := (`|d| %| `|m|)%N.
+
+Definition gcdz m n := (gcdn `|m| `|n|)%:Z.
+
+Definition egcdz m n : int * int :=
+ if m == 0 then (0, (-1) ^+ (n < 0)%R) else
+ let: (u, v) := egcdn `|m| `|n| in (sgz m * u, - (-1) ^+ (n < 0)%R * v%:Z).
+
+Definition coprimez m n := (gcdz m n == 1).
+
+Infix "%/" := divz : int_scope.
+Infix "%%" := modz : int_scope.
+Notation "d %| m" := (m \in dvdz d) : int_scope.
+Notation "m = n %[mod d ]" := (modz m d = modz n d) : int_scope.
+Notation "m == n %[mod d ]" := (modz m d == modz n d) : int_scope.
+Notation "m <> n %[mod d ]" := (modz m d <> modz n d) : int_scope.
+Notation "m != n %[mod d ]" := (modz m d != modz n d) : int_scope.
+
+Lemma divz_nat (n d : nat) : (n %/ d)%Z = (n %/ d)%N.
+Proof. by case: d => // d; rewrite /divz /= mul1r. Qed.
+
+Lemma divzN m d : (m %/ - d)%Z = - (m %/ d)%Z.
+Proof. by case: m => n; rewrite /divz /= sgzN abszN mulNr. Qed.
+
+Lemma divz_abs m d : (m %/ `|d|)%Z = (-1) ^+ (d < 0)%R * (m %/ d)%Z.
+Proof.
+by rewrite {3}[d]intEsign !mulr_sign; case: ifP => -> //; rewrite divzN opprK.
+Qed.
+
+Lemma div0z d : (0 %/ d)%Z = 0.
+Proof.
+by rewrite -(canLR (signrMK _) (divz_abs _ _)) (divz_nat 0) div0n mulr0.
+Qed.
+
+Lemma divNz_nat m d : (d > 0)%N -> (Negz m %/ d)%Z = - (m %/ d).+1%:Z.
+Proof. by case: d => // d _; apply: mul1r. Qed.
+
+Lemma divz_eq m d : m = (m %/ d)%Z * d + (m %% d)%Z.
+Proof. by rewrite addrC subrK. Qed.
+
+Lemma modzN m d : (m %% - d)%Z = (m %% d)%Z.
+Proof. by rewrite /modz divzN mulrNN. Qed.
+
+Lemma modz_abs m d : (m %% `|d|%N)%Z = (m %% d)%Z.
+Proof. by rewrite {2}[d]intEsign mulr_sign; case: ifP; rewrite ?modzN. Qed.
+
+Lemma modz_nat (m d : nat) : (m %% d)%Z = (m %% d)%N.
+Proof.
+by apply: (canLR (addrK _)); rewrite addrC divz_nat {1}(divn_eq m d).
+Qed.
+
+Lemma modNz_nat m d : (d > 0)%N -> (Negz m %% d)%Z = d%:Z - 1 - (m %% d)%:Z.
+Proof.
+rewrite /modz => /divNz_nat->; apply: (canLR (addrK _)).
+rewrite -!addrA -!opprD -!PoszD -opprB mulnSr !addnA PoszD addrK.
+by rewrite addnAC -addnA mulnC -divn_eq.
+Qed.
+
+Lemma modz_ge0 m d : d != 0 -> 0 <= (m %% d)%Z.
+Proof.
+rewrite -absz_gt0 -modz_abs => d_gt0.
+case: m => n; rewrite ?modNz_nat ?modz_nat // -addrA -opprD subr_ge0.
+by rewrite lez_nat ltn_mod.
+Qed.
+
+Lemma divz0 m : (m %/ 0)%Z = 0. Proof. by case: m. Qed.
+Lemma mod0z d : (0 %% d)%Z = 0. Proof. by rewrite /modz div0z mul0r subrr. Qed.
+Lemma modz0 m : (m %% 0)%Z = m. Proof. by rewrite /modz mulr0 subr0. Qed.
+
+Lemma divz_small m d : 0 <= m < `|d|%:Z -> (m %/ d)%Z = 0.
+Proof.
+rewrite -(canLR (signrMK _) (divz_abs _ _)); case: m => // n /divn_small.
+by rewrite divz_nat => ->; rewrite mulr0.
+Qed.
+
+Lemma divzMDl q m d : d != 0 -> ((q * d + m) %/ d)%Z = q + (m %/ d)%Z.
+Proof.
+rewrite neqr_lt -oppr_gt0 => nz_d.
+wlog{nz_d} d_gt0: q d / d > 0; last case: d => // d in d_gt0 *.
+ move=> IH; case/orP: nz_d => /IH// /(_ (- q)).
+ by rewrite mulrNN !divzN -opprD => /oppr_inj.
+wlog q_gt0: q m / q >= 0; last case: q q_gt0 => // q _.
+ move=> IH; case: q => n; first exact: IH; rewrite NegzE mulNr.
+ by apply: canRL (addKr _) _; rewrite -IH ?addNKr.
+case: m => n; first by rewrite !divz_nat divnMDl.
+have [le_qd_n | lt_qd_n] := leqP (q * d) n.
+ rewrite divNz_nat // NegzE -(subnKC le_qd_n) divnMDl //.
+ by rewrite -!addnS !PoszD !opprD !addNKr divNz_nat.
+rewrite divNz_nat // NegzE -PoszM subzn // divz_nat.
+apply: canRL (addrK _) _; congr _%:Z; rewrite addnC -divnMDl // mulSnr.
+rewrite -{3}(subnKC (ltn_pmod n d_gt0)) addnA addnS -divn_eq addnAC.
+by rewrite subnKC // divnMDl // divn_small ?addn0 // subnSK ?ltn_mod ?leq_subr.
+Qed.
+
+Lemma mulzK m d : d != 0 -> (m * d %/ d)%Z = m.
+Proof. by move=> d_nz; rewrite -[m * d]addr0 divzMDl // div0z addr0. Qed.
+
+Lemma mulKz m d : d != 0 -> (d * m %/ d)%Z = m.
+Proof. by move=> d_nz; rewrite mulrC mulzK. Qed.
+
+Lemma expzB p m n : p != 0 -> (m >= n)%N -> p ^+ (m - n) = (p ^+ m %/ p ^+ n)%Z.
+Proof. by move=> p_nz /subnK{2}<-; rewrite exprD mulzK // expf_neq0. Qed.
+
+Lemma modz1 m : (m %% 1)%Z = 0.
+Proof. by case: m => n; rewrite (modNz_nat, modz_nat) ?modn1. Qed.
+
+Lemma divn1 m : (m %/ 1)%Z = m. Proof. by rewrite -{1}[m]mulr1 mulzK. Qed.
+
+Lemma divzz d : (d %/ d)%Z = (d != 0).
+Proof. by have [-> // | d_nz] := altP eqP; rewrite -{1}[d]mul1r mulzK. Qed.
+
+Lemma ltz_pmod m d : d > 0 -> (m %% d)%Z < d.
+Proof.
+case: m d => n [] // d d_gt0; first by rewrite modz_nat ltz_nat ltn_pmod.
+by rewrite modNz_nat // -lez_addr1 addrAC subrK ger_addl oppr_le0.
+Qed.
+
+Lemma ltz_mod m d : d != 0 -> (m %% d)%Z < `|d|.
+Proof. by rewrite -absz_gt0 -modz_abs => d_gt0; apply: ltz_pmod. Qed.
+
+Lemma divzMpl p m d : p > 0 -> (p * m %/ (p * d) = m %/ d)%Z.
+Proof.
+case: p => // p p_gt0; wlog d_gt0: d / d > 0; last case: d => // d in d_gt0 *.
+ by move=> IH; case/intP: d => [|d|d]; rewrite ?mulr0 ?divz0 ?mulrN ?divzN ?IH.
+rewrite {1}(divz_eq m d) mulrDr mulrCA divzMDl ?mulf_neq0 ?gtr_eqF // addrC.
+rewrite divz_small ?add0r // PoszM pmulr_rge0 ?modz_ge0 ?gtr_eqF //=.
+by rewrite ltr_pmul2l ?ltz_pmod.
+Qed.
+Implicit Arguments divzMpl [p m d].
+
+Lemma divzMpr p m d : p > 0 -> (m * p %/ (d * p) = m %/ d)%Z.
+Proof. by move=> p_gt0; rewrite -!(mulrC p) divzMpl. Qed.
+Implicit Arguments divzMpr [p m d].
+
+Lemma lez_floor m d : d != 0 -> (m %/ d)%Z * d <= m.
+Proof. by rewrite -subr_ge0; apply: modz_ge0. Qed.
+
+(* leq_mod does not extend to negative m. *)
+Lemma lez_div m d : (`|(m %/ d)%Z| <= `|m|)%N.
+Proof.
+wlog d_gt0: d / d > 0; last case: d d_gt0 => // d d_gt0.
+ by move=> IH; case/intP: d => [|n|n]; rewrite ?divz0 ?divzN ?abszN // IH.
+case: m => n; first by rewrite divz_nat leq_div.
+by rewrite divNz_nat // NegzE !abszN ltnS leq_div.
+Qed.
+
+Lemma ltz_ceil m d : d > 0 -> m < ((m %/ d)%Z + 1) * d.
+Proof.
+by case: d => // d d_gt0; rewrite mulrDl mul1r -ltr_subl_addl ltz_mod ?gtr_eqF.
+Qed.
+
+Lemma ltz_divLR m n d : d > 0 -> ((m %/ d)%Z < n) = (m < n * d).
+Proof.
+move=> d_gt0; apply/idP/idP.
+ by rewrite -lez_addr1 -(ler_pmul2r d_gt0); apply: ltr_le_trans (ltz_ceil _ _).
+rewrite -(ltr_pmul2r d_gt0 _ n) //; apply: ler_lt_trans (lez_floor _ _).
+by rewrite gtr_eqF.
+Qed.
+
+Lemma lez_divRL m n d : d > 0 -> (m <= (n %/ d)%Z) = (m * d <= n).
+Proof. by move=> d_gt0; rewrite !lerNgt ltz_divLR. Qed.
+
+Lemma divz_ge0 m d : d > 0 -> ((m %/ d)%Z >= 0) = (m >= 0).
+Proof. by case: d m => // d [] n d_gt0; rewrite (divz_nat, divNz_nat). Qed.
+
+Lemma divzMA_ge0 m n p : n >= 0 -> (m %/ (n * p) = (m %/ n)%Z %/ p)%Z.
+Proof.
+case: n => // [[|n]] _; first by rewrite mul0r !divz0 div0z.
+wlog p_gt0: p / p > 0; last case: p => // p in p_gt0 *.
+ by case/intP: p => [|p|p] IH; rewrite ?mulr0 ?divz0 ?mulrN ?divzN // IH.
+rewrite {2}(divz_eq m (n.+1%:Z * p)) mulrA mulrAC !divzMDl // ?gtr_eqF //.
+rewrite [rhs in _ + rhs]divz_small ?addr0 // ltz_divLR // divz_ge0 //.
+by rewrite mulrC ltz_pmod ?modz_ge0 ?gtr_eqF ?pmulr_lgt0.
+Qed.
+
+Lemma modz_small m d : 0 <= m < d -> (m %% d)%Z = m.
+Proof. by case: m d => //= m [] // d; rewrite modz_nat => /modn_small->. Qed.
+
+Lemma modz_mod m d : ((m %% d)%Z = m %[mod d])%Z.
+Proof.
+rewrite -!(modz_abs _ d); case: {d}`|d|%N => [|d]; first by rewrite !modz0.
+by rewrite modz_small ?modz_ge0 ?ltz_mod.
+Qed.
+
+Lemma modzMDl p m d : (p * d + m = m %[mod d])%Z.
+Proof.
+have [-> | d_nz] := eqVneq d 0; first by rewrite mulr0 add0r.
+by rewrite /modz divzMDl // mulrDl opprD addrACA subrr add0r.
+Qed.
+
+Lemma mulz_modr {p m d} : 0 < p -> p * (m %% d)%Z = ((p * m) %% (p * d))%Z.
+Proof.
+case: p => // p p_gt0; rewrite mulrBr; apply: canLR (addrK _) _.
+by rewrite mulrCA -(divzMpl p_gt0) subrK.
+Qed.
+
+Lemma mulz_modl {p m d} : 0 < p -> (m %% d)%Z * p = ((m * p) %% (d * p))%Z.
+Proof. by rewrite -!(mulrC p); apply: mulz_modr. Qed.
+
+Lemma modzDl m d : (d + m = m %[mod d])%Z.
+Proof. by rewrite -{1}[d]mul1r modzMDl. Qed.
+
+Lemma modzDr m d : (m + d = m %[mod d])%Z.
+Proof. by rewrite addrC modzDl. Qed.
+
+Lemma modzz d : (d %% d)%Z = 0.
+Proof. by rewrite -{1}[d]addr0 modzDl mod0z. Qed.
+
+Lemma modzMl p d : (p * d %% d)%Z = 0.
+Proof. by rewrite -[p * d]addr0 modzMDl mod0z. Qed.
+
+Lemma modzMr p d : (d * p %% d)%Z = 0.
+Proof. by rewrite mulrC modzMl. Qed.
+
+Lemma modzDml m n d : ((m %% d)%Z + n = m + n %[mod d])%Z.
+Proof. by rewrite {2}(divz_eq m d) -[_ * d + _ + n]addrA modzMDl. Qed.
+
+Lemma modzDmr m n d : (m + (n %% d)%Z = m + n %[mod d])%Z.
+Proof. by rewrite !(addrC m) modzDml. Qed.
+
+Lemma modzDm m n d : ((m %% d)%Z + (n %% d)%Z = m + n %[mod d])%Z.
+Proof. by rewrite modzDml modzDmr. Qed.
+
+Lemma eqz_modDl p m n d : (p + m == p + n %[mod d])%Z = (m == n %[mod d])%Z.
+Proof.
+have [-> | d_nz] := eqVneq d 0; first by rewrite !modz0 (inj_eq (addrI p)).
+apply/eqP/eqP=> eq_mn; last by rewrite -modzDmr eq_mn modzDmr.
+by rewrite -(addKr p m) -modzDmr eq_mn modzDmr addKr.
+Qed.
+
+Lemma eqz_modDr p m n d : (m + p == n + p %[mod d])%Z = (m == n %[mod d])%Z.
+Proof. by rewrite -!(addrC p) eqz_modDl. Qed.
+
+Lemma modzMml m n d : ((m %% d)%Z * n = m * n %[mod d])%Z.
+Proof. by rewrite {2}(divz_eq m d) mulrDl mulrAC modzMDl. Qed.
+
+Lemma modzMmr m n d : (m * (n %% d)%Z = m * n %[mod d])%Z.
+Proof. by rewrite !(mulrC m) modzMml. Qed.
+
+Lemma modzMm m n d : ((m %% d)%Z * (n %% d)%Z = m * n %[mod d])%Z.
+Proof. by rewrite modzMml modzMmr. Qed.
+
+Lemma modzXm k m d : ((m %% d)%Z ^+ k = m ^+ k %[mod d])%Z.
+Proof. by elim: k => // k IHk; rewrite !exprS -modzMmr IHk modzMm. Qed.
+
+Lemma modzNm m d : (- (m %% d)%Z = - m %[mod d])%Z.
+Proof. by rewrite -mulN1r modzMmr mulN1r. Qed.
+
+Lemma modz_absm m d : ((-1) ^+ (m < 0)%R * (m %% d)%Z = `|m|%:Z %[mod d])%Z.
+Proof. by rewrite modzMmr -abszEsign. Qed.
+
+(** Divisibility **)
+
+Fact dvdz_key d : pred_key (dvdz d). Proof. by []. Qed.
+Canonical dvdz_keyed d := KeyedPred (dvdz_key d).
+
+Lemma dvdzE d m : (d %| m)%Z = (`|d| %| `|m|)%N. Proof. by []. Qed.
+Lemma dvdz0 d : (d %| 0)%Z. Proof. exact: dvdn0. Qed.
+Lemma dvd0z n : (0 %| n)%Z = (n == 0). Proof. by rewrite -absz_eq0 -dvd0n. Qed.
+Lemma dvdz1 d : (d %| 1)%Z = (`|d|%N == 1%N). Proof. exact: dvdn1. Qed.
+Lemma dvd1z m : (1 %| m)%Z. Proof. exact: dvd1n. Qed.
+Lemma dvdzz m : (m %| m)%Z. Proof. exact: dvdnn. Qed.
+
+Lemma dvdz_mull d m n : (d %| n)%Z -> (d %| m * n)%Z.
+Proof. by rewrite !dvdzE abszM; apply: dvdn_mull. Qed.
+
+Lemma dvdz_mulr d m n : (d %| m)%Z -> (d %| m * n)%Z.
+Proof. by move=> d_m; rewrite mulrC dvdz_mull. Qed.
+Hint Resolve dvdz0 dvd1z dvdzz dvdz_mull dvdz_mulr.
+
+Lemma dvdz_mul d1 d2 m1 m2 : (d1 %| m1 -> d2 %| m2 -> d1 * d2 %| m1 * m2)%Z.
+Proof. by rewrite !dvdzE !abszM; apply: dvdn_mul. Qed.
+
+Lemma dvdz_trans n d m : (d %| n -> n %| m -> d %| m)%Z.
+Proof. by rewrite !dvdzE; apply: dvdn_trans. Qed.
+
+Lemma dvdzP d m : reflect (exists q, m = q * d) (d %| m)%Z.
+Proof.
+apply: (iffP dvdnP) => [] [q Dm]; last by exists `|q|%N; rewrite Dm abszM.
+exists ((-1) ^+ (m < 0)%R * q%:Z * (-1) ^+ (d < 0)%R).
+by rewrite -!mulrA -abszEsign -PoszM -Dm -intEsign.
+Qed.
+Implicit Arguments dvdzP [d m].
+
+Lemma dvdz_mod0P d m : reflect (m %% d = 0)%Z (d %| m)%Z.
+Proof.
+apply: (iffP dvdzP) => [[q ->] | md0]; first by rewrite modzMl.
+by rewrite (divz_eq m d) md0 addr0; exists (m %/ d)%Z.
+Qed.
+Implicit Arguments dvdz_mod0P [d m].
+
+Lemma dvdz_eq d m : (d %| m)%Z = ((m %/ d)%Z * d == m).
+Proof. by rewrite (sameP dvdz_mod0P eqP) subr_eq0 eq_sym. Qed.
+
+Lemma divzK d m : (d %| m)%Z -> (m %/ d)%Z * d = m.
+Proof. by rewrite dvdz_eq => /eqP. Qed.
+
+Lemma lez_divLR d m n : 0 < d -> (d %| m)%Z -> ((m %/ d)%Z <= n) = (m <= n * d).
+Proof. by move=> /ler_pmul2r <- /divzK->. Qed.
+
+Lemma ltz_divRL d m n : 0 < d -> (d %| m)%Z -> (n < m %/ d)%Z = (n * d < m).
+Proof. by move=> /ltr_pmul2r <- /divzK->. Qed.
+
+Lemma eqz_div d m n : d != 0 -> (d %| m)%Z -> (n == m %/ d)%Z = (n * d == m).
+Proof. by move=> /mulIf/inj_eq <- /divzK->. Qed.
+
+Lemma eqz_mul d m n : d != 0 -> (d %| m)%Z -> (m == n * d) = (m %/ d == n)%Z.
+Proof. by move=> d_gt0 dv_d_m; rewrite eq_sym -eqz_div // eq_sym. Qed.
+
+Lemma divz_mulAC d m n : (d %| m)%Z -> (m %/ d)%Z * n = (m * n %/ d)%Z.
+Proof.
+have [-> | d_nz] := eqVneq d 0; first by rewrite !divz0 mul0r.
+by move/divzK=> {2} <-; rewrite mulrAC mulzK.
+Qed.
+
+Lemma mulz_divA d m n : (d %| n)%Z -> m * (n %/ d)%Z = (m * n %/ d)%Z.
+Proof. by move=> dv_d_m; rewrite !(mulrC m) divz_mulAC. Qed.
+
+Lemma mulz_divCA d m n :
+ (d %| m)%Z -> (d %| n)%Z -> m * (n %/ d)%Z = n * (m %/ d)%Z.
+Proof. by move=> dv_d_m dv_d_n; rewrite mulrC divz_mulAC ?mulz_divA. Qed.
+
+Lemma divzA m n p : (p %| n -> n %| m * p -> m %/ (n %/ p)%Z = m * p %/ n)%Z.
+Proof.
+move/divzK=> p_dv_n; have [->|] := eqVneq n 0; first by rewrite div0z !divz0.
+rewrite -{1 2}p_dv_n mulf_eq0 => /norP[pn_nz p_nz] /divzK; rewrite mulrA p_dv_n.
+by move/mulIf=> {1} <- //; rewrite mulzK.
+Qed.
+
+Lemma divzMA m n p : (n * p %| m -> m %/ (n * p) = (m %/ n)%Z %/ p)%Z.
+Proof.
+have [-> | nz_p] := eqVneq p 0; first by rewrite mulr0 !divz0.
+have [-> | nz_n] := eqVneq n 0; first by rewrite mul0r !divz0 div0z.
+by move/divzK=> {2} <-; rewrite mulrA mulrAC !mulzK.
+Qed.
+
+Lemma divzAC m n p : (n * p %| m -> (m %/ n)%Z %/ p = (m %/ p)%Z %/ n)%Z.
+Proof. by move=> np_dv_mn; rewrite -!divzMA // mulrC. Qed.
+
+Lemma divzMl p m d : p != 0 -> (d %| m -> p * m %/ (p * d) = m %/ d)%Z.
+Proof.
+have [-> | nz_d nz_p] := eqVneq d 0; first by rewrite mulr0 !divz0.
+by move/divzK=> {1}<-; rewrite mulrCA mulzK ?mulf_neq0.
+Qed.
+
+Lemma divzMr p m d : p != 0 -> (d %| m -> m * p %/ (d * p) = m %/ d)%Z.
+Proof. by rewrite -!(mulrC p); apply: divzMl. Qed.
+
+Lemma dvdz_mul2l p d m : p != 0 -> (p * d %| p * m)%Z = (d %| m)%Z.
+Proof. by rewrite !dvdzE -absz_gt0 !abszM; apply: dvdn_pmul2l. Qed.
+Implicit Arguments dvdz_mul2l [p m d].
+
+Lemma dvdz_mul2r p d m : p != 0 -> (d * p %| m * p)%Z = (d %| m)%Z.
+Proof. by rewrite !dvdzE -absz_gt0 !abszM; apply: dvdn_pmul2r. Qed.
+Implicit Arguments dvdz_mul2r [p m d].
+
+Lemma dvdz_exp2l p m n : (m <= n)%N -> (p ^+ m %| p ^+ n)%Z.
+Proof. by rewrite dvdzE !abszX; apply: dvdn_exp2l. Qed.
+
+Lemma dvdz_Pexp2l p m n : `|p| > 1 -> (p ^+ m %| p ^+ n)%Z = (m <= n)%N.
+Proof. by rewrite dvdzE !abszX ltz_nat; apply: dvdn_Pexp2l. Qed.
+
+Lemma dvdz_exp2r m n k : (m %| n -> m ^+ k %| n ^+ k)%Z.
+Proof. by rewrite !dvdzE !abszX; apply: dvdn_exp2r. Qed.
+
+Fact dvdz_zmod_closed d : zmod_closed (dvdz d).
+Proof.
+split=> [|_ _ /dvdzP[p ->] /dvdzP[q ->]]; first exact: dvdz0.
+by rewrite -mulrBl dvdz_mull.
+Qed.
+Canonical dvdz_addPred d := AddrPred (dvdz_zmod_closed d).
+Canonical dvdz_oppPred d := OpprPred (dvdz_zmod_closed d).
+Canonical dvdz_zmodPred d := ZmodPred (dvdz_zmod_closed d).
+
+Lemma dvdz_exp k d m : (0 < k)%N -> (d %| m -> d %| m ^+ k)%Z.
+Proof. by case: k => // k _ d_dv_m; rewrite exprS dvdz_mulr. Qed.
+
+Lemma eqz_mod_dvd d m n : (m == n %[mod d])%Z = (d %| m - n)%Z.
+Proof.
+apply/eqP/dvdz_mod0P=> eq_mn.
+ by rewrite -modzDml eq_mn modzDml subrr mod0z.
+by rewrite -(subrK n m) -modzDml eq_mn add0r.
+Qed.
+
+Lemma divzDl m n d :
+ (d %| m)%Z -> ((m + n) %/ d)%Z = (m %/ d)%Z + (n %/ d)%Z.
+Proof.
+have [-> | d_nz] := eqVneq d 0; first by rewrite !divz0.
+by move/divzK=> {1}<-; rewrite divzMDl.
+Qed.
+
+Lemma divzDr m n d :
+ (d %| n)%Z -> ((m + n) %/ d)%Z = (m %/ d)%Z + (n %/ d)%Z.
+Proof. by move=> dv_n; rewrite addrC divzDl // addrC. Qed.
+
+Lemma Qint_dvdz (m d : int) : (d %| m)%Z -> ((m%:~R / d%:~R : rat) \is a Qint).
+Proof.
+case/dvdzP=> z ->; rewrite rmorphM /=; case: (altP (d =P 0)) => [->|dn0].
+ by rewrite mulr0 mul0r.
+by rewrite mulfK ?intr_eq0 // rpred_int.
+Qed.
+
+Lemma Qnat_dvd (m d : nat) : (d %| m)%N -> ((m%:R / d%:R : rat) \is a Qnat).
+Proof.
+move=> h; rewrite Qnat_def divr_ge0 ?ler0n // -[m%:R]/(m%:~R) -[d%:R]/(d%:~R).
+by rewrite Qint_dvdz.
+Qed.
+
+(* Greatest common divisor *)
+
+Lemma gcdzz m : gcdz m m = `|m|%:Z. Proof. by rewrite /gcdz gcdnn. Qed.
+Lemma gcdzC : commutative gcdz. Proof. by move=> m n; rewrite /gcdz gcdnC. Qed.
+Lemma gcd0z m : gcdz 0 m = `|m|%:Z. Proof. by rewrite /gcdz gcd0n. Qed.
+Lemma gcdz0 m : gcdz m 0 = `|m|%:Z. Proof. by rewrite /gcdz gcdn0. Qed.
+Lemma gcd1z : left_zero 1 gcdz. Proof. by move=> m; rewrite /gcdz gcd1n. Qed.
+Lemma gcdz1 : right_zero 1 gcdz. Proof. by move=> m; rewrite /gcdz gcdn1. Qed.
+Lemma dvdz_gcdr m n : (gcdz m n %| n)%Z. Proof. exact: dvdn_gcdr. Qed.
+Lemma dvdz_gcdl m n : (gcdz m n %| m)%Z. Proof. exact: dvdn_gcdl. Qed.
+Lemma gcdz_eq0 m n : (gcdz m n == 0) = (m == 0) && (n == 0).
+Proof. by rewrite -absz_eq0 eqn0Ngt gcdn_gt0 !negb_or -!eqn0Ngt !absz_eq0. Qed.
+Lemma gcdNz m n : gcdz (- m) n = gcdz m n. Proof. by rewrite /gcdz abszN. Qed.
+Lemma gcdzN m n : gcdz m (- n) = gcdz m n. Proof. by rewrite /gcdz abszN. Qed.
+
+Lemma gcdz_modr m n : gcdz m (n %% m)%Z = gcdz m n.
+Proof.
+rewrite -modz_abs /gcdz; move/absz: m => m.
+have [-> | m_gt0] := posnP m; first by rewrite modz0.
+case: n => n; first by rewrite modz_nat gcdn_modr.
+rewrite modNz_nat // NegzE abszN {2}(divn_eq n m) -addnS gcdnMDl.
+rewrite -addrA -opprD -intS /=; set m1 := _.+1.
+have le_m1m: (m1 <= m)%N by exact: ltn_pmod.
+by rewrite subzn // !(gcdnC m) -{2 3}(subnK le_m1m) gcdnDl gcdnDr gcdnC.
+Qed.
+
+Lemma gcdz_modl m n : gcdz (m %% n)%Z n = gcdz m n.
+Proof. by rewrite -!(gcdzC n) gcdz_modr. Qed.
+
+Lemma gcdzMDl q m n : gcdz m (q * m + n) = gcdz m n.
+Proof. by rewrite -gcdz_modr modzMDl gcdz_modr. Qed.
+
+Lemma gcdzDl m n : gcdz m (m + n) = gcdz m n.
+Proof. by rewrite -{2}(mul1r m) gcdzMDl. Qed.
+
+Lemma gcdzDr m n : gcdz m (n + m) = gcdz m n.
+Proof. by rewrite addrC gcdzDl. Qed.
+
+Lemma gcdzMl n m : gcdz n (m * n) = `|n|%:Z.
+Proof. by rewrite -[m * n]addr0 gcdzMDl gcdz0. Qed.
+
+Lemma gcdzMr n m : gcdz n (n * m) = `|n|%:Z.
+Proof. by rewrite mulrC gcdzMl. Qed.
+
+Lemma gcdz_idPl {m n} : reflect (gcdz m n = `|m|%:Z) (m %| n)%Z.
+Proof. by apply: (iffP gcdn_idPl) => [<- | []]. Qed.
+
+Lemma gcdz_idPr {m n} : reflect (gcdz m n = `|n|%:Z) (n %| m)%Z.
+Proof. by rewrite gcdzC; apply: gcdz_idPl. Qed.
+
+Lemma expz_min e m n : e >= 0 -> e ^+ minn m n = gcdz (e ^+ m) (e ^+ n).
+Proof.
+by case: e => // e _; rewrite /gcdz !abszX -expn_min -natz -natrX !natz.
+Qed.
+
+Lemma dvdz_gcd p m n : (p %| gcdz m n)%Z = (p %| m)%Z && (p %| n)%Z.
+Proof. exact: dvdn_gcd. Qed.
+
+Lemma gcdzAC : right_commutative gcdz.
+Proof. by move=> m n p; rewrite /gcdz gcdnAC. Qed.
+
+Lemma gcdzA : associative gcdz.
+Proof. by move=> m n p; rewrite /gcdz gcdnA. Qed.
+
+Lemma gcdzCA : left_commutative gcdz.
+Proof. by move=> m n p; rewrite /gcdz gcdnCA. Qed.
+
+Lemma gcdzACA : interchange gcdz gcdz.
+Proof. by move=> m n p q; rewrite /gcdz gcdnACA. Qed.
+
+Lemma mulz_gcdr m n p : `|m|%:Z * gcdz n p = gcdz (m * n) (m * p).
+Proof. by rewrite -PoszM muln_gcdr -!abszM. Qed.
+
+Lemma mulz_gcdl m n p : gcdz m n * `|p|%:Z = gcdz (m * p) (n * p).
+Proof. by rewrite -PoszM muln_gcdl -!abszM. Qed.
+
+Lemma mulz_divCA_gcd n m : n * (m %/ gcdz n m)%Z = m * (n %/ gcdz n m)%Z.
+Proof. by rewrite mulz_divCA ?dvdz_gcdl ?dvdz_gcdr. Qed.
+
+(* Not including lcm theory, for now. *)
+
+(* Coprime factors *)
+
+Lemma coprimezE m n : coprimez m n = coprime `|m| `|n|. Proof. by []. Qed.
+
+Lemma coprimez_sym : symmetric coprimez.
+Proof. by move=> m n; apply: coprime_sym. Qed.
+
+Lemma coprimeNz m n : coprimez (- m) n = coprimez m n.
+Proof. by rewrite coprimezE abszN. Qed.
+
+Lemma coprimezN m n : coprimez m (- n) = coprimez m n.
+Proof. by rewrite coprimezE abszN. Qed.
+
+CoInductive egcdz_spec m n : int * int -> Type :=
+ EgcdzSpec u v of u * m + v * n = gcdz m n & coprimez u v
+ : egcdz_spec m n (u, v).
+
+Lemma egcdzP m n : egcdz_spec m n (egcdz m n).
+Proof.
+rewrite /egcdz; have [-> | m_nz] := altP eqP.
+ by split; [rewrite -abszEsign gcd0z | rewrite coprimezE absz_sign].
+have m_gt0 : (`|m| > 0)%N by rewrite absz_gt0.
+case: egcdnP (coprime_egcdn `|n| m_gt0) => //= u v Duv _ co_uv; split.
+ rewrite !mulNr -!mulrA mulrCA -abszEsg mulrCA -abszEsign.
+ by rewrite -!PoszM Duv addnC PoszD addrK.
+by rewrite coprimezE abszM absz_sg m_nz mul1n mulNr abszN abszMsign.
+Qed.
+
+Lemma Bezoutz m n : {u : int & {v : int | u * m + v * n = gcdz m n}}.
+Proof. by exists (egcdz m n).1, (egcdz m n).2; case: egcdzP. Qed.
+
+Lemma coprimezP m n :
+ reflect (exists uv, uv.1 * m + uv.2 * n = 1) (coprimez m n).
+Proof.
+apply: (iffP eqP) => [<-| [[u v] /= Duv]].
+ by exists (egcdz m n); case: egcdzP.
+congr _%:Z; apply: gcdn_def; rewrite ?dvd1n // => d dv_d_n dv_d_m.
+by rewrite -(dvdzE d 1) -Duv [m]intEsg [n]intEsg rpredD ?dvdz_mull.
+Qed.
+
+Lemma Gauss_dvdz m n p :
+ coprimez m n -> (m * n %| p)%Z = (m %| p)%Z && (n %| p)%Z.
+Proof. by move/Gauss_dvd <-; rewrite -abszM. Qed.
+
+Lemma Gauss_dvdzr m n p : coprimez m n -> (m %| n * p)%Z = (m %| p)%Z.
+Proof. by rewrite dvdzE abszM => /Gauss_dvdr->. Qed.
+
+Lemma Gauss_dvdzl m n p : coprimez m p -> (m %| n * p)%Z = (m %| n)%Z.
+Proof. by rewrite mulrC; apply: Gauss_dvdzr. Qed.
+
+Lemma Gauss_gcdzr p m n : coprimez p m -> gcdz p (m * n) = gcdz p n.
+Proof. by rewrite /gcdz abszM => /Gauss_gcdr->. Qed.
+
+Lemma Gauss_gcdzl p m n : coprimez p n -> gcdz p (m * n) = gcdz p m.
+Proof. by move=> co_pn; rewrite mulrC Gauss_gcdzr. Qed.
+
+Lemma coprimez_mulr p m n : coprimez p (m * n) = coprimez p m && coprimez p n.
+Proof. by rewrite -coprime_mulr -abszM. Qed.
+
+Lemma coprimez_mull p m n : coprimez (m * n) p = coprimez m p && coprimez n p.
+Proof. by rewrite -coprime_mull -abszM. Qed.
+
+Lemma coprimez_pexpl k m n : (0 < k)%N -> coprimez (m ^+ k) n = coprimez m n.
+Proof. by rewrite /coprimez /gcdz abszX; apply: coprime_pexpl. Qed.
+
+Lemma coprimez_pexpr k m n : (0 < k)%N -> coprimez m (n ^+ k) = coprimez m n.
+Proof. by move=> k_gt0; rewrite !(coprimez_sym m) coprimez_pexpl. Qed.
+
+Lemma coprimez_expl k m n : coprimez m n -> coprimez (m ^+ k) n.
+Proof. by rewrite /coprimez /gcdz abszX; apply: coprime_expl. Qed.
+
+Lemma coprimez_expr k m n : coprimez m n -> coprimez m (n ^+ k).
+Proof. by rewrite !(coprimez_sym m); apply: coprimez_expl. Qed.
+
+Lemma coprimez_dvdl m n p : (m %| n)%N -> coprimez n p -> coprimez m p.
+Proof. exact: coprime_dvdl. Qed.
+
+Lemma coprimez_dvdr m n p : (m %| n)%N -> coprimez p n -> coprimez p m.
+Proof. exact: coprime_dvdr. Qed.
+
+Lemma dvdz_pexp2r m n k : (k > 0)%N -> (m ^+ k %| n ^+ k)%Z = (m %| n)%Z.
+Proof. by rewrite dvdzE !abszX; apply: dvdn_pexp2r. Qed.
+
+Section Chinese.
+
+(***********************************************************************)
+(* The chinese remainder theorem *)
+(***********************************************************************)
+
+Variables m1 m2 : int.
+Hypothesis co_m12 : coprimez m1 m2.
+
+Lemma zchinese_remainder x y :
+ (x == y %[mod m1 * m2])%Z = (x == y %[mod m1])%Z && (x == y %[mod m2])%Z.
+Proof. by rewrite !eqz_mod_dvd Gauss_dvdz. Qed.
+
+(***********************************************************************)
+(* A function that solves the chinese remainder problem *)
+(***********************************************************************)
+
+Definition zchinese r1 r2 :=
+ r1 * m2 * (egcdz m1 m2).2 + r2 * m1 * (egcdz m1 m2).1.
+
+Lemma zchinese_modl r1 r2 : (zchinese r1 r2 = r1 %[mod m1])%Z.
+Proof.
+rewrite /zchinese; have [u v /= Duv _] := egcdzP m1 m2.
+rewrite -{2}[r1]mulr1 -((gcdz _ _ =P 1) co_m12) -Duv.
+by rewrite mulrDr mulrAC addrC (mulrAC r2) !mulrA !modzMDl.
+Qed.
+
+Lemma zchinese_modr r1 r2 : (zchinese r1 r2 = r2 %[mod m2])%Z.
+Proof.
+rewrite /zchinese; have [u v /= Duv _] := egcdzP m1 m2.
+rewrite -{2}[r2]mulr1 -((gcdz _ _ =P 1) co_m12) -Duv.
+by rewrite mulrAC modzMDl mulrAC addrC mulrDr !mulrA modzMDl.
+Qed.
+
+Lemma zchinese_mod x : (x = zchinese (x %% m1)%Z (x %% m2)%Z %[mod m1 * m2])%Z.
+Proof.
+apply/eqP; rewrite zchinese_remainder //.
+by rewrite zchinese_modl zchinese_modr !modz_mod !eqxx.
+Qed.
+
+End Chinese.
+
+Section ZpolyScale.
+
+Definition zcontents p :=
+ sgz (lead_coef p) * \big[gcdn/0%N]_(i < size p) `|(p`_i)%R|%N.
+
+Lemma sgz_contents p : sgz (zcontents p) = sgz (lead_coef p).
+Proof.
+rewrite /zcontents mulrC sgzM sgz_id; set d := _%:Z.
+have [-> | nz_p] := eqVneq p 0; first by rewrite lead_coef0 mulr0.
+rewrite gtr0_sgz ?mul1r // ltz_nat polySpred ?big_ord_recr //= -lead_coefE.
+by rewrite gcdn_gt0 orbC absz_gt0 lead_coef_eq0 nz_p.
+Qed.
+
+Lemma zcontents_eq0 p : (zcontents p == 0) = (p == 0).
+Proof. by rewrite -sgz_eq0 sgz_contents sgz_eq0 lead_coef_eq0. Qed.
+
+Lemma zcontents0 : zcontents 0 = 0.
+Proof. by apply/eqP; rewrite zcontents_eq0. Qed.
+
+Lemma zcontentsZ a p : zcontents (a *: p) = a * zcontents p.
+Proof.
+have [-> | nz_a] := eqVneq a 0; first by rewrite scale0r mul0r zcontents0.
+rewrite {2}[a]intEsg mulrCA -mulrA -PoszM big_distrr /= mulrCA mulrA -sgzM.
+rewrite -lead_coefZ; congr (_ * _%:Z); rewrite size_scale //.
+by apply: eq_bigr => i _; rewrite coefZ abszM.
+Qed.
+
+Lemma zcontents_monic p : p \is monic -> zcontents p = 1.
+Proof.
+move=> mon_p; rewrite /zcontents polySpred ?monic_neq0 //.
+by rewrite big_ord_recr /= -lead_coefE (monicP mon_p) gcdn1.
+Qed.
+
+Lemma dvdz_contents a p : (a %| zcontents p)%Z = (p \is a polyOver (dvdz a)).
+Proof.
+rewrite dvdzE abszM absz_sg lead_coef_eq0.
+have [-> | nz_p] := altP eqP; first by rewrite mul0n dvdn0 rpred0.
+rewrite mul1n; apply/dvdn_biggcdP/(all_nthP 0)=> a_dv_p i ltip /=.
+ exact: (a_dv_p (Ordinal ltip)).
+exact: a_dv_p.
+Qed.
+
+Lemma map_poly_divzK a p :
+ p \is a polyOver (dvdz a) -> a *: map_poly (divz^~ a) p = p.
+Proof.
+move/polyOverP=> a_dv_p; apply/polyP=> i.
+by rewrite coefZ coef_map_id0 ?div0z // mulrC divzK.
+Qed.
+
+Lemma polyOver_dvdzP a p :
+ reflect (exists q, p = a *: q) (p \is a polyOver (dvdz a)).
+Proof.
+apply: (iffP idP) => [/map_poly_divzK | [q ->]].
+ by exists (map_poly (divz^~ a) p).
+by apply/polyOverP=> i; rewrite coefZ dvdz_mulr.
+Qed.
+
+Definition zprimitive p := map_poly (divz^~ (zcontents p)) p.
+
+Lemma zpolyEprim p : p = zcontents p *: zprimitive p.
+Proof. by rewrite map_poly_divzK // -dvdz_contents. Qed.
+
+Lemma zprimitive0 : zprimitive 0 = 0.
+Proof.
+by apply/polyP=> i; rewrite coef0 coef_map_id0 ?div0z // zcontents0 divz0.
+Qed.
+
+Lemma zprimitive_eq0 p : (zprimitive p == 0) = (p == 0).
+Proof.
+apply/idP/idP=> /eqP p0; first by rewrite [p]zpolyEprim p0 scaler0.
+by rewrite p0 zprimitive0.
+Qed.
+
+Lemma size_zprimitive p : size (zprimitive p) = size p.
+Proof.
+have [-> | ] := eqVneq p 0; first by rewrite zprimitive0.
+by rewrite {1 3}[p]zpolyEprim scale_poly_eq0 => /norP[/size_scale-> _].
+Qed.
+
+Lemma sgz_lead_primitive p : sgz (lead_coef (zprimitive p)) = (p != 0).
+Proof.
+have [-> | nz_p] := altP eqP; first by rewrite zprimitive0 lead_coef0.
+apply: (@mulfI _ (sgz (zcontents p))); first by rewrite sgz_eq0 zcontents_eq0.
+by rewrite -sgzM mulr1 -lead_coefZ -zpolyEprim sgz_contents.
+Qed.
+
+Lemma zcontents_primitive p : zcontents (zprimitive p) = (p != 0).
+Proof.
+have [-> | nz_p] := altP eqP; first by rewrite zprimitive0 zcontents0.
+apply: (@mulfI _ (zcontents p)); first by rewrite zcontents_eq0.
+by rewrite mulr1 -zcontentsZ -zpolyEprim.
+Qed.
+
+Lemma zprimitive_id p : zprimitive (zprimitive p) = zprimitive p.
+Proof.
+have [-> | nz_p] := eqVneq p 0; first by rewrite !zprimitive0.
+by rewrite {2}[zprimitive p]zpolyEprim zcontents_primitive nz_p scale1r.
+Qed.
+
+Lemma zprimitive_monic p : p \in monic -> zprimitive p = p.
+Proof. by move=> mon_p; rewrite {2}[p]zpolyEprim zcontents_monic ?scale1r. Qed.
+
+Lemma zprimitiveZ a p : a != 0 -> zprimitive (a *: p) = zprimitive p.
+Proof.
+have [-> | nz_p nz_a] := eqVneq p 0; first by rewrite scaler0.
+apply: (@mulfI _ (a * zcontents p)%:P).
+ by rewrite polyC_eq0 mulf_neq0 ?zcontents_eq0.
+by rewrite -{1}zcontentsZ !mul_polyC -zpolyEprim -scalerA -zpolyEprim.
+Qed.
+
+Lemma zprimitive_min p a q :
+ p != 0 -> p = a *: q ->
+ {b | sgz b = sgz (lead_coef q) & q = b *: zprimitive p}.
+Proof.
+move=> nz_p Dp; have /dvdzP/sig_eqW[b Db]: (a %| zcontents p)%Z.
+ by rewrite dvdz_contents; apply/polyOver_dvdzP; exists q.
+suffices ->: q = b *: zprimitive p.
+ by rewrite lead_coefZ sgzM sgz_lead_primitive nz_p mulr1; exists b.
+apply: (@mulfI _ a%:P).
+ by apply: contraNneq nz_p; rewrite Dp -mul_polyC => ->; rewrite mul0r.
+by rewrite !mul_polyC -Dp scalerA mulrC -Db -zpolyEprim.
+Qed.
+
+Lemma zprimitive_irr p a q :
+ p != 0 -> zprimitive p = a *: q -> a = sgz (lead_coef q).
+Proof.
+move=> nz_p Dp; have: p = (a * zcontents p) *: q.
+ by rewrite mulrC -scalerA -Dp -zpolyEprim.
+case/zprimitive_min=> // b <- /eqP.
+rewrite Dp -{1}[q]scale1r scalerA -subr_eq0 -scalerBl scale_poly_eq0 subr_eq0.
+have{Dp} /negPf->: q != 0.
+ by apply: contraNneq nz_p; rewrite -zprimitive_eq0 Dp => ->; rewrite scaler0.
+by case: b a => [[|[|b]] | [|b]] [[|[|a]] | [|a]] //; rewrite mulr0.
+Qed.
+
+Lemma zcontentsM p q : zcontents (p * q) = zcontents p * zcontents q.
+Proof.
+have [-> | nz_p] := eqVneq p 0; first by rewrite !(mul0r, zcontents0).
+have [-> | nz_q] := eqVneq q 0; first by rewrite !(mulr0, zcontents0).
+rewrite -[zcontents q]mulr1 {1}[p]zpolyEprim {1}[q]zpolyEprim.
+rewrite -scalerAl -scalerAr !zcontentsZ; congr (_ * (_ * _)).
+rewrite [zcontents _]intEsg sgz_contents lead_coefM sgzM !sgz_lead_primitive.
+apply/eqP; rewrite nz_p nz_q !mul1r [_ == _]eqn_leq absz_gt0 zcontents_eq0.
+rewrite mulf_neq0 ?zprimitive_eq0 // andbT leqNgt.
+apply/negP=> /pdivP[r r_pr r_dv_d]; pose to_r : int -> 'F_r := intr.
+have nz_prim_r q1: q1 != 0 -> map_poly to_r (zprimitive q1) != 0.
+ move=> nz_q1; apply: contraTneq (prime_gt1 r_pr) => r_dv_q1.
+ rewrite -leqNgt dvdn_leq // -(dvdzE r true) -nz_q1 -zcontents_primitive.
+ rewrite dvdz_contents; apply/polyOverP=> i /=; rewrite dvdzE /=.
+ have /polyP/(_ i)/eqP := r_dv_q1; rewrite coef_map coef0 /=.
+ rewrite {1}[_`_i]intEsign rmorphM rmorph_sign /= mulf_eq0 signr_eq0 /=.
+ by rewrite -val_eqE /= val_Fp_nat.
+suffices{nz_prim_r} /idPn[]: map_poly to_r (zprimitive p * zprimitive q) == 0.
+ by rewrite rmorphM mulf_neq0 ?nz_prim_r.
+rewrite [_ * _]zpolyEprim [zcontents _]intEsign mulrC -scalerA map_polyZ /=.
+by rewrite scale_poly_eq0 -val_eqE /= val_Fp_nat ?(eqnP r_dv_d).
+Qed.
+
+
+Lemma zprimitiveM p q : zprimitive (p * q) = zprimitive p * zprimitive q.
+Proof.
+have [pq_0|] := eqVneq (p * q) 0.
+ rewrite pq_0; move/eqP: pq_0; rewrite mulf_eq0.
+ by case/pred2P=> ->; rewrite !zprimitive0 (mul0r, mulr0).
+rewrite -zcontents_eq0 -polyC_eq0 => /mulfI; apply; rewrite !mul_polyC.
+by rewrite -zpolyEprim zcontentsM -scalerA scalerAr scalerAl -!zpolyEprim.
+Qed.
+
+Lemma dvdpP_int p q : p %| q -> {r | q = zprimitive p * r}.
+Proof.
+case/Pdiv.Idomain.dvdpP/sig2_eqW=> [[c r] /= nz_c Dpr].
+exists (zcontents q *: zprimitive r); rewrite -scalerAr.
+by rewrite -zprimitiveM mulrC -Dpr zprimitiveZ // -zpolyEprim.
+Qed.
+
+Local Notation pZtoQ := (map_poly (intr : int -> rat)).
+
+Lemma size_rat_int_poly p : size (pZtoQ p) = size p.
+Proof. by apply: size_map_inj_poly; first exact: intr_inj. Qed.
+
+Lemma rat_poly_scale (p : {poly rat}) :
+ {q : {poly int} & {a | a != 0 & p = a%:~R^-1 *: pZtoQ q}}.
+Proof.
+pose a := \prod_(i < size p) denq p`_i.
+have nz_a: a != 0 by apply/prodf_neq0=> i _; exact: denq_neq0.
+exists (map_poly numq (a%:~R *: p)), a => //.
+apply: canRL (scalerK _) _; rewrite ?intr_eq0 //.
+apply/polyP=> i; rewrite !(coefZ, coef_map_id0) // numqK // Qint_def mulrC.
+have [ltip | /(nth_default 0)->] := ltnP i (size p); last by rewrite mul0r.
+by rewrite [a](bigD1 (Ordinal ltip)) // rmorphM mulrA -numqE -rmorphM denq_int.
+Qed.
+
+Lemma dvdp_rat_int p q : (pZtoQ p %| pZtoQ q) = (p %| q).
+Proof.
+apply/dvdpP/Pdiv.Idomain.dvdpP=> [[/= r1 Dq] | [[/= a r] nz_a Dq]]; last first.
+ exists (a%:~R^-1 *: pZtoQ r); rewrite -scalerAl -rmorphM -Dq.
+ by rewrite -{2}[a]intz scaler_int rmorphMz -scaler_int scalerK ?intr_eq0.
+have [r [a nz_a Dr1]] := rat_poly_scale r1; exists (a, r) => //=.
+apply: (map_inj_poly _ _ : injective pZtoQ) => //; first exact: intr_inj.
+rewrite -[a]intz scaler_int rmorphMz -scaler_int /= Dq Dr1.
+by rewrite -scalerAl -rmorphM scalerKV ?intr_eq0.
+Qed.
+
+Lemma dvdpP_rat_int p q :
+ p %| pZtoQ q ->
+ {p1 : {poly int} & {a | a != 0 & p = a *: pZtoQ p1} & {r | q = p1 * r}}.
+Proof.
+have{p} [p [a nz_a ->]] := rat_poly_scale p.
+rewrite dvdp_scalel ?invr_eq0 ?intr_eq0 // dvdp_rat_int => dv_p_q.
+exists (zprimitive p); last exact: dvdpP_int.
+have [-> | nz_p] := eqVneq p 0.
+ by exists 1; rewrite ?oner_eq0 // zprimitive0 map_poly0 !scaler0.
+exists ((zcontents p)%:~R / a%:~R).
+ by rewrite mulf_neq0 ?invr_eq0 ?intr_eq0 ?zcontents_eq0.
+by rewrite mulrC -scalerA -map_polyZ -zpolyEprim.
+Qed.
+
+End ZpolyScale.
+
+(* Integral spans. *)
+
+Lemma int_Smith_normal_form m n (M : 'M[int]_(m, n)) :
+ {L : 'M[int]_m & L \in unitmx &
+ {R : 'M[int]_n & R \in unitmx &
+ {d : seq int | sorted dvdz d &
+ M = L *m (\matrix_(i, j) (d`_i *+ (i == j :> nat))) *m R}}}.
+Proof.
+move: {2}_.+1 (ltnSn (m + n)) => mn.
+elim: mn => // mn IHmn in m n M *; rewrite ltnS => le_mn.
+have [[i j] nzMij | no_ij] := pickP (fun k => M k.1 k.2 != 0%N); last first.
+ do 2![exists 1%:M; first exact: unitmx1]; exists nil => //=.
+ apply/matrixP=> i j; apply/eqP; rewrite mulmx1 mul1mx mxE nth_nil mul0rn.
+ exact: negbFE (no_ij (i, j)).
+do [case: m i => [[]//|m] i; case: n j => [[]//|n] j /=] in M nzMij le_mn *.
+wlog Dj: j M nzMij / j = 0; last rewrite {j}Dj in nzMij.
+ case/(_ 0 (xcol j 0 M)); rewrite ?mxE ?tpermR // => L uL [R uR [d dvD dM]].
+ exists L => //; exists (xcol j 0 R); last exists d => //=.
+ by rewrite xcolE unitmx_mul uR unitmx_perm.
+ by rewrite xcolE !mulmxA -dM xcolE -mulmxA -perm_mxM tperm2 perm_mx1 mulmx1.
+move Da: (M i 0) nzMij => a nz_a.
+elim: {a}_.+1 {-2}a (ltnSn `|a|) => // A IHa a leA in m n M i Da nz_a le_mn *.
+wlog [j a'Mij]: m n M i Da le_mn / {j | ~~ (a %| M i j)%Z}; last first.
+ have nz_j: j != 0 by apply: contraNneq a'Mij => ->; rewrite Da.
+ case: n => [[[]//]|n] in j le_mn nz_j M a'Mij Da *.
+ wlog{nz_j} Dj: j M a'Mij Da / j = 1; last rewrite {j}Dj in a'Mij.
+ case/(_ 1 (xcol j 1 M)); rewrite ?mxE ?tpermR ?tpermD //.
+ move=> L uL [R uR [d dvD dM]]; exists L => //.
+ exists (xcol j 1 R); first by rewrite xcolE unitmx_mul uR unitmx_perm.
+ exists d; rewrite //= xcolE !mulmxA -dM xcolE -mulmxA -perm_mxM tperm2.
+ by rewrite perm_mx1 mulmx1.
+ have [u [v]] := Bezoutz a (M i 1); set b := gcdz _ _ => Db.
+ have{leA} ltA: (`|b| < A)%N.
+ rewrite -ltnS (leq_trans _ leA) // ltnS ltn_neqAle andbC.
+ rewrite dvdn_leq ?absz_gt0 ? dvdn_gcdl //=.
+ by rewrite (contraNneq _ a'Mij) ?dvdzE // => <-; exact: dvdn_gcdr.
+ pose t2 := [fun j : 'I_2 => [tuple _; _]`_j : int]; pose a1 := M i 1.
+ pose Uul := \matrix_(k, j) t2 (t2 u (- (a1 %/ b)%Z) j) (t2 v (a %/ b)%Z j) k.
+ pose U : 'M_(2 + n) := block_mx Uul 0 0 1%:M; pose M1 := M *m U.
+ have{nz_a} nz_b: b != 0 by rewrite gcdz_eq0 (negPf nz_a).
+ have uU: U \in unitmx.
+ rewrite unitmxE det_ublock det1 (expand_det_col _ 0) big_ord_recl big_ord1.
+ do 2!rewrite /cofactor [row' _ _]mx11_scalar !mxE det_scalar1 /=.
+ rewrite mulr1 mul1r mulN1r opprK -[_ + _](mulzK _ nz_b) mulrDl.
+ by rewrite -!mulrA !divzK ?dvdz_gcdl ?dvdz_gcdr // Db divzz nz_b unitr1.
+ have{Db} Db: M1 i 0 = b.
+ rewrite /M1 -(lshift0 n 1) [U]block_mxEh mul_mx_row row_mxEl.
+ rewrite -[M](@hsubmxK _ _ 2) (@mul_row_col _ _ 2) mulmx0 addr0 !mxE /=.
+ rewrite big_ord_recl big_ord1 !mxE /= [lshift _ _]((_ =P 0) _) // Da.
+ by rewrite [lshift _ _]((_ =P 1) _) // mulrC -(mulrC v).
+ have [L uL [R uR [d dvD dM1]]] := IHa b ltA _ _ M1 i Db nz_b le_mn.
+ exists L => //; exists (R *m invmx U); last exists d => //.
+ by rewrite unitmx_mul uR unitmx_inv.
+ by rewrite mulmxA -dM1 mulmxK.
+move=> {A leA IHa} IHa; wlog Di: i M Da / i = 0; last rewrite {i}Di in Da.
+ case/(_ 0 (xrow i 0 M)); rewrite ?mxE ?tpermR // => L uL [R uR [d dvD dM]].
+ exists (xrow i 0 L); first by rewrite xrowE unitmx_mul unitmx_perm.
+ exists R => //; exists d; rewrite //= xrowE -!mulmxA (mulmxA L) -dM xrowE.
+ by rewrite mulmxA -perm_mxM tperm2 perm_mx1 mul1mx.
+without loss /forallP a_dvM0: / [forall j, a %| M 0 j]%Z.
+ have [_|] := altP forallP; first exact; rewrite negb_forall => /existsP/sigW.
+ by move/IHa=> IH _; apply: IH.
+without loss{Da a_dvM0} Da: M / forall j, M 0 j = a.
+ pose Uur := col' 0 (\row_j (1 - (M 0 j %/ a)%Z)).
+ pose U : 'M_(1 + n) := block_mx 1 Uur 0 1%:M; pose M1 := M *m U.
+ have uU: U \in unitmx by rewrite unitmxE det_ublock !det1 mulr1.
+ case/(_ (M *m U)) => [j | L uL [R uR [d dvD dM]]].
+ rewrite -(lshift0 m 0) -[M](@submxK _ 1 _ 1) (@mulmx_block _ 1 m 1).
+ rewrite (@col_mxEu _ 1) !mulmx1 mulmx0 addr0 [ulsubmx _]mx11_scalar.
+ rewrite mul_scalar_mx !mxE !lshift0 Da.
+ case: splitP => [j0 _ | j1 Dj]; rewrite ?ord1 !mxE // lshift0 rshift1.
+ by rewrite mulrBr mulr1 mulrC divzK ?subrK.
+ exists L => //; exists (R * U^-1); first by rewrite unitmx_mul uR unitmx_inv.
+ by exists d; rewrite //= mulmxA -dM mulmxK.
+without loss{IHa} /forallP/(_ (_, _))/= a_dvM: / [forall k, a %| M k.1 k.2]%Z.
+ have [_|] := altP forallP; first exact; rewrite negb_forall => /existsP/sigW.
+ case=> [[i j] /= a'Mij] _.
+ have [|||L uL [R uR [d dvD dM]]] := IHa _ _ M^T j; rewrite ?mxE 1?addnC //.
+ by exists i; rewrite mxE.
+ exists R^T; last exists L^T; rewrite ?unitmx_tr //; exists d => //.
+ rewrite -[M]trmxK dM !trmx_mul mulmxA; congr (_ *m _ *m _).
+ by apply/matrixP=> i1 j1; rewrite !mxE eq_sym; case: eqP => // ->.
+without loss{nz_a a_dvM} a1: M a Da / a = 1.
+ pose M1 := map_mx (divz^~ a) M; case/(_ M1 1)=> // [k|L uL [R uR [d dvD dM]]].
+ by rewrite !mxE Da divzz nz_a.
+ exists L => //; exists R => //; exists [seq a * x | x <- d].
+ case: d dvD {dM} => //= x d; elim: d x => //= y d IHd x /andP[dv_xy /IHd].
+ by rewrite [dvdz _ _]dvdz_mul2l ?[_ \in _]dv_xy.
+ have ->: M = a *: M1 by apply/matrixP=> i j; rewrite !mxE mulrC divzK ?a_dvM.
+ rewrite dM scalemxAl scalemxAr; congr (_ *m _ *m _).
+ apply/matrixP=> i j; rewrite !mxE mulrnAr; congr (_ *+ _).
+ have [lt_i_d | le_d_i] := ltnP i (size d); first by rewrite (nth_map 0).
+ by rewrite !nth_default ?size_map ?mulr0.
+rewrite {a}a1 -[m.+1]/(1 + m)%N -[n.+1]/(1 + n)%N in M Da *.
+pose Mu := ursubmx M; pose Ml := dlsubmx M.
+have{Da} Da: ulsubmx M = 1 by rewrite [_ M]mx11_scalar !mxE !lshift0 Da.
+pose M1 := - (Ml *m Mu) + drsubmx M.
+have [|L uL [R uR [d dvD dM1]]] := IHmn m n M1; first by rewrite -addnS ltnW.
+exists (block_mx 1 0 Ml L).
+ by rewrite unitmxE det_lblock det_scalar1 mul1r.
+exists (block_mx 1 Mu 0 R).
+ by rewrite unitmxE det_ublock det_scalar1 mul1r.
+exists (1 :: d); set D1 := \matrix_(i, j) _ in dM1.
+ by rewrite /= path_min_sorted // => g _; exact: dvd1n.
+rewrite [D in _ *m D *m _](_ : _ = block_mx 1 0 0 D1); last first.
+ by apply/matrixP=> i j; do 3?[rewrite ?mxE ?ord1 //=; case: splitP => ? ->].
+rewrite !mulmx_block !(mul0mx, mulmx0, addr0) !mulmx1 add0r mul1mx -Da -dM1.
+by rewrite addNKr submxK.
+Qed.
+
+Definition inIntSpan (V : zmodType) m (s : m.-tuple V) v :=
+ exists a : int ^ m, v = \sum_(i < m) s`_i *~ a i.
+
+Lemma dec_Qint_span (vT : vectType rat) m (s : m.-tuple vT) v :
+ decidable (inIntSpan s v).
+Proof.
+have s_s (i : 'I_m): s`_i \in <<s>>%VS by rewrite memv_span ?memt_nth.
+have s_Zs a: \sum_(i < m) s`_i *~ a i \in <<s>>%VS.
+ by rewrite memv_suml // => i _; rewrite -scaler_int memvZ.
+case s_v: (v \in <<s>>%VS); last by right=> [[a Dv]]; rewrite Dv s_Zs in s_v.
+pose S := \matrix_(i < m, j < _) coord (vbasis <<s>>) j s`_i.
+pose r := \rank S; pose k := (m - r)%N; pose Em := erefl m; pose Ek := erefl k.
+have Dm: (m = k + r)%N by rewrite subnK ?rank_leq_row.
+have [K kerK]: {K : 'M_(k, m) | map_mx intr K == kermx S}%MS.
+ pose B := row_base (kermx S); pose d := \prod_ij denq (B ij.1 ij.2).
+ exists (castmx (mxrank_ker S, Em) (map_mx numq (intr d *: B))).
+ rewrite /k; case: _ / (mxrank_ker S); set B1 := map_mx _ _.
+ have ->: B1 = (intr d *: B).
+ apply/matrixP=> i j; rewrite 3!mxE mulrC [d](bigD1 (i, j)) // rmorphM mulrA.
+ by rewrite -numqE -rmorphM numq_int.
+ suffices nz_d: d%:Q != 0 by rewrite !eqmx_scale // !eq_row_base andbb.
+ by rewrite intr_eq0; apply/prodf_neq0 => i _; exact: denq_neq0.
+have [L _ [G uG [D _ defK]]] := int_Smith_normal_form K.
+pose Gud := castmx (Dm, Em) G; pose G'lr := castmx (Em, Dm) (invmx G).
+have{K L D defK kerK} kerGu: map_mx intr (usubmx Gud) *m S = 0.
+ pose Kl : 'M[rat]_k:= map_mx intr (lsubmx (castmx (Ek, Dm) (K *m invmx G))).
+ have{defK} defK: map_mx intr K = row_mx Kl 0 *m map_mx intr Gud.
+ rewrite -[K](mulmxKV uG) -{2}[G](castmxK Dm Em) -/Gud.
+ rewrite -[K *m _](castmxK Ek Dm) map_mxM map_castmx.
+ rewrite -(hsubmxK (castmx _ _)) map_row_mx -/Kl map_castmx /Em.
+ set Kr := map_mx _ _; case: _ / (esym Dm) (map_mx _ _) => /= GudQ.
+ congr (row_mx _ _ *m _); apply/matrixP=> i j; rewrite !mxE defK mulmxK //=.
+ rewrite castmxE mxE big1 //= => j1 _; rewrite mxE /= eqn_leq andbC.
+ by rewrite leqNgt (leq_trans (valP j1)) ?mulr0 ?leq_addr.
+ have /row_full_inj: row_full Kl; last apply.
+ rewrite /row_full eqn_leq rank_leq_row /= -{1}[k](mxrank_ker S).
+ rewrite -(eqmxP kerK) defK map_castmx mxrankMfree; last first.
+ case: _ / (Dm); apply/row_freeP; exists (map_mx intr (invmx G)).
+ by rewrite -map_mxM mulmxV ?map_mx1.
+ by rewrite -mxrank_tr tr_row_mx trmx0 -addsmxE addsmx0 mxrank_tr.
+ rewrite mulmx0 mulmxA (sub_kermxP _) // -(eqmxP kerK) defK.
+ by rewrite -{2}[Gud]vsubmxK map_col_mx mul_row_col mul0mx addr0.
+pose T := map_mx intr (dsubmx Gud) *m S.
+have{kerGu} defS: map_mx intr (rsubmx G'lr) *m T = S.
+ have: G'lr *m Gud = 1%:M by rewrite /G'lr /Gud; case: _ / (Dm); exact: mulVmx.
+ rewrite -{1}[G'lr]hsubmxK -[Gud]vsubmxK mulmxA mul_row_col -map_mxM.
+ move/(canRL (addKr _))->; rewrite -mulNmx raddfD /= map_mx1 map_mxM /=.
+ by rewrite mulmxDl -mulmxA kerGu mulmx0 add0r mul1mx.
+pose vv := \row_j coord (vbasis <<s>>) j v.
+have uS: row_full S.
+ apply/row_fullP; exists (\matrix_(i, j) coord s j (vbasis <<s>>)`_i).
+ apply/matrixP=> j1 j2; rewrite !mxE.
+ rewrite -(coord_free _ _ (basis_free (vbasisP _))).
+ rewrite -!tnth_nth (coord_span (vbasis_mem (mem_tnth j1 _))) linear_sum.
+ by apply: eq_bigr => i _; rewrite !mxE (tnth_nth 0) !linearZ.
+have eqST: (S :=: T)%MS by apply/eqmxP; rewrite -{1}defS !submxMl.
+case Zv: (map_mx denq (vv *m pinvmx T) == const_mx 1).
+ pose a := map_mx numq (vv *m pinvmx T) *m dsubmx Gud.
+ left; exists [ffun j => a 0 j].
+ transitivity (\sum_j (map_mx intr a *m S) 0 j *: (vbasis <<s>>)`_j).
+ rewrite {1}(coord_vbasis s_v); apply: eq_bigr => j _; congr (_ *: _).
+ have ->: map_mx intr a = vv *m pinvmx T *m map_mx intr (dsubmx Gud).
+ rewrite map_mxM /=; congr (_ *m _); apply/rowP=> i; rewrite 2!mxE numqE.
+ by have /eqP/rowP/(_ i) := Zv; rewrite !mxE => ->; rewrite mulr1.
+ by rewrite -(mulmxA _ _ S) mulmxKpV ?mxE // -eqST submx_full.
+ rewrite (coord_vbasis (s_Zs _)); apply: eq_bigr => j _; congr (_ *: _).
+ rewrite linear_sum mxE; apply: eq_bigr => i _.
+ by rewrite -scaler_int linearZ [a]lock !mxE ffunE.
+right=> [[a Dv]]; case/eqP: Zv; apply/rowP.
+have ->: vv = map_mx intr (\row_i a i) *m S.
+ apply/rowP=> j; rewrite !mxE Dv linear_sum.
+ by apply: eq_bigr => i _; rewrite -scaler_int linearZ !mxE.
+rewrite -defS -2!mulmxA; have ->: T *m pinvmx T = 1%:M.
+ have uT: row_free T by rewrite /row_free -eqST.
+ by apply: (row_free_inj uT); rewrite mul1mx mulmxKpV.
+by move=> i; rewrite mulmx1 -map_mxM 2!mxE denq_int mxE.
+Qed.
+
diff --git a/mathcomp/algebra/interval.v b/mathcomp/algebra/interval.v
new file mode 100644
index 0000000..1f59ce2
--- /dev/null
+++ b/mathcomp/algebra/interval.v
@@ -0,0 +1,403 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div choice fintype.
+Require Import bigop ssralg finset fingroup zmodp ssrint ssrnum.
+
+(*****************************************************************************)
+(* This file provide support for intervals in numerical and real domains. *)
+(* The datatype (interval R) gives a formal characterization of an *)
+(* interval, as the pair of its right and left bounds. *)
+(* interval R == the type of formal intervals on R. *)
+(* x \in i == when i is a formal interval on a numeric domain, *)
+(* \in can be used to test membership. *)
+(* itvP x_in_i == where x_in_i has type x \in i, if i is ground, *)
+(* gives a set of rewrite rules that x_in_i imply. *)
+(* x <= y ?< if c == x is smaller than y, and strictly if c is true *)
+(* *)
+(* We provide a set of notations to write intervals (see below) *)
+(* `[a, b], `]a, b], ..., `]-oo, a], ..., `]-oo, +oo[ *)
+(* We also provide the lemma subitvP which computes the inequalities one *)
+(* needs to prove when trying to prove the inclusion of intervals. *)
+(* *)
+(* Remark that we cannot implement a boolean comparison test for intervals *)
+(* on an arbitrary numeric domains, for this problem might be undecidable. *)
+(* Note also that type (interval R) may contain several inhabitants coding *)
+(* for the same interval. However, this pathological issues do nor arise *)
+(* when R is a real domain: we could provide a specific theory for this *)
+(* important case. *)
+(* *)
+(* See also "Formal proofs in real algebraic geometry: from ordered fields *)
+(* to quantifier elimination", LMCS journal, 2012 *)
+(* by Cyril Cohen and Assia Mahboubi *)
+(* *)
+(* And "Formalized algebraic numbers: construction and first-order theory" *)
+(* Cyril Cohen, PhD, 2012, section 4.3. *)
+(*****************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Local Open Scope ring_scope.
+Import GRing.Theory Num.Theory.
+
+Local Notation mid x y := ((x + y) / 2%:R).
+
+Section IntervalPo.
+
+CoInductive itv_bound (T : Type) : Type := BOpen_if of bool & T | BInfty.
+Notation BOpen := (BOpen_if true).
+Notation BClose := (BOpen_if false).
+CoInductive interval (T : Type) := Interval of itv_bound T & itv_bound T.
+
+Variable (R : numDomainType).
+
+Definition pred_of_itv (i : interval R) : pred R :=
+ [pred x | let: Interval l u := i in
+ match l with
+ | BOpen a => a < x
+ | BClose a => a <= x
+ | BInfty => true
+ end &&
+ match u with
+ | BOpen b => x < b
+ | BClose b => x <= b
+ | BInfty => true
+ end].
+Canonical Structure itvPredType := Eval hnf in mkPredType pred_of_itv.
+
+(* We provide the 9 following notations to help writing formal intervals *)
+Notation "`[ a , b ]" := (Interval (BClose a) (BClose b))
+ (at level 0, a, b at level 9 , format "`[ a , b ]") : ring_scope.
+Notation "`] a , b ]" := (Interval (BOpen a) (BClose b))
+ (at level 0, a, b at level 9 , format "`] a , b ]") : ring_scope.
+Notation "`[ a , b [" := (Interval (BClose a) (BOpen b))
+ (at level 0, a, b at level 9 , format "`[ a , b [") : ring_scope.
+Notation "`] a , b [" := (Interval (BOpen a) (BOpen b))
+ (at level 0, a, b at level 9 , format "`] a , b [") : ring_scope.
+Notation "`] '-oo' , b ]" := (Interval (BInfty _) (BClose b))
+ (at level 0, b at level 9 , format "`] '-oo' , b ]") : ring_scope.
+Notation "`] '-oo' , b [" := (Interval (BInfty _) (BOpen b))
+ (at level 0, b at level 9 , format "`] '-oo' , b [") : ring_scope.
+Notation "`[ a , '+oo' [" := (Interval (BClose a) (BInfty _))
+ (at level 0, a at level 9 , format "`[ a , '+oo' [") : ring_scope.
+Notation "`] a , '+oo' [" := (Interval (BOpen a) (BInfty _))
+ (at level 0, a at level 9 , format "`] a , '+oo' [") : ring_scope.
+Notation "`] -oo , '+oo' [" := (Interval (BInfty _) (BInfty _))
+ (at level 0, format "`] -oo , '+oo' [") : ring_scope.
+
+(* we compute a set of rewrite rules associated to an interval *)
+Definition itv_rewrite (i : interval R) x : Type :=
+ let: Interval l u := i in
+ (match l with
+ | BClose a => (a <= x) * (x < a = false)
+ | BOpen a => (a <= x) * (a < x) * (x <= a = false)
+ | BInfty => forall x : R, x == x
+ end *
+ match u with
+ | BClose b => (x <= b) * (b < x = false)
+ | BOpen b => (x <= b) * (x < b) * (b <= x = false)
+ | BInfty => forall x : R, x == x
+ end *
+ match l, u with
+ | BClose a, BClose b =>
+ (a <= b) * (b < a = false) * (a \in `[a, b]) * (b \in `[a, b])
+ | BClose a, BOpen b =>
+ (a <= b) * (a < b) * (b <= a = false)
+ * (a \in `[a, b]) * (a \in `[a, b[)* (b \in `[a, b]) * (b \in `]a, b])
+ | BOpen a, BClose b =>
+ (a <= b) * (a < b) * (b <= a = false)
+ * (a \in `[a, b]) * (a \in `[a, b[)* (b \in `[a, b]) * (b \in `]a, b])
+ | BOpen a, BOpen b =>
+ (a <= b) * (a < b) * (b <= a = false)
+ * (a \in `[a, b]) * (a \in `[a, b[)* (b \in `[a, b]) * (b \in `]a, b])
+ | _, _ => forall x : R, x == x
+ end)%type.
+
+Definition itv_decompose (i : interval R) x : Prop :=
+ let: Interval l u := i in
+ ((match l with
+ | BClose a => (a <= x) : Prop
+ | BOpen a => (a < x) : Prop
+ | BInfty => True
+ end : Prop) *
+ (match u with
+ | BClose b => (x <= b) : Prop
+ | BOpen b => (x < b) : Prop
+ | BInfty => True
+ end : Prop))%type.
+
+Lemma itv_dec : forall (x : R) (i : interval R),
+ reflect (itv_decompose i x) (x \in i).
+Proof. by move=> x [[[] a|] [[] b|]]; apply: (iffP andP); case. Qed.
+
+Implicit Arguments itv_dec [x i].
+
+Definition lersif (x y : R) b := if b then x < y else x <= y.
+
+Local Notation "x <= y ?< 'if' b" := (lersif x y b)
+ (at level 70, y at next level,
+ format "x '[hv' <= y '/' ?< 'if' b ']'") : ring_scope.
+
+Lemma lersifxx x b: (x <= x ?< if b) = ~~ b.
+Proof. by case: b; rewrite /= lterr. Qed.
+
+Lemma lersif_trans x y z b1 b2 :
+ x <= y ?< if b1 -> y <= z ?< if b2 -> x <= z ?< if b1 || b2.
+Proof.
+move: b1 b2 => [] [] //=;
+by [exact: ler_trans|exact: ler_lt_trans|exact: ltr_le_trans|exact: ltr_trans].
+Qed.
+
+Lemma lersifW b x y : x <= y ?< if b -> x <= y.
+Proof. by case: b => //; move/ltrW. Qed.
+
+Lemma lersifNF x y b : y <= x ?< if ~~ b -> x <= y ?< if b = false.
+Proof. by case: b => /= [/ler_gtF|/ltr_geF]. Qed.
+
+Lemma lersifS x y b : x < y -> x <= y ?< if b.
+Proof. by case: b => //= /ltrW. Qed.
+
+Lemma lersifT x y : x <= y ?< if true = (x < y). Proof. by []. Qed.
+
+Lemma lersifF x y : x <= y ?< if false = (x <= y). Proof. by []. Qed.
+
+Definition le_boundl b1 b2 :=
+ match b1, b2 with
+ | BOpen_if b1 x1, BOpen_if b2 x2 => x1 <= x2 ?< if (~~ b2 && b1)
+ | BOpen_if _ _, BInfty => false
+ | _, _ => true
+ end.
+
+Definition le_boundr b1 b2 :=
+ match b1, b2 with
+ | BOpen_if b1 x1, BOpen_if b2 x2 => x1 <= x2 ?< if (~~ b1 && b2)
+ | BInfty, BOpen_if _ _ => false
+ | _, _ => true
+ end.
+
+Lemma itv_boundlr bl br x :
+ (x \in Interval bl br) =
+ (le_boundl bl (BClose x)) && (le_boundr (BClose x) br).
+Proof. by move: bl br => [[] a|] [[] b|]. Qed.
+
+Lemma le_boundr_refl : reflexive le_boundr.
+Proof. by move=> [[] b|]; rewrite /le_boundr /= ?lerr. Qed.
+
+Hint Resolve le_boundr_refl.
+
+Lemma le_boundl_refl : reflexive le_boundl.
+Proof. by move=> [[] b|]; rewrite /le_boundl /= ?lerr. Qed.
+
+Hint Resolve le_boundl_refl.
+
+Lemma le_boundl_bb x b1 b2 :
+ le_boundl (BOpen_if b1 x) (BOpen_if b2 x) = (b1 ==> b2).
+Proof. by rewrite /le_boundl lersifxx andbC negb_and negbK implybE. Qed.
+
+Lemma le_boundr_bb x b1 b2 :
+ le_boundr (BOpen_if b1 x) (BOpen_if b2 x) = (b2 ==> b1).
+Proof. by rewrite /le_boundr lersifxx andbC negb_and negbK implybE. Qed.
+
+Lemma itv_xx x bl br :
+ Interval (BOpen_if bl x) (BOpen_if br x) =i
+ if ~~ (bl || br) then pred1 x else pred0.
+Proof.
+by move: bl br => [] [] y /=; rewrite !inE 1?eq_sym (eqr_le, lter_anti).
+Qed.
+
+Lemma itv_gte ba xa bb xb : xb <= xa ?< if ~~ (ba || bb)
+ -> Interval (BOpen_if ba xa) (BOpen_if bb xb) =i pred0.
+Proof.
+move=> hx y; rewrite itv_boundlr inE /=.
+by apply/negP => /andP [] /lersif_trans hy /hy {hy}; rewrite lersifNF.
+Qed.
+
+Lemma boundl_in_itv : forall ba xa b,
+ xa \in Interval (BOpen_if ba xa) b =
+ if ba then false else le_boundr (BClose xa) b.
+Proof. by move=> [] xa [[] xb|] //=; rewrite inE lterr. Qed.
+
+Lemma boundr_in_itv : forall bb xb a,
+ xb \in Interval a (BOpen_if bb xb) =
+ if bb then false else le_boundl a (BClose xb).
+Proof. by move=> [] xb [[] xa|] //=; rewrite inE lterr ?andbT ?andbF. Qed.
+
+Definition bound_in_itv := (boundl_in_itv, boundr_in_itv).
+
+Lemma itvP : forall (x : R) (i : interval R), (x \in i) -> itv_rewrite i x.
+Proof.
+move=> x [[[] a|] [[] b|]]; move/itv_dec=> //= [hl hu];do ?[split=> //;
+ do ?[by rewrite ltrW | by rewrite ltrWN | by rewrite ltrNW |
+ by rewrite (ltr_geF, ler_gtF)]];
+ rewrite ?(bound_in_itv) /le_boundl /le_boundr //=; do ?
+ [ by rewrite (@ler_trans _ x)
+ | by rewrite 1?ltrW // (@ltr_le_trans _ x)
+ | by rewrite 1?ltrW // (@ler_lt_trans _ x) // 1?ltrW
+ | by apply: negbTE; rewrite ler_gtF // (@ler_trans _ x)
+ | by apply: negbTE; rewrite ltr_geF // (@ltr_le_trans _ x) // 1?ltrW
+ | by apply: negbTE; rewrite ltr_geF // (@ler_lt_trans _ x)].
+Qed.
+
+Hint Rewrite intP.
+Implicit Arguments itvP [x i].
+
+Definition subitv (i1 i2 : interval R) :=
+ match i1, i2 with
+ | Interval a1 b1, Interval a2 b2 => le_boundl a2 a1 && le_boundr b1 b2
+ end.
+
+Lemma subitvP : forall (i2 i1 : interval R),
+ (subitv i1 i2) -> {subset i1 <= i2}.
+Proof.
+by move=> [[[] a2|] [[] b2|]] [[[] a1|] [[] b1|]];
+ rewrite /subitv //; case/andP=> /= ha hb; move=> x hx; rewrite ?inE;
+ rewrite ?(ler_trans ha) ?(ler_lt_trans ha) ?(ltr_le_trans ha) //;
+ rewrite ?(ler_trans _ hb) ?(ltr_le_trans _ hb) ?(ler_lt_trans _ hb) //;
+ rewrite ?(itvP hx) //.
+Qed.
+
+Lemma subitvPr : forall (a b1 b2 : itv_bound R),
+ le_boundr b1 b2 -> {subset (Interval a b1) <= (Interval a b2)}.
+Proof. by move=> a b1 b2 hb; apply: subitvP=> /=; rewrite hb andbT. Qed.
+
+Lemma subitvPl : forall (a1 a2 b : itv_bound R),
+ le_boundl a2 a1 -> {subset (Interval a1 b) <= (Interval a2 b)}.
+Proof. by move=> a1 a2 b ha; apply: subitvP=> /=; rewrite ha /=. Qed.
+
+Lemma lersif_in_itv : forall ba bb xa xb x,
+ x \in Interval (BOpen_if ba xa) (BOpen_if bb xb) ->
+ xa <= xb ?< if ba || bb.
+Proof.
+by move=> ba bb xa xb y; rewrite itv_boundlr; case/andP; apply: lersif_trans.
+Qed.
+
+Lemma ltr_in_itv : forall ba bb xa xb x, ba || bb ->
+ x \in Interval (BOpen_if ba xa) (BOpen_if bb xb) -> xa < xb.
+Proof.
+move=> ba bb xa xb x; case bab: (_ || _) => // _.
+by move/lersif_in_itv; rewrite bab.
+Qed.
+
+Lemma ler_in_itv : forall ba bb xa xb x,
+ x \in Interval (BOpen_if ba xa) (BOpen_if bb xb) -> xa <= xb.
+Proof. by move=> ba bb xa xb x; move/lersif_in_itv; move/lersifW. Qed.
+
+Lemma mem0_itvcc_xNx : forall x, (0 \in `[-x, x]) = (0 <= x).
+Proof.
+by move=> x; rewrite !inE; case hx: (0 <= _); rewrite (andbT, andbF) ?ge0_cp.
+Qed.
+
+Lemma mem0_itvoo_xNx : forall x, 0 \in `](-x), x[ = (0 < x).
+Proof.
+by move=> x; rewrite !inE; case hx: (0 < _); rewrite (andbT, andbF) ?gt0_cp.
+Qed.
+
+Lemma itv_splitI : forall a b, forall x,
+ x \in Interval a b = (x \in Interval a (BInfty _)) && (x \in Interval (BInfty _) b).
+Proof. by move=> [[] a|] [[] b|] x; rewrite ?inE ?andbT. Qed.
+
+
+Lemma real_lersifN x y b : x \in Num.real -> y \in Num.real ->
+ x <= y ?< if ~~b = ~~ (y <= x ?< if b).
+Proof. by case: b => [] xR yR /=; rewrite (real_ltrNge, real_lerNgt). Qed.
+
+Lemma oppr_itv ba bb (xa xb x : R) :
+ (-x \in Interval (BOpen_if ba xa) (BOpen_if bb xb)) =
+ (x \in Interval (BOpen_if bb (-xb)) (BOpen_if ba (-xa))).
+Proof. by move: ba bb => [] []; rewrite ?inE lter_oppr andbC lter_oppl. Qed.
+
+Lemma oppr_itvoo (a b x : R) : (-x \in `]a, b[) = (x \in `](-b), (-a)[).
+Proof. exact: oppr_itv. Qed.
+
+Lemma oppr_itvco (a b x : R) : (-x \in `[a, b[) = (x \in `](-b), (-a)]).
+Proof. exact: oppr_itv. Qed.
+
+Lemma oppr_itvoc (a b x : R) : (-x \in `]a, b]) = (x \in `[(-b), (-a)[).
+Proof. exact: oppr_itv. Qed.
+
+Lemma oppr_itvcc (a b x : R) : (-x \in `[a, b]) = (x \in `[(-b), (-a)]).
+Proof. exact: oppr_itv. Qed.
+
+End IntervalPo.
+
+Notation BOpen := (BOpen_if true).
+Notation BClose := (BOpen_if false).
+Notation "`[ a , b ]" := (Interval (BClose a) (BClose b))
+ (at level 0, a, b at level 9 , format "`[ a , b ]") : ring_scope.
+Notation "`] a , b ]" := (Interval (BOpen a) (BClose b))
+ (at level 0, a, b at level 9 , format "`] a , b ]") : ring_scope.
+Notation "`[ a , b [" := (Interval (BClose a) (BOpen b))
+ (at level 0, a, b at level 9 , format "`[ a , b [") : ring_scope.
+Notation "`] a , b [" := (Interval (BOpen a) (BOpen b))
+ (at level 0, a, b at level 9 , format "`] a , b [") : ring_scope.
+Notation "`] '-oo' , b ]" := (Interval (BInfty _) (BClose b))
+ (at level 0, b at level 9 , format "`] '-oo' , b ]") : ring_scope.
+Notation "`] '-oo' , b [" := (Interval (BInfty _) (BOpen b))
+ (at level 0, b at level 9 , format "`] '-oo' , b [") : ring_scope.
+Notation "`[ a , '+oo' [" := (Interval (BClose a) (BInfty _))
+ (at level 0, a at level 9 , format "`[ a , '+oo' [") : ring_scope.
+Notation "`] a , '+oo' [" := (Interval (BOpen a) (BInfty _))
+ (at level 0, a at level 9 , format "`] a , '+oo' [") : ring_scope.
+Notation "`] -oo , '+oo' [" := (Interval (BInfty _) (BInfty _))
+ (at level 0, format "`] -oo , '+oo' [") : ring_scope.
+
+Notation "x <= y ?< 'if' b" := (lersif x y b)
+ (at level 70, y at next level,
+ format "x '[hv' <= y '/' ?< 'if' b ']'") : ring_scope.
+
+Section IntervalOrdered.
+
+Variable R : realDomainType.
+
+Lemma lersifN (x y : R) b : (x <= y ?< if ~~ b) = ~~ (y <= x ?< if b).
+Proof. by rewrite real_lersifN ?num_real. Qed.
+
+Lemma itv_splitU (xc : R) bc a b : xc \in Interval a b ->
+ forall y, y \in Interval a b =
+ (y \in Interval a (BOpen_if (~~ bc) xc))
+ || (y \in Interval (BOpen_if bc xc) b).
+Proof.
+move=> hxc y; rewrite !itv_boundlr [le_boundr]lock /=.
+have [la /=|nla /=] := boolP (le_boundl a _); rewrite -lock.
+ have [lb /=|nlb /=] := boolP (le_boundr _ b); rewrite ?andbT ?andbF ?orbF //.
+ by case: bc => //=; case: ltrgtP.
+ symmetry; apply: contraNF nlb; rewrite /le_boundr /=.
+ case: b hxc => // bb xb hxc hyc.
+ suff /(lersif_trans hyc) : xc <= xb ?< if bb.
+ by case: bc {hyc} => //= /lersifS.
+ by case: a bb hxc {la} => [[] ?|] [] /= /itvP->.
+symmetry; apply: contraNF nla => /andP [hc _].
+case: a hxc hc => [[] xa|] hxc; rewrite /le_boundl //=.
+ by move=> /lersifW /(ltr_le_trans _) -> //; move: b hxc=> [[] ?|] /itvP->.
+by move=> /lersifW /(ler_trans _) -> //; move: b hxc=> [[] ?|] /itvP->.
+Qed.
+
+Lemma itv_splitU2 (x : R) a b : x \in Interval a b ->
+ forall y, y \in Interval a b =
+ [|| (y \in Interval a (BOpen x)), (y == x)
+ | (y \in Interval (BOpen x) b)].
+Proof.
+move=> xab y; rewrite (itv_splitU false xab y); congr (_ || _).
+rewrite (@itv_splitU x true _ _ _ y); first by rewrite itv_xx inE.
+by move: xab; rewrite boundl_in_itv itv_boundlr => /andP [].
+Qed.
+
+End IntervalOrdered.
+
+Section IntervalField.
+
+Variable R : realFieldType.
+
+Lemma mid_in_itv : forall ba bb (xa xb : R), xa <= xb ?< if (ba || bb)
+ -> mid xa xb \in Interval (BOpen_if ba xa) (BOpen_if bb xb).
+Proof.
+by move=> [] [] xa xb /= hx; apply/itv_dec=> /=; rewrite ?midf_lte // ?ltrW.
+Qed.
+
+Lemma mid_in_itvoo : forall (xa xb : R), xa < xb -> mid xa xb \in `]xa, xb[.
+Proof. by move=> xa xb hx; apply: mid_in_itv. Qed.
+
+Lemma mid_in_itvcc : forall (xa xb : R), xa <= xb -> mid xa xb \in `[xa, xb].
+Proof. by move=> xa xb hx; apply: mid_in_itv. Qed.
+
+End IntervalField.
diff --git a/mathcomp/algebra/matrix.v b/mathcomp/algebra/matrix.v
new file mode 100644
index 0000000..ecf5801
--- /dev/null
+++ b/mathcomp/algebra/matrix.v
@@ -0,0 +1,2872 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div choice fintype.
+Require Import finfun bigop prime binomial ssralg finset fingroup finalg.
+Require Import perm zmodp.
+
+(******************************************************************************)
+(* Basic concrete linear algebra : definition of type for matrices, and all *)
+(* basic matrix operations including determinant, trace and support for block *)
+(* decomposition. Matrices are represented by a row-major list of their *)
+(* coefficients but this implementation is hidden by three levels of wrappers *)
+(* (Matrix/Finfun/Tuple) so the matrix type should be treated as abstract and *)
+(* handled using only the operations described below: *)
+(* 'M[R]_(m, n) == the type of m rows by n columns matrices with *)
+(* 'M_(m, n) coefficients in R; the [R] is optional and is usually *)
+(* omitted. *)
+(* 'M[R]_n, 'M_n == the type of n x n square matrices. *)
+(* 'rV[R]_n, 'rV_n == the type of 1 x n row vectors. *)
+(* 'cV[R]_n, 'cV_n == the type of n x 1 column vectors. *)
+(* \matrix_(i < m, j < n) Expr(i, j) == *)
+(* the m x n matrix with general coefficient Expr(i, j), *)
+(* with i : 'I_m and j : 'I_n. the < m bound can be omitted *)
+(* if it is equal to n, though usually both bounds are *)
+(* omitted as they can be inferred from the context. *)
+(* \row_(j < n) Expr(j), \col_(i < m) Expr(i) *)
+(* the row / column vectors with general term Expr; the *)
+(* parentheses can be omitted along with the bound. *)
+(* \matrix_(i < m) RowExpr(i) == *)
+(* the m x n matrix with row i given by RowExpr(i) : 'rV_n. *)
+(* A i j == the coefficient of matrix A : 'M_(m, n) in column j of *)
+(* row i, where i : 'I_m, and j : 'I_n (via the coercion *)
+(* fun_of_matrix : matrix >-> Funclass). *)
+(* const_mx a == the constant matrix whose entries are all a (dimensions *)
+(* should be determined by context). *)
+(* map_mx f A == the pointwise image of A by f, i.e., the matrix Af *)
+(* congruent to A with Af i j = f (A i j) for all i and j. *)
+(* A^T == the matrix transpose of A. *)
+(* row i A == the i'th row of A (this is a row vector). *)
+(* col j A == the j'th column of A (a column vector). *)
+(* row' i A == A with the i'th row spliced out. *)
+(* col' i A == A with the j'th column spliced out. *)
+(* xrow i1 i2 A == A with rows i1 and i2 interchanged. *)
+(* xcol j1 j2 A == A with columns j1 and j2 interchanged. *)
+(* row_perm s A == A : 'M_(m, n) with rows permuted by s : 'S_m. *)
+(* col_perm s A == A : 'M_(m, n) with columns permuted by s : 'S_n. *)
+(* row_mx Al Ar == the row block matrix <Al Ar> obtained by contatenating *)
+(* two matrices Al and Ar of the same height. *)
+(* col_mx Au Ad == the column block matrix / Au \ (Au and Ad must have the *)
+(* same width). \ Ad / *)
+(* block_mx Aul Aur Adl Adr == the block matrix / Aul Aur \ *)
+(* \ Adl Adr / *)
+(* [l|r]submx A == the left/right submatrices of a row block matrix A. *)
+(* Note that the type of A, 'M_(m, n1 + n2) indicates how A *)
+(* should be decomposed. *)
+(* [u|d]submx A == the up/down submatrices of a column block matrix A. *)
+(* [u|d][l|r]submx A == the upper left, etc submatrices of a block matrix A. *)
+(* castmx eq_mn A == A : 'M_(m, n) cast to 'M_(m', n') using the equation *)
+(* pair eq_mn : (m = m') * (n = n'). This is the usual *)
+(* workaround for the syntactic limitations of dependent *)
+(* types in Coq, and can be used to introduce a block *)
+(* decomposition. It simplifies to A when eq_mn is the *)
+(* pair (erefl m, erefl n) (using rewrite /castmx /=). *)
+(* conform_mx B A == A if A and B have the same dimensions, else B. *)
+(* mxvec A == a row vector of width m * n holding all the entries of *)
+(* the m x n matrix A. *)
+(* mxvec_index i j == the index of A i j in mxvec A. *)
+(* vec_mx v == the inverse of mxvec, reshaping a vector of width m * n *)
+(* back into into an m x n rectangular matrix. *)
+(* In 'M[R]_(m, n), R can be any type, but 'M[R]_(m, n) inherits the eqType, *)
+(* choiceType, countType, finType, zmodType structures of R; 'M[R]_(m, n) *)
+(* also has a natural lmodType R structure when R has a ringType structure. *)
+(* Because the type of matrices specifies their dimension, only non-trivial *)
+(* square matrices (of type 'M[R]_n.+1) can inherit the ring structure of R; *)
+(* indeed they then have an algebra structure (lalgType R, or algType R if R *)
+(* is a comRingType, or even unitAlgType if R is a comUnitRingType). *)
+(* We thus provide separate syntax for the general matrix multiplication, *)
+(* and other operations for matrices over a ringType R: *)
+(* A *m B == the matrix product of A and B; the width of A must be *)
+(* equal to the height of B. *)
+(* a%:M == the scalar matrix with a's on the main diagonal; in *)
+(* particular 1%:M denotes the identity matrix, and is is *)
+(* equal to 1%R when n is of the form n'.+1 (e.g., n >= 1). *)
+(* is_scalar_mx A <=> A is a scalar matrix (A = a%:M for some A). *)
+(* diag_mx d == the diagonal matrix whose main diagonal is d : 'rV_n. *)
+(* delta_mx i j == the matrix with a 1 in row i, column j and 0 elsewhere. *)
+(* pid_mx r == the partial identity matrix with 1s only on the r first *)
+(* coefficients of the main diagonal; the dimensions of *)
+(* pid_mx r are determined by the context, and pid_mx r can *)
+(* be rectangular. *)
+(* copid_mx r == the complement to 1%:M of pid_mx r: a square diagonal *)
+(* matrix with 1s on all but the first r coefficients on *)
+(* its main diagonal. *)
+(* perm_mx s == the n x n permutation matrix for s : 'S_n. *)
+(* tperm_mx i1 i2 == the permutation matrix that exchanges i1 i2 : 'I_n. *)
+(* is_perm_mx A == A is a permutation matrix. *)
+(* lift0_mx A == the 1 + n square matrix block_mx 1 0 0 A when A : 'M_n. *)
+(* \tr A == the trace of a square matrix A. *)
+(* \det A == the determinant of A, using the Leibnitz formula. *)
+(* cofactor i j A == the i, j cofactor of A (the signed i, j minor of A), *)
+(* \adj A == the adjugate matrix of A (\adj A i j = cofactor j i A). *)
+(* A \in unitmx == A is invertible (R must be a comUnitRingType). *)
+(* invmx A == the inverse matrix of A if A \in unitmx A, otherwise A. *)
+(* The following operations provide a correspondance between linear functions *)
+(* and matrices: *)
+(* lin1_mx f == the m x n matrix that emulates via right product *)
+(* a (linear) function f : 'rV_m -> 'rV_n on ROW VECTORS *)
+(* lin_mx f == the (m1 * n1) x (m2 * n2) matrix that emulates, via the *)
+(* right multiplication on the mxvec encodings, a linear *)
+(* function f : 'M_(m1, n1) -> 'M_(m2, n2) *)
+(* lin_mul_row u := lin1_mx (mulmx u \o vec_mx) (applies a row-encoded *)
+(* function to the row-vector u). *)
+(* mulmx A == partially applied matrix multiplication (mulmx A B is *)
+(* displayed as A *m B), with, for A : 'M_(m, n), a *)
+(* canonical {linear 'M_(n, p) -> 'M(m, p}} structure. *)
+(* mulmxr A == self-simplifying right-hand matrix multiplication, i.e., *)
+(* mulmxr A B simplifies to B *m A, with, for A : 'M_(n, p), *)
+(* a canonical {linear 'M_(m, n) -> 'M(m, p}} structure. *)
+(* lin_mulmx A := lin_mx (mulmx A). *)
+(* lin_mulmxr A := lin_mx (mulmxr A). *)
+(* We also extend any finType structure of R to 'M[R]_(m, n), and define: *)
+(* {'GL_n[R]} == the finGroupType of units of 'M[R]_n.-1.+1. *)
+(* 'GL_n[R] == the general linear group of all matrices in {'GL_n(R)}. *)
+(* 'GL_n(p) == 'GL_n['F_p], the general linear group of a prime field. *)
+(* GLval u == the coercion of u : {'GL_n(R)} to a matrix. *)
+(* In addition to the lemmas relevant to these definitions, this file also *)
+(* proves several classic results, including : *)
+(* - The determinant is a multilinear alternate form. *)
+(* - The Laplace determinant expansion formulas: expand_det_[row|col]. *)
+(* - The Cramer rule : mul_mx_adj & mul_adj_mx. *)
+(* Finally, as an example of the use of block products, we program and prove *)
+(* the correctness of a classical linear algebra algorithm: *)
+(* cormenLUP A == the triangular decomposition (L, U, P) of a nontrivial *)
+(* square matrix A into a lower triagular matrix L with 1s *)
+(* on the main diagonal, an upper matrix U, and a *)
+(* permutation matrix P, such that P * A = L * U. *)
+(* This is example only; we use a different, more precise algorithm to *)
+(* develop the theory of matrix ranks and row spaces in mxalgebra.v *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+Import GRing.Theory.
+Open Local Scope ring_scope.
+
+Reserved Notation "''M_' n" (at level 8, n at level 2, format "''M_' n").
+Reserved Notation "''rV_' n" (at level 8, n at level 2, format "''rV_' n").
+Reserved Notation "''cV_' n" (at level 8, n at level 2, format "''cV_' n").
+Reserved Notation "''M_' ( n )" (at level 8, only parsing).
+Reserved Notation "''M_' ( m , n )" (at level 8, format "''M_' ( m , n )").
+Reserved Notation "''M[' R ]_ n" (at level 8, n at level 2, only parsing).
+Reserved Notation "''rV[' R ]_ n" (at level 8, n at level 2, only parsing).
+Reserved Notation "''cV[' R ]_ n" (at level 8, n at level 2, only parsing).
+Reserved Notation "''M[' R ]_ ( n )" (at level 8, only parsing).
+Reserved Notation "''M[' R ]_ ( m , n )" (at level 8, only parsing).
+
+Reserved Notation "\matrix_ i E"
+ (at level 36, E at level 36, i at level 2,
+ format "\matrix_ i E").
+Reserved Notation "\matrix_ ( i < n ) E"
+ (at level 36, E at level 36, i, n at level 50, only parsing).
+Reserved Notation "\matrix_ ( i , j ) E"
+ (at level 36, E at level 36, i, j at level 50,
+ format "\matrix_ ( i , j ) E").
+Reserved Notation "\matrix[ k ]_ ( i , j ) E"
+ (at level 36, E at level 36, i, j at level 50,
+ format "\matrix[ k ]_ ( i , j ) E").
+Reserved Notation "\matrix_ ( i < m , j < n ) E"
+ (at level 36, E at level 36, i, m, j, n at level 50, only parsing).
+Reserved Notation "\matrix_ ( i , j < n ) E"
+ (at level 36, E at level 36, i, j, n at level 50, only parsing).
+Reserved Notation "\row_ j E"
+ (at level 36, E at level 36, j at level 2,
+ format "\row_ j E").
+Reserved Notation "\row_ ( j < n ) E"
+ (at level 36, E at level 36, j, n at level 50, only parsing).
+Reserved Notation "\col_ j E"
+ (at level 36, E at level 36, j at level 2,
+ format "\col_ j E").
+Reserved Notation "\col_ ( j < n ) E"
+ (at level 36, E at level 36, j, n at level 50, only parsing).
+
+Reserved Notation "x %:M" (at level 8, format "x %:M").
+Reserved Notation "A *m B" (at level 40, left associativity, format "A *m B").
+Reserved Notation "A ^T" (at level 8, format "A ^T").
+Reserved Notation "\tr A" (at level 10, A at level 8, format "\tr A").
+Reserved Notation "\det A" (at level 10, A at level 8, format "\det A").
+Reserved Notation "\adj A" (at level 10, A at level 8, format "\adj A").
+
+Notation Local simp := (Monoid.Theory.simpm, oppr0).
+
+(*****************************************************************************)
+(****************************Type Definition**********************************)
+(*****************************************************************************)
+
+Section MatrixDef.
+
+Variable R : Type.
+Variables m n : nat.
+
+(* Basic linear algebra (matrices). *)
+(* We use dependent types (ordinals) for the indices so that ranges are *)
+(* mostly inferred automatically *)
+
+Inductive matrix : predArgType := Matrix of {ffun 'I_m * 'I_n -> R}.
+
+Definition mx_val A := let: Matrix g := A in g.
+
+Canonical matrix_subType := Eval hnf in [newType for mx_val].
+
+Fact matrix_key : unit. Proof. by []. Qed.
+Definition matrix_of_fun_def F := Matrix [ffun ij => F ij.1 ij.2].
+Definition matrix_of_fun k := locked_with k matrix_of_fun_def.
+Canonical matrix_unlockable k := [unlockable fun matrix_of_fun k].
+
+Definition fun_of_matrix A (i : 'I_m) (j : 'I_n) := mx_val A (i, j).
+
+Coercion fun_of_matrix : matrix >-> Funclass.
+
+Lemma mxE k F : matrix_of_fun k F =2 F.
+Proof. by move=> i j; rewrite unlock /fun_of_matrix /= ffunE. Qed.
+
+Lemma matrixP (A B : matrix) : A =2 B <-> A = B.
+Proof.
+rewrite /fun_of_matrix; split=> [/= eqAB | -> //].
+by apply/val_inj/ffunP=> [[i j]]; exact: eqAB.
+Qed.
+
+End MatrixDef.
+
+Bind Scope ring_scope with matrix.
+
+Notation "''M[' R ]_ ( m , n )" := (matrix R m n) (only parsing): type_scope.
+Notation "''rV[' R ]_ n" := 'M[R]_(1, n) (only parsing) : type_scope.
+Notation "''cV[' R ]_ n" := 'M[R]_(n, 1) (only parsing) : type_scope.
+Notation "''M[' R ]_ n" := 'M[R]_(n, n) (only parsing) : type_scope.
+Notation "''M[' R ]_ ( n )" := 'M[R]_n (only parsing) : type_scope.
+Notation "''M_' ( m , n )" := 'M[_]_(m, n) : type_scope.
+Notation "''rV_' n" := 'M_(1, n) : type_scope.
+Notation "''cV_' n" := 'M_(n, 1) : type_scope.
+Notation "''M_' n" := 'M_(n, n) : type_scope.
+Notation "''M_' ( n )" := 'M_n (only parsing) : type_scope.
+
+Notation "\matrix[ k ]_ ( i , j ) E" := (matrix_of_fun k (fun i j => E))
+ (at level 36, E at level 36, i, j at level 50): ring_scope.
+
+Notation "\matrix_ ( i < m , j < n ) E" :=
+ (@matrix_of_fun _ m n matrix_key (fun i j => E)) (only parsing) : ring_scope.
+
+Notation "\matrix_ ( i , j < n ) E" :=
+ (\matrix_(i < n, j < n) E) (only parsing) : ring_scope.
+
+Notation "\matrix_ ( i , j ) E" := (\matrix_(i < _, j < _) E) : ring_scope.
+
+Notation "\matrix_ ( i < m ) E" :=
+ (\matrix_(i < m, j < _) @fun_of_matrix _ 1 _ E 0 j)
+ (only parsing) : ring_scope.
+Notation "\matrix_ i E" := (\matrix_(i < _) E) : ring_scope.
+
+Notation "\col_ ( i < n ) E" := (@matrix_of_fun _ n 1 matrix_key (fun i _ => E))
+ (only parsing) : ring_scope.
+Notation "\col_ i E" := (\col_(i < _) E) : ring_scope.
+
+Notation "\row_ ( j < n ) E" := (@matrix_of_fun _ 1 n matrix_key (fun _ j => E))
+ (only parsing) : ring_scope.
+Notation "\row_ j E" := (\row_(j < _) E) : ring_scope.
+
+Definition matrix_eqMixin (R : eqType) m n :=
+ Eval hnf in [eqMixin of 'M[R]_(m, n) by <:].
+Canonical matrix_eqType (R : eqType) m n:=
+ Eval hnf in EqType 'M[R]_(m, n) (matrix_eqMixin R m n).
+Definition matrix_choiceMixin (R : choiceType) m n :=
+ [choiceMixin of 'M[R]_(m, n) by <:].
+Canonical matrix_choiceType (R : choiceType) m n :=
+ Eval hnf in ChoiceType 'M[R]_(m, n) (matrix_choiceMixin R m n).
+Definition matrix_countMixin (R : countType) m n :=
+ [countMixin of 'M[R]_(m, n) by <:].
+Canonical matrix_countType (R : countType) m n :=
+ Eval hnf in CountType 'M[R]_(m, n) (matrix_countMixin R m n).
+Canonical matrix_subCountType (R : countType) m n :=
+ Eval hnf in [subCountType of 'M[R]_(m, n)].
+Definition matrix_finMixin (R : finType) m n :=
+ [finMixin of 'M[R]_(m, n) by <:].
+Canonical matrix_finType (R : finType) m n :=
+ Eval hnf in FinType 'M[R]_(m, n) (matrix_finMixin R m n).
+Canonical matrix_subFinType (R : finType) m n :=
+ Eval hnf in [subFinType of 'M[R]_(m, n)].
+
+Lemma card_matrix (F : finType) m n : (#|{: 'M[F]_(m, n)}| = #|F| ^ (m * n))%N.
+Proof. by rewrite card_sub card_ffun card_prod !card_ord. Qed.
+
+(*****************************************************************************)
+(****** Matrix structural operations (transpose, permutation, blocks) ********)
+(*****************************************************************************)
+
+Section MatrixStructural.
+
+Variable R : Type.
+
+(* Constant matrix *)
+Fact const_mx_key : unit. Proof. by []. Qed.
+Definition const_mx m n a : 'M[R]_(m, n) := \matrix[const_mx_key]_(i, j) a.
+Implicit Arguments const_mx [[m] [n]].
+
+Section FixedDim.
+(* Definitions and properties for which we can work with fixed dimensions. *)
+
+Variables m n : nat.
+Implicit Type A : 'M[R]_(m, n).
+
+(* Reshape a matrix, to accomodate the block functions for instance. *)
+Definition castmx m' n' (eq_mn : (m = m') * (n = n')) A : 'M_(m', n') :=
+ let: erefl in _ = m' := eq_mn.1 return 'M_(m', n') in
+ let: erefl in _ = n' := eq_mn.2 return 'M_(m, n') in A.
+
+Definition conform_mx m' n' B A :=
+ match m =P m', n =P n' with
+ | ReflectT eq_m, ReflectT eq_n => castmx (eq_m, eq_n) A
+ | _, _ => B
+ end.
+
+(* Transpose a matrix *)
+Fact trmx_key : unit. Proof. by []. Qed.
+Definition trmx A := \matrix[trmx_key]_(i, j) A j i.
+
+(* Permute a matrix vertically (rows) or horizontally (columns) *)
+Fact row_perm_key : unit. Proof. by []. Qed.
+Definition row_perm (s : 'S_m) A := \matrix[row_perm_key]_(i, j) A (s i) j.
+Fact col_perm_key : unit. Proof. by []. Qed.
+Definition col_perm (s : 'S_n) A := \matrix[col_perm_key]_(i, j) A i (s j).
+
+(* Exchange two rows/columns of a matrix *)
+Definition xrow i1 i2 := row_perm (tperm i1 i2).
+Definition xcol j1 j2 := col_perm (tperm j1 j2).
+
+(* Row/Column sub matrices of a matrix *)
+Definition row i0 A := \row_j A i0 j.
+Definition col j0 A := \col_i A i j0.
+
+(* Removing a row/column from a matrix *)
+Definition row' i0 A := \matrix_(i, j) A (lift i0 i) j.
+Definition col' j0 A := \matrix_(i, j) A i (lift j0 j).
+
+Lemma castmx_const m' n' (eq_mn : (m = m') * (n = n')) a :
+ castmx eq_mn (const_mx a) = const_mx a.
+Proof. by case: eq_mn; case: m' /; case: n' /. Qed.
+
+Lemma trmx_const a : trmx (const_mx a) = const_mx a.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma row_perm_const s a : row_perm s (const_mx a) = const_mx a.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma col_perm_const s a : col_perm s (const_mx a) = const_mx a.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma xrow_const i1 i2 a : xrow i1 i2 (const_mx a) = const_mx a.
+Proof. exact: row_perm_const. Qed.
+
+Lemma xcol_const j1 j2 a : xcol j1 j2 (const_mx a) = const_mx a.
+Proof. exact: col_perm_const. Qed.
+
+Lemma rowP (u v : 'rV[R]_n) : u 0 =1 v 0 <-> u = v.
+Proof. by split=> [eq_uv | -> //]; apply/matrixP=> i; rewrite ord1. Qed.
+
+Lemma rowK u_ i0 : row i0 (\matrix_i u_ i) = u_ i0.
+Proof. by apply/rowP=> i'; rewrite !mxE. Qed.
+
+Lemma row_matrixP A B : (forall i, row i A = row i B) <-> A = B.
+Proof.
+split=> [eqAB | -> //]; apply/matrixP=> i j.
+by move/rowP/(_ j): (eqAB i); rewrite !mxE.
+Qed.
+
+Lemma colP (u v : 'cV[R]_m) : u^~ 0 =1 v^~ 0 <-> u = v.
+Proof. by split=> [eq_uv | -> //]; apply/matrixP=> i j; rewrite ord1. Qed.
+
+Lemma row_const i0 a : row i0 (const_mx a) = const_mx a.
+Proof. by apply/rowP=> j; rewrite !mxE. Qed.
+
+Lemma col_const j0 a : col j0 (const_mx a) = const_mx a.
+Proof. by apply/colP=> i; rewrite !mxE. Qed.
+
+Lemma row'_const i0 a : row' i0 (const_mx a) = const_mx a.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma col'_const j0 a : col' j0 (const_mx a) = const_mx a.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma col_perm1 A : col_perm 1 A = A.
+Proof. by apply/matrixP=> i j; rewrite mxE perm1. Qed.
+
+Lemma row_perm1 A : row_perm 1 A = A.
+Proof. by apply/matrixP=> i j; rewrite mxE perm1. Qed.
+
+Lemma col_permM s t A : col_perm (s * t) A = col_perm s (col_perm t A).
+Proof. by apply/matrixP=> i j; rewrite !mxE permM. Qed.
+
+Lemma row_permM s t A : row_perm (s * t) A = row_perm s (row_perm t A).
+Proof. by apply/matrixP=> i j; rewrite !mxE permM. Qed.
+
+Lemma col_row_permC s t A :
+ col_perm s (row_perm t A) = row_perm t (col_perm s A).
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+End FixedDim.
+
+Local Notation "A ^T" := (trmx A) : ring_scope.
+
+Lemma castmx_id m n erefl_mn (A : 'M_(m, n)) : castmx erefl_mn A = A.
+Proof. by case: erefl_mn => e_m e_n; rewrite [e_m]eq_axiomK [e_n]eq_axiomK. Qed.
+
+Lemma castmx_comp m1 n1 m2 n2 m3 n3 (eq_m1 : m1 = m2) (eq_n1 : n1 = n2)
+ (eq_m2 : m2 = m3) (eq_n2 : n2 = n3) A :
+ castmx (eq_m2, eq_n2) (castmx (eq_m1, eq_n1) A)
+ = castmx (etrans eq_m1 eq_m2, etrans eq_n1 eq_n2) A.
+Proof.
+by case: m2 / eq_m1 eq_m2; case: m3 /; case: n2 / eq_n1 eq_n2; case: n3 /.
+Qed.
+
+Lemma castmxK m1 n1 m2 n2 (eq_m : m1 = m2) (eq_n : n1 = n2) :
+ cancel (castmx (eq_m, eq_n)) (castmx (esym eq_m, esym eq_n)).
+Proof. by case: m2 / eq_m; case: n2 / eq_n. Qed.
+
+Lemma castmxKV m1 n1 m2 n2 (eq_m : m1 = m2) (eq_n : n1 = n2) :
+ cancel (castmx (esym eq_m, esym eq_n)) (castmx (eq_m, eq_n)).
+Proof. by case: m2 / eq_m; case: n2 / eq_n. Qed.
+
+(* This can be use to reverse an equation that involves a cast. *)
+Lemma castmx_sym m1 n1 m2 n2 (eq_m : m1 = m2) (eq_n : n1 = n2) A1 A2 :
+ A1 = castmx (eq_m, eq_n) A2 -> A2 = castmx (esym eq_m, esym eq_n) A1.
+Proof. by move/(canLR (castmxK _ _)). Qed.
+
+Lemma castmxE m1 n1 m2 n2 (eq_mn : (m1 = m2) * (n1 = n2)) A i j :
+ castmx eq_mn A i j =
+ A (cast_ord (esym eq_mn.1) i) (cast_ord (esym eq_mn.2) j).
+Proof.
+by do [case: eq_mn; case: m2 /; case: n2 /] in A i j *; rewrite !cast_ord_id.
+Qed.
+
+Lemma conform_mx_id m n (B A : 'M_(m, n)) : conform_mx B A = A.
+Proof. by rewrite /conform_mx; do 2!case: eqP => // *; rewrite castmx_id. Qed.
+
+Lemma nonconform_mx m m' n n' (B : 'M_(m', n')) (A : 'M_(m, n)) :
+ (m != m') || (n != n') -> conform_mx B A = B.
+Proof. by rewrite /conform_mx; do 2!case: eqP. Qed.
+
+Lemma conform_castmx m1 n1 m2 n2 m3 n3
+ (e_mn : (m2 = m3) * (n2 = n3)) (B : 'M_(m1, n1)) A :
+ conform_mx B (castmx e_mn A) = conform_mx B A.
+Proof. by do [case: e_mn; case: m3 /; case: n3 /] in A *. Qed.
+
+Lemma trmxK m n : cancel (@trmx m n) (@trmx n m).
+Proof. by move=> A; apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma trmx_inj m n : injective (@trmx m n).
+Proof. exact: can_inj (@trmxK m n). Qed.
+
+Lemma trmx_cast m1 n1 m2 n2 (eq_mn : (m1 = m2) * (n1 = n2)) A :
+ (castmx eq_mn A)^T = castmx (eq_mn.2, eq_mn.1) A^T.
+Proof.
+by case: eq_mn => eq_m eq_n; apply/matrixP=> i j; rewrite !(mxE, castmxE).
+Qed.
+
+Lemma tr_row_perm m n s (A : 'M_(m, n)) : (row_perm s A)^T = col_perm s A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma tr_col_perm m n s (A : 'M_(m, n)) : (col_perm s A)^T = row_perm s A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma tr_xrow m n i1 i2 (A : 'M_(m, n)) : (xrow i1 i2 A)^T = xcol i1 i2 A^T.
+Proof. exact: tr_row_perm. Qed.
+
+Lemma tr_xcol m n j1 j2 (A : 'M_(m, n)) : (xcol j1 j2 A)^T = xrow j1 j2 A^T.
+Proof. exact: tr_col_perm. Qed.
+
+Lemma row_id n i (V : 'rV_n) : row i V = V.
+Proof. by apply/rowP=> j; rewrite mxE [i]ord1. Qed.
+
+Lemma col_id n j (V : 'cV_n) : col j V = V.
+Proof. by apply/colP=> i; rewrite mxE [j]ord1. Qed.
+
+Lemma row_eq m1 m2 n i1 i2 (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) :
+ row i1 A1 = row i2 A2 -> A1 i1 =1 A2 i2.
+Proof. by move/rowP=> eqA12 j; have:= eqA12 j; rewrite !mxE. Qed.
+
+Lemma col_eq m n1 n2 j1 j2 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) :
+ col j1 A1 = col j2 A2 -> A1^~ j1 =1 A2^~ j2.
+Proof. by move/colP=> eqA12 i; have:= eqA12 i; rewrite !mxE. Qed.
+
+Lemma row'_eq m n i0 (A B : 'M_(m, n)) :
+ row' i0 A = row' i0 B -> {in predC1 i0, A =2 B}.
+Proof.
+move/matrixP=> eqAB' i; rewrite !inE eq_sym; case/unlift_some=> i' -> _ j.
+by have:= eqAB' i' j; rewrite !mxE.
+Qed.
+
+Lemma col'_eq m n j0 (A B : 'M_(m, n)) :
+ col' j0 A = col' j0 B -> forall i, {in predC1 j0, A i =1 B i}.
+Proof.
+move/matrixP=> eqAB' i j; rewrite !inE eq_sym; case/unlift_some=> j' -> _.
+by have:= eqAB' i j'; rewrite !mxE.
+Qed.
+
+Lemma tr_row m n i0 (A : 'M_(m, n)) : (row i0 A)^T = col i0 A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma tr_row' m n i0 (A : 'M_(m, n)) : (row' i0 A)^T = col' i0 A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma tr_col m n j0 (A : 'M_(m, n)) : (col j0 A)^T = row j0 A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma tr_col' m n j0 (A : 'M_(m, n)) : (col' j0 A)^T = row' j0 A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Ltac split_mxE := apply/matrixP=> i j; do ![rewrite mxE | case: split => ?].
+
+Section CutPaste.
+
+Variables m m1 m2 n n1 n2 : nat.
+
+(* Concatenating two matrices, in either direction. *)
+
+Fact row_mx_key : unit. Proof. by []. Qed.
+Definition row_mx (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) : 'M[R]_(m, n1 + n2) :=
+ \matrix[row_mx_key]_(i, j)
+ match split j with inl j1 => A1 i j1 | inr j2 => A2 i j2 end.
+
+Fact col_mx_key : unit. Proof. by []. Qed.
+Definition col_mx (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) : 'M[R]_(m1 + m2, n) :=
+ \matrix[col_mx_key]_(i, j)
+ match split i with inl i1 => A1 i1 j | inr i2 => A2 i2 j end.
+
+(* Left/Right | Up/Down submatrices of a rows | columns matrix. *)
+(* The shape of the (dependent) width parameters of the type of A *)
+(* determines which submatrix is selected. *)
+
+Fact lsubmx_key : unit. Proof. by []. Qed.
+Definition lsubmx (A : 'M[R]_(m, n1 + n2)) :=
+ \matrix[lsubmx_key]_(i, j) A i (lshift n2 j).
+
+Fact rsubmx_key : unit. Proof. by []. Qed.
+Definition rsubmx (A : 'M[R]_(m, n1 + n2)) :=
+ \matrix[rsubmx_key]_(i, j) A i (rshift n1 j).
+
+Fact usubmx_key : unit. Proof. by []. Qed.
+Definition usubmx (A : 'M[R]_(m1 + m2, n)) :=
+ \matrix[usubmx_key]_(i, j) A (lshift m2 i) j.
+
+Fact dsubmx_key : unit. Proof. by []. Qed.
+Definition dsubmx (A : 'M[R]_(m1 + m2, n)) :=
+ \matrix[dsubmx_key]_(i, j) A (rshift m1 i) j.
+
+Lemma row_mxEl A1 A2 i j : row_mx A1 A2 i (lshift n2 j) = A1 i j.
+Proof. by rewrite mxE (unsplitK (inl _ _)). Qed.
+
+Lemma row_mxKl A1 A2 : lsubmx (row_mx A1 A2) = A1.
+Proof. by apply/matrixP=> i j; rewrite mxE row_mxEl. Qed.
+
+Lemma row_mxEr A1 A2 i j : row_mx A1 A2 i (rshift n1 j) = A2 i j.
+Proof. by rewrite mxE (unsplitK (inr _ _)). Qed.
+
+Lemma row_mxKr A1 A2 : rsubmx (row_mx A1 A2) = A2.
+Proof. by apply/matrixP=> i j; rewrite mxE row_mxEr. Qed.
+
+Lemma hsubmxK A : row_mx (lsubmx A) (rsubmx A) = A.
+Proof.
+apply/matrixP=> i j; rewrite !mxE.
+case: splitP => k Dk //=; rewrite !mxE //=; congr (A _ _); exact: val_inj.
+Qed.
+
+Lemma col_mxEu A1 A2 i j : col_mx A1 A2 (lshift m2 i) j = A1 i j.
+Proof. by rewrite mxE (unsplitK (inl _ _)). Qed.
+
+Lemma col_mxKu A1 A2 : usubmx (col_mx A1 A2) = A1.
+Proof. by apply/matrixP=> i j; rewrite mxE col_mxEu. Qed.
+
+Lemma col_mxEd A1 A2 i j : col_mx A1 A2 (rshift m1 i) j = A2 i j.
+Proof. by rewrite mxE (unsplitK (inr _ _)). Qed.
+
+Lemma col_mxKd A1 A2 : dsubmx (col_mx A1 A2) = A2.
+Proof. by apply/matrixP=> i j; rewrite mxE col_mxEd. Qed.
+
+Lemma eq_row_mx A1 A2 B1 B2 : row_mx A1 A2 = row_mx B1 B2 -> A1 = B1 /\ A2 = B2.
+Proof.
+move=> eqAB; move: (congr1 lsubmx eqAB) (congr1 rsubmx eqAB).
+by rewrite !(row_mxKl, row_mxKr).
+Qed.
+
+Lemma eq_col_mx A1 A2 B1 B2 : col_mx A1 A2 = col_mx B1 B2 -> A1 = B1 /\ A2 = B2.
+Proof.
+move=> eqAB; move: (congr1 usubmx eqAB) (congr1 dsubmx eqAB).
+by rewrite !(col_mxKu, col_mxKd).
+Qed.
+
+Lemma row_mx_const a : row_mx (const_mx a) (const_mx a) = const_mx a.
+Proof. by split_mxE. Qed.
+
+Lemma col_mx_const a : col_mx (const_mx a) (const_mx a) = const_mx a.
+Proof. by split_mxE. Qed.
+
+End CutPaste.
+
+Lemma trmx_lsub m n1 n2 (A : 'M_(m, n1 + n2)) : (lsubmx A)^T = usubmx A^T.
+Proof. by split_mxE. Qed.
+
+Lemma trmx_rsub m n1 n2 (A : 'M_(m, n1 + n2)) : (rsubmx A)^T = dsubmx A^T.
+Proof. by split_mxE. Qed.
+
+Lemma tr_row_mx m n1 n2 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) :
+ (row_mx A1 A2)^T = col_mx A1^T A2^T.
+Proof. by split_mxE. Qed.
+
+Lemma tr_col_mx m1 m2 n (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) :
+ (col_mx A1 A2)^T = row_mx A1^T A2^T.
+Proof. by split_mxE. Qed.
+
+Lemma trmx_usub m1 m2 n (A : 'M_(m1 + m2, n)) : (usubmx A)^T = lsubmx A^T.
+Proof. by split_mxE. Qed.
+
+Lemma trmx_dsub m1 m2 n (A : 'M_(m1 + m2, n)) : (dsubmx A)^T = rsubmx A^T.
+Proof. by split_mxE. Qed.
+
+Lemma vsubmxK m1 m2 n (A : 'M_(m1 + m2, n)) : col_mx (usubmx A) (dsubmx A) = A.
+Proof. by apply: trmx_inj; rewrite tr_col_mx trmx_usub trmx_dsub hsubmxK. Qed.
+
+Lemma cast_row_mx m m' n1 n2 (eq_m : m = m') A1 A2 :
+ castmx (eq_m, erefl _) (row_mx A1 A2)
+ = row_mx (castmx (eq_m, erefl n1) A1) (castmx (eq_m, erefl n2) A2).
+Proof. by case: m' / eq_m. Qed.
+
+Lemma cast_col_mx m1 m2 n n' (eq_n : n = n') A1 A2 :
+ castmx (erefl _, eq_n) (col_mx A1 A2)
+ = col_mx (castmx (erefl m1, eq_n) A1) (castmx (erefl m2, eq_n) A2).
+Proof. by case: n' / eq_n. Qed.
+
+(* This lemma has Prenex Implicits to help RL rewrititng with castmx_sym. *)
+Lemma row_mxA m n1 n2 n3 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) (A3 : 'M_(m, n3)) :
+ let cast := (erefl m, esym (addnA n1 n2 n3)) in
+ row_mx A1 (row_mx A2 A3) = castmx cast (row_mx (row_mx A1 A2) A3).
+Proof.
+apply: (canRL (castmxKV _ _)); apply/matrixP=> i j.
+rewrite castmxE !mxE cast_ord_id; case: splitP => j1 /= def_j.
+ have: (j < n1 + n2) && (j < n1) by rewrite def_j lshift_subproof /=.
+ by move: def_j; do 2![case: splitP => // ? ->; rewrite ?mxE] => /ord_inj->.
+case: splitP def_j => j2 ->{j} def_j; rewrite !mxE.
+ have: ~~ (j2 < n1) by rewrite -leqNgt def_j leq_addr.
+ have: j1 < n2 by rewrite -(ltn_add2l n1) -def_j.
+ by move: def_j; do 2![case: splitP => // ? ->] => /addnI/val_inj->.
+have: ~~ (j1 < n2) by rewrite -leqNgt -(leq_add2l n1) -def_j leq_addr.
+by case: splitP def_j => // ? ->; rewrite addnA => /addnI/val_inj->.
+Qed.
+Definition row_mxAx := row_mxA. (* bypass Prenex Implicits. *)
+
+(* This lemma has Prenex Implicits to help RL rewrititng with castmx_sym. *)
+Lemma col_mxA m1 m2 m3 n (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) (A3 : 'M_(m3, n)) :
+ let cast := (esym (addnA m1 m2 m3), erefl n) in
+ col_mx A1 (col_mx A2 A3) = castmx cast (col_mx (col_mx A1 A2) A3).
+Proof. by apply: trmx_inj; rewrite trmx_cast !tr_col_mx -row_mxA. Qed.
+Definition col_mxAx := col_mxA. (* bypass Prenex Implicits. *)
+
+Lemma row_row_mx m n1 n2 i0 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) :
+ row i0 (row_mx A1 A2) = row_mx (row i0 A1) (row i0 A2).
+Proof.
+by apply/matrixP=> i j; rewrite !mxE; case: (split j) => j'; rewrite mxE.
+Qed.
+
+Lemma col_col_mx m1 m2 n j0 (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) :
+ col j0 (col_mx A1 A2) = col_mx (col j0 A1) (col j0 A2).
+Proof. by apply: trmx_inj; rewrite !(tr_col, tr_col_mx, row_row_mx). Qed.
+
+Lemma row'_row_mx m n1 n2 i0 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) :
+ row' i0 (row_mx A1 A2) = row_mx (row' i0 A1) (row' i0 A2).
+Proof.
+by apply/matrixP=> i j; rewrite !mxE; case: (split j) => j'; rewrite mxE.
+Qed.
+
+Lemma col'_col_mx m1 m2 n j0 (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) :
+ col' j0 (col_mx A1 A2) = col_mx (col' j0 A1) (col' j0 A2).
+Proof. by apply: trmx_inj; rewrite !(tr_col', tr_col_mx, row'_row_mx). Qed.
+
+Lemma colKl m n1 n2 j1 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) :
+ col (lshift n2 j1) (row_mx A1 A2) = col j1 A1.
+Proof. by apply/matrixP=> i j; rewrite !(row_mxEl, mxE). Qed.
+
+Lemma colKr m n1 n2 j2 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) :
+ col (rshift n1 j2) (row_mx A1 A2) = col j2 A2.
+Proof. by apply/matrixP=> i j; rewrite !(row_mxEr, mxE). Qed.
+
+Lemma rowKu m1 m2 n i1 (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) :
+ row (lshift m2 i1) (col_mx A1 A2) = row i1 A1.
+Proof. by apply/matrixP=> i j; rewrite !(col_mxEu, mxE). Qed.
+
+Lemma rowKd m1 m2 n i2 (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) :
+ row (rshift m1 i2) (col_mx A1 A2) = row i2 A2.
+Proof. by apply/matrixP=> i j; rewrite !(col_mxEd, mxE). Qed.
+
+Lemma col'Kl m n1 n2 j1 (A1 : 'M_(m, n1.+1)) (A2 : 'M_(m, n2)) :
+ col' (lshift n2 j1) (row_mx A1 A2) = row_mx (col' j1 A1) A2.
+Proof.
+apply/matrixP=> i /= j; symmetry; rewrite 2!mxE.
+case: splitP => j' def_j'.
+ rewrite mxE -(row_mxEl _ A2); congr (row_mx _ _ _); apply: ord_inj.
+ by rewrite /= def_j'.
+rewrite -(row_mxEr A1); congr (row_mx _ _ _); apply: ord_inj => /=.
+by rewrite /bump def_j' -ltnS -addSn ltn_addr.
+Qed.
+
+Lemma row'Ku m1 m2 n i1 (A1 : 'M_(m1.+1, n)) (A2 : 'M_(m2, n)) :
+ row' (lshift m2 i1) (@col_mx m1.+1 m2 n A1 A2) = col_mx (row' i1 A1) A2.
+Proof.
+by apply: trmx_inj; rewrite tr_col_mx !(@tr_row' _.+1) (@tr_col_mx _.+1) col'Kl.
+Qed.
+
+Lemma mx'_cast m n : 'I_n -> (m + n.-1)%N = (m + n).-1.
+Proof. by case=> j /ltn_predK <-; rewrite addnS. Qed.
+
+Lemma col'Kr m n1 n2 j2 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) :
+ col' (rshift n1 j2) (@row_mx m n1 n2 A1 A2)
+ = castmx (erefl m, mx'_cast n1 j2) (row_mx A1 (col' j2 A2)).
+Proof.
+apply/matrixP=> i j; symmetry; rewrite castmxE mxE cast_ord_id.
+case: splitP => j' /= def_j.
+ rewrite mxE -(row_mxEl _ A2); congr (row_mx _ _ _); apply: ord_inj.
+ by rewrite /= def_j /bump leqNgt ltn_addr.
+rewrite 2!mxE -(row_mxEr A1); congr (row_mx _ _ _ _); apply: ord_inj.
+by rewrite /= def_j /bump leq_add2l addnCA.
+Qed.
+
+Lemma row'Kd m1 m2 n i2 (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) :
+ row' (rshift m1 i2) (col_mx A1 A2)
+ = castmx (mx'_cast m1 i2, erefl n) (col_mx A1 (row' i2 A2)).
+Proof. by apply: trmx_inj; rewrite trmx_cast !(tr_row', tr_col_mx) col'Kr. Qed.
+
+Section Block.
+
+Variables m1 m2 n1 n2 : nat.
+
+(* Building a block matrix from 4 matrices : *)
+(* up left, up right, down left and down right components *)
+
+Definition block_mx Aul Aur Adl Adr : 'M_(m1 + m2, n1 + n2) :=
+ col_mx (row_mx Aul Aur) (row_mx Adl Adr).
+
+Lemma eq_block_mx Aul Aur Adl Adr Bul Bur Bdl Bdr :
+ block_mx Aul Aur Adl Adr = block_mx Bul Bur Bdl Bdr ->
+ [/\ Aul = Bul, Aur = Bur, Adl = Bdl & Adr = Bdr].
+Proof. by case/eq_col_mx; do 2!case/eq_row_mx=> -> ->. Qed.
+
+Lemma block_mx_const a :
+ block_mx (const_mx a) (const_mx a) (const_mx a) (const_mx a) = const_mx a.
+Proof. by split_mxE. Qed.
+
+Section CutBlock.
+
+Variable A : matrix R (m1 + m2) (n1 + n2).
+
+Definition ulsubmx := lsubmx (usubmx A).
+Definition ursubmx := rsubmx (usubmx A).
+Definition dlsubmx := lsubmx (dsubmx A).
+Definition drsubmx := rsubmx (dsubmx A).
+
+Lemma submxK : block_mx ulsubmx ursubmx dlsubmx drsubmx = A.
+Proof. by rewrite /block_mx !hsubmxK vsubmxK. Qed.
+
+End CutBlock.
+
+Section CatBlock.
+
+Variables (Aul : 'M[R]_(m1, n1)) (Aur : 'M[R]_(m1, n2)).
+Variables (Adl : 'M[R]_(m2, n1)) (Adr : 'M[R]_(m2, n2)).
+
+Let A := block_mx Aul Aur Adl Adr.
+
+Lemma block_mxEul i j : A (lshift m2 i) (lshift n2 j) = Aul i j.
+Proof. by rewrite col_mxEu row_mxEl. Qed.
+Lemma block_mxKul : ulsubmx A = Aul.
+Proof. by rewrite /ulsubmx col_mxKu row_mxKl. Qed.
+
+Lemma block_mxEur i j : A (lshift m2 i) (rshift n1 j) = Aur i j.
+Proof. by rewrite col_mxEu row_mxEr. Qed.
+Lemma block_mxKur : ursubmx A = Aur.
+Proof. by rewrite /ursubmx col_mxKu row_mxKr. Qed.
+
+Lemma block_mxEdl i j : A (rshift m1 i) (lshift n2 j) = Adl i j.
+Proof. by rewrite col_mxEd row_mxEl. Qed.
+Lemma block_mxKdl : dlsubmx A = Adl.
+Proof. by rewrite /dlsubmx col_mxKd row_mxKl. Qed.
+
+Lemma block_mxEdr i j : A (rshift m1 i) (rshift n1 j) = Adr i j.
+Proof. by rewrite col_mxEd row_mxEr. Qed.
+Lemma block_mxKdr : drsubmx A = Adr.
+Proof. by rewrite /drsubmx col_mxKd row_mxKr. Qed.
+
+Lemma block_mxEv : A = col_mx (row_mx Aul Aur) (row_mx Adl Adr).
+Proof. by []. Qed.
+
+End CatBlock.
+
+End Block.
+
+Section TrCutBlock.
+
+Variables m1 m2 n1 n2 : nat.
+Variable A : 'M[R]_(m1 + m2, n1 + n2).
+
+Lemma trmx_ulsub : (ulsubmx A)^T = ulsubmx A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma trmx_ursub : (ursubmx A)^T = dlsubmx A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma trmx_dlsub : (dlsubmx A)^T = ursubmx A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma trmx_drsub : (drsubmx A)^T = drsubmx A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+End TrCutBlock.
+
+Section TrBlock.
+Variables m1 m2 n1 n2 : nat.
+Variables (Aul : 'M[R]_(m1, n1)) (Aur : 'M[R]_(m1, n2)).
+Variables (Adl : 'M[R]_(m2, n1)) (Adr : 'M[R]_(m2, n2)).
+
+Lemma tr_block_mx :
+ (block_mx Aul Aur Adl Adr)^T = block_mx Aul^T Adl^T Aur^T Adr^T.
+Proof.
+rewrite -[_^T]submxK -trmx_ulsub -trmx_ursub -trmx_dlsub -trmx_drsub.
+by rewrite block_mxKul block_mxKur block_mxKdl block_mxKdr.
+Qed.
+
+Lemma block_mxEh :
+ block_mx Aul Aur Adl Adr = row_mx (col_mx Aul Adl) (col_mx Aur Adr).
+Proof. by apply: trmx_inj; rewrite tr_block_mx tr_row_mx 2!tr_col_mx. Qed.
+End TrBlock.
+
+(* This lemma has Prenex Implicits to help RL rewrititng with castmx_sym. *)
+Lemma block_mxA m1 m2 m3 n1 n2 n3
+ (A11 : 'M_(m1, n1)) (A12 : 'M_(m1, n2)) (A13 : 'M_(m1, n3))
+ (A21 : 'M_(m2, n1)) (A22 : 'M_(m2, n2)) (A23 : 'M_(m2, n3))
+ (A31 : 'M_(m3, n1)) (A32 : 'M_(m3, n2)) (A33 : 'M_(m3, n3)) :
+ let cast := (esym (addnA m1 m2 m3), esym (addnA n1 n2 n3)) in
+ let row1 := row_mx A12 A13 in let col1 := col_mx A21 A31 in
+ let row3 := row_mx A31 A32 in let col3 := col_mx A13 A23 in
+ block_mx A11 row1 col1 (block_mx A22 A23 A32 A33)
+ = castmx cast (block_mx (block_mx A11 A12 A21 A22) col3 row3 A33).
+Proof.
+rewrite /= block_mxEh !col_mxA -cast_row_mx -block_mxEv -block_mxEh.
+rewrite block_mxEv block_mxEh !row_mxA -cast_col_mx -block_mxEh -block_mxEv.
+by rewrite castmx_comp etrans_id.
+Qed.
+Definition block_mxAx := block_mxA. (* Bypass Prenex Implicits *)
+
+(* Bijections mxvec : 'M_(m, n) <----> 'rV_(m * n) : vec_mx *)
+Section VecMatrix.
+
+Variables m n : nat.
+
+Lemma mxvec_cast : #|{:'I_m * 'I_n}| = (m * n)%N.
+Proof. by rewrite card_prod !card_ord. Qed.
+
+Definition mxvec_index (i : 'I_m) (j : 'I_n) :=
+ cast_ord mxvec_cast (enum_rank (i, j)).
+
+CoInductive is_mxvec_index : 'I_(m * n) -> Type :=
+ IsMxvecIndex i j : is_mxvec_index (mxvec_index i j).
+
+Lemma mxvec_indexP k : is_mxvec_index k.
+Proof.
+rewrite -[k](cast_ordK (esym mxvec_cast)) esymK.
+by rewrite -[_ k]enum_valK; case: (enum_val _).
+Qed.
+
+Coercion pair_of_mxvec_index k (i_k : is_mxvec_index k) :=
+ let: IsMxvecIndex i j := i_k in (i, j).
+
+Definition mxvec (A : 'M[R]_(m, n)) :=
+ castmx (erefl _, mxvec_cast) (\row_k A (enum_val k).1 (enum_val k).2).
+
+Fact vec_mx_key : unit. Proof. by []. Qed.
+Definition vec_mx (u : 'rV[R]_(m * n)) :=
+ \matrix[vec_mx_key]_(i, j) u 0 (mxvec_index i j).
+
+Lemma mxvecE A i j : mxvec A 0 (mxvec_index i j) = A i j.
+Proof. by rewrite castmxE mxE cast_ordK enum_rankK. Qed.
+
+Lemma mxvecK : cancel mxvec vec_mx.
+Proof. by move=> A; apply/matrixP=> i j; rewrite mxE mxvecE. Qed.
+
+Lemma vec_mxK : cancel vec_mx mxvec.
+Proof.
+by move=> u; apply/rowP=> k; case/mxvec_indexP: k => i j; rewrite mxvecE mxE.
+Qed.
+
+Lemma curry_mxvec_bij : {on 'I_(m * n), bijective (prod_curry mxvec_index)}.
+Proof.
+exists (enum_val \o cast_ord (esym mxvec_cast)) => [[i j] _ | k _] /=.
+ by rewrite cast_ordK enum_rankK.
+by case/mxvec_indexP: k => i j /=; rewrite cast_ordK enum_rankK.
+Qed.
+
+End VecMatrix.
+
+End MatrixStructural.
+
+Implicit Arguments const_mx [R m n].
+Implicit Arguments row_mxA [R m n1 n2 n3 A1 A2 A3].
+Implicit Arguments col_mxA [R m1 m2 m3 n A1 A2 A3].
+Implicit Arguments block_mxA
+ [R m1 m2 m3 n1 n2 n3 A11 A12 A13 A21 A22 A23 A31 A32 A33].
+Prenex Implicits const_mx castmx trmx lsubmx rsubmx usubmx dsubmx row_mx col_mx.
+Prenex Implicits block_mx ulsubmx ursubmx dlsubmx drsubmx.
+Prenex Implicits row_mxA col_mxA block_mxA.
+Prenex Implicits mxvec vec_mx mxvec_indexP mxvecK vec_mxK.
+
+Notation "A ^T" := (trmx A) : ring_scope.
+
+(* Matrix parametricity. *)
+Section MapMatrix.
+
+Variables (aT rT : Type) (f : aT -> rT).
+
+Fact map_mx_key : unit. Proof. by []. Qed.
+Definition map_mx m n (A : 'M_(m, n)) := \matrix[map_mx_key]_(i, j) f (A i j).
+
+Notation "A ^f" := (map_mx A) : ring_scope.
+
+Section OneMatrix.
+
+Variables (m n : nat) (A : 'M[aT]_(m, n)).
+
+Lemma map_trmx : A^f^T = A^T^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_const_mx a : (const_mx a)^f = const_mx (f a) :> 'M_(m, n).
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_row i : (row i A)^f = row i A^f.
+Proof. by apply/rowP=> j; rewrite !mxE. Qed.
+
+Lemma map_col j : (col j A)^f = col j A^f.
+Proof. by apply/colP=> i; rewrite !mxE. Qed.
+
+Lemma map_row' i0 : (row' i0 A)^f = row' i0 A^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_col' j0 : (col' j0 A)^f = col' j0 A^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_row_perm s : (row_perm s A)^f = row_perm s A^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_col_perm s : (col_perm s A)^f = col_perm s A^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_xrow i1 i2 : (xrow i1 i2 A)^f = xrow i1 i2 A^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_xcol j1 j2 : (xcol j1 j2 A)^f = xcol j1 j2 A^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_castmx m' n' c : (castmx c A)^f = castmx c A^f :> 'M_(m', n').
+Proof. by apply/matrixP=> i j; rewrite !(castmxE, mxE). Qed.
+
+Lemma map_conform_mx m' n' (B : 'M_(m', n')) :
+ (conform_mx B A)^f = conform_mx B^f A^f.
+Proof.
+move: B; have [[<- <-] B|] := eqVneq (m, n) (m', n').
+ by rewrite !conform_mx_id.
+by rewrite negb_and => neq_mn B; rewrite !nonconform_mx.
+Qed.
+
+Lemma map_mxvec : (mxvec A)^f = mxvec A^f.
+Proof. by apply/rowP=> i; rewrite !(castmxE, mxE). Qed.
+
+Lemma map_vec_mx (v : 'rV_(m * n)) : (vec_mx v)^f = vec_mx v^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+End OneMatrix.
+
+Section Block.
+
+Variables m1 m2 n1 n2 : nat.
+Variables (Aul : 'M[aT]_(m1, n1)) (Aur : 'M[aT]_(m1, n2)).
+Variables (Adl : 'M[aT]_(m2, n1)) (Adr : 'M[aT]_(m2, n2)).
+Variables (Bh : 'M[aT]_(m1, n1 + n2)) (Bv : 'M[aT]_(m1 + m2, n1)).
+Variable B : 'M[aT]_(m1 + m2, n1 + n2).
+
+Lemma map_row_mx : (row_mx Aul Aur)^f = row_mx Aul^f Aur^f.
+Proof. by apply/matrixP=> i j; do 2![rewrite !mxE //; case: split => ?]. Qed.
+
+Lemma map_col_mx : (col_mx Aul Adl)^f = col_mx Aul^f Adl^f.
+Proof. by apply/matrixP=> i j; do 2![rewrite !mxE //; case: split => ?]. Qed.
+
+Lemma map_block_mx :
+ (block_mx Aul Aur Adl Adr)^f = block_mx Aul^f Aur^f Adl^f Adr^f.
+Proof. by apply/matrixP=> i j; do 3![rewrite !mxE //; case: split => ?]. Qed.
+
+Lemma map_lsubmx : (lsubmx Bh)^f = lsubmx Bh^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_rsubmx : (rsubmx Bh)^f = rsubmx Bh^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_usubmx : (usubmx Bv)^f = usubmx Bv^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_dsubmx : (dsubmx Bv)^f = dsubmx Bv^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_ulsubmx : (ulsubmx B)^f = ulsubmx B^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_ursubmx : (ursubmx B)^f = ursubmx B^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_dlsubmx : (dlsubmx B)^f = dlsubmx B^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma map_drsubmx : (drsubmx B)^f = drsubmx B^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+End Block.
+
+End MapMatrix.
+
+(*****************************************************************************)
+(********************* Matrix Zmodule (additive) structure *******************)
+(*****************************************************************************)
+
+Section MatrixZmodule.
+
+Variable V : zmodType.
+
+Section FixedDim.
+
+Variables m n : nat.
+Implicit Types A B : 'M[V]_(m, n).
+
+Fact oppmx_key : unit. Proof. by []. Qed.
+Fact addmx_key : unit. Proof. by []. Qed.
+Definition oppmx A := \matrix[oppmx_key]_(i, j) (- A i j).
+Definition addmx A B := \matrix[addmx_key]_(i, j) (A i j + B i j).
+(* In principle, diag_mx and scalar_mx could be defined here, but since they *)
+(* only make sense with the graded ring operations, we defer them to the *)
+(* next section. *)
+
+Lemma addmxA : associative addmx.
+Proof. by move=> A B C; apply/matrixP=> i j; rewrite !mxE addrA. Qed.
+
+Lemma addmxC : commutative addmx.
+Proof. by move=> A B; apply/matrixP=> i j; rewrite !mxE addrC. Qed.
+
+Lemma add0mx : left_id (const_mx 0) addmx.
+Proof. by move=> A; apply/matrixP=> i j; rewrite !mxE add0r. Qed.
+
+Lemma addNmx : left_inverse (const_mx 0) oppmx addmx.
+Proof. by move=> A; apply/matrixP=> i j; rewrite !mxE addNr. Qed.
+
+Definition matrix_zmodMixin := ZmodMixin addmxA addmxC add0mx addNmx.
+
+Canonical matrix_zmodType := Eval hnf in ZmodType 'M[V]_(m, n) matrix_zmodMixin.
+
+Lemma mulmxnE A d i j : (A *+ d) i j = A i j *+ d.
+Proof. by elim: d => [|d IHd]; rewrite ?mulrS mxE ?IHd. Qed.
+
+Lemma summxE I r (P : pred I) (E : I -> 'M_(m, n)) i j :
+ (\sum_(k <- r | P k) E k) i j = \sum_(k <- r | P k) E k i j.
+Proof. by apply: (big_morph (fun A => A i j)) => [A B|]; rewrite mxE. Qed.
+
+Lemma const_mx_is_additive : additive const_mx.
+Proof. by move=> a b; apply/matrixP=> i j; rewrite !mxE. Qed.
+Canonical const_mx_additive := Additive const_mx_is_additive.
+
+End FixedDim.
+
+Section Additive.
+
+Variables (m n p q : nat) (f : 'I_p -> 'I_q -> 'I_m) (g : 'I_p -> 'I_q -> 'I_n).
+
+Definition swizzle_mx k (A : 'M[V]_(m, n)) :=
+ \matrix[k]_(i, j) A (f i j) (g i j).
+
+Lemma swizzle_mx_is_additive k : additive (swizzle_mx k).
+Proof. by move=> A B; apply/matrixP=> i j; rewrite !mxE. Qed.
+Canonical swizzle_mx_additive k := Additive (swizzle_mx_is_additive k).
+
+End Additive.
+
+Local Notation SwizzleAdd op := [additive of op as swizzle_mx _ _ _].
+
+Canonical trmx_additive m n := SwizzleAdd (@trmx V m n).
+Canonical row_additive m n i := SwizzleAdd (@row V m n i).
+Canonical col_additive m n j := SwizzleAdd (@col V m n j).
+Canonical row'_additive m n i := SwizzleAdd (@row' V m n i).
+Canonical col'_additive m n j := SwizzleAdd (@col' V m n j).
+Canonical row_perm_additive m n s := SwizzleAdd (@row_perm V m n s).
+Canonical col_perm_additive m n s := SwizzleAdd (@col_perm V m n s).
+Canonical xrow_additive m n i1 i2 := SwizzleAdd (@xrow V m n i1 i2).
+Canonical xcol_additive m n j1 j2 := SwizzleAdd (@xcol V m n j1 j2).
+Canonical lsubmx_additive m n1 n2 := SwizzleAdd (@lsubmx V m n1 n2).
+Canonical rsubmx_additive m n1 n2 := SwizzleAdd (@rsubmx V m n1 n2).
+Canonical usubmx_additive m1 m2 n := SwizzleAdd (@usubmx V m1 m2 n).
+Canonical dsubmx_additive m1 m2 n := SwizzleAdd (@dsubmx V m1 m2 n).
+Canonical vec_mx_additive m n := SwizzleAdd (@vec_mx V m n).
+Canonical mxvec_additive m n :=
+ Additive (can2_additive (@vec_mxK V m n) mxvecK).
+
+Lemma flatmx0 n : all_equal_to (0 : 'M_(0, n)).
+Proof. by move=> A; apply/matrixP=> [] []. Qed.
+
+Lemma thinmx0 n : all_equal_to (0 : 'M_(n, 0)).
+Proof. by move=> A; apply/matrixP=> i []. Qed.
+
+Lemma trmx0 m n : (0 : 'M_(m, n))^T = 0.
+Proof. exact: trmx_const. Qed.
+
+Lemma row0 m n i0 : row i0 (0 : 'M_(m, n)) = 0.
+Proof. exact: row_const. Qed.
+
+Lemma col0 m n j0 : col j0 (0 : 'M_(m, n)) = 0.
+Proof. exact: col_const. Qed.
+
+Lemma mxvec_eq0 m n (A : 'M_(m, n)) : (mxvec A == 0) = (A == 0).
+Proof. by rewrite (can2_eq mxvecK vec_mxK) raddf0. Qed.
+
+Lemma vec_mx_eq0 m n (v : 'rV_(m * n)) : (vec_mx v == 0) = (v == 0).
+Proof. by rewrite (can2_eq vec_mxK mxvecK) raddf0. Qed.
+
+Lemma row_mx0 m n1 n2 : row_mx 0 0 = 0 :> 'M_(m, n1 + n2).
+Proof. exact: row_mx_const. Qed.
+
+Lemma col_mx0 m1 m2 n : col_mx 0 0 = 0 :> 'M_(m1 + m2, n).
+Proof. exact: col_mx_const. Qed.
+
+Lemma block_mx0 m1 m2 n1 n2 : block_mx 0 0 0 0 = 0 :> 'M_(m1 + m2, n1 + n2).
+Proof. exact: block_mx_const. Qed.
+
+Ltac split_mxE := apply/matrixP=> i j; do ![rewrite mxE | case: split => ?].
+
+Lemma opp_row_mx m n1 n2 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) :
+ - row_mx A1 A2 = row_mx (- A1) (- A2).
+Proof. by split_mxE. Qed.
+
+Lemma opp_col_mx m1 m2 n (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) :
+ - col_mx A1 A2 = col_mx (- A1) (- A2).
+Proof. by split_mxE. Qed.
+
+Lemma opp_block_mx m1 m2 n1 n2 (Aul : 'M_(m1, n1)) Aur Adl (Adr : 'M_(m2, n2)) :
+ - block_mx Aul Aur Adl Adr = block_mx (- Aul) (- Aur) (- Adl) (- Adr).
+Proof. by rewrite opp_col_mx !opp_row_mx. Qed.
+
+Lemma add_row_mx m n1 n2 (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) B1 B2 :
+ row_mx A1 A2 + row_mx B1 B2 = row_mx (A1 + B1) (A2 + B2).
+Proof. by split_mxE. Qed.
+
+Lemma add_col_mx m1 m2 n (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) B1 B2 :
+ col_mx A1 A2 + col_mx B1 B2 = col_mx (A1 + B1) (A2 + B2).
+Proof. by split_mxE. Qed.
+
+Lemma add_block_mx m1 m2 n1 n2 (Aul : 'M_(m1, n1)) Aur Adl (Adr : 'M_(m2, n2))
+ Bul Bur Bdl Bdr :
+ let A := block_mx Aul Aur Adl Adr in let B := block_mx Bul Bur Bdl Bdr in
+ A + B = block_mx (Aul + Bul) (Aur + Bur) (Adl + Bdl) (Adr + Bdr).
+Proof. by rewrite /= add_col_mx !add_row_mx. Qed.
+
+Definition nz_row m n (A : 'M_(m, n)) :=
+ oapp (fun i => row i A) 0 [pick i | row i A != 0].
+
+Lemma nz_row_eq0 m n (A : 'M_(m, n)) : (nz_row A == 0) = (A == 0).
+Proof.
+rewrite /nz_row; symmetry; case: pickP => [i /= nzAi | Ai0].
+ by rewrite (negbTE nzAi); apply: contraTF nzAi => /eqP->; rewrite row0 eqxx.
+by rewrite eqxx; apply/eqP/row_matrixP=> i; move/eqP: (Ai0 i) ->; rewrite row0.
+Qed.
+
+End MatrixZmodule.
+
+Section FinZmodMatrix.
+Variables (V : finZmodType) (m n : nat).
+Local Notation MV := 'M[V]_(m, n).
+
+Canonical matrix_finZmodType := Eval hnf in [finZmodType of MV].
+Canonical matrix_baseFinGroupType :=
+ Eval hnf in [baseFinGroupType of MV for +%R].
+Canonical matrix_finGroupType := Eval hnf in [finGroupType of MV for +%R].
+End FinZmodMatrix.
+
+(* Parametricity over the additive structure. *)
+Section MapZmodMatrix.
+
+Variables (aR rR : zmodType) (f : {additive aR -> rR}) (m n : nat).
+Local Notation "A ^f" := (map_mx f A) : ring_scope.
+Implicit Type A : 'M[aR]_(m, n).
+
+Lemma map_mx0 : 0^f = 0 :> 'M_(m, n).
+Proof. by rewrite map_const_mx raddf0. Qed.
+
+Lemma map_mxN A : (- A)^f = - A^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE raddfN. Qed.
+
+Lemma map_mxD A B : (A + B)^f = A^f + B^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE raddfD. Qed.
+
+Lemma map_mx_sub A B : (A - B)^f = A^f - B^f.
+Proof. by rewrite map_mxD map_mxN. Qed.
+
+Definition map_mx_sum := big_morph _ map_mxD map_mx0.
+
+Canonical map_mx_additive := Additive map_mx_sub.
+
+End MapZmodMatrix.
+
+(*****************************************************************************)
+(*********** Matrix ring module, graded ring, and ring structures ************)
+(*****************************************************************************)
+
+Section MatrixAlgebra.
+
+Variable R : ringType.
+
+Section RingModule.
+
+(* The ring module/vector space structure *)
+
+Variables m n : nat.
+Implicit Types A B : 'M[R]_(m, n).
+
+Fact scalemx_key : unit. Proof. by []. Qed.
+Definition scalemx x A := \matrix[scalemx_key]_(i, j) (x * A i j).
+
+(* Basis *)
+Fact delta_mx_key : unit. Proof. by []. Qed.
+Definition delta_mx i0 j0 : 'M[R]_(m, n) :=
+ \matrix[delta_mx_key]_(i, j) ((i == i0) && (j == j0))%:R.
+
+Local Notation "x *m: A" := (scalemx x A) (at level 40) : ring_scope.
+
+Lemma scale1mx A : 1 *m: A = A.
+Proof. by apply/matrixP=> i j; rewrite !mxE mul1r. Qed.
+
+Lemma scalemxDl A x y : (x + y) *m: A = x *m: A + y *m: A.
+Proof. by apply/matrixP=> i j; rewrite !mxE mulrDl. Qed.
+
+Lemma scalemxDr x A B : x *m: (A + B) = x *m: A + x *m: B.
+Proof. by apply/matrixP=> i j; rewrite !mxE mulrDr. Qed.
+
+Lemma scalemxA x y A : x *m: (y *m: A) = (x * y) *m: A.
+Proof. by apply/matrixP=> i j; rewrite !mxE mulrA. Qed.
+
+Definition matrix_lmodMixin :=
+ LmodMixin scalemxA scale1mx scalemxDr scalemxDl.
+
+Canonical matrix_lmodType :=
+ Eval hnf in LmodType R 'M[R]_(m, n) matrix_lmodMixin.
+
+Lemma scalemx_const a b : a *: const_mx b = const_mx (a * b).
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma matrix_sum_delta A :
+ A = \sum_(i < m) \sum_(j < n) A i j *: delta_mx i j.
+Proof.
+apply/matrixP=> i j.
+rewrite summxE (bigD1 i) // summxE (bigD1 j) //= !mxE !eqxx mulr1.
+rewrite !big1 ?addr0 //= => [i' | j']; rewrite eq_sym => /negbTE diff.
+ by rewrite summxE big1 // => j' _; rewrite !mxE diff mulr0.
+by rewrite !mxE eqxx diff mulr0.
+Qed.
+
+End RingModule.
+
+Section StructuralLinear.
+
+Lemma swizzle_mx_is_scalable m n p q f g k :
+ scalable (@swizzle_mx R m n p q f g k).
+Proof. by move=> a A; apply/matrixP=> i j; rewrite !mxE. Qed.
+Canonical swizzle_mx_scalable m n p q f g k :=
+ AddLinear (@swizzle_mx_is_scalable m n p q f g k).
+
+Local Notation SwizzleLin op := [linear of op as swizzle_mx _ _ _].
+
+Canonical trmx_linear m n := SwizzleLin (@trmx R m n).
+Canonical row_linear m n i := SwizzleLin (@row R m n i).
+Canonical col_linear m n j := SwizzleLin (@col R m n j).
+Canonical row'_linear m n i := SwizzleLin (@row' R m n i).
+Canonical col'_linear m n j := SwizzleLin (@col' R m n j).
+Canonical row_perm_linear m n s := SwizzleLin (@row_perm R m n s).
+Canonical col_perm_linear m n s := SwizzleLin (@col_perm R m n s).
+Canonical xrow_linear m n i1 i2 := SwizzleLin (@xrow R m n i1 i2).
+Canonical xcol_linear m n j1 j2 := SwizzleLin (@xcol R m n j1 j2).
+Canonical lsubmx_linear m n1 n2 := SwizzleLin (@lsubmx R m n1 n2).
+Canonical rsubmx_linear m n1 n2 := SwizzleLin (@rsubmx R m n1 n2).
+Canonical usubmx_linear m1 m2 n := SwizzleLin (@usubmx R m1 m2 n).
+Canonical dsubmx_linear m1 m2 n := SwizzleLin (@dsubmx R m1 m2 n).
+Canonical vec_mx_linear m n := SwizzleLin (@vec_mx R m n).
+Definition mxvec_is_linear m n := can2_linear (@vec_mxK R m n) mxvecK.
+Canonical mxvec_linear m n := AddLinear (@mxvec_is_linear m n).
+
+End StructuralLinear.
+
+Lemma trmx_delta m n i j : (delta_mx i j)^T = delta_mx j i :> 'M[R]_(n, m).
+Proof. by apply/matrixP=> i' j'; rewrite !mxE andbC. Qed.
+
+Lemma row_sum_delta n (u : 'rV_n) : u = \sum_(j < n) u 0 j *: delta_mx 0 j.
+Proof. by rewrite {1}[u]matrix_sum_delta big_ord1. Qed.
+
+Lemma delta_mx_lshift m n1 n2 i j :
+ delta_mx i (lshift n2 j) = row_mx (delta_mx i j) 0 :> 'M_(m, n1 + n2).
+Proof.
+apply/matrixP=> i' j'; rewrite !mxE -(can_eq (@splitK _ _)).
+by rewrite (unsplitK (inl _ _)); case: split => ?; rewrite mxE ?andbF.
+Qed.
+
+Lemma delta_mx_rshift m n1 n2 i j :
+ delta_mx i (rshift n1 j) = row_mx 0 (delta_mx i j) :> 'M_(m, n1 + n2).
+Proof.
+apply/matrixP=> i' j'; rewrite !mxE -(can_eq (@splitK _ _)).
+by rewrite (unsplitK (inr _ _)); case: split => ?; rewrite mxE ?andbF.
+Qed.
+
+Lemma delta_mx_ushift m1 m2 n i j :
+ delta_mx (lshift m2 i) j = col_mx (delta_mx i j) 0 :> 'M_(m1 + m2, n).
+Proof.
+apply/matrixP=> i' j'; rewrite !mxE -(can_eq (@splitK _ _)).
+by rewrite (unsplitK (inl _ _)); case: split => ?; rewrite mxE.
+Qed.
+
+Lemma delta_mx_dshift m1 m2 n i j :
+ delta_mx (rshift m1 i) j = col_mx 0 (delta_mx i j) :> 'M_(m1 + m2, n).
+Proof.
+apply/matrixP=> i' j'; rewrite !mxE -(can_eq (@splitK _ _)).
+by rewrite (unsplitK (inr _ _)); case: split => ?; rewrite mxE.
+Qed.
+
+Lemma vec_mx_delta m n i j :
+ vec_mx (delta_mx 0 (mxvec_index i j)) = delta_mx i j :> 'M_(m, n).
+Proof.
+by apply/matrixP=> i' j'; rewrite !mxE /= [_ == _](inj_eq enum_rank_inj).
+Qed.
+
+Lemma mxvec_delta m n i j :
+ mxvec (delta_mx i j) = delta_mx 0 (mxvec_index i j) :> 'rV_(m * n).
+Proof. by rewrite -vec_mx_delta vec_mxK. Qed.
+
+Ltac split_mxE := apply/matrixP=> i j; do ![rewrite mxE | case: split => ?].
+
+Lemma scale_row_mx m n1 n2 a (A1 : 'M_(m, n1)) (A2 : 'M_(m, n2)) :
+ a *: row_mx A1 A2 = row_mx (a *: A1) (a *: A2).
+Proof. by split_mxE. Qed.
+
+Lemma scale_col_mx m1 m2 n a (A1 : 'M_(m1, n)) (A2 : 'M_(m2, n)) :
+ a *: col_mx A1 A2 = col_mx (a *: A1) (a *: A2).
+Proof. by split_mxE. Qed.
+
+Lemma scale_block_mx m1 m2 n1 n2 a (Aul : 'M_(m1, n1)) (Aur : 'M_(m1, n2))
+ (Adl : 'M_(m2, n1)) (Adr : 'M_(m2, n2)) :
+ a *: block_mx Aul Aur Adl Adr
+ = block_mx (a *: Aul) (a *: Aur) (a *: Adl) (a *: Adr).
+Proof. by rewrite scale_col_mx !scale_row_mx. Qed.
+
+(* Diagonal matrices *)
+
+Fact diag_mx_key : unit. Proof. by []. Qed.
+Definition diag_mx n (d : 'rV[R]_n) :=
+ \matrix[diag_mx_key]_(i, j) (d 0 i *+ (i == j)).
+
+Lemma tr_diag_mx n (d : 'rV_n) : (diag_mx d)^T = diag_mx d.
+Proof. by apply/matrixP=> i j; rewrite !mxE eq_sym; case: eqP => // ->. Qed.
+
+Lemma diag_mx_is_linear n : linear (@diag_mx n).
+Proof.
+by move=> a A B; apply/matrixP=> i j; rewrite !mxE mulrnAr mulrnDl.
+Qed.
+Canonical diag_mx_additive n := Additive (@diag_mx_is_linear n).
+Canonical diag_mx_linear n := Linear (@diag_mx_is_linear n).
+
+Lemma diag_mx_sum_delta n (d : 'rV_n) :
+ diag_mx d = \sum_i d 0 i *: delta_mx i i.
+Proof.
+apply/matrixP=> i j; rewrite summxE (bigD1 i) //= !mxE eqxx /=.
+rewrite eq_sym mulr_natr big1 ?addr0 // => i' ne_i'i.
+by rewrite !mxE eq_sym (negbTE ne_i'i) mulr0.
+Qed.
+
+(* Scalar matrix : a diagonal matrix with a constant on the diagonal *)
+Section ScalarMx.
+
+Variable n : nat.
+
+Fact scalar_mx_key : unit. Proof. by []. Qed.
+Definition scalar_mx x : 'M[R]_n :=
+ \matrix[scalar_mx_key]_(i , j) (x *+ (i == j)).
+Notation "x %:M" := (scalar_mx x) : ring_scope.
+
+Lemma diag_const_mx a : diag_mx (const_mx a) = a%:M :> 'M_n.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma tr_scalar_mx a : (a%:M)^T = a%:M.
+Proof. by apply/matrixP=> i j; rewrite !mxE eq_sym. Qed.
+
+Lemma trmx1 : (1%:M)^T = 1%:M. Proof. exact: tr_scalar_mx. Qed.
+
+Lemma scalar_mx_is_additive : additive scalar_mx.
+Proof. by move=> a b; rewrite -!diag_const_mx !raddfB. Qed.
+Canonical scalar_mx_additive := Additive scalar_mx_is_additive.
+
+Lemma scale_scalar_mx a1 a2 : a1 *: a2%:M = (a1 * a2)%:M :> 'M_n.
+Proof. by apply/matrixP=> i j; rewrite !mxE mulrnAr. Qed.
+
+Lemma scalemx1 a : a *: 1%:M = a%:M.
+Proof. by rewrite scale_scalar_mx mulr1. Qed.
+
+Lemma scalar_mx_sum_delta a : a%:M = \sum_i a *: delta_mx i i.
+Proof.
+by rewrite -diag_const_mx diag_mx_sum_delta; apply: eq_bigr => i _; rewrite mxE.
+Qed.
+
+Lemma mx1_sum_delta : 1%:M = \sum_i delta_mx i i.
+Proof. by rewrite [1%:M]scalar_mx_sum_delta -scaler_sumr scale1r. Qed.
+
+Lemma row1 i : row i 1%:M = delta_mx 0 i.
+Proof. by apply/rowP=> j; rewrite !mxE eq_sym. Qed.
+
+Definition is_scalar_mx (A : 'M[R]_n) :=
+ if insub 0%N is Some i then A == (A i i)%:M else true.
+
+Lemma is_scalar_mxP A : reflect (exists a, A = a%:M) (is_scalar_mx A).
+Proof.
+rewrite /is_scalar_mx; case: insubP => [i _ _ | ].
+ by apply: (iffP eqP) => [|[a ->]]; [exists (A i i) | rewrite mxE eqxx].
+rewrite -eqn0Ngt => /eqP n0; left; exists 0.
+by rewrite raddf0; rewrite n0 in A *; rewrite [A]flatmx0.
+Qed.
+
+Lemma scalar_mx_is_scalar a : is_scalar_mx a%:M.
+Proof. by apply/is_scalar_mxP; exists a. Qed.
+
+Lemma mx0_is_scalar : is_scalar_mx 0.
+Proof. by apply/is_scalar_mxP; exists 0; rewrite raddf0. Qed.
+
+End ScalarMx.
+
+Notation "x %:M" := (scalar_mx _ x) : ring_scope.
+
+Lemma mx11_scalar (A : 'M_1) : A = (A 0 0)%:M.
+Proof. by apply/rowP=> j; rewrite ord1 mxE. Qed.
+
+Lemma scalar_mx_block n1 n2 a : a%:M = block_mx a%:M 0 0 a%:M :> 'M_(n1 + n2).
+Proof.
+apply/matrixP=> i j; rewrite !mxE -val_eqE /=.
+by do 2![case: splitP => ? ->; rewrite !mxE];
+ rewrite ?eqn_add2l // -?(eq_sym (n1 + _)%N) eqn_leq leqNgt lshift_subproof.
+Qed.
+
+(* Matrix multiplication using bigops. *)
+Fact mulmx_key : unit. Proof. by []. Qed.
+Definition mulmx {m n p} (A : 'M_(m, n)) (B : 'M_(n, p)) : 'M[R]_(m, p) :=
+ \matrix[mulmx_key]_(i, k) \sum_j (A i j * B j k).
+
+Local Notation "A *m B" := (mulmx A B) : ring_scope.
+
+Lemma mulmxA m n p q (A : 'M_(m, n)) (B : 'M_(n, p)) (C : 'M_(p, q)) :
+ A *m (B *m C) = A *m B *m C.
+Proof.
+apply/matrixP=> i l; rewrite !mxE.
+transitivity (\sum_j (\sum_k (A i j * (B j k * C k l)))).
+ by apply: eq_bigr => j _; rewrite mxE big_distrr.
+rewrite exchange_big; apply: eq_bigr => j _; rewrite mxE big_distrl /=.
+by apply: eq_bigr => k _; rewrite mulrA.
+Qed.
+
+Lemma mul0mx m n p (A : 'M_(n, p)) : 0 *m A = 0 :> 'M_(m, p).
+Proof.
+by apply/matrixP=> i k; rewrite !mxE big1 //= => j _; rewrite mxE mul0r.
+Qed.
+
+Lemma mulmx0 m n p (A : 'M_(m, n)) : A *m 0 = 0 :> 'M_(m, p).
+Proof.
+by apply/matrixP=> i k; rewrite !mxE big1 // => j _; rewrite mxE mulr0.
+Qed.
+
+Lemma mulmxN m n p (A : 'M_(m, n)) (B : 'M_(n, p)) : A *m (- B) = - (A *m B).
+Proof.
+apply/matrixP=> i k; rewrite !mxE -sumrN.
+by apply: eq_bigr => j _; rewrite mxE mulrN.
+Qed.
+
+Lemma mulNmx m n p (A : 'M_(m, n)) (B : 'M_(n, p)) : - A *m B = - (A *m B).
+Proof.
+apply/matrixP=> i k; rewrite !mxE -sumrN.
+by apply: eq_bigr => j _; rewrite mxE mulNr.
+Qed.
+
+Lemma mulmxDl m n p (A1 A2 : 'M_(m, n)) (B : 'M_(n, p)) :
+ (A1 + A2) *m B = A1 *m B + A2 *m B.
+Proof.
+apply/matrixP=> i k; rewrite !mxE -big_split /=.
+by apply: eq_bigr => j _; rewrite !mxE -mulrDl.
+Qed.
+
+Lemma mulmxDr m n p (A : 'M_(m, n)) (B1 B2 : 'M_(n, p)) :
+ A *m (B1 + B2) = A *m B1 + A *m B2.
+Proof.
+apply/matrixP=> i k; rewrite !mxE -big_split /=.
+by apply: eq_bigr => j _; rewrite mxE mulrDr.
+Qed.
+
+Lemma mulmxBl m n p (A1 A2 : 'M_(m, n)) (B : 'M_(n, p)) :
+ (A1 - A2) *m B = A1 *m B - A2 *m B.
+Proof. by rewrite mulmxDl mulNmx. Qed.
+
+Lemma mulmxBr m n p (A : 'M_(m, n)) (B1 B2 : 'M_(n, p)) :
+ A *m (B1 - B2) = A *m B1 - A *m B2.
+Proof. by rewrite mulmxDr mulmxN. Qed.
+
+Lemma mulmx_suml m n p (A : 'M_(n, p)) I r P (B_ : I -> 'M_(m, n)) :
+ (\sum_(i <- r | P i) B_ i) *m A = \sum_(i <- r | P i) B_ i *m A.
+Proof.
+by apply: (big_morph (mulmx^~ A)) => [B C|]; rewrite ?mul0mx ?mulmxDl.
+Qed.
+
+Lemma mulmx_sumr m n p (A : 'M_(m, n)) I r P (B_ : I -> 'M_(n, p)) :
+ A *m (\sum_(i <- r | P i) B_ i) = \sum_(i <- r | P i) A *m B_ i.
+Proof.
+by apply: (big_morph (mulmx A)) => [B C|]; rewrite ?mulmx0 ?mulmxDr.
+Qed.
+
+Lemma scalemxAl m n p a (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ a *: (A *m B) = (a *: A) *m B.
+Proof.
+apply/matrixP=> i k; rewrite !mxE big_distrr /=.
+by apply: eq_bigr => j _; rewrite mulrA mxE.
+Qed.
+(* Right scaling associativity requires a commutative ring *)
+
+Lemma rowE m n i (A : 'M_(m, n)) : row i A = delta_mx 0 i *m A.
+Proof.
+apply/rowP=> j; rewrite !mxE (bigD1 i) //= mxE !eqxx mul1r.
+by rewrite big1 ?addr0 // => i' ne_i'i; rewrite mxE /= (negbTE ne_i'i) mul0r.
+Qed.
+
+Lemma row_mul m n p (i : 'I_m) A (B : 'M_(n, p)) :
+ row i (A *m B) = row i A *m B.
+Proof. by rewrite !rowE mulmxA. Qed.
+
+Lemma mulmx_sum_row m n (u : 'rV_m) (A : 'M_(m, n)) :
+ u *m A = \sum_i u 0 i *: row i A.
+Proof.
+by apply/rowP=> j; rewrite mxE summxE; apply: eq_bigr => i _; rewrite !mxE.
+Qed.
+
+Lemma mul_delta_mx_cond m n p (j1 j2 : 'I_n) (i1 : 'I_m) (k2 : 'I_p) :
+ delta_mx i1 j1 *m delta_mx j2 k2 = delta_mx i1 k2 *+ (j1 == j2).
+Proof.
+apply/matrixP=> i k; rewrite !mxE (bigD1 j1) //=.
+rewrite mulmxnE !mxE !eqxx andbT -natrM -mulrnA !mulnb !andbA andbAC.
+by rewrite big1 ?addr0 // => j; rewrite !mxE andbC -natrM; move/negbTE->.
+Qed.
+
+Lemma mul_delta_mx m n p (j : 'I_n) (i : 'I_m) (k : 'I_p) :
+ delta_mx i j *m delta_mx j k = delta_mx i k.
+Proof. by rewrite mul_delta_mx_cond eqxx. Qed.
+
+Lemma mul_delta_mx_0 m n p (j1 j2 : 'I_n) (i1 : 'I_m) (k2 : 'I_p) :
+ j1 != j2 -> delta_mx i1 j1 *m delta_mx j2 k2 = 0.
+Proof. by rewrite mul_delta_mx_cond => /negbTE->. Qed.
+
+Lemma mul_diag_mx m n d (A : 'M_(m, n)) :
+ diag_mx d *m A = \matrix_(i, j) (d 0 i * A i j).
+Proof.
+apply/matrixP=> i j; rewrite !mxE (bigD1 i) //= mxE eqxx big1 ?addr0 // => i'.
+by rewrite mxE eq_sym mulrnAl => /negbTE->.
+Qed.
+
+Lemma mul_mx_diag m n (A : 'M_(m, n)) d :
+ A *m diag_mx d = \matrix_(i, j) (A i j * d 0 j).
+Proof.
+apply/matrixP=> i j; rewrite !mxE (bigD1 j) //= mxE eqxx big1 ?addr0 // => i'.
+by rewrite mxE eq_sym mulrnAr; move/negbTE->.
+Qed.
+
+Lemma mulmx_diag n (d e : 'rV_n) :
+ diag_mx d *m diag_mx e = diag_mx (\row_j (d 0 j * e 0 j)).
+Proof. by apply/matrixP=> i j; rewrite mul_diag_mx !mxE mulrnAr. Qed.
+
+Lemma mul_scalar_mx m n a (A : 'M_(m, n)) : a%:M *m A = a *: A.
+Proof.
+by rewrite -diag_const_mx mul_diag_mx; apply/matrixP=> i j; rewrite !mxE.
+Qed.
+
+Lemma scalar_mxM n a b : (a * b)%:M = a%:M *m b%:M :> 'M_n.
+Proof. by rewrite mul_scalar_mx scale_scalar_mx. Qed.
+
+Lemma mul1mx m n (A : 'M_(m, n)) : 1%:M *m A = A.
+Proof. by rewrite mul_scalar_mx scale1r. Qed.
+
+Lemma mulmx1 m n (A : 'M_(m, n)) : A *m 1%:M = A.
+Proof.
+rewrite -diag_const_mx mul_mx_diag.
+by apply/matrixP=> i j; rewrite !mxE mulr1.
+Qed.
+
+Lemma mul_col_perm m n p s (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ col_perm s A *m B = A *m row_perm s^-1 B.
+Proof.
+apply/matrixP=> i k; rewrite !mxE (reindex_inj (@perm_inj _ s^-1)).
+by apply: eq_bigr => j _ /=; rewrite !mxE permKV.
+Qed.
+
+Lemma mul_row_perm m n p s (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ A *m row_perm s B = col_perm s^-1 A *m B.
+Proof. by rewrite mul_col_perm invgK. Qed.
+
+Lemma mul_xcol m n p j1 j2 (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ xcol j1 j2 A *m B = A *m xrow j1 j2 B.
+Proof. by rewrite mul_col_perm tpermV. Qed.
+
+(* Permutation matrix *)
+
+Definition perm_mx n s : 'M_n := row_perm s 1%:M.
+
+Definition tperm_mx n i1 i2 : 'M_n := perm_mx (tperm i1 i2).
+
+Lemma col_permE m n s (A : 'M_(m, n)) : col_perm s A = A *m perm_mx s^-1.
+Proof. by rewrite mul_row_perm mulmx1 invgK. Qed.
+
+Lemma row_permE m n s (A : 'M_(m, n)) : row_perm s A = perm_mx s *m A.
+Proof.
+by rewrite -[perm_mx _]mul1mx mul_row_perm mulmx1 -mul_row_perm mul1mx.
+Qed.
+
+Lemma xcolE m n j1 j2 (A : 'M_(m, n)) : xcol j1 j2 A = A *m tperm_mx j1 j2.
+Proof. by rewrite /xcol col_permE tpermV. Qed.
+
+Lemma xrowE m n i1 i2 (A : 'M_(m, n)) : xrow i1 i2 A = tperm_mx i1 i2 *m A.
+Proof. exact: row_permE. Qed.
+
+Lemma tr_perm_mx n (s : 'S_n) : (perm_mx s)^T = perm_mx s^-1.
+Proof. by rewrite -[_^T]mulmx1 tr_row_perm mul_col_perm trmx1 mul1mx. Qed.
+
+Lemma tr_tperm_mx n i1 i2 : (tperm_mx i1 i2)^T = tperm_mx i1 i2 :> 'M_n.
+Proof. by rewrite tr_perm_mx tpermV. Qed.
+
+Lemma perm_mx1 n : perm_mx 1 = 1%:M :> 'M_n.
+Proof. exact: row_perm1. Qed.
+
+Lemma perm_mxM n (s t : 'S_n) : perm_mx (s * t) = perm_mx s *m perm_mx t.
+Proof. by rewrite -row_permE -row_permM. Qed.
+
+Definition is_perm_mx n (A : 'M_n) := [exists s, A == perm_mx s].
+
+Lemma is_perm_mxP n (A : 'M_n) :
+ reflect (exists s, A = perm_mx s) (is_perm_mx A).
+Proof. by apply: (iffP existsP) => [] [s /eqP]; exists s. Qed.
+
+Lemma perm_mx_is_perm n (s : 'S_n) : is_perm_mx (perm_mx s).
+Proof. by apply/is_perm_mxP; exists s. Qed.
+
+Lemma is_perm_mx1 n : is_perm_mx (1%:M : 'M_n).
+Proof. by rewrite -perm_mx1 perm_mx_is_perm. Qed.
+
+Lemma is_perm_mxMl n (A B : 'M_n) :
+ is_perm_mx A -> is_perm_mx (A *m B) = is_perm_mx B.
+Proof.
+case/is_perm_mxP=> s ->.
+apply/is_perm_mxP/is_perm_mxP=> [[t def_t] | [t ->]]; last first.
+ by exists (s * t)%g; rewrite perm_mxM.
+exists (s^-1 * t)%g.
+by rewrite perm_mxM -def_t -!row_permE -row_permM mulVg row_perm1.
+Qed.
+
+Lemma is_perm_mx_tr n (A : 'M_n) : is_perm_mx A^T = is_perm_mx A.
+Proof.
+apply/is_perm_mxP/is_perm_mxP=> [[t def_t] | [t ->]]; exists t^-1%g.
+ by rewrite -tr_perm_mx -def_t trmxK.
+by rewrite tr_perm_mx.
+Qed.
+
+Lemma is_perm_mxMr n (A B : 'M_n) :
+ is_perm_mx B -> is_perm_mx (A *m B) = is_perm_mx A.
+Proof.
+case/is_perm_mxP=> s ->.
+rewrite -[s]invgK -col_permE -is_perm_mx_tr tr_col_perm row_permE.
+by rewrite is_perm_mxMl (perm_mx_is_perm, is_perm_mx_tr).
+Qed.
+
+(* Partial identity matrix (used in rank decomposition). *)
+
+Fact pid_mx_key : unit. Proof. by []. Qed.
+Definition pid_mx {m n} r : 'M[R]_(m, n) :=
+ \matrix[pid_mx_key]_(i, j) ((i == j :> nat) && (i < r))%:R.
+
+Lemma pid_mx_0 m n : pid_mx 0 = 0 :> 'M_(m, n).
+Proof. by apply/matrixP=> i j; rewrite !mxE andbF. Qed.
+
+Lemma pid_mx_1 r : pid_mx r = 1%:M :> 'M_r.
+Proof. by apply/matrixP=> i j; rewrite !mxE ltn_ord andbT. Qed.
+
+Lemma pid_mx_row n r : pid_mx r = row_mx 1%:M 0 :> 'M_(r, r + n).
+Proof.
+apply/matrixP=> i j; rewrite !mxE ltn_ord andbT.
+case: splitP => j' ->; rewrite !mxE // .
+by rewrite eqn_leq andbC leqNgt lshift_subproof.
+Qed.
+
+Lemma pid_mx_col m r : pid_mx r = col_mx 1%:M 0 :> 'M_(r + m, r).
+Proof.
+apply/matrixP=> i j; rewrite !mxE andbC.
+by case: splitP => i' ->; rewrite !mxE // eq_sym.
+Qed.
+
+Lemma pid_mx_block m n r : pid_mx r = block_mx 1%:M 0 0 0 :> 'M_(r + m, r + n).
+Proof.
+apply/matrixP=> i j; rewrite !mxE row_mx0 andbC.
+case: splitP => i' ->; rewrite !mxE //; case: splitP => j' ->; rewrite !mxE //=.
+by rewrite eqn_leq andbC leqNgt lshift_subproof.
+Qed.
+
+Lemma tr_pid_mx m n r : (pid_mx r)^T = pid_mx r :> 'M_(n, m).
+Proof. by apply/matrixP=> i j; rewrite !mxE eq_sym; case: eqP => // ->. Qed.
+
+Lemma pid_mx_minv m n r : pid_mx (minn m r) = pid_mx r :> 'M_(m, n).
+Proof. by apply/matrixP=> i j; rewrite !mxE leq_min ltn_ord. Qed.
+
+Lemma pid_mx_minh m n r : pid_mx (minn n r) = pid_mx r :> 'M_(m, n).
+Proof. by apply: trmx_inj; rewrite !tr_pid_mx pid_mx_minv. Qed.
+
+Lemma mul_pid_mx m n p q r :
+ (pid_mx q : 'M_(m, n)) *m (pid_mx r : 'M_(n, p)) = pid_mx (minn n (minn q r)).
+Proof.
+apply/matrixP=> i k; rewrite !mxE !leq_min.
+have [le_n_i | lt_i_n] := leqP n i.
+ rewrite andbF big1 // => j _.
+ by rewrite -pid_mx_minh !mxE leq_min ltnNge le_n_i andbF mul0r.
+rewrite (bigD1 (Ordinal lt_i_n)) //= big1 ?addr0 => [|j].
+ by rewrite !mxE eqxx /= -natrM mulnb andbCA.
+by rewrite -val_eqE /= !mxE eq_sym -natrM => /negbTE->.
+Qed.
+
+Lemma pid_mx_id m n p r :
+ r <= n -> (pid_mx r : 'M_(m, n)) *m (pid_mx r : 'M_(n, p)) = pid_mx r.
+Proof. by move=> le_r_n; rewrite mul_pid_mx minnn (minn_idPr _). Qed.
+
+Definition copid_mx {n} r : 'M_n := 1%:M - pid_mx r.
+
+Lemma mul_copid_mx_pid m n r :
+ r <= m -> copid_mx r *m pid_mx r = 0 :> 'M_(m, n).
+Proof. by move=> le_r_m; rewrite mulmxBl mul1mx pid_mx_id ?subrr. Qed.
+
+Lemma mul_pid_mx_copid m n r :
+ r <= n -> pid_mx r *m copid_mx r = 0 :> 'M_(m, n).
+Proof. by move=> le_r_n; rewrite mulmxBr mulmx1 pid_mx_id ?subrr. Qed.
+
+Lemma copid_mx_id n r :
+ r <= n -> copid_mx r *m copid_mx r = copid_mx r :> 'M_n.
+Proof.
+by move=> le_r_n; rewrite mulmxBl mul1mx mul_pid_mx_copid // oppr0 addr0.
+Qed.
+
+(* Block products; we cover all 1 x 2, 2 x 1, and 2 x 2 block products. *)
+Lemma mul_mx_row m n p1 p2 (A : 'M_(m, n)) (Bl : 'M_(n, p1)) (Br : 'M_(n, p2)) :
+ A *m row_mx Bl Br = row_mx (A *m Bl) (A *m Br).
+Proof.
+apply/matrixP=> i k; rewrite !mxE.
+by case defk: (split k); rewrite mxE; apply: eq_bigr => j _; rewrite mxE defk.
+Qed.
+
+Lemma mul_col_mx m1 m2 n p (Au : 'M_(m1, n)) (Ad : 'M_(m2, n)) (B : 'M_(n, p)) :
+ col_mx Au Ad *m B = col_mx (Au *m B) (Ad *m B).
+Proof.
+apply/matrixP=> i k; rewrite !mxE.
+by case defi: (split i); rewrite mxE; apply: eq_bigr => j _; rewrite mxE defi.
+Qed.
+
+Lemma mul_row_col m n1 n2 p (Al : 'M_(m, n1)) (Ar : 'M_(m, n2))
+ (Bu : 'M_(n1, p)) (Bd : 'M_(n2, p)) :
+ row_mx Al Ar *m col_mx Bu Bd = Al *m Bu + Ar *m Bd.
+Proof.
+apply/matrixP=> i k; rewrite !mxE big_split_ord /=.
+congr (_ + _); apply: eq_bigr => j _; first by rewrite row_mxEl col_mxEu.
+by rewrite row_mxEr col_mxEd.
+Qed.
+
+Lemma mul_col_row m1 m2 n p1 p2 (Au : 'M_(m1, n)) (Ad : 'M_(m2, n))
+ (Bl : 'M_(n, p1)) (Br : 'M_(n, p2)) :
+ col_mx Au Ad *m row_mx Bl Br
+ = block_mx (Au *m Bl) (Au *m Br) (Ad *m Bl) (Ad *m Br).
+Proof. by rewrite mul_col_mx !mul_mx_row. Qed.
+
+Lemma mul_row_block m n1 n2 p1 p2 (Al : 'M_(m, n1)) (Ar : 'M_(m, n2))
+ (Bul : 'M_(n1, p1)) (Bur : 'M_(n1, p2))
+ (Bdl : 'M_(n2, p1)) (Bdr : 'M_(n2, p2)) :
+ row_mx Al Ar *m block_mx Bul Bur Bdl Bdr
+ = row_mx (Al *m Bul + Ar *m Bdl) (Al *m Bur + Ar *m Bdr).
+Proof. by rewrite block_mxEh mul_mx_row !mul_row_col. Qed.
+
+Lemma mul_block_col m1 m2 n1 n2 p (Aul : 'M_(m1, n1)) (Aur : 'M_(m1, n2))
+ (Adl : 'M_(m2, n1)) (Adr : 'M_(m2, n2))
+ (Bu : 'M_(n1, p)) (Bd : 'M_(n2, p)) :
+ block_mx Aul Aur Adl Adr *m col_mx Bu Bd
+ = col_mx (Aul *m Bu + Aur *m Bd) (Adl *m Bu + Adr *m Bd).
+Proof. by rewrite mul_col_mx !mul_row_col. Qed.
+
+Lemma mulmx_block m1 m2 n1 n2 p1 p2 (Aul : 'M_(m1, n1)) (Aur : 'M_(m1, n2))
+ (Adl : 'M_(m2, n1)) (Adr : 'M_(m2, n2))
+ (Bul : 'M_(n1, p1)) (Bur : 'M_(n1, p2))
+ (Bdl : 'M_(n2, p1)) (Bdr : 'M_(n2, p2)) :
+ block_mx Aul Aur Adl Adr *m block_mx Bul Bur Bdl Bdr
+ = block_mx (Aul *m Bul + Aur *m Bdl) (Aul *m Bur + Aur *m Bdr)
+ (Adl *m Bul + Adr *m Bdl) (Adl *m Bur + Adr *m Bdr).
+Proof. by rewrite mul_col_mx !mul_row_block. Qed.
+
+(* Correspondance between matrices and linear function on row vectors. *)
+Section LinRowVector.
+
+Variables m n : nat.
+
+Fact lin1_mx_key : unit. Proof. by []. Qed.
+Definition lin1_mx (f : 'rV[R]_m -> 'rV[R]_n) :=
+ \matrix[lin1_mx_key]_(i, j) f (delta_mx 0 i) 0 j.
+
+Variable f : {linear 'rV[R]_m -> 'rV[R]_n}.
+
+Lemma mul_rV_lin1 u : u *m lin1_mx f = f u.
+Proof.
+rewrite {2}[u]matrix_sum_delta big_ord1 linear_sum; apply/rowP=> i.
+by rewrite mxE summxE; apply: eq_bigr => j _; rewrite linearZ !mxE.
+Qed.
+
+End LinRowVector.
+
+(* Correspondance between matrices and linear function on matrices. *)
+Section LinMatrix.
+
+Variables m1 n1 m2 n2 : nat.
+
+Definition lin_mx (f : 'M[R]_(m1, n1) -> 'M[R]_(m2, n2)) :=
+ lin1_mx (mxvec \o f \o vec_mx).
+
+Variable f : {linear 'M[R]_(m1, n1) -> 'M[R]_(m2, n2)}.
+
+Lemma mul_rV_lin u : u *m lin_mx f = mxvec (f (vec_mx u)).
+Proof. exact: mul_rV_lin1. Qed.
+
+Lemma mul_vec_lin A : mxvec A *m lin_mx f = mxvec (f A).
+Proof. by rewrite mul_rV_lin mxvecK. Qed.
+
+Lemma mx_rV_lin u : vec_mx (u *m lin_mx f) = f (vec_mx u).
+Proof. by rewrite mul_rV_lin mxvecK. Qed.
+
+Lemma mx_vec_lin A : vec_mx (mxvec A *m lin_mx f) = f A.
+Proof. by rewrite mul_rV_lin !mxvecK. Qed.
+
+End LinMatrix.
+
+Canonical mulmx_additive m n p A := Additive (@mulmxBr m n p A).
+
+Section Mulmxr.
+
+Variables m n p : nat.
+Implicit Type A : 'M[R]_(m, n).
+Implicit Type B : 'M[R]_(n, p).
+
+Definition mulmxr_head t B A := let: tt := t in A *m B.
+Local Notation mulmxr := (mulmxr_head tt).
+
+Definition lin_mulmxr B := lin_mx (mulmxr B).
+
+Lemma mulmxr_is_linear B : linear (mulmxr B).
+Proof. by move=> a A1 A2; rewrite /= mulmxDl scalemxAl. Qed.
+Canonical mulmxr_additive B := Additive (mulmxr_is_linear B).
+Canonical mulmxr_linear B := Linear (mulmxr_is_linear B).
+
+Lemma lin_mulmxr_is_linear : linear lin_mulmxr.
+Proof.
+move=> a A B; apply/row_matrixP; case/mxvec_indexP=> i j.
+rewrite linearP /= !rowE !mul_rV_lin /= vec_mx_delta -linearP mulmxDr.
+congr (mxvec (_ + _)); apply/row_matrixP=> k.
+rewrite linearZ /= !row_mul rowE mul_delta_mx_cond.
+by case: (k == i); [rewrite -!rowE linearZ | rewrite !mul0mx raddf0].
+Qed.
+Canonical lin_mulmxr_additive := Additive lin_mulmxr_is_linear.
+Canonical lin_mulmxr_linear := Linear lin_mulmxr_is_linear.
+
+End Mulmxr.
+
+(* The trace. *)
+Section Trace.
+
+Variable n : nat.
+
+Definition mxtrace (A : 'M[R]_n) := \sum_i A i i.
+Local Notation "'\tr' A" := (mxtrace A) : ring_scope.
+
+Lemma mxtrace_tr A : \tr A^T = \tr A.
+Proof. by apply: eq_bigr=> i _; rewrite mxE. Qed.
+
+Lemma mxtrace_is_scalar : scalar mxtrace.
+Proof.
+move=> a A B; rewrite mulr_sumr -big_split /=; apply: eq_bigr=> i _.
+by rewrite !mxE.
+Qed.
+Canonical mxtrace_additive := Additive mxtrace_is_scalar.
+Canonical mxtrace_linear := Linear mxtrace_is_scalar.
+
+Lemma mxtrace0 : \tr 0 = 0. Proof. exact: raddf0. Qed.
+Lemma mxtraceD A B : \tr (A + B) = \tr A + \tr B. Proof. exact: raddfD. Qed.
+Lemma mxtraceZ a A : \tr (a *: A) = a * \tr A. Proof. exact: scalarZ. Qed.
+
+Lemma mxtrace_diag D : \tr (diag_mx D) = \sum_j D 0 j.
+Proof. by apply: eq_bigr => j _; rewrite mxE eqxx. Qed.
+
+Lemma mxtrace_scalar a : \tr a%:M = a *+ n.
+Proof.
+rewrite -diag_const_mx mxtrace_diag.
+by rewrite (eq_bigr _ (fun j _ => mxE _ _ 0 j)) sumr_const card_ord.
+Qed.
+
+Lemma mxtrace1 : \tr 1%:M = n%:R. Proof. exact: mxtrace_scalar. Qed.
+
+End Trace.
+Local Notation "'\tr' A" := (mxtrace A) : ring_scope.
+
+Lemma trace_mx11 (A : 'M_1) : \tr A = A 0 0.
+Proof. by rewrite {1}[A]mx11_scalar mxtrace_scalar. Qed.
+
+Lemma mxtrace_block n1 n2 (Aul : 'M_n1) Aur Adl (Adr : 'M_n2) :
+ \tr (block_mx Aul Aur Adl Adr) = \tr Aul + \tr Adr.
+Proof.
+rewrite /(\tr _) big_split_ord /=.
+by congr (_ + _); apply: eq_bigr => i _; rewrite (block_mxEul, block_mxEdr).
+Qed.
+
+(* The matrix ring structure requires a strutural condition (dimension of the *)
+(* form n.+1) to statisfy the nontriviality condition we have imposed. *)
+Section MatrixRing.
+
+Variable n' : nat.
+Local Notation n := n'.+1.
+
+Lemma matrix_nonzero1 : 1%:M != 0 :> 'M_n.
+Proof. by apply/eqP=> /matrixP/(_ 0 0)/eqP; rewrite !mxE oner_eq0. Qed.
+
+Definition matrix_ringMixin :=
+ RingMixin (@mulmxA n n n n) (@mul1mx n n) (@mulmx1 n n)
+ (@mulmxDl n n n) (@mulmxDr n n n) matrix_nonzero1.
+
+Canonical matrix_ringType := Eval hnf in RingType 'M[R]_n matrix_ringMixin.
+Canonical matrix_lAlgType := Eval hnf in LalgType R 'M[R]_n (@scalemxAl n n n).
+
+Lemma mulmxE : mulmx = *%R. Proof. by []. Qed.
+Lemma idmxE : 1%:M = 1 :> 'M_n. Proof. by []. Qed.
+
+Lemma scalar_mx_is_multiplicative : multiplicative (@scalar_mx n).
+Proof. by split=> //; exact: scalar_mxM. Qed.
+Canonical scalar_mx_rmorphism := AddRMorphism scalar_mx_is_multiplicative.
+
+End MatrixRing.
+
+Section LiftPerm.
+
+(* Block expresssion of a lifted permutation matrix, for the Cormen LUP. *)
+
+Variable n : nat.
+
+(* These could be in zmodp, but that would introduce a dependency on perm. *)
+
+Definition lift0_perm s : 'S_n.+1 := lift_perm 0 0 s.
+
+Lemma lift0_perm0 s : lift0_perm s 0 = 0.
+Proof. exact: lift_perm_id. Qed.
+
+Lemma lift0_perm_lift s k' :
+ lift0_perm s (lift 0 k') = lift (0 : 'I_n.+1) (s k').
+Proof. exact: lift_perm_lift. Qed.
+
+Lemma lift0_permK s : cancel (lift0_perm s) (lift0_perm s^-1).
+Proof. by move=> i; rewrite /lift0_perm -lift_permV permK. Qed.
+
+Lemma lift0_perm_eq0 s i : (lift0_perm s i == 0) = (i == 0).
+Proof. by rewrite (canF_eq (lift0_permK s)) lift0_perm0. Qed.
+
+(* Block expresssion of a lifted permutation matrix *)
+
+Definition lift0_mx A : 'M_(1 + n) := block_mx 1 0 0 A.
+
+Lemma lift0_mx_perm s : lift0_mx (perm_mx s) = perm_mx (lift0_perm s).
+Proof.
+apply/matrixP=> /= i j; rewrite !mxE split1 /=; case: unliftP => [i'|] -> /=.
+ rewrite lift0_perm_lift !mxE split1 /=.
+ by case: unliftP => [j'|] ->; rewrite ?(inj_eq (@lift_inj _ _)) /= !mxE.
+rewrite lift0_perm0 !mxE split1 /=.
+by case: unliftP => [j'|] ->; rewrite /= mxE.
+Qed.
+
+Lemma lift0_mx_is_perm s : is_perm_mx (lift0_mx (perm_mx s)).
+Proof. by rewrite lift0_mx_perm perm_mx_is_perm. Qed.
+
+End LiftPerm.
+
+(* Determinants and adjugates are defined here, but most of their properties *)
+(* only hold for matrices over a commutative ring, so their theory is *)
+(* deferred to that section. *)
+
+(* The determinant, in one line with the Leibniz Formula *)
+Definition determinant n (A : 'M_n) : R :=
+ \sum_(s : 'S_n) (-1) ^+ s * \prod_i A i (s i).
+
+(* The cofactor of a matrix on the indexes i and j *)
+Definition cofactor n A (i j : 'I_n) : R :=
+ (-1) ^+ (i + j) * determinant (row' i (col' j A)).
+
+(* The adjugate matrix : defined as the transpose of the matrix of cofactors *)
+Fact adjugate_key : unit. Proof. by []. Qed.
+Definition adjugate n (A : 'M_n) := \matrix[adjugate_key]_(i, j) cofactor A j i.
+
+End MatrixAlgebra.
+
+Implicit Arguments delta_mx [R m n].
+Implicit Arguments scalar_mx [R n].
+Implicit Arguments perm_mx [R n].
+Implicit Arguments tperm_mx [R n].
+Implicit Arguments pid_mx [R m n].
+Implicit Arguments copid_mx [R n].
+Implicit Arguments lin_mulmxr [R m n p].
+Prenex Implicits delta_mx diag_mx scalar_mx is_scalar_mx perm_mx tperm_mx.
+Prenex Implicits pid_mx copid_mx mulmx lin_mulmxr.
+Prenex Implicits mxtrace determinant cofactor adjugate.
+
+Implicit Arguments is_scalar_mxP [R n A].
+Implicit Arguments mul_delta_mx [R m n p].
+Prenex Implicits mul_delta_mx.
+
+Notation "a %:M" := (scalar_mx a) : ring_scope.
+Notation "A *m B" := (mulmx A B) : ring_scope.
+Notation mulmxr := (mulmxr_head tt).
+Notation "\tr A" := (mxtrace A) : ring_scope.
+Notation "'\det' A" := (determinant A) : ring_scope.
+Notation "'\adj' A" := (adjugate A) : ring_scope.
+
+(* Non-commutative transpose requires multiplication in the converse ring. *)
+Lemma trmx_mul_rev (R : ringType) m n p (A : 'M[R]_(m, n)) (B : 'M[R]_(n, p)) :
+ (A *m B)^T = (B : 'M[R^c]_(n, p))^T *m (A : 'M[R^c]_(m, n))^T.
+Proof.
+by apply/matrixP=> k i; rewrite !mxE; apply: eq_bigr => j _; rewrite !mxE.
+Qed.
+
+Canonical matrix_finRingType (R : finRingType) n' :=
+ Eval hnf in [finRingType of 'M[R]_n'.+1].
+
+(* Parametricity over the algebra structure. *)
+Section MapRingMatrix.
+
+Variables (aR rR : ringType) (f : {rmorphism aR -> rR}).
+Local Notation "A ^f" := (map_mx f A) : ring_scope.
+
+Section FixedSize.
+
+Variables m n p : nat.
+Implicit Type A : 'M[aR]_(m, n).
+
+Lemma map_mxZ a A : (a *: A)^f = f a *: A^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE rmorphM. Qed.
+
+Lemma map_mxM A B : (A *m B)^f = A^f *m B^f :> 'M_(m, p).
+Proof.
+apply/matrixP=> i k; rewrite !mxE rmorph_sum //.
+by apply: eq_bigr => j; rewrite !mxE rmorphM.
+Qed.
+
+Lemma map_delta_mx i j : (delta_mx i j)^f = delta_mx i j :> 'M_(m, n).
+Proof. by apply/matrixP=> i' j'; rewrite !mxE rmorph_nat. Qed.
+
+Lemma map_diag_mx d : (diag_mx d)^f = diag_mx d^f :> 'M_n.
+Proof. by apply/matrixP=> i j; rewrite !mxE rmorphMn. Qed.
+
+Lemma map_scalar_mx a : a%:M^f = (f a)%:M :> 'M_n.
+Proof. by apply/matrixP=> i j; rewrite !mxE rmorphMn. Qed.
+
+Lemma map_mx1 : 1%:M^f = 1%:M :> 'M_n.
+Proof. by rewrite map_scalar_mx rmorph1. Qed.
+
+Lemma map_perm_mx (s : 'S_n) : (perm_mx s)^f = perm_mx s.
+Proof. by apply/matrixP=> i j; rewrite !mxE rmorph_nat. Qed.
+
+Lemma map_tperm_mx (i1 i2 : 'I_n) : (tperm_mx i1 i2)^f = tperm_mx i1 i2.
+Proof. exact: map_perm_mx. Qed.
+
+Lemma map_pid_mx r : (pid_mx r)^f = pid_mx r :> 'M_(m, n).
+Proof. by apply/matrixP=> i j; rewrite !mxE rmorph_nat. Qed.
+
+Lemma trace_map_mx (A : 'M_n) : \tr A^f = f (\tr A).
+Proof. by rewrite rmorph_sum; apply: eq_bigr => i _; rewrite mxE. Qed.
+
+Lemma det_map_mx n' (A : 'M_n') : \det A^f = f (\det A).
+Proof.
+rewrite rmorph_sum //; apply: eq_bigr => s _.
+rewrite rmorphM rmorph_sign rmorph_prod; congr (_ * _).
+by apply: eq_bigr => i _; rewrite mxE.
+Qed.
+
+Lemma cofactor_map_mx (A : 'M_n) i j : cofactor A^f i j = f (cofactor A i j).
+Proof. by rewrite rmorphM rmorph_sign -det_map_mx map_row' map_col'. Qed.
+
+Lemma map_mx_adj (A : 'M_n) : (\adj A)^f = \adj A^f.
+Proof. by apply/matrixP=> i j; rewrite !mxE cofactor_map_mx. Qed.
+
+End FixedSize.
+
+Lemma map_copid_mx n r : (copid_mx r)^f = copid_mx r :> 'M_n.
+Proof. by rewrite map_mx_sub map_mx1 map_pid_mx. Qed.
+
+Lemma map_mx_is_multiplicative n' (n := n'.+1) :
+ multiplicative ((map_mx f) n n).
+Proof. by split; [exact: map_mxM | exact: map_mx1]. Qed.
+
+Canonical map_mx_rmorphism n' := AddRMorphism (map_mx_is_multiplicative n').
+
+Lemma map_lin1_mx m n (g : 'rV_m -> 'rV_n) gf :
+ (forall v, (g v)^f = gf v^f) -> (lin1_mx g)^f = lin1_mx gf.
+Proof.
+by move=> def_gf; apply/matrixP=> i j; rewrite !mxE -map_delta_mx -def_gf mxE.
+Qed.
+
+Lemma map_lin_mx m1 n1 m2 n2 (g : 'M_(m1, n1) -> 'M_(m2, n2)) gf :
+ (forall A, (g A)^f = gf A^f) -> (lin_mx g)^f = lin_mx gf.
+Proof.
+move=> def_gf; apply: map_lin1_mx => A /=.
+by rewrite map_mxvec def_gf map_vec_mx.
+Qed.
+
+End MapRingMatrix.
+
+Section ComMatrix.
+(* Lemmas for matrices with coefficients in a commutative ring *)
+Variable R : comRingType.
+
+Section AssocLeft.
+
+Variables m n p : nat.
+Implicit Type A : 'M[R]_(m, n).
+Implicit Type B : 'M[R]_(n, p).
+
+Lemma trmx_mul A B : (A *m B)^T = B^T *m A^T.
+Proof.
+rewrite trmx_mul_rev; apply/matrixP=> k i; rewrite !mxE.
+by apply: eq_bigr => j _; rewrite mulrC.
+Qed.
+
+Lemma scalemxAr a A B : a *: (A *m B) = A *m (a *: B).
+Proof. by apply: trmx_inj; rewrite trmx_mul !linearZ /= trmx_mul scalemxAl. Qed.
+
+Lemma mulmx_is_scalable A : scalable (@mulmx _ m n p A).
+Proof. by move=> a B; rewrite scalemxAr. Qed.
+Canonical mulmx_linear A := AddLinear (mulmx_is_scalable A).
+
+Definition lin_mulmx A : 'M[R]_(n * p, m * p) := lin_mx (mulmx A).
+
+Lemma lin_mulmx_is_linear : linear lin_mulmx.
+Proof.
+move=> a A B; apply/row_matrixP=> i; rewrite linearP /= !rowE !mul_rV_lin /=.
+by rewrite [_ *m _](linearP (mulmxr_linear _ _)) linearP.
+Qed.
+Canonical lin_mulmx_additive := Additive lin_mulmx_is_linear.
+Canonical lin_mulmx_linear := Linear lin_mulmx_is_linear.
+
+End AssocLeft.
+
+Section LinMulRow.
+
+Variables m n : nat.
+
+Definition lin_mul_row u : 'M[R]_(m * n, n) := lin1_mx (mulmx u \o vec_mx).
+
+Lemma lin_mul_row_is_linear : linear lin_mul_row.
+Proof.
+move=> a u v; apply/row_matrixP=> i; rewrite linearP /= !rowE !mul_rV_lin1 /=.
+by rewrite [_ *m _](linearP (mulmxr_linear _ _)).
+Qed.
+Canonical lin_mul_row_additive := Additive lin_mul_row_is_linear.
+Canonical lin_mul_row_linear := Linear lin_mul_row_is_linear.
+
+Lemma mul_vec_lin_row A u : mxvec A *m lin_mul_row u = u *m A.
+Proof. by rewrite mul_rV_lin1 /= mxvecK. Qed.
+
+End LinMulRow.
+
+Lemma mxvec_dotmul m n (A : 'M[R]_(m, n)) u v :
+ mxvec (u^T *m v) *m (mxvec A)^T = u *m A *m v^T.
+Proof.
+transitivity (\sum_i \sum_j (u 0 i * A i j *: row j v^T)).
+ apply/rowP=> i; rewrite {i}ord1 mxE (reindex _ (curry_mxvec_bij _ _)) /=.
+ rewrite pair_bigA summxE; apply: eq_bigr => [[i j]] /= _.
+ by rewrite !mxE !mxvecE mxE big_ord1 mxE mulrAC.
+rewrite mulmx_sum_row exchange_big; apply: eq_bigr => j _ /=.
+by rewrite mxE -scaler_suml.
+Qed.
+
+Section MatrixAlgType.
+
+Variable n' : nat.
+Local Notation n := n'.+1.
+
+Canonical matrix_algType :=
+ Eval hnf in AlgType R 'M[R]_n (fun k => scalemxAr k).
+
+End MatrixAlgType.
+
+Lemma diag_mxC n (d e : 'rV[R]_n) :
+ diag_mx d *m diag_mx e = diag_mx e *m diag_mx d.
+Proof.
+by rewrite !mulmx_diag; congr (diag_mx _); apply/rowP=> i; rewrite !mxE mulrC.
+Qed.
+
+Lemma diag_mx_comm n' (d e : 'rV[R]_n'.+1) : GRing.comm (diag_mx d) (diag_mx e).
+Proof. exact: diag_mxC. Qed.
+
+Lemma scalar_mxC m n a (A : 'M[R]_(m, n)) : A *m a%:M = a%:M *m A.
+Proof.
+by apply: trmx_inj; rewrite trmx_mul tr_scalar_mx !mul_scalar_mx linearZ.
+Qed.
+
+Lemma scalar_mx_comm n' a (A : 'M[R]_n'.+1) : GRing.comm A a%:M.
+Proof. exact: scalar_mxC. Qed.
+
+Lemma mul_mx_scalar m n a (A : 'M[R]_(m, n)) : A *m a%:M = a *: A.
+Proof. by rewrite scalar_mxC mul_scalar_mx. Qed.
+
+Lemma mxtrace_mulC m n (A : 'M[R]_(m, n)) B :
+ \tr (A *m B) = \tr (B *m A).
+Proof.
+have expand_trM C D: \tr (C *m D) = \sum_i \sum_j C i j * D j i.
+ by apply: eq_bigr => i _; rewrite mxE.
+rewrite !{}expand_trM exchange_big /=.
+by do 2!apply: eq_bigr => ? _; apply: mulrC.
+Qed.
+
+(* The theory of determinants *)
+
+Lemma determinant_multilinear n (A B C : 'M[R]_n) i0 b c :
+ row i0 A = b *: row i0 B + c *: row i0 C ->
+ row' i0 B = row' i0 A ->
+ row' i0 C = row' i0 A ->
+ \det A = b * \det B + c * \det C.
+Proof.
+rewrite -[_ + _](row_id 0); move/row_eq=> ABC.
+move/row'_eq=> BA; move/row'_eq=> CA.
+rewrite !big_distrr -big_split; apply: eq_bigr => s _ /=.
+rewrite -!(mulrCA (_ ^+s)) -mulrDr; congr (_ * _).
+rewrite !(bigD1 i0 (_ : predT i0)) //= {}ABC !mxE mulrDl !mulrA.
+by congr (_ * _ + _ * _); apply: eq_bigr => i i0i; rewrite ?BA ?CA.
+Qed.
+
+Lemma determinant_alternate n (A : 'M[R]_n) i1 i2 :
+ i1 != i2 -> A i1 =1 A i2 -> \det A = 0.
+Proof.
+move=> neq_i12 eqA12; pose t := tperm i1 i2.
+have oddMt s: (t * s)%g = ~~ s :> bool by rewrite odd_permM odd_tperm neq_i12.
+rewrite [\det A](bigID (@odd_perm _)) /=.
+apply: canLR (subrK _) _; rewrite add0r -sumrN.
+rewrite (reindex_inj (mulgI t)); apply: eq_big => //= s.
+rewrite oddMt => /negPf->; rewrite mulN1r mul1r; congr (- _).
+rewrite (reindex_inj (@perm_inj _ t)); apply: eq_bigr => /= i _.
+by rewrite permM tpermK /t; case: tpermP => // ->; rewrite eqA12.
+Qed.
+
+Lemma det_tr n (A : 'M[R]_n) : \det A^T = \det A.
+Proof.
+rewrite [\det A^T](reindex_inj (@invg_inj _)) /=.
+apply: eq_bigr => s _ /=; rewrite !odd_permV (reindex_inj (@perm_inj _ s)) /=.
+by congr (_ * _); apply: eq_bigr => i _; rewrite mxE permK.
+Qed.
+
+Lemma det_perm n (s : 'S_n) : \det (perm_mx s) = (-1) ^+ s :> R.
+Proof.
+rewrite [\det _](bigD1 s) //= big1 => [|i _]; last by rewrite /= !mxE eqxx.
+rewrite mulr1 big1 ?addr0 => //= t Dst.
+case: (pickP (fun i => s i != t i)) => [i ist | Est].
+ by rewrite (bigD1 i) // mulrCA /= !mxE (negbTE ist) mul0r.
+by case/eqP: Dst; apply/permP => i; move/eqP: (Est i).
+Qed.
+
+Lemma det1 n : \det (1%:M : 'M[R]_n) = 1.
+Proof. by rewrite -perm_mx1 det_perm odd_perm1. Qed.
+
+Lemma det_mx00 (A : 'M[R]_0) : \det A = 1.
+Proof. by rewrite flatmx0 -(flatmx0 1%:M) det1. Qed.
+
+Lemma detZ n a (A : 'M[R]_n) : \det (a *: A) = a ^+ n * \det A.
+Proof.
+rewrite big_distrr /=; apply: eq_bigr => s _; rewrite mulrCA; congr (_ * _).
+rewrite -[n in a ^+ n]card_ord -prodr_const -big_split /=.
+by apply: eq_bigr=> i _; rewrite mxE.
+Qed.
+
+Lemma det0 n' : \det (0 : 'M[R]_n'.+1) = 0.
+Proof. by rewrite -(scale0r 0) detZ exprS !mul0r. Qed.
+
+Lemma det_scalar n a : \det (a%:M : 'M[R]_n) = a ^+ n.
+Proof. by rewrite -{1}(mulr1 a) -scale_scalar_mx detZ det1 mulr1. Qed.
+
+Lemma det_scalar1 a : \det (a%:M : 'M[R]_1) = a.
+Proof. exact: det_scalar. Qed.
+
+Lemma det_mulmx n (A B : 'M[R]_n) : \det (A *m B) = \det A * \det B.
+Proof.
+rewrite big_distrl /=.
+pose F := ('I_n ^ n)%type; pose AB s i j := A i j * B j (s i).
+transitivity (\sum_(f : F) \sum_(s : 'S_n) (-1) ^+ s * \prod_i AB s i (f i)).
+ rewrite exchange_big; apply: eq_bigr => /= s _; rewrite -big_distrr /=.
+ congr (_ * _); rewrite -(bigA_distr_bigA (AB s)) /=.
+ by apply: eq_bigr => x _; rewrite mxE.
+rewrite (bigID (fun f : F => injectiveb f)) /= addrC big1 ?add0r => [|f Uf].
+ rewrite (reindex (@pval _)) /=; last first.
+ pose in_Sn := insubd (1%g : 'S_n).
+ by exists in_Sn => /= f Uf; first apply: val_inj; exact: insubdK.
+ apply: eq_big => /= [s | s _]; rewrite ?(valP s) // big_distrr /=.
+ rewrite (reindex_inj (mulgI s)); apply: eq_bigr => t _ /=.
+ rewrite big_split /= mulrA mulrCA mulrA mulrCA mulrA.
+ rewrite -signr_addb odd_permM !pvalE; congr (_ * _); symmetry.
+ by rewrite (reindex_inj (@perm_inj _ s)); apply: eq_bigr => i; rewrite permM.
+transitivity (\det (\matrix_(i, j) B (f i) j) * \prod_i A i (f i)).
+ rewrite mulrC big_distrr /=; apply: eq_bigr => s _.
+ rewrite mulrCA big_split //=; congr (_ * (_ * _)).
+ by apply: eq_bigr => x _; rewrite mxE.
+case/injectivePn: Uf => i1 [i2 Di12 Ef12].
+by rewrite (determinant_alternate Di12) ?simp //= => j; rewrite !mxE Ef12.
+Qed.
+
+Lemma detM n' (A B : 'M[R]_n'.+1) : \det (A * B) = \det A * \det B.
+Proof. exact: det_mulmx. Qed.
+
+Lemma det_diag n (d : 'rV[R]_n) : \det (diag_mx d) = \prod_i d 0 i.
+Proof.
+rewrite /(\det _) (bigD1 1%g) //= addrC big1 => [|p p1].
+ by rewrite add0r odd_perm1 mul1r; apply: eq_bigr => i; rewrite perm1 mxE eqxx.
+have{p1}: ~~ perm_on set0 p.
+ apply: contra p1; move/subsetP=> p1; apply/eqP; apply/permP=> i.
+ by rewrite perm1; apply/eqP; apply/idPn; move/p1; rewrite inE.
+case/subsetPn=> i; rewrite !inE eq_sym; move/negbTE=> p_i _.
+by rewrite (bigD1 i) //= mulrCA mxE p_i mul0r.
+Qed.
+
+(* Laplace expansion lemma *)
+Lemma expand_cofactor n (A : 'M[R]_n) i j :
+ cofactor A i j =
+ \sum_(s : 'S_n | s i == j) (-1) ^+ s * \prod_(k | i != k) A k (s k).
+Proof.
+case: n A i j => [|n] A i0 j0; first by case: i0.
+rewrite (reindex (lift_perm i0 j0)); last first.
+ pose ulsf i (s : 'S_n.+1) k := odflt k (unlift (s i) (s (lift i k))).
+ have ulsfK i (s : 'S_n.+1) k: lift (s i) (ulsf i s k) = s (lift i k).
+ rewrite /ulsf; have:= neq_lift i k.
+ by rewrite -(inj_eq (@perm_inj _ s)) => /unlift_some[] ? ? ->.
+ have inj_ulsf: injective (ulsf i0 _).
+ move=> s; apply: can_inj (ulsf (s i0) s^-1%g) _ => k'.
+ by rewrite {1}/ulsf ulsfK !permK liftK.
+ exists (fun s => perm (inj_ulsf s)) => [s _ | s].
+ by apply/permP=> k'; rewrite permE /ulsf lift_perm_lift lift_perm_id liftK.
+ move/(s _ =P _) => si0; apply/permP=> k.
+ case: (unliftP i0 k) => [k'|] ->; rewrite ?lift_perm_id //.
+ by rewrite lift_perm_lift -si0 permE ulsfK.
+rewrite /cofactor big_distrr /=.
+apply: eq_big => [s | s _]; first by rewrite lift_perm_id eqxx.
+rewrite -signr_odd mulrA -signr_addb odd_add -odd_lift_perm; congr (_ * _).
+case: (pickP 'I_n) => [k0 _ | n0]; last first.
+ by rewrite !big1 // => [j /unlift_some[i] | i _]; have:= n0 i.
+rewrite (reindex (lift i0)).
+ by apply: eq_big => [k | k _] /=; rewrite ?neq_lift // !mxE lift_perm_lift.
+exists (fun k => odflt k0 (unlift i0 k)) => k; first by rewrite liftK.
+by case/unlift_some=> k' -> ->.
+Qed.
+
+Lemma expand_det_row n (A : 'M[R]_n) i0 :
+ \det A = \sum_j A i0 j * cofactor A i0 j.
+Proof.
+rewrite /(\det A) (partition_big (fun s : 'S_n => s i0) predT) //=.
+apply: eq_bigr => j0 _; rewrite expand_cofactor big_distrr /=.
+apply: eq_bigr => s /eqP Dsi0.
+rewrite mulrCA (bigID (pred1 i0)) /= big_pred1_eq Dsi0; congr (_ * (_ * _)).
+by apply: eq_bigl => i; rewrite eq_sym.
+Qed.
+
+Lemma cofactor_tr n (A : 'M[R]_n) i j : cofactor A^T i j = cofactor A j i.
+Proof.
+rewrite /cofactor addnC; congr (_ * _).
+rewrite -tr_row' -tr_col' det_tr; congr (\det _).
+by apply/matrixP=> ? ?; rewrite !mxE.
+Qed.
+
+Lemma cofactorZ n a (A : 'M[R]_n) i j :
+ cofactor (a *: A) i j = a ^+ n.-1 * cofactor A i j.
+Proof. by rewrite {1}/cofactor !linearZ detZ mulrCA mulrA. Qed.
+
+Lemma expand_det_col n (A : 'M[R]_n) j0 :
+ \det A = \sum_i (A i j0 * cofactor A i j0).
+Proof.
+rewrite -det_tr (expand_det_row _ j0).
+by apply: eq_bigr => i _; rewrite cofactor_tr mxE.
+Qed.
+
+Lemma trmx_adj n (A : 'M[R]_n) : (\adj A)^T = \adj A^T.
+Proof. by apply/matrixP=> i j; rewrite !mxE cofactor_tr. Qed.
+
+Lemma adjZ n a (A : 'M[R]_n) : \adj (a *: A) = a^+n.-1 *: \adj A.
+Proof. by apply/matrixP=> i j; rewrite !mxE cofactorZ. Qed.
+
+(* Cramer Rule : adjugate on the left *)
+Lemma mul_mx_adj n (A : 'M[R]_n) : A *m \adj A = (\det A)%:M.
+Proof.
+apply/matrixP=> i1 i2; rewrite !mxE; case Di: (i1 == i2).
+ rewrite (eqP Di) (expand_det_row _ i2) //=.
+ by apply: eq_bigr => j _; congr (_ * _); rewrite mxE.
+pose B := \matrix_(i, j) (if i == i2 then A i1 j else A i j).
+have EBi12: B i1 =1 B i2 by move=> j; rewrite /= !mxE Di eq_refl.
+rewrite -[_ *+ _](determinant_alternate (negbT Di) EBi12) (expand_det_row _ i2).
+apply: eq_bigr => j _; rewrite !mxE eq_refl; congr (_ * (_ * _)).
+apply: eq_bigr => s _; congr (_ * _); apply: eq_bigr => i _.
+by rewrite !mxE eq_sym -if_neg neq_lift.
+Qed.
+
+(* Cramer rule : adjugate on the right *)
+Lemma mul_adj_mx n (A : 'M[R]_n) : \adj A *m A = (\det A)%:M.
+Proof.
+by apply: trmx_inj; rewrite trmx_mul trmx_adj mul_mx_adj det_tr tr_scalar_mx.
+Qed.
+
+Lemma adj1 n : \adj (1%:M) = 1%:M :> 'M[R]_n.
+Proof. by rewrite -{2}(det1 n) -mul_adj_mx mulmx1. Qed.
+
+(* Left inverses are right inverses. *)
+Lemma mulmx1C n (A B : 'M[R]_n) : A *m B = 1%:M -> B *m A = 1%:M.
+Proof.
+move=> AB1; pose A' := \det B *: \adj A.
+suffices kA: A' *m A = 1%:M by rewrite -[B]mul1mx -kA -(mulmxA A') AB1 mulmx1.
+by rewrite -scalemxAl mul_adj_mx scale_scalar_mx mulrC -det_mulmx AB1 det1.
+Qed.
+
+(* Only tall matrices have inverses. *)
+Lemma mulmx1_min m n (A : 'M[R]_(m, n)) B : A *m B = 1%:M -> m <= n.
+Proof.
+move=> AB1; rewrite leqNgt; apply/negP=> /subnKC; rewrite addSnnS.
+move: (_ - _)%N => m' def_m; move: AB1; rewrite -{m}def_m in A B *.
+rewrite -(vsubmxK A) -(hsubmxK B) mul_col_row scalar_mx_block.
+case/eq_block_mx=> /mulmx1C BlAu1 AuBr0 _ => /eqP/idPn[].
+by rewrite -[_ B]mul1mx -BlAu1 -mulmxA AuBr0 !mulmx0 eq_sym oner_neq0.
+Qed.
+
+Lemma det_ublock n1 n2 Aul (Aur : 'M[R]_(n1, n2)) Adr :
+ \det (block_mx Aul Aur 0 Adr) = \det Aul * \det Adr.
+Proof.
+elim: n1 => [|n1 IHn1] in Aul Aur *.
+ have ->: Aul = 1%:M by apply/matrixP=> i [].
+ rewrite det1 mul1r; congr (\det _); apply/matrixP=> i j.
+ by do 2![rewrite !mxE; case: splitP => [[]|k] //=; move/val_inj=> <- {k}].
+rewrite (expand_det_col _ (lshift n2 0)) big_split_ord /=.
+rewrite addrC big1 1?simp => [|i _]; last by rewrite block_mxEdl mxE simp.
+rewrite (expand_det_col _ 0) big_distrl /=; apply eq_bigr=> i _.
+rewrite block_mxEul -!mulrA; do 2!congr (_ * _).
+by rewrite col'_col_mx !col'Kl raddf0 row'Ku row'_row_mx IHn1.
+Qed.
+
+Lemma det_lblock n1 n2 Aul (Adl : 'M[R]_(n2, n1)) Adr :
+ \det (block_mx Aul 0 Adl Adr) = \det Aul * \det Adr.
+Proof. by rewrite -det_tr tr_block_mx trmx0 det_ublock !det_tr. Qed.
+
+End ComMatrix.
+
+Implicit Arguments lin_mul_row [R m n].
+Implicit Arguments lin_mulmx [R m n p].
+Prenex Implicits lin_mul_row lin_mulmx.
+
+(*****************************************************************************)
+(********************** Matrix unit ring and inverse matrices ****************)
+(*****************************************************************************)
+
+Section MatrixInv.
+
+Variables R : comUnitRingType.
+
+Section Defs.
+
+Variable n : nat.
+Implicit Type A : 'M[R]_n.
+
+Definition unitmx : pred 'M[R]_n := fun A => \det A \is a GRing.unit.
+Definition invmx A := if A \in unitmx then (\det A)^-1 *: \adj A else A.
+
+Lemma unitmxE A : (A \in unitmx) = (\det A \is a GRing.unit).
+Proof. by []. Qed.
+
+Lemma unitmx1 : 1%:M \in unitmx. Proof. by rewrite unitmxE det1 unitr1. Qed.
+
+Lemma unitmx_perm s : perm_mx s \in unitmx.
+Proof. by rewrite unitmxE det_perm unitrX ?unitrN ?unitr1. Qed.
+
+Lemma unitmx_tr A : (A^T \in unitmx) = (A \in unitmx).
+Proof. by rewrite unitmxE det_tr. Qed.
+
+Lemma unitmxZ a A : a \is a GRing.unit -> (a *: A \in unitmx) = (A \in unitmx).
+Proof. by move=> Ua; rewrite !unitmxE detZ unitrM unitrX. Qed.
+
+Lemma invmx1 : invmx 1%:M = 1%:M.
+Proof. by rewrite /invmx det1 invr1 scale1r adj1 if_same. Qed.
+
+Lemma invmxZ a A : a *: A \in unitmx -> invmx (a *: A) = a^-1 *: invmx A.
+Proof.
+rewrite /invmx !unitmxE detZ unitrM => /andP[Ua U_A].
+rewrite Ua U_A adjZ !scalerA invrM {U_A}//=.
+case: (posnP n) A => [-> | n_gt0] A; first by rewrite flatmx0 [_ *: _]flatmx0.
+rewrite unitrX_pos // in Ua; rewrite -[_ * _](mulrK Ua) mulrC -!mulrA.
+by rewrite -exprSr prednK // !mulrA divrK ?unitrX.
+Qed.
+
+Lemma invmx_scalar a : invmx (a%:M) = a^-1%:M.
+Proof.
+case Ua: (a%:M \in unitmx).
+ by rewrite -scalemx1 in Ua *; rewrite invmxZ // invmx1 scalemx1.
+rewrite /invmx Ua; have [->|n_gt0] := posnP n; first by rewrite ![_%:M]flatmx0.
+by rewrite unitmxE det_scalar unitrX_pos // in Ua; rewrite invr_out ?Ua.
+Qed.
+
+Lemma mulVmx : {in unitmx, left_inverse 1%:M invmx mulmx}.
+Proof.
+by move=> A nsA; rewrite /invmx nsA -scalemxAl mul_adj_mx scale_scalar_mx mulVr.
+Qed.
+
+Lemma mulmxV : {in unitmx, right_inverse 1%:M invmx mulmx}.
+Proof.
+by move=> A nsA; rewrite /invmx nsA -scalemxAr mul_mx_adj scale_scalar_mx mulVr.
+Qed.
+
+Lemma mulKmx m : {in unitmx, @left_loop _ 'M_(n, m) invmx mulmx}.
+Proof. by move=> A uA /= B; rewrite mulmxA mulVmx ?mul1mx. Qed.
+
+Lemma mulKVmx m : {in unitmx, @rev_left_loop _ 'M_(n, m) invmx mulmx}.
+Proof. by move=> A uA /= B; rewrite mulmxA mulmxV ?mul1mx. Qed.
+
+Lemma mulmxK m : {in unitmx, @right_loop 'M_(m, n) _ invmx mulmx}.
+Proof. by move=> A uA /= B; rewrite -mulmxA mulmxV ?mulmx1. Qed.
+
+Lemma mulmxKV m : {in unitmx, @rev_right_loop 'M_(m, n) _ invmx mulmx}.
+Proof. by move=> A uA /= B; rewrite -mulmxA mulVmx ?mulmx1. Qed.
+
+Lemma det_inv A : \det (invmx A) = (\det A)^-1.
+Proof.
+case uA: (A \in unitmx); last by rewrite /invmx uA invr_out ?negbT.
+by apply: (mulrI uA); rewrite -det_mulmx mulmxV ?divrr ?det1.
+Qed.
+
+Lemma unitmx_inv A : (invmx A \in unitmx) = (A \in unitmx).
+Proof. by rewrite !unitmxE det_inv unitrV. Qed.
+
+Lemma unitmx_mul A B : (A *m B \in unitmx) = (A \in unitmx) && (B \in unitmx).
+Proof. by rewrite -unitrM -det_mulmx. Qed.
+
+Lemma trmx_inv (A : 'M_n) : (invmx A)^T = invmx (A^T).
+Proof. by rewrite (fun_if trmx) linearZ /= trmx_adj -unitmx_tr -det_tr. Qed.
+
+Lemma invmxK : involutive invmx.
+Proof.
+move=> A; case uA : (A \in unitmx); last by rewrite /invmx !uA.
+by apply: (can_inj (mulKVmx uA)); rewrite mulVmx // mulmxV ?unitmx_inv.
+Qed.
+
+Lemma mulmx1_unit A B : A *m B = 1%:M -> A \in unitmx /\ B \in unitmx.
+Proof. by move=> AB1; apply/andP; rewrite -unitmx_mul AB1 unitmx1. Qed.
+
+Lemma intro_unitmx A B : B *m A = 1%:M /\ A *m B = 1%:M -> unitmx A.
+Proof. by case=> _ /mulmx1_unit[]. Qed.
+
+Lemma invmx_out : {in [predC unitmx], invmx =1 id}.
+Proof. by move=> A; rewrite inE /= /invmx -if_neg => ->. Qed.
+
+End Defs.
+
+Variable n' : nat.
+Local Notation n := n'.+1.
+
+Definition matrix_unitRingMixin :=
+ UnitRingMixin (@mulVmx n) (@mulmxV n) (@intro_unitmx n) (@invmx_out n).
+Canonical matrix_unitRing :=
+ Eval hnf in UnitRingType 'M[R]_n matrix_unitRingMixin.
+Canonical matrix_unitAlg := Eval hnf in [unitAlgType R of 'M[R]_n].
+
+(* Lemmas requiring that the coefficients are in a unit ring *)
+
+Lemma detV (A : 'M_n) : \det A^-1 = (\det A)^-1.
+Proof. exact: det_inv. Qed.
+
+Lemma unitr_trmx (A : 'M_n) : (A^T \is a GRing.unit) = (A \is a GRing.unit).
+Proof. exact: unitmx_tr. Qed.
+
+Lemma trmxV (A : 'M_n) : A^-1^T = (A^T)^-1.
+Proof. exact: trmx_inv. Qed.
+
+Lemma perm_mxV (s : 'S_n) : perm_mx s^-1 = (perm_mx s)^-1.
+Proof.
+rewrite -[_^-1]mul1r; apply: (canRL (mulmxK (unitmx_perm s))).
+by rewrite -perm_mxM mulVg perm_mx1.
+Qed.
+
+Lemma is_perm_mxV (A : 'M_n) : is_perm_mx A^-1 = is_perm_mx A.
+Proof.
+apply/is_perm_mxP/is_perm_mxP=> [] [s defA]; exists s^-1%g.
+ by rewrite -(invrK A) defA perm_mxV.
+by rewrite defA perm_mxV.
+Qed.
+
+End MatrixInv.
+
+Prenex Implicits unitmx invmx.
+
+(* Finite inversible matrices and the general linear group. *)
+Section FinUnitMatrix.
+
+Variables (n : nat) (R : finComUnitRingType).
+
+Canonical matrix_finUnitRingType n' :=
+ Eval hnf in [finUnitRingType of 'M[R]_n'.+1].
+
+Definition GLtype of phant R := {unit 'M[R]_n.-1.+1}.
+
+Coercion GLval ph (u : GLtype ph) : 'M[R]_n.-1.+1 :=
+ let: FinRing.Unit A _ := u in A.
+
+End FinUnitMatrix.
+
+Bind Scope group_scope with GLtype.
+Arguments Scope GLval [nat_scope _ _ group_scope].
+Prenex Implicits GLval.
+
+Notation "{ ''GL_' n [ R ] }" := (GLtype n (Phant R))
+ (at level 0, n at level 2, format "{ ''GL_' n [ R ] }") : type_scope.
+Notation "{ ''GL_' n ( p ) }" := {'GL_n['F_p]}
+ (at level 0, n at level 2, p at level 10,
+ format "{ ''GL_' n ( p ) }") : type_scope.
+
+Section GL_unit.
+
+Variables (n : nat) (R : finComUnitRingType).
+
+Canonical GL_subType := [subType of {'GL_n[R]} for GLval].
+Definition GL_eqMixin := Eval hnf in [eqMixin of {'GL_n[R]} by <:].
+Canonical GL_eqType := Eval hnf in EqType {'GL_n[R]} GL_eqMixin.
+Canonical GL_choiceType := Eval hnf in [choiceType of {'GL_n[R]}].
+Canonical GL_countType := Eval hnf in [countType of {'GL_n[R]}].
+Canonical GL_subCountType := Eval hnf in [subCountType of {'GL_n[R]}].
+Canonical GL_finType := Eval hnf in [finType of {'GL_n[R]}].
+Canonical GL_subFinType := Eval hnf in [subFinType of {'GL_n[R]}].
+Canonical GL_baseFinGroupType := Eval hnf in [baseFinGroupType of {'GL_n[R]}].
+Canonical GL_finGroupType := Eval hnf in [finGroupType of {'GL_n[R]}].
+Definition GLgroup of phant R := [set: {'GL_n[R]}].
+Canonical GLgroup_group ph := Eval hnf in [group of GLgroup ph].
+
+Implicit Types u v : {'GL_n[R]}.
+
+Lemma GL_1E : GLval 1 = 1. Proof. by []. Qed.
+Lemma GL_VE u : GLval u^-1 = (GLval u)^-1. Proof. by []. Qed.
+Lemma GL_VxE u : GLval u^-1 = invmx u. Proof. by []. Qed.
+Lemma GL_ME u v : GLval (u * v) = GLval u * GLval v. Proof. by []. Qed.
+Lemma GL_MxE u v : GLval (u * v) = u *m v. Proof. by []. Qed.
+Lemma GL_unit u : GLval u \is a GRing.unit. Proof. exact: valP. Qed.
+Lemma GL_unitmx u : val u \in unitmx. Proof. exact: GL_unit. Qed.
+
+Lemma GL_det u : \det u != 0.
+Proof.
+by apply: contraL (GL_unitmx u); rewrite unitmxE => /eqP->; rewrite unitr0.
+Qed.
+
+End GL_unit.
+
+Notation "''GL_' n [ R ]" := (GLgroup n (Phant R))
+ (at level 8, n at level 2, format "''GL_' n [ R ]") : group_scope.
+Notation "''GL_' n ( p )" := 'GL_n['F_p]
+ (at level 8, n at level 2, p at level 10,
+ format "''GL_' n ( p )") : group_scope.
+Notation "''GL_' n [ R ]" := (GLgroup_group n (Phant R)) : Group_scope.
+Notation "''GL_' n ( p )" := (GLgroup_group n (Phant 'F_p)) : Group_scope.
+
+(*****************************************************************************)
+(********************** Matrices over a domain *******************************)
+(*****************************************************************************)
+
+Section MatrixDomain.
+
+Variable R : idomainType.
+
+Lemma scalemx_eq0 m n a (A : 'M[R]_(m, n)) :
+ (a *: A == 0) = (a == 0) || (A == 0).
+Proof.
+case nz_a: (a == 0) / eqP => [-> | _]; first by rewrite scale0r eqxx.
+apply/eqP/eqP=> [aA0 | ->]; last exact: scaler0.
+apply/matrixP=> i j; apply/eqP; move/matrixP/(_ i j)/eqP: aA0.
+by rewrite !mxE mulf_eq0 nz_a.
+Qed.
+
+Lemma scalemx_inj m n a :
+ a != 0 -> injective ( *:%R a : 'M[R]_(m, n) -> 'M[R]_(m, n)).
+Proof.
+move=> nz_a A B eq_aAB; apply: contraNeq nz_a.
+rewrite -[A == B]subr_eq0 -[a == 0]orbF => /negPf<-.
+by rewrite -scalemx_eq0 linearB subr_eq0 /= eq_aAB.
+Qed.
+
+Lemma det0P n (A : 'M[R]_n) :
+ reflect (exists2 v : 'rV[R]_n, v != 0 & v *m A = 0) (\det A == 0).
+Proof.
+apply: (iffP eqP) => [detA0 | [v n0v vA0]]; last first.
+ apply: contraNeq n0v => nz_detA; rewrite -(inj_eq (scalemx_inj nz_detA)).
+ by rewrite scaler0 -mul_mx_scalar -mul_mx_adj mulmxA vA0 mul0mx.
+elim: n => [|n IHn] in A detA0 *.
+ by case/idP: (oner_eq0 R); rewrite -detA0 [A]thinmx0 -(thinmx0 1%:M) det1.
+have [{detA0}A'0 | nzA'] := eqVneq (row 0 (\adj A)) 0; last first.
+ exists (row 0 (\adj A)) => //; rewrite rowE -mulmxA mul_adj_mx detA0.
+ by rewrite mul_mx_scalar scale0r.
+pose A' := col' 0 A; pose vA := col 0 A.
+have defA: A = row_mx vA A'.
+ apply/matrixP=> i j; rewrite !mxE.
+ case: splitP => j' def_j; rewrite mxE; congr (A i _); apply: val_inj => //=.
+ by rewrite def_j [j']ord1.
+have{IHn} w_ j : exists w : 'rV_n.+1, [/\ w != 0, w 0 j = 0 & w *m A' = 0].
+ have [|wj nzwj wjA'0] := IHn (row' j A').
+ by apply/eqP; move/rowP/(_ j)/eqP: A'0; rewrite !mxE mulf_eq0 signr_eq0.
+ exists (\row_k oapp (wj 0) 0 (unlift j k)).
+ rewrite !mxE unlift_none -wjA'0; split=> //.
+ apply: contraNneq nzwj => w0; apply/eqP/rowP=> k'.
+ by move/rowP/(_ (lift j k')): w0; rewrite !mxE liftK.
+ apply/rowP=> k; rewrite !mxE (bigD1 j) //= mxE unlift_none mul0r add0r.
+ rewrite (reindex_onto (lift j) (odflt k \o unlift j)) /= => [|k'].
+ by apply: eq_big => k'; rewrite ?mxE liftK eq_sym neq_lift eqxx.
+ by rewrite eq_sym; case/unlift_some=> ? ? ->.
+have [w0 [nz_w0 w00_0 w0A']] := w_ 0; pose a0 := (w0 *m vA) 0 0.
+have [j {nz_w0}/= nz_w0j | w00] := pickP [pred j | w0 0 j != 0]; last first.
+ by case/eqP: nz_w0; apply/rowP=> j; rewrite mxE; move/eqP: (w00 j).
+have{w_} [wj [nz_wj wj0_0 wjA']] := w_ j; pose aj := (wj *m vA) 0 0.
+have [aj0 | nz_aj] := eqVneq aj 0.
+ exists wj => //; rewrite defA (@mul_mx_row _ _ _ 1) [_ *m _]mx11_scalar -/aj.
+ by rewrite aj0 raddf0 wjA' row_mx0.
+exists (aj *: w0 - a0 *: wj).
+ apply: contraNneq nz_aj; move/rowP/(_ j)/eqP; rewrite !mxE wj0_0 mulr0 subr0.
+ by rewrite mulf_eq0 (negPf nz_w0j) orbF.
+rewrite defA (@mul_mx_row _ _ _ 1) !mulmxBl -!scalemxAl w0A' wjA' !linear0.
+by rewrite -mul_mx_scalar -mul_scalar_mx -!mx11_scalar subrr addr0 row_mx0.
+Qed.
+
+End MatrixDomain.
+
+Implicit Arguments det0P [R n A].
+
+(* Parametricity at the field level (mx_is_scalar, unit and inverse are only *)
+(* mapped at this level). *)
+Section MapFieldMatrix.
+
+Variables (aF : fieldType) (rF : comUnitRingType) (f : {rmorphism aF -> rF}).
+Local Notation "A ^f" := (map_mx f A) : ring_scope.
+
+Lemma map_mx_inj m n : injective ((map_mx f) m n).
+Proof.
+move=> A B eq_AB; apply/matrixP=> i j.
+by move/matrixP/(_ i j): eq_AB; rewrite !mxE; exact: fmorph_inj.
+Qed.
+
+Lemma map_mx_is_scalar n (A : 'M_n) : is_scalar_mx A^f = is_scalar_mx A.
+Proof.
+rewrite /is_scalar_mx; case: (insub _) => // i.
+by rewrite mxE -map_scalar_mx inj_eq //; exact: map_mx_inj.
+Qed.
+
+Lemma map_unitmx n (A : 'M_n) : (A^f \in unitmx) = (A \in unitmx).
+Proof. by rewrite unitmxE det_map_mx // fmorph_unit // -unitfE. Qed.
+
+Lemma map_mx_unit n' (A : 'M_n'.+1) :
+ (A^f \is a GRing.unit) = (A \is a GRing.unit).
+Proof. exact: map_unitmx. Qed.
+
+Lemma map_invmx n (A : 'M_n) : (invmx A)^f = invmx A^f.
+Proof.
+rewrite /invmx map_unitmx (fun_if ((map_mx f) n n)).
+by rewrite map_mxZ map_mx_adj det_map_mx fmorphV.
+Qed.
+
+Lemma map_mx_inv n' (A : 'M_n'.+1) : A^-1^f = A^f^-1.
+Proof. exact: map_invmx. Qed.
+
+Lemma map_mx_eq0 m n (A : 'M_(m, n)) : (A^f == 0) = (A == 0).
+Proof. by rewrite -(inj_eq (@map_mx_inj m n)) raddf0. Qed.
+
+End MapFieldMatrix.
+
+(*****************************************************************************)
+(****************************** LUP decomposion ******************************)
+(*****************************************************************************)
+
+Section CormenLUP.
+
+Variable F : fieldType.
+
+(* Decomposition of the matrix A to P A = L U with *)
+(* - P a permutation matrix *)
+(* - L a unipotent lower triangular matrix *)
+(* - U an upper triangular matrix *)
+
+Fixpoint cormen_lup {n} :=
+ match n return let M := 'M[F]_n.+1 in M -> M * M * M with
+ | 0 => fun A => (1, 1, A)
+ | _.+1 => fun A =>
+ let k := odflt 0 [pick k | A k 0 != 0] in
+ let A1 : 'M_(1 + _) := xrow 0 k A in
+ let P1 : 'M_(1 + _) := tperm_mx 0 k in
+ let Schur := ((A k 0)^-1 *: dlsubmx A1) *m ursubmx A1 in
+ let: (P2, L2, U2) := cormen_lup (drsubmx A1 - Schur) in
+ let P := block_mx 1 0 0 P2 *m P1 in
+ let L := block_mx 1 0 ((A k 0)^-1 *: (P2 *m dlsubmx A1)) L2 in
+ let U := block_mx (ulsubmx A1) (ursubmx A1) 0 U2 in
+ (P, L, U)
+ end.
+
+Lemma cormen_lup_perm n (A : 'M_n.+1) : is_perm_mx (cormen_lup A).1.1.
+Proof.
+elim: n => [|n IHn] /= in A *; first exact: is_perm_mx1.
+set A' := _ - _; move/(_ A'): IHn; case: cormen_lup => [[P L U]] {A'}/=.
+rewrite (is_perm_mxMr _ (perm_mx_is_perm _ _)).
+case/is_perm_mxP => s ->; exact: lift0_mx_is_perm.
+Qed.
+
+Lemma cormen_lup_correct n (A : 'M_n.+1) :
+ let: (P, L, U) := cormen_lup A in P * A = L * U.
+Proof.
+elim: n => [|n IHn] /= in A *; first by rewrite !mul1r.
+set k := odflt _ _; set A1 : 'M_(1 + _) := xrow _ _ _.
+set A' := _ - _; move/(_ A'): IHn; case: cormen_lup => [[P' L' U']] /= IHn.
+rewrite -mulrA -!mulmxE -xrowE -/A1 /= -[n.+2]/(1 + n.+1)%N -{1}(submxK A1).
+rewrite !mulmx_block !mul0mx !mulmx0 !add0r !addr0 !mul1mx -{L' U'}[L' *m _]IHn.
+rewrite -scalemxAl !scalemxAr -!mulmxA addrC -mulrDr {A'}subrK.
+congr (block_mx _ _ (_ *m _) _).
+rewrite [_ *: _]mx11_scalar !mxE lshift0 tpermL {}/A1 {}/k.
+case: pickP => /= [k nzAk0 | no_k]; first by rewrite mulVf ?mulmx1.
+rewrite (_ : dlsubmx _ = 0) ?mul0mx //; apply/colP=> i.
+by rewrite !mxE lshift0 (elimNf eqP (no_k _)).
+Qed.
+
+Lemma cormen_lup_detL n (A : 'M_n.+1) : \det (cormen_lup A).1.2 = 1.
+Proof.
+elim: n => [|n IHn] /= in A *; first by rewrite det1.
+set A' := _ - _; move/(_ A'): IHn; case: cormen_lup => [[P L U]] {A'}/= detL.
+by rewrite (@det_lblock _ 1) det1 mul1r.
+Qed.
+
+Lemma cormen_lup_lower n A (i j : 'I_n.+1) :
+ i <= j -> (cormen_lup A).1.2 i j = (i == j)%:R.
+Proof.
+elim: n => [|n IHn] /= in A i j *; first by rewrite [i]ord1 [j]ord1 mxE.
+set A' := _ - _; move/(_ A'): IHn; case: cormen_lup => [[P L U]] {A'}/= Ll.
+rewrite !mxE split1; case: unliftP => [i'|] -> /=; rewrite !mxE split1.
+ by case: unliftP => [j'|] -> //; exact: Ll.
+by case: unliftP => [j'|] ->; rewrite /= mxE.
+Qed.
+
+Lemma cormen_lup_upper n A (i j : 'I_n.+1) :
+ j < i -> (cormen_lup A).2 i j = 0 :> F.
+Proof.
+elim: n => [|n IHn] /= in A i j *; first by rewrite [i]ord1.
+set A' := _ - _; move/(_ A'): IHn; case: cormen_lup => [[P L U]] {A'}/= Uu.
+rewrite !mxE split1; case: unliftP => [i'|] -> //=; rewrite !mxE split1.
+by case: unliftP => [j'|] ->; [exact: Uu | rewrite /= mxE].
+Qed.
+
+End CormenLUP.
diff --git a/mathcomp/algebra/mxalgebra.v b/mathcomp/algebra/mxalgebra.v
new file mode 100644
index 0000000..54d5588
--- /dev/null
+++ b/mathcomp/algebra/mxalgebra.v
@@ -0,0 +1,2764 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div choice fintype.
+Require Import finfun bigop prime binomial ssralg finset fingroup finalg.
+Require Import perm zmodp matrix.
+
+(*****************************************************************************)
+(* In this file we develop the rank and row space theory of matrices, based *)
+(* on an extended Gaussian elimination procedure similar to LUP *)
+(* decomposition. This provides us with a concrete but generic model of *)
+(* finite dimensional vector spaces and F-algebras, in which vectors, linear *)
+(* functions, families, bases, subspaces, ideals and subrings are all *)
+(* represented using matrices. This model can be used as a foundation for *)
+(* the usual theory of abstract linear algebra, but it can also be used to *)
+(* develop directly substantial theories, such as the theory of finite group *)
+(* linear representation. *)
+(* Here we define the following concepts and notations: *)
+(* Gaussian_elimination A == a permuted triangular decomposition (L, U, r) *)
+(* of A, with L a column permutation of a lower triangular *)
+(* invertible matrix, U a row permutation of an upper *)
+(* triangular invertible matrix, and r the rank of A, all *)
+(* satisfying the identity L *m pid_mx r *m U = A. *)
+(* \rank A == the rank of A. *)
+(* row_free A <=> the rows of A are linearly free (i.e., the rank and *)
+(* height of A are equal). *)
+(* row_full A <=> the row-space of A spans all row-vectors (i.e., the *)
+(* rank and width of A are equal). *)
+(* col_ebase A == the extended column basis of A (the first matrix L *)
+(* returned by Gaussian_elimination A). *)
+(* row_ebase A == the extended row base of A (the second matrix U *)
+(* returned by Gaussian_elimination A). *)
+(* col_base A == a basis for the columns of A: a row-full matrix *)
+(* consisting of the first \rank A columns of col_ebase A. *)
+(* row_base A == a basis for the rows of A: a row-free matrix consisting *)
+(* of the first \rank A rows of row_ebase A. *)
+(* pinvmx A == a partial inverse for A in its row space (or on its *)
+(* column space, equivalently). In particular, if u is a *)
+(* row vector in the row_space of A, then u *m pinvmx A is *)
+(* the row vector of the coefficients of a decomposition *)
+(* of u as a sub of rows of A. *)
+(* kermx A == the row kernel of A : a square matrix whose row space *)
+(* consists of all u such that u *m A = 0 (it consists of *)
+(* the inverse of col_ebase A, with the top \rank A rows *)
+(* zeroed out). Also, kermx A is a partial right inverse *)
+(* to col_ebase A, in the row space anihilated by A. *)
+(* cokermx A == the cokernel of A : a square matrix whose column space *)
+(* consists of all v such that A *m v = 0 (it consists of *)
+(* the inverse of row_ebase A, with the leftmost \rank A *)
+(* columns zeroed out). *)
+(* eigenvalue g a <=> a is an eigenvalue of the square matrix g. *)
+(* eigenspace g a == a square matrix whose row space is the eigenspace of *)
+(* the eigenvalue a of g (or 0 if a is not an eigenvalue). *)
+(* We use a different scope %MS for matrix row-space set-like operations; to *)
+(* avoid confusion, this scope should not be opened globally. Note that the *)
+(* the arguments of \rank _ and the operations below have default scope %MS. *)
+(* (A <= B)%MS <=> the row-space of A is included in the row-space of B. *)
+(* We test for this by testing if cokermx B anihilates A. *)
+(* (A < B)%MS <=> the row-space of A is properly included in the *)
+(* row-space of B. *)
+(* (A <= B <= C)%MS == (A <= B)%MS && (B <= C)%MS, and similarly for *)
+(* (A < B <= C)%MS, (A < B <= C)%MS and (A < B < C)%MS. *)
+(* (A == B)%MS == (A <= B <= A)%MS (A and B have the same row-space). *)
+(* (A :=: B)%MS == A and B behave identically wrt. \rank and <=. This *)
+(* triple rewrite rule is the Prop version of (A == B)%MS. *)
+(* Note that :=: cannot be treated as a setoid-style *)
+(* Equivalence because its arguments can have different *)
+(* types: A and B need not have the same number of rows, *)
+(* and often don't (e.g., in row_base A :=: A). *)
+(* <<A>>%MS == a square matrix with the same row-space as A; <<A>>%MS *)
+(* is a canonical representation of the subspace generated *)
+(* by A, viewed as a list of row-vectors: if (A == B)%MS, *)
+(* then <<A>>%MS = <<B>>%MS. *)
+(* (A + B)%MS == a square matrix whose row-space is the sum of the *)
+(* row-spaces of A and B; thus (A + B == col_mx A B)%MS. *)
+(* (\sum_i <expr i>)%MS == the "big" version of (_ + _)%MS; as the latter *)
+(* has a canonical abelian monoid structure, most generic *)
+(* bigop lemmas apply (the other bigop indexing notations *)
+(* are also defined). *)
+(* (A :&: B)%MS == a square matrix whose row-space is the intersection of *)
+(* the row-spaces of A and B. *)
+(* (\bigcap_i <expr i>)%MS == the "big" version of (_ :&: _)%MS, which also *)
+(* has a canonical abelian monoid structure. *)
+(* A^C%MS == a square matrix whose row-space is a complement to the *)
+(* the row-space of A (it consists of row_ebase A with the *)
+(* top \rank A rows zeroed out). *)
+(* (A :\: B)%MS == a square matrix whose row-space is a complement of the *)
+(* the row-space of (A :&: B)%MS in the row-space of A. *)
+(* We have (A :\: B := A :&: (capmx_gen A B)^C)%MS, where *)
+(* capmx_gen A B is a rectangular matrix equivalent to *)
+(* (A :&: B)%MS, i.e., (capmx_gen A B == A :&: B)%MS. *)
+(* proj_mx A B == a square matrix that projects (A + B)%MS onto A *)
+(* parellel to B, when (A :&: B)%MS = 0 (A and B must also *)
+(* be square). *)
+(* mxdirect S == the sum expression S is a direct sum. This is a NON *)
+(* EXTENSIONAL notation: the exact boolean expression is *)
+(* inferred from the syntactic form of S (expanding *)
+(* definitions, however); both (\sum_(i | _) _)%MS and *)
+(* (_ + _)%MS sums are recognized. This construct uses a *)
+(* variant of the reflexive ("quote") canonical structure, *)
+(* mxsum_expr. The structure also recognizes sums of *)
+(* matrix ranks, so that lemmas concerning the rank of *)
+(* direct sums can be used bidirectionally. *)
+(* The next set of definitions let us represent F-algebras using matrices: *)
+(* 'A[F]_(m, n) == the type of matrices encoding (sub)algebras of square *)
+(* n x n matrices, via mxvec; as in the matrix type *)
+(* notation, m and F can be omitted (m defaults to n ^ 2). *)
+(* := 'M[F]_(m, n ^ 2). *)
+(* (A \in R)%MS <=> the square matrix A belongs to the linear set of *)
+(* matrices (most often, a sub-algebra) encoded by the *)
+(* row space of R. This is simply notation, so all the *)
+(* lemmas and rewrite rules for (_ <= _)%MS can apply. *)
+(* := (mxvec A <= R)%MS. *)
+(* (R * S)%MS == a square n^2 x n^2 matrix whose row-space encodes the *)
+(* linear set of n x n matrices generated by the pointwise *)
+(* product of the sets of matrices encoded by R and S. *)
+(* 'C(R)%MS == a square matric encoding the centraliser of the set of *)
+(* square matrices encoded by R. *)
+(* 'C_S(R)%MS := (S :&: 'C(R))%MS (the centraliser of R in S). *)
+(* 'Z(R)%MS == the center of R (i.e., 'C_R(R)%MS). *)
+(* left_mx_ideal R S <=> S is a left ideal for R (R * S <= S)%MS. *)
+(* right_mx_ideal R S <=> S is a right ideal for R (S * R <= S)%MS. *)
+(* mx_ideal R S <=> S is a bilateral ideal for R. *)
+(* mxring_id R e <-> e is an identity element for R (Prop predicate). *)
+(* has_mxring_id R <=> R has a nonzero identity element (bool predicate). *)
+(* mxring R <=> R encodes a nontrivial subring. *)
+(*****************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+Import GRing.Theory.
+Open Local Scope ring_scope.
+
+Reserved Notation "\rank A" (at level 10, A at level 8, format "\rank A").
+Reserved Notation "A ^C" (at level 8, format "A ^C").
+
+Notation "''A_' ( m , n )" := 'M_(m, n ^ 2)
+ (at level 8, format "''A_' ( m , n )") : type_scope.
+
+Notation "''A_' ( n )" := 'A_(n ^ 2, n)
+ (at level 8, only parsing) : type_scope.
+
+Notation "''A_' n" := 'A_(n)
+ (at level 8, n at next level, format "''A_' n") : type_scope.
+
+Notation "''A' [ F ]_ ( m , n )" := 'M[F]_(m, n ^ 2)
+ (at level 8, only parsing) : type_scope.
+
+Notation "''A' [ F ]_ ( n )" := 'A[F]_(n ^ 2, n)
+ (at level 8, only parsing) : type_scope.
+
+Notation "''A' [ F ]_ n" := 'A[F]_(n)
+ (at level 8, n at level 2, only parsing) : type_scope.
+
+Delimit Scope matrix_set_scope with MS.
+
+Notation Local simp := (Monoid.Theory.simpm, oppr0).
+
+(*****************************************************************************)
+(******************** Rank and row-space theory ******************************)
+(*****************************************************************************)
+
+Section RowSpaceTheory.
+
+Variable F : fieldType.
+Implicit Types m n p r : nat.
+
+Local Notation "''M_' ( m , n )" := 'M[F]_(m, n) : type_scope.
+Local Notation "''M_' n" := 'M[F]_(n, n) : type_scope.
+
+(* Decomposition with double pivoting; computes the rank, row and column *)
+(* images, kernels, and complements of a matrix. *)
+
+Fixpoint Gaussian_elimination {m n} : 'M_(m, n) -> 'M_m * 'M_n * nat :=
+ match m, n with
+ | _.+1, _.+1 => fun A : 'M_(1 + _, 1 + _) =>
+ if [pick ij | A ij.1 ij.2 != 0] is Some (i, j) then
+ let a := A i j in let A1 := xrow i 0 (xcol j 0 A) in
+ let u := ursubmx A1 in let v := a^-1 *: dlsubmx A1 in
+ let: (L, U, r) := Gaussian_elimination (drsubmx A1 - v *m u) in
+ (xrow i 0 (block_mx 1 0 v L), xcol j 0 (block_mx a%:M u 0 U), r.+1)
+ else (1%:M, 1%:M, 0%N)
+ | _, _ => fun _ => (1%:M, 1%:M, 0%N)
+ end.
+
+Section Defs.
+
+Variables (m n : nat) (A : 'M_(m, n)).
+
+Fact Gaussian_elimination_key : unit. Proof. by []. Qed.
+
+Let LUr := locked_with Gaussian_elimination_key (@Gaussian_elimination) m n A.
+
+Definition col_ebase := LUr.1.1.
+Definition row_ebase := LUr.1.2.
+Definition mxrank := if [|| m == 0 | n == 0]%N then 0%N else LUr.2.
+
+Definition row_free := mxrank == m.
+Definition row_full := mxrank == n.
+
+Definition row_base : 'M_(mxrank, n) := pid_mx mxrank *m row_ebase.
+Definition col_base : 'M_(m, mxrank) := col_ebase *m pid_mx mxrank.
+
+Definition complmx : 'M_n := copid_mx mxrank *m row_ebase.
+Definition kermx : 'M_m := copid_mx mxrank *m invmx col_ebase.
+Definition cokermx : 'M_n := invmx row_ebase *m copid_mx mxrank.
+
+Definition pinvmx : 'M_(n, m) :=
+ invmx row_ebase *m pid_mx mxrank *m invmx col_ebase.
+
+End Defs.
+
+Arguments Scope mxrank [nat_scope nat_scope matrix_set_scope].
+Local Notation "\rank A" := (mxrank A) : nat_scope.
+Arguments Scope complmx [nat_scope nat_scope matrix_set_scope].
+Local Notation "A ^C" := (complmx A) : matrix_set_scope.
+
+Definition submx_def := idfun (fun m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) =>
+ A *m cokermx B == 0).
+Fact submx_key : unit. Proof. by []. Qed.
+Definition submx := locked_with submx_key submx_def.
+Canonical submx_unlockable := [unlockable fun submx].
+
+Arguments Scope submx
+ [nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Prenex Implicits submx.
+Local Notation "A <= B" := (submx A B) : matrix_set_scope.
+Local Notation "A <= B <= C" := ((A <= B) && (B <= C))%MS : matrix_set_scope.
+Local Notation "A == B" := (A <= B <= A)%MS : matrix_set_scope.
+
+Definition ltmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :=
+ (A <= B)%MS && ~~ (B <= A)%MS.
+Arguments Scope ltmx
+ [nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Prenex Implicits ltmx.
+Local Notation "A < B" := (ltmx A B) : matrix_set_scope.
+
+Definition eqmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :=
+ prod (\rank A = \rank B)
+ (forall m3 (C : 'M_(m3, n)),
+ ((A <= C) = (B <= C)) * ((C <= A) = (C <= B)))%MS.
+Arguments Scope eqmx
+ [nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Local Notation "A :=: B" := (eqmx A B) : matrix_set_scope.
+
+Section LtmxIdentities.
+
+Variables (m1 m2 n : nat) (A : 'M_(m1, n)) (B : 'M_(m2, n)).
+
+Lemma ltmxE : (A < B)%MS = ((A <= B)%MS && ~~ (B <= A)%MS). Proof. by []. Qed.
+
+Lemma ltmxW : (A < B)%MS -> (A <= B)%MS. Proof. by case/andP. Qed.
+
+Lemma ltmxEneq : (A < B)%MS = (A <= B)%MS && ~~ (A == B)%MS.
+Proof. by apply: andb_id2l => ->. Qed.
+
+Lemma submxElt : (A <= B)%MS = (A == B)%MS || (A < B)%MS.
+Proof. by rewrite -andb_orr orbN andbT. Qed.
+
+End LtmxIdentities.
+
+(* The definition of the row-space operator is rigged to return the identity *)
+(* matrix for full matrices. To allow for further tweaks that will make the *)
+(* row-space intersection operator strictly commutative and monoidal, we *)
+(* slightly generalize some auxiliary definitions: we parametrize the *)
+(* "equivalent subspace and identity" choice predicate equivmx by a boolean *)
+(* determining whether the matrix should be the identity (so for genmx A its *)
+(* value is row_full A), and introduce a "quasi-identity" predicate qidmx *)
+(* that selects non-square full matrices along with the identity matrix 1%:M *)
+(* (this does not affect genmx, which chooses a square matrix). *)
+(* The choice witness for genmx A is either 1%:M for a row-full A, or else *)
+(* row_base A padded with null rows. *)
+Let qidmx m n (A : 'M_(m, n)) :=
+ if m == n then A == pid_mx n else row_full A.
+Let equivmx m n (A : 'M_(m, n)) idA (B : 'M_n) :=
+ (B == A)%MS && (qidmx B == idA).
+Let equivmx_spec m n (A : 'M_(m, n)) idA (B : 'M_n) :=
+ prod (B :=: A)%MS (qidmx B = idA).
+Definition genmx_witness m n (A : 'M_(m, n)) : 'M_n :=
+ if row_full A then 1%:M else pid_mx (\rank A) *m row_ebase A.
+Definition genmx_def := idfun (fun m n (A : 'M_(m, n)) =>
+ choose (equivmx A (row_full A)) (genmx_witness A) : 'M_n).
+Fact genmx_key : unit. Proof. by []. Qed.
+Definition genmx := locked_with genmx_key genmx_def.
+Canonical genmx_unlockable := [unlockable fun genmx].
+Local Notation "<< A >>" := (genmx A) : matrix_set_scope.
+
+(* The setwise sum is tweaked so that 0 is a strict identity element for *)
+(* square matrices, because this lets us use the bigop component. As a result *)
+(* setwise sum is not quite strictly extensional. *)
+Let addsmx_nop m n (A : 'M_(m, n)) := conform_mx <<A>>%MS A.
+Definition addsmx_def := idfun (fun m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) =>
+ if A == 0 then addsmx_nop B else if B == 0 then addsmx_nop A else
+ <<col_mx A B>>%MS : 'M_n).
+Fact addsmx_key : unit. Proof. by []. Qed.
+Definition addsmx := locked_with addsmx_key addsmx_def.
+Canonical addsmx_unlockable := [unlockable fun addsmx].
+Arguments Scope addsmx
+ [nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Prenex Implicits addsmx.
+Local Notation "A + B" := (addsmx A B) : matrix_set_scope.
+Local Notation "\sum_ ( i | P ) B" := (\big[addsmx/0]_(i | P) B%MS)
+ : matrix_set_scope.
+Local Notation "\sum_ ( i <- r | P ) B" := (\big[addsmx/0]_(i <- r | P) B%MS)
+ : matrix_set_scope.
+
+(* The set intersection is similarly biased so that the identity matrix is a *)
+(* strict identity. This is somewhat more delicate than for the sum, because *)
+(* the test for the identity is non-extensional. This forces us to actually *)
+(* bias the choice operator so that it does not accidentally map an *)
+(* intersection of non-identity matrices to 1%:M; this would spoil *)
+(* associativity: if B :&: C = 1%:M but B and C are not identity, then for a *)
+(* square matrix A we have A :&: (B :&: C) = A != (A :&: B) :&: C in general. *)
+(* To complicate matters there may not be a square non-singular matrix *)
+(* different than 1%:M, since we could be dealing with 'M['F_2]_1. We *)
+(* sidestep the issue by making all non-square row-full matrices identities, *)
+(* and choosing a normal representative that preserves the qidmx property. *)
+(* Thus A :&: B = 1%:M iff A and B are both identities, and this suffices for *)
+(* showing that associativity is strict. *)
+Let capmx_witness m n (A : 'M_(m, n)) :=
+ if row_full A then conform_mx 1%:M A else <<A>>%MS.
+Let capmx_norm m n (A : 'M_(m, n)) :=
+ choose (equivmx A (qidmx A)) (capmx_witness A).
+Let capmx_nop m n (A : 'M_(m, n)) := conform_mx (capmx_norm A) A.
+Definition capmx_gen m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :=
+ lsubmx (kermx (col_mx A B)) *m A.
+Definition capmx_def := idfun (fun m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) =>
+ if qidmx A then capmx_nop B else
+ if qidmx B then capmx_nop A else
+ if row_full B then capmx_norm A else capmx_norm (capmx_gen A B) : 'M_n).
+Fact capmx_key : unit. Proof. by []. Qed.
+Definition capmx := locked_with capmx_key capmx_def.
+Canonical capmx_unlockable := [unlockable fun capmx].
+Arguments Scope capmx
+ [nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Prenex Implicits capmx.
+Local Notation "A :&: B" := (capmx A B) : matrix_set_scope.
+Local Notation "\bigcap_ ( i | P ) B" := (\big[capmx/1%:M]_(i | P) B)
+ : matrix_set_scope.
+
+Definition diffmx_def := idfun (fun m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) =>
+ <<capmx_gen A (capmx_gen A B)^C>>%MS : 'M_n).
+Fact diffmx_key : unit. Proof. by []. Qed.
+Definition diffmx := locked_with diffmx_key diffmx_def.
+Canonical diffmx_unlockable := [unlockable fun diffmx].
+Arguments Scope diffmx
+ [nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Prenex Implicits diffmx.
+Local Notation "A :\: B" := (diffmx A B) : matrix_set_scope.
+
+Definition proj_mx n (U V : 'M_n) : 'M_n := pinvmx (col_mx U V) *m col_mx U 0.
+
+Local Notation GaussE := Gaussian_elimination.
+
+Fact mxrankE m n (A : 'M_(m, n)) : \rank A = (GaussE A).2.
+Proof. by rewrite /mxrank unlock /=; case: m n A => [|m] [|n]. Qed.
+
+Lemma rank_leq_row m n (A : 'M_(m, n)) : \rank A <= m.
+Proof.
+rewrite mxrankE.
+elim: m n A => [|m IHm] [|n] //= A; case: pickP => [[i j] _|] //=.
+by move: (_ - _) => B; case: GaussE (IHm _ B) => [[L U] r] /=.
+Qed.
+
+Lemma row_leq_rank m n (A : 'M_(m, n)) : (m <= \rank A) = row_free A.
+Proof. by rewrite /row_free eqn_leq rank_leq_row. Qed.
+
+Lemma rank_leq_col m n (A : 'M_(m, n)) : \rank A <= n.
+Proof.
+rewrite mxrankE.
+elim: m n A => [|m IHm] [|n] //= A; case: pickP => [[i j] _|] //=.
+by move: (_ - _) => B; case: GaussE (IHm _ B) => [[L U] r] /=.
+Qed.
+
+Lemma col_leq_rank m n (A : 'M_(m, n)) : (n <= \rank A) = row_full A.
+Proof. by rewrite /row_full eqn_leq rank_leq_col. Qed.
+
+Let unitmx1F := @unitmx1 F.
+Lemma row_ebase_unit m n (A : 'M_(m, n)) : row_ebase A \in unitmx.
+Proof.
+rewrite /row_ebase unlock; elim: m n A => [|m IHm] [|n] //= A.
+case: pickP => [[i j] /= nzAij | //=]; move: (_ - _) => B.
+case: GaussE (IHm _ B) => [[L U] r] /= uU.
+rewrite unitmxE xcolE det_mulmx (@det_ublock _ 1) det_scalar1 !unitrM.
+by rewrite unitfE nzAij -!unitmxE uU unitmx_perm.
+Qed.
+
+Lemma col_ebase_unit m n (A : 'M_(m, n)) : col_ebase A \in unitmx.
+Proof.
+rewrite /col_ebase unlock; elim: m n A => [|m IHm] [|n] //= A.
+case: pickP => [[i j] _|] //=; move: (_ - _) => B.
+case: GaussE (IHm _ B) => [[L U] r] /= uL.
+rewrite unitmxE xrowE det_mulmx (@det_lblock _ 1) det1 mul1r unitrM.
+by rewrite -unitmxE unitmx_perm.
+Qed.
+Hint Resolve rank_leq_row rank_leq_col row_ebase_unit col_ebase_unit.
+
+Lemma mulmx_ebase m n (A : 'M_(m, n)) :
+ col_ebase A *m pid_mx (\rank A) *m row_ebase A = A.
+Proof.
+rewrite mxrankE /col_ebase /row_ebase unlock.
+elim: m n A => [n A | m IHm]; first by rewrite [A]flatmx0 [_ *m _]flatmx0.
+case=> [A | n]; first by rewrite [_ *m _]thinmx0 [A]thinmx0.
+rewrite -(add1n m) -?(add1n n) => A /=.
+case: pickP => [[i0 j0] | A0] /=; last first.
+ apply/matrixP=> i j; rewrite pid_mx_0 mulmx0 mul0mx mxE.
+ by move/eqP: (A0 (i, j)).
+set a := A i0 j0 => nz_a; set A1 := xrow _ _ _.
+set u := ursubmx _; set v := _ *: _; set B : 'M_(m, n) := _ - _.
+move: (rank_leq_col B) (rank_leq_row B) {IHm}(IHm n B); rewrite mxrankE.
+case: (GaussE B) => [[L U] r] /= r_m r_n defB.
+have ->: pid_mx (1 + r) = block_mx 1 0 0 (pid_mx r) :> 'M[F]_(1 + m, 1 + n).
+ rewrite -(subnKC r_m) -(subnKC r_n) pid_mx_block -col_mx0 -row_mx0.
+ by rewrite block_mxA castmx_id col_mx0 row_mx0 -scalar_mx_block -pid_mx_block.
+rewrite xcolE xrowE mulmxA -xcolE -!mulmxA.
+rewrite !(addr0, add0r, mulmx0, mul0mx, mulmx_block, mul1mx) mulmxA defB.
+rewrite addrC subrK mul_mx_scalar scalerA divff // scale1r.
+have ->: a%:M = ulsubmx A1 by rewrite [_ A1]mx11_scalar !mxE !lshift0 !tpermR.
+rewrite submxK /A1 xrowE !xcolE -!mulmxA mulmxA -!perm_mxM !tperm2 !perm_mx1.
+by rewrite mulmx1 mul1mx.
+Qed.
+
+Lemma mulmx_base m n (A : 'M_(m, n)) : col_base A *m row_base A = A.
+Proof. by rewrite mulmxA -[col_base A *m _]mulmxA pid_mx_id ?mulmx_ebase. Qed.
+
+Lemma mulmx1_min_rank r m n (A : 'M_(m, n)) M N :
+ M *m A *m N = 1%:M :> 'M_r -> r <= \rank A.
+Proof. by rewrite -{1}(mulmx_base A) mulmxA -mulmxA; move/mulmx1_min. Qed.
+Implicit Arguments mulmx1_min_rank [r m n A].
+
+Lemma mulmx_max_rank r m n (M : 'M_(m, r)) (N : 'M_(r, n)) :
+ \rank (M *m N) <= r.
+Proof.
+set MN := M *m N; set rMN := \rank _.
+pose L : 'M_(rMN, m) := pid_mx rMN *m invmx (col_ebase MN).
+pose U : 'M_(n, rMN) := invmx (row_ebase MN) *m pid_mx rMN.
+suffices: L *m M *m (N *m U) = 1%:M by exact: mulmx1_min.
+rewrite mulmxA -(mulmxA L) -[M *m N]mulmx_ebase -/MN.
+by rewrite !mulmxA mulmxKV // mulmxK // !pid_mx_id /rMN ?pid_mx_1.
+Qed.
+Implicit Arguments mulmx_max_rank [r m n].
+
+Lemma mxrank_tr m n (A : 'M_(m, n)) : \rank A^T = \rank A.
+Proof.
+apply/eqP; rewrite eqn_leq -{3}[A]trmxK -{1}(mulmx_base A) -{1}(mulmx_base A^T).
+by rewrite !trmx_mul !mulmx_max_rank.
+Qed.
+
+Lemma mxrank_add m n (A B : 'M_(m, n)) : \rank (A + B)%R <= \rank A + \rank B.
+Proof.
+by rewrite -{1}(mulmx_base A) -{1}(mulmx_base B) -mul_row_col mulmx_max_rank.
+Qed.
+
+Lemma mxrankM_maxl m n p (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ \rank (A *m B) <= \rank A.
+Proof. by rewrite -{1}(mulmx_base A) -mulmxA mulmx_max_rank. Qed.
+
+Lemma mxrankM_maxr m n p (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ \rank (A *m B) <= \rank B.
+Proof. by rewrite -mxrank_tr -(mxrank_tr B) trmx_mul mxrankM_maxl. Qed.
+
+Lemma mxrank_scale m n a (A : 'M_(m, n)) : \rank (a *: A) <= \rank A.
+Proof. by rewrite -mul_scalar_mx mxrankM_maxr. Qed.
+
+Lemma mxrank_scale_nz m n a (A : 'M_(m, n)) :
+ a != 0 -> \rank (a *: A) = \rank A.
+Proof.
+move=> nza; apply/eqP; rewrite eqn_leq -{3}[A]scale1r -(mulVf nza).
+by rewrite -scalerA !mxrank_scale.
+Qed.
+
+Lemma mxrank_opp m n (A : 'M_(m, n)) : \rank (- A) = \rank A.
+Proof. by rewrite -scaleN1r mxrank_scale_nz // oppr_eq0 oner_eq0. Qed.
+
+Lemma mxrank0 m n : \rank (0 : 'M_(m, n)) = 0%N.
+Proof. by apply/eqP; rewrite -leqn0 -(@mulmx0 _ m 0 n 0) mulmx_max_rank. Qed.
+
+Lemma mxrank_eq0 m n (A : 'M_(m, n)) : (\rank A == 0%N) = (A == 0).
+Proof.
+apply/eqP/eqP=> [rA0 | ->{A}]; last exact: mxrank0.
+move: (col_base A) (row_base A) (mulmx_base A); rewrite rA0 => Ac Ar <-.
+by rewrite [Ac]thinmx0 mul0mx.
+Qed.
+
+Lemma mulmx_coker m n (A : 'M_(m, n)) : A *m cokermx A = 0.
+Proof.
+by rewrite -{1}[A]mulmx_ebase -!mulmxA mulKVmx // mul_pid_mx_copid ?mulmx0.
+Qed.
+
+Lemma submxE m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A <= B)%MS = (A *m cokermx B == 0).
+Proof. by rewrite unlock. Qed.
+
+Lemma mulmxKpV m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A <= B)%MS -> A *m pinvmx B *m B = A.
+Proof.
+rewrite submxE !mulmxA mulmxBr mulmx1 subr_eq0 => /eqP defA.
+rewrite -{4}[B]mulmx_ebase -!mulmxA mulKmx //.
+by rewrite (mulmxA (pid_mx _)) pid_mx_id // !mulmxA -{}defA mulmxKV.
+Qed.
+
+Lemma submxP m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (exists D, A = D *m B) (A <= B)%MS.
+Proof.
+apply: (iffP idP) => [/mulmxKpV | [D ->]]; first by exists (A *m pinvmx B).
+by rewrite submxE -mulmxA mulmx_coker mulmx0.
+Qed.
+Implicit Arguments submxP [m1 m2 n A B].
+
+Lemma submx_refl m n (A : 'M_(m, n)) : (A <= A)%MS.
+Proof. by rewrite submxE mulmx_coker. Qed.
+Hint Resolve submx_refl.
+
+Lemma submxMl m n p (D : 'M_(m, n)) (A : 'M_(n, p)) : (D *m A <= A)%MS.
+Proof. by rewrite submxE -mulmxA mulmx_coker mulmx0. Qed.
+
+Lemma submxMr m1 m2 n p (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(n, p)) :
+ (A <= B)%MS -> (A *m C <= B *m C)%MS.
+Proof. by case/submxP=> D ->; rewrite -mulmxA submxMl. Qed.
+
+Lemma mulmx_sub m n1 n2 p (C : 'M_(m, n1)) A (B : 'M_(n2, p)) :
+ (A <= B -> C *m A <= B)%MS.
+Proof. by case/submxP=> D ->; rewrite mulmxA submxMl. Qed.
+
+Lemma submx_trans m1 m2 m3 n
+ (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ (A <= B -> B <= C -> A <= C)%MS.
+Proof. by case/submxP=> D ->{A}; exact: mulmx_sub. Qed.
+
+Lemma ltmx_sub_trans m1 m2 m3 n
+ (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ (A < B)%MS -> (B <= C)%MS -> (A < C)%MS.
+Proof.
+case/andP=> sAB ltAB sBC; rewrite ltmxE (submx_trans sAB) //.
+by apply: contra ltAB; exact: submx_trans.
+Qed.
+
+Lemma sub_ltmx_trans m1 m2 m3 n
+ (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ (A <= B)%MS -> (B < C)%MS -> (A < C)%MS.
+Proof.
+move=> sAB /andP[sBC ltBC]; rewrite ltmxE (submx_trans sAB) //.
+by apply: contra ltBC => sCA; exact: submx_trans sAB.
+Qed.
+
+Lemma ltmx_trans m n : transitive (@ltmx m m n).
+Proof. by move=> A B C; move/ltmxW; exact: sub_ltmx_trans. Qed.
+
+Lemma ltmx_irrefl m n : irreflexive (@ltmx m m n).
+Proof. by move=> A; rewrite /ltmx submx_refl andbF. Qed.
+
+Lemma sub0mx m1 m2 n (A : 'M_(m2, n)) : ((0 : 'M_(m1, n)) <= A)%MS.
+Proof. by rewrite submxE mul0mx. Qed.
+
+Lemma submx0null m1 m2 n (A : 'M[F]_(m1, n)) :
+ (A <= (0 : 'M_(m2, n)))%MS -> A = 0.
+Proof. by case/submxP=> D; rewrite mulmx0. Qed.
+
+Lemma submx0 m n (A : 'M_(m, n)) : (A <= (0 : 'M_n))%MS = (A == 0).
+Proof. by apply/idP/eqP=> [|->]; [exact: submx0null | exact: sub0mx]. Qed.
+
+Lemma lt0mx m n (A : 'M_(m, n)) : ((0 : 'M_n) < A)%MS = (A != 0).
+Proof. by rewrite /ltmx sub0mx submx0. Qed.
+
+Lemma ltmx0 m n (A : 'M[F]_(m, n)) : (A < (0 : 'M_n))%MS = false.
+Proof. by rewrite /ltmx sub0mx andbF. Qed.
+
+Lemma eqmx0P m n (A : 'M_(m, n)) : reflect (A = 0) (A == (0 : 'M_n))%MS.
+Proof. by rewrite submx0 sub0mx andbT; exact: eqP. Qed.
+
+Lemma eqmx_eq0 m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A :=: B)%MS -> (A == 0) = (B == 0).
+Proof. by move=> eqAB; rewrite -!submx0 eqAB. Qed.
+
+Lemma addmx_sub m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m1, n)) (C : 'M_(m2, n)) :
+ (A <= C)%MS -> (B <= C)%MS -> ((A + B)%R <= C)%MS.
+Proof.
+by case/submxP=> A' ->; case/submxP=> B' ->; rewrite -mulmxDl submxMl.
+Qed.
+
+Lemma summx_sub m1 m2 n (B : 'M_(m2, n))
+ I (r : seq I) (P : pred I) (A_ : I -> 'M_(m1, n)) :
+ (forall i, P i -> A_ i <= B)%MS -> ((\sum_(i <- r | P i) A_ i)%R <= B)%MS.
+Proof.
+move=> leAB; elim/big_ind: _ => // [|A1 A2]; [exact: sub0mx | exact: addmx_sub].
+Qed.
+
+Lemma scalemx_sub m1 m2 n a (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A <= B)%MS -> (a *: A <= B)%MS.
+Proof. by case/submxP=> A' ->; rewrite scalemxAl submxMl. Qed.
+
+Lemma row_sub m n i (A : 'M_(m, n)) : (row i A <= A)%MS.
+Proof. by rewrite rowE submxMl. Qed.
+
+Lemma eq_row_sub m n v (A : 'M_(m, n)) i : row i A = v -> (v <= A)%MS.
+Proof. by move <-; rewrite row_sub. Qed.
+
+Lemma nz_row_sub m n (A : 'M_(m, n)) : (nz_row A <= A)%MS.
+Proof. by rewrite /nz_row; case: pickP => [i|] _; rewrite ?row_sub ?sub0mx. Qed.
+
+Lemma row_subP m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (forall i, row i A <= B)%MS (A <= B)%MS.
+Proof.
+apply: (iffP idP) => [sAB i|sAB].
+ by apply: submx_trans sAB; exact: row_sub.
+rewrite submxE; apply/eqP/row_matrixP=> i; apply/eqP.
+by rewrite row_mul row0 -submxE.
+Qed.
+Implicit Arguments row_subP [m1 m2 n A B].
+
+Lemma rV_subP m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (forall v : 'rV_n, v <= A -> v <= B)%MS (A <= B)%MS.
+Proof.
+apply: (iffP idP) => [sAB v Av | sAB]; first exact: submx_trans sAB.
+by apply/row_subP=> i; rewrite sAB ?row_sub.
+Qed.
+Implicit Arguments rV_subP [m1 m2 n A B].
+
+Lemma row_subPn m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (exists i, ~~ (row i A <= B)%MS) (~~ (A <= B)%MS).
+Proof. by rewrite (sameP row_subP forallP) negb_forall; exact: existsP. Qed.
+
+Lemma sub_rVP n (u v : 'rV_n) : reflect (exists a, u = a *: v) (u <= v)%MS.
+Proof.
+apply: (iffP submxP) => [[w ->] | [a ->]].
+ by exists (w 0 0); rewrite -mul_scalar_mx -mx11_scalar.
+by exists a%:M; rewrite mul_scalar_mx.
+Qed.
+
+Lemma rank_rV n (v : 'rV_n) : \rank v = (v != 0).
+Proof.
+case: eqP => [-> | nz_v]; first by rewrite mxrank0.
+by apply/eqP; rewrite eqn_leq rank_leq_row lt0n mxrank_eq0; exact/eqP.
+Qed.
+
+Lemma rowV0Pn m n (A : 'M_(m, n)) :
+ reflect (exists2 v : 'rV_n, v <= A & v != 0)%MS (A != 0).
+Proof.
+rewrite -submx0; apply: (iffP idP) => [| [v svA]]; last first.
+ by rewrite -submx0; exact: contra (submx_trans _).
+by case/row_subPn=> i; rewrite submx0; exists (row i A); rewrite ?row_sub.
+Qed.
+
+Lemma rowV0P m n (A : 'M_(m, n)) :
+ reflect (forall v : 'rV_n, v <= A -> v = 0)%MS (A == 0).
+Proof.
+rewrite -[A == 0]negbK; case: rowV0Pn => IH.
+ by right; case: IH => v svA nzv IH; case/eqP: nzv; exact: IH.
+by left=> v svA; apply/eqP; apply/idPn=> nzv; case: IH; exists v.
+Qed.
+
+Lemma submx_full m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ row_full B -> (A <= B)%MS.
+Proof.
+by rewrite submxE /cokermx =>/eqnP->; rewrite /copid_mx pid_mx_1 subrr !mulmx0.
+Qed.
+
+Lemma row_fullP m n (A : 'M_(m, n)) :
+ reflect (exists B, B *m A = 1%:M) (row_full A).
+Proof.
+apply: (iffP idP) => [Afull | [B kA]].
+ by exists (1%:M *m pinvmx A); apply: mulmxKpV (submx_full _ Afull).
+by rewrite [_ A]eqn_leq rank_leq_col (mulmx1_min_rank B 1%:M) ?mulmx1.
+Qed.
+Implicit Arguments row_fullP [m n A].
+
+Lemma row_full_inj m n p A : row_full A -> injective (@mulmx _ m n p A).
+Proof.
+case/row_fullP=> A' A'K; apply: can_inj (mulmx A') _ => B.
+by rewrite mulmxA A'K mul1mx.
+Qed.
+
+Lemma row_freeP m n (A : 'M_(m, n)) :
+ reflect (exists B, A *m B = 1%:M) (row_free A).
+Proof.
+rewrite /row_free -mxrank_tr.
+apply: (iffP row_fullP) => [] [B kA];
+ by exists B^T; rewrite -trmx1 -kA trmx_mul ?trmxK.
+Qed.
+
+Lemma row_free_inj m n p A : row_free A -> injective ((@mulmx _ m n p)^~ A).
+Proof.
+case/row_freeP=> A' AK; apply: can_inj (mulmx^~ A') _ => B.
+by rewrite -mulmxA AK mulmx1.
+Qed.
+
+Lemma row_free_unit n (A : 'M_n) : row_free A = (A \in unitmx).
+Proof.
+apply/row_fullP/idP=> [[A'] | uA]; first by case/mulmx1_unit.
+by exists (invmx A); rewrite mulVmx.
+Qed.
+
+Lemma row_full_unit n (A : 'M_n) : row_full A = (A \in unitmx).
+Proof. exact: row_free_unit. Qed.
+
+Lemma mxrank_unit n (A : 'M_n) : A \in unitmx -> \rank A = n.
+Proof. by rewrite -row_full_unit =>/eqnP. Qed.
+
+Lemma mxrank1 n : \rank (1%:M : 'M_n) = n.
+Proof. by apply: mxrank_unit; exact: unitmx1. Qed.
+
+Lemma mxrank_delta m n i j : \rank (delta_mx i j : 'M_(m, n)) = 1%N.
+Proof.
+apply/eqP; rewrite eqn_leq lt0n mxrank_eq0.
+rewrite -{1}(mul_delta_mx (0 : 'I_1)) mulmx_max_rank.
+by apply/eqP; move/matrixP; move/(_ i j); move/eqP; rewrite !mxE !eqxx oner_eq0.
+Qed.
+
+Lemma mxrankS m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A <= B)%MS -> \rank A <= \rank B.
+Proof. by case/submxP=> D ->; rewrite mxrankM_maxr. Qed.
+
+Lemma submx1 m n (A : 'M_(m, n)) : (A <= 1%:M)%MS.
+Proof. by rewrite submx_full // row_full_unit unitmx1. Qed.
+
+Lemma sub1mx m n (A : 'M_(m, n)) : (1%:M <= A)%MS = row_full A.
+Proof.
+apply/idP/idP; last exact: submx_full.
+by move/mxrankS; rewrite mxrank1 col_leq_rank.
+Qed.
+
+Lemma ltmx1 m n (A : 'M_(m, n)) : (A < 1%:M)%MS = ~~ row_full A.
+Proof. by rewrite /ltmx sub1mx submx1. Qed.
+
+Lemma lt1mx m n (A : 'M_(m, n)) : (1%:M < A)%MS = false.
+Proof. by rewrite /ltmx submx1 andbF. Qed.
+
+Lemma eqmxP m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (A :=: B)%MS (A == B)%MS.
+Proof.
+apply: (iffP andP) => [[sAB sBA] | eqAB]; last by rewrite !eqAB.
+split=> [|m3 C]; first by apply/eqP; rewrite eqn_leq !mxrankS.
+split; first by apply/idP/idP; exact: submx_trans.
+by apply/idP/idP=> sC; exact: submx_trans sC _.
+Qed.
+Implicit Arguments eqmxP [m1 m2 n A B].
+
+Lemma rV_eqP m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (forall u : 'rV_n, (u <= A) = (u <= B))%MS (A == B)%MS.
+Proof.
+apply: (iffP idP) => [eqAB u | eqAB]; first by rewrite (eqmxP eqAB).
+by apply/andP; split; apply/rV_subP=> u; rewrite eqAB.
+Qed.
+
+Lemma eqmx_refl m1 n (A : 'M_(m1, n)) : (A :=: A)%MS.
+Proof. by []. Qed.
+
+Lemma eqmx_sym m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A :=: B)%MS -> (B :=: A)%MS.
+Proof. by move=> eqAB; split=> [|m3 C]; rewrite !eqAB. Qed.
+
+Lemma eqmx_trans m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ (A :=: B)%MS -> (B :=: C)%MS -> (A :=: C)%MS.
+Proof. by move=> eqAB eqBC; split=> [|m4 D]; rewrite !eqAB !eqBC. Qed.
+
+Lemma eqmx_rank m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A == B)%MS -> \rank A = \rank B.
+Proof. by move/eqmxP->. Qed.
+
+Lemma lt_eqmx m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A :=: B)%MS ->
+ forall C : 'M_(m3, n), (((A < C) = (B < C))%MS * ((C < A) = (C < B))%MS)%type.
+Proof. by move=> eqAB C; rewrite /ltmx !eqAB. Qed.
+
+Lemma eqmxMr m1 m2 n p (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(n, p)) :
+ (A :=: B)%MS -> (A *m C :=: B *m C)%MS.
+Proof. by move=> eqAB; apply/eqmxP; rewrite !submxMr ?eqAB. Qed.
+
+Lemma eqmxMfull m n p (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ row_full A -> (A *m B :=: B)%MS.
+Proof.
+case/row_fullP=> A' A'A; apply/eqmxP; rewrite submxMl /=.
+by apply/submxP; exists A'; rewrite mulmxA A'A mul1mx.
+Qed.
+
+Lemma eqmx0 m n : ((0 : 'M[F]_(m, n)) :=: (0 : 'M_n))%MS.
+Proof. by apply/eqmxP; rewrite !sub0mx. Qed.
+
+Lemma eqmx_scale m n a (A : 'M_(m, n)) : a != 0 -> (a *: A :=: A)%MS.
+Proof.
+move=> nz_a; apply/eqmxP; rewrite scalemx_sub //.
+by rewrite -{1}[A]scale1r -(mulVf nz_a) -scalerA scalemx_sub.
+Qed.
+
+Lemma eqmx_opp m n (A : 'M_(m, n)) : (- A :=: A)%MS.
+Proof.
+by rewrite -scaleN1r; apply: eqmx_scale => //; rewrite oppr_eq0 oner_eq0.
+Qed.
+
+Lemma submxMfree m1 m2 n p (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(n, p)) :
+ row_free C -> (A *m C <= B *m C)%MS = (A <= B)%MS.
+Proof.
+case/row_freeP=> C' C_C'_1; apply/idP/idP=> sAB; last exact: submxMr.
+by rewrite -[A]mulmx1 -[B]mulmx1 -C_C'_1 !mulmxA submxMr.
+Qed.
+
+Lemma eqmxMfree m1 m2 n p (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(n, p)) :
+ row_free C -> (A *m C :=: B *m C)%MS -> (A :=: B)%MS.
+Proof.
+by move=> Cfree eqAB; apply/eqmxP; move/eqmxP: eqAB; rewrite !submxMfree.
+Qed.
+
+Lemma mxrankMfree m n p (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ row_free B -> \rank (A *m B) = \rank A.
+Proof.
+by move=> Bfree; rewrite -mxrank_tr trmx_mul eqmxMfull /row_full mxrank_tr.
+Qed.
+
+Lemma eq_row_base m n (A : 'M_(m, n)) : (row_base A :=: A)%MS.
+Proof.
+apply/eqmxP; apply/andP; split; apply/submxP.
+ exists (pid_mx (\rank A) *m invmx (col_ebase A)).
+ by rewrite -{8}[A]mulmx_ebase !mulmxA mulmxKV // pid_mx_id.
+exists (col_ebase A *m pid_mx (\rank A)).
+by rewrite mulmxA -(mulmxA _ _ (pid_mx _)) pid_mx_id // mulmx_ebase.
+Qed.
+
+Let qidmx_eq1 n (A : 'M_n) : qidmx A = (A == 1%:M).
+Proof. by rewrite /qidmx eqxx pid_mx_1. Qed.
+
+Let genmx_witnessP m n (A : 'M_(m, n)) :
+ equivmx A (row_full A) (genmx_witness A).
+Proof.
+rewrite /equivmx qidmx_eq1 /genmx_witness.
+case fullA: (row_full A); first by rewrite eqxx sub1mx submx1 fullA.
+set B := _ *m _; have defB : (B == A)%MS.
+ apply/andP; split; apply/submxP.
+ exists (pid_mx (\rank A) *m invmx (col_ebase A)).
+ by rewrite -{3}[A]mulmx_ebase !mulmxA mulmxKV // pid_mx_id.
+ exists (col_ebase A *m pid_mx (\rank A)).
+ by rewrite mulmxA -(mulmxA _ _ (pid_mx _)) pid_mx_id // mulmx_ebase.
+rewrite defB -negb_add addbF; case: eqP defB => // ->.
+by rewrite sub1mx fullA.
+Qed.
+
+Lemma genmxE m n (A : 'M_(m, n)) : (<<A>> :=: A)%MS.
+Proof.
+by rewrite unlock; apply/eqmxP; case/andP: (chooseP (genmx_witnessP A)).
+Qed.
+
+Lemma eq_genmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A :=: B -> <<A>> = <<B>>)%MS.
+Proof.
+move=> eqAB; rewrite unlock.
+have{eqAB} eqAB: equivmx A (row_full A) =1 equivmx B (row_full B).
+ by move=> C; rewrite /row_full /equivmx !eqAB.
+rewrite (eq_choose eqAB) (choose_id _ (genmx_witnessP B)) //.
+by rewrite -eqAB genmx_witnessP.
+Qed.
+
+Lemma genmxP m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (<<A>> = <<B>>)%MS (A == B)%MS.
+Proof.
+apply: (iffP idP) => eqAB; first exact: eq_genmx (eqmxP _).
+by rewrite -!(genmxE A) eqAB !genmxE andbb.
+Qed.
+Implicit Arguments genmxP [m1 m2 n A B].
+
+Lemma genmx0 m n : <<0 : 'M_(m, n)>>%MS = 0.
+Proof. by apply/eqP; rewrite -submx0 genmxE sub0mx. Qed.
+
+Lemma genmx1 n : <<1%:M : 'M_n>>%MS = 1%:M.
+Proof.
+rewrite unlock; case/andP: (chooseP (@genmx_witnessP n n 1%:M)) => _ /eqP.
+by rewrite qidmx_eq1 row_full_unit unitmx1 => /eqP.
+Qed.
+
+Lemma genmx_id m n (A : 'M_(m, n)) : (<<<<A>>>> = <<A>>)%MS.
+Proof. by apply: eq_genmx; exact: genmxE. Qed.
+
+Lemma row_base_free m n (A : 'M_(m, n)) : row_free (row_base A).
+Proof. by apply/eqnP; rewrite eq_row_base. Qed.
+
+Lemma mxrank_gen m n (A : 'M_(m, n)) : \rank <<A>> = \rank A.
+Proof. by rewrite genmxE. Qed.
+
+Lemma col_base_full m n (A : 'M_(m, n)) : row_full (col_base A).
+Proof.
+apply/row_fullP; exists (pid_mx (\rank A) *m invmx (col_ebase A)).
+by rewrite !mulmxA mulmxKV // pid_mx_id // pid_mx_1.
+Qed.
+Hint Resolve row_base_free col_base_full.
+
+Lemma mxrank_leqif_sup m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A <= B)%MS -> \rank A <= \rank B ?= iff (B <= A)%MS.
+Proof.
+move=> sAB; split; first by rewrite mxrankS.
+apply/idP/idP=> [| sBA]; last by rewrite eqn_leq !mxrankS.
+case/submxP: sAB => D ->; rewrite -{-2}(mulmx_base B) mulmxA.
+rewrite mxrankMfree // => /row_fullP[E kE].
+by rewrite -{1}[row_base B]mul1mx -kE -(mulmxA E) (mulmxA _ E) submxMl.
+Qed.
+
+Lemma mxrank_leqif_eq m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A <= B)%MS -> \rank A <= \rank B ?= iff (A == B)%MS.
+Proof. by move=> sAB; rewrite sAB; exact: mxrank_leqif_sup. Qed.
+
+Lemma ltmxErank m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A < B)%MS = (A <= B)%MS && (\rank A < \rank B).
+Proof.
+by apply: andb_id2l => sAB; rewrite (ltn_leqif (mxrank_leqif_sup sAB)).
+Qed.
+
+Lemma rank_ltmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A < B)%MS -> \rank A < \rank B.
+Proof. by rewrite ltmxErank => /andP[]. Qed.
+
+Lemma eqmx_cast m1 m2 n (A : 'M_(m1, n)) e :
+ ((castmx e A : 'M_(m2, n)) :=: A)%MS.
+Proof. by case: e A; case: m2 / => A e; rewrite castmx_id. Qed.
+
+Lemma eqmx_conform m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (conform_mx A B :=: A \/ conform_mx A B :=: B)%MS.
+Proof.
+case: (eqVneq m2 m1) => [-> | neqm12] in B *.
+ by right; rewrite conform_mx_id.
+by left; rewrite nonconform_mx ?neqm12.
+Qed.
+
+Let eqmx_sum_nop m n (A : 'M_(m, n)) : (addsmx_nop A :=: A)%MS.
+Proof.
+case: (eqmx_conform <<A>>%MS A) => // eq_id_gen.
+exact: eqmx_trans (genmxE A).
+Qed.
+
+Section AddsmxSub.
+
+Variable (m1 m2 n : nat) (A : 'M[F]_(m1, n)) (B : 'M[F]_(m2, n)).
+
+Lemma col_mx_sub m3 (C : 'M_(m3, n)) :
+ (col_mx A B <= C)%MS = (A <= C)%MS && (B <= C)%MS.
+Proof.
+rewrite !submxE mul_col_mx -col_mx0.
+by apply/eqP/andP; [case/eq_col_mx=> -> -> | case; do 2!move/eqP->].
+Qed.
+
+Lemma addsmxE : (A + B :=: col_mx A B)%MS.
+Proof.
+have:= submx_refl (col_mx A B); rewrite col_mx_sub; case/andP=> sAS sBS.
+rewrite unlock; do 2?case: eqP => [AB0 | _]; last exact: genmxE.
+ by apply/eqmxP; rewrite !eqmx_sum_nop sBS col_mx_sub AB0 sub0mx /=.
+by apply/eqmxP; rewrite !eqmx_sum_nop sAS col_mx_sub AB0 sub0mx andbT /=.
+Qed.
+
+Lemma addsmx_sub m3 (C : 'M_(m3, n)) :
+ (A + B <= C)%MS = (A <= C)%MS && (B <= C)%MS.
+Proof. by rewrite addsmxE col_mx_sub. Qed.
+
+Lemma addsmxSl : (A <= A + B)%MS.
+Proof. by have:= submx_refl (A + B)%MS; rewrite addsmx_sub; case/andP. Qed.
+
+Lemma addsmxSr : (B <= A + B)%MS.
+Proof. by have:= submx_refl (A + B)%MS; rewrite addsmx_sub; case/andP. Qed.
+
+Lemma addsmx_idPr : reflect (A + B :=: B)%MS (A <= B)%MS.
+Proof.
+have:= @eqmxP _ _ _ (A + B)%MS B.
+by rewrite addsmxSr addsmx_sub submx_refl !andbT.
+Qed.
+
+Lemma addsmx_idPl : reflect (A + B :=: A)%MS (B <= A)%MS.
+Proof.
+have:= @eqmxP _ _ _ (A + B)%MS A.
+by rewrite addsmxSl addsmx_sub submx_refl !andbT.
+Qed.
+
+End AddsmxSub.
+
+Lemma adds0mx m1 m2 n (B : 'M_(m2, n)) : ((0 : 'M_(m1, n)) + B :=: B)%MS.
+Proof. by apply/eqmxP; rewrite addsmx_sub sub0mx addsmxSr /= andbT. Qed.
+
+Lemma addsmx0 m1 m2 n (A : 'M_(m1, n)) : (A + (0 : 'M_(m2, n)) :=: A)%MS.
+Proof. by apply/eqmxP; rewrite addsmx_sub sub0mx addsmxSl /= !andbT. Qed.
+
+Let addsmx_nop_eq0 m n (A : 'M_(m, n)) : (addsmx_nop A == 0) = (A == 0).
+Proof. by rewrite -!submx0 eqmx_sum_nop. Qed.
+
+Let addsmx_nop0 m n : addsmx_nop (0 : 'M_(m, n)) = 0.
+Proof. by apply/eqP; rewrite addsmx_nop_eq0. Qed.
+
+Let addsmx_nop_id n (A : 'M_n) : addsmx_nop A = A.
+Proof. exact: conform_mx_id. Qed.
+
+Lemma addsmxC m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : (A + B = B + A)%MS.
+Proof.
+have: (A + B == B + A)%MS.
+ by apply/andP; rewrite !addsmx_sub andbC -addsmx_sub andbC -addsmx_sub.
+move/genmxP; rewrite [@addsmx]unlock -!submx0 !submx0.
+by do 2!case: eqP => [// -> | _]; rewrite ?genmx_id ?addsmx_nop0.
+Qed.
+
+Lemma adds0mx_id m1 n (B : 'M_n) : ((0 : 'M_(m1, n)) + B)%MS = B.
+Proof. by rewrite unlock eqxx addsmx_nop_id. Qed.
+
+Lemma addsmx0_id m2 n (A : 'M_n) : (A + (0 : 'M_(m2, n)))%MS = A.
+Proof. by rewrite addsmxC adds0mx_id. Qed.
+
+Lemma addsmxA m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ (A + (B + C) = A + B + C)%MS.
+Proof.
+have: (A + (B + C) :=: A + B + C)%MS.
+ by apply/eqmxP/andP; rewrite !addsmx_sub -andbA andbA -!addsmx_sub.
+rewrite {1 3}[in @addsmx m1]unlock [in @addsmx n]unlock !addsmx_nop_id -!submx0.
+rewrite !addsmx_sub ![@addsmx]unlock -!submx0; move/eq_genmx.
+by do 3!case: (_ <= 0)%MS; rewrite //= !genmx_id.
+Qed.
+
+Canonical addsmx_monoid n :=
+ Monoid.Law (@addsmxA n n n n) (@adds0mx_id n n) (@addsmx0_id n n).
+Canonical addsmx_comoid n := Monoid.ComLaw (@addsmxC n n n).
+
+Lemma addsmxMr m1 m2 n p (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(n, p)) :
+ ((A + B)%MS *m C :=: A *m C + B *m C)%MS.
+Proof. by apply/eqmxP; rewrite !addsmxE -!mul_col_mx !submxMr ?addsmxE. Qed.
+
+Lemma addsmxS m1 m2 m3 m4 n (A : 'M_(m1, n)) (B : 'M_(m2, n))
+ (C : 'M_(m3, n)) (D : 'M_(m4, n)) :
+ (A <= C -> B <= D -> A + B <= C + D)%MS.
+Proof.
+move=> sAC sBD.
+by rewrite addsmx_sub {1}addsmxC !(submx_trans _ (addsmxSr _ _)).
+Qed.
+
+Lemma addmx_sub_adds m m1 m2 n (A : 'M_(m, n)) (B : 'M_(m, n))
+ (C : 'M_(m1, n)) (D : 'M_(m2, n)) :
+ (A <= C -> B <= D -> (A + B)%R <= C + D)%MS.
+Proof.
+move=> sAC; move/(addsmxS sAC); apply: submx_trans.
+by rewrite addmx_sub ?addsmxSl ?addsmxSr.
+Qed.
+
+Lemma addsmx_addKl n m1 m2 (A : 'M_(m1, n)) (B C : 'M_(m2, n)) :
+ (B <= A)%MS -> (A + (B + C)%R :=: A + C)%MS.
+Proof.
+move=> sBA; apply/eqmxP; rewrite !addsmx_sub !addsmxSl.
+by rewrite -{3}[C](addKr B) !addmx_sub_adds ?eqmx_opp.
+Qed.
+
+Lemma addsmx_addKr n m1 m2 (A B : 'M_(m1, n)) (C : 'M_(m2, n)) :
+ (B <= C)%MS -> ((A + B)%R + C :=: A + C)%MS.
+Proof. by rewrite -!(addsmxC C) addrC; exact: addsmx_addKl. Qed.
+
+Lemma adds_eqmx m1 m2 m3 m4 n (A : 'M_(m1, n)) (B : 'M_(m2, n))
+ (C : 'M_(m3, n)) (D : 'M_(m4, n)) :
+ (A :=: C -> B :=: D -> A + B :=: C + D)%MS.
+Proof. by move=> eqAC eqBD; apply/eqmxP; rewrite !addsmxS ?eqAC ?eqBD. Qed.
+
+Lemma genmx_adds m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (<<(A + B)%MS>> = <<A>> + <<B>>)%MS.
+Proof.
+rewrite -(eq_genmx (adds_eqmx (genmxE A) (genmxE B))).
+by rewrite [@addsmx]unlock !addsmx_nop_id !(fun_if (@genmx _ _)) !genmx_id.
+Qed.
+
+Lemma sub_addsmxP m1 m2 m3 n
+ (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ reflect (exists u, A = u.1 *m B + u.2 *m C) (A <= B + C)%MS.
+Proof.
+apply: (iffP idP) => [|[u ->]]; last by rewrite addmx_sub_adds ?submxMl.
+rewrite addsmxE; case/submxP=> u ->; exists (lsubmx u, rsubmx u).
+by rewrite -mul_row_col hsubmxK.
+Qed.
+Implicit Arguments sub_addsmxP [m1 m2 m3 n A B C].
+
+Variable I : finType.
+Implicit Type P : pred I.
+
+Lemma genmx_sums P n (B_ : I -> 'M_n) :
+ <<(\sum_(i | P i) B_ i)%MS>>%MS = (\sum_(i | P i) <<B_ i>>)%MS.
+Proof. exact: (big_morph _ (@genmx_adds n n n) (@genmx0 n n)). Qed.
+
+Lemma sumsmx_sup i0 P m n (A : 'M_(m, n)) (B_ : I -> 'M_n) :
+ P i0 -> (A <= B_ i0)%MS -> (A <= \sum_(i | P i) B_ i)%MS.
+Proof.
+by move=> Pi0 sAB; apply: submx_trans sAB _; rewrite (bigD1 i0) // addsmxSl.
+Qed.
+Implicit Arguments sumsmx_sup [P m n A B_].
+
+Lemma sumsmx_subP P m n (A_ : I -> 'M_n) (B : 'M_(m, n)) :
+ reflect (forall i, P i -> A_ i <= B)%MS (\sum_(i | P i) A_ i <= B)%MS.
+Proof.
+apply: (iffP idP) => [sAB i Pi | sAB].
+ by apply: submx_trans sAB; apply: sumsmx_sup Pi _.
+by elim/big_rec: _ => [|i Ai Pi sAiB]; rewrite ?sub0mx // addsmx_sub sAB.
+Qed.
+
+Lemma summx_sub_sums P m n (A : I -> 'M[F]_(m, n)) B :
+ (forall i, P i -> A i <= B i)%MS ->
+ ((\sum_(i | P i) A i)%R <= \sum_(i | P i) B i)%MS.
+Proof.
+by move=> sAB; apply: summx_sub => i Pi; rewrite (sumsmx_sup i) ?sAB.
+Qed.
+
+Lemma sumsmxS P n (A B : I -> 'M[F]_n) :
+ (forall i, P i -> A i <= B i)%MS ->
+ (\sum_(i | P i) A i <= \sum_(i | P i) B i)%MS.
+Proof.
+by move=> sAB; apply/sumsmx_subP=> i Pi; rewrite (sumsmx_sup i) ?sAB.
+Qed.
+
+Lemma eqmx_sums P n (A B : I -> 'M[F]_n) :
+ (forall i, P i -> A i :=: B i)%MS ->
+ (\sum_(i | P i) A i :=: \sum_(i | P i) B i)%MS.
+Proof. by move=> eqAB; apply/eqmxP; rewrite !sumsmxS // => i; move/eqAB->. Qed.
+
+Lemma sub_sumsmxP P m n (A : 'M_(m, n)) (B_ : I -> 'M_n) :
+ reflect (exists u_, A = \sum_(i | P i) u_ i *m B_ i)
+ (A <= \sum_(i | P i) B_ i)%MS.
+Proof.
+apply: (iffP idP) => [| [u_ ->]]; last first.
+ by apply: summx_sub_sums => i _; exact: submxMl.
+elim: {P}_.+1 {-2}P A (ltnSn #|P|) => // b IHb P A.
+case: (pickP P) => [i Pi | P0 _]; last first.
+ rewrite big_pred0 //; move/submx0null->.
+ by exists (fun _ => 0); rewrite big_pred0.
+rewrite (cardD1x Pi) (bigD1 i) //= => /IHb{b IHb} /= IHi /sub_addsmxP[u ->].
+have [u_ ->] := IHi _ (submxMl u.2 _).
+exists [eta u_ with i |-> u.1]; rewrite (bigD1 i Pi) /= eqxx; congr (_ + _).
+by apply: eq_bigr => j /andP[_ /negPf->].
+Qed.
+
+Lemma sumsmxMr_gen P m n A (B : 'M[F]_(m, n)) :
+ ((\sum_(i | P i) A i)%MS *m B :=: \sum_(i | P i) <<A i *m B>>)%MS.
+Proof.
+apply/eqmxP/andP; split; last first.
+ by apply/sumsmx_subP=> i Pi; rewrite genmxE submxMr ?(sumsmx_sup i).
+have [u ->] := sub_sumsmxP _ _ _ (submx_refl (\sum_(i | P i) A i)%MS).
+by rewrite mulmx_suml summx_sub_sums // => i _; rewrite genmxE -mulmxA submxMl.
+Qed.
+
+Lemma sumsmxMr P n (A_ : I -> 'M[F]_n) (B : 'M_n) :
+ ((\sum_(i | P i) A_ i)%MS *m B :=: \sum_(i | P i) (A_ i *m B))%MS.
+Proof.
+by apply: eqmx_trans (sumsmxMr_gen _ _ _) (eqmx_sums _) => i _; exact: genmxE.
+Qed.
+
+Lemma rank_pid_mx m n r : r <= m -> r <= n -> \rank (pid_mx r : 'M_(m, n)) = r.
+Proof.
+do 2!move/subnKC <-; rewrite pid_mx_block block_mxEv row_mx0 -addsmxE addsmx0.
+by rewrite -mxrank_tr tr_row_mx trmx0 trmx1 -addsmxE addsmx0 mxrank1.
+Qed.
+
+Lemma rank_copid_mx n r : r <= n -> \rank (copid_mx r : 'M_n) = (n - r)%N.
+Proof.
+move/subnKC <-; rewrite /copid_mx pid_mx_block scalar_mx_block.
+rewrite opp_block_mx !oppr0 add_block_mx !addr0 subrr block_mxEv row_mx0.
+rewrite -addsmxE adds0mx -mxrank_tr tr_row_mx trmx0 trmx1.
+by rewrite -addsmxE adds0mx mxrank1 addKn.
+Qed.
+
+Lemma mxrank_compl m n (A : 'M_(m, n)) : \rank A^C = (n - \rank A)%N.
+Proof. by rewrite mxrankMfree ?row_free_unit ?rank_copid_mx. Qed.
+
+Lemma mxrank_ker m n (A : 'M_(m, n)) : \rank (kermx A) = (m - \rank A)%N.
+Proof. by rewrite mxrankMfree ?row_free_unit ?unitmx_inv ?rank_copid_mx. Qed.
+
+Lemma kermx_eq0 n m (A : 'M_(m, n)) : (kermx A == 0) = row_free A.
+Proof. by rewrite -mxrank_eq0 mxrank_ker subn_eq0 row_leq_rank. Qed.
+
+Lemma mxrank_coker m n (A : 'M_(m, n)) : \rank (cokermx A) = (n - \rank A)%N.
+Proof. by rewrite eqmxMfull ?row_full_unit ?unitmx_inv ?rank_copid_mx. Qed.
+
+Lemma cokermx_eq0 n m (A : 'M_(m, n)) : (cokermx A == 0) = row_full A.
+Proof. by rewrite -mxrank_eq0 mxrank_coker subn_eq0 col_leq_rank. Qed.
+
+Lemma mulmx_ker m n (A : 'M_(m, n)) : kermx A *m A = 0.
+Proof.
+by rewrite -{2}[A]mulmx_ebase !mulmxA mulmxKV // mul_copid_mx_pid ?mul0mx.
+Qed.
+
+Lemma mulmxKV_ker m n p (A : 'M_(n, p)) (B : 'M_(m, n)) :
+ B *m A = 0 -> B *m col_ebase A *m kermx A = B.
+Proof.
+rewrite mulmxA mulmxBr mulmx1 mulmxBl mulmxK //.
+rewrite -{1}[A]mulmx_ebase !mulmxA => /(canRL (mulmxK (row_ebase_unit A))).
+rewrite mul0mx // => BA0; apply: (canLR (addrK _)).
+by rewrite -(pid_mx_id _ _ n (rank_leq_col A)) mulmxA BA0 !mul0mx addr0.
+Qed.
+
+Lemma sub_kermxP p m n (A : 'M_(m, n)) (B : 'M_(p, m)) :
+ reflect (B *m A = 0) (B <= kermx A)%MS.
+Proof.
+apply: (iffP submxP) => [[D ->]|]; first by rewrite -mulmxA mulmx_ker mulmx0.
+by move/mulmxKV_ker; exists (B *m col_ebase A).
+Qed.
+
+Lemma mulmx0_rank_max m n p (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ A *m B = 0 -> \rank A + \rank B <= n.
+Proof.
+move=> AB0; rewrite -{3}(subnK (rank_leq_row B)) leq_add2r.
+rewrite -mxrank_ker mxrankS //; exact/sub_kermxP.
+Qed.
+
+Lemma mxrank_Frobenius m n p q (A : 'M_(m, n)) B (C : 'M_(p, q)) :
+ \rank (A *m B) + \rank (B *m C) <= \rank B + \rank (A *m B *m C).
+Proof.
+rewrite -{2}(mulmx_base (A *m B)) -mulmxA (eqmxMfull _ (col_base_full _)).
+set C2 := row_base _ *m C.
+rewrite -{1}(subnK (rank_leq_row C2)) -(mxrank_ker C2) addnAC leq_add2r.
+rewrite addnC -{1}(mulmx_base B) -mulmxA eqmxMfull //.
+set C1 := _ *m C; rewrite -{2}(subnKC (rank_leq_row C1)) leq_add2l -mxrank_ker.
+rewrite -(mxrankMfree _ (row_base_free (A *m B))).
+have: (row_base (A *m B) <= row_base B)%MS by rewrite !eq_row_base submxMl.
+case/submxP=> D defD; rewrite defD mulmxA mxrankMfree ?mxrankS //.
+by apply/sub_kermxP; rewrite -mulmxA (mulmxA D) -defD -/C2 mulmx_ker.
+Qed.
+
+Lemma mxrank_mul_min m n p (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ \rank A + \rank B - n <= \rank (A *m B).
+Proof.
+by have:= mxrank_Frobenius A 1%:M B; rewrite mulmx1 mul1mx mxrank1 leq_subLR.
+Qed.
+
+Lemma addsmx_compl_full m n (A : 'M_(m, n)) : row_full (A + A^C)%MS.
+Proof.
+rewrite /row_full addsmxE; apply/row_fullP.
+exists (row_mx (pinvmx A) (cokermx A)); rewrite mul_row_col.
+rewrite -{2}[A]mulmx_ebase -!mulmxA mulKmx // -mulmxDr !mulmxA.
+by rewrite pid_mx_id ?copid_mx_id // -mulmxDl addrC subrK mul1mx mulVmx.
+Qed.
+
+Lemma sub_capmx_gen m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ (A <= capmx_gen B C)%MS = (A <= B)%MS && (A <= C)%MS.
+Proof.
+apply/idP/andP=> [sAI | [/submxP[B' ->{A}] /submxP[C' eqBC']]].
+ rewrite !(submx_trans sAI) ?submxMl // /capmx_gen.
+ have:= mulmx_ker (col_mx B C); set K := kermx _.
+ rewrite -{1}[K]hsubmxK mul_row_col; move/(canRL (addrK _))->.
+ by rewrite add0r -mulNmx submxMl.
+have: (row_mx B' (- C') <= kermx (col_mx B C))%MS.
+ by apply/sub_kermxP; rewrite mul_row_col eqBC' mulNmx subrr.
+case/submxP=> D; rewrite -[kermx _]hsubmxK mul_mx_row.
+by case/eq_row_mx=> -> _; rewrite -mulmxA submxMl.
+Qed.
+
+Let capmx_witnessP m n (A : 'M_(m, n)) : equivmx A (qidmx A) (capmx_witness A).
+Proof.
+rewrite /equivmx qidmx_eq1 /qidmx /capmx_witness.
+rewrite -sub1mx; case s1A: (1%:M <= A)%MS => /=; last first.
+ rewrite !genmxE submx_refl /= -negb_add; apply: contra {s1A}(negbT s1A).
+ case: eqP => [<- _| _]; first by rewrite genmxE.
+ by case: eqP A => //= -> A; move/eqP->; rewrite pid_mx_1.
+case: (m =P n) => [-> | ne_mn] in A s1A *.
+ by rewrite conform_mx_id submx_refl pid_mx_1 eqxx.
+by rewrite nonconform_mx ?submx1 ?s1A ?eqxx //; case: eqP.
+Qed.
+
+Let capmx_normP m n (A : 'M_(m, n)) : equivmx_spec A (qidmx A) (capmx_norm A).
+Proof. by case/andP: (chooseP (capmx_witnessP A)) => /eqmxP defN /eqP. Qed.
+
+Let capmx_norm_eq m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ qidmx A = qidmx B -> (A == B)%MS -> capmx_norm A = capmx_norm B.
+Proof.
+move=> eqABid /eqmxP eqAB.
+have{eqABid eqAB} eqAB: equivmx A (qidmx A) =1 equivmx B (qidmx B).
+ by move=> C; rewrite /equivmx eqABid !eqAB.
+rewrite {1}/capmx_norm (eq_choose eqAB).
+by apply: choose_id; first rewrite -eqAB; exact: capmx_witnessP.
+Qed.
+
+Let capmx_nopP m n (A : 'M_(m, n)) : equivmx_spec A (qidmx A) (capmx_nop A).
+Proof.
+rewrite /capmx_nop; case: (eqVneq m n) => [-> | ne_mn] in A *.
+ by rewrite conform_mx_id.
+rewrite nonconform_mx ?ne_mn //; exact: capmx_normP.
+Qed.
+
+Let sub_qidmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ qidmx B -> (A <= B)%MS.
+Proof.
+rewrite /qidmx => idB; apply: {A}submx_trans (submx1 A) _.
+by case: eqP B idB => [-> _ /eqP-> | _ B]; rewrite (=^~ sub1mx, pid_mx_1).
+Qed.
+
+Let qidmx_cap m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ qidmx (A :&: B)%MS = qidmx A && qidmx B.
+Proof.
+rewrite unlock -sub1mx.
+case idA: (qidmx A); case idB: (qidmx B); try by rewrite capmx_nopP.
+case s1B: (_ <= B)%MS; first by rewrite capmx_normP.
+apply/idP=> /(sub_qidmx 1%:M).
+by rewrite capmx_normP sub_capmx_gen s1B andbF.
+Qed.
+
+Let capmx_eq_norm m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ qidmx A = qidmx B -> (A :&: B)%MS = capmx_norm (A :&: B)%MS.
+Proof.
+move=> eqABid; rewrite unlock -sub1mx {}eqABid.
+have norm_id m (C : 'M_(m, n)) (N := capmx_norm C) : capmx_norm N = N.
+ by apply: capmx_norm_eq; rewrite ?capmx_normP ?andbb.
+case idB: (qidmx B); last by case: ifP; rewrite norm_id.
+rewrite /capmx_nop; case: (eqVneq m2 n) => [-> | neqm2n] in B idB *.
+ have idN := idB; rewrite -{1}capmx_normP !qidmx_eq1 in idN idB.
+ by rewrite conform_mx_id (eqP idN) (eqP idB).
+by rewrite nonconform_mx ?neqm2n ?norm_id.
+Qed.
+
+Lemma capmxE m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A :&: B :=: capmx_gen A B)%MS.
+Proof.
+rewrite unlock -sub1mx; apply/eqmxP.
+have:= submx_refl (capmx_gen A B); rewrite !sub_capmx_gen => /andP[sIA sIB].
+case idA: (qidmx A); first by rewrite !capmx_nopP submx_refl sub_qidmx.
+case idB: (qidmx B); first by rewrite !capmx_nopP submx_refl sub_qidmx.
+case s1B: (1%:M <= B)%MS; rewrite !capmx_normP ?sub_capmx_gen sIA ?sIB //=.
+by rewrite submx_refl (submx_trans (submx1 _)).
+Qed.
+
+Lemma capmxSl m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : (A :&: B <= A)%MS.
+Proof. by rewrite capmxE submxMl. Qed.
+
+Lemma sub_capmx m m1 m2 n (A : 'M_(m, n)) (B : 'M_(m1, n)) (C : 'M_(m2, n)) :
+ (A <= B :&: C)%MS = (A <= B)%MS && (A <= C)%MS.
+Proof. by rewrite capmxE sub_capmx_gen. Qed.
+
+Lemma capmxC m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : (A :&: B = B :&: A)%MS.
+Proof.
+have [eqAB|] := eqVneq (qidmx A) (qidmx B).
+ rewrite (capmx_eq_norm eqAB) (capmx_eq_norm (esym eqAB)).
+ apply: capmx_norm_eq; first by rewrite !qidmx_cap andbC.
+ by apply/andP; split; rewrite !sub_capmx andbC -sub_capmx.
+by rewrite negb_eqb !unlock => /addbP <-; case: (qidmx A).
+Qed.
+
+Lemma capmxSr m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : (A :&: B <= B)%MS.
+Proof. by rewrite capmxC capmxSl. Qed.
+
+Lemma capmx_idPr n m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (A :&: B :=: B)%MS (B <= A)%MS.
+Proof.
+have:= @eqmxP _ _ _ (A :&: B)%MS B.
+by rewrite capmxSr sub_capmx submx_refl !andbT.
+Qed.
+
+Lemma capmx_idPl n m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (A :&: B :=: A)%MS (A <= B)%MS.
+Proof. by rewrite capmxC; exact: capmx_idPr. Qed.
+
+Lemma capmxS m1 m2 m3 m4 n (A : 'M_(m1, n)) (B : 'M_(m2, n))
+ (C : 'M_(m3, n)) (D : 'M_(m4, n)) :
+ (A <= C -> B <= D -> A :&: B <= C :&: D)%MS.
+Proof.
+by move=> sAC sBD; rewrite sub_capmx {1}capmxC !(submx_trans (capmxSr _ _)).
+Qed.
+
+Lemma cap_eqmx m1 m2 m3 m4 n (A : 'M_(m1, n)) (B : 'M_(m2, n))
+ (C : 'M_(m3, n)) (D : 'M_(m4, n)) :
+ (A :=: C -> B :=: D -> A :&: B :=: C :&: D)%MS.
+Proof. by move=> eqAC eqBD; apply/eqmxP; rewrite !capmxS ?eqAC ?eqBD. Qed.
+
+Lemma capmxMr m1 m2 n p (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(n, p)) :
+ ((A :&: B) *m C <= A *m C :&: B *m C)%MS.
+Proof. by rewrite sub_capmx !submxMr ?capmxSl ?capmxSr. Qed.
+
+Lemma cap0mx m1 m2 n (A : 'M_(m2, n)) : ((0 : 'M_(m1, n)) :&: A)%MS = 0.
+Proof. exact: submx0null (capmxSl _ _). Qed.
+
+Lemma capmx0 m1 m2 n (A : 'M_(m1, n)) : (A :&: (0 : 'M_(m2, n)))%MS = 0.
+Proof. exact: submx0null (capmxSr _ _). Qed.
+
+Lemma capmxT m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ row_full B -> (A :&: B :=: A)%MS.
+Proof.
+rewrite -sub1mx => s1B; apply/eqmxP.
+by rewrite capmxSl sub_capmx submx_refl (submx_trans (submx1 A)).
+Qed.
+
+Lemma capTmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ row_full A -> (A :&: B :=: B)%MS.
+Proof. by move=> Afull; apply/eqmxP; rewrite capmxC !capmxT ?andbb. Qed.
+
+Let capmx_nop_id n (A : 'M_n) : capmx_nop A = A.
+Proof. by rewrite /capmx_nop conform_mx_id. Qed.
+
+Lemma cap1mx n (A : 'M_n) : (1%:M :&: A = A)%MS.
+Proof. by rewrite unlock qidmx_eq1 eqxx capmx_nop_id. Qed.
+
+Lemma capmx1 n (A : 'M_n) : (A :&: 1%:M = A)%MS.
+Proof. by rewrite capmxC cap1mx. Qed.
+
+Lemma genmx_cap m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ <<A :&: B>>%MS = (<<A>> :&: <<B>>)%MS.
+Proof.
+rewrite -(eq_genmx (cap_eqmx (genmxE A) (genmxE B))).
+case idAB: (qidmx <<A>> || qidmx <<B>>)%MS.
+ rewrite [@capmx]unlock !capmx_nop_id !(fun_if (@genmx _ _)) !genmx_id.
+ by case: (qidmx _) idAB => //= ->.
+case idA: (qidmx _) idAB => //= idB; rewrite {2}capmx_eq_norm ?idA //.
+set C := (_ :&: _)%MS; have eq_idC: row_full C = qidmx C.
+ rewrite qidmx_cap idA -sub1mx sub_capmx genmxE; apply/andP=> [[s1A]].
+ by case/idP: idA; rewrite qidmx_eq1 -genmx1 (sameP eqP genmxP) submx1.
+rewrite unlock /capmx_norm eq_idC.
+by apply: choose_id (capmx_witnessP _); rewrite -eq_idC genmx_witnessP.
+Qed.
+
+Lemma capmxA m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ (A :&: (B :&: C) = A :&: B :&: C)%MS.
+Proof.
+rewrite (capmxC A B) capmxC; wlog idA: m1 m3 A C / qidmx A.
+ move=> IH; case idA: (qidmx A); first exact: IH.
+ case idC: (qidmx C); first by rewrite -IH.
+ rewrite (@capmx_eq_norm n m3) ?qidmx_cap ?idA ?idC ?andbF //.
+ rewrite capmx_eq_norm ?qidmx_cap ?idA ?idC ?andbF //.
+ apply: capmx_norm_eq; first by rewrite !qidmx_cap andbAC.
+ by apply/andP; split; rewrite !sub_capmx andbAC -!sub_capmx.
+rewrite -!(capmxC A) [in @capmx m1]unlock idA capmx_nop_id.
+have [eqBC |] :=eqVneq (qidmx B) (qidmx C).
+ rewrite (@capmx_eq_norm n) ?capmx_nopP // capmx_eq_norm //.
+ by apply: capmx_norm_eq; rewrite ?qidmx_cap ?capmxS ?capmx_nopP.
+by rewrite !unlock capmx_nopP capmx_nop_id; do 2?case: (qidmx _) => //.
+Qed.
+
+Canonical capmx_monoid n :=
+ Monoid.Law (@capmxA n n n n) (@cap1mx n) (@capmx1 n).
+Canonical capmx_comoid n := Monoid.ComLaw (@capmxC n n n).
+
+Lemma bigcapmx_inf i0 P m n (A_ : I -> 'M_n) (B : 'M_(m, n)) :
+ P i0 -> (A_ i0 <= B -> \bigcap_(i | P i) A_ i <= B)%MS.
+Proof. by move=> Pi0; apply: submx_trans; rewrite (bigD1 i0) // capmxSl. Qed.
+
+Lemma sub_bigcapmxP P m n (A : 'M_(m, n)) (B_ : I -> 'M_n) :
+ reflect (forall i, P i -> A <= B_ i)%MS (A <= \bigcap_(i | P i) B_ i)%MS.
+Proof.
+apply: (iffP idP) => [sAB i Pi | sAB].
+ by apply: (submx_trans sAB); rewrite (bigcapmx_inf Pi).
+by elim/big_rec: _ => [|i Pi C sAC]; rewrite ?submx1 // sub_capmx sAB.
+Qed.
+
+Lemma genmx_bigcap P n (A_ : I -> 'M_n) :
+ (<<\bigcap_(i | P i) A_ i>> = \bigcap_(i | P i) <<A_ i>>)%MS.
+Proof. exact: (big_morph _ (@genmx_cap n n n) (@genmx1 n)). Qed.
+
+Lemma matrix_modl m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ (A <= C -> A + (B :&: C) :=: (A + B) :&: C)%MS.
+Proof.
+move=> sAC; set D := ((A + B) :&: C)%MS; apply/eqmxP.
+rewrite sub_capmx addsmxS ?capmxSl // addsmx_sub sAC capmxSr /=.
+have: (D <= B + A)%MS by rewrite addsmxC capmxSl.
+case/sub_addsmxP=> u defD; rewrite defD addrC addmx_sub_adds ?submxMl //.
+rewrite sub_capmx submxMl -[_ *m B](addrK (u.2 *m A)) -defD.
+by rewrite addmx_sub ?capmxSr // eqmx_opp mulmx_sub.
+Qed.
+
+Lemma matrix_modr m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) :
+ (C <= A -> (A :&: B) + C :=: A :&: (B + C))%MS.
+Proof. by rewrite !(capmxC A) -!(addsmxC C); exact: matrix_modl. Qed.
+
+Lemma capmx_compl m n (A : 'M_(m, n)) : (A :&: A^C)%MS = 0.
+Proof.
+set D := (A :&: A^C)%MS; have: (D <= D)%MS by [].
+rewrite sub_capmx andbC => /andP[/submxP[B defB]].
+rewrite submxE => /eqP; rewrite defB -!mulmxA mulKVmx ?copid_mx_id //.
+by rewrite mulmxA => ->; rewrite mul0mx.
+Qed.
+
+Lemma mxrank_mul_ker m n p (A : 'M_(m, n)) (B : 'M_(n, p)) :
+ (\rank (A *m B) + \rank (A :&: kermx B))%N = \rank A.
+Proof.
+apply/eqP; set K := kermx B; set C := (A :&: K)%MS.
+rewrite -(eqmxMr B (eq_row_base A)); set K' := _ *m B.
+rewrite -{2}(subnKC (rank_leq_row K')) -mxrank_ker eqn_add2l.
+rewrite -(mxrankMfree _ (row_base_free A)) mxrank_leqif_sup.
+ rewrite sub_capmx -(eq_row_base A) submxMl.
+ by apply/sub_kermxP; rewrite -mulmxA mulmx_ker.
+have /submxP[C' defC]: (C <= row_base A)%MS by rewrite eq_row_base capmxSl.
+rewrite defC submxMr //; apply/sub_kermxP.
+by rewrite mulmxA -defC; apply/sub_kermxP; rewrite capmxSr.
+Qed.
+
+Lemma mxrank_injP m n p (A : 'M_(m, n)) (f : 'M_(n, p)) :
+ reflect (\rank (A *m f) = \rank A) ((A :&: kermx f)%MS == 0).
+Proof.
+rewrite -mxrank_eq0 -(eqn_add2l (\rank (A *m f))).
+by rewrite mxrank_mul_ker addn0 eq_sym; exact: eqP.
+Qed.
+
+Lemma mxrank_disjoint_sum m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A :&: B)%MS = 0 -> \rank (A + B)%MS = (\rank A + \rank B)%N.
+Proof.
+move=> AB0; pose Ar := row_base A; pose Br := row_base B.
+have [Afree Bfree]: row_free Ar /\ row_free Br by rewrite !row_base_free.
+have: (Ar :&: Br <= A :&: B)%MS by rewrite capmxS ?eq_row_base.
+rewrite {}AB0 submx0 -mxrank_eq0 capmxE mxrankMfree //.
+set Cr := col_mx Ar Br; set Crl := lsubmx _; rewrite mxrank_eq0 => /eqP Crl0.
+rewrite -(adds_eqmx (eq_row_base _) (eq_row_base _)) addsmxE -/Cr.
+suffices K0: kermx Cr = 0.
+ by apply/eqP; rewrite eqn_leq rank_leq_row -subn_eq0 -mxrank_ker K0 mxrank0.
+move/eqP: (mulmx_ker Cr); rewrite -[kermx Cr]hsubmxK mul_row_col -/Crl Crl0.
+rewrite mul0mx add0r -mxrank_eq0 mxrankMfree // mxrank_eq0 => /eqP->.
+exact: row_mx0.
+Qed.
+
+Lemma diffmxE m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A :\: B :=: A :&: (capmx_gen A B)^C)%MS.
+Proof. by rewrite unlock; apply/eqmxP; rewrite !genmxE !capmxE andbb. Qed.
+
+Lemma genmx_diff m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (<<A :\: B>> = A :\: B)%MS.
+Proof. by rewrite [@diffmx]unlock genmx_id. Qed.
+
+Lemma diffmxSl m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : (A :\: B <= A)%MS.
+Proof. by rewrite diffmxE capmxSl. Qed.
+
+Lemma capmx_diff m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ ((A :\: B) :&: B)%MS = 0.
+Proof.
+apply/eqP; pose C := capmx_gen A B; rewrite -submx0 -(capmx_compl C).
+by rewrite sub_capmx -capmxE sub_capmx andbAC -sub_capmx -diffmxE -sub_capmx.
+Qed.
+
+Lemma addsmx_diff_cap_eq m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A :\: B + A :&: B :=: A)%MS.
+Proof.
+apply/eqmxP; rewrite addsmx_sub capmxSl diffmxSl /=.
+set C := (A :\: B)%MS; set D := capmx_gen A B.
+suffices sACD: (A <= C + D)%MS.
+ by rewrite (submx_trans sACD) ?addsmxS ?capmxE.
+have:= addsmx_compl_full D; rewrite /row_full addsmxE.
+case/row_fullP=> U /(congr1 (mulmx A)); rewrite mulmx1.
+rewrite -[U]hsubmxK mul_row_col mulmxDr addrC 2!mulmxA.
+set V := _ *m _ => defA; rewrite -defA; move/(canRL (addrK _)): defA => defV.
+suffices /submxP[W ->]: (V <= C)%MS by rewrite -mul_row_col addsmxE submxMl.
+rewrite diffmxE sub_capmx {1}defV -mulNmx addmx_sub 1?mulmx_sub //.
+by rewrite -capmxE capmxSl.
+Qed.
+
+Lemma mxrank_cap_compl m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (\rank (A :&: B) + \rank (A :\: B))%N = \rank A.
+Proof.
+rewrite addnC -mxrank_disjoint_sum ?addsmx_diff_cap_eq //.
+by rewrite (capmxC A) capmxA capmx_diff cap0mx.
+Qed.
+
+Lemma mxrank_sum_cap m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (\rank (A + B) + \rank (A :&: B) = \rank A + \rank B)%N.
+Proof.
+set C := (A :&: B)%MS; set D := (A :\: B)%MS.
+have rDB: \rank (A + B)%MS = \rank (D + B)%MS.
+ apply/eqP; rewrite mxrank_leqif_sup; first by rewrite addsmxS ?diffmxSl.
+ by rewrite addsmx_sub addsmxSr -(addsmx_diff_cap_eq A B) addsmxS ?capmxSr.
+rewrite {1}rDB mxrank_disjoint_sum ?capmx_diff //.
+by rewrite addnC addnA mxrank_cap_compl.
+Qed.
+
+Lemma mxrank_adds_leqif m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ \rank (A + B) <= \rank A + \rank B ?= iff (A :&: B <= (0 : 'M_n))%MS.
+Proof.
+rewrite -mxrank_sum_cap; split; first exact: leq_addr.
+by rewrite addnC (@eqn_add2r _ 0) eq_sym mxrank_eq0 -submx0.
+Qed.
+
+(* Subspace projection matrix *)
+
+Lemma proj_mx_sub m n U V (W : 'M_(m, n)) : (W *m proj_mx U V <= U)%MS.
+Proof. by rewrite !mulmx_sub // -addsmxE addsmx0. Qed.
+
+Lemma proj_mx_compl_sub m n U V (W : 'M_(m, n)) :
+ (W <= U + V -> W - W *m proj_mx U V <= V)%MS.
+Proof.
+rewrite addsmxE => sWUV; rewrite mulmxA -{1}(mulmxKpV sWUV) -mulmxBr.
+by rewrite mulmx_sub // opp_col_mx add_col_mx subrr subr0 -addsmxE adds0mx.
+Qed.
+
+Lemma proj_mx_id m n U V (W : 'M_(m, n)) :
+ (U :&: V = 0)%MS -> (W <= U)%MS -> W *m proj_mx U V = W.
+Proof.
+move=> dxUV sWU; apply/eqP; rewrite -subr_eq0 -submx0 -dxUV.
+rewrite sub_capmx addmx_sub ?eqmx_opp ?proj_mx_sub //= -eqmx_opp opprB.
+by rewrite proj_mx_compl_sub // (submx_trans sWU) ?addsmxSl.
+Qed.
+
+Lemma proj_mx_0 m n U V (W : 'M_(m, n)) :
+ (U :&: V = 0)%MS -> (W <= V)%MS -> W *m proj_mx U V = 0.
+Proof.
+move=> dxUV sWV; apply/eqP; rewrite -submx0 -dxUV.
+rewrite sub_capmx proj_mx_sub /= -[_ *m _](subrK W) addmx_sub // -eqmx_opp.
+by rewrite opprB proj_mx_compl_sub // (submx_trans sWV) ?addsmxSr.
+Qed.
+
+Lemma add_proj_mx m n U V (W : 'M_(m, n)) :
+ (U :&: V = 0)%MS -> (W <= U + V)%MS ->
+ W *m proj_mx U V + W *m proj_mx V U = W.
+Proof.
+move=> dxUV sWUV; apply/eqP; rewrite -subr_eq0 -submx0 -dxUV.
+rewrite -addrA sub_capmx {2}addrCA -!(opprB W).
+by rewrite !{1}addmx_sub ?proj_mx_sub ?eqmx_opp ?proj_mx_compl_sub // addsmxC.
+Qed.
+
+Lemma proj_mx_proj n (U V : 'M_n) :
+ let P := proj_mx U V in (U :&: V = 0)%MS -> P *m P = P.
+Proof. by move=> P dxUV; rewrite -{-2}[P]mul1mx proj_mx_id ?proj_mx_sub. Qed.
+
+(* Completing a partially injective matrix to get a unit matrix. *)
+
+Lemma complete_unitmx m n (U : 'M_(m, n)) (f : 'M_n) :
+ \rank (U *m f) = \rank U -> {g : 'M_n | g \in unitmx & U *m f = U *m g}.
+Proof.
+move=> injfU; pose V := <<U>>%MS; pose W := V *m f.
+pose g := proj_mx V (V^C)%MS *m f + cokermx V *m row_ebase W.
+have defW: V *m g = W.
+ rewrite mulmxDr mulmxA proj_mx_id ?genmxE ?capmx_compl //.
+ by rewrite mulmxA mulmx_coker mul0mx addr0.
+exists g; last first.
+ have /submxP[u ->]: (U <= V)%MS by rewrite genmxE.
+ by rewrite -!mulmxA defW.
+rewrite -row_full_unit -sub1mx; apply/submxP.
+have: (invmx (col_ebase W) *m W <= V *m g)%MS by rewrite defW submxMl.
+case/submxP=> v def_v; exists (invmx (row_ebase W) *m (v *m V + (V^C)%MS)).
+rewrite -mulmxA mulmxDl -mulmxA -def_v -{3}[W]mulmx_ebase -mulmxA.
+rewrite mulKmx ?col_ebase_unit // [_ *m g]mulmxDr mulmxA.
+rewrite (proj_mx_0 (capmx_compl _)) // mul0mx add0r 2!mulmxA.
+rewrite mulmxK ?row_ebase_unit // copid_mx_id ?rank_leq_row //.
+rewrite (eqmxMr _ (genmxE U)) injfU genmxE addrC -mulmxDl subrK.
+by rewrite mul1mx mulVmx ?row_ebase_unit.
+Qed.
+
+(* Mapping between two subspaces with the same dimension. *)
+
+Lemma eq_rank_unitmx m1 m2 n (U : 'M_(m1, n)) (V : 'M_(m2, n)) :
+ \rank U = \rank V -> {f : 'M_n | f \in unitmx & V :=: U *m f}%MS.
+Proof.
+move=> eqrUV; pose f := invmx (row_ebase <<U>>%MS) *m row_ebase <<V>>%MS.
+have defUf: (<<U>> *m f :=: <<V>>)%MS.
+ rewrite -[<<U>>%MS]mulmx_ebase mulmxA mulmxK ?row_ebase_unit // -mulmxA.
+ rewrite genmxE eqrUV -genmxE -{3}[<<V>>%MS]mulmx_ebase -mulmxA.
+ move: (pid_mx _ *m _) => W; apply/eqmxP.
+ by rewrite !eqmxMfull ?andbb // row_full_unit col_ebase_unit.
+have{defUf} defV: (V :=: U *m f)%MS.
+ by apply/eqmxP; rewrite -!(eqmxMr f (genmxE U)) !defUf !genmxE andbb.
+have injfU: \rank (U *m f) = \rank U by rewrite -defV eqrUV.
+by have [g injg defUg] := complete_unitmx injfU; exists g; rewrite -?defUg.
+Qed.
+
+Section SumExpr.
+
+(* This is the infrastructure to support the mxdirect predicate. We use a *)
+(* bespoke canonical structure to decompose a matrix expression into binary *)
+(* and n-ary products, using some of the "quote" technology. This lets us *)
+(* characterize direct sums as set sums whose rank is equal to the sum of the *)
+(* ranks of the individual terms. The mxsum_expr/proper_mxsum_expr structures *)
+(* below supply both the decomposition and the calculation of the rank sum. *)
+(* The mxsum_spec dependent predicate family expresses the consistency of *)
+(* these two decompositions. *)
+(* The main technical difficulty we need to overcome is the fact that *)
+(* the "catch-all" case of canonical structures has a priority lower than *)
+(* constant expansion. However, it is undesireable that local abbreviations *)
+(* be opaque for the direct-sum predicate, e.g., not be able to handle *)
+(* let S := (\sum_(i | P i) LargeExpression i)%MS in mxdirect S -> ...). *)
+(* As in "quote", we use the interleaving of constant expansion and *)
+(* canonical projection matching to achieve our goal: we use a "wrapper" type *)
+(* (indeed, the wrapped T type defined in ssrfun.v) with a self-inserting *)
+(* non-primitive constructor to gain finer control over the type and *)
+(* structure inference process. The innermost, primitive, constructor flags *)
+(* trivial sums; it is initially hidden by an eta-expansion, which has been *)
+(* made into a (default) canonical structure -- this lets type inference *)
+(* automatically insert this outer tag. *)
+(* In detail, we define three types *)
+(* mxsum_spec S r <-> There exists a finite list of matrices A1, ..., Ak *)
+(* such that S is the set sum of the Ai, and r is the sum *)
+(* of the ranks of the Ai, i.e., S = (A1 + ... + Ak)%MS *)
+(* and r = \rank A1 + ... + \rank Ak. Note that *)
+(* mxsum_spec is a recursive dependent predicate family *)
+(* whose elimination rewrites simultaneaously S, r and *)
+(* the height of S. *)
+(* proper_mxsum_expr n == The interface for proper sum expressions; this is *)
+(* a double-entry interface, keyed on both the matrix sum *)
+(* value and the rank sum. The matrix value is restricted *)
+(* to square matrices, as the "+"%MS operator always *)
+(* returns a square matrix. This interface has two *)
+(* canonical insances, for binary and n-ary sums. *)
+(* mxsum_expr m n == The interface for general sum expressions, comprising *)
+(* both proper sums and trivial sums consisting of a *)
+(* single matrix. The key values are WRAPPED as this lets *)
+(* us give priority to the "proper sum" interpretation *)
+(* (see below). To allow for trivial sums, the matrix key *)
+(* can have any dimension. The mxsum_expr interface has *)
+(* two canonical instances, for trivial and proper sums, *)
+(* keyed to the Wrap and wrap constructors, respectively. *)
+(* The projections for the two interfaces above are *)
+(* proper_mxsum_val, mxsum_val : these are respectively coercions to 'M_n *)
+(* and wrapped 'M_(m, n); thus, the matrix sum for an *)
+(* S : mxsum_expr m n can be written unwrap S. *)
+(* proper_mxsum_rank, mxsum_rank : projections to the nat and wrapped nat, *)
+(* respectively; the rank sum for S : mxsum_expr m n is *)
+(* thus written unwrap (mxsum_rank S). *)
+(* The mxdirect A predicate actually gets A in a phantom argument, which is *)
+(* used to infer an (implicit) S : mxsum_expr such that unwrap S = A; the *)
+(* actual definition is \rank (unwrap S) == unwrap (mxsum_rank S). *)
+(* Note that the inference of S is inherently ambiguous: ANY matrix can be *)
+(* viewed as a trivial sum, including one whose description is manifestly a *)
+(* proper sum. We use the wrapped type and the interaction between delta *)
+(* reduction and canonical structure inference to resolve this ambiguity in *)
+(* favor of proper sums, as follows: *)
+(* - The phantom type sets up a unification problem of the form *)
+(* unwrap (mxsum_val ?S) = A *)
+(* with unknown evar ?S : mxsum_expr m n. *)
+(* - As the constructor wrap is also a default Canonical instance for the *)
+(* wrapped type, so A is immediately replaced with unwrap (wrap A) and *)
+(* we get the residual unification problem *)
+(* mxsum_val ?S = wrap A *)
+(* - Now Coq tries to apply the proper sum Canonical instance, which has *)
+(* key projection wrap (proper_mxsum_val ?PS) where ?PS is a fresh evar *)
+(* (of type proper_mxsum_expr n). This can only succeed if m = n, and if *)
+(* a solution can be found to the recursive unification problem *)
+(* proper_mxsum_val ?PS = A *)
+(* This causes Coq to look for one of the two canonical constants for *)
+(* proper_mxsum_val (addsmx or bigop) at the head of A, delta-expanding *)
+(* A as needed, and then inferring recursively mxsum_expr structures for *)
+(* the last argument(s) of that constant. *)
+(* - If the above step fails then the wrap constant is expanded, revealing *)
+(* the primitive Wrap constructor; the unification problem now becomes *)
+(* mxsum_val ?S = Wrap A *)
+(* which fits perfectly the trivial sum canonical structure, whose key *)
+(* projection is Wrap ?B where ?B is a fresh evar. Thus the inference *)
+(* succeeds, and returns the trivial sum. *)
+(* Note that the rank projections also register canonical values, so that the *)
+(* same process can be used to infer a sum structure from the rank sum. In *)
+(* that case, however, there is no ambiguity and the inference can fail, *)
+(* because the rank sum for a trivial sum is not an arbitrary integer -- it *)
+(* must be of the form \rank ?B. It is nevertheless necessary to use the *)
+(* wrapped nat type for the rank sums, because in the non-trivial case the *)
+(* head constant of the nat expression is determined by the proper_mxsum_expr *)
+(* canonical structure, so the mxsum_expr structure must use a generic *)
+(* constant, namely wrap. *)
+
+Inductive mxsum_spec n : forall m, 'M[F]_(m, n) -> nat -> Prop :=
+ | TrivialMxsum m A
+ : @mxsum_spec n m A (\rank A)
+ | ProperMxsum m1 m2 T1 T2 r1 r2 of
+ @mxsum_spec n m1 T1 r1 & @mxsum_spec n m2 T2 r2
+ : mxsum_spec (T1 + T2)%MS (r1 + r2)%N.
+Arguments Scope mxsum_spec [nat_scope nat_scope matrix_set_scope nat_scope].
+
+Structure mxsum_expr m n := Mxsum {
+ mxsum_val :> wrapped 'M_(m, n);
+ mxsum_rank : wrapped nat;
+ _ : mxsum_spec (unwrap mxsum_val) (unwrap mxsum_rank)
+}.
+
+Canonical trivial_mxsum m n A :=
+ @Mxsum m n (Wrap A) (Wrap (\rank A)) (TrivialMxsum A).
+
+Structure proper_mxsum_expr n := ProperMxsumExpr {
+ proper_mxsum_val :> 'M_n;
+ proper_mxsum_rank : nat;
+ _ : mxsum_spec proper_mxsum_val proper_mxsum_rank
+}.
+
+Definition proper_mxsumP n (S : proper_mxsum_expr n) :=
+ let: ProperMxsumExpr _ _ termS := S return mxsum_spec S (proper_mxsum_rank S)
+ in termS.
+
+Canonical sum_mxsum n (S : proper_mxsum_expr n) :=
+ @Mxsum n n (wrap (S : 'M_n)) (wrap (proper_mxsum_rank S)) (proper_mxsumP S).
+
+Section Binary.
+Variable (m1 m2 n : nat) (S1 : mxsum_expr m1 n) (S2 : mxsum_expr m2 n).
+Fact binary_mxsum_proof :
+ mxsum_spec (unwrap S1 + unwrap S2)
+ (unwrap (mxsum_rank S1) + unwrap (mxsum_rank S2)).
+Proof. by case: S1 S2 => [A1 r1 A1P] [A2 r2 A2P]; right. Qed.
+Canonical binary_mxsum_expr := ProperMxsumExpr binary_mxsum_proof.
+End Binary.
+
+Section Nary.
+Context J (r : seq J) (P : pred J) n (S_ : J -> mxsum_expr n n).
+Fact nary_mxsum_proof :
+ mxsum_spec (\sum_(j <- r | P j) unwrap (S_ j))
+ (\sum_(j <- r | P j) unwrap (mxsum_rank (S_ j))).
+Proof.
+elim/big_rec2: _ => [|j]; first by rewrite -(mxrank0 n n); left.
+by case: (S_ j); right.
+Qed.
+Canonical nary_mxsum_expr := ProperMxsumExpr nary_mxsum_proof.
+End Nary.
+
+Definition mxdirect_def m n T of phantom 'M_(m, n) (unwrap (mxsum_val T)) :=
+ \rank (unwrap T) == unwrap (mxsum_rank T).
+
+End SumExpr.
+
+Notation mxdirect A := (mxdirect_def (Phantom 'M_(_,_) A%MS)).
+
+Lemma mxdirectP n (S : proper_mxsum_expr n) :
+ reflect (\rank S = proper_mxsum_rank S) (mxdirect S).
+Proof. exact: eqnP. Qed.
+Implicit Arguments mxdirectP [n S].
+
+Lemma mxdirect_trivial m n A : mxdirect (unwrap (@trivial_mxsum m n A)).
+Proof. exact: eqxx. Qed.
+
+Lemma mxrank_sum_leqif m n (S : mxsum_expr m n) :
+ \rank (unwrap S) <= unwrap (mxsum_rank S) ?= iff mxdirect (unwrap S).
+Proof.
+rewrite /mxdirect_def; case: S => [[A] [r] /= defAr]; split=> //=.
+elim: m A r / defAr => // m1 m2 A1 A2 r1 r2 _ leAr1 _ leAr2.
+by apply: leq_trans (leq_add leAr1 leAr2); rewrite mxrank_adds_leqif.
+Qed.
+
+Lemma mxdirectE m n (S : mxsum_expr m n) :
+ mxdirect (unwrap S) = (\rank (unwrap S) == unwrap (mxsum_rank S)).
+Proof. by []. Qed.
+
+Lemma mxdirectEgeq m n (S : mxsum_expr m n) :
+ mxdirect (unwrap S) = (\rank (unwrap S) >= unwrap (mxsum_rank S)).
+Proof. by rewrite (geq_leqif (mxrank_sum_leqif S)). Qed.
+
+Section BinaryDirect.
+
+Variables m1 m2 n : nat.
+
+Lemma mxdirect_addsE (S1 : mxsum_expr m1 n) (S2 : mxsum_expr m2 n) :
+ mxdirect (unwrap S1 + unwrap S2)
+ = [&& mxdirect (unwrap S1), mxdirect (unwrap S2)
+ & unwrap S1 :&: unwrap S2 == 0]%MS.
+Proof.
+rewrite (@mxdirectE n) /=.
+have:= leqif_add (mxrank_sum_leqif S1) (mxrank_sum_leqif S2).
+move/(leqif_trans (mxrank_adds_leqif (unwrap S1) (unwrap S2)))=> ->.
+by rewrite andbC -andbA submx0.
+Qed.
+
+Lemma mxdirect_addsP (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ reflect (A :&: B = 0)%MS (mxdirect (A + B)).
+Proof. by rewrite mxdirect_addsE !mxdirect_trivial; exact: eqP. Qed.
+
+End BinaryDirect.
+
+Section NaryDirect.
+
+Variables (P : pred I) (n : nat).
+
+Let TIsum A_ i := (A_ i :&: (\sum_(j | P j && (j != i)) A_ j) = 0 :> 'M_n)%MS.
+
+Let mxdirect_sums_recP (S_ : I -> mxsum_expr n n) :
+ reflect (forall i, P i -> mxdirect (unwrap (S_ i)) /\ TIsum (unwrap \o S_) i)
+ (mxdirect (\sum_(i | P i) (unwrap (S_ i)))).
+Proof.
+rewrite /TIsum; apply: (iffP eqnP) => /= [dxS i Pi | dxS].
+ set Si' := (\sum_(j | _) unwrap (S_ j))%MS.
+ have: mxdirect (unwrap (S_ i) + Si') by apply/eqnP; rewrite /= -!(bigD1 i).
+ by rewrite mxdirect_addsE => /and3P[-> _ /eqP].
+elim: _.+1 {-2 4}P (subxx P) (ltnSn #|P|) => // m IHm Q; move/subsetP=> sQP.
+case: (pickP Q) => [i Qi | Q0]; last by rewrite !big_pred0 ?mxrank0.
+rewrite (cardD1x Qi) !((bigD1 i) Q) //=.
+move/IHm=> <- {IHm}/=; last by apply/subsetP=> j /andP[/sQP].
+case: (dxS i (sQP i Qi)) => /eqnP=> <- TiQ_0; rewrite mxrank_disjoint_sum //.
+apply/eqP; rewrite -submx0 -{2}TiQ_0 capmxS //=.
+by apply/sumsmx_subP=> j /= /andP[Qj i'j]; rewrite (sumsmx_sup j) ?[P j]sQP.
+Qed.
+
+Lemma mxdirect_sumsP (A_ : I -> 'M_n) :
+ reflect (forall i, P i -> A_ i :&: (\sum_(j | P j && (j != i)) A_ j) = 0)%MS
+ (mxdirect (\sum_(i | P i) A_ i)).
+Proof.
+apply: (iffP (mxdirect_sums_recP _)) => dxA i /dxA; first by case.
+by rewrite mxdirect_trivial.
+Qed.
+
+Lemma mxdirect_sumsE (S_ : I -> mxsum_expr n n) (xunwrap := unwrap) :
+ reflect (and (forall i, P i -> mxdirect (unwrap (S_ i)))
+ (mxdirect (\sum_(i | P i) (xunwrap (S_ i)))))
+ (mxdirect (\sum_(i | P i) (unwrap (S_ i)))).
+Proof.
+apply: (iffP (mxdirect_sums_recP _)) => [dxS | [dxS_ dxS] i Pi].
+ by do [split; last apply/mxdirect_sumsP] => i; case/dxS.
+by split; [exact: dxS_ | exact: mxdirect_sumsP Pi].
+Qed.
+
+End NaryDirect.
+
+Section SubDaddsmx.
+
+Variables m m1 m2 n : nat.
+Variables (A : 'M[F]_(m, n)) (B1 : 'M[F]_(m1, n)) (B2 : 'M[F]_(m2, n)).
+
+CoInductive sub_daddsmx_spec : Prop :=
+ SubDaddsmxSpec A1 A2 of (A1 <= B1)%MS & (A2 <= B2)%MS & A = A1 + A2
+ & forall C1 C2, (C1 <= B1)%MS -> (C2 <= B2)%MS ->
+ A = C1 + C2 -> C1 = A1 /\ C2 = A2.
+
+Lemma sub_daddsmx : (B1 :&: B2 = 0)%MS -> (A <= B1 + B2)%MS -> sub_daddsmx_spec.
+Proof.
+move=> dxB /sub_addsmxP[u defA].
+exists (u.1 *m B1) (u.2 *m B2); rewrite ?submxMl // => C1 C2 sCB1 sCB2.
+move/(canLR (addrK _)) => defC1.
+suffices: (C2 - u.2 *m B2 <= B1 :&: B2)%MS.
+ by rewrite dxB submx0 subr_eq0 -defC1 defA; move/eqP->; rewrite addrK.
+rewrite sub_capmx -opprB -{1}(canLR (addKr _) defA) -addrA defC1.
+by rewrite !(eqmx_opp, addmx_sub) ?submxMl.
+Qed.
+
+End SubDaddsmx.
+
+Section SubDsumsmx.
+
+Variables (P : pred I) (m n : nat) (A : 'M[F]_(m, n)) (B : I -> 'M[F]_n).
+
+CoInductive sub_dsumsmx_spec : Prop :=
+ SubDsumsmxSpec A_ of forall i, P i -> (A_ i <= B i)%MS
+ & A = \sum_(i | P i) A_ i
+ & forall C, (forall i, P i -> C i <= B i)%MS ->
+ A = \sum_(i | P i) C i -> {in SimplPred P, C =1 A_}.
+
+Lemma sub_dsumsmx :
+ mxdirect (\sum_(i | P i) B i) -> (A <= \sum_(i | P i) B i)%MS ->
+ sub_dsumsmx_spec.
+Proof.
+move/mxdirect_sumsP=> dxB /sub_sumsmxP[u defA].
+pose A_ i := u i *m B i.
+exists A_ => //= [i _ | C sCB defAC i Pi]; first exact: submxMl.
+apply/eqP; rewrite -subr_eq0 -submx0 -{dxB}(dxB i Pi) /=.
+rewrite sub_capmx addmx_sub ?eqmx_opp ?submxMl ?sCB //=.
+rewrite -(subrK A (C i)) -addrA -opprB addmx_sub ?eqmx_opp //.
+ rewrite addrC defAC (bigD1 i) // addKr /= summx_sub // => j Pi'j.
+ by rewrite (sumsmx_sup j) ?sCB //; case/andP: Pi'j.
+rewrite addrC defA (bigD1 i) // addKr /= summx_sub // => j Pi'j.
+by rewrite (sumsmx_sup j) ?submxMl.
+Qed.
+
+End SubDsumsmx.
+
+Section Eigenspace.
+
+Variables (n : nat) (g : 'M_n).
+
+Definition eigenspace a := kermx (g - a%:M).
+Definition eigenvalue : pred F := fun a => eigenspace a != 0.
+
+Lemma eigenspaceP a m (W : 'M_(m, n)) :
+ reflect (W *m g = a *: W) (W <= eigenspace a)%MS.
+Proof.
+rewrite (sameP (sub_kermxP _ _) eqP).
+by rewrite mulmxBr subr_eq0 mul_mx_scalar; exact: eqP.
+Qed.
+
+Lemma eigenvalueP a :
+ reflect (exists2 v : 'rV_n, v *m g = a *: v & v != 0) (eigenvalue a).
+Proof. by apply: (iffP (rowV0Pn _)) => [] [v]; move/eigenspaceP; exists v. Qed.
+
+Lemma mxdirect_sum_eigenspace (P : pred I) a_ :
+ {in P &, injective a_} -> mxdirect (\sum_(i | P i) eigenspace (a_ i)).
+Proof.
+elim: {P}_.+1 {-2}P (ltnSn #|P|) => // m IHm P lePm inj_a.
+apply/mxdirect_sumsP=> i Pi; apply/eqP/rowV0P => v.
+rewrite sub_capmx => /andP[/eigenspaceP def_vg].
+set Vi' := (\sum_(i | _) _)%MS => Vi'v.
+have dxVi': mxdirect Vi'.
+ rewrite (cardD1x Pi) in lePm; apply: IHm => //.
+ by apply: sub_in2 inj_a => j /andP[].
+case/sub_dsumsmx: Vi'v => // u Vi'u def_v _.
+rewrite def_v big1 // => j Pi'j; apply/eqP.
+have nz_aij: a_ i - a_ j != 0.
+ by case/andP: Pi'j => Pj ne_ji; rewrite subr_eq0 eq_sym (inj_in_eq inj_a).
+case: (sub_dsumsmx dxVi' (sub0mx 1 _)) => C _ _ uniqC.
+rewrite -(eqmx_eq0 (eqmx_scale _ nz_aij)).
+rewrite (uniqC (fun k => (a_ i - a_ k) *: u k)) => // [|k Pi'k|].
+- by rewrite -(uniqC (fun _ => 0)) ?big1 // => k Pi'k; exact: sub0mx.
+- by rewrite scalemx_sub ?Vi'u.
+rewrite -{1}(subrr (v *m g)) {1}def_vg def_v scaler_sumr mulmx_suml -sumrB.
+by apply: eq_bigr => k /Vi'u/eigenspaceP->; rewrite scalerBl.
+Qed.
+
+End Eigenspace.
+
+End RowSpaceTheory.
+
+Hint Resolve submx_refl.
+Implicit Arguments submxP [F m1 m2 n A B].
+Implicit Arguments eq_row_sub [F m n v A].
+Implicit Arguments row_subP [F m1 m2 n A B].
+Implicit Arguments rV_subP [F m1 m2 n A B].
+Implicit Arguments row_subPn [F m1 m2 n A B].
+Implicit Arguments sub_rVP [F n u v].
+Implicit Arguments rV_eqP [F m1 m2 n A B].
+Implicit Arguments rowV0Pn [F m n A].
+Implicit Arguments rowV0P [F m n A].
+Implicit Arguments eqmx0P [F m n A].
+Implicit Arguments row_fullP [F m n A].
+Implicit Arguments row_freeP [F m n A].
+Implicit Arguments eqmxP [F m1 m2 n A B].
+Implicit Arguments genmxP [F m1 m2 n A B].
+Implicit Arguments addsmx_idPr [F m1 m2 n A B].
+Implicit Arguments addsmx_idPl [F m1 m2 n A B].
+Implicit Arguments sub_addsmxP [F m1 m2 m3 n A B C].
+Implicit Arguments sumsmx_sup [F I P m n A B_].
+Implicit Arguments sumsmx_subP [F I P m n A_ B].
+Implicit Arguments sub_sumsmxP [F I P m n A B_].
+Implicit Arguments sub_kermxP [F p m n A B].
+Implicit Arguments capmx_idPr [F m1 m2 n A B].
+Implicit Arguments capmx_idPl [F m1 m2 n A B].
+Implicit Arguments bigcapmx_inf [F I P m n A_ B].
+Implicit Arguments sub_bigcapmxP [F I P m n A B_].
+Implicit Arguments mxrank_injP [F m n A f].
+Implicit Arguments mxdirectP [F n S].
+Implicit Arguments mxdirect_addsP [F m1 m2 n A B].
+Implicit Arguments mxdirect_sumsP [F I P n A_].
+Implicit Arguments mxdirect_sumsE [F I P n S_].
+Implicit Arguments eigenspaceP [F n g a m W].
+Implicit Arguments eigenvalueP [F n g a].
+
+Arguments Scope mxrank [_ nat_scope nat_scope matrix_set_scope].
+Arguments Scope complmx [_ nat_scope nat_scope matrix_set_scope].
+Arguments Scope row_full [_ nat_scope nat_scope matrix_set_scope].
+Arguments Scope submx
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Arguments Scope ltmx
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Arguments Scope eqmx
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Arguments Scope addsmx
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Arguments Scope capmx
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Arguments Scope diffmx
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Prenex Implicits mxrank genmx complmx submx ltmx addsmx capmx.
+Notation "\rank A" := (mxrank A) : nat_scope.
+Notation "<< A >>" := (genmx A) : matrix_set_scope.
+Notation "A ^C" := (complmx A) : matrix_set_scope.
+Notation "A <= B" := (submx A B) : matrix_set_scope.
+Notation "A < B" := (ltmx A B) : matrix_set_scope.
+Notation "A <= B <= C" := ((submx A B) && (submx B C)) : matrix_set_scope.
+Notation "A < B <= C" := (ltmx A B && submx B C) : matrix_set_scope.
+Notation "A <= B < C" := (submx A B && ltmx B C) : matrix_set_scope.
+Notation "A < B < C" := (ltmx A B && ltmx B C) : matrix_set_scope.
+Notation "A == B" := ((submx A B) && (submx B A)) : matrix_set_scope.
+Notation "A :=: B" := (eqmx A B) : matrix_set_scope.
+Notation "A + B" := (addsmx A B) : matrix_set_scope.
+Notation "A :&: B" := (capmx A B) : matrix_set_scope.
+Notation "A :\: B" := (diffmx A B) : matrix_set_scope.
+Notation mxdirect S := (mxdirect_def (Phantom 'M_(_,_) S%MS)).
+
+Notation "\sum_ ( i <- r | P ) B" :=
+ (\big[addsmx/0%R]_(i <- r | P%B) B%MS) : matrix_set_scope.
+Notation "\sum_ ( i <- r ) B" :=
+ (\big[addsmx/0%R]_(i <- r) B%MS) : matrix_set_scope.
+Notation "\sum_ ( m <= i < n | P ) B" :=
+ (\big[addsmx/0%R]_(m <= i < n | P%B) B%MS) : matrix_set_scope.
+Notation "\sum_ ( m <= i < n ) B" :=
+ (\big[addsmx/0%R]_(m <= i < n) B%MS) : matrix_set_scope.
+Notation "\sum_ ( i | P ) B" :=
+ (\big[addsmx/0%R]_(i | P%B) B%MS) : matrix_set_scope.
+Notation "\sum_ i B" :=
+ (\big[addsmx/0%R]_i B%MS) : matrix_set_scope.
+Notation "\sum_ ( i : t | P ) B" :=
+ (\big[addsmx/0%R]_(i : t | P%B) B%MS) (only parsing) : matrix_set_scope.
+Notation "\sum_ ( i : t ) B" :=
+ (\big[addsmx/0%R]_(i : t) B%MS) (only parsing) : matrix_set_scope.
+Notation "\sum_ ( i < n | P ) B" :=
+ (\big[addsmx/0%R]_(i < n | P%B) B%MS) : matrix_set_scope.
+Notation "\sum_ ( i < n ) B" :=
+ (\big[addsmx/0%R]_(i < n) B%MS) : matrix_set_scope.
+Notation "\sum_ ( i 'in' A | P ) B" :=
+ (\big[addsmx/0%R]_(i in A | P%B) B%MS) : matrix_set_scope.
+Notation "\sum_ ( i 'in' A ) B" :=
+ (\big[addsmx/0%R]_(i in A) B%MS) : matrix_set_scope.
+
+Notation "\bigcap_ ( i <- r | P ) B" :=
+ (\big[capmx/1%:M]_(i <- r | P%B) B%MS) : matrix_set_scope.
+Notation "\bigcap_ ( i <- r ) B" :=
+ (\big[capmx/1%:M]_(i <- r) B%MS) : matrix_set_scope.
+Notation "\bigcap_ ( m <= i < n | P ) B" :=
+ (\big[capmx/1%:M]_(m <= i < n | P%B) B%MS) : matrix_set_scope.
+Notation "\bigcap_ ( m <= i < n ) B" :=
+ (\big[capmx/1%:M]_(m <= i < n) B%MS) : matrix_set_scope.
+Notation "\bigcap_ ( i | P ) B" :=
+ (\big[capmx/1%:M]_(i | P%B) B%MS) : matrix_set_scope.
+Notation "\bigcap_ i B" :=
+ (\big[capmx/1%:M]_i B%MS) : matrix_set_scope.
+Notation "\bigcap_ ( i : t | P ) B" :=
+ (\big[capmx/1%:M]_(i : t | P%B) B%MS) (only parsing) : matrix_set_scope.
+Notation "\bigcap_ ( i : t ) B" :=
+ (\big[capmx/1%:M]_(i : t) B%MS) (only parsing) : matrix_set_scope.
+Notation "\bigcap_ ( i < n | P ) B" :=
+ (\big[capmx/1%:M]_(i < n | P%B) B%MS) : matrix_set_scope.
+Notation "\bigcap_ ( i < n ) B" :=
+ (\big[capmx/1%:M]_(i < n) B%MS) : matrix_set_scope.
+Notation "\bigcap_ ( i 'in' A | P ) B" :=
+ (\big[capmx/1%:M]_(i in A | P%B) B%MS) : matrix_set_scope.
+Notation "\bigcap_ ( i 'in' A ) B" :=
+ (\big[capmx/1%:M]_(i in A) B%MS) : matrix_set_scope.
+
+Section CardGL.
+
+Variable F : finFieldType.
+
+Lemma card_GL n : n > 0 ->
+ #|'GL_n[F]| = (#|F| ^ 'C(n, 2) * \prod_(1 <= i < n.+1) (#|F| ^ i - 1))%N.
+Proof.
+case: n => // n' _; set n := n'.+1; set p := #|F|.
+rewrite big_nat_rev big_add1 -triangular_sum expn_sum -big_split /=.
+pose fr m := [pred A : 'M[F]_(m, n) | \rank A == m].
+set m := {-7}n; transitivity #|fr m|.
+ by rewrite cardsT /= card_sub; apply: eq_card => A; rewrite -row_free_unit.
+elim: m (leqnn m : m <= n) => [_|m IHm]; last move/ltnW=> le_mn.
+ rewrite (@eq_card1 _ (0 : 'M_(0, n))) ?big_geq //= => A.
+ by rewrite flatmx0 !inE !eqxx.
+rewrite big_nat_recr // -{}IHm //= !subSS mulnBr muln1 -expnD subnKC //.
+rewrite -sum_nat_const /= -sum1_card -add1n.
+rewrite (partition_big dsubmx (fr m)) /= => [|A]; last first.
+ rewrite !inE -{1}(vsubmxK A); move: {A}(_ A) (_ A) => Ad Au Afull.
+ rewrite eqn_leq rank_leq_row -(leq_add2l (\rank Au)) -mxrank_sum_cap.
+ rewrite {1 3}[@mxrank]lock addsmxE (eqnP Afull) -lock -addnA.
+ by rewrite leq_add ?rank_leq_row ?leq_addr.
+apply: eq_bigr => A rAm; rewrite (reindex (col_mx^~ A)) /=; last first.
+ exists usubmx => [v _ | vA]; first by rewrite col_mxKu.
+ by case/andP=> _ /eqP <-; rewrite vsubmxK.
+transitivity #|~: [set v *m A | v in 'rV_m]|; last first.
+ rewrite cardsCs setCK card_imset ?card_matrix ?card_ord ?mul1n //.
+ have [B AB1] := row_freeP rAm; apply: can_inj (mulmx^~ B) _ => v.
+ by rewrite -mulmxA AB1 mulmx1.
+rewrite -sum1_card; apply: eq_bigl => v; rewrite !inE col_mxKd eqxx.
+rewrite andbT eqn_leq rank_leq_row /= -(leq_add2r (\rank (v :&: A)%MS)).
+rewrite -addsmxE mxrank_sum_cap (eqnP rAm) addnAC leq_add2r.
+rewrite (ltn_leqif (mxrank_leqif_sup _)) ?capmxSl // sub_capmx submx_refl.
+by congr (~~ _); apply/submxP/imsetP=> [] [u]; exists u.
+Qed.
+
+(* An alternate, somewhat more elementary proof, that does not rely on the *)
+(* row-space theory, but directly performs the LUP decomposition. *)
+Lemma LUP_card_GL n : n > 0 ->
+ #|'GL_n[F]| = (#|F| ^ 'C(n, 2) * \prod_(1 <= i < n.+1) (#|F| ^ i - 1))%N.
+Proof.
+case: n => // n' _; set n := n'.+1; set p := #|F|.
+rewrite cardsT /= card_sub /GRing.unit /= big_add1 /= -triangular_sum -/n.
+elim: {n'}n => [|n IHn].
+ rewrite !big_geq // mul1n (@eq_card _ _ predT) ?card_matrix //= => M.
+ by rewrite {1}[M]flatmx0 -(flatmx0 1%:M) unitmx1.
+rewrite !big_nat_recr //= expnD mulnAC mulnA -{}IHn -mulnA mulnC.
+set LHS := #|_|; rewrite -[n.+1]muln1 -{2}[n]mul1n {}/LHS.
+rewrite -!card_matrix subn1 -(cardC1 0) -mulnA; set nzC := predC1 _.
+rewrite -sum1_card (partition_big lsubmx nzC) => [|A]; last first.
+ rewrite unitmxE unitfE; apply: contra; move/eqP=> v0.
+ rewrite -[A]hsubmxK v0 -[n.+1]/(1 + n)%N -col_mx0.
+ rewrite -[rsubmx _]vsubmxK -det_tr tr_row_mx !tr_col_mx !trmx0.
+ by rewrite det_lblock [0]mx11_scalar det_scalar1 mxE mul0r.
+rewrite -sum_nat_const; apply: eq_bigr; rewrite /= -[n.+1]/(1 + n)%N => v nzv.
+case: (pickP (fun i => v i 0 != 0)) => [k nza | v0]; last first.
+ by case/eqP: nzv; apply/colP=> i; move/eqP: (v0 i); rewrite mxE.
+have xrkK: involutive (@xrow F _ _ 0 k).
+ by move=> m A /=; rewrite /xrow -row_permM tperm2 row_perm1.
+rewrite (reindex_inj (inv_inj (xrkK (1 + n)%N))) /= -[n.+1]/(1 + n)%N.
+rewrite (partition_big ursubmx xpredT) //= -sum_nat_const.
+apply: eq_bigr => u _; set a : F := v _ _ in nza.
+set v1 : 'cV_(1 + n) := xrow 0 k v.
+have def_a: usubmx v1 = a%:M.
+ by rewrite [_ v1]mx11_scalar mxE lshift0 mxE tpermL.
+pose Schur := dsubmx v1 *m (a^-1 *: u).
+pose L : 'M_(1 + n) := block_mx a%:M 0 (dsubmx v1) 1%:M.
+pose U B : 'M_(1 + n) := block_mx 1 (a^-1 *: u) 0 B.
+rewrite (reindex (fun B => L *m U B)); last first.
+ exists (fun A1 => drsubmx A1 - Schur) => [B _ | A1].
+ by rewrite mulmx_block block_mxKdr mul1mx addrC addKr.
+ rewrite !inE mulmx_block !mulmx0 mul0mx !mulmx1 !addr0 mul1mx addrC subrK.
+ rewrite mul_scalar_mx scalerA divff // scale1r andbC; case/and3P => /eqP <- _.
+ rewrite -{1}(hsubmxK A1) xrowE mul_mx_row row_mxKl -xrowE => /eqP def_v.
+ rewrite -def_a block_mxEh vsubmxK /v1 -def_v xrkK.
+ apply: trmx_inj; rewrite tr_row_mx tr_col_mx trmx_ursub trmx_drsub trmx_lsub.
+ by rewrite hsubmxK vsubmxK.
+rewrite -sum1_card; apply: eq_bigl => B; rewrite xrowE unitmxE.
+rewrite !det_mulmx unitrM -unitmxE unitmx_perm det_lblock det_ublock.
+rewrite !det_scalar1 det1 mulr1 mul1r unitrM unitfE nza -unitmxE.
+rewrite mulmx_block !mulmx0 mul0mx !addr0 !mulmx1 mul1mx block_mxKur.
+rewrite mul_scalar_mx scalerA divff // scale1r eqxx andbT.
+by rewrite block_mxEh mul_mx_row row_mxKl -def_a vsubmxK -xrowE xrkK eqxx andbT.
+Qed.
+
+Lemma card_GL_1 : #|'GL_1[F]| = #|F|.-1.
+Proof. by rewrite card_GL // mul1n big_nat1 expn1 subn1. Qed.
+
+Lemma card_GL_2 : #|'GL_2[F]| = (#|F| * #|F|.-1 ^ 2 * #|F|.+1)%N.
+Proof.
+rewrite card_GL // big_ltn // big_nat1 expn1 -(addn1 #|F|) -subn1 -!mulnA.
+by rewrite -subn_sqr.
+Qed.
+
+End CardGL.
+
+Lemma logn_card_GL_p n p : prime p -> logn p #|'GL_n(p)| = 'C(n, 2).
+Proof.
+move=> p_pr; have p_gt1 := prime_gt1 p_pr.
+have p_i_gt0: p ^ _ > 0 by move=> i; rewrite expn_gt0 ltnW.
+rewrite (card_GL _ (ltn0Sn n.-1)) card_ord Fp_cast // big_add1 /=.
+pose p'gt0 m := m > 0 /\ logn p m = 0%N.
+suffices [Pgt0 p'P]: p'gt0 (\prod_(0 <= i < n.-1.+1) (p ^ i.+1 - 1))%N.
+ by rewrite lognM // p'P pfactorK //; case n.
+apply big_ind => [|m1 m2 [m10 p'm1] [m20]|i _]; rewrite {}/p'gt0 ?logn1 //.
+ by rewrite muln_gt0 m10 lognM ?p'm1.
+rewrite lognE -if_neg subn_gt0 p_pr /= -{1 2}(exp1n i.+1) ltn_exp2r // p_gt1.
+by rewrite dvdn_subr ?dvdn_exp // gtnNdvd.
+Qed.
+
+Section MatrixAlgebra.
+
+Variables F : fieldType.
+
+Local Notation "A \in R" := (@submx F _ _ _ (mxvec A) R).
+
+Lemma mem0mx m n (R : 'A_(m, n)) : 0 \in R.
+Proof. by rewrite linear0 sub0mx. Qed.
+
+Lemma memmx0 n A : (A \in (0 : 'A_n)) -> A = 0.
+Proof. by rewrite submx0 mxvec_eq0; move/eqP. Qed.
+
+Lemma memmx1 n (A : 'M_n) : (A \in mxvec 1%:M) = is_scalar_mx A.
+Proof.
+apply/sub_rVP/is_scalar_mxP=> [[a] | [a ->]].
+ by rewrite -linearZ scale_scalar_mx mulr1 => /(can_inj mxvecK); exists a.
+by exists a; rewrite -linearZ scale_scalar_mx mulr1.
+Qed.
+
+Lemma memmx_subP m1 m2 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) :
+ reflect (forall A, A \in R1 -> A \in R2) (R1 <= R2)%MS.
+Proof.
+apply: (iffP idP) => [sR12 A R1_A | sR12]; first exact: submx_trans sR12.
+by apply/rV_subP=> vA; rewrite -(vec_mxK vA); exact: sR12.
+Qed.
+Implicit Arguments memmx_subP [m1 m2 n R1 R2].
+
+Lemma memmx_eqP m1 m2 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) :
+ reflect (forall A, (A \in R1) = (A \in R2)) (R1 == R2)%MS.
+Proof.
+apply: (iffP eqmxP) => [eqR12 A | eqR12]; first by rewrite eqR12.
+by apply/eqmxP; apply/rV_eqP=> vA; rewrite -(vec_mxK vA) eqR12.
+Qed.
+Implicit Arguments memmx_eqP [m1 m2 n R1 R2].
+
+Lemma memmx_addsP m1 m2 n A (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) :
+ reflect (exists D, [/\ D.1 \in R1, D.2 \in R2 & A = D.1 + D.2])
+ (A \in R1 + R2)%MS.
+Proof.
+apply: (iffP sub_addsmxP) => [[u /(canRL mxvecK)->] | [D []]].
+ exists (vec_mx (u.1 *m R1), vec_mx (u.2 *m R2)).
+ by rewrite /= linearD !vec_mxK !submxMl.
+case/submxP=> u1 defD1 /submxP[u2 defD2] ->.
+by exists (u1, u2); rewrite linearD /= defD1 defD2.
+Qed.
+Implicit Arguments memmx_addsP [m1 m2 n A R1 R2].
+
+Lemma memmx_sumsP (I : finType) (P : pred I) n (A : 'M_n) R_ :
+ reflect (exists2 A_, A = \sum_(i | P i) A_ i & forall i, A_ i \in R_ i)
+ (A \in \sum_(i | P i) R_ i)%MS.
+Proof.
+apply: (iffP sub_sumsmxP) => [[C defA] | [A_ -> R_A] {A}].
+ exists (fun i => vec_mx (C i *m R_ i)) => [|i].
+ by rewrite -linear_sum -defA /= mxvecK.
+ by rewrite vec_mxK submxMl.
+exists (fun i => mxvec (A_ i) *m pinvmx (R_ i)).
+by rewrite linear_sum; apply: eq_bigr => i _; rewrite mulmxKpV.
+Qed.
+Implicit Arguments memmx_sumsP [I P n A R_].
+
+Lemma has_non_scalar_mxP m n (R : 'A_(m, n)) :
+ (1%:M \in R)%MS ->
+ reflect (exists2 A, A \in R & ~~ is_scalar_mx A)%MS (1 < \rank R).
+Proof.
+case: (posnP n) => [-> | n_gt0] in R *; set S := mxvec _ => sSR.
+ by rewrite [R]thinmx0 mxrank0; right; case; rewrite /is_scalar_mx ?insubF.
+have rankS: \rank S = 1%N.
+ apply/eqP; rewrite eqn_leq rank_leq_row lt0n mxrank_eq0 mxvec_eq0.
+ by rewrite -mxrank_eq0 mxrank1 -lt0n.
+rewrite -{2}rankS (ltn_leqif (mxrank_leqif_sup sSR)).
+apply: (iffP idP) => [/row_subPn[i] | [A sAR]].
+ rewrite -[row i R]vec_mxK memmx1; set A := vec_mx _ => nsA.
+ by exists A; rewrite // vec_mxK row_sub.
+by rewrite -memmx1; apply: contra; exact: submx_trans.
+Qed.
+
+Definition mulsmx m1 m2 n (R1 : 'A[F]_(m1, n)) (R2 : 'A_(m2, n)) :=
+ (\sum_i <<R1 *m lin_mx (mulmxr (vec_mx (row i R2)))>>)%MS.
+
+Arguments Scope mulsmx
+ [nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+
+Local Notation "R1 * R2" := (mulsmx R1 R2) : matrix_set_scope.
+
+Lemma genmx_muls m1 m2 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) :
+ <<(R1 * R2)%MS>>%MS = (R1 * R2)%MS.
+Proof. by rewrite genmx_sums; apply: eq_bigr => i; rewrite genmx_id. Qed.
+
+Lemma mem_mulsmx m1 m2 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) A1 A2 :
+ (A1 \in R1 -> A2 \in R2 -> A1 *m A2 \in R1 * R2)%MS.
+Proof.
+move=> R_A1 R_A2; rewrite -[A2]mxvecK; case/submxP: R_A2 => a ->{A2}.
+rewrite mulmx_sum_row !linear_sum summx_sub // => i _.
+rewrite !linearZ scalemx_sub {a}//= (sumsmx_sup i) // genmxE.
+rewrite -[A1]mxvecK; case/submxP: R_A1 => a ->{A1}.
+by apply/submxP; exists a; rewrite mulmxA mul_rV_lin.
+Qed.
+
+Lemma mulsmx_subP m1 m2 m n
+ (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) (R : 'A_(m, n)) :
+ reflect (forall A1 A2, A1 \in R1 -> A2 \in R2 -> A1 *m A2 \in R)
+ (R1 * R2 <= R)%MS.
+Proof.
+apply: (iffP memmx_subP) => [sR12R A1 A2 R_A1 R_A2 | sR12R A].
+ by rewrite sR12R ?mem_mulsmx.
+case/memmx_sumsP=> A_ -> R_A; rewrite linear_sum summx_sub //= => j _.
+rewrite (submx_trans (R_A _)) // genmxE; apply/row_subP=> i.
+by rewrite row_mul mul_rV_lin sR12R ?vec_mxK ?row_sub.
+Qed.
+Implicit Arguments mulsmx_subP [m1 m2 m n R1 R2 R].
+
+Lemma mulsmxS m1 m2 m3 m4 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n))
+ (R3 : 'A_(m3, n)) (R4 : 'A_(m4, n)) :
+ (R1 <= R3 -> R2 <= R4 -> R1 * R2 <= R3 * R4)%MS.
+Proof.
+move=> sR13 sR24; apply/mulsmx_subP=> A1 A2 R_A1 R_A2.
+by apply: mem_mulsmx; [exact: submx_trans sR13 | exact: submx_trans sR24].
+Qed.
+
+Lemma muls_eqmx m1 m2 m3 m4 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n))
+ (R3 : 'A_(m3, n)) (R4 : 'A_(m4, n)) :
+ (R1 :=: R3 -> R2 :=: R4 -> R1 * R2 = R3 * R4)%MS.
+Proof.
+move=> eqR13 eqR24; rewrite -(genmx_muls R1 R2) -(genmx_muls R3 R4).
+by apply/genmxP; rewrite !mulsmxS ?eqR13 ?eqR24.
+Qed.
+
+Lemma mulsmxP m1 m2 n A (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) :
+ reflect (exists2 A1, forall i, A1 i \in R1
+ & exists2 A2, forall i, A2 i \in R2
+ & A = \sum_(i < n ^ 2) A1 i *m A2 i)
+ (A \in R1 * R2)%MS.
+Proof.
+apply: (iffP idP) => [R_A|[A1 R_A1 [A2 R_A2 ->{A}]]]; last first.
+ by rewrite linear_sum summx_sub // => i _; rewrite mem_mulsmx.
+have{R_A}: (A \in R1 * <<R2>>)%MS.
+ by apply: memmx_subP R_A; rewrite mulsmxS ?genmxE.
+case/memmx_sumsP=> A_ -> R_A; pose A2_ i := vec_mx (row i <<R2>>%MS).
+pose A1_ i := mxvec (A_ i) *m pinvmx (R1 *m lin_mx (mulmxr (A2_ i))) *m R1.
+exists (vec_mx \o A1_) => [i|]; first by rewrite vec_mxK submxMl.
+exists A2_ => [i|]; first by rewrite vec_mxK -(genmxE R2) row_sub.
+apply: eq_bigr => i _; rewrite -[_ *m _](mx_rV_lin (mulmxr_linear _ _)).
+by rewrite -mulmxA mulmxKpV ?mxvecK // -(genmxE (_ *m _)) R_A.
+Qed.
+Implicit Arguments mulsmxP [m1 m2 n A R1 R2].
+
+Lemma mulsmxA m1 m2 m3 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) (R3 : 'A_(m3, n)) :
+ (R1 * (R2 * R3) = R1 * R2 * R3)%MS.
+Proof.
+rewrite -(genmx_muls (_ * _)%MS) -genmx_muls; apply/genmxP; apply/andP; split.
+ apply/mulsmx_subP=> A1 A23 R_A1; case/mulsmxP=> A2 R_A2 [A3 R_A3 ->{A23}].
+ by rewrite !linear_sum summx_sub //= => i _; rewrite mulmxA !mem_mulsmx.
+apply/mulsmx_subP=> _ A3 /mulsmxP[A1 R_A1 [A2 R_A2 ->]] R_A3.
+rewrite mulmx_suml linear_sum summx_sub //= => i _.
+by rewrite -mulmxA !mem_mulsmx.
+Qed.
+
+Lemma mulsmx_addl m1 m2 m3 n
+ (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) (R3 : 'A_(m3, n)) :
+ ((R1 + R2) * R3 = R1 * R3 + R2 * R3)%MS.
+Proof.
+rewrite -(genmx_muls R2 R3) -(genmx_muls R1 R3) -genmx_muls -genmx_adds.
+apply/genmxP; rewrite andbC addsmx_sub !mulsmxS ?addsmxSl ?addsmxSr //=.
+apply/mulsmx_subP=> _ A3 /memmx_addsP[A [R_A1 R_A2 ->]] R_A3.
+by rewrite mulmxDl linearD addmx_sub_adds ?mem_mulsmx.
+Qed.
+
+Lemma mulsmx_addr m1 m2 m3 n
+ (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) (R3 : 'A_(m3, n)) :
+ (R1 * (R2 + R3) = R1 * R2 + R1 * R3)%MS.
+Proof.
+rewrite -(genmx_muls R1 R3) -(genmx_muls R1 R2) -genmx_muls -genmx_adds.
+apply/genmxP; rewrite andbC addsmx_sub !mulsmxS ?addsmxSl ?addsmxSr //=.
+apply/mulsmx_subP=> A1 _ R_A1 /memmx_addsP[A [R_A2 R_A3 ->]].
+by rewrite mulmxDr linearD addmx_sub_adds ?mem_mulsmx.
+Qed.
+
+Lemma mulsmx0 m1 m2 n (R1 : 'A_(m1, n)) : (R1 * (0 : 'A_(m2, n)) = 0)%MS.
+Proof.
+apply/eqP; rewrite -submx0; apply/mulsmx_subP=> A1 A0 _.
+by rewrite [A0 \in 0]eqmx0 => /memmx0->; rewrite mulmx0 mem0mx.
+Qed.
+
+Lemma muls0mx m1 m2 n (R2 : 'A_(m2, n)) : ((0 : 'A_(m1, n)) * R2 = 0)%MS.
+Proof.
+apply/eqP; rewrite -submx0; apply/mulsmx_subP=> A0 A2.
+by rewrite [A0 \in 0]eqmx0 => /memmx0->; rewrite mul0mx mem0mx.
+Qed.
+
+Definition left_mx_ideal m1 m2 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) :=
+ (R1 * R2 <= R2)%MS.
+
+Definition right_mx_ideal m1 m2 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) :=
+ (R2 * R1 <= R2)%MS.
+
+Definition mx_ideal m1 m2 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) :=
+ left_mx_ideal R1 R2 && right_mx_ideal R1 R2.
+
+Definition mxring_id m n (R : 'A_(m, n)) e :=
+ [/\ e != 0,
+ e \in R,
+ forall A, A \in R -> e *m A = A
+ & forall A, A \in R -> A *m e = A]%MS.
+
+Definition has_mxring_id m n (R : 'A[F]_(m , n)) :=
+ (R != 0) &&
+ (row_mx 0 (row_mx (mxvec R) (mxvec R))
+ <= row_mx (cokermx R) (row_mx (lin_mx (mulmx R \o lin_mulmx))
+ (lin_mx (mulmx R \o lin_mulmxr))))%MS.
+
+Definition mxring m n (R : 'A_(m, n)) :=
+ left_mx_ideal R R && has_mxring_id R.
+
+Lemma mxring_idP m n (R : 'A_(m, n)) :
+ reflect (exists e, mxring_id R e) (has_mxring_id R).
+Proof.
+apply: (iffP andP) => [[nzR] | [e [nz_e Re ideR idRe]]].
+ case/submxP=> v; rewrite -[v]vec_mxK; move/vec_mx: v => e.
+ rewrite !mul_mx_row; case/eq_row_mx => /eqP.
+ rewrite eq_sym -submxE => Re.
+ case/eq_row_mx; rewrite !{1}mul_rV_lin1 /= mxvecK.
+ set u := (_ *m _) => /(can_inj mxvecK) idRe /(can_inj mxvecK) ideR.
+ exists e; split=> // [ | A /submxP[a defA] | A /submxP[a defA]].
+ - by apply: contra nzR; rewrite ideR => /eqP->; rewrite !linear0.
+ - by rewrite -{2}[A]mxvecK defA idRe mulmxA mx_rV_lin -defA /= mxvecK.
+ by rewrite -{2}[A]mxvecK defA ideR mulmxA mx_rV_lin -defA /= mxvecK.
+split.
+ by apply: contraNneq nz_e => R0; rewrite R0 eqmx0 in Re; rewrite (memmx0 Re).
+apply/submxP; exists (mxvec e); rewrite !mul_mx_row !{1}mul_rV_lin1.
+rewrite submxE in Re; rewrite {Re}(eqP Re).
+congr (row_mx 0 (row_mx (mxvec _) (mxvec _))); apply/row_matrixP=> i.
+ by rewrite !row_mul !mul_rV_lin1 /= mxvecK ideR vec_mxK ?row_sub.
+by rewrite !row_mul !mul_rV_lin1 /= mxvecK idRe vec_mxK ?row_sub.
+Qed.
+Implicit Arguments mxring_idP [m n R].
+
+Section CentMxDef.
+
+Variables (m n : nat) (R : 'A[F]_(m, n)).
+
+Definition cent_mx_fun (B : 'M[F]_n) := R *m lin_mx (mulmxr B \- mulmx B).
+
+Lemma cent_mx_fun_is_linear : linear cent_mx_fun.
+Proof.
+move=> a A B; apply/row_matrixP=> i; rewrite linearP row_mul mul_rV_lin.
+rewrite /= {-3}[row]lock row_mul mul_rV_lin -lock row_mul mul_rV_lin.
+by rewrite -linearP -(linearP [linear of mulmx _ \- mulmxr _]).
+Qed.
+Canonical cent_mx_fun_additive := Additive cent_mx_fun_is_linear.
+Canonical cent_mx_fun_linear := Linear cent_mx_fun_is_linear.
+
+Definition cent_mx := kermx (lin_mx cent_mx_fun).
+
+Definition center_mx := (R :&: cent_mx)%MS.
+
+End CentMxDef.
+
+Local Notation "''C' ( R )" := (cent_mx R) : matrix_set_scope.
+Local Notation "''Z' ( R )" := (center_mx R) : matrix_set_scope.
+
+Lemma cent_rowP m n B (R : 'A_(m, n)) :
+ reflect (forall i (A := vec_mx (row i R)), A *m B = B *m A) (B \in 'C(R))%MS.
+Proof.
+apply: (iffP sub_kermxP); rewrite mul_vec_lin => cBE.
+ move/(canRL mxvecK): cBE => cBE i A /=; move/(congr1 (row i)): cBE.
+ rewrite row_mul mul_rV_lin -/A; move/(canRL mxvecK).
+ by move/(canRL (subrK _)); rewrite !linear0 add0r.
+apply: (canLR vec_mxK); apply/row_matrixP=> i.
+by rewrite row_mul mul_rV_lin /= cBE subrr !linear0.
+Qed.
+Implicit Arguments cent_rowP [m n B R].
+
+Lemma cent_mxP m n B (R : 'A_(m, n)) :
+ reflect (forall A, A \in R -> A *m B = B *m A) (B \in 'C(R))%MS.
+Proof.
+apply: (iffP cent_rowP) => cEB => [A sAE | i A].
+ rewrite -[A]mxvecK -(mulmxKpV sAE); move: (mxvec A *m _) => u.
+ rewrite !mulmx_sum_row !linear_sum mulmx_suml; apply: eq_bigr => i _ /=.
+ by rewrite !linearZ -scalemxAl /= cEB.
+by rewrite cEB // vec_mxK row_sub.
+Qed.
+Implicit Arguments cent_mxP [m n B R].
+
+Lemma scalar_mx_cent m n a (R : 'A_(m, n)) : (a%:M \in 'C(R))%MS.
+Proof. by apply/cent_mxP=> A _; exact: scalar_mxC. Qed.
+
+Lemma center_mx_sub m n (R : 'A_(m, n)) : ('Z(R) <= R)%MS.
+Proof. exact: capmxSl. Qed.
+
+Lemma center_mxP m n A (R : 'A_(m, n)) :
+ reflect (A \in R /\ forall B, B \in R -> B *m A = A *m B)
+ (A \in 'Z(R))%MS.
+Proof.
+rewrite sub_capmx; case R_A: (A \in R); last by right; case.
+by apply: (iffP cent_mxP) => [cAR | [_ cAR]].
+Qed.
+Implicit Arguments center_mxP [m n A R].
+
+Lemma mxring_id_uniq m n (R : 'A_(m, n)) e1 e2 :
+ mxring_id R e1 -> mxring_id R e2 -> e1 = e2.
+Proof.
+by case=> [_ Re1 idRe1 _] [_ Re2 _ ide2R]; rewrite -(idRe1 _ Re2) ide2R.
+Qed.
+
+Lemma cent_mx_ideal m n (R : 'A_(m, n)) : left_mx_ideal 'C(R)%MS 'C(R)%MS.
+Proof.
+apply/mulsmx_subP=> A1 A2 C_A1 C_A2; apply/cent_mxP=> B R_B.
+by rewrite mulmxA (cent_mxP C_A1) // -!mulmxA (cent_mxP C_A2).
+Qed.
+
+Lemma cent_mx_ring m n (R : 'A_(m, n)) : n > 0 -> mxring 'C(R)%MS.
+Proof.
+move=> n_gt0; rewrite /mxring cent_mx_ideal; apply/mxring_idP.
+exists 1%:M; split=> [||A _|A _]; rewrite ?mulmx1 ?mul1mx ?scalar_mx_cent //.
+by rewrite -mxrank_eq0 mxrank1 -lt0n.
+Qed.
+
+Lemma mxdirect_adds_center m1 m2 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) :
+ mx_ideal (R1 + R2)%MS R1 -> mx_ideal (R1 + R2)%MS R2 ->
+ mxdirect (R1 + R2) ->
+ ('Z((R1 + R2)%MS) :=: 'Z(R1) + 'Z(R2))%MS.
+Proof.
+case/andP=> idlR1 idrR1 /andP[idlR2 idrR2] /mxdirect_addsP dxR12.
+apply/eqmxP/andP; split.
+ apply/memmx_subP=> z0; rewrite sub_capmx => /andP[].
+ case/memmx_addsP=> z [R1z1 R2z2 ->{z0}] Cz.
+ rewrite linearD addmx_sub_adds //= ?sub_capmx ?R1z1 ?R2z2 /=.
+ apply/cent_mxP=> A R1_A; have R_A := submx_trans R1_A (addsmxSl R1 R2).
+ have Rz2 := submx_trans R2z2 (addsmxSr R1 R2).
+ rewrite -{1}[z.1](addrK z.2) mulmxBr (cent_mxP Cz) // mulmxDl.
+ rewrite [A *m z.2]memmx0 1?[z.2 *m A]memmx0 ?addrK //.
+ by rewrite -dxR12 sub_capmx (mulsmx_subP idlR1) // (mulsmx_subP idrR2).
+ by rewrite -dxR12 sub_capmx (mulsmx_subP idrR1) // (mulsmx_subP idlR2).
+ apply/cent_mxP=> A R2_A; have R_A := submx_trans R2_A (addsmxSr R1 R2).
+ have Rz1 := submx_trans R1z1 (addsmxSl R1 R2).
+ rewrite -{1}[z.2](addKr z.1) mulmxDr (cent_mxP Cz) // mulmxDl.
+ rewrite mulmxN [A *m z.1]memmx0 1?[z.1 *m A]memmx0 ?addKr //.
+ by rewrite -dxR12 sub_capmx (mulsmx_subP idrR1) // (mulsmx_subP idlR2).
+ by rewrite -dxR12 sub_capmx (mulsmx_subP idlR1) // (mulsmx_subP idrR2).
+rewrite addsmx_sub; apply/andP; split.
+ apply/memmx_subP=> z; rewrite sub_capmx => /andP[R1z cR1z].
+ have Rz := submx_trans R1z (addsmxSl R1 R2).
+ rewrite sub_capmx Rz; apply/cent_mxP=> A0.
+ case/memmx_addsP=> A [R1_A1 R2_A2] ->{A0}.
+ have R_A2 := submx_trans R2_A2 (addsmxSr R1 R2).
+ rewrite mulmxDl mulmxDr (cent_mxP cR1z) //; congr (_ + _).
+ rewrite [A.2 *m z]memmx0 1?[z *m A.2]memmx0 //.
+ by rewrite -dxR12 sub_capmx (mulsmx_subP idrR1) // (mulsmx_subP idlR2).
+ by rewrite -dxR12 sub_capmx (mulsmx_subP idlR1) // (mulsmx_subP idrR2).
+apply/memmx_subP=> z; rewrite !sub_capmx => /andP[R2z cR2z].
+have Rz := submx_trans R2z (addsmxSr R1 R2); rewrite Rz.
+apply/cent_mxP=> _ /memmx_addsP[A [R1_A1 R2_A2 ->]].
+rewrite mulmxDl mulmxDr (cent_mxP cR2z _ R2_A2) //; congr (_ + _).
+have R_A1 := submx_trans R1_A1 (addsmxSl R1 R2).
+rewrite [A.1 *m z]memmx0 1?[z *m A.1]memmx0 //.
+ by rewrite -dxR12 sub_capmx (mulsmx_subP idlR1) // (mulsmx_subP idrR2).
+by rewrite -dxR12 sub_capmx (mulsmx_subP idrR1) // (mulsmx_subP idlR2).
+Qed.
+
+Lemma mxdirect_sums_center (I : finType) m n (R : 'A_(m, n)) R_ :
+ (\sum_i R_ i :=: R)%MS -> mxdirect (\sum_i R_ i) ->
+ (forall i : I, mx_ideal R (R_ i)) ->
+ ('Z(R) :=: \sum_i 'Z(R_ i))%MS.
+Proof.
+move=> defR dxR idealR.
+have sR_R: (R_ _ <= R)%MS by move=> i; rewrite -defR (sumsmx_sup i).
+have anhR i j A B : i != j -> A \in R_ i -> B \in R_ j -> A *m B = 0.
+ move=> ne_ij RiA RjB; apply: memmx0.
+ have [[_ idRiR] [idRRj _]] := (andP (idealR i), andP (idealR j)).
+ rewrite -(mxdirect_sumsP dxR j) // sub_capmx (sumsmx_sup i) //.
+ by rewrite (mulsmx_subP idRRj) // (memmx_subP (sR_R i)).
+ by rewrite (mulsmx_subP idRiR) // (memmx_subP (sR_R j)).
+apply/eqmxP/andP; split.
+ apply/memmx_subP=> Z; rewrite sub_capmx => /andP[].
+ rewrite -{1}defR => /memmx_sumsP[z ->{Z} Rz cRz].
+ apply/memmx_sumsP; exists z => // i; rewrite sub_capmx Rz.
+ apply/cent_mxP=> A RiA; have:= cent_mxP cRz A (memmx_subP (sR_R i) A RiA).
+ rewrite (bigD1 i) //= mulmxDl mulmxDr mulmx_suml mulmx_sumr.
+ by rewrite !big1 ?addr0 // => j; last rewrite eq_sym; move/anhR->.
+apply/sumsmx_subP => i _; apply/memmx_subP=> z; rewrite sub_capmx.
+case/andP=> Riz cRiz; rewrite sub_capmx (memmx_subP (sR_R i)) //=.
+apply/cent_mxP=> A; rewrite -{1}defR; case/memmx_sumsP=> a -> R_a.
+rewrite (bigD1 i) // mulmxDl mulmxDr mulmx_suml mulmx_sumr.
+rewrite !big1 => [|j|j]; first by rewrite !addr0 (cent_mxP cRiz).
+ by rewrite eq_sym => /anhR->.
+by move/anhR->.
+Qed.
+
+End MatrixAlgebra.
+
+Arguments Scope mulsmx
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Arguments Scope left_mx_ideal
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Arguments Scope right_mx_ideal
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Arguments Scope mx_ideal
+ [_ nat_scope nat_scope nat_scope matrix_set_scope matrix_set_scope].
+Arguments Scope mxring_id
+ [_ nat_scope nat_scope ring_scope matrix_set_scope].
+Arguments Scope has_mxring_id
+ [_ nat_scope nat_scope ring_scope matrix_set_scope].
+Arguments Scope mxring
+ [_ nat_scope nat_scope matrix_set_scope].
+Arguments Scope cent_mx
+ [_ nat_scope nat_scope matrix_set_scope].
+Arguments Scope center_mx
+ [_ nat_scope nat_scope matrix_set_scope].
+
+Prenex Implicits mulsmx.
+
+Notation "A \in R" := (submx (mxvec A) R) : matrix_set_scope.
+Notation "R * S" := (mulsmx R S) : matrix_set_scope.
+Notation "''C' ( R )" := (cent_mx R) : matrix_set_scope.
+Notation "''C_' R ( S )" := (R :&: 'C(S))%MS : matrix_set_scope.
+Notation "''C_' ( R ) ( S )" := ('C_R(S))%MS (only parsing) : matrix_set_scope.
+Notation "''Z' ( R )" := (center_mx R) : matrix_set_scope.
+
+Implicit Arguments memmx_subP [F m1 m2 n R1 R2].
+Implicit Arguments memmx_eqP [F m1 m2 n R1 R2].
+Implicit Arguments memmx_addsP [F m1 m2 n R1 R2].
+Implicit Arguments memmx_sumsP [F I P n A R_].
+Implicit Arguments mulsmx_subP [F m1 m2 m n R1 R2 R].
+Implicit Arguments mulsmxP [F m1 m2 n A R1 R2].
+Implicit Arguments mxring_idP [m n R].
+Implicit Arguments cent_rowP [F m n B R].
+Implicit Arguments cent_mxP [F m n B R].
+Implicit Arguments center_mxP [F m n A R].
+
+(* Parametricity for the row-space/F-algebra theory. *)
+Section MapMatrixSpaces.
+
+Variables (aF rF : fieldType) (f : {rmorphism aF -> rF}).
+Local Notation "A ^f" := (map_mx f A) : ring_scope.
+
+Lemma Gaussian_elimination_map m n (A : 'M_(m, n)) :
+ Gaussian_elimination A^f = ((col_ebase A)^f, (row_ebase A)^f, \rank A).
+Proof.
+rewrite mxrankE /row_ebase /col_ebase unlock.
+elim: m n A => [|m IHm] [|n] A /=; rewrite ?map_mx1 //.
+set pAnz := [pred k | A k.1 k.2 != 0].
+rewrite (@eq_pick _ _ pAnz) => [|k]; last by rewrite /= mxE fmorph_eq0.
+case: {+}(pick _) => [[i j]|]; last by rewrite !map_mx1.
+rewrite mxE -fmorphV -map_xcol -map_xrow -map_dlsubmx -map_drsubmx.
+rewrite -map_ursubmx -map_mxZ -map_mxM -map_mx_sub {}IHm /=.
+case: {+}(Gaussian_elimination _) => [[L U] r] /=; rewrite map_xrow map_xcol.
+by rewrite !(@map_block_mx _ _ f 1 _ 1) !map_mx0 ?map_mx1 ?map_scalar_mx.
+Qed.
+
+Lemma mxrank_map m n (A : 'M_(m, n)) : \rank A^f = \rank A.
+Proof. by rewrite mxrankE Gaussian_elimination_map. Qed.
+
+Lemma row_free_map m n (A : 'M_(m, n)) : row_free A^f = row_free A.
+Proof. by rewrite /row_free mxrank_map. Qed.
+
+Lemma row_full_map m n (A : 'M_(m, n)) : row_full A^f = row_full A.
+Proof. by rewrite /row_full mxrank_map. Qed.
+
+Lemma map_row_ebase m n (A : 'M_(m, n)) : (row_ebase A)^f = row_ebase A^f.
+Proof. by rewrite {2}/row_ebase unlock Gaussian_elimination_map. Qed.
+
+Lemma map_col_ebase m n (A : 'M_(m, n)) : (col_ebase A)^f = col_ebase A^f.
+Proof. by rewrite {2}/col_ebase unlock Gaussian_elimination_map. Qed.
+
+Lemma map_row_base m n (A : 'M_(m, n)) :
+ (row_base A)^f = castmx (mxrank_map A, erefl n) (row_base A^f).
+Proof.
+move: (mxrank_map A); rewrite {2}/row_base mxrank_map => eqrr.
+by rewrite castmx_id map_mxM map_pid_mx map_row_ebase.
+Qed.
+
+Lemma map_col_base m n (A : 'M_(m, n)) :
+ (col_base A)^f = castmx (erefl m, mxrank_map A) (col_base A^f).
+Proof.
+move: (mxrank_map A); rewrite {2}/col_base mxrank_map => eqrr.
+by rewrite castmx_id map_mxM map_pid_mx map_col_ebase.
+Qed.
+
+Lemma map_pinvmx m n (A : 'M_(m, n)) : (pinvmx A)^f = pinvmx A^f.
+Proof.
+rewrite !map_mxM !map_invmx map_row_ebase map_col_ebase.
+by rewrite map_pid_mx -mxrank_map.
+Qed.
+
+Lemma map_kermx m n (A : 'M_(m, n)) : (kermx A)^f = kermx A^f.
+Proof.
+by rewrite !map_mxM map_invmx map_col_ebase -mxrank_map map_copid_mx.
+Qed.
+
+Lemma map_cokermx m n (A : 'M_(m, n)) : (cokermx A)^f = cokermx A^f.
+Proof.
+by rewrite !map_mxM map_invmx map_row_ebase -mxrank_map map_copid_mx.
+Qed.
+
+Lemma map_submx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A^f <= B^f)%MS = (A <= B)%MS.
+Proof. by rewrite !submxE -map_cokermx -map_mxM map_mx_eq0. Qed.
+
+Lemma map_ltmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A^f < B^f)%MS = (A < B)%MS.
+Proof. by rewrite /ltmx !map_submx. Qed.
+
+Lemma map_eqmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (A^f :=: B^f)%MS <-> (A :=: B)%MS.
+Proof.
+split=> [/eqmxP|eqAB]; first by rewrite !map_submx => /eqmxP.
+by apply/eqmxP; rewrite !map_submx !eqAB !submx_refl.
+Qed.
+
+Lemma map_genmx m n (A : 'M_(m, n)) : (<<A>>^f :=: <<A^f>>)%MS.
+Proof. by apply/eqmxP; rewrite !(genmxE, map_submx) andbb. Qed.
+
+Lemma map_addsmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (((A + B)%MS)^f :=: A^f + B^f)%MS.
+Proof.
+by apply/eqmxP; rewrite !addsmxE -map_col_mx !map_submx !addsmxE andbb.
+Qed.
+
+Lemma map_capmx_gen m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ (capmx_gen A B)^f = capmx_gen A^f B^f.
+Proof. by rewrite map_mxM map_lsubmx map_kermx map_col_mx. Qed.
+
+Lemma map_capmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ ((A :&: B)^f :=: A^f :&: B^f)%MS.
+Proof.
+by apply/eqmxP; rewrite !capmxE -map_capmx_gen !map_submx -!capmxE andbb.
+Qed.
+
+Lemma map_complmx m n (A : 'M_(m, n)) : (A^C^f = A^f^C)%MS.
+Proof. by rewrite map_mxM map_row_ebase -mxrank_map map_copid_mx. Qed.
+
+Lemma map_diffmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) :
+ ((A :\: B)^f :=: A^f :\: B^f)%MS.
+Proof.
+apply/eqmxP; rewrite !diffmxE -map_capmx_gen -map_complmx.
+by rewrite -!map_capmx !map_submx -!diffmxE andbb.
+Qed.
+
+Lemma map_eigenspace n (g : 'M_n) a : (eigenspace g a)^f = eigenspace g^f (f a).
+Proof. by rewrite map_kermx map_mx_sub ?map_scalar_mx. Qed.
+
+Lemma eigenvalue_map n (g : 'M_n) a : eigenvalue g^f (f a) = eigenvalue g a.
+Proof. by rewrite /eigenvalue -map_eigenspace map_mx_eq0. Qed.
+
+Lemma memmx_map m n A (E : 'A_(m, n)) : (A^f \in E^f)%MS = (A \in E)%MS.
+Proof. by rewrite -map_mxvec map_submx. Qed.
+
+Lemma map_mulsmx m1 m2 n (E1 : 'A_(m1, n)) (E2 : 'A_(m2, n)) :
+ ((E1 * E2)%MS^f :=: E1^f * E2^f)%MS.
+Proof.
+rewrite /mulsmx; elim/big_rec2: _ => [|i A Af _ eqA]; first by rewrite map_mx0.
+apply: (eqmx_trans (map_addsmx _ _)); apply: adds_eqmx {A Af}eqA.
+apply/eqmxP; rewrite !map_genmx !genmxE map_mxM.
+apply/rV_eqP=> u; congr (u <= _ *m _)%MS.
+by apply: map_lin_mx => //= A; rewrite map_mxM // map_vec_mx map_row.
+Qed.
+
+Lemma map_cent_mx m n (E : 'A_(m, n)) : ('C(E)%MS)^f = 'C(E^f)%MS.
+Proof.
+rewrite map_kermx //; congr (kermx _); apply: map_lin_mx => // A.
+rewrite map_mxM //; congr (_ *m _); apply: map_lin_mx => //= B.
+by rewrite map_mx_sub ? map_mxM.
+Qed.
+
+Lemma map_center_mx m n (E : 'A_(m, n)) : (('Z(E))^f :=: 'Z(E^f))%MS.
+Proof. by rewrite /center_mx -map_cent_mx; exact: map_capmx. Qed.
+
+End MapMatrixSpaces.
+
+
diff --git a/mathcomp/algebra/mxpoly.v b/mathcomp/algebra/mxpoly.v
new file mode 100644
index 0000000..7e889a0
--- /dev/null
+++ b/mathcomp/algebra/mxpoly.v
@@ -0,0 +1,1109 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div fintype tuple.
+Require Import finfun bigop fingroup perm ssralg zmodp matrix mxalgebra.
+Require Import poly polydiv.
+
+(******************************************************************************)
+(* This file provides basic support for formal computation with matrices, *)
+(* mainly results combining matrices and univariate polynomials, such as the *)
+(* Cayley-Hamilton theorem; it also contains an extension of the first order *)
+(* representation of algebra introduced in ssralg (GRing.term/formula). *)
+(* rVpoly v == the little-endian decoding of the row vector v as a *)
+(* polynomial p = \sum_i (v 0 i)%:P * 'X^i. *)
+(* poly_rV p == the partial inverse to rVpoly, for polynomials of degree *)
+(* less than d to 'rV_d (d is inferred from the context). *)
+(* Sylvester_mx p q == the Sylvester matrix of p and q. *)
+(* resultant p q == the resultant of p and q, i.e., \det (Sylvester_mx p q). *)
+(* horner_mx A == the morphism from {poly R} to 'M_n (n of the form n'.+1) *)
+(* mapping a (scalar) polynomial p to the value of its *)
+(* scalar matrix interpretation at A (this is an instance of *)
+(* the generic horner_morph construct defined in poly). *)
+(* powers_mx A d == the d x (n ^ 2) matrix whose rows are the mxvec encodings *)
+(* of the first d powers of A (n of the form n'.+1). Thus, *)
+(* vec_mx (v *m powers_mx A d) = horner_mx A (rVpoly v). *)
+(* char_poly A == the characteristic polynomial of A. *)
+(* char_poly_mx A == a matrix whose detereminant is char_poly A. *)
+(* mxminpoly A == the minimal polynomial of A, i.e., the smallest monic *)
+(* polynomial that annihilates A (A must be nontrivial). *)
+(* degree_mxminpoly A == the (positive) degree of mxminpoly A. *)
+(* mx_inv_horner A == the inverse of horner_mx A for polynomials of degree *)
+(* smaller than degree_mxminpoly A. *)
+(* integralOver RtoK u <-> u is in the integral closure of the image of R *)
+(* under RtoK : R -> K, i.e. u is a root of the image of a *)
+(* monic polynomial in R. *)
+(* algebraicOver FtoE u <-> u : E is algebraic over E; it is a root of the *)
+(* image of a nonzero polynomial under FtoE; as F must be a *)
+(* fieldType, this is equivalent to integralOver FtoE u. *)
+(* integralRange RtoK <-> the integral closure of the image of R contains *)
+(* all of K (:= forall u, integralOver RtoK u). *)
+(* This toolkit for building formal matrix expressions is packaged in the *)
+(* MatrixFormula submodule, and comprises the following: *)
+(* eval_mx e == GRing.eval lifted to matrices (:= map_mx (GRing.eval e)). *)
+(* mx_term A == GRing.Const lifted to matrices. *)
+(* mulmx_term A B == the formal product of two matrices of terms. *)
+(* mxrank_form m A == a GRing.formula asserting that the interpretation of *)
+(* the term matrix A has rank m. *)
+(* submx_form A B == a GRing.formula asserting that the row space of the *)
+(* interpretation of the term matrix A is included in the *)
+(* row space of the interpretation of B. *)
+(* seq_of_rV v == the seq corresponding to a row vector. *)
+(* row_env e == the flattening of a tensored environment e : seq 'rV_d. *)
+(* row_var F d k == the term vector of width d such that for e : seq 'rV[F]_d *)
+(* we have eval e 'X_k = eval_mx (row_env e) (row_var d k). *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GRing.Theory.
+Import Monoid.Theory.
+
+Open Local Scope ring_scope.
+
+Import Pdiv.Idomain.
+(* Row vector <-> bounded degree polynomial bijection *)
+Section RowPoly.
+
+Variables (R : ringType) (d : nat).
+Implicit Types u v : 'rV[R]_d.
+Implicit Types p q : {poly R}.
+
+Definition rVpoly v := \poly_(k < d) (if insub k is Some i then v 0 i else 0).
+Definition poly_rV p := \row_(i < d) p`_i.
+
+Lemma coef_rVpoly v k : (rVpoly v)`_k = if insub k is Some i then v 0 i else 0.
+Proof. by rewrite coef_poly; case: insubP => [i ->|]; rewrite ?if_same. Qed.
+
+Lemma coef_rVpoly_ord v (i : 'I_d) : (rVpoly v)`_i = v 0 i.
+Proof. by rewrite coef_rVpoly valK. Qed.
+
+Lemma rVpoly_delta i : rVpoly (delta_mx 0 i) = 'X^i.
+Proof.
+apply/polyP=> j; rewrite coef_rVpoly coefXn.
+case: insubP => [k _ <- | j_ge_d]; first by rewrite mxE.
+by case: eqP j_ge_d => // ->; rewrite ltn_ord.
+Qed.
+
+Lemma rVpolyK : cancel rVpoly poly_rV.
+Proof. by move=> u; apply/rowP=> i; rewrite mxE coef_rVpoly_ord. Qed.
+
+Lemma poly_rV_K p : size p <= d -> rVpoly (poly_rV p) = p.
+Proof.
+move=> le_p_d; apply/polyP=> k; rewrite coef_rVpoly.
+case: insubP => [i _ <- | ]; first by rewrite mxE.
+by rewrite -ltnNge => le_d_l; rewrite nth_default ?(leq_trans le_p_d).
+Qed.
+
+Lemma poly_rV_is_linear : linear poly_rV.
+Proof. by move=> a p q; apply/rowP=> i; rewrite !mxE coefD coefZ. Qed.
+Canonical poly_rV_additive := Additive poly_rV_is_linear.
+Canonical poly_rV_linear := Linear poly_rV_is_linear.
+
+Lemma rVpoly_is_linear : linear rVpoly.
+Proof.
+move=> a u v; apply/polyP=> k; rewrite coefD coefZ !coef_rVpoly.
+by case: insubP => [i _ _ | _]; rewrite ?mxE // mulr0 addr0.
+Qed.
+Canonical rVpoly_additive := Additive rVpoly_is_linear.
+Canonical rVpoly_linear := Linear rVpoly_is_linear.
+
+End RowPoly.
+
+Implicit Arguments poly_rV [R d].
+Prenex Implicits rVpoly poly_rV.
+
+Section Resultant.
+
+Variables (R : ringType) (p q : {poly R}).
+
+Let dS := ((size q).-1 + (size p).-1)%N.
+Local Notation band r := (lin1_mx (poly_rV \o r \o* rVpoly)).
+
+Definition Sylvester_mx : 'M[R]_dS := col_mx (band p) (band q).
+
+Lemma Sylvester_mxE (i j : 'I_dS) :
+ let S_ r k := r`_(j - k) *+ (k <= j) in
+ Sylvester_mx i j = match split i with inl k => S_ p k | inr k => S_ q k end.
+Proof.
+move=> S_; rewrite mxE; case: {i}(split i) => i; rewrite !mxE /=;
+ by rewrite rVpoly_delta coefXnM ltnNge if_neg -mulrb.
+Qed.
+
+Definition resultant := \det Sylvester_mx.
+
+End Resultant.
+
+Lemma resultant_in_ideal (R : comRingType) (p q : {poly R}) :
+ size p > 1 -> size q > 1 ->
+ {uv : {poly R} * {poly R} | size uv.1 < size q /\ size uv.2 < size p
+ & (resultant p q)%:P = uv.1 * p + uv.2 * q}.
+Proof.
+move=> p_nc q_nc; pose dp := (size p).-1; pose dq := (size q).-1.
+pose S := Sylvester_mx p q; pose dS := (dq + dp)%N.
+have dS_gt0: dS > 0 by rewrite /dS /dq -(subnKC q_nc).
+pose j0 := Ordinal dS_gt0.
+pose Ss0 := col_mx (p *: \col_(i < dq) 'X^i) (q *: \col_(i < dp) 'X^i).
+pose Ss := \matrix_(i, j) (if j == j0 then Ss0 i 0 else (S i j)%:P).
+pose u ds s := \sum_(i < ds) cofactor Ss (s i) j0 * 'X^i.
+exists (u _ (lshift dp), u _ ((rshift dq) _)).
+ suffices sz_u ds s: ds > 1 -> size (u ds.-1 s) < ds by rewrite !sz_u.
+ move/ltn_predK=> {2}<-; apply: leq_trans (size_sum _ _ _) _.
+ apply/bigmax_leqP=> i _.
+ have ->: cofactor Ss (s i) j0 = (cofactor S (s i) j0)%:P.
+ rewrite rmorphM rmorph_sign -det_map_mx; congr (_ * \det _).
+ by apply/matrixP=> i' j'; rewrite !mxE.
+ apply: leq_trans (size_mul_leq _ _) (leq_trans _ (valP i)).
+ by rewrite size_polyC size_polyXn addnS /= -add1n leq_add2r leq_b1.
+transitivity (\det Ss); last first.
+ rewrite (expand_det_col Ss j0) big_split_ord !big_distrl /=.
+ by congr (_ + _); apply: eq_bigr => i _;
+ rewrite mxE eqxx (col_mxEu, col_mxEd) !mxE mulrC mulrA mulrAC.
+pose S_ j1 := map_mx polyC (\matrix_(i, j) S i (if j == j0 then j1 else j)).
+pose Ss0_ i dj := \poly_(j < dj) S i (insubd j0 j).
+pose Ss_ dj := \matrix_(i, j) (if j == j0 then Ss0_ i dj else (S i j)%:P).
+have{Ss u} ->: Ss = Ss_ dS.
+ apply/matrixP=> i j; rewrite mxE [in X in _ = X]mxE; case: (j == j0) => {j}//.
+ apply/polyP=> k; rewrite coef_poly Sylvester_mxE mxE.
+ have [k_ge_dS | k_lt_dS] := leqP dS k.
+ case: (split i) => {i}i; rewrite !mxE coefMXn;
+ case: ifP => // /negbT; rewrite -ltnNge ltnS => hi.
+ apply: (leq_sizeP _ _ (leqnn (size p))); rewrite -(ltn_predK p_nc).
+ by rewrite ltn_subRL (leq_trans _ k_ge_dS) // ltn_add2r.
+ - apply: (leq_sizeP _ _ (leqnn (size q))); rewrite -(ltn_predK q_nc).
+ by rewrite ltn_subRL (leq_trans _ k_ge_dS) // addnC ltn_add2l.
+ by rewrite insubdK //; case: (split i) => {i}i;
+ rewrite !mxE coefMXn; case: leqP.
+elim: {-2}dS (leqnn dS) (dS_gt0) => // dj IHj dj_lt_dS _.
+pose j1 := Ordinal dj_lt_dS; pose rj0T (A : 'M[{poly R}]_dS) := row j0 A^T.
+have: rj0T (Ss_ dj.+1) = 'X^dj *: rj0T (S_ j1) + 1 *: rj0T (Ss_ dj).
+ apply/rowP=> i; apply/polyP=> k; rewrite scale1r !(Sylvester_mxE, mxE) eqxx.
+ rewrite coefD coefXnM coefC !coef_poly ltnS subn_eq0 ltn_neqAle andbC.
+ case: (leqP k dj) => [k_le_dj | k_gt_dj] /=; last by rewrite addr0.
+ rewrite Sylvester_mxE insubdK; last exact: leq_ltn_trans (dj_lt_dS).
+ by case: eqP => [-> | _]; rewrite (addr0, add0r).
+rewrite -det_tr => /determinant_multilinear->;
+ try by apply/matrixP=> i j; rewrite !mxE eq_sym (negPf (neq_lift _ _)).
+have [dj0 | dj_gt0] := posnP dj; rewrite ?dj0 !mul1r.
+ rewrite !det_tr det_map_mx addrC (expand_det_col _ j0) big1 => [|i _].
+ rewrite add0r; congr (\det _)%:P.
+ apply/matrixP=> i j; rewrite [in X in _ = X]mxE; case: eqP => // ->.
+ by congr (S i _); apply: val_inj.
+ by rewrite mxE /= [Ss0_ _ _]poly_def big_ord0 mul0r.
+have /determinant_alternate->: j1 != j0 by rewrite -val_eqE -lt0n.
+ by rewrite mulr0 add0r det_tr IHj // ltnW.
+by move=> i; rewrite !mxE if_same.
+Qed.
+
+Lemma resultant_eq0 (R : idomainType) (p q : {poly R}) :
+ (resultant p q == 0) = (size (gcdp p q) > 1).
+Proof.
+have dvdpp := dvdpp; set r := gcdp p q.
+pose dp := (size p).-1; pose dq := (size q).-1.
+have /andP[r_p r_q]: (r %| p) && (r %| q) by rewrite -dvdp_gcd.
+apply/det0P/idP=> [[uv nz_uv] | r_nonC].
+ have [p0 _ | p_nz] := eqVneq p 0.
+ have: dq + dp > 0.
+ rewrite lt0n; apply: contraNneq nz_uv => dqp0.
+ by rewrite dqp0 in uv *; rewrite [uv]thinmx0.
+ by rewrite /dp /dq /r p0 size_poly0 addn0 gcd0p -subn1 subn_gt0.
+ do [rewrite -[uv]hsubmxK -{1}row_mx0 mul_row_col !mul_rV_lin1 /=] in nz_uv *.
+ set u := rVpoly _; set v := rVpoly _; pose m := gcdp (v * p) (v * q).
+ have lt_vp: size v < size p by rewrite (polySpred p_nz) ltnS size_poly.
+ move/(congr1 rVpoly)/eqP; rewrite -linearD linear0 poly_rV_K; last first.
+ rewrite (leq_trans (size_add _ _)) // geq_max.
+ rewrite !(leq_trans (size_mul_leq _ _)) // -subn1 leq_subLR.
+ by rewrite addnC addnA leq_add ?leqSpred ?size_poly.
+ by rewrite addnCA leq_add ?leqSpred ?size_poly.
+ rewrite addrC addr_eq0 => /eqP vq_up.
+ have nz_v: v != 0.
+ apply: contraNneq nz_uv => v0; apply/eqP.
+ congr row_mx; apply: (can_inj (@rVpolyK _ _)); rewrite linear0 // -/u.
+ by apply: contra_eq vq_up; rewrite v0 mul0r -addr_eq0 add0r => /mulf_neq0->.
+ have r_nz: r != 0 := dvdpN0 r_p p_nz.
+ have /dvdpP [[c w] /= nz_c wv]: v %| m by rewrite dvdp_gcd !dvdp_mulr.
+ have m_wd d: m %| v * d -> w %| d.
+ case/dvdpP=> [[k f]] /= nz_k /(congr1 ( *:%R c)).
+ rewrite mulrC scalerA scalerAl scalerAr wv mulrA => /(mulIf nz_v)def_fw.
+ by apply/dvdpP; exists (c * k, f); rewrite //= mulf_neq0.
+ have w_r: w %| r by rewrite dvdp_gcd !m_wd ?dvdp_gcdl ?dvdp_gcdr.
+ have w_nz: w != 0 := dvdpN0 w_r r_nz.
+ have p_m: p %| m by rewrite dvdp_gcd vq_up -mulNr !dvdp_mull.
+ rewrite (leq_trans _ (dvdp_leq r_nz w_r)) // -(ltn_add2l (size v)).
+ rewrite addnC -ltn_subRL subn1 -size_mul // mulrC -wv size_scale //.
+ rewrite (leq_trans lt_vp) // dvdp_leq // -size_poly_eq0.
+ by rewrite -(size_scale _ nz_c) size_poly_eq0 wv mulf_neq0.
+have [[c p'] /= nz_c p'r] := dvdpP _ _ r_p.
+have [[k q'] /= nz_k q'r] := dvdpP _ _ r_q.
+have def_r := subnKC r_nonC; have r_nz: r != 0 by rewrite -size_poly_eq0 -def_r.
+have le_p'_dp: size p' <= dp.
+ have [-> | nz_p'] := eqVneq p' 0; first by rewrite size_poly0.
+ by rewrite /dp -(size_scale p nz_c) p'r size_mul // addnC -def_r leq_addl.
+have le_q'_dq: size q' <= dq.
+ have [-> | nz_q'] := eqVneq q' 0; first by rewrite size_poly0.
+ by rewrite /dq -(size_scale q nz_k) q'r size_mul // addnC -def_r leq_addl.
+exists (row_mx (- c *: poly_rV q') (k *: poly_rV p')).
+ apply: contraNneq r_nz; rewrite -row_mx0; case/eq_row_mx=> q0 p0.
+ have{p0} p0: p = 0.
+ apply/eqP; rewrite -size_poly_eq0 -(size_scale p nz_c) p'r.
+ rewrite -(size_scale _ nz_k) scalerAl -(poly_rV_K le_p'_dp) -linearZ p0.
+ by rewrite linear0 mul0r size_poly0.
+ rewrite /r p0 gcd0p -size_poly_eq0 -(size_scale q nz_k) q'r.
+ rewrite -(size_scale _ nz_c) scalerAl -(poly_rV_K le_q'_dq) -linearZ.
+ by rewrite -[c]opprK scaleNr q0 !linear0 mul0r size_poly0.
+rewrite mul_row_col scaleNr mulNmx !mul_rV_lin1 /= !linearZ /= !poly_rV_K //.
+by rewrite !scalerCA p'r q'r mulrCA addNr.
+Qed.
+
+Section HornerMx.
+
+Variables (R : comRingType) (n' : nat).
+Local Notation n := n'.+1.
+Variable A : 'M[R]_n.
+Implicit Types p q : {poly R}.
+
+Definition horner_mx := horner_morph (fun a => scalar_mx_comm a A).
+Canonical horner_mx_additive := [additive of horner_mx].
+Canonical horner_mx_rmorphism := [rmorphism of horner_mx].
+
+Lemma horner_mx_C a : horner_mx a%:P = a%:M.
+Proof. exact: horner_morphC. Qed.
+
+Lemma horner_mx_X : horner_mx 'X = A. Proof. exact: horner_morphX. Qed.
+
+Lemma horner_mxZ : scalable horner_mx.
+Proof.
+move=> a p /=; rewrite -mul_polyC rmorphM /=.
+by rewrite horner_mx_C [_ * _]mul_scalar_mx.
+Qed.
+
+Canonical horner_mx_linear := AddLinear horner_mxZ.
+Canonical horner_mx_lrmorphism := [lrmorphism of horner_mx].
+
+Definition powers_mx d := \matrix_(i < d) mxvec (A ^+ i).
+
+Lemma horner_rVpoly m (u : 'rV_m) :
+ horner_mx (rVpoly u) = vec_mx (u *m powers_mx m).
+Proof.
+rewrite mulmx_sum_row linear_sum [rVpoly u]poly_def rmorph_sum.
+apply: eq_bigr => i _.
+by rewrite valK !linearZ rmorphX /= horner_mx_X rowK /= mxvecK.
+Qed.
+
+End HornerMx.
+
+Section CharPoly.
+
+Variables (R : ringType) (n : nat) (A : 'M[R]_n).
+Implicit Types p q : {poly R}.
+
+Definition char_poly_mx := 'X%:M - map_mx (@polyC R) A.
+Definition char_poly := \det char_poly_mx.
+
+Let diagA := [seq A i i | i : 'I_n].
+Let size_diagA : size diagA = n.
+Proof. by rewrite size_image card_ord. Qed.
+
+Let split_diagA :
+ exists2 q, \prod_(x <- diagA) ('X - x%:P) + q = char_poly & size q <= n.-1.
+Proof.
+rewrite [char_poly](bigD1 1%g) //=; set q := \sum_(s | _) _; exists q.
+ congr (_ + _); rewrite odd_perm1 mul1r big_map enumT; apply: eq_bigr => i _.
+ by rewrite !mxE perm1 eqxx.
+apply: leq_trans {q}(size_sum _ _ _) _; apply/bigmax_leqP=> s nt_s.
+have{nt_s} [i nfix_i]: exists i, s i != i.
+ apply/existsP; rewrite -negb_forall; apply: contra nt_s => s_1.
+ by apply/eqP; apply/permP=> i; apply/eqP; rewrite perm1 (forallP s_1).
+apply: leq_trans (_ : #|[pred j | s j == j]|.+1 <= n.-1).
+ rewrite -sum1_card (@big_mkcond nat) /= size_Msign.
+ apply: (big_ind2 (fun p m => size p <= m.+1)) => [| p mp q mq IHp IHq | j _].
+ - by rewrite size_poly1.
+ - apply: leq_trans (size_mul_leq _ _) _.
+ by rewrite -subn1 -addnS leq_subLR addnA leq_add.
+ rewrite !mxE eq_sym !inE; case: (s j == j); first by rewrite polyseqXsubC.
+ by rewrite sub0r size_opp size_polyC leq_b1.
+rewrite -{8}[n]card_ord -(cardC (pred2 (s i) i)) card2 nfix_i !ltnS.
+apply: subset_leq_card; apply/subsetP=> j; move/(_ =P j)=> fix_j.
+rewrite !inE -{1}fix_j (inj_eq (@perm_inj _ s)) orbb.
+by apply: contraNneq nfix_i => <-; rewrite fix_j.
+Qed.
+
+Lemma size_char_poly : size char_poly = n.+1.
+Proof.
+have [q <- lt_q_n] := split_diagA; have le_q_n := leq_trans lt_q_n (leq_pred n).
+by rewrite size_addl size_prod_XsubC size_diagA.
+Qed.
+
+Lemma char_poly_monic : char_poly \is monic.
+Proof.
+rewrite monicE -(monicP (monic_prod_XsubC diagA xpredT id)).
+rewrite !lead_coefE size_char_poly.
+have [q <- lt_q_n] := split_diagA; have le_q_n := leq_trans lt_q_n (leq_pred n).
+by rewrite size_prod_XsubC size_diagA coefD (nth_default 0 le_q_n) addr0.
+Qed.
+
+Lemma char_poly_trace : n > 0 -> char_poly`_n.-1 = - \tr A.
+Proof.
+move=> n_gt0; have [q <- lt_q_n] := split_diagA; set p := \prod_(x <- _) _.
+rewrite coefD {q lt_q_n}(nth_default 0 lt_q_n) addr0.
+have{n_gt0} ->: p`_n.-1 = ('X * p)`_n by rewrite coefXM eqn0Ngt n_gt0.
+have ->: \tr A = \sum_(x <- diagA) x by rewrite big_map enumT.
+rewrite -size_diagA {}/p; elim: diagA => [|x d IHd].
+ by rewrite !big_nil mulr1 coefX oppr0.
+rewrite !big_cons coefXM mulrBl coefB IHd opprD addrC; congr (- _ + _).
+rewrite mul_polyC coefZ [size _]/= -(size_prod_XsubC _ id) -lead_coefE.
+by rewrite (monicP _) ?monic_prod_XsubC ?mulr1.
+Qed.
+
+Lemma char_poly_det : char_poly`_0 = (- 1) ^+ n * \det A.
+Proof.
+rewrite big_distrr coef_sum [0%N]lock /=; apply: eq_bigr => s _.
+rewrite -{1}rmorphN -rmorphX mul_polyC coefZ /=.
+rewrite mulrA -exprD addnC exprD -mulrA -lock; congr (_ * _).
+transitivity (\prod_(i < n) - A i (s i)); last by rewrite prodrN card_ord.
+elim: (index_enum _) => [|i e IHe]; rewrite !(big_nil, big_cons) ?coef1 //.
+by rewrite coefM big_ord1 IHe !mxE coefB coefC coefMn coefX mul0rn sub0r.
+Qed.
+
+End CharPoly.
+
+Lemma mx_poly_ring_isom (R : ringType) n' (n := n'.+1) :
+ exists phi : {rmorphism 'M[{poly R}]_n -> {poly 'M[R]_n}},
+ [/\ bijective phi,
+ forall p, phi p%:M = map_poly scalar_mx p,
+ forall A, phi (map_mx polyC A) = A%:P
+ & forall A i j k, (phi A)`_k i j = (A i j)`_k].
+Proof.
+set M_RX := 'M[{poly R}]_n; set MR_X := ({poly 'M[R]_n}).
+pose Msize (A : M_RX) := \max_i \max_j size (A i j).
+pose phi (A : M_RX) := \poly_(k < Msize A) \matrix_(i, j) (A i j)`_k.
+have coef_phi A i j k: (phi A)`_k i j = (A i j)`_k.
+ rewrite coef_poly; case: (ltnP k _) => le_m_k; rewrite mxE // nth_default //.
+ apply: leq_trans (leq_trans (leq_bigmax i) le_m_k); exact: (leq_bigmax j).
+have phi_is_rmorphism : rmorphism phi.
+ do 2?[split=> [A B|]]; apply/polyP=> k; apply/matrixP=> i j; last 1 first.
+ - rewrite coef_phi mxE coefMn !coefC.
+ by case: (k == _); rewrite ?mxE ?mul0rn.
+ - by rewrite !(coef_phi, mxE, coefD, coefN).
+ rewrite !coef_phi !mxE !coefM summxE coef_sum.
+ pose F k1 k2 := (A i k1)`_k2 * (B k1 j)`_(k - k2).
+ transitivity (\sum_k1 \sum_(k2 < k.+1) F k1 k2); rewrite {}/F.
+ by apply: eq_bigr=> k1 _; rewrite coefM.
+ rewrite exchange_big /=; apply: eq_bigr => k2 _.
+ by rewrite mxE; apply: eq_bigr => k1 _; rewrite !coef_phi.
+have bij_phi: bijective phi.
+ exists (fun P : MR_X => \matrix_(i, j) \poly_(k < size P) P`_k i j) => [A|P].
+ apply/matrixP=> i j; rewrite mxE; apply/polyP=> k.
+ rewrite coef_poly -coef_phi.
+ by case: leqP => // P_le_k; rewrite nth_default ?mxE.
+ apply/polyP=> k; apply/matrixP=> i j; rewrite coef_phi mxE coef_poly.
+ by case: leqP => // P_le_k; rewrite nth_default ?mxE.
+exists (RMorphism phi_is_rmorphism).
+split=> // [p | A]; apply/polyP=> k; apply/matrixP=> i j.
+ by rewrite coef_phi coef_map !mxE coefMn.
+by rewrite coef_phi !mxE !coefC; case k; last rewrite /= mxE.
+Qed.
+
+Theorem Cayley_Hamilton (R : comRingType) n' (A : 'M[R]_n'.+1) :
+ horner_mx A (char_poly A) = 0.
+Proof.
+have [phi [_ phiZ phiC _]] := mx_poly_ring_isom R n'.
+apply/rootP/factor_theorem; rewrite -phiZ -mul_adj_mx rmorphM.
+by move: (phi _) => q; exists q; rewrite rmorphB phiC phiZ map_polyX.
+Qed.
+
+Lemma eigenvalue_root_char (F : fieldType) n (A : 'M[F]_n) a :
+ eigenvalue A a = root (char_poly A) a.
+Proof.
+transitivity (\det (a%:M - A) == 0).
+ apply/eigenvalueP/det0P=> [[v Av_av v_nz] | [v v_nz Av_av]]; exists v => //.
+ by rewrite mulmxBr Av_av mul_mx_scalar subrr.
+ by apply/eqP; rewrite -mul_mx_scalar eq_sym -subr_eq0 -mulmxBr Av_av.
+congr (_ == 0); rewrite horner_sum; apply: eq_bigr => s _.
+rewrite hornerM horner_exp !hornerE; congr (_ * _).
+rewrite (big_morph _ (fun p q => hornerM p q a) (hornerC 1 a)).
+by apply: eq_bigr => i _; rewrite !mxE !(hornerE, hornerMn).
+Qed.
+
+Section MinPoly.
+
+Variables (F : fieldType) (n' : nat).
+Local Notation n := n'.+1.
+Variable A : 'M[F]_n.
+Implicit Types p q : {poly F}.
+
+Fact degree_mxminpoly_proof : exists d, \rank (powers_mx A d.+1) <= d.
+Proof. by exists (n ^ 2)%N; rewrite rank_leq_col. Qed.
+Definition degree_mxminpoly := ex_minn degree_mxminpoly_proof.
+Local Notation d := degree_mxminpoly.
+Local Notation Ad := (powers_mx A d).
+
+Lemma mxminpoly_nonconstant : d > 0.
+Proof.
+rewrite /d; case: ex_minnP; case=> //; rewrite leqn0 mxrank_eq0; move/eqP.
+move/row_matrixP; move/(_ 0); move/eqP; rewrite rowK row0 mxvec_eq0.
+by rewrite -mxrank_eq0 mxrank1.
+Qed.
+
+Lemma minpoly_mx1 : (1%:M \in Ad)%MS.
+Proof.
+by apply: (eq_row_sub (Ordinal mxminpoly_nonconstant)); rewrite rowK.
+Qed.
+
+Lemma minpoly_mx_free : row_free Ad.
+Proof.
+have:= mxminpoly_nonconstant; rewrite /d; case: ex_minnP; case=> // d' _.
+move/(_ d'); move/implyP; rewrite ltnn implybF -ltnS ltn_neqAle.
+by rewrite rank_leq_row andbT negbK.
+Qed.
+
+Lemma horner_mx_mem p : (horner_mx A p \in Ad)%MS.
+Proof.
+elim/poly_ind: p => [|p a IHp]; first by rewrite rmorph0 // linear0 sub0mx.
+rewrite rmorphD rmorphM /= horner_mx_C horner_mx_X.
+rewrite addrC -scalemx1 linearP /= -(mul_vec_lin (mulmxr_linear _ A)).
+case/submxP: IHp => u ->{p}.
+have: (powers_mx A (1 + d) <= Ad)%MS.
+ rewrite -(geq_leqif (mxrank_leqif_sup _)).
+ by rewrite (eqnP minpoly_mx_free) /d; case: ex_minnP.
+ rewrite addnC; apply/row_subP=> i.
+ by apply: eq_row_sub (lshift 1 i) _; rewrite !rowK.
+apply: submx_trans; rewrite addmx_sub ?scalemx_sub //.
+ by apply: (eq_row_sub 0); rewrite rowK.
+rewrite -mulmxA mulmx_sub {u}//; apply/row_subP=> i.
+rewrite row_mul rowK mul_vec_lin /= mulmxE -exprSr.
+by apply: (eq_row_sub (rshift 1 i)); rewrite rowK.
+Qed.
+
+Definition mx_inv_horner B := rVpoly (mxvec B *m pinvmx Ad).
+
+Lemma mx_inv_horner0 : mx_inv_horner 0 = 0.
+Proof. by rewrite /mx_inv_horner !(linear0, mul0mx). Qed.
+
+Lemma mx_inv_hornerK B : (B \in Ad)%MS -> horner_mx A (mx_inv_horner B) = B.
+Proof. by move=> sBAd; rewrite horner_rVpoly mulmxKpV ?mxvecK. Qed.
+
+Lemma minpoly_mxM B C : (B \in Ad -> C \in Ad -> B * C \in Ad)%MS.
+Proof.
+move=> AdB AdC; rewrite -(mx_inv_hornerK AdB) -(mx_inv_hornerK AdC).
+by rewrite -rmorphM ?horner_mx_mem.
+Qed.
+
+Lemma minpoly_mx_ring : mxring Ad.
+Proof.
+apply/andP; split; first by apply/mulsmx_subP; exact: minpoly_mxM.
+apply/mxring_idP; exists 1%:M; split=> *; rewrite ?mulmx1 ?mul1mx //.
+ by rewrite -mxrank_eq0 mxrank1.
+exact: minpoly_mx1.
+Qed.
+
+Definition mxminpoly := 'X^d - mx_inv_horner (A ^+ d).
+Local Notation p_A := mxminpoly.
+
+Lemma size_mxminpoly : size p_A = d.+1.
+Proof. by rewrite size_addl ?size_polyXn // size_opp ltnS size_poly. Qed.
+
+Lemma mxminpoly_monic : p_A \is monic.
+Proof.
+rewrite monicE /lead_coef size_mxminpoly coefB coefXn eqxx /=.
+by rewrite nth_default ?size_poly // subr0.
+Qed.
+
+Lemma size_mod_mxminpoly p : size (p %% p_A) <= d.
+Proof.
+by rewrite -ltnS -size_mxminpoly ltn_modp // -size_poly_eq0 size_mxminpoly.
+Qed.
+
+Lemma mx_root_minpoly : horner_mx A p_A = 0.
+Proof.
+rewrite rmorphB -{3}(horner_mx_X A) -rmorphX /=.
+by rewrite mx_inv_hornerK ?subrr ?horner_mx_mem.
+Qed.
+
+Lemma horner_rVpolyK (u : 'rV_d) :
+ mx_inv_horner (horner_mx A (rVpoly u)) = rVpoly u.
+Proof.
+congr rVpoly; rewrite horner_rVpoly vec_mxK.
+by apply: (row_free_inj minpoly_mx_free); rewrite mulmxKpV ?submxMl.
+Qed.
+
+Lemma horner_mxK p : mx_inv_horner (horner_mx A p) = p %% p_A.
+Proof.
+rewrite {1}(Pdiv.IdomainMonic.divp_eq mxminpoly_monic p) rmorphD rmorphM /=.
+rewrite mx_root_minpoly mulr0 add0r.
+by rewrite -(poly_rV_K (size_mod_mxminpoly _)) horner_rVpolyK.
+Qed.
+
+Lemma mxminpoly_min p : horner_mx A p = 0 -> p_A %| p.
+Proof. by move=> pA0; rewrite /dvdp -horner_mxK pA0 mx_inv_horner0. Qed.
+
+Lemma horner_rVpoly_inj : @injective 'M_n 'rV_d (horner_mx A \o rVpoly).
+Proof.
+apply: can_inj (poly_rV \o mx_inv_horner) _ => u.
+by rewrite /= horner_rVpolyK rVpolyK.
+Qed.
+
+Lemma mxminpoly_linear_is_scalar : (d <= 1) = is_scalar_mx A.
+Proof.
+have scalP := has_non_scalar_mxP minpoly_mx1.
+rewrite leqNgt -(eqnP minpoly_mx_free); apply/scalP/idP=> [|[[B]]].
+ case scalA: (is_scalar_mx A); [by right | left].
+ by exists A; rewrite ?scalA // -{1}(horner_mx_X A) horner_mx_mem.
+move/mx_inv_hornerK=> <- nsB; case/is_scalar_mxP=> a defA; case/negP: nsB.
+move: {B}(_ B); apply: poly_ind => [|p c].
+ by rewrite rmorph0 ?mx0_is_scalar.
+rewrite rmorphD ?rmorphM /= horner_mx_X defA; case/is_scalar_mxP=> b ->.
+by rewrite -rmorphM horner_mx_C -rmorphD /= scalar_mx_is_scalar.
+Qed.
+
+Lemma mxminpoly_dvd_char : p_A %| char_poly A.
+Proof. by apply: mxminpoly_min; exact: Cayley_Hamilton. Qed.
+
+Lemma eigenvalue_root_min a : eigenvalue A a = root p_A a.
+Proof.
+apply/idP/idP=> Aa; last first.
+ rewrite eigenvalue_root_char !root_factor_theorem in Aa *.
+ exact: dvdp_trans Aa mxminpoly_dvd_char.
+have{Aa} [v Av_av v_nz] := eigenvalueP Aa.
+apply: contraR v_nz => pa_nz; rewrite -{pa_nz}(eqmx_eq0 (eqmx_scale _ pa_nz)).
+apply/eqP; rewrite -(mulmx0 _ v) -mx_root_minpoly.
+elim/poly_ind: p_A => [|p c IHp].
+ by rewrite rmorph0 horner0 scale0r mulmx0.
+rewrite !hornerE rmorphD rmorphM /= horner_mx_X horner_mx_C scalerDl.
+by rewrite -scalerA mulmxDr mul_mx_scalar mulmxA -IHp -scalemxAl Av_av.
+Qed.
+
+End MinPoly.
+
+(* Parametricity. *)
+Section MapRingMatrix.
+
+Variables (aR rR : ringType) (f : {rmorphism aR -> rR}).
+Local Notation "A ^f" := (map_mx (GRing.RMorphism.apply f) A) : ring_scope.
+Local Notation fp := (map_poly (GRing.RMorphism.apply f)).
+Variables (d n : nat) (A : 'M[aR]_n).
+
+Lemma map_rVpoly (u : 'rV_d) : fp (rVpoly u) = rVpoly u^f.
+Proof.
+apply/polyP=> k; rewrite coef_map !coef_rVpoly.
+by case: (insub k) => [i|]; rewrite /= ?rmorph0 // mxE.
+Qed.
+
+Lemma map_poly_rV p : (poly_rV p)^f = poly_rV (fp p) :> 'rV_d.
+Proof. by apply/rowP=> j; rewrite !mxE coef_map. Qed.
+
+Lemma map_char_poly_mx : map_mx fp (char_poly_mx A) = char_poly_mx A^f.
+Proof.
+rewrite raddfB /= map_scalar_mx /= map_polyX; congr (_ - _).
+by apply/matrixP=> i j; rewrite !mxE map_polyC.
+Qed.
+
+Lemma map_char_poly : fp (char_poly A) = char_poly A^f.
+Proof. by rewrite -det_map_mx map_char_poly_mx. Qed.
+
+End MapRingMatrix.
+
+Section MapResultant.
+
+Lemma map_resultant (aR rR : ringType) (f : {rmorphism {poly aR} -> rR}) p q :
+ f (lead_coef p) != 0 -> f (lead_coef q) != 0 ->
+ f (resultant p q)= resultant (map_poly f p) (map_poly f q).
+Proof.
+move=> nz_fp nz_fq; rewrite /resultant /Sylvester_mx !size_map_poly_id0 //.
+rewrite -det_map_mx /= map_col_mx; congr (\det (col_mx _ _));
+ by apply: map_lin1_mx => v; rewrite map_poly_rV rmorphM /= map_rVpoly.
+Qed.
+
+End MapResultant.
+
+Section MapComRing.
+
+Variables (aR rR : comRingType) (f : {rmorphism aR -> rR}).
+Local Notation "A ^f" := (map_mx f A) : ring_scope.
+Local Notation fp := (map_poly f).
+Variables (n' : nat) (A : 'M[aR]_n'.+1).
+
+Lemma map_powers_mx e : (powers_mx A e)^f = powers_mx A^f e.
+Proof. by apply/row_matrixP=> i; rewrite -map_row !rowK map_mxvec rmorphX. Qed.
+
+Lemma map_horner_mx p : (horner_mx A p)^f = horner_mx A^f (fp p).
+Proof.
+rewrite -[p](poly_rV_K (leqnn _)) map_rVpoly.
+by rewrite !horner_rVpoly map_vec_mx map_mxM map_powers_mx.
+Qed.
+
+End MapComRing.
+
+Section MapField.
+
+Variables (aF rF : fieldType) (f : {rmorphism aF -> rF}).
+Local Notation "A ^f" := (map_mx f A) : ring_scope.
+Local Notation fp := (map_poly f).
+Variables (n' : nat) (A : 'M[aF]_n'.+1).
+
+Lemma degree_mxminpoly_map : degree_mxminpoly A^f = degree_mxminpoly A.
+Proof. by apply: eq_ex_minn => e; rewrite -map_powers_mx mxrank_map. Qed.
+
+Lemma mxminpoly_map : mxminpoly A^f = fp (mxminpoly A).
+Proof.
+rewrite rmorphB; congr (_ - _).
+ by rewrite /= map_polyXn degree_mxminpoly_map.
+rewrite degree_mxminpoly_map -rmorphX /=.
+apply/polyP=> i; rewrite coef_map //= !coef_rVpoly degree_mxminpoly_map.
+case/insub: i => [i|]; last by rewrite rmorph0.
+by rewrite -map_powers_mx -map_pinvmx // -map_mxvec -map_mxM // mxE.
+Qed.
+
+Lemma map_mx_inv_horner u : fp (mx_inv_horner A u) = mx_inv_horner A^f u^f.
+Proof.
+rewrite map_rVpoly map_mxM map_mxvec map_pinvmx map_powers_mx.
+by rewrite /mx_inv_horner degree_mxminpoly_map.
+Qed.
+
+End MapField.
+
+Section IntegralOverRing.
+
+Definition integralOver (R K : ringType) (RtoK : R -> K) (z : K) :=
+ exists2 p, p \is monic & root (map_poly RtoK p) z.
+
+Definition integralRange R K RtoK := forall z, @integralOver R K RtoK z.
+
+Variables (B R K : ringType) (BtoR : B -> R) (RtoK : {rmorphism R -> K}).
+
+Lemma integral_rmorph x :
+ integralOver BtoR x -> integralOver (RtoK \o BtoR) (RtoK x).
+Proof. by case=> p; exists p; rewrite // map_poly_comp rmorph_root. Qed.
+
+Lemma integral_id x : integralOver RtoK (RtoK x).
+Proof. by exists ('X - x%:P); rewrite ?monicXsubC ?rmorph_root ?root_XsubC. Qed.
+
+Lemma integral_nat n : integralOver RtoK n%:R.
+Proof. by rewrite -(rmorph_nat RtoK); apply: integral_id. Qed.
+
+Lemma integral0 : integralOver RtoK 0. Proof. exact: (integral_nat 0). Qed.
+
+Lemma integral1 : integralOver RtoK 1. Proof. exact: (integral_nat 1). Qed.
+
+Lemma integral_poly (p : {poly K}) :
+ (forall i, integralOver RtoK p`_i) <-> {in p : seq K, integralRange RtoK}.
+Proof.
+split=> intRp => [_ /(nthP 0)[i _ <-] // | i]; rewrite -[p]coefK coef_poly.
+by case: ifP => [ltip | _]; [apply/intRp/mem_nth | apply: integral0].
+Qed.
+
+End IntegralOverRing.
+
+Section IntegralOverComRing.
+
+Variables (R K : comRingType) (RtoK : {rmorphism R -> K}).
+
+Lemma integral_horner_root w (p q : {poly K}) :
+ p \is monic -> root p w ->
+ {in p : seq K, integralRange RtoK} -> {in q : seq K, integralRange RtoK} ->
+ integralOver RtoK q.[w].
+Proof.
+move=> mon_p pw0 intRp intRq.
+pose memR y := exists x, y = RtoK x.
+have memRid x: memR (RtoK x) by exists x.
+have memR_nat n: memR n%:R by rewrite -(rmorph_nat RtoK).
+have [memR0 memR1]: memR 0 * memR 1 := (memR_nat 0%N, memR_nat 1%N).
+have memRN1: memR (- 1) by exists (- 1); rewrite rmorphN1.
+pose rVin (E : K -> Prop) n (a : 'rV[K]_n) := forall i, E (a 0 i).
+pose pXin (E : K -> Prop) (r : {poly K}) := forall i, E r`_i.
+pose memM E n (X : 'rV_n) y := exists a, rVin E n a /\ y = (a *m X^T) 0 0.
+pose finM E S := exists n, exists X, forall y, memM E n X y <-> S y.
+have tensorM E n1 n2 X Y: finM E (memM (memM E n2 Y) n1 X).
+ exists (n1 * n2)%N, (mxvec (X^T *m Y)) => y.
+ split=> [[a [Ea Dy]] | [a1 [/fin_all_exists[a /all_and2[Ea Da1]] ->]]].
+ exists (Y *m (vec_mx a)^T); split=> [i|].
+ exists (row i (vec_mx a)); split=> [j|]; first by rewrite !mxE; apply: Ea.
+ by rewrite -row_mul -{1}[Y]trmxK -trmx_mul !mxE.
+ by rewrite -[Y]trmxK -!trmx_mul mulmxA -mxvec_dotmul trmx_mul trmxK vec_mxK.
+ exists (mxvec (\matrix_i a i)); split.
+ by case/mxvec_indexP=> i j; rewrite mxvecE mxE; apply: Ea.
+ rewrite -[mxvec _]trmxK -trmx_mul mxvec_dotmul -mulmxA trmx_mul !mxE.
+ apply: eq_bigr => i _; rewrite Da1 !mxE; congr (_ * _).
+ by apply: eq_bigr => j _; rewrite !mxE.
+suffices [m [X [[u [_ Du]] idealM]]]: exists m,
+ exists X, let M := memM memR m X in M 1 /\ forall y, M y -> M (q.[w] * y).
+- do [set M := memM _ m X; move: q.[w] => z] in idealM *.
+ have MX i: M (X 0 i).
+ by exists (delta_mx 0 i); split=> [j|]; rewrite -?rowE !mxE.
+ have /fin_all_exists[a /all_and2[Fa Da1]] i := idealM _ (MX i).
+ have /fin_all_exists[r Dr] i := fin_all_exists (Fa i).
+ pose A := \matrix_(i, j) r j i; pose B := z%:M - map_mx RtoK A.
+ have XB0: X *m B = 0.
+ apply/eqP; rewrite mulmxBr mul_mx_scalar subr_eq0; apply/eqP/rowP=> i.
+ by rewrite !mxE Da1 mxE; apply: eq_bigr=> j _; rewrite !mxE mulrC Dr.
+ exists (char_poly A); first exact: char_poly_monic.
+ have: (\det B *: (u *m X^T)) 0 0 == 0.
+ rewrite scalemxAr -linearZ -mul_mx_scalar -mul_mx_adj mulmxA XB0 /=.
+ by rewrite mul0mx trmx0 mulmx0 mxE.
+ rewrite mxE -Du mulr1 rootE -horner_evalE -!det_map_mx; congr (\det _ == 0).
+ rewrite !raddfB /= !map_scalar_mx /= map_polyX horner_evalE hornerX.
+ by apply/matrixP=> i j; rewrite !mxE map_polyC /horner_eval hornerC.
+pose gen1 x E y := exists2 r, pXin E r & y = r.[x]; pose gen := foldr gen1 memR.
+have gen1S (E : K -> Prop) x y: E 0 -> E y -> gen1 x E y.
+ by exists y%:P => [i|]; rewrite ?hornerC ?coefC //; case: ifP.
+have genR S y: memR y -> gen S y.
+ by elim: S => //= x S IH in y * => /IH; apply: gen1S; apply: IH.
+have gen0 := genR _ 0 memR0; have gen_1 := genR _ 1 memR1.
+have{gen1S} genS S y: y \in S -> gen S y.
+ elim: S => //= x S IH /predU1P[-> | /IH//]; last exact: gen1S.
+ by exists 'X => [i|]; rewrite ?hornerX // coefX; apply: genR.
+pose propD (R : K -> Prop) := forall x y, R x -> R y -> R (x + y).
+have memRD: propD memR.
+ by move=> _ _ [a ->] [b ->]; exists (a + b); rewrite rmorphD.
+have genD S: propD (gen S).
+ elim: S => //= x S IH _ _ [r1 Sr1 ->] [r2 Sr2 ->]; rewrite -hornerD.
+ by exists (r1 + r2) => // i; rewrite coefD; apply: IH.
+have gen_sum S := big_ind _ (gen0 S) (genD S).
+pose propM (R : K -> Prop) := forall x y, R x -> R y -> R (x * y).
+have memRM: propM memR.
+ by move=> _ _ [a ->] [b ->]; exists (a * b); rewrite rmorphM.
+have genM S: propM (gen S).
+ elim: S => //= x S IH _ _ [r1 Sr1 ->] [r2 Sr2 ->]; rewrite -hornerM.
+ by exists (r1 * r2) => // i; rewrite coefM; apply: gen_sum => j _; apply: IH.
+have gen_horner S r y: pXin (gen S) r -> gen S y -> gen S r.[y].
+ move=> Sq Sy; rewrite horner_coef; apply: gen_sum => [[i _] /= _].
+ by elim: {2}i => [|n IHn]; rewrite ?mulr1 // exprSr mulrA; apply: genM.
+pose S := w :: q ++ p; suffices [m [X defX]]: finM memR (gen S).
+ exists m, X => M; split=> [|y /defX Xy]; first exact/defX.
+ apply/defX/genM => //; apply: gen_horner => // [i|]; last exact/genS/mem_head.
+ rewrite -[q]coefK coef_poly; case: ifP => // lt_i_q.
+ by apply: genS; rewrite inE mem_cat mem_nth ?orbT.
+pose intR R y := exists r, [/\ r \is monic, root r y & pXin R r].
+pose fix genI s := if s is y :: s1 then intR (gen s1) y /\ genI s1 else True.
+have{mon_p pw0 intRp intRq}: genI S.
+ split; set S1 := _ ++ _; first exists p.
+ split=> // i; rewrite -[p]coefK coef_poly; case: ifP => // lt_i_p.
+ by apply: genS; rewrite mem_cat orbC mem_nth.
+ have: all (mem S1) S1 by exact/allP.
+ elim: {-1}S1 => //= y S2 IH /andP[S1y S12]; split; last exact: IH.
+ have{q S S1 IH S1y S12 intRp intRq} [q mon_q qx0]: integralOver RtoK y.
+ by move: S1y; rewrite mem_cat => /orP[]; [apply: intRq | apply: intRp].
+ exists (map_poly RtoK q); split=> // [|i]; first exact: monic_map.
+ by rewrite coef_map /=; apply: genR.
+elim: {w p q}S => /= [_|x S IH [[p [mon_p px0 Sp]] /IH{IH}[m2 [X2 defS]]]].
+ exists 1%N, 1 => y; split=> [[a [Fa ->]] | Fy].
+ by rewrite tr_scalar_mx mulmx1; apply: Fa.
+ by exists y%:M; split=> [i|]; rewrite 1?ord1 ?tr_scalar_mx ?mulmx1 mxE.
+pose m1 := (size p).-1; pose X1 := \row_(i < m1) x ^+ i.
+have [m [X defM]] := tensorM memR m1 m2 X1 X2; set M := memM _ _ _ in defM.
+exists m, X => y; rewrite -/M; split=> [/defM[a [M2a]] | [q Sq]] -> {y}.
+ exists (rVpoly a) => [i|].
+ by rewrite coef_rVpoly; case/insub: i => // i; apply/defS/M2a.
+ rewrite mxE (horner_coef_wide _ (size_poly _ _)) -/(rVpoly a).
+ by apply: eq_bigr => i _; rewrite coef_rVpoly_ord !mxE.
+have M_0: M 0 by exists 0; split=> [i|]; rewrite ?mul0mx mxE.
+have M_D: propD M.
+ move=> _ _ [a [Fa ->]] [b [Fb ->]]; exists (a + b).
+ by rewrite mulmxDl !mxE; split=> // i; rewrite mxE; apply: memRD.
+have{M_0 M_D} Msum := big_ind _ M_0 M_D.
+rewrite horner_coef; apply: (Msum) => i _; case: i q`_i {Sq}(Sq i) => /=.
+elim: {q}(size q) => // n IHn i i_le_n y Sy.
+have [i_lt_m1 | m1_le_i] := ltnP i m1.
+ apply/defM; exists (y *: delta_mx 0 (Ordinal i_lt_m1)); split=> [j|].
+ by apply/defS; rewrite !mxE /= mulr_natr; case: eqP.
+ by rewrite -scalemxAl -rowE !mxE.
+rewrite -(subnK m1_le_i) exprD -[x ^+ m1]subr0 -(rootP px0) horner_coef.
+rewrite polySpred ?monic_neq0 // -/m1 big_ord_recr /= -lead_coefE.
+rewrite opprD addrC (monicP mon_p) mul1r subrK !mulrN -mulNr !mulr_sumr.
+apply: Msum => j _; rewrite mulrA mulrACA -exprD; apply: IHn.
+ by rewrite -addnS addnC addnBA // leq_subLR leq_add.
+by rewrite -mulN1r; do 2!apply: (genM) => //; apply: genR.
+Qed.
+
+Lemma integral_root_monic u p :
+ p \is monic -> root p u -> {in p : seq K, integralRange RtoK} ->
+ integralOver RtoK u.
+Proof.
+move=> mon_p pu0 intRp; rewrite -[u]hornerX.
+apply: integral_horner_root mon_p pu0 intRp _.
+by apply/integral_poly => i; rewrite coefX; apply: integral_nat.
+Qed.
+
+Hint Resolve (integral0 RtoK) (integral1 RtoK) (@monicXsubC K).
+
+Let XsubC0 (u : K) : root ('X - u%:P) u. Proof. by rewrite root_XsubC. Qed.
+Let intR_XsubC u :
+ integralOver RtoK (- u) -> {in 'X - u%:P : seq K, integralRange RtoK}.
+Proof. by move=> intRu v; rewrite polyseqXsubC !inE => /pred2P[]->. Qed.
+
+Lemma integral_opp u : integralOver RtoK u -> integralOver RtoK (- u).
+Proof. by rewrite -{1}[u]opprK => /intR_XsubC/integral_root_monic; apply. Qed.
+
+Lemma integral_horner (p : {poly K}) u :
+ {in p : seq K, integralRange RtoK} -> integralOver RtoK u ->
+ integralOver RtoK p.[u].
+Proof. by move=> ? /integral_opp/intR_XsubC/integral_horner_root; apply. Qed.
+
+Lemma integral_sub u v :
+ integralOver RtoK u -> integralOver RtoK v -> integralOver RtoK (u - v).
+Proof.
+move=> intRu /integral_opp/intR_XsubC/integral_horner/(_ intRu).
+by rewrite !hornerE.
+Qed.
+
+Lemma integral_add u v :
+ integralOver RtoK u -> integralOver RtoK v -> integralOver RtoK (u + v).
+Proof. by rewrite -{2}[v]opprK => intRu /integral_opp; apply: integral_sub. Qed.
+
+Lemma integral_mul u v :
+ integralOver RtoK u -> integralOver RtoK v -> integralOver RtoK (u * v).
+Proof.
+rewrite -{2}[v]hornerX -hornerZ => intRu; apply: integral_horner.
+by apply/integral_poly=> i; rewrite coefZ coefX mulr_natr mulrb; case: ifP.
+Qed.
+
+End IntegralOverComRing.
+
+Section IntegralOverField.
+
+Variables (F E : fieldType) (FtoE : {rmorphism F -> E}).
+
+Definition algebraicOver (fFtoE : F -> E) u :=
+ exists2 p, p != 0 & root (map_poly fFtoE p) u.
+
+Notation mk_mon p := ((lead_coef p)^-1 *: p).
+
+Lemma integral_algebraic u : algebraicOver FtoE u <-> integralOver FtoE u.
+Proof.
+split=> [] [p p_nz pu0]; last by exists p; rewrite ?monic_neq0.
+exists (mk_mon p); first by rewrite monicE lead_coefZ mulVf ?lead_coef_eq0.
+by rewrite linearZ rootE hornerZ (rootP pu0) mulr0.
+Qed.
+
+Lemma algebraic_id a : algebraicOver FtoE (FtoE a).
+Proof. exact/integral_algebraic/integral_id. Qed.
+
+Lemma algebraic0 : algebraicOver FtoE 0.
+Proof. exact/integral_algebraic/integral0. Qed.
+
+Lemma algebraic1 : algebraicOver FtoE 1.
+Proof. exact/integral_algebraic/integral1. Qed.
+
+Lemma algebraic_opp x : algebraicOver FtoE x -> algebraicOver FtoE (- x).
+Proof. by move/integral_algebraic/integral_opp/integral_algebraic. Qed.
+
+Lemma algebraic_add x y :
+ algebraicOver FtoE x -> algebraicOver FtoE y -> algebraicOver FtoE (x + y).
+Proof.
+move/integral_algebraic=> intFx /integral_algebraic intFy.
+exact/integral_algebraic/integral_add.
+Qed.
+
+Lemma algebraic_sub x y :
+ algebraicOver FtoE x -> algebraicOver FtoE y -> algebraicOver FtoE (x - y).
+Proof. by move=> algFx /algebraic_opp; apply: algebraic_add. Qed.
+
+Lemma algebraic_mul x y :
+ algebraicOver FtoE x -> algebraicOver FtoE y -> algebraicOver FtoE (x * y).
+Proof.
+move/integral_algebraic=> intFx /integral_algebraic intFy.
+exact/integral_algebraic/integral_mul.
+Qed.
+
+Lemma algebraic_inv u : algebraicOver FtoE u -> algebraicOver FtoE u^-1.
+Proof.
+have [-> | /expf_neq0 nz_u_n] := eqVneq u 0; first by rewrite invr0.
+case=> p nz_p pu0; exists (Poly (rev p)).
+ apply/eqP=> /polyP/(_ 0%N); rewrite coef_Poly coef0 nth_rev ?size_poly_gt0 //.
+ by apply/eqP; rewrite subn1 lead_coef_eq0.
+apply/eqP/(mulfI (nz_u_n (size p).-1)); rewrite mulr0 -(rootP pu0).
+rewrite (@horner_coef_wide _ (size p)); last first.
+ by rewrite size_map_poly -(size_rev p) size_Poly.
+rewrite horner_coef mulr_sumr size_map_poly.
+rewrite [rhs in _ = rhs](reindex_inj rev_ord_inj) /=.
+apply: eq_bigr => i _; rewrite !coef_map coef_Poly nth_rev // mulrCA.
+by congr (_ * _); rewrite -{1}(subnKC (valP i)) addSn addnC exprD exprVn ?mulfK.
+Qed.
+
+Lemma algebraic_div x y :
+ algebraicOver FtoE x -> algebraicOver FtoE y -> algebraicOver FtoE (x / y).
+Proof. by move=> algFx /algebraic_inv; apply: algebraic_mul. Qed.
+
+Lemma integral_inv x : integralOver FtoE x -> integralOver FtoE x^-1.
+Proof. by move/integral_algebraic/algebraic_inv/integral_algebraic. Qed.
+
+Lemma integral_div x y :
+ integralOver FtoE x -> integralOver FtoE y -> integralOver FtoE (x / y).
+Proof. by move=> algFx /integral_inv; apply: integral_mul. Qed.
+
+Lemma integral_root p u :
+ p != 0 -> root p u -> {in p : seq E, integralRange FtoE} ->
+ integralOver FtoE u.
+Proof.
+move=> nz_p pu0 algFp.
+have mon_p1: mk_mon p \is monic.
+ by rewrite monicE lead_coefZ mulVf ?lead_coef_eq0.
+have p1u0: root (mk_mon p) u by rewrite rootE hornerZ (rootP pu0) mulr0.
+apply: integral_root_monic mon_p1 p1u0 _ => _ /(nthP 0)[i ltip <-].
+rewrite coefZ mulrC; rewrite size_scale ?invr_eq0 ?lead_coef_eq0 // in ltip.
+by apply: integral_div; apply/algFp/mem_nth; rewrite -?polySpred.
+Qed.
+
+End IntegralOverField.
+
+(* Lifting term, formula, envs and eval to matrices. Wlog, and for the sake *)
+(* of simplicity, we only lift (tensor) envs to row vectors; we can always *)
+(* use mxvec/vec_mx to store and retrieve matrices. *)
+(* We don't provide definitions for addition, subtraction, scaling, etc, *)
+(* because they have simple matrix expressions. *)
+Module MatrixFormula.
+
+Section MatrixFormula.
+
+Variable F : fieldType.
+
+Local Notation False := GRing.False.
+Local Notation True := GRing.True.
+Local Notation And := GRing.And (only parsing).
+Local Notation Add := GRing.Add (only parsing).
+Local Notation Bool b := (GRing.Bool b%bool).
+Local Notation term := (GRing.term F).
+Local Notation form := (GRing.formula F).
+Local Notation eval := GRing.eval.
+Local Notation holds := GRing.holds.
+Local Notation qf_form := GRing.qf_form.
+Local Notation qf_eval := GRing.qf_eval.
+
+Definition eval_mx (e : seq F) := map_mx (eval e).
+
+Definition mx_term := map_mx (@GRing.Const F).
+
+Lemma eval_mx_term e m n (A : 'M_(m, n)) : eval_mx e (mx_term A) = A.
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Definition mulmx_term m n p (A : 'M[term]_(m, n)) (B : 'M_(n, p)) :=
+ \matrix_(i, k) (\big[Add/0]_j (A i j * B j k))%T.
+
+Lemma eval_mulmx e m n p (A : 'M[term]_(m, n)) (B : 'M_(n, p)) :
+ eval_mx e (mulmx_term A B) = eval_mx e A *m eval_mx e B.
+Proof.
+apply/matrixP=> i k; rewrite !mxE /= ((big_morph (eval e)) 0 +%R) //=.
+by apply: eq_bigr => j _; rewrite /= !mxE.
+Qed.
+
+Local Notation morphAnd f := ((big_morph f) true andb).
+
+Let Schur m n (A : 'M[term]_(1 + m, 1 + n)) (a := A 0 0) :=
+ \matrix_(i, j) (drsubmx A i j - a^-1 * dlsubmx A i 0%R * ursubmx A 0%R j)%T.
+
+Fixpoint mxrank_form (r m n : nat) : 'M_(m, n) -> form :=
+ match m, n return 'M_(m, n) -> form with
+ | m'.+1, n'.+1 => fun A : 'M_(1 + m', 1 + n') =>
+ let nzA k := A k.1 k.2 != 0 in
+ let xSchur k := Schur (xrow k.1 0%R (xcol k.2 0%R A)) in
+ let recf k := Bool (r > 0) /\ mxrank_form r.-1 (xSchur k) in
+ GRing.Pick nzA recf (Bool (r == 0%N))
+ | _, _ => fun _ => Bool (r == 0%N)
+ end%T.
+
+Lemma mxrank_form_qf r m n (A : 'M_(m, n)) : qf_form (mxrank_form r A).
+Proof.
+by elim: m r n A => [|m IHm] r [|n] A //=; rewrite GRing.Pick_form_qf /=.
+Qed.
+
+Lemma eval_mxrank e r m n (A : 'M_(m, n)) :
+ qf_eval e (mxrank_form r A) = (\rank (eval_mx e A) == r).
+Proof.
+elim: m r n A => [|m IHm] r [|n] A /=; try by case r.
+rewrite GRing.eval_Pick /mxrank unlock /=; set pf := fun _ => _.
+rewrite -(@eq_pick _ pf) => [|k]; rewrite {}/pf ?mxE // eq_sym.
+case: pick => [[i j]|] //=; set B := _ - _; have:= mxrankE B.
+case: (Gaussian_elimination B) r => [[_ _] _] [|r] //= <-; rewrite {}IHm eqSS.
+by congr (\rank _ == r); apply/matrixP=> k l; rewrite !(mxE, big_ord1) !tpermR.
+Qed.
+
+Lemma eval_vec_mx e m n (u : 'rV_(m * n)) :
+ eval_mx e (vec_mx u) = vec_mx (eval_mx e u).
+Proof. by apply/matrixP=> i j; rewrite !mxE. Qed.
+
+Lemma eval_mxvec e m n (A : 'M_(m, n)) :
+ eval_mx e (mxvec A) = mxvec (eval_mx e A).
+Proof. by rewrite -{2}[A]mxvecK eval_vec_mx vec_mxK. Qed.
+
+Section Subsetmx.
+
+Variables (m1 m2 n : nat) (A : 'M[term]_(m1, n)) (B : 'M[term]_(m2, n)).
+
+Definition submx_form :=
+ \big[And/True]_(r < n.+1) (mxrank_form r (col_mx A B) ==> mxrank_form r B)%T.
+
+Lemma eval_col_mx e :
+ eval_mx e (col_mx A B) = col_mx (eval_mx e A) (eval_mx e B).
+Proof. by apply/matrixP=> i j; do 2![rewrite !mxE //; case: split => ?]. Qed.
+
+Lemma submx_form_qf : qf_form submx_form.
+Proof.
+by rewrite (morphAnd (@qf_form _)) ?big1 //= => r _; rewrite !mxrank_form_qf.
+Qed.
+
+Lemma eval_submx e : qf_eval e submx_form = (eval_mx e A <= eval_mx e B)%MS.
+Proof.
+rewrite (morphAnd (qf_eval e)) //= big_andE /=.
+apply/forallP/idP=> /= [|sAB d]; last first.
+ rewrite !eval_mxrank eval_col_mx -addsmxE; apply/implyP=> /eqP <-.
+ by rewrite mxrank_leqif_sup ?addsmxSr // addsmx_sub sAB /=.
+move/(_ (inord (\rank (eval_mx e (col_mx A B))))).
+rewrite inordK ?ltnS ?rank_leq_col // !eval_mxrank eqxx /= eval_col_mx.
+by rewrite -addsmxE mxrank_leqif_sup ?addsmxSr // addsmx_sub; case/andP.
+Qed.
+
+End Subsetmx.
+
+Section Env.
+
+Variable d : nat.
+
+Definition seq_of_rV (v : 'rV_d) : seq F := fgraph [ffun i => v 0 i].
+
+Lemma size_seq_of_rV v : size (seq_of_rV v) = d.
+Proof. by rewrite tuple.size_tuple card_ord. Qed.
+
+Lemma nth_seq_of_rV x0 v (i : 'I_d) : nth x0 (seq_of_rV v) i = v 0 i.
+Proof. by rewrite nth_fgraph_ord ffunE. Qed.
+
+Definition row_var k : 'rV[term]_d := \row_i ('X_(k * d + i))%T.
+
+Definition row_env (e : seq 'rV_d) := flatten (map seq_of_rV e).
+
+Lemma nth_row_env e k (i : 'I_d) : (row_env e)`_(k * d + i) = e`_k 0 i.
+Proof.
+elim: e k => [|v e IHe] k; first by rewrite !nth_nil mxE.
+rewrite /row_env /= nth_cat size_seq_of_rV.
+case: k => [|k]; first by rewrite (valP i) nth_seq_of_rV.
+by rewrite mulSn -addnA -if_neg -leqNgt leq_addr addKn IHe.
+Qed.
+
+Lemma eval_row_var e k : eval_mx (row_env e) (row_var k) = e`_k :> 'rV_d.
+Proof. by apply/rowP=> i; rewrite !mxE /= nth_row_env. Qed.
+
+Definition Exists_row_form k (f : form) :=
+ foldr GRing.Exists f (codom (fun i : 'I_d => k * d + i)%N).
+
+Lemma Exists_rowP e k f :
+ d > 0 ->
+ ((exists v : 'rV[F]_d, holds (row_env (set_nth 0 e k v)) f)
+ <-> holds (row_env e) (Exists_row_form k f)).
+Proof.
+move=> d_gt0; pose i_ j := Ordinal (ltn_pmod j d_gt0).
+have d_eq j: (j = j %/ d * d + i_ j)%N := divn_eq j d.
+split=> [[v f_v] | ]; last case/GRing.foldExistsP=> e' ee' f_e'.
+ apply/GRing.foldExistsP; exists (row_env (set_nth 0 e k v)) => {f f_v}// j.
+ rewrite [j]d_eq !nth_row_env nth_set_nth /=; case: eqP => // ->.
+ by case/imageP; exists (i_ j).
+exists (\row_i e'`_(k * d + i)); apply: eq_holds f_e' => j /=.
+move/(_ j): ee'; rewrite [j]d_eq !nth_row_env nth_set_nth /=.
+case: eqP => [-> | ne_j_k -> //]; first by rewrite mxE.
+apply/mapP=> [[r lt_r_d]]; rewrite -d_eq => def_j; case: ne_j_k.
+by rewrite def_j divnMDl // divn_small ?addn0.
+Qed.
+
+End Env.
+
+End MatrixFormula.
+
+End MatrixFormula.
diff --git a/mathcomp/algebra/poly.v b/mathcomp/algebra/poly.v
new file mode 100644
index 0000000..813f70d
--- /dev/null
+++ b/mathcomp/algebra/poly.v
@@ -0,0 +1,2591 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div choice fintype.
+Require Import bigop ssralg binomial.
+
+(******************************************************************************)
+(* This file provides a library for univariate polynomials over ring *)
+(* structures; it also provides an extended theory for polynomials whose *)
+(* coefficients range over commutative rings and integral domains. *)
+(* *)
+(* {poly R} == the type of polynomials with coefficients of type R, *)
+(* represented as lists with a non zero last element *)
+(* (big endian representation); the coeficient type R *)
+(* must have a canonical ringType structure cR. In fact *)
+(* {poly R} denotes the concrete type polynomial cR; R *)
+(* is just a phantom argument that lets type inference *)
+(* reconstruct the (hidden) ringType structure cR. *)
+(* p : seq R == the big-endian sequence of coefficients of p, via *)
+(* the coercion polyseq : polynomial >-> seq. *)
+(* Poly s == the polynomial with coefficient sequence s (ignoring *)
+(* trailing zeroes). *)
+(* \poly_(i < n) E(i) == the polynomial of degree at most n - 1 whose *)
+(* coefficients are given by the general term E(i) *)
+(* 0, 1, - p, p + q, == the usual ring operations: {poly R} has a canonical *)
+(* p * q, p ^+ n, ... ringType structure, which is commutative / integral *)
+(* when R is commutative / integral, respectively. *)
+(* polyC c, c%:P == the constant polynomial c *)
+(* 'X == the (unique) variable *)
+(* 'X^n == a power of 'X; 'X^0 is 1, 'X^1 is convertible to 'X *)
+(* p`_i == the coefficient of 'X^i in p; this is in fact just *)
+(* the ring_scope notation generic seq-indexing using *)
+(* nth 0%R, combined with the polyseq coercion. *)
+(* coefp i == the linear function p |-> p`_i (self-exapanding). *)
+(* size p == 1 + the degree of p, or 0 if p = 0 (this is the *)
+(* generic seq function combined with polyseq). *)
+(* lead_coef p == the coefficient of the highest monomial in p, or 0 *)
+(* if p = 0 (hence lead_coef p = 0 iff p = 0) *)
+(* p \is monic <=> lead_coef p == 1 (0 is not monic). *)
+(* p \is a polyOver S <=> the coefficients of p satisfy S; S should have a *)
+(* key that should be (at least) an addrPred. *)
+(* p.[x] == the evaluation of a polynomial p at a point x using *)
+(* the Horner scheme *)
+(* *** The multi-rule hornerE (resp., hornerE_comm) unwinds *)
+(* horner evaluation of a polynomial expression (resp., *)
+(* in a non commutative ring, with side conditions). *)
+(* p^`() == formal derivative of p *)
+(* p^`(n) == formal n-derivative of p *)
+(* p^`N(n) == formal n-derivative of p divided by n! *)
+(* p \Po q == polynomial composition; because this is naturally a *)
+(* a linear morphism in the first argument, this *)
+(* notation is transposed (q comes before p for redex *)
+(* selection, etc). *)
+(* := \sum(i < size p) p`_i *: q ^+ i *)
+(* comm_poly p x == x and p.[x] commute; this is a sufficient condition *)
+(* for evaluating (q * p).[x] as q.[x] * p.[x] when R *)
+(* is not commutative. *)
+(* comm_coef p x == x commutes with all the coefficients of p (clearly, *)
+(* this implies comm_poly p x). *)
+(* root p x == x is a root of p, i.e., p.[x] = 0 *)
+(* n.-unity_root x == x is an nth root of unity, i.e., a root of 'X^n - 1 *)
+(* n.-primitive_root x == x is a primitive nth root of unity, i.e., n is the *)
+(* least positive integer m > 0 such that x ^+ m = 1. *)
+(* *** The submodule poly.UnityRootTheory can be used to *)
+(* import selectively the part of the theory of roots *)
+(* of unity that doesn't mention polynomials explicitly *)
+(* map_poly f p == the image of the polynomial by the function f (which *)
+(* (locally, p^f) is usually a ring morphism). *)
+(* p^:P == p lifted to {poly {poly R}} (:= map_poly polyC p). *)
+(* commr_rmorph f u == u commutes with the image of f (i.e., with all f x). *)
+(* horner_morph cfu == given cfu : commr_rmorph f u, the function mapping p *)
+(* to the value of map_poly f p at u; this is a ring *)
+(* morphism from {poly R} to the codomain of f when f *)
+(* is a ring morphism. *)
+(* horner_eval u == the function mapping p to p.[u]; this function can *)
+(* only be used for u in a commutative ring, so it is *)
+(* always a linear ring morphism from {poly R} to R. *)
+(* diff_roots x y == x and y are distinct roots; if R is a field, this *)
+(* just means x != y, but this concept is generalized *)
+(* to the case where R is only a ring with units (i.e., *)
+(* a unitRingType); in which case it means that x and y *)
+(* commute, and that the difference x - y is a unit *)
+(* (i.e., has a multiplicative inverse) in R. *)
+(* to just x != y). *)
+(* uniq_roots s == s is a sequence or pairwise distinct roots, in the *)
+(* sense of diff_roots p above. *)
+(* *** We only show that these operations and properties are transferred by *)
+(* morphisms whose domain is a field (thus ensuring injectivity). *)
+(* We prove the factor_theorem, and the max_poly_roots inequality relating *)
+(* the number of distinct roots of a polynomial and its size. *)
+(* The some polynomial lemmas use following suffix interpretation : *)
+(* C - constant polynomial (as in polyseqC : a%:P = nseq (a != 0) a). *)
+(* X - the polynomial variable 'X (as in coefX : 'X`_i = (i == 1%N)). *)
+(* Xn - power of 'X (as in monicXn : monic 'X^n). *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GRing.Theory.
+Open Local Scope ring_scope.
+
+Reserved Notation "{ 'poly' T }" (at level 0, format "{ 'poly' T }").
+Reserved Notation "c %:P" (at level 2, format "c %:P").
+Reserved Notation "p ^:P" (at level 2, format "p ^:P").
+Reserved Notation "'X" (at level 0).
+Reserved Notation "''X^' n" (at level 3, n at level 2, format "''X^' n").
+Reserved Notation "\poly_ ( i < n ) E"
+ (at level 36, E at level 36, i, n at level 50,
+ format "\poly_ ( i < n ) E").
+Reserved Notation "p \Po q" (at level 50).
+Reserved Notation "p ^`N ( n )" (at level 8, format "p ^`N ( n )").
+Reserved Notation "n .-unity_root" (at level 2, format "n .-unity_root").
+Reserved Notation "n .-primitive_root"
+ (at level 2, format "n .-primitive_root").
+
+Local Notation simp := Monoid.simpm.
+
+Section Polynomial.
+
+Variable R : ringType.
+
+(* Defines a polynomial as a sequence with <> 0 last element *)
+Record polynomial := Polynomial {polyseq :> seq R; _ : last 1 polyseq != 0}.
+
+Canonical polynomial_subType := Eval hnf in [subType for polyseq].
+Definition polynomial_eqMixin := Eval hnf in [eqMixin of polynomial by <:].
+Canonical polynomial_eqType := Eval hnf in EqType polynomial polynomial_eqMixin.
+Definition polynomial_choiceMixin := [choiceMixin of polynomial by <:].
+Canonical polynomial_choiceType :=
+ Eval hnf in ChoiceType polynomial polynomial_choiceMixin.
+
+Lemma poly_inj : injective polyseq. Proof. exact: val_inj. Qed.
+
+Definition poly_of of phant R := polynomial.
+Identity Coercion type_poly_of : poly_of >-> polynomial.
+
+Definition coefp_head h i (p : poly_of (Phant R)) := let: tt := h in p`_i.
+
+End Polynomial.
+
+(* We need to break off the section here to let the argument scope *)
+(* directives take effect. *)
+Bind Scope ring_scope with poly_of.
+Bind Scope ring_scope with polynomial.
+Arguments Scope polyseq [_ ring_scope].
+Arguments Scope poly_inj [_ ring_scope ring_scope _].
+Arguments Scope coefp_head [_ _ nat_scope ring_scope _].
+Notation "{ 'poly' T }" := (poly_of (Phant T)).
+Notation coefp i := (coefp_head tt i).
+
+Section PolynomialTheory.
+
+Variable R : ringType.
+Implicit Types (a b c x y z : R) (p q r d : {poly R}).
+
+Canonical poly_subType := Eval hnf in [subType of {poly R}].
+Canonical poly_eqType := Eval hnf in [eqType of {poly R}].
+Canonical poly_choiceType := Eval hnf in [choiceType of {poly R}].
+
+Definition lead_coef p := p`_(size p).-1.
+Lemma lead_coefE p : lead_coef p = p`_(size p).-1. Proof. by []. Qed.
+
+Definition poly_nil := @Polynomial R [::] (oner_neq0 R).
+Definition polyC c : {poly R} := insubd poly_nil [:: c].
+
+Local Notation "c %:P" := (polyC c).
+
+(* Remember the boolean (c != 0) is coerced to 1 if true and 0 if false *)
+Lemma polyseqC c : c%:P = nseq (c != 0) c :> seq R.
+Proof. by rewrite val_insubd /=; case: (c == 0). Qed.
+
+Lemma size_polyC c : size c%:P = (c != 0).
+Proof. by rewrite polyseqC size_nseq. Qed.
+
+Lemma coefC c i : c%:P`_i = if i == 0%N then c else 0.
+Proof. by rewrite polyseqC; case: i => [|[]]; case: eqP. Qed.
+
+Lemma polyCK : cancel polyC (coefp 0).
+Proof. by move=> c; rewrite [coefp 0 _]coefC. Qed.
+
+Lemma polyC_inj : injective polyC.
+Proof. by move=> c1 c2 eqc12; have:= coefC c2 0; rewrite -eqc12 coefC. Qed.
+
+Lemma lead_coefC c : lead_coef c%:P = c.
+Proof. by rewrite /lead_coef polyseqC; case: eqP. Qed.
+
+(* Extensional interpretation (poly <=> nat -> R) *)
+Lemma polyP p q : nth 0 p =1 nth 0 q <-> p = q.
+Proof.
+split=> [eq_pq | -> //]; apply: poly_inj.
+without loss lt_pq: p q eq_pq / size p < size q.
+ move=> IH; case: (ltngtP (size p) (size q)); try by move/IH->.
+ move/(@eq_from_nth _ 0); exact.
+case: q => q nz_q /= in lt_pq eq_pq *; case/eqP: nz_q.
+by rewrite (last_nth 0) -(subnKC lt_pq) /= -eq_pq nth_default ?leq_addr.
+Qed.
+
+Lemma size1_polyC p : size p <= 1 -> p = (p`_0)%:P.
+Proof.
+move=> le_p_1; apply/polyP=> i; rewrite coefC.
+by case: i => // i; rewrite nth_default // (leq_trans le_p_1).
+Qed.
+
+(* Builds a polynomial by extension. *)
+Definition cons_poly c p : {poly R} :=
+ if p is Polynomial ((_ :: _) as s) ns then
+ @Polynomial R (c :: s) ns
+ else c%:P.
+
+Lemma polyseq_cons c p :
+ cons_poly c p = (if ~~ nilp p then c :: p else c%:P) :> seq R.
+Proof. by case: p => [[]]. Qed.
+
+Lemma size_cons_poly c p :
+ size (cons_poly c p) = (if nilp p && (c == 0) then 0%N else (size p).+1).
+Proof. by case: p => [[|c' s] _] //=; rewrite size_polyC; case: eqP. Qed.
+
+Lemma coef_cons c p i : (cons_poly c p)`_i = if i == 0%N then c else p`_i.-1.
+Proof.
+by case: p i => [[|c' s] _] [] //=; rewrite polyseqC; case: eqP => //= _ [].
+Qed.
+
+(* Build a polynomial directly from a list of coefficients. *)
+Definition Poly := foldr cons_poly 0%:P.
+
+Lemma PolyK c s : last c s != 0 -> Poly s = s :> seq R.
+Proof.
+case: s => {c}/= [_ |c s]; first by rewrite polyseqC eqxx.
+elim: s c => /= [|a s IHs] c nz_c; rewrite polyseq_cons ?{}IHs //.
+by rewrite !polyseqC !eqxx nz_c.
+Qed.
+
+Lemma polyseqK p : Poly p = p.
+Proof. by apply: poly_inj; exact: PolyK (valP p). Qed.
+
+Lemma size_Poly s : size (Poly s) <= size s.
+Proof.
+elim: s => [|c s IHs] /=; first by rewrite polyseqC eqxx.
+by rewrite polyseq_cons; case: ifP => // _; rewrite size_polyC; case: (~~ _).
+Qed.
+
+Lemma coef_Poly s i : (Poly s)`_i = s`_i.
+Proof.
+by elim: s i => [|c s IHs] /= [|i]; rewrite !(coefC, eqxx, coef_cons) /=.
+Qed.
+
+(* Build a polynomial from an infinite sequence of coefficients and a bound. *)
+Definition poly_expanded_def n E := Poly (mkseq E n).
+Fact poly_key : unit. Proof. by []. Qed.
+Definition poly := locked_with poly_key poly_expanded_def.
+Canonical poly_unlockable := [unlockable fun poly].
+Local Notation "\poly_ ( i < n ) E" := (poly n (fun i : nat => E)).
+
+Lemma polyseq_poly n E :
+ E n.-1 != 0 -> \poly_(i < n) E i = mkseq [eta E] n :> seq R.
+Proof.
+rewrite unlock; case: n => [|n] nzEn; first by rewrite polyseqC eqxx.
+by rewrite (@PolyK 0) // -nth_last nth_mkseq size_mkseq.
+Qed.
+
+Lemma size_poly n E : size (\poly_(i < n) E i) <= n.
+Proof. by rewrite unlock (leq_trans (size_Poly _)) ?size_mkseq. Qed.
+
+Lemma size_poly_eq n E : E n.-1 != 0 -> size (\poly_(i < n) E i) = n.
+Proof. by move/polyseq_poly->; apply: size_mkseq. Qed.
+
+Lemma coef_poly n E k : (\poly_(i < n) E i)`_k = (if k < n then E k else 0).
+Proof.
+rewrite unlock coef_Poly.
+have [lt_kn | le_nk] := ltnP k n; first by rewrite nth_mkseq.
+by rewrite nth_default // size_mkseq.
+Qed.
+
+Lemma lead_coef_poly n E :
+ n > 0 -> E n.-1 != 0 -> lead_coef (\poly_(i < n) E i) = E n.-1.
+Proof.
+by case: n => // n _ nzE; rewrite /lead_coef size_poly_eq // coef_poly leqnn.
+Qed.
+
+Lemma coefK p : \poly_(i < size p) p`_i = p.
+Proof.
+by apply/polyP=> i; rewrite coef_poly; case: ltnP => // /(nth_default 0)->.
+Qed.
+
+(* Zmodule structure for polynomial *)
+Definition add_poly_def p q := \poly_(i < maxn (size p) (size q)) (p`_i + q`_i).
+Fact add_poly_key : unit. Proof. by []. Qed.
+Definition add_poly := locked_with add_poly_key add_poly_def.
+Canonical add_poly_unlockable := [unlockable fun add_poly].
+
+Definition opp_poly_def p := \poly_(i < size p) - p`_i.
+Fact opp_poly_key : unit. Proof. by []. Qed.
+Definition opp_poly := locked_with opp_poly_key opp_poly_def.
+Canonical opp_poly_unlockable := [unlockable fun opp_poly].
+
+Fact coef_add_poly p q i : (add_poly p q)`_i = p`_i + q`_i.
+Proof.
+rewrite unlock coef_poly; case: leqP => //.
+by rewrite geq_max => /andP[le_p_i le_q_i]; rewrite !nth_default ?add0r.
+Qed.
+
+Fact coef_opp_poly p i : (opp_poly p)`_i = - p`_i.
+Proof.
+rewrite unlock coef_poly /=.
+by case: leqP => // le_p_i; rewrite nth_default ?oppr0.
+Qed.
+
+Fact add_polyA : associative add_poly.
+Proof. by move=> p q r; apply/polyP=> i; rewrite !coef_add_poly addrA. Qed.
+
+Fact add_polyC : commutative add_poly.
+Proof. by move=> p q; apply/polyP=> i; rewrite !coef_add_poly addrC. Qed.
+
+Fact add_poly0 : left_id 0%:P add_poly.
+Proof.
+by move=> p; apply/polyP=> i; rewrite coef_add_poly coefC if_same add0r.
+Qed.
+
+Fact add_polyN : left_inverse 0%:P opp_poly add_poly.
+Proof.
+move=> p; apply/polyP=> i.
+by rewrite coef_add_poly coef_opp_poly coefC if_same addNr.
+Qed.
+
+Definition poly_zmodMixin :=
+ ZmodMixin add_polyA add_polyC add_poly0 add_polyN.
+
+Canonical poly_zmodType := Eval hnf in ZmodType {poly R} poly_zmodMixin.
+Canonical polynomial_zmodType :=
+ Eval hnf in ZmodType (polynomial R) poly_zmodMixin.
+
+(* Properties of the zero polynomial *)
+Lemma polyC0 : 0%:P = 0 :> {poly R}. Proof. by []. Qed.
+
+Lemma polyseq0 : (0 : {poly R}) = [::] :> seq R.
+Proof. by rewrite polyseqC eqxx. Qed.
+
+Lemma size_poly0 : size (0 : {poly R}) = 0%N.
+Proof. by rewrite polyseq0. Qed.
+
+Lemma coef0 i : (0 : {poly R})`_i = 0.
+Proof. by rewrite coefC if_same. Qed.
+
+Lemma lead_coef0 : lead_coef 0 = 0 :> R. Proof. exact: lead_coefC. Qed.
+
+Lemma size_poly_eq0 p : (size p == 0%N) = (p == 0).
+Proof. by rewrite size_eq0 -polyseq0. Qed.
+
+Lemma size_poly_leq0 p : (size p <= 0) = (p == 0).
+Proof. by rewrite leqn0 size_poly_eq0. Qed.
+
+Lemma size_poly_leq0P p : reflect (p = 0) (size p <= 0%N).
+Proof. by apply: (iffP idP); rewrite size_poly_leq0; move/eqP. Qed.
+
+Lemma size_poly_gt0 p : (0 < size p) = (p != 0).
+Proof. by rewrite lt0n size_poly_eq0. Qed.
+
+Lemma nil_poly p : nilp p = (p == 0).
+Proof. exact: size_poly_eq0. Qed.
+
+Lemma poly0Vpos p : {p = 0} + {size p > 0}.
+Proof. by rewrite lt0n size_poly_eq0; exact: eqVneq. Qed.
+
+Lemma polySpred p : p != 0 -> size p = (size p).-1.+1.
+Proof. by rewrite -size_poly_eq0 -lt0n => /prednK. Qed.
+
+Lemma lead_coef_eq0 p : (lead_coef p == 0) = (p == 0).
+Proof.
+rewrite -nil_poly /lead_coef nth_last.
+by case: p => [[|x s] /= /negbTE // _]; rewrite eqxx.
+Qed.
+
+Lemma polyC_eq0 (c : R) : (c%:P == 0) = (c == 0).
+Proof. by rewrite -nil_poly polyseqC; case: (c == 0). Qed.
+
+Lemma size_poly1P p : reflect (exists2 c, c != 0 & p = c%:P) (size p == 1%N).
+Proof.
+apply: (iffP eqP) => [pC | [c nz_c ->]]; last by rewrite size_polyC nz_c.
+have def_p: p = (p`_0)%:P by rewrite -size1_polyC ?pC.
+by exists p`_0; rewrite // -polyC_eq0 -def_p -size_poly_eq0 pC.
+Qed.
+
+Lemma leq_sizeP p i : reflect (forall j, i <= j -> p`_j = 0) (size p <= i).
+Proof.
+apply: (iffP idP) => [hp j hij| hp].
+ by apply: nth_default; apply: leq_trans hij.
+case p0: (p == 0); first by rewrite (eqP p0) size_poly0.
+move: (lead_coef_eq0 p); rewrite p0 leqNgt; move/negbT; apply: contra => hs.
+by apply/eqP; apply: hp; rewrite -ltnS (ltn_predK hs).
+Qed.
+
+(* Size, leading coef, morphism properties of coef *)
+
+Lemma coefD p q i : (p + q)`_i = p`_i + q`_i.
+Proof. exact: coef_add_poly. Qed.
+
+Lemma coefN p i : (- p)`_i = - p`_i.
+Proof. exact: coef_opp_poly. Qed.
+
+Lemma coefB p q i : (p - q)`_i = p`_i - q`_i.
+Proof. by rewrite coefD coefN. Qed.
+
+Canonical coefp_additive i :=
+ Additive ((fun p => (coefB p)^~ i) : additive (coefp i)).
+
+Lemma coefMn p n i : (p *+ n)`_i = p`_i *+ n.
+Proof. exact: (raddfMn (coefp_additive i)). Qed.
+
+Lemma coefMNn p n i : (p *- n)`_i = p`_i *- n.
+Proof. by rewrite coefN coefMn. Qed.
+
+Lemma coef_sum I (r : seq I) (P : pred I) (F : I -> {poly R}) k :
+ (\sum_(i <- r | P i) F i)`_k = \sum_(i <- r | P i) (F i)`_k.
+Proof. exact: (raddf_sum (coefp_additive k)). Qed.
+
+Lemma polyC_add : {morph polyC : a b / a + b}.
+Proof. by move=> a b; apply/polyP=> [[|i]]; rewrite coefD !coefC ?addr0. Qed.
+
+Lemma polyC_opp : {morph polyC : c / - c}.
+Proof. by move=> c; apply/polyP=> [[|i]]; rewrite coefN !coefC ?oppr0. Qed.
+
+Lemma polyC_sub : {morph polyC : a b / a - b}.
+Proof. by move=> a b; rewrite polyC_add polyC_opp. Qed.
+
+Canonical polyC_additive := Additive polyC_sub.
+
+Lemma polyC_muln n : {morph polyC : c / c *+ n}.
+Proof. exact: raddfMn. Qed.
+
+Lemma size_opp p : size (- p) = size p.
+Proof.
+by apply/eqP; rewrite eqn_leq -{3}(opprK p) -[-%R]/opp_poly unlock !size_poly.
+Qed.
+
+Lemma lead_coef_opp p : lead_coef (- p) = - lead_coef p.
+Proof. by rewrite /lead_coef size_opp coefN. Qed.
+
+Lemma size_add p q : size (p + q) <= maxn (size p) (size q).
+Proof. by rewrite -[+%R]/add_poly unlock; apply: size_poly. Qed.
+
+Lemma size_addl p q : size p > size q -> size (p + q) = size p.
+Proof.
+move=> ltqp; rewrite -[+%R]/add_poly unlock size_poly_eq (maxn_idPl (ltnW _))//.
+by rewrite addrC nth_default ?simp ?nth_last //; case: p ltqp => [[]].
+Qed.
+
+Lemma size_sum I (r : seq I) (P : pred I) (F : I -> {poly R}) :
+ size (\sum_(i <- r | P i) F i) <= \max_(i <- r | P i) size (F i).
+Proof.
+elim/big_rec2: _ => [|i p q _ IHp]; first by rewrite size_poly0.
+by rewrite -(maxn_idPr IHp) maxnA leq_max size_add.
+Qed.
+
+Lemma lead_coefDl p q : size p > size q -> lead_coef (p + q) = lead_coef p.
+Proof.
+move=> ltqp; rewrite /lead_coef coefD size_addl //.
+by rewrite addrC nth_default ?simp // -ltnS (ltn_predK ltqp).
+Qed.
+
+(* Polynomial ring structure. *)
+
+Definition mul_poly_def p q :=
+ \poly_(i < (size p + size q).-1) (\sum_(j < i.+1) p`_j * q`_(i - j)).
+Fact mul_poly_key : unit. Proof. by []. Qed.
+Definition mul_poly := locked_with mul_poly_key mul_poly_def.
+Canonical mul_poly_unlockable := [unlockable fun mul_poly].
+
+Fact coef_mul_poly p q i :
+ (mul_poly p q)`_i = \sum_(j < i.+1) p`_j * q`_(i - j)%N.
+Proof.
+rewrite unlock coef_poly -subn1 ltn_subRL add1n; case: leqP => // le_pq_i1.
+rewrite big1 // => j _; have [lq_q_ij | gt_q_ij] := leqP (size q) (i - j).
+ by rewrite [q`__]nth_default ?mulr0.
+rewrite nth_default ?mul0r // -(leq_add2r (size q)) (leq_trans le_pq_i1) //.
+by rewrite -leq_subLR -subnSK.
+Qed.
+
+Fact coef_mul_poly_rev p q i :
+ (mul_poly p q)`_i = \sum_(j < i.+1) p`_(i - j)%N * q`_j.
+Proof.
+rewrite coef_mul_poly (reindex_inj rev_ord_inj) /=.
+by apply: eq_bigr => j _; rewrite (sub_ordK j).
+Qed.
+
+Fact mul_polyA : associative mul_poly.
+Proof.
+move=> p q r; apply/polyP=> i; rewrite coef_mul_poly coef_mul_poly_rev.
+pose coef3 j k := p`_j * (q`_(i - j - k)%N * r`_k).
+transitivity (\sum_(j < i.+1) \sum_(k < i.+1 | k <= i - j) coef3 j k).
+ apply: eq_bigr => /= j _; rewrite coef_mul_poly_rev big_distrr /=.
+ by rewrite (big_ord_narrow_leq (leq_subr _ _)).
+rewrite (exchange_big_dep predT) //=; apply: eq_bigr => k _.
+transitivity (\sum_(j < i.+1 | j <= i - k) coef3 j k).
+ apply: eq_bigl => j; rewrite -ltnS -(ltnS j) -!subSn ?leq_ord //.
+ by rewrite -subn_gt0 -(subn_gt0 j) -!subnDA addnC.
+rewrite (big_ord_narrow_leq (leq_subr _ _)) coef_mul_poly big_distrl /=.
+by apply: eq_bigr => j _; rewrite /coef3 -!subnDA addnC mulrA.
+Qed.
+
+Fact mul_1poly : left_id 1%:P mul_poly.
+Proof.
+move=> p; apply/polyP => i; rewrite coef_mul_poly big_ord_recl subn0.
+by rewrite big1 => [|j _]; rewrite coefC !simp.
+Qed.
+
+Fact mul_poly1 : right_id 1%:P mul_poly.
+Proof.
+move=> p; apply/polyP => i; rewrite coef_mul_poly_rev big_ord_recl subn0.
+by rewrite big1 => [|j _]; rewrite coefC !simp.
+Qed.
+
+Fact mul_polyDl : left_distributive mul_poly +%R.
+Proof.
+move=> p q r; apply/polyP=> i; rewrite coefD !coef_mul_poly -big_split.
+by apply: eq_bigr => j _; rewrite coefD mulrDl.
+Qed.
+
+Fact mul_polyDr : right_distributive mul_poly +%R.
+Proof.
+move=> p q r; apply/polyP=> i; rewrite coefD !coef_mul_poly -big_split.
+by apply: eq_bigr => j _; rewrite coefD mulrDr.
+Qed.
+
+Fact poly1_neq0 : 1%:P != 0 :> {poly R}.
+Proof. by rewrite polyC_eq0 oner_neq0. Qed.
+
+Definition poly_ringMixin :=
+ RingMixin mul_polyA mul_1poly mul_poly1 mul_polyDl mul_polyDr poly1_neq0.
+
+Canonical poly_ringType := Eval hnf in RingType {poly R} poly_ringMixin.
+Canonical polynomial_ringType :=
+ Eval hnf in RingType (polynomial R) poly_ringMixin.
+
+Lemma polyC1 : 1%:P = 1 :> {poly R}. Proof. by []. Qed.
+
+Lemma polyseq1 : (1 : {poly R}) = [:: 1] :> seq R.
+Proof. by rewrite polyseqC oner_neq0. Qed.
+
+Lemma size_poly1 : size (1 : {poly R}) = 1%N.
+Proof. by rewrite polyseq1. Qed.
+
+Lemma coef1 i : (1 : {poly R})`_i = (i == 0%N)%:R.
+Proof. by case: i => [|i]; rewrite polyseq1 /= ?nth_nil. Qed.
+
+Lemma lead_coef1 : lead_coef 1 = 1 :> R. Proof. exact: lead_coefC. Qed.
+
+Lemma coefM p q i : (p * q)`_i = \sum_(j < i.+1) p`_j * q`_(i - j)%N.
+Proof. exact: coef_mul_poly. Qed.
+
+Lemma coefMr p q i : (p * q)`_i = \sum_(j < i.+1) p`_(i - j)%N * q`_j.
+Proof. exact: coef_mul_poly_rev. Qed.
+
+Lemma size_mul_leq p q : size (p * q) <= (size p + size q).-1.
+Proof. by rewrite -[*%R]/mul_poly unlock size_poly. Qed.
+
+Lemma mul_lead_coef p q :
+ lead_coef p * lead_coef q = (p * q)`_(size p + size q).-2.
+Proof.
+pose dp := (size p).-1; pose dq := (size q).-1.
+have [-> | nz_p] := eqVneq p 0; first by rewrite lead_coef0 !mul0r coef0.
+have [-> | nz_q] := eqVneq q 0; first by rewrite lead_coef0 !mulr0 coef0.
+have ->: (size p + size q).-2 = (dp + dq)%N.
+ by do 2! rewrite polySpred // addSn addnC.
+have lt_p_pq: dp < (dp + dq).+1 by rewrite ltnS leq_addr.
+rewrite coefM (bigD1 (Ordinal lt_p_pq)) ?big1 ?simp ?addKn //= => i.
+rewrite -val_eqE neq_ltn /= => /orP[lt_i_p | gt_i_p]; last first.
+ by rewrite nth_default ?mul0r //; rewrite -polySpred in gt_i_p.
+rewrite [q`__]nth_default ?mulr0 //= -subSS -{1}addnS -polySpred //.
+by rewrite addnC -addnBA ?leq_addr.
+Qed.
+
+Lemma size_proper_mul p q :
+ lead_coef p * lead_coef q != 0 -> size (p * q) = (size p + size q).-1.
+Proof.
+apply: contraNeq; rewrite mul_lead_coef eqn_leq size_mul_leq -ltnNge => lt_pq.
+by rewrite nth_default // -subn1 -(leq_add2l 1) -leq_subLR leq_sub2r.
+Qed.
+
+Lemma lead_coef_proper_mul p q :
+ let c := lead_coef p * lead_coef q in c != 0 -> lead_coef (p * q) = c.
+Proof. by move=> /= nz_c; rewrite mul_lead_coef -size_proper_mul. Qed.
+
+Lemma size_prod_leq (I : finType) (P : pred I) (F : I -> {poly R}) :
+ size (\prod_(i | P i) F i) <= (\sum_(i | P i) size (F i)).+1 - #|P|.
+Proof.
+rewrite -sum1_card.
+elim/big_rec3: _ => [|i n m p _ IHp]; first by rewrite size_poly1.
+have [-> | nz_p] := eqVneq p 0; first by rewrite mulr0 size_poly0.
+rewrite (leq_trans (size_mul_leq _ _)) // subnS -!subn1 leq_sub2r //.
+rewrite -addnS -addnBA ?leq_add2l // ltnW // -subn_gt0 (leq_trans _ IHp) //.
+by rewrite polySpred.
+Qed.
+
+Lemma coefCM c p i : (c%:P * p)`_i = c * p`_i.
+Proof.
+rewrite coefM big_ord_recl subn0.
+by rewrite big1 => [|j _]; rewrite coefC !simp.
+Qed.
+
+Lemma coefMC c p i : (p * c%:P)`_i = p`_i * c.
+Proof.
+rewrite coefMr big_ord_recl subn0.
+by rewrite big1 => [|j _]; rewrite coefC !simp.
+Qed.
+
+Lemma polyC_mul : {morph polyC : a b / a * b}.
+Proof. by move=> a b; apply/polyP=> [[|i]]; rewrite coefCM !coefC ?simp. Qed.
+
+Fact polyC_multiplicative : multiplicative polyC.
+Proof. by split; first exact: polyC_mul. Qed.
+Canonical polyC_rmorphism := AddRMorphism polyC_multiplicative.
+
+Lemma polyC_exp n : {morph polyC : c / c ^+ n}.
+Proof. exact: rmorphX. Qed.
+
+Lemma size_exp_leq p n : size (p ^+ n) <= ((size p).-1 * n).+1.
+Proof.
+elim: n => [|n IHn]; first by rewrite size_poly1.
+have [-> | nzp] := poly0Vpos p; first by rewrite exprS mul0r size_poly0.
+rewrite exprS (leq_trans (size_mul_leq _ _)) //.
+by rewrite -{1}(prednK nzp) mulnS -addnS leq_add2l.
+Qed.
+
+Lemma size_Msign p n : size ((-1) ^+ n * p) = size p.
+Proof.
+by rewrite -signr_odd; case: (odd n); rewrite ?mul1r // mulN1r size_opp.
+Qed.
+
+Fact coefp0_multiplicative : multiplicative (coefp 0 : {poly R} -> R).
+Proof.
+split=> [p q|]; last by rewrite polyCK.
+by rewrite [coefp 0 _]coefM big_ord_recl big_ord0 addr0.
+Qed.
+
+Canonical coefp0_rmorphism := AddRMorphism coefp0_multiplicative.
+
+(* Algebra structure of polynomials. *)
+Definition scale_poly_def a (p : {poly R}) := \poly_(i < size p) (a * p`_i).
+Fact scale_poly_key : unit. Proof. by []. Qed.
+Definition scale_poly := locked_with scale_poly_key scale_poly_def.
+Canonical scale_poly_unlockable := [unlockable fun scale_poly].
+
+Fact scale_polyE a p : scale_poly a p = a%:P * p.
+Proof.
+apply/polyP=> n; rewrite unlock coef_poly coefCM.
+by case: leqP => // le_p_n; rewrite nth_default ?mulr0.
+Qed.
+
+Fact scale_polyA a b p : scale_poly a (scale_poly b p) = scale_poly (a * b) p.
+Proof. by rewrite !scale_polyE mulrA polyC_mul. Qed.
+
+Fact scale_1poly : left_id 1 scale_poly.
+Proof. by move=> p; rewrite scale_polyE mul1r. Qed.
+
+Fact scale_polyDr a : {morph scale_poly a : p q / p + q}.
+Proof. by move=> p q; rewrite !scale_polyE mulrDr. Qed.
+
+Fact scale_polyDl p : {morph scale_poly^~ p : a b / a + b}.
+Proof. by move=> a b /=; rewrite !scale_polyE raddfD mulrDl. Qed.
+
+Fact scale_polyAl a p q : scale_poly a (p * q) = scale_poly a p * q.
+Proof. by rewrite !scale_polyE mulrA. Qed.
+
+Definition poly_lmodMixin :=
+ LmodMixin scale_polyA scale_1poly scale_polyDr scale_polyDl.
+
+Canonical poly_lmodType :=
+ Eval hnf in LmodType R {poly R} poly_lmodMixin.
+Canonical polynomial_lmodType :=
+ Eval hnf in LmodType R (polynomial R) poly_lmodMixin.
+Canonical poly_lalgType :=
+ Eval hnf in LalgType R {poly R} scale_polyAl.
+Canonical polynomial_lalgType :=
+ Eval hnf in LalgType R (polynomial R) scale_polyAl.
+
+Lemma mul_polyC a p : a%:P * p = a *: p.
+Proof. by rewrite -scale_polyE. Qed.
+
+Lemma alg_polyC a : a%:A = a%:P :> {poly R}.
+Proof. by rewrite -mul_polyC mulr1. Qed.
+
+Lemma coefZ a p i : (a *: p)`_i = a * p`_i.
+Proof.
+rewrite -[*:%R]/scale_poly unlock coef_poly.
+by case: leqP => // le_p_n; rewrite nth_default ?mulr0.
+Qed.
+
+Lemma size_scale_leq a p : size (a *: p) <= size p.
+Proof. by rewrite -[*:%R]/scale_poly unlock size_poly. Qed.
+
+Canonical coefp_linear i : {scalar {poly R}} :=
+ AddLinear ((fun a => (coefZ a) ^~ i) : scalable_for *%R (coefp i)).
+Canonical coefp0_lrmorphism := [lrmorphism of coefp 0].
+
+(* The indeterminate, at last! *)
+Definition polyX_def := Poly [:: 0; 1].
+Fact polyX_key : unit. Proof. by []. Qed.
+Definition polyX : {poly R} := locked_with polyX_key polyX_def.
+Canonical polyX_unlockable := [unlockable of polyX].
+Local Notation "'X" := polyX.
+
+Lemma polyseqX : 'X = [:: 0; 1] :> seq R.
+Proof. by rewrite unlock !polyseq_cons nil_poly eqxx /= polyseq1. Qed.
+
+Lemma size_polyX : size 'X = 2. Proof. by rewrite polyseqX. Qed.
+
+Lemma polyX_eq0 : ('X == 0) = false.
+Proof. by rewrite -size_poly_eq0 size_polyX. Qed.
+
+Lemma coefX i : 'X`_i = (i == 1%N)%:R.
+Proof. by case: i => [|[|i]]; rewrite polyseqX //= nth_nil. Qed.
+
+Lemma lead_coefX : lead_coef 'X = 1.
+Proof. by rewrite /lead_coef polyseqX. Qed.
+
+Lemma commr_polyX p : GRing.comm p 'X.
+Proof.
+apply/polyP=> i; rewrite coefMr coefM.
+by apply: eq_bigr => j _; rewrite coefX commr_nat.
+Qed.
+
+Lemma coefMX p i : (p * 'X)`_i = (if (i == 0)%N then 0 else p`_i.-1).
+Proof.
+rewrite coefMr big_ord_recl coefX ?simp.
+case: i => [|i]; rewrite ?big_ord0 //= big_ord_recl polyseqX subn1 /=.
+by rewrite big1 ?simp // => j _; rewrite nth_nil !simp.
+Qed.
+
+Lemma coefXM p i : ('X * p)`_i = (if (i == 0)%N then 0 else p`_i.-1).
+Proof. by rewrite -commr_polyX coefMX. Qed.
+
+Lemma cons_poly_def p a : cons_poly a p = p * 'X + a%:P.
+Proof.
+apply/polyP=> i; rewrite coef_cons coefD coefMX coefC.
+by case: ifP; rewrite !simp.
+Qed.
+
+Lemma poly_ind (K : {poly R} -> Type) :
+ K 0 -> (forall p c, K p -> K (p * 'X + c%:P)) -> (forall p, K p).
+Proof.
+move=> K0 Kcons p; rewrite -[p]polyseqK.
+elim: {p}(p : seq R) => //= p c IHp; rewrite cons_poly_def; exact: Kcons.
+Qed.
+
+Lemma polyseqXsubC a : 'X - a%:P = [:: - a; 1] :> seq R.
+Proof.
+by rewrite -['X]mul1r -polyC_opp -cons_poly_def polyseq_cons polyseq1.
+Qed.
+
+Lemma size_XsubC a : size ('X - a%:P) = 2%N.
+Proof. by rewrite polyseqXsubC. Qed.
+
+Lemma size_XaddC b : size ('X + b%:P) = 2.
+Proof. by rewrite -[b]opprK rmorphN size_XsubC. Qed.
+
+Lemma lead_coefXsubC a : lead_coef ('X - a%:P) = 1.
+Proof. by rewrite lead_coefE polyseqXsubC. Qed.
+
+Lemma polyXsubC_eq0 a : ('X - a%:P == 0) = false.
+Proof. by rewrite -nil_poly polyseqXsubC. Qed.
+
+Lemma size_MXaddC p c :
+ size (p * 'X + c%:P) = (if (p == 0) && (c == 0) then 0%N else (size p).+1).
+Proof. by rewrite -cons_poly_def size_cons_poly nil_poly. Qed.
+
+Lemma polyseqMX p : p != 0 -> p * 'X = 0 :: p :> seq R.
+Proof.
+by move=> nz_p; rewrite -[p * _]addr0 -cons_poly_def polyseq_cons nil_poly nz_p.
+Qed.
+
+Lemma size_mulX p : p != 0 -> size (p * 'X) = (size p).+1.
+Proof. by move/polyseqMX->. Qed.
+
+Lemma lead_coefMX p : lead_coef (p * 'X) = lead_coef p.
+Proof.
+have [-> | nzp] := eqVneq p 0; first by rewrite mul0r.
+by rewrite /lead_coef !nth_last polyseqMX.
+Qed.
+
+Lemma size_XmulC a : a != 0 -> size ('X * a%:P) = 2.
+Proof.
+by move=> nz_a; rewrite -commr_polyX size_mulX ?polyC_eq0 ?size_polyC nz_a.
+Qed.
+
+Local Notation "''X^' n" := ('X ^+ n).
+
+Lemma coefXn n i : 'X^n`_i = (i == n)%:R.
+Proof.
+by elim: n i => [|n IHn] [|i]; rewrite ?coef1 // exprS coefXM ?IHn.
+Qed.
+
+Lemma polyseqXn n : 'X^n = rcons (nseq n 0) 1 :> seq R.
+Proof.
+elim: n => [|n IHn]; rewrite ?polyseq1 // exprSr.
+by rewrite polyseqMX -?size_poly_eq0 IHn ?size_rcons.
+Qed.
+
+Lemma size_polyXn n : size 'X^n = n.+1.
+Proof. by rewrite polyseqXn size_rcons size_nseq. Qed.
+
+Lemma commr_polyXn p n : GRing.comm p 'X^n.
+Proof. by apply: commrX; exact: commr_polyX. Qed.
+
+Lemma lead_coefXn n : lead_coef 'X^n = 1.
+Proof. by rewrite /lead_coef nth_last polyseqXn last_rcons. Qed.
+
+Lemma polyseqMXn n p : p != 0 -> p * 'X^n = ncons n 0 p :> seq R.
+Proof.
+case: n => [|n] nz_p; first by rewrite mulr1.
+elim: n => [|n IHn]; first exact: polyseqMX.
+by rewrite exprSr mulrA polyseqMX -?nil_poly IHn.
+Qed.
+
+Lemma coefMXn n p i : (p * 'X^n)`_i = if i < n then 0 else p`_(i - n).
+Proof.
+have [-> | /polyseqMXn->] := eqVneq p 0; last exact: nth_ncons.
+by rewrite mul0r !coef0 if_same.
+Qed.
+
+Lemma coefXnM n p i : ('X^n * p)`_i = if i < n then 0 else p`_(i - n).
+Proof. by rewrite -commr_polyXn coefMXn. Qed.
+
+(* Expansion of a polynomial as an indexed sum *)
+Lemma poly_def n E : \poly_(i < n) E i = \sum_(i < n) E i *: 'X^i.
+Proof.
+rewrite unlock; elim: n => [|n IHn] in E *; first by rewrite big_ord0.
+rewrite big_ord_recl /= cons_poly_def addrC expr0 alg_polyC.
+congr (_ + _); rewrite (iota_addl 1 0) -map_comp IHn big_distrl /=.
+by apply: eq_bigr => i _; rewrite -scalerAl exprSr.
+Qed.
+
+(* Monic predicate *)
+Definition monic := [qualify p | lead_coef p == 1].
+Fact monic_key : pred_key monic. Proof. by []. Qed.
+Canonical monic_keyed := KeyedQualifier monic_key.
+
+Lemma monicE p : (p \is monic) = (lead_coef p == 1). Proof. by []. Qed.
+Lemma monicP p : reflect (lead_coef p = 1) (p \is monic).
+Proof. exact: eqP. Qed.
+
+Lemma monic1 : 1 \is monic. Proof. exact/eqP/lead_coef1. Qed.
+Lemma monicX : 'X \is monic. Proof. exact/eqP/lead_coefX. Qed.
+Lemma monicXn n : 'X^n \is monic. Proof. exact/eqP/lead_coefXn. Qed.
+
+Lemma monic_neq0 p : p \is monic -> p != 0.
+Proof. by rewrite -lead_coef_eq0 => /eqP->; exact: oner_neq0. Qed.
+
+Lemma lead_coef_monicM p q : p \is monic -> lead_coef (p * q) = lead_coef q.
+Proof.
+have [-> | nz_q] := eqVneq q 0; first by rewrite mulr0.
+by move/monicP=> mon_p; rewrite lead_coef_proper_mul mon_p mul1r ?lead_coef_eq0.
+Qed.
+
+Lemma lead_coef_Mmonic p q : q \is monic -> lead_coef (p * q) = lead_coef p.
+Proof.
+have [-> | nz_p] := eqVneq p 0; first by rewrite mul0r.
+by move/monicP=> mon_q; rewrite lead_coef_proper_mul mon_q mulr1 ?lead_coef_eq0.
+Qed.
+
+Lemma size_monicM p q :
+ p \is monic -> q != 0 -> size (p * q) = (size p + size q).-1.
+Proof.
+move/monicP=> mon_p nz_q.
+by rewrite size_proper_mul // mon_p mul1r lead_coef_eq0.
+Qed.
+
+Lemma size_Mmonic p q :
+ p != 0 -> q \is monic -> size (p * q) = (size p + size q).-1.
+Proof.
+move=> nz_p /monicP mon_q.
+by rewrite size_proper_mul // mon_q mulr1 lead_coef_eq0.
+Qed.
+
+Lemma monicMl p q : p \is monic -> (p * q \is monic) = (q \is monic).
+Proof. by move=> mon_p; rewrite !monicE lead_coef_monicM. Qed.
+
+Lemma monicMr p q : q \is monic -> (p * q \is monic) = (p \is monic).
+Proof. by move=> mon_q; rewrite !monicE lead_coef_Mmonic. Qed.
+
+Fact monic_mulr_closed : mulr_closed monic.
+Proof. by split=> [|p q mon_p]; rewrite (monic1, monicMl). Qed.
+Canonical monic_mulrPred := MulrPred monic_mulr_closed.
+
+Lemma monic_exp p n : p \is monic -> p ^+ n \is monic.
+Proof. exact: rpredX. Qed.
+
+Lemma monic_prod I rI (P : pred I) (F : I -> {poly R}):
+ (forall i, P i -> F i \is monic) -> \prod_(i <- rI | P i) F i \is monic.
+Proof. exact: rpred_prod. Qed.
+
+Lemma monicXsubC c : 'X - c%:P \is monic.
+Proof. exact/eqP/lead_coefXsubC. Qed.
+
+Lemma monic_prod_XsubC I rI (P : pred I) (F : I -> R) :
+ \prod_(i <- rI | P i) ('X - (F i)%:P) \is monic.
+Proof. by apply: monic_prod => i _; exact: monicXsubC. Qed.
+
+Lemma size_prod_XsubC I rI (F : I -> R) :
+ size (\prod_(i <- rI) ('X - (F i)%:P)) = (size rI).+1.
+Proof.
+elim: rI => [|i r /= <-]; rewrite ?big_nil ?size_poly1 // big_cons.
+rewrite size_monicM ?monicXsubC ?monic_neq0 ?monic_prod_XsubC //.
+by rewrite size_XsubC.
+Qed.
+
+Lemma size_exp_XsubC n a : size (('X - a%:P) ^+ n) = n.+1.
+Proof. by rewrite -[n]card_ord -prodr_const size_prod_XsubC cardE enumT. Qed.
+
+(* Some facts about regular elements. *)
+
+Lemma lreg_lead p : GRing.lreg (lead_coef p) -> GRing.lreg p.
+Proof.
+move/mulrI_eq0=> reg_p; apply: mulrI0_lreg => q /eqP; apply: contraTeq => nz_q.
+by rewrite -lead_coef_eq0 lead_coef_proper_mul reg_p lead_coef_eq0.
+Qed.
+
+Lemma rreg_lead p : GRing.rreg (lead_coef p) -> GRing.rreg p.
+Proof.
+move/mulIr_eq0=> reg_p; apply: mulIr0_rreg => q /eqP; apply: contraTeq => nz_q.
+by rewrite -lead_coef_eq0 lead_coef_proper_mul reg_p lead_coef_eq0.
+Qed.
+
+Lemma lreg_lead0 p : GRing.lreg (lead_coef p) -> p != 0.
+Proof. by move/lreg_neq0; rewrite lead_coef_eq0. Qed.
+
+Lemma rreg_lead0 p : GRing.rreg (lead_coef p) -> p != 0.
+Proof. by move/rreg_neq0; rewrite lead_coef_eq0. Qed.
+
+Lemma lreg_size c p : GRing.lreg c -> size (c *: p) = size p.
+Proof.
+move=> reg_c; have [-> | nz_p] := eqVneq p 0; first by rewrite scaler0.
+rewrite -mul_polyC size_proper_mul; first by rewrite size_polyC lreg_neq0.
+by rewrite lead_coefC mulrI_eq0 ?lead_coef_eq0.
+Qed.
+
+Lemma lreg_polyZ_eq0 c p : GRing.lreg c -> (c *: p == 0) = (p == 0).
+Proof. by rewrite -!size_poly_eq0 => /lreg_size->. Qed.
+
+Lemma lead_coef_lreg c p :
+ GRing.lreg c -> lead_coef (c *: p) = c * lead_coef p.
+Proof. by move=> reg_c; rewrite !lead_coefE coefZ lreg_size. Qed.
+
+Lemma rreg_size c p : GRing.rreg c -> size (p * c%:P) = size p.
+Proof.
+move=> reg_c; have [-> | nz_p] := eqVneq p 0; first by rewrite mul0r.
+rewrite size_proper_mul; first by rewrite size_polyC rreg_neq0 ?addn1.
+by rewrite lead_coefC mulIr_eq0 ?lead_coef_eq0.
+Qed.
+
+Lemma rreg_polyMC_eq0 c p : GRing.rreg c -> (p * c%:P == 0) = (p == 0).
+Proof. by rewrite -!size_poly_eq0 => /rreg_size->. Qed.
+
+Lemma rreg_div0 q r d :
+ GRing.rreg (lead_coef d) -> size r < size d ->
+ (q * d + r == 0) = (q == 0) && (r == 0).
+Proof.
+move=> reg_d lt_r_d; rewrite addrC addr_eq0.
+have [-> | nz_q] := altP (q =P 0); first by rewrite mul0r oppr0.
+apply: contraTF lt_r_d => /eqP->; rewrite -leqNgt size_opp.
+rewrite size_proper_mul ?mulIr_eq0 ?lead_coef_eq0 //.
+by rewrite (polySpred nz_q) leq_addl.
+Qed.
+
+Lemma monic_comreg p :
+ p \is monic -> GRing.comm p (lead_coef p)%:P /\ GRing.rreg (lead_coef p).
+Proof. by move/monicP->; split; [exact: commr1 | exact: rreg1]. Qed.
+
+(* Horner evaluation of polynomials *)
+Implicit Types s rs : seq R.
+Fixpoint horner_rec s x := if s is a :: s' then horner_rec s' x * x + a else 0.
+Definition horner p := horner_rec p.
+
+Local Notation "p .[ x ]" := (horner p x) : ring_scope.
+
+Lemma horner0 x : (0 : {poly R}).[x] = 0.
+Proof. by rewrite /horner polyseq0. Qed.
+
+Lemma hornerC c x : (c%:P).[x] = c.
+Proof. by rewrite /horner polyseqC; case: eqP; rewrite /= ?simp. Qed.
+
+Lemma hornerX x : 'X.[x] = x.
+Proof. by rewrite /horner polyseqX /= !simp. Qed.
+
+Lemma horner_cons p c x : (cons_poly c p).[x] = p.[x] * x + c.
+Proof.
+rewrite /horner polyseq_cons; case: nilP => //= ->.
+by rewrite !simp -/(_.[x]) hornerC.
+Qed.
+
+Lemma horner_coef0 p : p.[0] = p`_0.
+Proof. by rewrite /horner; case: (p : seq R) => //= c p'; rewrite !simp. Qed.
+
+Lemma hornerMXaddC p c x : (p * 'X + c%:P).[x] = p.[x] * x + c.
+Proof. by rewrite -cons_poly_def horner_cons. Qed.
+
+Lemma hornerMX p x : (p * 'X).[x] = p.[x] * x.
+Proof. by rewrite -[p * 'X]addr0 hornerMXaddC addr0. Qed.
+
+Lemma horner_Poly s x : (Poly s).[x] = horner_rec s x.
+Proof. by elim: s => [|a s /= <-]; rewrite (horner0, horner_cons). Qed.
+
+Lemma horner_coef p x : p.[x] = \sum_(i < size p) p`_i * x ^+ i.
+Proof.
+rewrite /horner.
+elim: {p}(p : seq R) => /= [|a s ->]; first by rewrite big_ord0.
+rewrite big_ord_recl simp addrC big_distrl /=.
+by congr (_ + _); apply: eq_bigr => i _; rewrite -mulrA exprSr.
+Qed.
+
+Lemma horner_coef_wide n p x :
+ size p <= n -> p.[x] = \sum_(i < n) p`_i * x ^+ i.
+Proof.
+move=> le_p_n.
+rewrite horner_coef (big_ord_widen n (fun i => p`_i * x ^+ i)) // big_mkcond.
+by apply: eq_bigr => i _; case: ltnP => // le_p_i; rewrite nth_default ?simp.
+Qed.
+
+Lemma horner_poly n E x : (\poly_(i < n) E i).[x] = \sum_(i < n) E i * x ^+ i.
+Proof.
+rewrite (@horner_coef_wide n) ?size_poly //.
+by apply: eq_bigr => i _; rewrite coef_poly ltn_ord.
+Qed.
+
+Lemma hornerN p x : (- p).[x] = - p.[x].
+Proof.
+rewrite -[-%R]/opp_poly unlock horner_poly horner_coef -sumrN /=.
+by apply: eq_bigr => i _; rewrite mulNr.
+Qed.
+
+Lemma hornerD p q x : (p + q).[x] = p.[x] + q.[x].
+Proof.
+rewrite -[+%R]/add_poly unlock horner_poly; set m := maxn _ _.
+rewrite !(@horner_coef_wide m) ?leq_max ?leqnn ?orbT // -big_split /=.
+by apply: eq_bigr => i _; rewrite -mulrDl.
+Qed.
+
+Lemma hornerXsubC a x : ('X - a%:P).[x] = x - a.
+Proof. by rewrite hornerD hornerN hornerC hornerX. Qed.
+
+Lemma horner_sum I (r : seq I) (P : pred I) F x :
+ (\sum_(i <- r | P i) F i).[x] = \sum_(i <- r | P i) (F i).[x].
+Proof. by elim/big_rec2: _ => [|i _ p _ <-]; rewrite (horner0, hornerD). Qed.
+
+Lemma hornerCM a p x : (a%:P * p).[x] = a * p.[x].
+Proof.
+elim/poly_ind: p => [|p c IHp]; first by rewrite !(mulr0, horner0).
+by rewrite mulrDr mulrA -polyC_mul !hornerMXaddC IHp mulrDr mulrA.
+Qed.
+
+Lemma hornerZ c p x : (c *: p).[x] = c * p.[x].
+Proof. by rewrite -mul_polyC hornerCM. Qed.
+
+Lemma hornerMn n p x : (p *+ n).[x] = p.[x] *+ n.
+Proof. by elim: n => [| n IHn]; rewrite ?horner0 // !mulrS hornerD IHn. Qed.
+
+Definition comm_coef p x := forall i, p`_i * x = x * p`_i.
+
+Definition comm_poly p x := x * p.[x] = p.[x] * x.
+
+Lemma comm_coef_poly p x : comm_coef p x -> comm_poly p x.
+Proof.
+move=> cpx; rewrite /comm_poly !horner_coef big_distrl big_distrr /=.
+by apply: eq_bigr => i _; rewrite /= mulrA -cpx -!mulrA commrX.
+Qed.
+
+Lemma comm_poly0 x : comm_poly 0 x.
+Proof. by rewrite /comm_poly !horner0 !simp. Qed.
+
+Lemma comm_poly1 x : comm_poly 1 x.
+Proof. by rewrite /comm_poly !hornerC !simp. Qed.
+
+Lemma comm_polyX x : comm_poly 'X x.
+Proof. by rewrite /comm_poly !hornerX. Qed.
+
+Lemma hornerM_comm p q x : comm_poly q x -> (p * q).[x] = p.[x] * q.[x].
+Proof.
+move=> comm_qx.
+elim/poly_ind: p => [|p c IHp]; first by rewrite !(simp, horner0).
+rewrite mulrDl hornerD hornerCM -mulrA -commr_polyX mulrA hornerMX.
+by rewrite {}IHp -mulrA -comm_qx mulrA -mulrDl hornerMXaddC.
+Qed.
+
+Lemma horner_exp_comm p x n : comm_poly p x -> (p ^+ n).[x] = p.[x] ^+ n.
+Proof.
+move=> comm_px; elim: n => [|n IHn]; first by rewrite hornerC.
+by rewrite !exprSr -IHn hornerM_comm.
+Qed.
+
+Lemma hornerXn x n : ('X^n).[x] = x ^+ n.
+Proof. by rewrite horner_exp_comm /comm_poly hornerX. Qed.
+
+Definition hornerE_comm :=
+ (hornerD, hornerN, hornerX, hornerC, horner_cons,
+ simp, hornerCM, hornerZ,
+ (fun p x => hornerM_comm p (comm_polyX x))).
+
+Definition root p : pred R := fun x => p.[x] == 0.
+
+Lemma mem_root p x : x \in root p = (p.[x] == 0).
+Proof. by []. Qed.
+
+Lemma rootE p x : (root p x = (p.[x] == 0)) * ((x \in root p) = (p.[x] == 0)).
+Proof. by []. Qed.
+
+Lemma rootP p x : reflect (p.[x] = 0) (root p x).
+Proof. exact: eqP. Qed.
+
+Lemma rootPt p x : reflect (p.[x] == 0) (root p x).
+Proof. exact: idP. Qed.
+
+Lemma rootPf p x : reflect ((p.[x] == 0) = false) (~~ root p x).
+Proof. exact: negPf. Qed.
+
+Lemma rootC a x : root a%:P x = (a == 0).
+Proof. by rewrite rootE hornerC. Qed.
+
+Lemma root0 x : root 0 x.
+Proof. by rewrite rootC. Qed.
+
+Lemma root1 x : ~~ root 1 x.
+Proof. by rewrite rootC oner_eq0. Qed.
+
+Lemma rootX x : root 'X x = (x == 0).
+Proof. by rewrite rootE hornerX. Qed.
+
+Lemma rootN p x : root (- p) x = root p x.
+Proof. by rewrite rootE hornerN oppr_eq0. Qed.
+
+Lemma root_size_gt1 a p : p != 0 -> root p a -> 1 < size p.
+Proof.
+rewrite ltnNge => nz_p; apply: contraL => /size1_polyC Dp.
+by rewrite Dp rootC -polyC_eq0 -Dp.
+Qed.
+
+Lemma root_XsubC a x : root ('X - a%:P) x = (x == a).
+Proof. by rewrite rootE hornerXsubC subr_eq0. Qed.
+
+Lemma root_XaddC a x : root ('X + a%:P) x = (x == - a).
+Proof. by rewrite -root_XsubC rmorphN opprK. Qed.
+
+Theorem factor_theorem p a : reflect (exists q, p = q * ('X - a%:P)) (root p a).
+Proof.
+apply: (iffP eqP) => [pa0 | [q ->]]; last first.
+ by rewrite hornerM_comm /comm_poly hornerXsubC subrr ?simp.
+exists (\poly_(i < size p) horner_rec (drop i.+1 p) a).
+apply/polyP=> i; rewrite mulrBr coefB coefMX coefMC !coef_poly.
+apply: canRL (addrK _) _; rewrite addrC; have [le_p_i | lt_i_p] := leqP.
+ rewrite nth_default // !simp drop_oversize ?if_same //.
+ exact: leq_trans (leqSpred _).
+case: i => [|i] in lt_i_p *; last by rewrite ltnW // (drop_nth 0 lt_i_p).
+by rewrite drop1 /= -{}pa0 /horner; case: (p : seq R) lt_i_p.
+Qed.
+
+Lemma multiplicity_XsubC p a :
+ {m | exists2 q, (p != 0) ==> ~~ root q a & p = q * ('X - a%:P) ^+ m}.
+Proof.
+elim: {p}(size p) {-2}p (eqxx (size p)) => [|n IHn] p.
+ by rewrite size_poly_eq0 => ->; exists 0%N, p; rewrite ?mulr1.
+have [/sig_eqW[{p}p ->] sz_p | nz_pa] := altP (factor_theorem p a); last first.
+ by exists 0%N, p; rewrite ?mulr1 ?nz_pa ?implybT.
+have nz_p: p != 0 by apply: contraTneq sz_p => ->; rewrite mul0r size_poly0.
+rewrite size_Mmonic ?monicXsubC // size_XsubC addn2 eqSS in sz_p.
+have [m /sig2_eqW[q nz_qa Dp]] := IHn p sz_p; rewrite nz_p /= in nz_qa.
+by exists m.+1, q; rewrite ?nz_qa ?implybT // exprSr mulrA -Dp.
+Qed.
+
+(* Roots of unity. *)
+
+Lemma size_Xn_sub_1 n : n > 0 -> size ('X^n - 1 : {poly R}) = n.+1.
+Proof.
+by move=> n_gt0; rewrite size_addl size_polyXn // size_opp size_poly1.
+Qed.
+
+Lemma monic_Xn_sub_1 n : n > 0 -> 'X^n - 1 \is monic.
+Proof.
+move=> n_gt0; rewrite monicE lead_coefE size_Xn_sub_1 // coefB.
+by rewrite coefXn coef1 eqxx eqn0Ngt n_gt0 subr0.
+Qed.
+
+Definition root_of_unity n : pred R := root ('X^n - 1).
+Local Notation "n .-unity_root" := (root_of_unity n) : ring_scope.
+
+Lemma unity_rootE n z : n.-unity_root z = (z ^+ n == 1).
+Proof.
+by rewrite /root_of_unity rootE hornerD hornerN hornerXn hornerC subr_eq0.
+Qed.
+
+Lemma unity_rootP n z : reflect (z ^+ n = 1) (n.-unity_root z).
+Proof. by rewrite unity_rootE; exact: eqP. Qed.
+
+Definition primitive_root_of_unity n z :=
+ (n > 0) && [forall i : 'I_n, i.+1.-unity_root z == (i.+1 == n)].
+Local Notation "n .-primitive_root" := (primitive_root_of_unity n) : ring_scope.
+
+Lemma prim_order_exists n z :
+ n > 0 -> z ^+ n = 1 -> {m | m.-primitive_root z & (m %| n)}.
+Proof.
+move=> n_gt0 zn1.
+have: exists m, (m > 0) && (z ^+ m == 1) by exists n; rewrite n_gt0 /= zn1.
+case/ex_minnP=> m /andP[m_gt0 /eqP zm1] m_min.
+exists m.
+ apply/andP; split=> //; apply/eqfunP=> [[i]] /=.
+ rewrite leq_eqVlt unity_rootE.
+ case: eqP => [-> _ | _]; first by rewrite zm1 eqxx.
+ by apply: contraTF => zi1; rewrite -leqNgt m_min.
+have: n %% m < m by rewrite ltn_mod.
+apply: contraLR; rewrite -lt0n -leqNgt => nm_gt0; apply: m_min.
+by rewrite nm_gt0 /= expr_mod ?zn1.
+Qed.
+
+Section OnePrimitive.
+
+Variables (n : nat) (z : R).
+Hypothesis prim_z : n.-primitive_root z.
+
+Lemma prim_order_gt0 : n > 0. Proof. by case/andP: prim_z. Qed.
+Let n_gt0 := prim_order_gt0.
+
+Lemma prim_expr_order : z ^+ n = 1.
+Proof.
+case/andP: prim_z => _; rewrite -(prednK n_gt0) => /forallP/(_ ord_max).
+by rewrite unity_rootE eqxx eqb_id => /eqP.
+Qed.
+
+Lemma prim_expr_mod i : z ^+ (i %% n) = z ^+ i.
+Proof. exact: expr_mod prim_expr_order. Qed.
+
+Lemma prim_order_dvd i : (n %| i) = (z ^+ i == 1).
+Proof.
+move: n_gt0; rewrite -prim_expr_mod /dvdn -(ltn_mod i).
+case: {i}(i %% n)%N => [|i] lt_i; first by rewrite !eqxx.
+case/andP: prim_z => _ /forallP/(_ (Ordinal (ltnW lt_i))).
+by move/eqP; rewrite unity_rootE eqn_leq andbC leqNgt lt_i.
+Qed.
+
+Lemma eq_prim_root_expr i j : (z ^+ i == z ^+ j) = (i == j %[mod n]).
+Proof.
+wlog le_ji: i j / j <= i.
+ move=> IH; case: (leqP j i); last move/ltnW; move/IH=> //.
+ by rewrite eq_sym (eq_sym (j %% n)%N).
+rewrite -{1}(subnKC le_ji) exprD -prim_expr_mod eqn_mod_dvd //.
+rewrite prim_order_dvd; apply/eqP/eqP=> [|->]; last by rewrite mulr1.
+move/(congr1 ( *%R (z ^+ (n - j %% n)))); rewrite mulrA -exprD.
+by rewrite subnK ?prim_expr_order ?mul1r // ltnW ?ltn_mod.
+Qed.
+
+Lemma exp_prim_root k : (n %/ gcdn k n).-primitive_root (z ^+ k).
+Proof.
+set d := gcdn k n; have d_gt0: (0 < d)%N by rewrite gcdn_gt0 orbC n_gt0.
+have [d_dv_k d_dv_n]: (d %| k /\ d %| n)%N by rewrite dvdn_gcdl dvdn_gcdr.
+set q := (n %/ d)%N; rewrite /q.-primitive_root ltn_divRL // n_gt0.
+apply/forallP=> i; rewrite unity_rootE -exprM -prim_order_dvd.
+rewrite -(divnK d_dv_n) -/q -(divnK d_dv_k) mulnAC dvdn_pmul2r //.
+apply/eqP; apply/idP/idP=> [|/eqP->]; last by rewrite dvdn_mull.
+rewrite Gauss_dvdr; first by rewrite eqn_leq ltn_ord; exact: dvdn_leq.
+by rewrite /coprime gcdnC -(eqn_pmul2r d_gt0) mul1n muln_gcdl !divnK.
+Qed.
+
+Lemma dvdn_prim_root m : (m %| n)%N -> m.-primitive_root (z ^+ (n %/ m)).
+Proof.
+set k := (n %/ m)%N => m_dv_n; rewrite -{1}(mulKn m n_gt0) -divnA // -/k.
+by rewrite -{1}(@gcdn_idPl k n _) ?exp_prim_root // -(divnK m_dv_n) dvdn_mulr.
+Qed.
+
+End OnePrimitive.
+
+Lemma prim_root_exp_coprime n z k :
+ n.-primitive_root z -> n.-primitive_root (z ^+ k) = coprime k n.
+Proof.
+move=> prim_z;have n_gt0 := prim_order_gt0 prim_z.
+apply/idP/idP=> [prim_zk | co_k_n].
+ set d := gcdn k n; have dv_d_n: (d %| n)%N := dvdn_gcdr _ _.
+ rewrite /coprime -/d -(eqn_pmul2r n_gt0) mul1n -{2}(gcdnMl n d).
+ rewrite -{2}(divnK dv_d_n) (mulnC _ d) -muln_gcdr (gcdn_idPr _) //.
+ rewrite (prim_order_dvd prim_zk) -exprM -(prim_order_dvd prim_z).
+ by rewrite muln_divCA_gcd dvdn_mulr.
+have zkn_1: z ^+ k ^+ n = 1 by rewrite exprAC (prim_expr_order prim_z) expr1n.
+have{zkn_1} [m prim_zk dv_m_n]:= prim_order_exists n_gt0 zkn_1.
+suffices /eqP <-: m == n by [].
+rewrite eqn_dvd dv_m_n -(@Gauss_dvdr n k m) 1?coprime_sym //=.
+by rewrite (prim_order_dvd prim_z) exprM (prim_expr_order prim_zk).
+Qed.
+
+(* Lifting a ring predicate to polynomials. *)
+
+Definition polyOver (S : pred_class) :=
+ [qualify a p : {poly R} | all (mem S) p].
+
+Fact polyOver_key S : pred_key (polyOver S). Proof. by []. Qed.
+Canonical polyOver_keyed S := KeyedQualifier (polyOver_key S).
+
+Lemma polyOverS (S1 S2 : pred_class) :
+ {subset S1 <= S2} -> {subset polyOver S1 <= polyOver S2}.
+Proof.
+by move=> sS12 p /(all_nthP 0)S1p; apply/(all_nthP 0)=> i /S1p; apply: sS12.
+Qed.
+
+Lemma polyOver0 S : 0 \is a polyOver S.
+Proof. by rewrite qualifE polyseq0. Qed.
+
+Lemma polyOver_poly (S : pred_class) n E :
+ (forall i, i < n -> E i \in S) -> \poly_(i < n) E i \is a polyOver S.
+Proof.
+move=> S_E; apply/(all_nthP 0)=> i lt_i_p /=; rewrite coef_poly.
+by case: ifP => [/S_E// | /idP[]]; apply: leq_trans lt_i_p (size_poly n E).
+Qed.
+
+Section PolyOverAdd.
+
+Variables (S : predPredType R) (addS : addrPred S) (kS : keyed_pred addS).
+
+Lemma polyOverP {p} : reflect (forall i, p`_i \in kS) (p \in polyOver kS).
+Proof.
+apply: (iffP (all_nthP 0)) => [Sp i | Sp i _]; last exact: Sp.
+by have [/Sp // | /(nth_default 0)->] := ltnP i (size p); apply: rpred0.
+Qed.
+
+Lemma polyOverC c : (c%:P \in polyOver kS) = (c \in kS).
+Proof.
+by rewrite qualifE polyseqC; case: eqP => [->|] /=; rewrite ?andbT ?rpred0.
+Qed.
+
+Fact polyOver_addr_closed : addr_closed (polyOver kS).
+Proof.
+split=> [|p q Sp Sq]; first exact: polyOver0.
+by apply/polyOverP=> i; rewrite coefD rpredD ?(polyOverP _).
+Qed.
+Canonical polyOver_addrPred := AddrPred polyOver_addr_closed.
+
+End PolyOverAdd.
+
+Fact polyOverNr S (addS : zmodPred S) (kS : keyed_pred addS) :
+ oppr_closed (polyOver kS).
+Proof.
+by move=> p /polyOverP Sp; apply/polyOverP=> i; rewrite coefN rpredN.
+Qed.
+Canonical polyOver_opprPred S addS kS := OpprPred (@polyOverNr S addS kS).
+Canonical polyOver_zmodPred S addS kS := ZmodPred (@polyOverNr S addS kS).
+
+Section PolyOverSemiring.
+
+Context (S : pred_class) (ringS : @semiringPred R S) (kS : keyed_pred ringS).
+
+Fact polyOver_mulr_closed : mulr_closed (polyOver kS).
+Proof.
+split=> [|p q /polyOverP Sp /polyOverP Sq]; first by rewrite polyOverC rpred1.
+by apply/polyOverP=> i; rewrite coefM rpred_sum // => j _; apply: rpredM.
+Qed.
+Canonical polyOver_mulrPred := MulrPred polyOver_mulr_closed.
+Canonical polyOver_semiringPred := SemiringPred polyOver_mulr_closed.
+
+Lemma polyOverZ : {in kS & polyOver kS, forall c p, c *: p \is a polyOver kS}.
+Proof.
+by move=> c p Sc /polyOverP Sp; apply/polyOverP=> i; rewrite coefZ rpredM ?Sp.
+Qed.
+
+Lemma polyOverX : 'X \in polyOver kS.
+Proof. by rewrite qualifE polyseqX /= rpred0 rpred1. Qed.
+
+Lemma rpred_horner : {in polyOver kS & kS, forall p x, p.[x] \in kS}.
+Proof.
+move=> p x /polyOverP Sp Sx; rewrite horner_coef rpred_sum // => i _.
+by rewrite rpredM ?rpredX.
+Qed.
+
+End PolyOverSemiring.
+
+Section PolyOverRing.
+
+Context (S : pred_class) (ringS : @subringPred R S) (kS : keyed_pred ringS).
+Canonical polyOver_smulrPred := SmulrPred (polyOver_mulr_closed kS).
+Canonical polyOver_subringPred := SubringPred (polyOver_mulr_closed kS).
+
+Lemma polyOverXsubC c : ('X - c%:P \in polyOver kS) = (c \in kS).
+Proof. by rewrite rpredBl ?polyOverX ?polyOverC. Qed.
+
+End PolyOverRing.
+
+(* Single derivative. *)
+
+Definition deriv p := \poly_(i < (size p).-1) (p`_i.+1 *+ i.+1).
+
+Local Notation "a ^` ()" := (deriv a).
+
+Lemma coef_deriv p i : p^`()`_i = p`_i.+1 *+ i.+1.
+Proof.
+rewrite coef_poly -subn1 ltn_subRL.
+by case: leqP => // /(nth_default 0) ->; rewrite mul0rn.
+Qed.
+
+Lemma polyOver_deriv S (ringS : semiringPred S) (kS : keyed_pred ringS) :
+ {in polyOver kS, forall p, p^`() \is a polyOver kS}.
+Proof.
+by move=> p /polyOverP Kp; apply/polyOverP=> i; rewrite coef_deriv rpredMn ?Kp.
+Qed.
+
+Lemma derivC c : c%:P^`() = 0.
+Proof. by apply/polyP=> i; rewrite coef_deriv coef0 coefC mul0rn. Qed.
+
+Lemma derivX : ('X)^`() = 1.
+Proof. by apply/polyP=> [[|i]]; rewrite coef_deriv coef1 coefX ?mul0rn. Qed.
+
+Lemma derivXn n : 'X^n^`() = 'X^n.-1 *+ n.
+Proof.
+case: n => [|n]; first exact: derivC.
+apply/polyP=> i; rewrite coef_deriv coefMn !coefXn eqSS.
+by case: eqP => [-> // | _]; rewrite !mul0rn.
+Qed.
+
+Fact deriv_is_linear : linear deriv.
+Proof.
+move=> k p q; apply/polyP=> i.
+by rewrite !(coef_deriv, coefD, coefZ) mulrnDl mulrnAr.
+Qed.
+Canonical deriv_additive := Additive deriv_is_linear.
+Canonical deriv_linear := Linear deriv_is_linear.
+
+Lemma deriv0 : 0^`() = 0.
+Proof. exact: linear0. Qed.
+
+Lemma derivD : {morph deriv : p q / p + q}.
+Proof. exact: linearD. Qed.
+
+Lemma derivN : {morph deriv : p / - p}.
+Proof. exact: linearN. Qed.
+
+Lemma derivB : {morph deriv : p q / p - q}.
+Proof. exact: linearB. Qed.
+
+Lemma derivXsubC (a : R) : ('X - a%:P)^`() = 1.
+Proof. by rewrite derivB derivX derivC subr0. Qed.
+
+Lemma derivMn n p : (p *+ n)^`() = p^`() *+ n.
+Proof. exact: linearMn. Qed.
+
+Lemma derivMNn n p : (p *- n)^`() = p^`() *- n.
+Proof. exact: linearMNn. Qed.
+
+Lemma derivZ c p : (c *: p)^`() = c *: p^`().
+Proof. by rewrite linearZ. Qed.
+
+Lemma deriv_mulC c p : (c%:P * p)^`() = c%:P * p^`().
+Proof. by rewrite !mul_polyC derivZ. Qed.
+
+Lemma derivMXaddC p c : (p * 'X + c%:P)^`() = p + p^`() * 'X.
+Proof.
+apply/polyP=> i; rewrite raddfD /= derivC addr0 coefD !(coefMX, coef_deriv).
+by case: i; rewrite ?addr0.
+Qed.
+
+Lemma derivM p q : (p * q)^`() = p^`() * q + p * q^`().
+Proof.
+elim/poly_ind: p => [|p b IHp]; first by rewrite !(mul0r, add0r, derivC).
+rewrite mulrDl -mulrA -commr_polyX mulrA -[_ * 'X]addr0 raddfD /= !derivMXaddC.
+by rewrite deriv_mulC IHp !mulrDl -!mulrA !commr_polyX !addrA.
+Qed.
+
+Definition derivE := Eval lazy beta delta [morphism_2 morphism_1] in
+ (derivZ, deriv_mulC, derivC, derivX, derivMXaddC, derivXsubC, derivM, derivB,
+ derivD, derivN, derivXn, derivM, derivMn).
+
+(* Iterated derivative. *)
+Definition derivn n p := iter n deriv p.
+
+Local Notation "a ^` ( n )" := (derivn n a) : ring_scope.
+
+Lemma derivn0 p : p^`(0) = p.
+Proof. by []. Qed.
+
+Lemma derivn1 p : p^`(1) = p^`().
+Proof. by []. Qed.
+
+Lemma derivnS p n : p^`(n.+1) = p^`(n)^`().
+Proof. by []. Qed.
+
+Lemma derivSn p n : p^`(n.+1) = p^`()^`(n).
+Proof. exact: iterSr. Qed.
+
+Lemma coef_derivn n p i : p^`(n)`_i = p`_(n + i) *+ (n + i) ^_ n.
+Proof.
+elim: n i => [|n IHn] i; first by rewrite ffactn0 mulr1n.
+by rewrite derivnS coef_deriv IHn -mulrnA ffactnSr addSnnS addKn.
+Qed.
+
+Lemma polyOver_derivn S (ringS : semiringPred S) (kS : keyed_pred ringS) :
+ {in polyOver kS, forall p n, p^`(n) \is a polyOver kS}.
+Proof.
+move=> p /polyOverP Kp /= n; apply/polyOverP=> i.
+by rewrite coef_derivn rpredMn.
+Qed.
+
+Fact derivn_is_linear n : linear (derivn n).
+Proof. by elim: n => // n IHn a p q; rewrite derivnS IHn linearP. Qed.
+Canonical derivn_additive n := Additive (derivn_is_linear n).
+Canonical derivn_linear n := Linear (derivn_is_linear n).
+
+Lemma derivnC c n : c%:P^`(n) = if n == 0%N then c%:P else 0.
+Proof. by case: n => // n; rewrite derivSn derivC linear0. Qed.
+
+Lemma derivnD n : {morph derivn n : p q / p + q}.
+Proof. exact: linearD. Qed.
+
+Lemma derivn_sub n : {morph derivn n : p q / p - q}.
+Proof. exact: linearB. Qed.
+
+Lemma derivnMn n m p : (p *+ m)^`(n) = p^`(n) *+ m.
+Proof. exact: linearMn. Qed.
+
+Lemma derivnMNn n m p : (p *- m)^`(n) = p^`(n) *- m.
+Proof. exact: linearMNn. Qed.
+
+Lemma derivnN n : {morph derivn n : p / - p}.
+Proof. exact: linearN. Qed.
+
+Lemma derivnZ n : scalable (derivn n).
+Proof. exact: linearZZ. Qed.
+
+Lemma derivnXn m n : 'X^m^`(n) = 'X^(m - n) *+ m ^_ n.
+Proof.
+apply/polyP=>i; rewrite coef_derivn coefMn !coefXn.
+case: (ltnP m n) => [lt_m_n | le_m_n].
+ by rewrite eqn_leq leqNgt ltn_addr // mul0rn ffact_small.
+by rewrite -{1 3}(subnKC le_m_n) eqn_add2l; case: eqP => [->|]; rewrite ?mul0rn.
+Qed.
+
+Lemma derivnMXaddC n p c :
+ (p * 'X + c%:P)^`(n.+1) = p^`(n) *+ n.+1 + p^`(n.+1) * 'X.
+Proof.
+elim: n => [|n IHn]; first by rewrite derivn1 derivMXaddC.
+rewrite derivnS IHn derivD derivM derivX mulr1 derivMn -!derivnS.
+by rewrite addrA addrAC -mulrSr.
+Qed.
+
+Lemma derivn_poly0 p n : size p <= n -> p^`(n) = 0.
+Proof.
+move=> le_p_n; apply/polyP=> i; rewrite coef_derivn.
+rewrite nth_default; first by rewrite mul0rn coef0.
+by apply: leq_trans le_p_n _; apply leq_addr.
+Qed.
+
+Lemma lt_size_deriv (p : {poly R}) : p != 0 -> size p^`() < size p.
+Proof. by move=> /polySpred->; exact: size_poly. Qed.
+
+(* A normalising version of derivation to get the division by n! in Taylor *)
+
+Definition nderivn n p := \poly_(i < size p - n) (p`_(n + i) *+ 'C(n + i, n)).
+
+Local Notation "a ^`N ( n )" := (nderivn n a) : ring_scope.
+
+Lemma coef_nderivn n p i : p^`N(n)`_i = p`_(n + i) *+ 'C(n + i, n).
+Proof.
+rewrite coef_poly ltn_subRL; case: leqP => // le_p_ni.
+by rewrite nth_default ?mul0rn.
+Qed.
+
+(* Here is the division by n! *)
+Lemma nderivn_def n p : p^`(n) = p^`N(n) *+ n`!.
+Proof.
+by apply/polyP=> i; rewrite coefMn coef_nderivn coef_derivn -mulrnA bin_ffact.
+Qed.
+
+Lemma polyOver_nderivn S (ringS : semiringPred S) (kS : keyed_pred ringS) :
+ {in polyOver kS, forall p n, p^`N(n) \in polyOver kS}.
+Proof.
+move=> p /polyOverP Sp /= n; apply/polyOverP=> i.
+by rewrite coef_nderivn rpredMn.
+Qed.
+
+Lemma nderivn0 p : p^`N(0) = p.
+Proof. by rewrite -[p^`N(0)](nderivn_def 0). Qed.
+
+Lemma nderivn1 p : p^`N(1) = p^`().
+Proof. by rewrite -[p^`N(1)](nderivn_def 1). Qed.
+
+Lemma nderivnC c n : (c%:P)^`N(n) = if n == 0%N then c%:P else 0.
+Proof.
+apply/polyP=> i; rewrite coef_nderivn.
+by case: n => [|n]; rewrite ?bin0 // coef0 coefC mul0rn.
+Qed.
+
+Lemma nderivnXn m n : 'X^m^`N(n) = 'X^(m - n) *+ 'C(m, n).
+Proof.
+apply/polyP=> i; rewrite coef_nderivn coefMn !coefXn.
+have [lt_m_n | le_n_m] := ltnP m n.
+ by rewrite eqn_leq leqNgt ltn_addr // mul0rn bin_small.
+by rewrite -{1 3}(subnKC le_n_m) eqn_add2l; case: eqP => [->|]; rewrite ?mul0rn.
+Qed.
+
+Fact nderivn_is_linear n : linear (nderivn n).
+Proof.
+move=> k p q; apply/polyP=> i.
+by rewrite !(coef_nderivn, coefD, coefZ) mulrnDl mulrnAr.
+Qed.
+Canonical nderivn_additive n := Additive(nderivn_is_linear n).
+Canonical nderivn_linear n := Linear (nderivn_is_linear n).
+
+Lemma nderivnD n : {morph nderivn n : p q / p + q}.
+Proof. exact: linearD. Qed.
+
+Lemma nderivnB n : {morph nderivn n : p q / p - q}.
+Proof. exact: linearB. Qed.
+
+Lemma nderivnMn n m p : (p *+ m)^`N(n) = p^`N(n) *+ m.
+Proof. exact: linearMn. Qed.
+
+Lemma nderivnMNn n m p : (p *- m)^`N(n) = p^`N(n) *- m.
+Proof. exact: linearMNn. Qed.
+
+Lemma nderivnN n : {morph nderivn n : p / - p}.
+Proof. exact: linearN. Qed.
+
+Lemma nderivnZ n : scalable (nderivn n).
+Proof. exact: linearZZ. Qed.
+
+Lemma nderivnMXaddC n p c :
+ (p * 'X + c%:P)^`N(n.+1) = p^`N(n) + p^`N(n.+1) * 'X.
+Proof.
+apply/polyP=> i; rewrite coef_nderivn !coefD !coefMX coefC.
+rewrite !addSn /= !coef_nderivn addr0 binS mulrnDr addrC; congr (_ + _).
+by rewrite addSnnS; case: i; rewrite // addn0 bin_small.
+Qed.
+
+Lemma nderivn_poly0 p n : size p <= n -> p^`N(n) = 0.
+Proof.
+move=> le_p_n; apply/polyP=> i; rewrite coef_nderivn.
+rewrite nth_default; first by rewrite mul0rn coef0.
+by apply: leq_trans le_p_n _; apply leq_addr.
+Qed.
+
+Lemma nderiv_taylor p x h :
+ GRing.comm x h -> p.[x + h] = \sum_(i < size p) p^`N(i).[x] * h ^+ i.
+Proof.
+move/commrX=> cxh; elim/poly_ind: p => [|p c IHp].
+ by rewrite size_poly0 big_ord0 horner0.
+rewrite hornerMXaddC size_MXaddC.
+have [-> | nz_p] := altP (p =P 0).
+ rewrite horner0 !simp; have [-> | _] := c =P 0; first by rewrite big_ord0.
+ by rewrite size_poly0 big_ord_recl big_ord0 nderivn0 hornerC !simp.
+rewrite big_ord_recl nderivn0 !simp hornerMXaddC addrAC; congr (_ + _).
+rewrite mulrDr {}IHp !big_distrl polySpred //= big_ord_recl /= mulr1 -addrA.
+rewrite nderivn0 /bump /(addn 1) /=; congr (_ + _).
+rewrite !big_ord_recr /= nderivnMXaddC -mulrA -exprSr -polySpred // !addrA.
+congr (_ + _); last by rewrite (nderivn_poly0 (leqnn _)) !simp.
+rewrite addrC -big_split /=; apply: eq_bigr => i _.
+by rewrite nderivnMXaddC !hornerE_comm /= mulrDl -!mulrA -exprSr cxh.
+Qed.
+
+Lemma nderiv_taylor_wide n p x h :
+ GRing.comm x h -> size p <= n ->
+ p.[x + h] = \sum_(i < n) p^`N(i).[x] * h ^+ i.
+Proof.
+move/nderiv_taylor=> -> le_p_n.
+rewrite (big_ord_widen n (fun i => p^`N(i).[x] * h ^+ i)) // big_mkcond.
+apply: eq_bigr => i _; case: leqP => // /nderivn_poly0->.
+by rewrite horner0 simp.
+Qed.
+
+End PolynomialTheory.
+
+Prenex Implicits polyC Poly lead_coef root horner polyOver.
+Implicit Arguments monic [[R]].
+Notation "\poly_ ( i < n ) E" := (poly n (fun i => E)) : ring_scope.
+Notation "c %:P" := (polyC c) : ring_scope.
+Notation "'X" := (polyX _) : ring_scope.
+Notation "''X^' n" := ('X ^+ n) : ring_scope.
+Notation "p .[ x ]" := (horner p x) : ring_scope.
+Notation "n .-unity_root" := (root_of_unity n) : ring_scope.
+Notation "n .-primitive_root" := (primitive_root_of_unity n) : ring_scope.
+Notation "a ^` ()" := (deriv a) : ring_scope.
+Notation "a ^` ( n )" := (derivn n a) : ring_scope.
+Notation "a ^`N ( n )" := (nderivn n a) : ring_scope.
+
+Implicit Arguments monicP [R p].
+Implicit Arguments rootP [R p x].
+Implicit Arguments rootPf [R p x].
+Implicit Arguments rootPt [R p x].
+Implicit Arguments unity_rootP [R n z].
+Implicit Arguments polyOverP [[R] [S0] [addS] [kS] [p]].
+
+(* Container morphism. *)
+Section MapPoly.
+
+Section Definitions.
+
+Variables (aR rR : ringType) (f : aR -> rR).
+
+Definition map_poly (p : {poly aR}) := \poly_(i < size p) f p`_i.
+
+(* Alternative definition; the one above is more convenient because it lets *)
+(* us use the lemmas on \poly, e.g., size (map_poly p) <= size p is an *)
+(* instance of size_poly. *)
+Lemma map_polyE p : map_poly p = Poly (map f p).
+Proof.
+rewrite /map_poly unlock; congr Poly.
+apply: (@eq_from_nth _ 0); rewrite size_mkseq ?size_map // => i lt_i_p.
+by rewrite (nth_map 0) ?nth_mkseq.
+Qed.
+
+Definition commr_rmorph u := forall x, GRing.comm u (f x).
+
+Definition horner_morph u of commr_rmorph u := fun p => (map_poly p).[u].
+
+End Definitions.
+
+Variables aR rR : ringType.
+
+Section Combinatorial.
+
+Variables (iR : ringType) (f : aR -> rR).
+Local Notation "p ^f" := (map_poly f p) : ring_scope.
+
+Lemma map_poly0 : 0^f = 0.
+Proof. by rewrite map_polyE polyseq0. Qed.
+
+Lemma eq_map_poly (g : aR -> rR) : f =1 g -> map_poly f =1 map_poly g.
+Proof. by move=> eq_fg p; rewrite !map_polyE (eq_map eq_fg). Qed.
+
+Lemma map_poly_id g (p : {poly iR}) :
+ {in (p : seq iR), g =1 id} -> map_poly g p = p.
+Proof. by move=> g_id; rewrite map_polyE map_id_in ?polyseqK. Qed.
+
+Lemma coef_map_id0 p i : f 0 = 0 -> (p^f)`_i = f p`_i.
+Proof.
+by move=> f0; rewrite coef_poly; case: ltnP => // le_p_i; rewrite nth_default.
+Qed.
+
+Lemma map_Poly_id0 s : f 0 = 0 -> (Poly s)^f = Poly (map f s).
+Proof.
+move=> f0; apply/polyP=> j; rewrite coef_map_id0 ?coef_Poly //.
+have [/(nth_map 0 0)->// | le_s_j] := ltnP j (size s).
+by rewrite !nth_default ?size_map.
+Qed.
+
+Lemma map_poly_comp_id0 (g : iR -> aR) p :
+ f 0 = 0 -> map_poly (f \o g) p = (map_poly g p)^f.
+Proof. by move=> f0; rewrite map_polyE map_comp -map_Poly_id0 -?map_polyE. Qed.
+
+Lemma size_map_poly_id0 p : f (lead_coef p) != 0 -> size p^f = size p.
+Proof. by move=> nz_fp; apply: size_poly_eq. Qed.
+
+Lemma map_poly_eq0_id0 p : f (lead_coef p) != 0 -> (p^f == 0) = (p == 0).
+Proof. by rewrite -!size_poly_eq0 => /size_map_poly_id0->. Qed.
+
+Lemma lead_coef_map_id0 p :
+ f 0 = 0 -> f (lead_coef p) != 0 -> lead_coef p^f = f (lead_coef p).
+Proof.
+by move=> f0 nz_fp; rewrite lead_coefE coef_map_id0 ?size_map_poly_id0.
+Qed.
+
+Hypotheses (inj_f : injective f) (f_0 : f 0 = 0).
+
+Lemma size_map_inj_poly p : size p^f = size p.
+Proof.
+have [-> | nz_p] := eqVneq p 0; first by rewrite map_poly0 !size_poly0.
+by rewrite size_map_poly_id0 // -f_0 (inj_eq inj_f) lead_coef_eq0.
+Qed.
+
+Lemma map_inj_poly : injective (map_poly f).
+Proof.
+move=> p q /polyP eq_pq; apply/polyP=> i; apply: inj_f.
+by rewrite -!coef_map_id0 ?eq_pq.
+Qed.
+
+Lemma lead_coef_map_inj p : lead_coef p^f = f (lead_coef p).
+Proof. by rewrite !lead_coefE size_map_inj_poly coef_map_id0. Qed.
+
+End Combinatorial.
+
+Lemma map_polyK (f : aR -> rR) g :
+ cancel g f -> f 0 = 0 -> cancel (map_poly g) (map_poly f).
+Proof.
+by move=> gK f_0 p; rewrite /= -map_poly_comp_id0 ?map_poly_id // => x _ //=.
+Qed.
+
+Section Additive.
+
+Variables (iR : ringType) (f : {additive aR -> rR}).
+
+Local Notation "p ^f" := (map_poly (GRing.Additive.apply f) p) : ring_scope.
+
+Lemma coef_map p i : p^f`_i = f p`_i.
+Proof. exact: coef_map_id0 (raddf0 f). Qed.
+
+Lemma map_Poly s : (Poly s)^f = Poly (map f s).
+Proof. exact: map_Poly_id0 (raddf0 f). Qed.
+
+Lemma map_poly_comp (g : iR -> aR) p :
+ map_poly (f \o g) p = map_poly f (map_poly g p).
+Proof. exact: map_poly_comp_id0 (raddf0 f). Qed.
+
+Fact map_poly_is_additive : additive (map_poly f).
+Proof. by move=> p q; apply/polyP=> i; rewrite !(coef_map, coefB) raddfB. Qed.
+Canonical map_poly_additive := Additive map_poly_is_additive.
+
+Lemma map_polyC a : (a%:P)^f = (f a)%:P.
+Proof. by apply/polyP=> i; rewrite !(coef_map, coefC) -!mulrb raddfMn. Qed.
+
+Lemma lead_coef_map_eq p :
+ f (lead_coef p) != 0 -> lead_coef p^f = f (lead_coef p).
+Proof. exact: lead_coef_map_id0 (raddf0 f). Qed.
+
+End Additive.
+
+Variable f : {rmorphism aR -> rR}.
+Implicit Types p : {poly aR}.
+
+Local Notation "p ^f" := (map_poly (GRing.RMorphism.apply f) p) : ring_scope.
+
+Fact map_poly_is_rmorphism : rmorphism (map_poly f).
+Proof.
+split; first exact: map_poly_is_additive.
+split=> [p q|]; apply/polyP=> i; last first.
+ by rewrite !(coef_map, coef1) /= rmorph_nat.
+rewrite coef_map /= !coefM /= !rmorph_sum; apply: eq_bigr => j _.
+by rewrite !coef_map rmorphM.
+Qed.
+Canonical map_poly_rmorphism := RMorphism map_poly_is_rmorphism.
+
+Lemma map_polyZ c p : (c *: p)^f = f c *: p^f.
+Proof. by apply/polyP=> i; rewrite !(coef_map, coefZ) /= rmorphM. Qed.
+Canonical map_poly_linear :=
+ AddLinear (map_polyZ : scalable_for (f \; *:%R) (map_poly f)).
+Canonical map_poly_lrmorphism := [lrmorphism of map_poly f].
+
+Lemma map_polyX : ('X)^f = 'X.
+Proof. by apply/polyP=> i; rewrite coef_map !coefX /= rmorph_nat. Qed.
+
+Lemma map_polyXn n : ('X^n)^f = 'X^n.
+Proof. by rewrite rmorphX /= map_polyX. Qed.
+
+Lemma monic_map p : p \is monic -> p^f \is monic.
+Proof.
+move/monicP=> mon_p; rewrite monicE.
+by rewrite lead_coef_map_eq mon_p /= rmorph1 ?oner_neq0.
+Qed.
+
+Lemma horner_map p x : p^f.[f x] = f p.[x].
+Proof.
+elim/poly_ind: p => [|p c IHp]; first by rewrite !(rmorph0, horner0).
+rewrite hornerMXaddC !rmorphD !rmorphM /=.
+by rewrite map_polyX map_polyC hornerMXaddC IHp.
+Qed.
+
+Lemma map_comm_poly p x : comm_poly p x -> comm_poly p^f (f x).
+Proof. by rewrite /comm_poly horner_map -!rmorphM // => ->. Qed.
+
+Lemma map_comm_coef p x : comm_coef p x -> comm_coef p^f (f x).
+Proof. by move=> cpx i; rewrite coef_map -!rmorphM ?cpx. Qed.
+
+Lemma rmorph_root p x : root p x -> root p^f (f x).
+Proof. by move/eqP=> px0; rewrite rootE horner_map px0 rmorph0. Qed.
+
+Lemma rmorph_unity_root n z : n.-unity_root z -> n.-unity_root (f z).
+Proof.
+move/rmorph_root; rewrite rootE rmorphB hornerD hornerN.
+by rewrite /= map_polyXn rmorph1 hornerC hornerXn subr_eq0 unity_rootE.
+Qed.
+
+Section HornerMorph.
+
+Variable u : rR.
+Hypothesis cfu : commr_rmorph f u.
+
+Lemma horner_morphC a : horner_morph cfu a%:P = f a.
+Proof. by rewrite /horner_morph map_polyC hornerC. Qed.
+
+Lemma horner_morphX : horner_morph cfu 'X = u.
+Proof. by rewrite /horner_morph map_polyX hornerX. Qed.
+
+Fact horner_is_lrmorphism : lrmorphism_for (f \; *%R) (horner_morph cfu).
+Proof.
+rewrite /horner_morph; split=> [|c p]; last by rewrite linearZ hornerZ.
+split=> [p q|]; first by rewrite /horner_morph rmorphB hornerD hornerN.
+split=> [p q|]; last by rewrite /horner_morph rmorph1 hornerC.
+rewrite /horner_morph rmorphM /= hornerM_comm //.
+by apply: comm_coef_poly => i; rewrite coef_map cfu.
+Qed.
+Canonical horner_additive := Additive horner_is_lrmorphism.
+Canonical horner_rmorphism := RMorphism horner_is_lrmorphism.
+Canonical horner_linear := AddLinear horner_is_lrmorphism.
+Canonical horner_lrmorphism := [lrmorphism of horner_morph cfu].
+
+End HornerMorph.
+
+Lemma deriv_map p : p^f^`() = (p^`())^f.
+Proof. by apply/polyP => i; rewrite !(coef_map, coef_deriv) //= rmorphMn. Qed.
+
+Lemma derivn_map p n : p^f^`(n) = (p^`(n))^f.
+Proof. by apply/polyP => i; rewrite !(coef_map, coef_derivn) //= rmorphMn. Qed.
+
+Lemma nderivn_map p n : p^f^`N(n) = (p^`N(n))^f.
+Proof. by apply/polyP => i; rewrite !(coef_map, coef_nderivn) //= rmorphMn. Qed.
+
+End MapPoly.
+
+(* Morphisms from the polynomial ring, and the initiality of polynomials *)
+(* with respect to these. *)
+Section MorphPoly.
+
+Variable (aR rR : ringType) (pf : {rmorphism {poly aR} -> rR}).
+
+Lemma poly_morphX_comm : commr_rmorph (pf \o polyC) (pf 'X).
+Proof. by move=> a; rewrite /GRing.comm /= -!rmorphM // commr_polyX. Qed.
+
+Lemma poly_initial : pf =1 horner_morph poly_morphX_comm.
+Proof.
+apply: poly_ind => [|p a IHp]; first by rewrite !rmorph0.
+by rewrite !rmorphD !rmorphM /= -{}IHp horner_morphC ?horner_morphX.
+Qed.
+
+End MorphPoly.
+
+Notation "p ^:P" := (map_poly polyC p) : ring_scope.
+
+Section PolyCompose.
+
+Variable R : ringType.
+Implicit Types p q : {poly R}.
+
+Definition comp_poly q p := p^:P.[q].
+
+Local Notation "p \Po q" := (comp_poly q p) : ring_scope.
+
+Lemma size_map_polyC p : size p^:P = size p.
+Proof. exact: size_map_inj_poly (@polyC_inj R) _ _. Qed.
+
+Lemma map_polyC_eq0 p : (p^:P == 0) = (p == 0).
+Proof. by rewrite -!size_poly_eq0 size_map_polyC. Qed.
+
+Lemma root_polyC p x : root p^:P x%:P = root p x.
+Proof. by rewrite rootE horner_map polyC_eq0. Qed.
+
+Lemma comp_polyE p q : p \Po q = \sum_(i < size p) p`_i *: q^+i.
+Proof.
+by rewrite [p \Po q]horner_poly; apply: eq_bigr => i _; rewrite mul_polyC.
+Qed.
+
+Lemma polyOver_comp S (ringS : semiringPred S) (kS : keyed_pred ringS) :
+ {in polyOver kS &, forall p q, p \Po q \in polyOver kS}.
+Proof.
+move=> p q /polyOverP Sp Sq; rewrite comp_polyE rpred_sum // => i _.
+by rewrite polyOverZ ?rpredX.
+Qed.
+
+Lemma comp_polyCr p c : p \Po c%:P = p.[c]%:P.
+Proof. exact: horner_map. Qed.
+
+Lemma comp_poly0r p : p \Po 0 = (p`_0)%:P.
+Proof. by rewrite comp_polyCr horner_coef0. Qed.
+
+Lemma comp_polyC c p : c%:P \Po p = c%:P.
+Proof. by rewrite /(_ \Po p) map_polyC hornerC. Qed.
+
+Fact comp_poly_is_linear p : linear (comp_poly p).
+Proof.
+move=> a q r.
+by rewrite /comp_poly rmorphD /= map_polyZ !hornerE_comm mul_polyC.
+Qed.
+Canonical comp_poly_additive p := Additive (comp_poly_is_linear p).
+Canonical comp_poly_linear p := Linear (comp_poly_is_linear p).
+
+Lemma comp_poly0 p : 0 \Po p = 0.
+Proof. exact: raddf0. Qed.
+
+Lemma comp_polyD p q r : (p + q) \Po r = (p \Po r) + (q \Po r).
+Proof. exact: raddfD. Qed.
+
+Lemma comp_polyB p q r : (p - q) \Po r = (p \Po r) - (q \Po r).
+Proof. exact: raddfB. Qed.
+
+Lemma comp_polyZ c p q : (c *: p) \Po q = c *: (p \Po q).
+Proof. exact: linearZZ. Qed.
+
+Lemma comp_polyXr p : p \Po 'X = p.
+Proof. by rewrite -{2}/(idfun p) poly_initial. Qed.
+
+Lemma comp_polyX p : 'X \Po p = p.
+Proof. by rewrite /(_ \Po p) map_polyX hornerX. Qed.
+
+Lemma comp_poly_MXaddC c p q : (p * 'X + c%:P) \Po q = (p \Po q) * q + c%:P.
+Proof.
+by rewrite /(_ \Po q) rmorphD rmorphM /= map_polyX map_polyC hornerMXaddC.
+Qed.
+
+Lemma comp_polyXaddC_K p z : (p \Po ('X + z%:P)) \Po ('X - z%:P) = p.
+Proof.
+have addzK: ('X + z%:P) \Po ('X - z%:P) = 'X.
+ by rewrite raddfD /= comp_polyC comp_polyX subrK.
+elim/poly_ind: p => [|p c IHp]; first by rewrite !comp_poly0.
+rewrite comp_poly_MXaddC linearD /= comp_polyC {1}/comp_poly rmorphM /=.
+by rewrite hornerM_comm /comm_poly -!/(_ \Po _) ?IHp ?addzK ?commr_polyX.
+Qed.
+
+Lemma size_comp_poly_leq p q :
+ size (p \Po q) <= ((size p).-1 * (size q).-1).+1.
+Proof.
+rewrite comp_polyE (leq_trans (size_sum _ _ _)) //; apply/bigmax_leqP => i _.
+rewrite (leq_trans (size_scale_leq _ _)) // (leq_trans (size_exp_leq _ _)) //.
+by rewrite ltnS mulnC leq_mul // -{2}(subnKC (valP i)) leq_addr.
+Qed.
+
+End PolyCompose.
+
+Notation "p \Po q" := (comp_poly q p) : ring_scope.
+
+Lemma map_comp_poly (aR rR : ringType) (f : {rmorphism aR -> rR}) p q :
+ map_poly f (p \Po q) = map_poly f p \Po map_poly f q.
+Proof.
+elim/poly_ind: p => [|p a IHp]; first by rewrite !raddf0.
+rewrite comp_poly_MXaddC !rmorphD !rmorphM /= !map_polyC map_polyX.
+by rewrite comp_poly_MXaddC -IHp.
+Qed.
+
+Section PolynomialComRing.
+
+Variable R : comRingType.
+Implicit Types p q : {poly R}.
+
+Fact poly_mul_comm p q : p * q = q * p.
+Proof.
+apply/polyP=> i; rewrite coefM coefMr.
+by apply: eq_bigr => j _; rewrite mulrC.
+Qed.
+
+Canonical poly_comRingType := Eval hnf in ComRingType {poly R} poly_mul_comm.
+Canonical polynomial_comRingType :=
+ Eval hnf in ComRingType (polynomial R) poly_mul_comm.
+Canonical poly_algType := Eval hnf in CommAlgType R {poly R}.
+Canonical polynomial_algType :=
+ Eval hnf in [algType R of polynomial R for poly_algType].
+
+Lemma hornerM p q x : (p * q).[x] = p.[x] * q.[x].
+Proof. by rewrite hornerM_comm //; exact: mulrC. Qed.
+
+Lemma horner_exp p x n : (p ^+ n).[x] = p.[x] ^+ n.
+Proof. by rewrite horner_exp_comm //; exact: mulrC. Qed.
+
+Lemma horner_prod I r (P : pred I) (F : I -> {poly R}) x :
+ (\prod_(i <- r | P i) F i).[x] = \prod_(i <- r | P i) (F i).[x].
+Proof. by elim/big_rec2: _ => [|i _ p _ <-]; rewrite (hornerM, hornerC). Qed.
+
+Definition hornerE :=
+ (hornerD, hornerN, hornerX, hornerC, horner_cons,
+ simp, hornerCM, hornerZ, hornerM).
+
+Definition horner_eval (x : R) := horner^~ x.
+Lemma horner_evalE x p : horner_eval x p = p.[x]. Proof. by []. Qed.
+
+Fact horner_eval_is_lrmorphism x : lrmorphism_for *%R (horner_eval x).
+Proof.
+have cxid: commr_rmorph idfun x by exact: mulrC.
+have evalE : horner_eval x =1 horner_morph cxid.
+ by move=> p; congr _.[x]; rewrite map_poly_id.
+split=> [|c p]; last by rewrite !evalE /= -linearZ.
+by do 2?split=> [p q|]; rewrite !evalE (rmorphB, rmorphM, rmorph1).
+Qed.
+Canonical horner_eval_additive x := Additive (horner_eval_is_lrmorphism x).
+Canonical horner_eval_rmorphism x := RMorphism (horner_eval_is_lrmorphism x).
+Canonical horner_eval_linear x := AddLinear (horner_eval_is_lrmorphism x).
+Canonical horner_eval_lrmorphism x := [lrmorphism of horner_eval x].
+
+Fact comp_poly_multiplicative q : multiplicative (comp_poly q).
+Proof.
+split=> [p1 p2|]; last by rewrite comp_polyC.
+by rewrite /comp_poly rmorphM hornerM_comm //; exact: mulrC.
+Qed.
+Canonical comp_poly_rmorphism q := AddRMorphism (comp_poly_multiplicative q).
+Canonical comp_poly_lrmorphism q := [lrmorphism of comp_poly q].
+
+Lemma comp_polyM p q r : (p * q) \Po r = (p \Po r) * (q \Po r).
+Proof. exact: rmorphM. Qed.
+
+Lemma comp_polyA p q r : p \Po (q \Po r) = (p \Po q) \Po r.
+Proof.
+elim/poly_ind: p => [|p c IHp]; first by rewrite !comp_polyC.
+by rewrite !comp_polyD !comp_polyM !comp_polyX IHp !comp_polyC.
+Qed.
+
+Lemma horner_comp p q x : (p \Po q).[x] = p.[q.[x]].
+Proof. by apply: polyC_inj; rewrite -!comp_polyCr comp_polyA. Qed.
+
+Lemma root_comp p q x : root (p \Po q) x = root p (q.[x]).
+Proof. by rewrite !rootE horner_comp. Qed.
+
+Lemma deriv_comp p q : (p \Po q) ^`() = (p ^`() \Po q) * q^`().
+Proof.
+elim/poly_ind: p => [|p c IHp]; first by rewrite !(deriv0, comp_poly0) mul0r.
+rewrite comp_poly_MXaddC derivD derivC derivM IHp derivMXaddC comp_polyD.
+by rewrite comp_polyM comp_polyX addr0 addrC mulrAC -mulrDl.
+Qed.
+
+Lemma deriv_exp p n : (p ^+ n)^`() = p^`() * p ^+ n.-1 *+ n.
+Proof.
+elim: n => [|n IHn]; first by rewrite expr0 mulr0n derivC.
+by rewrite exprS derivM {}IHn (mulrC p) mulrnAl -mulrA -exprSr mulrS; case n.
+Qed.
+
+Definition derivCE := (derivE, deriv_exp).
+
+End PolynomialComRing.
+
+Section PolynomialIdomain.
+
+(* Integral domain structure on poly *)
+Variable R : idomainType.
+
+Implicit Types (a b x y : R) (p q r m : {poly R}).
+
+Lemma size_mul p q : p != 0 -> q != 0 -> size (p * q) = (size p + size q).-1.
+Proof.
+by move=> nz_p nz_q; rewrite -size_proper_mul ?mulf_neq0 ?lead_coef_eq0.
+Qed.
+
+Fact poly_idomainAxiom p q : p * q = 0 -> (p == 0) || (q == 0).
+Proof.
+move=> pq0; apply/norP=> [[p_nz q_nz]]; move/eqP: (size_mul p_nz q_nz).
+by rewrite eq_sym pq0 size_poly0 (polySpred p_nz) (polySpred q_nz) addnS.
+Qed.
+
+Definition poly_unit : pred {poly R} :=
+ fun p => (size p == 1%N) && (p`_0 \in GRing.unit).
+
+Definition poly_inv p := if p \in poly_unit then (p`_0)^-1%:P else p.
+
+Fact poly_mulVp : {in poly_unit, left_inverse 1 poly_inv *%R}.
+Proof.
+move=> p Up; rewrite /poly_inv Up.
+by case/andP: Up => /size_poly1P[c _ ->]; rewrite coefC -polyC_mul => /mulVr->.
+Qed.
+
+Fact poly_intro_unit p q : q * p = 1 -> p \in poly_unit.
+Proof.
+move=> pq1; apply/andP; split; last first.
+ apply/unitrP; exists q`_0.
+ by rewrite 2!mulrC -!/(coefp 0 _) -rmorphM pq1 rmorph1.
+have: size (q * p) == 1%N by rewrite pq1 size_poly1.
+have [-> | nz_p] := eqVneq p 0; first by rewrite mulr0 size_poly0.
+have [-> | nz_q] := eqVneq q 0; first by rewrite mul0r size_poly0.
+rewrite size_mul // (polySpred nz_p) (polySpred nz_q) addnS addSn !eqSS.
+by rewrite addn_eq0 => /andP[].
+Qed.
+
+Fact poly_inv_out : {in [predC poly_unit], poly_inv =1 id}.
+Proof. by rewrite /poly_inv => p /negbTE/= ->. Qed.
+
+Definition poly_comUnitMixin :=
+ ComUnitRingMixin poly_mulVp poly_intro_unit poly_inv_out.
+
+Canonical poly_unitRingType :=
+ Eval hnf in UnitRingType {poly R} poly_comUnitMixin.
+Canonical polynomial_unitRingType :=
+ Eval hnf in [unitRingType of polynomial R for poly_unitRingType].
+
+Canonical poly_unitAlgType := Eval hnf in [unitAlgType R of {poly R}].
+Canonical polynomial_unitAlgType := Eval hnf in [unitAlgType R of polynomial R].
+
+Canonical poly_comUnitRingType := Eval hnf in [comUnitRingType of {poly R}].
+Canonical polynomial_comUnitRingType :=
+ Eval hnf in [comUnitRingType of polynomial R].
+
+Canonical poly_idomainType :=
+ Eval hnf in IdomainType {poly R} poly_idomainAxiom.
+Canonical polynomial_idomainType :=
+ Eval hnf in [idomainType of polynomial R for poly_idomainType].
+
+Lemma poly_unitE p :
+ (p \in GRing.unit) = (size p == 1%N) && (p`_0 \in GRing.unit).
+Proof. by []. Qed.
+
+Lemma poly_invE p : p ^-1 = if p \in GRing.unit then (p`_0)^-1%:P else p.
+Proof. by []. Qed.
+
+Lemma polyC_inv c : c%:P^-1 = (c^-1)%:P.
+Proof.
+have [/rmorphV-> // | nUc] := boolP (c \in GRing.unit).
+by rewrite !invr_out // poly_unitE coefC (negbTE nUc) andbF.
+Qed.
+
+Lemma rootM p q x : root (p * q) x = root p x || root q x.
+Proof. by rewrite !rootE hornerM mulf_eq0. Qed.
+
+Lemma rootZ x a p : a != 0 -> root (a *: p) x = root p x.
+Proof. by move=> nz_a; rewrite -mul_polyC rootM rootC (negPf nz_a). Qed.
+
+Lemma size_scale a p : a != 0 -> size (a *: p) = size p.
+Proof. by move/lregP/lreg_size->. Qed.
+
+Lemma size_Cmul a p : a != 0 -> size (a%:P * p) = size p.
+Proof. by rewrite mul_polyC => /size_scale->. Qed.
+
+Lemma lead_coefM p q : lead_coef (p * q) = lead_coef p * lead_coef q.
+Proof.
+have [-> | nz_p] := eqVneq p 0; first by rewrite !(mul0r, lead_coef0).
+have [-> | nz_q] := eqVneq q 0; first by rewrite !(mulr0, lead_coef0).
+by rewrite lead_coef_proper_mul // mulf_neq0 ?lead_coef_eq0.
+Qed.
+
+Lemma lead_coefZ a p : lead_coef (a *: p) = a * lead_coef p.
+Proof. by rewrite -mul_polyC lead_coefM lead_coefC. Qed.
+
+Lemma scale_poly_eq0 a p : (a *: p == 0) = (a == 0) || (p == 0).
+Proof. by rewrite -mul_polyC mulf_eq0 polyC_eq0. Qed.
+
+Lemma size_prod (I : finType) (P : pred I) (F : I -> {poly R}) :
+ (forall i, P i -> F i != 0) ->
+ size (\prod_(i | P i) F i) = ((\sum_(i | P i) size (F i)).+1 - #|P|)%N.
+Proof.
+move=> nzF; transitivity (\sum_(i | P i) (size (F i)).-1).+1; last first.
+ apply: canRL (addKn _) _; rewrite addnS -sum1_card -big_split /=.
+ by congr _.+1; apply: eq_bigr => i /nzF/polySpred.
+elim/big_rec2: _ => [|i d p /nzF nzFi IHp]; first by rewrite size_poly1.
+by rewrite size_mul // -?size_poly_eq0 IHp // addnS polySpred.
+Qed.
+
+Lemma size_exp p n : (size (p ^+ n)).-1 = ((size p).-1 * n)%N.
+Proof.
+elim: n => [|n IHn]; first by rewrite size_poly1 muln0.
+have [-> | nz_p] := eqVneq p 0; first by rewrite exprS mul0r size_poly0.
+rewrite exprS size_mul ?expf_neq0 // mulnS -{}IHn.
+by rewrite polySpred // [size (p ^+ n)]polySpred ?expf_neq0 ?addnS.
+Qed.
+
+Lemma lead_coef_exp p n : lead_coef (p ^+ n) = lead_coef p ^+ n.
+Proof.
+elim: n => [|n IHn]; first by rewrite !expr0 lead_coef1.
+by rewrite !exprS lead_coefM IHn.
+Qed.
+
+Lemma root_prod_XsubC rs x :
+ root (\prod_(a <- rs) ('X - a%:P)) x = (x \in rs).
+Proof.
+elim: rs => [|a rs IHrs]; first by rewrite rootE big_nil hornerC oner_eq0.
+by rewrite big_cons rootM IHrs root_XsubC.
+Qed.
+
+Lemma root_exp_XsubC n a x : root (('X - a%:P) ^+ n.+1) x = (x == a).
+Proof. by rewrite rootE horner_exp expf_eq0 [_ == 0]root_XsubC. Qed.
+
+Lemma size_comp_poly p q :
+ (size (p \Po q)).-1 = ((size p).-1 * (size q).-1)%N.
+Proof.
+have [-> | nz_p] := eqVneq p 0; first by rewrite comp_poly0 size_poly0.
+have [/size1_polyC-> | nc_q] := leqP (size q) 1.
+ by rewrite comp_polyCr !size_polyC -!sub1b -!subnS muln0.
+have nz_q: q != 0 by rewrite -size_poly_eq0 -(subnKC nc_q).
+rewrite mulnC comp_polyE (polySpred nz_p) /= big_ord_recr /= addrC.
+rewrite size_addl size_scale ?lead_coef_eq0 ?size_exp //=.
+rewrite [X in _ < X]polySpred ?expf_neq0 // ltnS size_exp.
+rewrite (leq_trans (size_sum _ _ _)) //; apply/bigmax_leqP => i _.
+rewrite (leq_trans (size_scale_leq _ _)) // polySpred ?expf_neq0 //.
+by rewrite size_exp -(subnKC nc_q) ltn_pmul2l.
+Qed.
+
+Lemma size_comp_poly2 p q : size q = 2 -> size (p \Po q) = size p.
+Proof.
+have [/size1_polyC->| p_gt1] := leqP (size p) 1; first by rewrite comp_polyC.
+move=> lin_q; have{lin_q} sz_pq: (size (p \Po q)).-1 = (size p).-1.
+ by rewrite size_comp_poly lin_q muln1.
+rewrite -(ltn_predK p_gt1) -sz_pq -polySpred // -size_poly_gt0 ltnW //.
+by rewrite -subn_gt0 subn1 sz_pq -subn1 subn_gt0.
+Qed.
+
+Lemma comp_poly2_eq0 p q : size q = 2 -> (p \Po q == 0) = (p == 0).
+Proof. by rewrite -!size_poly_eq0 => /size_comp_poly2->. Qed.
+
+Lemma lead_coef_comp p q :
+ size q > 1 -> lead_coef (p \Po q) = lead_coef p * lead_coef q ^+ (size p).-1.
+Proof.
+move=> q_gt1; have nz_q: q != 0 by rewrite -size_poly_gt0 ltnW.
+have [-> | nz_p] := eqVneq p 0; first by rewrite comp_poly0 !lead_coef0 mul0r.
+rewrite comp_polyE polySpred //= big_ord_recr /= addrC -lead_coefE.
+rewrite lead_coefDl; first by rewrite lead_coefZ lead_coef_exp.
+rewrite size_scale ?lead_coef_eq0 // (polySpred (expf_neq0 _ nz_q)) ltnS.
+apply/leq_sizeP=> i le_qp_i; rewrite coef_sum big1 // => j _.
+rewrite coefZ (nth_default 0 (leq_trans _ le_qp_i)) ?mulr0 //=.
+by rewrite polySpred ?expf_neq0 // !size_exp -(subnKC q_gt1) ltn_pmul2l.
+Qed.
+
+End PolynomialIdomain.
+
+Section MapFieldPoly.
+
+Variables (F : fieldType) (R : ringType) (f : {rmorphism F -> R}).
+
+Local Notation "p ^f" := (map_poly f p) : ring_scope.
+
+Lemma size_map_poly p : size p^f = size p.
+Proof.
+have [-> | nz_p] := eqVneq p 0; first by rewrite rmorph0 !size_poly0.
+by rewrite size_poly_eq // fmorph_eq0 // lead_coef_eq0.
+Qed.
+
+Lemma lead_coef_map p : lead_coef p^f = f (lead_coef p).
+Proof.
+have [-> | nz_p] := eqVneq p 0; first by rewrite !(rmorph0, lead_coef0).
+by rewrite lead_coef_map_eq // fmorph_eq0 // lead_coef_eq0.
+Qed.
+
+Lemma map_poly_eq0 p : (p^f == 0) = (p == 0).
+Proof. by rewrite -!size_poly_eq0 size_map_poly. Qed.
+
+Lemma map_poly_inj : injective (map_poly f).
+Proof.
+move=> p q eqfpq; apply/eqP; rewrite -subr_eq0 -map_poly_eq0.
+by rewrite rmorphB /= eqfpq subrr.
+Qed.
+
+Lemma map_monic p : (p^f \is monic) = (p \is monic).
+Proof. by rewrite monicE lead_coef_map fmorph_eq1. Qed.
+
+Lemma map_poly_com p x : comm_poly p^f (f x).
+Proof. exact: map_comm_poly (mulrC x _). Qed.
+
+Lemma fmorph_root p x : root p^f (f x) = root p x.
+Proof. by rewrite rootE horner_map // fmorph_eq0. Qed.
+
+Lemma fmorph_unity_root n z : n.-unity_root (f z) = n.-unity_root z.
+Proof. by rewrite !unity_rootE -(inj_eq (fmorph_inj f)) rmorphX ?rmorph1. Qed.
+
+Lemma fmorph_primitive_root n z :
+ n.-primitive_root (f z) = n.-primitive_root z.
+Proof.
+by congr (_ && _); apply: eq_forallb => i; rewrite fmorph_unity_root.
+Qed.
+
+End MapFieldPoly.
+
+Implicit Arguments map_poly_inj [[F] [R] x1 x2].
+
+Section MaxRoots.
+
+Variable R : unitRingType.
+Implicit Types (x y : R) (rs : seq R) (p : {poly R}).
+
+Definition diff_roots (x y : R) := (x * y == y * x) && (y - x \in GRing.unit).
+
+Fixpoint uniq_roots rs :=
+ if rs is x :: rs' then all (diff_roots x) rs' && uniq_roots rs' else true.
+
+Lemma uniq_roots_prod_XsubC p rs :
+ all (root p) rs -> uniq_roots rs ->
+ exists q, p = q * \prod_(z <- rs) ('X - z%:P).
+Proof.
+elim: rs => [|z rs IHrs] /=; first by rewrite big_nil; exists p; rewrite mulr1.
+case/andP=> rpz rprs /andP[drs urs]; case: IHrs => {urs rprs}// q def_p.
+have [|q' def_q] := factor_theorem q z _; last first.
+ by exists q'; rewrite big_cons mulrA -def_q.
+rewrite {p}def_p in rpz.
+elim/last_ind: rs drs rpz => [|rs t IHrs] /=; first by rewrite big_nil mulr1.
+rewrite all_rcons => /andP[/andP[/eqP czt Uzt] /IHrs {IHrs}IHrs].
+rewrite -cats1 big_cat big_seq1 /= mulrA rootE hornerM_comm; last first.
+ by rewrite /comm_poly hornerXsubC mulrBl mulrBr czt.
+rewrite hornerXsubC -opprB mulrN oppr_eq0 -(mul0r (t - z)).
+by rewrite (inj_eq (mulIr Uzt)) => /IHrs.
+Qed.
+
+Theorem max_ring_poly_roots p rs :
+ p != 0 -> all (root p) rs -> uniq_roots rs -> size rs < size p.
+Proof.
+move=> nz_p _ /(@uniq_roots_prod_XsubC p)[// | q def_p]; rewrite def_p in nz_p *.
+have nz_q: q != 0 by apply: contraNneq nz_p => ->; rewrite mul0r.
+rewrite size_Mmonic ?monic_prod_XsubC // (polySpred nz_q) addSn /=.
+by rewrite size_prod_XsubC leq_addl.
+Qed.
+
+Lemma all_roots_prod_XsubC p rs :
+ size p = (size rs).+1 -> all (root p) rs -> uniq_roots rs ->
+ p = lead_coef p *: \prod_(z <- rs) ('X - z%:P).
+Proof.
+move=> size_p /uniq_roots_prod_XsubC def_p Urs.
+case/def_p: Urs => q -> {p def_p} in size_p *.
+have [q0 | nz_q] := eqVneq q 0; first by rewrite q0 mul0r size_poly0 in size_p.
+have{q nz_q size_p} /size_poly1P[c _ ->]: size q == 1%N.
+ rewrite -(eqn_add2r (size rs)) add1n -size_p.
+ by rewrite size_Mmonic ?monic_prod_XsubC // size_prod_XsubC addnS.
+by rewrite lead_coef_Mmonic ?monic_prod_XsubC // lead_coefC mul_polyC.
+Qed.
+
+End MaxRoots.
+
+Section FieldRoots.
+
+Variable F : fieldType.
+Implicit Types (p : {poly F}) (rs : seq F).
+
+Lemma poly2_root p : size p = 2 -> {r | root p r}.
+Proof.
+case: p => [[|p0 [|p1 []]] //= nz_p1]; exists (- p0 / p1).
+by rewrite /root addr_eq0 /= mul0r add0r mulrC divfK ?opprK.
+Qed.
+
+Lemma uniq_rootsE rs : uniq_roots rs = uniq rs.
+Proof.
+elim: rs => //= r rs ->; congr (_ && _); rewrite -has_pred1 -all_predC.
+by apply: eq_all => t; rewrite /diff_roots mulrC eqxx unitfE subr_eq0.
+Qed.
+
+Theorem max_poly_roots p rs :
+ p != 0 -> all (root p) rs -> uniq rs -> size rs < size p.
+Proof. by rewrite -uniq_rootsE; exact: max_ring_poly_roots. Qed.
+
+Section UnityRoots.
+
+Variable n : nat.
+
+Lemma max_unity_roots rs :
+ n > 0 -> all n.-unity_root rs -> uniq rs -> size rs <= n.
+Proof.
+move=> n_gt0 rs_n_1 Urs; have szPn := size_Xn_sub_1 F n_gt0.
+by rewrite -ltnS -szPn max_poly_roots -?size_poly_eq0 ?szPn.
+Qed.
+
+Lemma mem_unity_roots rs :
+ n > 0 -> all n.-unity_root rs -> uniq rs -> size rs = n ->
+ n.-unity_root =i rs.
+Proof.
+move=> n_gt0 rs_n_1 Urs sz_rs_n x; rewrite -topredE /=.
+apply/idP/idP=> xn1; last exact: (allP rs_n_1).
+apply: contraFT (ltnn n) => not_rs_x.
+by rewrite -{1}sz_rs_n (@max_unity_roots (x :: rs)) //= ?xn1 ?not_rs_x.
+Qed.
+
+(* Showing the existence of a primitive root requires the theory in cyclic. *)
+
+Variable z : F.
+Hypothesis prim_z : n.-primitive_root z.
+
+Let zn := [seq z ^+ i | i <- index_iota 0 n].
+
+Lemma factor_Xn_sub_1 : \prod_(0 <= i < n) ('X - (z ^+ i)%:P) = 'X^n - 1.
+Proof.
+transitivity (\prod_(w <- zn) ('X - w%:P)); first by rewrite big_map.
+have n_gt0: n > 0 := prim_order_gt0 prim_z.
+rewrite (@all_roots_prod_XsubC _ ('X^n - 1) zn); first 1 last.
+- by rewrite size_Xn_sub_1 // size_map size_iota subn0.
+- apply/allP=> _ /mapP[i _ ->] /=; rewrite rootE !hornerE hornerXn.
+ by rewrite exprAC (prim_expr_order prim_z) expr1n subrr.
+- rewrite uniq_rootsE map_inj_in_uniq ?iota_uniq // => i j.
+ rewrite !mem_index_iota => ltin ltjn /eqP.
+ by rewrite (eq_prim_root_expr prim_z) !modn_small // => /eqP.
+by rewrite (monicP (monic_Xn_sub_1 F n_gt0)) scale1r.
+Qed.
+
+Lemma prim_rootP x : x ^+ n = 1 -> {i : 'I_n | x = z ^+ i}.
+Proof.
+move=> xn1; pose logx := [pred i : 'I_n | x == z ^+ i].
+case: (pickP logx) => [i /eqP-> | no_i]; first by exists i.
+case: notF; suffices{no_i}: x \in zn.
+ case/mapP=> i; rewrite mem_index_iota => lt_i_n def_x.
+ by rewrite -(no_i (Ordinal lt_i_n)) /= -def_x.
+rewrite -root_prod_XsubC big_map factor_Xn_sub_1.
+by rewrite [root _ x]unity_rootE xn1.
+Qed.
+
+End UnityRoots.
+
+End FieldRoots.
+
+Section MapPolyRoots.
+
+Variables (F : fieldType) (R : unitRingType) (f : {rmorphism F -> R}).
+
+Lemma map_diff_roots x y : diff_roots (f x) (f y) = (x != y).
+Proof.
+rewrite /diff_roots -rmorphB // fmorph_unit // subr_eq0 //.
+by rewrite rmorph_comm // eqxx eq_sym.
+Qed.
+
+Lemma map_uniq_roots s : uniq_roots (map f s) = uniq s.
+Proof.
+elim: s => //= x s ->; congr (_ && _); elim: s => //= y s ->.
+by rewrite map_diff_roots -negb_or.
+Qed.
+
+End MapPolyRoots.
+
+Section AutPolyRoot.
+(* The action of automorphisms on roots of unity. *)
+
+Variable F : fieldType.
+Implicit Types u v : {rmorphism F -> F}.
+
+Lemma aut_prim_rootP u z n :
+ n.-primitive_root z -> {k | coprime k n & u z = z ^+ k}.
+Proof.
+move=> prim_z; have:= prim_z; rewrite -(fmorph_primitive_root u) => prim_uz.
+have [[k _] /= def_uz] := prim_rootP prim_z (prim_expr_order prim_uz).
+by exists k; rewrite // -(prim_root_exp_coprime _ prim_z) -def_uz.
+Qed.
+
+Lemma aut_unity_rootP u z n : n > 0 -> z ^+ n = 1 -> {k | u z = z ^+ k}.
+Proof.
+by move=> _ /prim_order_exists[// | m /(aut_prim_rootP u)[k]]; exists k.
+Qed.
+
+Lemma aut_unity_rootC u v z n : n > 0 -> z ^+ n = 1 -> u (v z) = v (u z).
+Proof.
+move=> n_gt0 /(aut_unity_rootP _ n_gt0) def_z.
+have [[i def_uz] [j def_vz]] := (def_z u, def_z v).
+by rewrite !(def_uz, def_vz, rmorphX) exprAC.
+Qed.
+
+End AutPolyRoot.
+
+Module UnityRootTheory.
+
+Notation "n .-unity_root" := (root_of_unity n) : unity_root_scope.
+Notation "n .-primitive_root" := (primitive_root_of_unity n) : unity_root_scope.
+Open Scope unity_root_scope.
+
+Definition unity_rootE := unity_rootE.
+Definition unity_rootP := @unity_rootP.
+Implicit Arguments unity_rootP [R n z].
+
+Definition prim_order_exists := prim_order_exists.
+Notation prim_order_gt0 := prim_order_gt0.
+Notation prim_expr_order := prim_expr_order.
+Definition prim_expr_mod := prim_expr_mod.
+Definition prim_order_dvd := prim_order_dvd.
+Definition eq_prim_root_expr := eq_prim_root_expr.
+
+Definition rmorph_unity_root := rmorph_unity_root.
+Definition fmorph_unity_root := fmorph_unity_root.
+Definition fmorph_primitive_root := fmorph_primitive_root.
+Definition max_unity_roots := max_unity_roots.
+Definition mem_unity_roots := mem_unity_roots.
+Definition prim_rootP := prim_rootP.
+
+End UnityRootTheory.
+
+Module PreClosedField.
+Section UseAxiom.
+
+Variable F : fieldType.
+Hypothesis closedF : GRing.ClosedField.axiom F.
+Implicit Type p : {poly F}.
+
+Lemma closed_rootP p : reflect (exists x, root p x) (size p != 1%N).
+Proof.
+have [-> | nz_p] := eqVneq p 0.
+ by rewrite size_poly0; left; exists 0; rewrite root0.
+rewrite neq_ltn {1}polySpred //=.
+apply: (iffP idP) => [p_gt1 | [a]]; last exact: root_size_gt1.
+pose n := (size p).-1; have n_gt0: n > 0 by rewrite -ltnS -polySpred.
+have [a Dan] := closedF (fun i => - p`_i / lead_coef p) n_gt0.
+exists a; apply/rootP; rewrite horner_coef polySpred // big_ord_recr /= -/n.
+rewrite {}Dan mulr_sumr -big_split big1 //= => i _.
+by rewrite -!mulrA mulrCA mulNr mulVKf ?subrr ?lead_coef_eq0.
+Qed.
+
+Lemma closed_nonrootP p : reflect (exists x, ~~ root p x) (p != 0).
+Proof.
+apply: (iffP idP) => [nz_p | [x]]; last first.
+ by apply: contraNneq => ->; apply: root0.
+have [[x /rootP p1x0]|] := altP (closed_rootP (p - 1)).
+ by exists x; rewrite -[p](subrK 1) /root hornerD p1x0 add0r hornerC oner_eq0.
+rewrite negbK => /size_poly1P[c _ /(canRL (subrK 1)) Dp].
+by exists 0; rewrite Dp -raddfD polyC_eq0 rootC in nz_p *.
+Qed.
+
+End UseAxiom.
+End PreClosedField.
+
+Section ClosedField.
+
+Variable F : closedFieldType.
+Implicit Type p : {poly F}.
+
+Let closedF := @solve_monicpoly F.
+
+Lemma closed_rootP p : reflect (exists x, root p x) (size p != 1%N).
+Proof. exact: PreClosedField.closed_rootP. Qed.
+
+Lemma closed_nonrootP p : reflect (exists x, ~~ root p x) (p != 0).
+Proof. exact: PreClosedField.closed_nonrootP. Qed.
+
+Lemma closed_field_poly_normal p :
+ {r : seq F | p = lead_coef p *: \prod_(z <- r) ('X - z%:P)}.
+Proof.
+apply: sig_eqW; elim: {p}_.+1 {-2}p (ltnSn (size p)) => // n IHn p le_p_n.
+have [/size1_polyC-> | p_gt1] := leqP (size p) 1.
+ by exists nil; rewrite big_nil lead_coefC alg_polyC.
+have [|x /factor_theorem[q Dp]] := closed_rootP p _; first by rewrite gtn_eqF.
+have nz_p: p != 0 by rewrite -size_poly_eq0 -(subnKC p_gt1).
+have:= nz_p; rewrite Dp mulf_eq0 lead_coefM => /norP[nz_q nz_Xx].
+rewrite ltnS polySpred // Dp size_mul // size_XsubC addn2 in le_p_n.
+have [r {1}->] := IHn q le_p_n; exists (x :: r).
+by rewrite lead_coefXsubC mulr1 big_cons -scalerAl mulrC.
+Qed.
+
+End ClosedField.
diff --git a/mathcomp/algebra/polyXY.v b/mathcomp/algebra/polyXY.v
new file mode 100644
index 0000000..3d2292e
--- /dev/null
+++ b/mathcomp/algebra/polyXY.v
@@ -0,0 +1,405 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool choice eqtype ssrnat seq div fintype.
+Require Import tuple finfun bigop fingroup perm ssralg zmodp matrix mxalgebra.
+Require Import poly polydiv mxpoly binomial.
+
+(******************************************************************************)
+(* This file provides additional primitives and theory for bivariate *)
+(* polynomials (polynomials of two variables), represented as polynomials *)
+(* with (univariate) polynomial coefficients : *)
+(* 'Y == the (generic) second variable (:= 'X%:P). *)
+(* p^:P == the bivariate polynomial p['X], for p univariate. *)
+(* := map_poly polyC p (this notation is defined in poly.v). *)
+(* u.[x, y] == the bivariate polynomial u evaluated at 'X = x, 'Y = y. *)
+(* := u.[x%:P].[y]. *)
+(* sizeY u == the size of u in 'Y (1 + the 'Y-degree of u, if u != 0). *)
+(* := \max_(i < size u) size u`_i. *)
+(* swapXY u == the bivariate polynomial u['Y, 'X], for u bivariate. *)
+(* poly_XaY p == the bivariate polynomial p['X + 'Y], for p univariate. *)
+(* := p^:P \Po ('X + 'Y). *)
+(* poly_XmY p == the bivariate polynomial p['X * 'Y], for p univariate. *)
+(* := P^:P \Po ('X * 'Y). *)
+(* sub_annihilant p q == for univariate p, q != 0, a nonzero polynomial whose *)
+(* roots include all the differences of roots of p and q, in *)
+(* all field extensions (:= resultant (poly_XaY p) q^:P). *)
+(* div_annihilant p q == for polynomials p != 0, q with q.[0] != 0, a nonzero *)
+(* polynomial whose roots include all the quotients of roots *)
+(* of p by roots of q, in all field extensions *)
+(* (:= resultant (poly_XmY p) q^:P). *)
+(* The latter two "annhilants" provide uniform witnesses for an alternative *)
+(* proof of the closure of the algebraicOver predicate (see mxpoly.v). The *)
+(* fact that the annhilant does not depend on the particular choice of roots *)
+(* of p and q is crucial for the proof of the Primitive Element Theorem (file *)
+(* separable.v). *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Local Open Scope ring_scope.
+Import GRing.Theory.
+
+Local Notation "p ^ f" := (map_poly f p) : ring_scope.
+Local Notation eval := horner_eval.
+
+Notation "'Y" := 'X%:P : ring_scope.
+Notation "p ^:P" := (p ^ polyC) (at level 2, format "p ^:P") : ring_scope.
+Notation "p .[ x , y ]" := (p.[x%:P].[y])
+ (at level 2, left associativity, format "p .[ x , y ]") : ring_scope.
+
+Section PolyXY_Ring.
+
+Variable R : ringType.
+Implicit Types (u : {poly {poly R}}) (p q : {poly R}) (x : R).
+
+Fact swapXY_key : unit. Proof. by []. Qed.
+Definition swapXY_def u : {poly {poly R}} := (u ^ map_poly polyC).['Y].
+Definition swapXY := locked_with swapXY_key swapXY_def.
+Canonical swapXY_unlockable := [unlockable fun swapXY].
+
+Definition sizeY u : nat := \max_(i < size u) (size u`_i).
+Definition poly_XaY p : {poly {poly R}} := p^:P \Po ('X + 'Y).
+Definition poly_XmY p : {poly {poly R}} := p^:P \Po ('X * 'Y).
+Definition sub_annihilant p q := resultant (poly_XaY p) q^:P.
+Definition div_annihilant p q := resultant (poly_XmY p) q^:P.
+
+Lemma swapXY_polyC p : swapXY p%:P = p^:P.
+Proof. by rewrite unlock map_polyC hornerC. Qed.
+
+Lemma swapXY_X : swapXY 'X = 'Y.
+Proof. by rewrite unlock map_polyX hornerX. Qed.
+
+Lemma swapXY_Y : swapXY 'Y = 'X.
+Proof. by rewrite swapXY_polyC map_polyX. Qed.
+
+Lemma swapXY_is_additive : additive swapXY.
+Proof. by move=> u v; rewrite unlock rmorphB !hornerE. Qed.
+Canonical swapXY_addf := Additive swapXY_is_additive.
+
+Lemma coef_swapXY u i j : (swapXY u)`_i`_j = u`_j`_i.
+Proof.
+elim/poly_ind: u => [|u p IHu] in i j *; first by rewrite raddf0 !coef0.
+rewrite raddfD !coefD /= swapXY_polyC coef_map /= !coefC coefMX.
+rewrite !(fun_if (fun q : {poly R} => q`_i)) coef0 -IHu; congr (_ + _).
+by rewrite unlock rmorphM /= map_polyX hornerMX coefMC coefMX.
+Qed.
+
+Lemma swapXYK : involutive swapXY.
+Proof. by move=> u; apply/polyP=> i; apply/polyP=> j; rewrite !coef_swapXY. Qed.
+
+Lemma swapXY_map_polyC p : swapXY p^:P = p%:P.
+Proof. by rewrite -swapXY_polyC swapXYK. Qed.
+
+Lemma swapXY_eq0 u : (swapXY u == 0) = (u == 0).
+Proof. by rewrite (inv_eq swapXYK) raddf0. Qed.
+
+Lemma swapXY_is_multiplicative : multiplicative swapXY.
+Proof.
+split=> [u v|]; last by rewrite swapXY_polyC map_polyC.
+apply/polyP=> i; apply/polyP=> j; rewrite coef_swapXY !coefM !coef_sum.
+rewrite (eq_bigr _ (fun _ _ => coefM _ _ _)) exchange_big /=.
+apply: eq_bigr => j1 _; rewrite coefM; apply: eq_bigr=> i1 _.
+by rewrite !coef_swapXY.
+Qed.
+Canonical swapXY_rmorphism := AddRMorphism swapXY_is_multiplicative.
+
+Lemma swapXY_is_scalable : scalable_for (map_poly polyC \; *%R) swapXY.
+Proof. by move=> p u /=; rewrite -mul_polyC rmorphM /= swapXY_polyC. Qed.
+Canonical swapXY_linear := AddLinear swapXY_is_scalable.
+Canonical swapXY_lrmorphism := [lrmorphism of swapXY].
+
+Lemma swapXY_comp_poly p u : swapXY (p^:P \Po u) = p^:P \Po swapXY u.
+Proof.
+rewrite -horner_map; congr _.[_]; rewrite -!map_poly_comp /=.
+by apply: eq_map_poly => x; rewrite /= swapXY_polyC map_polyC.
+Qed.
+
+Lemma max_size_coefXY u i : size u`_i <= sizeY u.
+Proof.
+have [ltiu | /(nth_default 0)->] := ltnP i (size u); last by rewrite size_poly0.
+exact: (bigmax_sup (Ordinal ltiu)).
+Qed.
+
+Lemma max_size_lead_coefXY u : size (lead_coef u) <= sizeY u.
+Proof. by rewrite lead_coefE max_size_coefXY. Qed.
+
+Lemma max_size_evalX u : size u.['X] <= sizeY u + (size u).-1.
+Proof.
+rewrite horner_coef (leq_trans (size_sum _ _ _)) //; apply/bigmax_leqP=> i _.
+rewrite (leq_trans (size_mul_leq _ _)) // size_polyXn addnS.
+by rewrite leq_add ?max_size_coefXY //= -ltnS (leq_trans _ (leqSpred _)).
+Qed.
+
+Lemma max_size_evalC u x : size u.[x%:P] <= sizeY u.
+Proof.
+rewrite horner_coef (leq_trans (size_sum _ _ _)) //; apply/bigmax_leqP=> i _.
+rewrite (leq_trans (size_mul_leq _ _)) // -polyC_exp size_polyC addnC -subn1.
+by rewrite (leq_trans _ (max_size_coefXY _ i)) // leq_subLR leq_add2r leq_b1.
+Qed.
+
+Lemma sizeYE u : sizeY u = size (swapXY u).
+Proof.
+apply/eqP; rewrite eqn_leq; apply/andP; split.
+ apply/bigmax_leqP=> /= i _; apply/leq_sizeP => j /(nth_default 0) u_j_0.
+ by rewrite -coef_swapXY u_j_0 coef0.
+apply/leq_sizeP=> j le_uY_j; apply/polyP=> i; rewrite coef_swapXY coef0.
+by rewrite nth_default // (leq_trans _ le_uY_j) ?max_size_coefXY.
+Qed.
+
+Lemma sizeY_eq0 u : (sizeY u == 0%N) = (u == 0).
+Proof. by rewrite sizeYE size_poly_eq0 swapXY_eq0. Qed.
+
+Lemma sizeY_mulX u : sizeY (u * 'X) = sizeY u.
+Proof.
+rewrite !sizeYE rmorphM /= swapXY_X rreg_size //.
+by have /monic_comreg[_ /rreg_lead] := monicX R.
+Qed.
+
+Lemma swapXY_poly_XaY p : swapXY (poly_XaY p) = poly_XaY p.
+Proof. by rewrite swapXY_comp_poly rmorphD /= swapXY_X swapXY_Y addrC. Qed.
+
+Lemma swapXY_poly_XmY p : swapXY (poly_XmY p) = poly_XmY p.
+Proof.
+by rewrite swapXY_comp_poly rmorphM /= swapXY_X swapXY_Y commr_polyX.
+Qed.
+
+Lemma poly_XaY0 : poly_XaY 0 = 0.
+Proof. by rewrite /poly_XaY rmorph0 comp_poly0. Qed.
+
+Lemma poly_XmY0 : poly_XmY 0 = 0.
+Proof. by rewrite /poly_XmY rmorph0 comp_poly0. Qed.
+
+End PolyXY_Ring.
+
+Prenex Implicits swapXY sizeY poly_XaY poly_XmY sub_annihilant div_annihilant.
+Prenex Implicits swapXYK.
+
+Lemma swapXY_map (R S : ringType) (f : {additive R -> S}) u :
+ swapXY (u ^ map_poly f) = swapXY u ^ map_poly f.
+Proof.
+by apply/polyP=> i; apply/polyP=> j; rewrite !(coef_map, coef_swapXY).
+Qed.
+
+Section PolyXY_ComRing.
+
+Variable R : comRingType.
+Implicit Types (u : {poly {poly R}}) (p : {poly R}) (x y : R).
+
+Lemma horner_swapXY u x : (swapXY u).[x%:P] = u ^ eval x.
+Proof.
+apply/polyP=> i /=; rewrite coef_map /= /eval horner_coef coef_sum -sizeYE.
+rewrite (horner_coef_wide _ (max_size_coefXY u i)); apply: eq_bigr=> j _.
+by rewrite -polyC_exp coefMC coef_swapXY.
+Qed.
+
+Lemma horner_polyC u x : u.[x%:P] = swapXY u ^ eval x.
+Proof. by rewrite -horner_swapXY swapXYK. Qed.
+
+Lemma horner2_swapXY u x y : (swapXY u).[x, y] = u.[y, x].
+Proof. by rewrite horner_swapXY -{1}(hornerC y x) horner_map. Qed.
+
+Lemma horner_poly_XaY p v : (poly_XaY p).[v] = p \Po (v + 'X).
+Proof. by rewrite horner_comp !hornerE. Qed.
+
+Lemma horner_poly_XmY p v : (poly_XmY p).[v] = p \Po (v * 'X).
+Proof. by rewrite horner_comp !hornerE. Qed.
+
+End PolyXY_ComRing.
+
+Section PolyXY_Idomain.
+
+Variable R : idomainType.
+Implicit Types (p q : {poly R}) (x y : R).
+
+Lemma size_poly_XaY p : size (poly_XaY p) = size p.
+Proof. by rewrite size_comp_poly2 ?size_XaddC // size_map_polyC. Qed.
+
+Lemma poly_XaY_eq0 p : (poly_XaY p == 0) = (p == 0).
+Proof. by rewrite -!size_poly_eq0 size_poly_XaY. Qed.
+
+Lemma size_poly_XmY p : size (poly_XmY p) = size p.
+Proof. by rewrite size_comp_poly2 ?size_XmulC ?polyX_eq0 ?size_map_polyC. Qed.
+
+Lemma poly_XmY_eq0 p : (poly_XmY p == 0) = (p == 0).
+Proof. by rewrite -!size_poly_eq0 size_poly_XmY. Qed.
+
+Lemma lead_coef_poly_XaY p : lead_coef (poly_XaY p) = (lead_coef p)%:P.
+Proof.
+rewrite lead_coef_comp ?size_XaddC // -['Y]opprK -polyC_opp lead_coefXsubC.
+by rewrite expr1n mulr1 lead_coef_map_inj //; apply: polyC_inj.
+Qed.
+
+Lemma sub_annihilant_in_ideal p q :
+ 1 < size p -> 1 < size q ->
+ {uv : {poly {poly R}} * {poly {poly R}}
+ | size uv.1 < size q /\ size uv.2 < size p
+ & forall x y,
+ (sub_annihilant p q).[y] = uv.1.[x, y] * p.[x + y] + uv.2.[x, y] * q.[x]}.
+Proof.
+rewrite -size_poly_XaY -(size_map_polyC q) => p1_gt1 q1_gt1.
+have [uv /= [ub_u ub_v Dr]] := resultant_in_ideal p1_gt1 q1_gt1.
+exists uv => // x y; rewrite -[r in r.[y]](hornerC _ x%:P) Dr.
+by rewrite !(hornerE, horner_comp).
+Qed.
+
+Lemma sub_annihilantP p q x y :
+ p != 0 -> q != 0 -> p.[x] = 0 -> q.[y] = 0 ->
+ (sub_annihilant p q).[x - y] = 0.
+Proof.
+move=> nz_p nz_q px0 qy0.
+have p_gt1: size p > 1 by have /rootP/root_size_gt1-> := px0.
+have q_gt1: size q > 1 by have /rootP/root_size_gt1-> := qy0.
+have [uv /= _ /(_ y)->] := sub_annihilant_in_ideal p_gt1 q_gt1.
+by rewrite (addrC y) subrK px0 qy0 !mulr0 addr0.
+Qed.
+
+Lemma sub_annihilant_neq0 p q : p != 0 -> q != 0 -> sub_annihilant p q != 0.
+Proof.
+rewrite resultant_eq0; set p1 := poly_XaY p => nz_p nz_q.
+have [nz_p1 nz_q1]: p1 != 0 /\ q^:P != 0 by rewrite poly_XaY_eq0 map_polyC_eq0.
+rewrite -leqNgt eq_leq //; apply/eqP/Bezout_coprimepPn=> // [[[u v]]] /=.
+rewrite !size_poly_gt0 -andbA => /and4P[nz_u ltuq nz_v _] Duv.
+have /eqP/= := congr1 (size \o (lead_coef \o swapXY)) Duv.
+rewrite ltn_eqF // !rmorphM !lead_coefM (leq_trans (leq_ltn_trans _ ltuq)) //=.
+ rewrite -{2}[u]swapXYK -sizeYE swapXY_poly_XaY lead_coef_poly_XaY.
+ by rewrite mulrC mul_polyC size_scale ?max_size_lead_coefXY ?lead_coef_eq0.
+rewrite swapXY_map_polyC lead_coefC size_map_polyC.
+set v1 := lead_coef _; have nz_v1: v1 != 0 by rewrite lead_coef_eq0 swapXY_eq0.
+rewrite [in rhs in _ <= rhs]polySpred ?mulf_neq0 // size_mul //.
+by rewrite (polySpred nz_v1) addnC addnS polySpred // ltnS leq_addr.
+Qed.
+
+Lemma div_annihilant_in_ideal p q :
+ 1 < size p -> 1 < size q ->
+ {uv : {poly {poly R}} * {poly {poly R}}
+ | size uv.1 < size q /\ size uv.2 < size p
+ & forall x y,
+ (div_annihilant p q).[y] = uv.1.[x, y] * p.[x * y] + uv.2.[x, y] * q.[x]}.
+Proof.
+rewrite -size_poly_XmY -(size_map_polyC q) => p1_gt1 q1_gt1.
+have [uv /= [ub_u ub_v Dr]] := resultant_in_ideal p1_gt1 q1_gt1.
+exists uv => // x y; rewrite -[r in r.[y]](hornerC _ x%:P) Dr.
+by rewrite !(hornerE, horner_comp).
+Qed.
+
+Lemma div_annihilant_neq0 p q : p != 0 -> q.[0] != 0 -> div_annihilant p q != 0.
+Proof.
+have factorX u: u != 0 -> root u 0 -> exists2 v, v != 0 & u = v * 'X.
+ move=> nz_u /factor_theorem[v]; rewrite subr0 => Du; exists v => //.
+ by apply: contraNneq nz_u => v0; rewrite Du v0 mul0r.
+have nzX: 'X != 0 := monic_neq0 (monicX _); have rootC0 := root_polyC _ 0.
+rewrite resultant_eq0 -leqNgt -rootE // => nz_p nz_q0; apply/eq_leq/eqP.
+have nz_q: q != 0 by apply: contraNneq nz_q0 => ->; rewrite root0.
+apply/Bezout_coprimepPn; rewrite ?map_polyC_eq0 ?poly_XmY_eq0 // => [[uv]].
+rewrite !size_poly_gt0 -andbA ltnNge => /and4P[nz_u /negP ltuq nz_v _] Duv.
+pose u := swapXY uv.1; pose v := swapXY uv.2.
+suffices{ltuq}: size q <= sizeY u by rewrite sizeYE swapXYK -size_map_polyC.
+have{nz_u nz_v} [nz_u nz_v Dvu]: [/\ u != 0, v != 0 & q *: v = u * poly_XmY p].
+ rewrite !swapXY_eq0; split=> //; apply: (can_inj swapXYK).
+ by rewrite linearZ rmorphM /= !swapXYK swapXY_poly_XmY Duv mulrC.
+have{Duv} [n ltvn]: {n | size v < n} by exists (size v).+1.
+elim: n {uv} => // n IHn in p (v) (u) nz_u nz_v Dvu nz_p ltvn *.
+have Dp0: root (poly_XmY p) 0 = root p 0 by rewrite root_comp !hornerE rootC0.
+have Dv0: root u 0 || root p 0 = root v 0 by rewrite -Dp0 -rootM -Dvu rootZ.
+have [v0_0 | nz_v0] := boolP (root v 0); last first.
+ have nz_p0: ~~ root p 0 by apply: contra nz_v0; rewrite -Dv0 orbC => ->.
+ apply: (@leq_trans (size (q * v.[0]))).
+ by rewrite size_mul // (polySpred nz_v0) addnS leq_addr.
+ rewrite -hornerZ Dvu !(horner_comp, hornerE) horner_map mulrC size_Cmul //.
+ by rewrite horner_coef0 max_size_coefXY.
+have [v1 nz_v1 Dv] := factorX _ _ nz_v v0_0; rewrite Dv size_mulX // in ltvn.
+have /orP[/factorX[//|u1 nz_u1 Du] | p0_0]: root u 0 || root p 0 by rewrite Dv0.
+ rewrite Du sizeY_mulX; apply: IHn nz_u1 nz_v1 _ nz_p ltvn.
+ by apply: (mulIf (nzX _)); rewrite mulrAC -scalerAl -Du -Dv.
+have /factorX[|v2 nz_v2 Dv1]: root (swapXY v1) 0; rewrite ?swapXY_eq0 //.
+ suffices: root (swapXY v1 * 'Y) 0 by rewrite mulrC mul_polyC rootZ ?polyX_eq0.
+ have: root (swapXY (q *: v)) 0.
+ by rewrite Dvu rmorphM rootM /= swapXY_poly_XmY Dp0 p0_0 orbT.
+ by rewrite linearZ rootM rootC0 (negPf nz_q0) /= Dv rmorphM /= swapXY_X.
+rewrite ltnS (canRL swapXYK Dv1) -sizeYE sizeY_mulX sizeYE in ltvn.
+have [p1 nz_p1 Dp] := factorX _ _ nz_p p0_0.
+apply: IHn nz_u _ _ nz_p1 ltvn; first by rewrite swapXY_eq0.
+apply: (@mulIf _ ('X * 'Y)); first by rewrite mulf_neq0 ?polyC_eq0 ?nzX.
+rewrite -scalerAl mulrA mulrAC -{1}swapXY_X -rmorphM /= -Dv1 swapXYK -Dv Dvu.
+by rewrite /poly_XmY Dp rmorphM /= map_polyX comp_polyM comp_polyX mulrA.
+Qed.
+
+End PolyXY_Idomain.
+
+Section PolyXY_Field.
+
+Variables (F E : fieldType) (FtoE : {rmorphism F -> E}).
+
+Local Notation pFtoE := (map_poly (GRing.RMorphism.apply FtoE)).
+
+Lemma div_annihilantP (p q : {poly E}) (x y : E) :
+ p != 0 -> q != 0 -> y != 0 -> p.[x] = 0 -> q.[y] = 0 ->
+ (div_annihilant p q).[x / y] = 0.
+Proof.
+move=> nz_p nz_q nz_y px0 qy0.
+have p_gt1: size p > 1 by have /rootP/root_size_gt1-> := px0.
+have q_gt1: size q > 1 by have /rootP/root_size_gt1-> := qy0.
+have [uv /= _ /(_ y)->] := div_annihilant_in_ideal p_gt1 q_gt1.
+by rewrite (mulrC y) divfK // px0 qy0 !mulr0 addr0.
+Qed.
+
+Lemma map_sub_annihilantP (p q : {poly F}) (x y : E) :
+ p != 0 -> q != 0 ->(p ^ FtoE).[x] = 0 -> (q ^ FtoE).[y] = 0 ->
+ (sub_annihilant p q ^ FtoE).[x - y] = 0.
+Proof.
+move=> nz_p nz_q px0 qy0; have pFto0 := map_poly_eq0 FtoE.
+rewrite map_resultant ?pFto0 ?lead_coef_eq0 ?map_poly_eq0 ?poly_XaY_eq0 //.
+rewrite map_comp_poly rmorphD /= map_polyC /= !map_polyX -!map_poly_comp /=.
+by rewrite !(eq_map_poly (map_polyC _)) !map_poly_comp sub_annihilantP ?pFto0.
+Qed.
+
+Lemma map_div_annihilantP (p q : {poly F}) (x y : E) :
+ p != 0 -> q != 0 -> y != 0 -> (p ^ FtoE).[x] = 0 -> (q ^ FtoE).[y] = 0 ->
+ (div_annihilant p q ^ FtoE).[x / y] = 0.
+Proof.
+move=> nz_p nz_q nz_y px0 qy0; have pFto0 := map_poly_eq0 FtoE.
+rewrite map_resultant ?pFto0 ?lead_coef_eq0 ?map_poly_eq0 ?poly_XmY_eq0 //.
+rewrite map_comp_poly rmorphM /= map_polyC /= !map_polyX -!map_poly_comp /=.
+by rewrite !(eq_map_poly (map_polyC _)) !map_poly_comp div_annihilantP ?pFto0.
+Qed.
+
+Lemma root_annihilant x p (pEx := (p ^ pFtoE).[x%:P]) :
+ pEx != 0 -> algebraicOver FtoE x ->
+ exists2 r : {poly F}, r != 0 & forall y, root pEx y -> root (r ^ FtoE) y.
+Proof.
+move=> nz_px [q nz_q qx0].
+have [/size1_polyC Dp | p_gt1] := leqP (size p) 1.
+ by rewrite {}/pEx Dp map_polyC hornerC map_poly_eq0 in nz_px *; exists p`_0.
+have nz_p: p != 0 by rewrite -size_poly_gt0 ltnW.
+elim: {q}_.+1 {-2}q (ltnSn (size q)) => // m IHm q le_qm in nz_q qx0 *.
+have nz_q1: q^:P != 0 by rewrite map_poly_eq0.
+have sz_q1: size q^:P = size q by rewrite size_map_polyC.
+have q1_gt1: size q^:P > 1.
+ by rewrite sz_q1 -(size_map_poly FtoE) (root_size_gt1 _ qx0) ?map_poly_eq0.
+have [uv _ Dr] := resultant_in_ideal p_gt1 q1_gt1; set r := resultant p _ in Dr.
+have /eqP q1x0: (q^:P ^ pFtoE).[x%:P] == 0.
+ by rewrite -swapXY_polyC -swapXY_map horner_swapXY !map_polyC polyC_eq0.
+have [|r_nz] := boolP (r == 0); last first.
+ exists r => // y pxy0; rewrite -[r ^ _](hornerC _ x%:P) -map_polyC Dr.
+ by rewrite rmorphD !rmorphM !hornerE q1x0 mulr0 addr0 rootM pxy0 orbT.
+rewrite resultant_eq0 => /gtn_eqF/Bezout_coprimepPn[]// [q2 p1] /=.
+rewrite size_poly_gt0 sz_q1 => /andP[/andP[nz_q2 ltq2] _] Dq.
+pose n := (size (lead_coef q2)).-1; pose q3 := map_poly (coefp n) q2.
+have nz_q3: q3 != 0 by rewrite map_poly_eq0_id0 ?lead_coef_eq0.
+apply: (IHm q3); rewrite ?(leq_ltn_trans (size_poly _ _)) ?(leq_trans ltq2) //.
+have /polyP/(_ n)/eqP: (q2 ^ pFtoE).[x%:P] = 0.
+apply: (mulIf nz_px); rewrite -hornerM -rmorphM Dq rmorphM hornerM /= q1x0.
+ by rewrite mul0r mulr0.
+rewrite coef0; congr (_ == 0); rewrite !horner_coef coef_sum.
+rewrite size_map_poly !size_map_poly_id0 ?map_poly_eq0 ?lead_coef_eq0 //.
+by apply: eq_bigr => i _; rewrite -rmorphX coefMC !coef_map.
+Qed.
+
+Lemma algebraic_root_polyXY x y :
+ (let pEx p := (p ^ map_poly FtoE).[x%:P] in
+ exists2 p, pEx p != 0 & root (pEx p) y) ->
+ algebraicOver FtoE x -> algebraicOver FtoE y.
+Proof. by case=> p nz_px pxy0 /(root_annihilant nz_px)[r]; exists r; auto. Qed.
+
+End PolyXY_Field.
diff --git a/mathcomp/algebra/polydiv.v b/mathcomp/algebra/polydiv.v
new file mode 100644
index 0000000..eca9a27
--- /dev/null
+++ b/mathcomp/algebra/polydiv.v
@@ -0,0 +1,3418 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype.
+Require Import bigop ssralg poly.
+
+(******************************************************************************)
+(* This file provides a library for the basic theory of Euclidean and pseudo- *)
+(* Euclidean division for polynomials over ring structures. *)
+(* The library defines two versions of the pseudo-euclidean division: one for *)
+(* coefficients in a (not necessarily commutative) ring structure and one for *)
+(* coefficients equipped with a structure of integral domain. From the latter *)
+(* we derive the definition of the usual Euclidean division for coefficients *)
+(* in a field. Only the definition of the pseudo-division for coefficients in *)
+(* an integral domain is exported by default and benefits from notations. *)
+(* Also, the only theory exported by default is the one of division for *)
+(* polynomials with coefficients in a field. *)
+(* Other definitions and facts are qualified using name spaces indicating the *)
+(* hypotheses made on the structure of coefficients and the properties of the *)
+(* polynomial one divides with. *)
+(* *)
+(* Pdiv.Field (exported by the present library): *)
+(* edivp p q == pseudo-division of p by q with p q : {poly R} where *)
+(* R is an idomainType. *)
+(* Computes (k, quo, rem) : nat * {poly r} * {poly R}, *)
+(* such that size rem < size q and: *)
+(* + if lead_coef q is not a unit, then: *)
+(* (lead_coef q ^+ k) *: p = q * quo + rem *)
+(* + else if lead_coef q is a unit, then: *)
+(* p = q * quo + rem and k = 0 *)
+(* p %/ q == quotient (second component) computed by (edivp p q). *)
+(* p %% q == remainder (third component) computed by (edivp p q). *)
+(* scalp p q == exponent (first component) computed by (edivp p q). *)
+(* p %| q == tests the nullity of the remainder of the *)
+(* pseudo-division of p by q. *)
+(* rgcdp p q == Pseudo-greater common divisor obtained by performing *)
+(* the Euclidean algorithm on p and q using redivp as *)
+(* Euclidean division. *)
+(* p %= q == p and q are associate polynomials, i.e., p %| q and *)
+(* q %| p, or equivalently, p = c *: q for some nonzero *)
+(* constant c. *)
+(* gcdp p q == Pseudo-greater common divisor obtained by performing *)
+(* the Euclidean algorithm on p and q using edivp as *)
+(* Euclidean division. *)
+(* egcdp p q == The pair of Bezout coefficients: if e := egcdp p q, *)
+(* then size e.1 <= size q, size e.2 <= size p, and *)
+(* gcdp p q %= e.1 * p + e.2 * q *)
+(* coprimep p q == p and q are coprime, i.e., (gcdp p q) is a nonzero *)
+(* constant. *)
+(* gdcop q p == greatest divisor of p which is coprime to q. *)
+(* irreducible_poly p <-> p has only trivial (constant) divisors. *)
+(* *)
+(* Pdiv.Idomain: theory available for edivp and the related operation under *)
+(* the sole assumption that the ring of coefficients is canonically an *)
+(* integral domain (R : idomainType). *)
+(* *)
+(* Pdiv.IdomainMonic: theory available for edivp and the related operations *)
+(* under the assumption that the ring of coefficients is canonically *)
+(* and integral domain (R : idomainType) an the divisor is monic. *)
+(* *)
+(* Pdiv.IdomainUnit: theory available for edivp and the related operations *)
+(* under the assumption that the ring of coefficients is canonically an *)
+(* integral domain (R : idomainType) and the leading coefficient of the *)
+(* divisor is a unit. *)
+(* *)
+(* Pdiv.ClosedField: theory available for edivp and the related operation *)
+(* under the sole assumption that the ring of coefficients is canonically *)
+(* an algebraically closed field (R : closedField). *)
+(* *)
+(* Pdiv.Ring : *)
+(* redivp p q == pseudo-division of p by q with p q : {poly R} where R is *)
+(* a ringType. *)
+(* Computes (k, quo, rem) : nat * {poly r} * {poly R}, *)
+(* such that if rem = 0 then quo * q = p * (lead_coef q ^+ k) *)
+(* *)
+(* rdivp p q == quotient (second component) computed by (redivp p q). *)
+(* rmodp p q == remainder (third component) computed by (redivp p q). *)
+(* rscalp p q == exponent (first component) computed by (redivp p q). *)
+(* rdvdp p q == tests the nullity of the remainder of the pseudo-division *)
+(* of p by q. *)
+(* rgcdp p q == analogue of gcdp for coefficients in a ringType. *)
+(* rgdcop p q == analogue of gdcop for coefficients in a ringType. *)
+(*rcoprimep p q == analogue of coprimep p q for coefficients in a ringType. *)
+(* *)
+(* Pdiv.RingComRreg : theory of the operations defined in Pdiv.Ring, when the *)
+(* ring of coefficients is canonically commutative (R : comRingType) and *)
+(* the leading coefficient of the divisor is both right regular and *)
+(* commutes as a constant polynomial with the divisor itself *)
+(* *)
+(* Pdiv.RingMonic : theory of the operations defined in Pdiv.Ring, under the *)
+(* assumption that the divisor is monic. *)
+(* *)
+(* Pdiv.UnitRing: theory of the operations defined in Pdiv.Ring, when the *)
+(* ring R of coefficients is canonically with units (R : unitRingType). *)
+(* *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GRing.Theory.
+Local Open Scope ring_scope.
+
+Reserved Notation "p %= q" (at level 70, no associativity).
+
+Local Notation simp := Monoid.simpm.
+
+Module Pdiv.
+
+Module CommonRing.
+
+Section RingPseudoDivision.
+
+Variable R : ringType.
+Implicit Types d p q r : {poly R}.
+
+(* Pseudo division, defined on an arbitrary ring *)
+Definition redivp_rec (q : {poly R}) :=
+ let sq := size q in
+ let cq := lead_coef q in
+ fix loop (k : nat) (qq r : {poly R})(n : nat) {struct n} :=
+ if size r < sq then (k, qq, r) else
+ let m := (lead_coef r) *: 'X^(size r - sq) in
+ let qq1 := qq * cq%:P + m in
+ let r1 := r * cq%:P - m * q in
+ if n is n1.+1 then loop k.+1 qq1 r1 n1 else (k.+1, qq1, r1).
+
+Definition redivp_expanded_def p q :=
+ if q == 0 then (0%N, 0, p) else redivp_rec q 0 0 p (size p).
+Fact redivp_key : unit. Proof. by []. Qed.
+Definition redivp : {poly R} -> {poly R} -> nat * {poly R} * {poly R} :=
+ locked_with redivp_key redivp_expanded_def.
+Canonical redivp_unlockable := [unlockable fun redivp].
+
+Definition rdivp p q := ((redivp p q).1).2.
+Definition rmodp p q := (redivp p q).2.
+Definition rscalp p q := ((redivp p q).1).1.
+Definition rdvdp p q := rmodp q p == 0.
+(*Definition rmultp := [rel m d | rdvdp d m].*)
+Lemma redivp_def p q : redivp p q = (rscalp p q, rdivp p q, rmodp p q).
+Proof. by rewrite /rscalp /rdivp /rmodp; case: (redivp p q) => [[]] /=. Qed.
+
+Lemma rdiv0p p : rdivp 0 p = 0.
+Proof.
+rewrite /rdivp unlock; case: ifP => // Hp; rewrite /redivp_rec !size_poly0.
+by rewrite polySpred ?Hp.
+Qed.
+
+Lemma rdivp0 p : rdivp p 0 = 0.
+Proof. by rewrite /rdivp unlock eqxx. Qed.
+
+Lemma rdivp_small p q : size p < size q -> rdivp p q = 0.
+Proof.
+rewrite /rdivp unlock; case: eqP => Eq; first by rewrite Eq size_poly0.
+by case sp: (size p) => [| s] hs /=; rewrite sp hs.
+Qed.
+
+Lemma leq_rdivp p q : (size (rdivp p q) <= size p).
+Proof.
+case: (ltnP (size p) (size q)); first by move/rdivp_small->; rewrite size_poly0.
+rewrite /rdivp /rmodp /rscalp unlock; case q0 : (q == 0) => /=.
+ by rewrite size_poly0.
+have : (size (0 : {poly R})) <= size p by rewrite size_poly0.
+move: (leqnn (size p)); move: {2 3 4 6}(size p)=> A.
+elim: (size p) 0%N (0 : {poly R}) {1 3 4}p (leqnn (size p)) => [| n ihn] k q1 r.
+ by move/size_poly_leq0P->; rewrite /= size_poly0 lt0n size_poly_eq0 q0.
+move=> /= hrn hr hq1 hq; case: ltnP=> //= hqr.
+have sr: 0 < size r by apply: leq_trans hqr; rewrite size_poly_gt0 q0.
+have sq: 0 < size q by rewrite size_poly_gt0 q0.
+apply: ihn => //.
+- apply/leq_sizeP => j hnj.
+ rewrite coefB -scalerAl coefZ coefXnM ltn_subRL ltnNge.
+ have hj : (size r).-1 <= j.
+ by apply: leq_trans hnj; move: hrn; rewrite -{1}(prednK sr) ltnS.
+ rewrite polySpred -?size_poly_gt0 // (leq_ltn_trans hj) /=; last first.
+ by rewrite -{1}(add0n j) ltn_add2r.
+ move: (hj); rewrite leq_eqVlt; case/orP.
+ move/eqP<-; rewrite (@polySpred _ q) ?q0 // subSS coefMC.
+ rewrite subKn; first by rewrite lead_coefE subrr.
+ by rewrite -ltnS -!polySpred // ?q0 -?size_poly_gt0.
+ move=> {hj} hj; move: (hj); rewrite prednK // coefMC; move/leq_sizeP=> -> //.
+ suff: size q <= j - (size r - size q).
+ by rewrite mul0r sub0r; move/leq_sizeP=> -> //; rewrite mulr0 oppr0.
+ rewrite subnBA // addnC -(prednK sq) -(prednK sr) addSn subSS.
+ by rewrite -addnBA ?(ltnW hj) // -{1}[_.-1]addn0 ltn_add2l subn_gt0.
+- apply: leq_trans (size_add _ _) _; rewrite geq_max; apply/andP; split.
+ apply: leq_trans (size_mul_leq _ _) _.
+ by rewrite size_polyC lead_coef_eq0 q0 /= addn1.
+ rewrite size_opp; apply: leq_trans (size_mul_leq _ _) _.
+ apply: leq_trans hr; rewrite -subn1 leq_subLR -{2}(subnK hqr) addnA leq_add2r.
+ by rewrite add1n -(@size_polyXn R) size_scale_leq.
+apply: leq_trans (size_add _ _) _; rewrite geq_max; apply/andP; split.
+ apply: leq_trans (size_mul_leq _ _) _.
+ by rewrite size_polyC lead_coef_eq0 q0 /= addnS addn0.
+apply: leq_trans (size_scale_leq _ _) _; rewrite size_polyXn.
+by rewrite -subSn // leq_subLR -add1n leq_add.
+Qed.
+
+Lemma rmod0p p : rmodp 0 p = 0.
+Proof.
+rewrite /rmodp unlock; case: ifP => // Hp; rewrite /redivp_rec !size_poly0.
+by rewrite polySpred ?Hp.
+Qed.
+
+Lemma rmodp0 p : rmodp p 0 = p.
+Proof. by rewrite /rmodp unlock eqxx. Qed.
+
+Lemma rscalp_small p q : size p < size q -> rscalp p q = 0%N.
+Proof.
+rewrite /rscalp unlock; case: eqP => Eq // spq.
+by case sp: (size p) => [| s] /=; rewrite spq.
+Qed.
+
+Lemma ltn_rmodp p q : (size (rmodp p q) < size q) = (q != 0).
+Proof.
+rewrite /rdivp /rmodp /rscalp unlock; case q0 : (q == 0).
+ by rewrite (eqP q0) /= size_poly0 ltn0.
+elim: (size p) 0%N 0 {1 3}p (leqnn (size p)) => [|n ihn] k q1 r.
+ rewrite leqn0 size_poly_eq0; move/eqP->; rewrite /= size_poly0 /= lt0n.
+ by rewrite size_poly_eq0 q0 /= size_poly0 lt0n size_poly_eq0 q0.
+move=> hr /=; case: (@ltnP (size r) _) => //= hsrq; rewrite ihn //.
+apply/leq_sizeP => j hnj; rewrite coefB.
+have sr: 0 < size r.
+ by apply: leq_trans hsrq; apply: neq0_lt0n; rewrite size_poly_eq0.
+have sq: 0 < size q by rewrite size_poly_gt0 q0.
+have hj : (size r).-1 <= j.
+ by apply: leq_trans hnj; move: hr; rewrite -{1}(prednK sr) ltnS.
+rewrite -scalerAl !coefZ coefXnM ltn_subRL ltnNge; move: (sr).
+move/prednK => {1}<-.
+have -> /= : (size r).-1 < size q + j.
+ apply: (@leq_trans ((size q) + (size r).-1)); last by rewrite leq_add2l.
+ by rewrite -{1}[_.-1]add0n ltn_add2r.
+move: (hj); rewrite leq_eqVlt; case/orP.
+ move/eqP<-; rewrite -{1}(prednK sq) -{3}(prednK sr) subSS.
+ rewrite subKn; first by rewrite coefMC !lead_coefE subrr.
+ by move: hsrq; rewrite -{1}(prednK sq) -{1}(prednK sr) ltnS.
+move=> {hj} hj; move: (hj); rewrite prednK // coefMC; move/leq_sizeP=> -> //.
+suff: size q <= j - (size r - size q).
+ by rewrite mul0r sub0r; move/leq_sizeP=> -> //; rewrite mulr0 oppr0.
+rewrite subnBA // addnC -(prednK sq) -(prednK sr) addSn subSS.
+by rewrite -addnBA ?(ltnW hj) // -{1}[_.-1]addn0 ltn_add2l subn_gt0.
+Qed.
+
+Lemma ltn_rmodpN0 p q : q != 0 -> size (rmodp p q) < size q.
+Proof. by rewrite ltn_rmodp. Qed.
+
+Lemma rmodp1 p : rmodp p 1 = 0.
+Proof.
+case p0: (p == 0); first by rewrite (eqP p0) rmod0p.
+apply/eqP; rewrite -size_poly_eq0.
+by have := (ltn_rmodp p 1); rewrite size_polyC !oner_neq0 ltnS leqn0.
+Qed.
+
+Lemma rmodp_small p q : size p < size q -> rmodp p q = p.
+Proof.
+rewrite /rmodp unlock; case: eqP => Eq; first by rewrite Eq size_poly0.
+by case sp: (size p) => [| s] Hs /=; rewrite sp Hs /=.
+Qed.
+
+Lemma leq_rmodp m d : size (rmodp m d) <= size m.
+Proof.
+case: (ltnP (size m) (size d)) => [|h]; first by move/rmodp_small->.
+case d0: (d == 0); first by rewrite (eqP d0) rmodp0.
+by apply: leq_trans h; apply: ltnW; rewrite ltn_rmodp d0.
+Qed.
+
+Lemma rmodpC p c : c != 0 -> rmodp p c%:P = 0.
+Proof.
+move=> Hc; apply/eqP; rewrite -size_poly_eq0 -leqn0 -ltnS.
+have -> : 1%N = nat_of_bool (c != 0) by rewrite Hc.
+by rewrite -size_polyC ltn_rmodp polyC_eq0.
+Qed.
+
+Lemma rdvdp0 d : rdvdp d 0.
+Proof. by rewrite /rdvdp rmod0p. Qed.
+
+Lemma rdvd0p n : (rdvdp 0 n) = (n == 0).
+Proof. by rewrite /rdvdp rmodp0. Qed.
+
+Lemma rdvd0pP n : reflect (n = 0) (rdvdp 0 n).
+Proof. by apply: (iffP idP); rewrite rdvd0p; move/eqP. Qed.
+
+Lemma rdvdpN0 p q : rdvdp p q -> q != 0 -> p != 0.
+Proof. by move=> pq hq; apply: contraL pq => /eqP ->; rewrite rdvd0p. Qed.
+
+Lemma rdvdp1 d : (rdvdp d 1) = ((size d) == 1%N).
+Proof.
+rewrite /rdvdp; case d0: (d == 0).
+ by rewrite (eqP d0) rmodp0 size_poly0 (negPf (@oner_neq0 _)).
+have:= (size_poly_eq0 d); rewrite d0; move/negbT; rewrite -lt0n.
+rewrite leq_eqVlt; case/orP => hd; last first.
+ by rewrite rmodp_small ?size_poly1 // oner_eq0 -(subnKC hd).
+rewrite eq_sym in hd; rewrite hd; have [c cn0 ->] := size_poly1P _ hd.
+rewrite /rmodp unlock -size_poly_eq0 size_poly1 /= size_poly1 size_polyC cn0 /=.
+by rewrite polyC_eq0 (negPf cn0) !lead_coefC !scale1r subrr !size_poly0.
+Qed.
+
+Lemma rdvd1p m : rdvdp 1 m.
+Proof. by rewrite /rdvdp rmodp1. Qed.
+
+Lemma Nrdvdp_small (n d : {poly R}) :
+ n != 0 -> size n < size d -> (rdvdp d n) = false.
+Proof.
+by move=> nn0 hs; rewrite /rdvdp; rewrite (rmodp_small hs); apply: negPf.
+Qed.
+
+Lemma rmodp_eq0P p q : reflect (rmodp p q = 0) (rdvdp q p).
+Proof. exact: (iffP eqP). Qed.
+
+Lemma rmodp_eq0 p q : rdvdp q p -> rmodp p q = 0.
+Proof. by move/rmodp_eq0P. Qed.
+
+Lemma rdvdp_leq p q : rdvdp p q -> q != 0 -> size p <= size q.
+Proof. by move=> dvd_pq; rewrite leqNgt; apply: contra => /rmodp_small <-. Qed.
+
+Definition rgcdp p q :=
+ let: (p1, q1) := if size p < size q then (q, p) else (p, q) in
+ if p1 == 0 then q1 else
+ let fix loop (n : nat) (pp qq : {poly R}) {struct n} :=
+ let rr := rmodp pp qq in
+ if rr == 0 then qq else
+ if n is n1.+1 then loop n1 qq rr else rr in
+ loop (size p1) p1 q1.
+
+Lemma rgcd0p : left_id 0 rgcdp.
+Proof.
+move=> p; rewrite /rgcdp size_poly0 size_poly_gt0 if_neg.
+case: ifP => /= [_ | nzp]; first by rewrite eqxx.
+by rewrite polySpred !(rmodp0, nzp) //; case: _.-1 => [|m]; rewrite rmod0p eqxx.
+Qed.
+
+Lemma rgcdp0 : right_id 0 rgcdp.
+Proof.
+move=> p; have:= rgcd0p p; rewrite /rgcdp size_poly0 size_poly_gt0 if_neg.
+by case: ifP => /= p0; rewrite ?(eqxx, p0) // (eqP p0).
+Qed.
+
+Lemma rgcdpE p q :
+ rgcdp p q = if size p < size q
+ then rgcdp (rmodp q p) p else rgcdp (rmodp p q) q.
+Proof.
+pose rgcdp_rec := fix rgcdp_rec (n : nat) (pp qq : {poly R}) {struct n} :=
+ let rr := rmodp pp qq in
+ if rr == 0 then qq else
+ if n is n1.+1 then rgcdp_rec n1 qq rr else rr.
+have Irec: forall m n p q, size q <= m -> size q <= n
+ -> size q < size p -> rgcdp_rec m p q = rgcdp_rec n p q.
+ + elim=> [|m Hrec] [|n] //= p1 q1.
+ - rewrite leqn0 size_poly_eq0; move/eqP=> -> _.
+ rewrite size_poly0 size_poly_gt0 rmodp0 => nzp.
+ by rewrite (negPf nzp); case: n => [|n] /=; rewrite rmod0p eqxx.
+ - rewrite leqn0 size_poly_eq0 => _; move/eqP=> ->.
+ rewrite size_poly0 size_poly_gt0 rmodp0 => nzp.
+ by rewrite (negPf nzp); case: m {Hrec} => [|m] /=; rewrite rmod0p eqxx.
+ case: ifP => Epq Sm Sn Sq //; rewrite ?Epq //.
+ case: (eqVneq q1 0) => [->|nzq].
+ by case: n m {Sm Sn Hrec} => [|m] [|n] //=; rewrite rmod0p eqxx.
+ apply: Hrec; last by rewrite ltn_rmodp.
+ by rewrite -ltnS (leq_trans _ Sm) // ltn_rmodp.
+ by rewrite -ltnS (leq_trans _ Sn) // ltn_rmodp.
+case: (eqVneq p 0) => [-> | nzp].
+ by rewrite rmod0p rmodp0 rgcd0p rgcdp0 if_same.
+case: (eqVneq q 0) => [-> | nzq].
+ by rewrite rmod0p rmodp0 rgcd0p rgcdp0 if_same.
+rewrite /rgcdp -/rgcdp_rec.
+case: ltnP; rewrite (negPf nzp, negPf nzq) //=.
+ move=> ltpq; rewrite ltn_rmodp (negPf nzp) //=.
+ rewrite -(ltn_predK ltpq) /=; case: eqP => [->|].
+ by case: (size p) => [|[|s]]; rewrite /= rmodp0 (negPf nzp) // rmod0p eqxx.
+ move/eqP=> nzqp; rewrite (negPf nzp).
+ apply: Irec => //; last by rewrite ltn_rmodp.
+ by rewrite -ltnS (ltn_predK ltpq) (leq_trans _ ltpq) ?leqW // ltn_rmodp.
+ by rewrite ltnW // ltn_rmodp.
+move=> leqp; rewrite ltn_rmodp (negPf nzq) //=.
+have p_gt0: size p > 0 by rewrite size_poly_gt0.
+rewrite -(prednK p_gt0) /=; case: eqP => [->|].
+ by case: (size q) => [|[|s]]; rewrite /= rmodp0 (negPf nzq) // rmod0p eqxx.
+move/eqP=> nzpq; rewrite (negPf nzq).
+apply: Irec => //; last by rewrite ltn_rmodp.
+ by rewrite -ltnS (prednK p_gt0) (leq_trans _ leqp) // ltn_rmodp.
+by rewrite ltnW // ltn_rmodp.
+Qed.
+
+CoInductive comm_redivp_spec m d : nat * {poly R} * {poly R} -> Type :=
+ ComEdivnSpec k (q r : {poly R}) of
+ (GRing.comm d (lead_coef d)%:P -> m * (lead_coef d ^+ k)%:P = q * d + r) &
+ (d != 0 -> size r < size d) : comm_redivp_spec m d (k, q, r).
+
+Lemma comm_redivpP m d : comm_redivp_spec m d (redivp m d).
+Proof.
+rewrite unlock; case: (altP (d =P 0))=> [->| Hd].
+ by constructor; rewrite !(simp, eqxx).
+have: GRing.comm d (lead_coef d)%:P -> m * (lead_coef d ^+ 0)%:P = 0 * d + m.
+ by rewrite !simp.
+elim: (size m) 0%N 0 {1 4 6}m (leqnn (size m))=>
+ [|n IHn] k q r Hr /=.
+ have{Hr} ->: r = 0 by apply/eqP; rewrite -size_poly_eq0; move: Hr; case: size.
+ suff hsd: size (0: {poly R}) < size d by rewrite hsd => /= ?; constructor.
+ by rewrite size_polyC eqxx (polySpred Hd).
+case: ltP=> Hlt Heq; first by constructor=> // _; apply/ltP.
+apply: IHn=> [|Cda]; last first.
+ rewrite mulrDl addrAC -addrA subrK exprSr polyC_mul mulrA Heq //.
+ by rewrite mulrDl -mulrA Cda mulrA.
+apply/leq_sizeP => j Hj.
+rewrite coefD coefN coefMC -scalerAl coefZ coefXnM.
+move/ltP: Hlt; rewrite -leqNgt=> Hlt.
+move: Hj; rewrite leq_eqVlt; case/predU1P => [<-{j} | Hj]; last first.
+ rewrite nth_default ?(leq_trans Hqq) // ?simp; last by apply: (leq_trans Hr).
+ rewrite nth_default; first by rewrite if_same !simp oppr0.
+ by rewrite -{1}(subKn Hlt) leq_sub2r // (leq_trans Hr).
+move: Hr; rewrite leq_eqVlt ltnS; case/predU1P=> Hqq; last first.
+ rewrite !nth_default ?if_same ?simp ?oppr0 //.
+ by rewrite -{1}(subKn Hlt) leq_sub2r // (leq_trans Hqq).
+rewrite {2}/lead_coef Hqq polySpred // subSS ltnNge leq_subr /=.
+by rewrite subKn ?addrN // -subn1 leq_subLR add1n -Hqq.
+Qed.
+
+Lemma rmodpp p : GRing.comm p (lead_coef p)%:P -> rmodp p p = 0.
+Proof.
+move=> hC; rewrite /rmodp unlock; case: ifP => hp /=; first by rewrite (eqP hp).
+move: (hp); rewrite -size_poly_eq0 /redivp_rec; case sp: (size p)=> [|n] // _.
+rewrite mul0r sp ltnn add0r subnn expr0 hC alg_polyC subrr.
+by case: n sp => [|n] sp; rewrite size_polyC /= eqxx.
+Qed.
+
+Definition rcoprimep (p q : {poly R}) := size (rgcdp p q) == 1%N.
+
+Fixpoint rgdcop_rec q p n :=
+ if n is m.+1 then
+ if rcoprimep p q then p
+ else rgdcop_rec q (rdivp p (rgcdp p q)) m
+ else (q == 0)%:R.
+
+Definition rgdcop q p := rgdcop_rec q p (size p).
+
+Lemma rgdcop0 q : rgdcop q 0 = (q == 0)%:R.
+Proof. by rewrite /rgdcop size_poly0. Qed.
+
+End RingPseudoDivision.
+
+End CommonRing.
+
+Module RingComRreg.
+
+Import CommonRing.
+
+Section ComRegDivisor.
+
+Variable R : ringType.
+Variable d : {poly R}.
+Hypothesis Cdl : GRing.comm d (lead_coef d)%:P.
+Hypothesis Rreg : GRing.rreg (lead_coef d).
+
+Implicit Types p q r : {poly R}.
+
+Lemma redivp_eq q r :
+ size r < size d ->
+ let k := (redivp (q * d + r) d).1.1 in
+ let c := (lead_coef d ^+ k)%:P in
+ redivp (q * d + r) d = (k, q * c, r * c).
+Proof.
+move=> lt_rd; case: comm_redivpP=> k q1 r1; move/(_ Cdl)=> Heq.
+have: d != 0 by case: (size d) lt_rd (size_poly_eq0 d) => // n _ <-.
+move=> dn0; move/(_ dn0)=> Hs.
+have eC : q * d * (lead_coef d ^+ k)%:P = q * (lead_coef d ^+ k)%:P * d.
+ by rewrite -mulrA polyC_exp (GRing.commrX k Cdl) mulrA.
+suff e1 : q1 = q * (lead_coef d ^+ k)%:P.
+ congr (_, _, _) => //=; move/eqP: Heq; rewrite [_ + r1]addrC.
+ rewrite -subr_eq; move/eqP<-; rewrite e1 mulrDl addrAC -{2}(add0r (r * _)).
+ by rewrite eC subrr add0r.
+have : (q1 - q * (lead_coef d ^+ k)%:P) * d = r * (lead_coef d ^+ k)%:P - r1.
+ apply: (@addIr _ r1); rewrite subrK.
+ apply: (@addrI _ ((q * (lead_coef d ^+ k)%:P) * d)).
+ by rewrite mulrDl mulNr !addrA [_ + (q1 * d)]addrC addrK -eC -mulrDl.
+move/eqP; rewrite -[_ == _ - _]subr_eq0 rreg_div0 //.
+ by case/andP; rewrite subr_eq0; move/eqP.
+rewrite size_opp; apply: (leq_ltn_trans (size_add _ _)); rewrite size_opp.
+rewrite gtn_max Hs (leq_ltn_trans (size_mul_leq _ _)) //.
+rewrite size_polyC; case: (_ == _); last by rewrite addnS addn0.
+by rewrite addn0; apply: leq_ltn_trans lt_rd; case: size.
+Qed.
+
+(* this is a bad name *)
+Lemma rdivp_eq p :
+ p * (lead_coef d ^+ (rscalp p d))%:P = (rdivp p d) * d + (rmodp p d).
+Proof.
+rewrite /rdivp /rmodp /rscalp; case: comm_redivpP=> k q1 r1 Hc _; exact: Hc.
+Qed.
+
+(* section variables impose an inconvenient order on parameters *)
+Lemma eq_rdvdp k q1 p:
+ p * ((lead_coef d)^+ k)%:P = q1 * d -> rdvdp d p.
+Proof.
+move=> he.
+have Hnq0 := rreg_lead0 Rreg; set lq := lead_coef d.
+pose v := rscalp p d; pose m := maxn v k.
+rewrite /rdvdp -(rreg_polyMC_eq0 _ (@rregX _ _ (m - v) Rreg)).
+suff:
+ ((rdivp p d) * (lq ^+ (m - v))%:P - q1 * (lq ^+ (m - k))%:P) * d +
+ (rmodp p d) * (lq ^+ (m - v))%:P == 0.
+ rewrite rreg_div0 //; first by case/andP.
+ by rewrite rreg_size ?ltn_rmodp //; apply rregX.
+rewrite mulrDl addrAC mulNr -!mulrA polyC_exp -(GRing.commrX (m-v) Cdl).
+rewrite -polyC_exp mulrA -mulrDl -rdivp_eq // [(_ ^+ (m - k))%:P]polyC_exp.
+rewrite -(GRing.commrX (m-k) Cdl) -polyC_exp mulrA -he -!mulrA -!polyC_mul.
+rewrite -/v -!exprD addnC subnK ?leq_maxl //.
+by rewrite addnC subnK ?subrr ?leq_maxr.
+Qed.
+
+CoInductive rdvdp_spec p q : {poly R} -> bool -> Type :=
+ | Rdvdp k q1 & p * ((lead_coef q)^+ k)%:P = q1 * q : rdvdp_spec p q 0 true
+ | RdvdpN & rmodp p q != 0 : rdvdp_spec p q (rmodp p q) false.
+
+(* Is that version useable ? *)
+
+Lemma rdvdp_eqP p : rdvdp_spec p d (rmodp p d) (rdvdp d p).
+Proof.
+case hdvd: (rdvdp d p); last by apply: RdvdpN; move/rmodp_eq0P/eqP: hdvd.
+move/rmodp_eq0P: (hdvd)->; apply: (@Rdvdp _ _ (rscalp p d) (rdivp p d)).
+by rewrite rdivp_eq //; move/rmodp_eq0P: (hdvd)->; rewrite addr0.
+Qed.
+
+Lemma rdvdp_mull p : rdvdp d (p * d).
+Proof. by apply: (@eq_rdvdp 0%N p); rewrite expr0 mulr1. Qed.
+
+Lemma rmodp_mull p : rmodp (p * d) d = 0.
+Proof.
+case: (d =P 0)=> Hd; first by rewrite Hd simp rmod0p.
+by apply/eqP; apply: rdvdp_mull.
+Qed.
+
+Lemma rmodpp : rmodp d d = 0.
+Proof. by rewrite -{1}(mul1r d) rmodp_mull. Qed.
+
+Lemma rdivpp : rdivp d d = (lead_coef d ^+ rscalp d d)%:P.
+have dn0 : d != 0 by rewrite -lead_coef_eq0 rreg_neq0.
+move: (rdivp_eq d); rewrite rmodpp addr0.
+suff ->: GRing.comm d (lead_coef d ^+ rscalp d d)%:P by move/(rreg_lead Rreg)->.
+by rewrite polyC_exp; apply: commrX.
+Qed.
+
+Lemma rdvdpp : rdvdp d d.
+Proof. apply/eqP; exact: rmodpp. Qed.
+
+Lemma rdivpK p : rdvdp d p ->
+ (rdivp p d) * d = p * (lead_coef d ^+ rscalp p d)%:P.
+Proof. by rewrite rdivp_eq /rdvdp; move/eqP->; rewrite addr0. Qed.
+
+End ComRegDivisor.
+
+End RingComRreg.
+
+Module RingMonic.
+
+Import CommonRing.
+
+Import RingComRreg.
+
+Section MonicDivisor.
+
+Variable R : ringType.
+Implicit Types p q r : {poly R}.
+
+
+Variable d : {poly R}.
+Hypothesis mond : d \is monic.
+
+Lemma redivp_eq q r : size r < size d ->
+ let k := (redivp (q * d + r) d).1.1 in
+ redivp (q * d + r) d = (k, q, r).
+Proof.
+case: (monic_comreg mond)=> Hc Hr; move/(redivp_eq Hc Hr q).
+by rewrite (eqP mond); move=> -> /=; rewrite expr1n !mulr1.
+Qed.
+
+Lemma rdivp_eq p :
+ p = (rdivp p d) * d + (rmodp p d).
+Proof.
+rewrite -rdivp_eq; rewrite (eqP mond); last exact: commr1.
+by rewrite expr1n mulr1.
+Qed.
+
+Lemma rdivpp : rdivp d d = 1.
+Proof.
+by case: (monic_comreg mond) => hc hr; rewrite rdivpp // (eqP mond) expr1n.
+Qed.
+
+Lemma rdivp_addl_mul_small q r :
+ size r < size d -> rdivp (q * d + r) d = q.
+Proof.
+by move=> Hd; case: (monic_comreg mond)=> Hc Hr; rewrite /rdivp redivp_eq.
+Qed.
+
+Lemma rdivp_addl_mul q r : rdivp (q * d + r) d = q + rdivp r d.
+Proof.
+case: (monic_comreg mond)=> Hc Hr; rewrite {1}(rdivp_eq r) addrA.
+by rewrite -mulrDl rdivp_addl_mul_small // ltn_rmodp monic_neq0.
+Qed.
+
+Lemma rdivp_addl q r :
+ rdvdp d q -> rdivp (q + r) d = rdivp q d + rdivp r d.
+Proof.
+case: (monic_comreg mond)=> Hc Hr; rewrite {1}(rdivp_eq r) addrA.
+rewrite {2}(rdivp_eq q); move/rmodp_eq0P->; rewrite addr0.
+by rewrite -mulrDl rdivp_addl_mul_small // ltn_rmodp monic_neq0.
+Qed.
+
+Lemma rdivp_addr q r :
+ rdvdp d r -> rdivp (q + r) d = rdivp q d + rdivp r d.
+Proof. by rewrite addrC; move/rdivp_addl->; rewrite addrC. Qed.
+
+Lemma rdivp_mull p : rdivp (p * d) d = p.
+Proof. by rewrite -[p * d]addr0 rdivp_addl_mul rdiv0p addr0. Qed.
+
+Lemma rmodp_mull p : rmodp (p * d) d = 0.
+Proof.
+apply: rmodp_mull; rewrite (eqP mond); [exact: commr1 | exact: rreg1].
+Qed.
+
+Lemma rmodpp : rmodp d d = 0.
+Proof.
+apply: rmodpp; rewrite (eqP mond); [exact: commr1 | exact: rreg1].
+Qed.
+
+Lemma rmodp_addl_mul_small q r :
+ size r < size d -> rmodp (q * d + r) d = r.
+Proof.
+by move=> Hd; case: (monic_comreg mond)=> Hc Hr; rewrite /rmodp redivp_eq.
+Qed.
+
+Lemma rmodp_add p q : rmodp (p + q) d = rmodp p d + rmodp q d.
+Proof.
+rewrite {1}(rdivp_eq p) {1}(rdivp_eq q).
+rewrite addrCA 2!addrA -mulrDl (addrC (rdivp q d)) -addrA.
+rewrite rmodp_addl_mul_small //; apply: (leq_ltn_trans (size_add _ _)).
+by rewrite gtn_max !ltn_rmodp // monic_neq0.
+Qed.
+
+Lemma rmodp_mulmr p q : rmodp (p * (rmodp q d)) d = rmodp (p * q) d.
+Proof.
+have -> : rmodp q d = q - (rdivp q d) * d.
+ by rewrite {2}(rdivp_eq q) addrAC subrr add0r.
+rewrite mulrDr rmodp_add -mulNr mulrA.
+rewrite -{2}[rmodp _ _]addr0; congr (_ + _); exact: rmodp_mull.
+Qed.
+
+Lemma rdvdpp : rdvdp d d.
+Proof.
+apply: rdvdpp; rewrite (eqP mond); [exact: commr1 | exact: rreg1].
+Qed.
+
+(* section variables impose an inconvenient order on parameters *)
+Lemma eq_rdvdp q1 p : p = q1 * d -> rdvdp d p.
+Proof.
+(* this probably means I need to specify impl args for comm_rref_rdvdp *)
+move=> h; apply: (@eq_rdvdp _ _ _ _ 1%N q1); rewrite (eqP mond).
+- exact: commr1.
+- exact: rreg1.
+by rewrite expr1n mulr1.
+Qed.
+
+Lemma rdvdp_mull p : rdvdp d (p * d).
+Proof.
+apply: rdvdp_mull; rewrite (eqP mond) //; [exact: commr1 | exact: rreg1].
+Qed.
+
+Lemma rdvdpP p : reflect (exists qq, p = qq * d) (rdvdp d p).
+Proof.
+case: (monic_comreg mond)=> Hc Hr; apply: (iffP idP).
+ case: rdvdp_eqP=> // k qq; rewrite (eqP mond) expr1n mulr1; move=> -> _.
+ by exists qq.
+by case=> [qq]; move/eq_rdvdp.
+Qed.
+
+Lemma rdivpK p : rdvdp d p -> (rdivp p d) * d = p.
+Proof. by move=> dvddp; rewrite {2}[p]rdivp_eq rmodp_eq0 ?addr0. Qed.
+
+End MonicDivisor.
+End RingMonic.
+
+Module Ring.
+
+Include CommonRing.
+Import RingMonic.
+
+Section ExtraMonicDivisor.
+
+Variable R : ringType.
+
+Implicit Types d p q r : {poly R}.
+
+Lemma rdivp1 p : rdivp p 1 = p.
+Proof. by rewrite -{1}(mulr1 p) rdivp_mull // monic1. Qed.
+
+Lemma rdvdp_XsubCl p x : rdvdp ('X - x%:P) p = root p x.
+Proof.
+have [HcX Hr] := (monic_comreg (monicXsubC x)).
+apply/rmodp_eq0P/factor_theorem; last first.
+ case=> p1 ->; apply: rmodp_mull; exact: monicXsubC.
+move=> e0; exists (rdivp p ('X - x%:P)).
+by rewrite {1}(rdivp_eq (monicXsubC x) p) e0 addr0.
+Qed.
+
+Lemma polyXsubCP p x : reflect (p.[x] = 0) (rdvdp ('X - x%:P) p).
+Proof. by apply: (iffP idP); rewrite rdvdp_XsubCl; move/rootP. Qed.
+
+
+Lemma root_factor_theorem p x : root p x = (rdvdp ('X - x%:P) p).
+Proof. by rewrite rdvdp_XsubCl. Qed.
+
+End ExtraMonicDivisor.
+
+End Ring.
+
+Module ComRing.
+
+Import Ring.
+
+Import RingComRreg.
+
+Section CommutativeRingPseudoDivision.
+
+Variable R : comRingType.
+
+Implicit Types d p q m n r : {poly R}.
+
+CoInductive redivp_spec (m d : {poly R}) : nat * {poly R} * {poly R} -> Type :=
+ EdivnSpec k (q r: {poly R}) of
+ (lead_coef d ^+ k) *: m = q * d + r &
+ (d != 0 -> size r < size d) : redivp_spec m d (k, q, r).
+
+
+Lemma redivpP m d : redivp_spec m d (redivp m d).
+Proof.
+rewrite redivp_def; constructor; last by move=> dn0; rewrite ltn_rmodp.
+by rewrite -mul_polyC mulrC rdivp_eq //= /GRing.comm mulrC.
+Qed.
+
+Lemma rdivp_eq d p :
+ (lead_coef d ^+ (rscalp p d)) *: p = (rdivp p d) * d + (rmodp p d).
+Proof.
+rewrite /rdivp /rmodp /rscalp; case: redivpP=> k q1 r1 Hc _; exact: Hc.
+Qed.
+
+Lemma rdvdp_eqP d p : rdvdp_spec p d (rmodp p d) (rdvdp d p).
+Proof.
+case hdvd: (rdvdp d p); last by apply: RdvdpN; move/rmodp_eq0P/eqP: hdvd.
+move/rmodp_eq0P: (hdvd)->; apply: (@Rdvdp _ _ _ (rscalp p d) (rdivp p d)).
+by rewrite mulrC mul_polyC rdivp_eq; move/rmodp_eq0P: (hdvd)->; rewrite addr0.
+Qed.
+
+Lemma rdvdp_eq q p :
+ (rdvdp q p) = ((lead_coef q) ^+ (rscalp p q) *: p == (rdivp p q) * q).
+apply/rmodp_eq0P/eqP; rewrite rdivp_eq; first by move->; rewrite addr0.
+by move/eqP; rewrite eq_sym addrC -subr_eq subrr; move/eqP->.
+Qed.
+
+End CommutativeRingPseudoDivision.
+
+End ComRing.
+
+Module UnitRing.
+
+Import Ring.
+
+Section UnitRingPseudoDivision.
+
+Variable R : unitRingType.
+Implicit Type p q r d : {poly R}.
+
+Lemma uniq_roots_rdvdp p rs :
+ all (root p) rs -> uniq_roots rs ->
+ rdvdp (\prod_(z <- rs) ('X - z%:P)) p.
+Proof.
+move=> rrs; case/(uniq_roots_prod_XsubC rrs)=> q ->; apply: RingMonic.rdvdp_mull.
+exact: monic_prod_XsubC.
+Qed.
+
+End UnitRingPseudoDivision.
+
+End UnitRing.
+
+Module IdomainDefs.
+
+Import Ring.
+
+Section IDomainPseudoDivisionDefs.
+
+Variable R : idomainType.
+Implicit Type p q r d : {poly R}.
+
+Definition edivp_expanded_def p q :=
+ let: (k, d, r) as edvpq := redivp p q in
+ if lead_coef q \in GRing.unit then
+ (0%N, (lead_coef q)^-k *: d, (lead_coef q)^-k *: r)
+ else edvpq.
+Fact edivp_key : unit. Proof. by []. Qed.
+Definition edivp := locked_with edivp_key edivp_expanded_def.
+Canonical edivp_unlockable := [unlockable fun edivp].
+
+Definition divp p q := ((edivp p q).1).2.
+Definition modp p q := (edivp p q).2.
+Definition scalp p q := ((edivp p q).1).1.
+Definition dvdp p q := modp q p == 0.
+Definition eqp p q := (dvdp p q) && (dvdp q p).
+
+
+End IDomainPseudoDivisionDefs.
+
+Notation "m %/ d" := (divp m d) : ring_scope.
+Notation "m %% d" := (modp m d) : ring_scope.
+Notation "p %| q" := (dvdp p q) : ring_scope.
+Notation "p %= q" := (eqp p q) : ring_scope.
+End IdomainDefs.
+
+Module WeakIdomain.
+
+Import Ring ComRing UnitRing IdomainDefs.
+
+Section WeakTheoryForIDomainPseudoDivision.
+
+Variable R : idomainType.
+Implicit Type p q r d : {poly R}.
+
+
+Lemma edivp_def p q : edivp p q = (scalp p q, divp p q, modp p q).
+Proof. by rewrite /scalp /divp /modp; case: (edivp p q) => [[]] /=. Qed.
+
+Lemma edivp_redivp p q : (lead_coef q \in GRing.unit) = false ->
+ edivp p q = redivp p q.
+Proof. by move=> hu; rewrite unlock hu; case: (redivp p q) => [[? ?] ?]. Qed.
+
+Lemma divpE p q :
+ p %/ q = if lead_coef q \in GRing.unit
+ then (lead_coef q)^-(rscalp p q) *: (rdivp p q)
+ else rdivp p q.
+Proof.
+by case ulcq: (lead_coef q \in GRing.unit); rewrite /divp unlock redivp_def ulcq.
+Qed.
+
+Lemma modpE p q :
+ p %% q = if lead_coef q \in GRing.unit
+ then (lead_coef q)^-(rscalp p q) *: (rmodp p q)
+ else rmodp p q.
+Proof.
+by case ulcq: (lead_coef q \in GRing.unit); rewrite /modp unlock redivp_def ulcq.
+Qed.
+
+Lemma scalpE p q :
+ scalp p q = if lead_coef q \in GRing.unit then 0%N else rscalp p q.
+Proof.
+by case h: (lead_coef q \in GRing.unit); rewrite /scalp unlock redivp_def h.
+Qed.
+
+Lemma dvdpE p q : p %| q = rdvdp p q.
+Proof.
+rewrite /dvdp modpE /rdvdp; case ulcq: (lead_coef p \in GRing.unit)=> //.
+rewrite -[_ *: _ == 0]size_poly_eq0 size_scale ?size_poly_eq0 //.
+by rewrite invr_eq0 expf_neq0 //; apply: contraTneq ulcq => ->; rewrite unitr0.
+Qed.
+
+Lemma lc_expn_scalp_neq0 p q : lead_coef q ^+ scalp p q != 0.
+Proof.
+case: (eqVneq q 0) => [->|nzq]; last by rewrite expf_neq0 ?lead_coef_eq0.
+by rewrite /scalp 2!unlock /= eqxx lead_coef0 unitr0 /= oner_neq0.
+Qed.
+
+Hint Resolve lc_expn_scalp_neq0.
+
+CoInductive edivp_spec (m d : {poly R}) :
+ nat * {poly R} * {poly R} -> bool -> Type :=
+|Redivp_spec k (q r: {poly R}) of
+ (lead_coef d ^+ k) *: m = q * d + r & lead_coef d \notin GRing.unit &
+ (d != 0 -> size r < size d) : edivp_spec m d (k, q, r) false
+|Fedivp_spec (q r: {poly R}) of m = q * d + r & (lead_coef d \in GRing.unit) &
+ (d != 0 -> size r < size d) : edivp_spec m d (0%N, q, r) true.
+
+(* There are several ways to state this fact. The most appropriate statement*)
+(* might be polished in light of usage. *)
+Lemma edivpP m d : edivp_spec m d (edivp m d) (lead_coef d \in GRing.unit).
+Proof.
+have hC : GRing.comm d (lead_coef d)%:P by rewrite /GRing.comm mulrC.
+case ud: (lead_coef d \in GRing.unit); last first.
+ rewrite edivp_redivp // redivp_def; constructor; rewrite ?ltn_rmodp // ?ud //.
+ by rewrite rdivp_eq.
+have cdn0: lead_coef d != 0 by apply: contraTneq ud => ->; rewrite unitr0.
+rewrite unlock ud redivp_def; constructor => //.
+ rewrite -scalerAl -scalerDr -mul_polyC.
+ have hn0 : (lead_coef d ^+ rscalp m d)%:P != 0.
+ by rewrite polyC_eq0; apply: expf_neq0.
+ apply: (mulfI hn0); rewrite !mulrA -exprVn !polyC_exp -exprMn -polyC_mul.
+ by rewrite divrr // expr1n mul1r -polyC_exp mul_polyC rdivp_eq.
+move=> dn0; rewrite size_scale ?ltn_rmodp // -exprVn expf_eq0 negb_and.
+by rewrite invr_eq0 cdn0 orbT.
+Qed.
+
+Lemma edivp_eq d q r : size r < size d -> lead_coef d \in GRing.unit ->
+ edivp (q * d + r) d = (0%N, q, r).
+Proof.
+have hC : GRing.comm d (lead_coef d)%:P by exact: mulrC.
+move=> hsrd hu; rewrite unlock hu; case et: (redivp _ _) => [[s qq] rr].
+have cdn0 : lead_coef d != 0.
+ by move: hu; case d0: (lead_coef d == 0) => //; rewrite (eqP d0) unitr0.
+move: (et); rewrite RingComRreg.redivp_eq //; last by apply/rregP.
+rewrite et /=; case => e1 e2; rewrite -!mul_polyC -!exprVn !polyC_exp.
+suff h x y: x * (lead_coef d ^+ s)%:P = y -> ((lead_coef d)^-1)%:P ^+ s * y = x.
+ by congr (_, _, _); apply: h.
+have hn0 : (lead_coef d)%:P ^+ s != 0 by apply: expf_neq0; rewrite polyC_eq0.
+move=> hh; apply: (mulfI hn0); rewrite mulrA -exprMn -polyC_mul divrr //.
+by rewrite expr1n mul1r -polyC_exp mulrC; apply: sym_eq.
+Qed.
+
+Lemma divp_eq p q :
+ (lead_coef q ^+ (scalp p q)) *: p = (p %/ q) * q + (p %% q).
+Proof.
+rewrite divpE modpE scalpE.
+case uq: (lead_coef q \in GRing.unit); last by rewrite rdivp_eq.
+rewrite expr0 scale1r; case: (altP (q =P 0)) => [-> | qn0].
+ rewrite mulr0 add0r lead_coef0 rmodp0 /rscalp unlock eqxx expr0 invr1.
+ by rewrite scale1r.
+have hn0 : (lead_coef q ^+ rscalp p q)%:P != 0.
+ by rewrite polyC_eq0 expf_neq0 // lead_coef_eq0.
+apply: (mulfI hn0).
+rewrite -scalerAl -scalerDr !mul_polyC scalerA mulrV ?unitrX //.
+by rewrite scale1r rdivp_eq.
+Qed.
+
+
+Lemma dvdp_eq q p :
+ (q %| p) = ((lead_coef q) ^+ (scalp p q) *: p == (p %/ q) * q).
+Proof.
+rewrite dvdpE rdvdp_eq scalpE divpE; case: ifP => ulcq //.
+rewrite expr0 scale1r; apply/eqP/eqP.
+ by rewrite -scalerAl; move<-; rewrite scalerA mulVr ?scale1r // unitrX.
+by move=> {2}->; rewrite scalerAl scalerA mulrV ?scale1r // unitrX.
+Qed.
+
+Lemma divpK d p : d %| p -> p %/ d * d = ((lead_coef d) ^+ (scalp p d)) *: p.
+Proof. by rewrite dvdp_eq; move/eqP->. Qed.
+
+Lemma divpKC d p : d %| p -> d * (p %/ d) = ((lead_coef d) ^+ (scalp p d)) *: p.
+Proof. by move=> ?; rewrite mulrC divpK. Qed.
+
+Lemma dvdpP q p :
+ reflect (exists2 cqq, cqq.1 != 0 & cqq.1 *: p = cqq.2 * q) (q %| p).
+Proof.
+rewrite dvdp_eq; apply: (iffP eqP) => [e | [[c qq] cn0 e]].
+ by exists (lead_coef q ^+ scalp p q, p %/ q) => //=.
+apply/eqP; rewrite -dvdp_eq dvdpE.
+have Ecc: c%:P != 0 by rewrite polyC_eq0.
+case: (eqVneq p 0) => [->|nz_p]; first by rewrite rdvdp0.
+pose p1 : {poly R} := lead_coef q ^+ rscalp p q *: qq - c *: (rdivp p q).
+have E1: c *: (rmodp p q) = p1 * q.
+ rewrite mulrDl {1}mulNr -scalerAl -e scalerA mulrC -scalerA -scalerAl.
+ by rewrite -scalerBr rdivp_eq addrC addKr.
+rewrite /dvdp; apply/idPn=> m_nz.
+have: p1 * q != 0 by rewrite -E1 -mul_polyC mulf_neq0 // -/(dvdp q p) dvdpE.
+rewrite mulf_eq0; case/norP=> p1_nz q_nz; have:= ltn_rmodp p q.
+rewrite q_nz -(size_scale _ cn0) E1 size_mul //.
+by rewrite polySpred // ltnNge leq_addl.
+Qed.
+
+Lemma mulpK p q : q != 0 ->
+ p * q %/ q = lead_coef q ^+ scalp (p * q) q *: p.
+Proof.
+move=> qn0; move/rregP: (qn0); apply; rewrite -scalerAl divp_eq.
+suff -> : (p * q) %% q = 0 by rewrite addr0.
+rewrite modpE RingComRreg.rmodp_mull ?scaler0 ?if_same //.
+ by red; rewrite mulrC.
+by apply/rregP; rewrite lead_coef_eq0.
+Qed.
+
+Lemma mulKp p q : q != 0 ->
+ q * p %/ q = lead_coef q ^+ scalp (p * q) q *: p.
+Proof. move=> ?; rewrite mulrC; exact: mulpK. Qed.
+
+Lemma divpp p : p != 0 -> p %/ p = (lead_coef p ^+ scalp p p)%:P.
+Proof.
+move=> np0; have := (divp_eq p p).
+suff -> : p %% p = 0.
+ by rewrite addr0; move/eqP; rewrite -mul_polyC (inj_eq (mulIf np0)); move/eqP.
+rewrite modpE Ring.rmodpp; last by red; rewrite mulrC.
+by rewrite scaler0 if_same.
+Qed.
+
+End WeakTheoryForIDomainPseudoDivision.
+
+Hint Resolve lc_expn_scalp_neq0.
+
+End WeakIdomain.
+
+Module CommonIdomain.
+
+Import Ring ComRing UnitRing IdomainDefs WeakIdomain.
+
+Section IDomainPseudoDivision.
+
+Variable R : idomainType.
+Implicit Type p q r d m n : {poly R}.
+
+Lemma scalp0 p : scalp p 0 = 0%N.
+Proof. by rewrite /scalp unlock lead_coef0 unitr0 unlock eqxx. Qed.
+
+Lemma divp_small p q : size p < size q -> p %/ q = 0.
+Proof.
+move=> spq; rewrite /divp unlock redivp_def /=.
+by case: ifP; rewrite rdivp_small // scaler0.
+Qed.
+
+Lemma leq_divp p q : (size (p %/ q) <= size p).
+Proof.
+rewrite /divp unlock redivp_def /=; case: ifP=> /=; rewrite ?leq_rdivp //.
+move=> ulcq; rewrite size_scale ?leq_rdivp //.
+rewrite -exprVn expf_neq0 // invr_eq0.
+by move: ulcq; case lcq0: (lead_coef q == 0) => //; rewrite (eqP lcq0) unitr0.
+Qed.
+
+Lemma div0p p : 0 %/ p = 0.
+Proof.
+by rewrite /divp unlock redivp_def /=; case: ifP; rewrite rdiv0p // scaler0.
+Qed.
+
+Lemma divp0 p : p %/ 0 = 0.
+Proof.
+by rewrite /divp unlock redivp_def /=; case: ifP; rewrite rdivp0 // scaler0.
+Qed.
+
+Lemma divp1 m : m %/ 1 = m.
+Proof.
+by rewrite divpE lead_coefC unitr1 Ring.rdivp1 expr1n invr1 scale1r.
+Qed.
+
+Lemma modp0 p : p %% 0 = p.
+Proof.
+rewrite /modp unlock redivp_def; case: ifP; rewrite rmodp0 //= lead_coef0.
+by rewrite unitr0.
+Qed.
+
+Lemma mod0p p : 0 %% p = 0.
+Proof.
+by rewrite /modp unlock redivp_def /=; case: ifP; rewrite rmod0p // scaler0.
+Qed.
+
+Lemma modp1 p : p %% 1 = 0.
+Proof.
+by rewrite /modp unlock redivp_def /=; case: ifP; rewrite rmodp1 // scaler0.
+Qed.
+
+Hint Resolve divp0 divp1 mod0p modp0 modp1.
+
+Lemma modp_small p q : size p < size q -> p %% q = p.
+Proof.
+move=> spq; rewrite /modp unlock redivp_def; case: ifP; rewrite rmodp_small //.
+by rewrite /= rscalp_small // expr0 /= invr1 scale1r.
+Qed.
+
+Lemma modpC p c : c != 0 -> p %% c%:P = 0.
+Proof.
+move=> cn0; rewrite /modp unlock redivp_def /=; case: ifP; rewrite ?rmodpC //.
+by rewrite scaler0.
+Qed.
+
+Lemma modp_mull p q : (p * q) %% q = 0.
+Proof.
+case: (altP (q =P 0)) => [-> | nq0]; first by rewrite modp0 mulr0.
+have rlcq : (GRing.rreg (lead_coef q)) by apply/rregP; rewrite lead_coef_eq0.
+have hC : GRing.comm q (lead_coef q)%:P by red; rewrite mulrC.
+by rewrite modpE; case: ifP => ulcq; rewrite RingComRreg.rmodp_mull // scaler0.
+Qed.
+
+Lemma modp_mulr d p : (d * p) %% d = 0.
+Proof. by rewrite mulrC modp_mull. Qed.
+
+Lemma modpp d : d %% d = 0.
+Proof. by rewrite -{1}(mul1r d) modp_mull. Qed.
+
+Lemma ltn_modp p q : (size (p %% q) < size q) = (q != 0).
+Proof.
+rewrite /modp unlock redivp_def /=; case: ifP=> /=; rewrite ?ltn_rmodp //.
+move=> ulcq; rewrite size_scale ?ltn_rmodp //.
+rewrite -exprVn expf_neq0 // invr_eq0.
+by move: ulcq; case lcq0: (lead_coef q == 0) => //; rewrite (eqP lcq0) unitr0.
+Qed.
+
+Lemma ltn_divpl d q p : d != 0 ->
+ (size (q %/ d) < size p) = (size q < size (p * d)).
+Proof.
+move=> dn0; have sd : size d > 0 by rewrite size_poly_gt0 dn0.
+have: (lead_coef d) ^+ (scalp q d) != 0 by exact: lc_expn_scalp_neq0.
+move/size_scale; move/(_ q)<-; rewrite divp_eq; case quo0 : (q %/ d == 0).
+ rewrite (eqP quo0) mul0r add0r size_poly0.
+ case p0 : (p == 0); first by rewrite (eqP p0) mul0r size_poly0 ltnn ltn0.
+ have sp : size p > 0 by rewrite size_poly_gt0 p0.
+ rewrite /= size_mul ?p0 // sp; apply: sym_eq; move/prednK:(sp)<-.
+ by rewrite addSn /= ltn_addl // ltn_modp.
+rewrite size_addl; last first.
+ rewrite size_mul ?quo0 //; move/negbT: quo0; rewrite -size_poly_gt0.
+ by move/prednK<-; rewrite addSn /= ltn_addl // ltn_modp.
+case: (altP (p =P 0)) => [-> | pn0]; first by rewrite mul0r size_poly0 !ltn0.
+by rewrite !size_mul ?quo0 //; move/prednK: sd<-; rewrite !addnS ltn_add2r.
+Qed.
+
+Lemma leq_divpr d p q : d != 0 ->
+ (size p <= size (q %/ d)) = (size (p * d) <= size q).
+Proof. by move=> dn0; rewrite leqNgt ltn_divpl // -leqNgt. Qed.
+
+Lemma divpN0 d p : d != 0 -> (p %/ d != 0) = (size d <= size p).
+Proof.
+move=> dn0; rewrite -{2}(mul1r d) -leq_divpr // size_polyC oner_eq0 /=.
+by rewrite size_poly_gt0.
+Qed.
+
+Lemma size_divp p q : q != 0 -> size (p %/ q) = ((size p) - (size q).-1)%N.
+Proof.
+move=> nq0; case: (leqP (size q) (size p)) => sqp; last first.
+ move: (sqp); rewrite -{1}(ltn_predK sqp) ltnS -subn_eq0 divp_small //.
+ by move/eqP->; rewrite size_poly0.
+move: (nq0); rewrite -size_poly_gt0 => lt0sq.
+move: (sqp); move/(leq_trans lt0sq) => lt0sp.
+move: (lt0sp); rewrite size_poly_gt0=> p0.
+move: (divp_eq p q); move/(congr1 (size \o (@polyseq R)))=> /=.
+rewrite size_scale; last by rewrite expf_eq0 lead_coef_eq0 (negPf nq0) andbF.
+case: (eqVneq (p %/ q) 0) => [-> | qq0].
+ by rewrite mul0r add0r=> es; move: nq0; rewrite -(ltn_modp p) -es ltnNge sqp.
+move/negP:(qq0); move/negP; rewrite -size_poly_gt0 => lt0qq.
+rewrite size_addl.
+ rewrite size_mul ?qq0 // => ->.
+ apply/eqP; rewrite -(eqn_add2r ((size q).-1)).
+ rewrite subnK; first by rewrite -subn1 addnBA // subn1.
+ rewrite /leq -(subnDl 1%N) !add1n prednK // (@ltn_predK (size q)) //.
+ by rewrite addnC subnDA subnn sub0n.
+ by rewrite -[size q]add0n ltn_add2r.
+rewrite size_mul ?qq0 //.
+move: nq0; rewrite -(ltn_modp p); move/leq_trans; apply; move/prednK: lt0qq<-.
+by rewrite addSn /= leq_addl.
+Qed.
+
+Lemma ltn_modpN0 p q : q != 0 -> size (p %% q) < size q.
+Proof. by rewrite ltn_modp. Qed.
+
+Lemma modp_mod p q : (p %% q) %% q = p %% q.
+Proof.
+by case: (eqVneq q 0) => [-> | qn0]; rewrite ?modp0 // modp_small ?ltn_modp.
+Qed.
+
+Lemma leq_modp m d : size (m %% d) <= size m.
+Proof.
+rewrite /modp unlock redivp_def /=; case: ifP; rewrite ?leq_rmodp //.
+move=> ud; rewrite size_scale ?leq_rmodp // invr_eq0 expf_neq0 //.
+by apply: contraTneq ud => ->; rewrite unitr0.
+Qed.
+
+Lemma dvdp0 d : d %| 0.
+Proof. by rewrite /dvdp mod0p. Qed.
+
+Hint Resolve dvdp0.
+
+Lemma dvd0p p : (0 %| p) = (p == 0).
+Proof. by rewrite /dvdp modp0. Qed.
+
+Lemma dvd0pP p : reflect (p = 0) (0 %| p).
+Proof. by apply: (iffP idP); rewrite dvd0p; move/eqP. Qed.
+
+Lemma dvdpN0 p q : p %| q -> q != 0 -> p != 0.
+Proof. by move=> pq hq; apply: contraL pq=> /eqP ->; rewrite dvd0p. Qed.
+
+Lemma dvdp1 d : (d %| 1) = ((size d) == 1%N).
+Proof.
+rewrite /dvdp modpE; case ud: (lead_coef d \in GRing.unit); last exact: rdvdp1.
+rewrite -size_poly_eq0 size_scale; first by rewrite size_poly_eq0; exact: rdvdp1.
+by rewrite invr_eq0 expf_neq0 //; apply: contraTneq ud => ->; rewrite unitr0.
+Qed.
+
+Lemma dvd1p m : 1 %| m.
+Proof. by rewrite /dvdp modp1. Qed.
+
+Lemma gtNdvdp p q : p != 0 -> size p < size q -> (q %| p) = false.
+Proof.
+by move=> nn0 hs; rewrite /dvdp; rewrite (modp_small hs); apply: negPf.
+Qed.
+
+Lemma modp_eq0P p q : reflect (p %% q = 0) (q %| p).
+Proof. exact: (iffP eqP). Qed.
+
+Lemma modp_eq0 p q : (q %| p) -> p %% q = 0.
+Proof. by move/modp_eq0P. Qed.
+
+Lemma leq_divpl d p q :
+ d %| p -> (size (p %/ d) <= size q) = (size p <= size (q * d)).
+Proof.
+case: (eqVneq d 0) => [-> | nd0].
+ by move/dvd0pP->; rewrite divp0 size_poly0 !leq0n.
+move=> hd; rewrite leq_eqVlt ltn_divpl // (leq_eqVlt (size p)).
+case lhs: (size p < size (q * d)); rewrite ?orbT ?orbF //.
+have: (lead_coef d) ^+ (scalp p d) != 0 by rewrite expf_neq0 // lead_coef_eq0.
+move/size_scale; move/(_ p)<-; rewrite divp_eq.
+move/modp_eq0P: hd->; rewrite addr0; case: (altP (p %/ d =P 0))=> [-> | quon0].
+ rewrite mul0r size_poly0 eq_sym (eq_sym 0%N) size_poly_eq0.
+ case: (altP (q =P 0)) => [-> | nq0]; first by rewrite mul0r size_poly0 eqxx.
+ by rewrite size_poly_eq0 mulf_eq0 (negPf nq0) (negPf nd0).
+case: (altP (q =P 0)) => [-> | nq0].
+ by rewrite mul0r size_poly0 !size_poly_eq0 mulf_eq0 (negPf nd0) orbF.
+rewrite !size_mul //; move: nd0; rewrite -size_poly_gt0; move/prednK<-.
+by rewrite !addnS /= eqn_add2r.
+Qed.
+
+Lemma dvdp_leq p q : q != 0 -> p %| q -> size p <= size q.
+move=> nq0 /modp_eq0P => rpq; case: (ltnP (size p) (size q)).
+ by move/ltnW->.
+rewrite leq_eqVlt; case/orP; first by move/eqP->.
+by move/modp_small; rewrite rpq => h; move: nq0; rewrite h eqxx.
+Qed.
+
+Lemma eq_dvdp c quo q p : c != 0 -> c *: p = quo * q -> q %| p.
+Proof.
+move=> cn0; case: (eqVneq p 0) => [->|nz_quo def_quo] //.
+pose p1 : {poly R} := lead_coef q ^+ scalp p q *: quo - c *: (p %/ q).
+have E1: c *: (p %% q) = p1 * q.
+ rewrite mulrDl {1}mulNr-scalerAl -def_quo scalerA mulrC -scalerA.
+ by rewrite -scalerAl -scalerBr divp_eq addrAC subrr add0r.
+rewrite /dvdp; apply/idPn=> m_nz.
+have: p1 * q != 0 by rewrite -E1 -mul_polyC mulf_neq0 // polyC_eq0.
+rewrite mulf_eq0; case/norP=> p1_nz q_nz.
+have := (ltn_modp p q); rewrite q_nz -(size_scale (p %% q) cn0) E1.
+by rewrite size_mul // polySpred // ltnNge leq_addl.
+Qed.
+
+Lemma dvdpp d : d %| d.
+Proof. by rewrite /dvdp modpp. Qed.
+
+Hint Resolve dvdpp.
+
+Lemma divp_dvd p q : (p %| q) -> ((q %/ p) %| q).
+Proof.
+case: (eqVneq p 0) => [-> | np0]; first by rewrite divp0.
+rewrite dvdp_eq => /eqP h.
+apply: (@eq_dvdp ((lead_coef p)^+ (scalp q p)) p); last by rewrite mulrC.
+by rewrite expf_neq0 // lead_coef_eq0.
+Qed.
+
+Lemma dvdp_mull m d n : d %| n -> d %| m * n.
+Proof.
+case: (eqVneq d 0) => [-> |dn0]; first by move/dvd0pP->; rewrite mulr0 dvdpp.
+rewrite dvdp_eq => /eqP e.
+apply: (@eq_dvdp (lead_coef d ^+ scalp n d) (m * (n %/ d))).
+ by rewrite expf_neq0 // lead_coef_eq0.
+by rewrite scalerAr e mulrA.
+Qed.
+
+Lemma dvdp_mulr n d m : d %| m -> d %| m * n.
+Proof. by move=> hdm; rewrite mulrC dvdp_mull. Qed.
+
+Hint Resolve dvdp_mull dvdp_mulr.
+
+Lemma dvdp_mul d1 d2 m1 m2 : d1 %| m1 -> d2 %| m2 -> d1 * d2 %| m1 * m2.
+Proof.
+case: (eqVneq d1 0) => [-> |d1n0]; first by move/dvd0pP->; rewrite !mul0r dvdpp.
+case: (eqVneq d2 0) => [-> |d2n0]; first by move => _ /dvd0pP ->; rewrite !mulr0.
+rewrite dvdp_eq; set c1 := _ ^+ _; set q1 := _ %/ _; move/eqP=> Hq1.
+rewrite dvdp_eq; set c2 := _ ^+ _; set q2 := _ %/ _; move/eqP=> Hq2.
+apply: (@eq_dvdp (c1 * c2) (q1 * q2)).
+ by rewrite mulf_neq0 // expf_neq0 // lead_coef_eq0.
+rewrite -scalerA scalerAr scalerAl Hq1 Hq2 -!mulrA.
+by rewrite [d1 * (q2 * _)]mulrCA.
+Qed.
+
+Lemma dvdp_addr m d n : d %| m -> (d %| m + n) = (d %| n).
+Proof.
+case: (altP (d =P 0)) => [-> | dn0]; first by move/dvd0pP->; rewrite add0r.
+rewrite dvdp_eq; set c1 := _ ^+ _; set q1 := _ %/ _; move/eqP=> Eq1.
+apply/idP/idP; rewrite dvdp_eq; set c2 := _ ^+ _; set q2 := _ %/ _.
+ have sn0 : c1 * c2 != 0.
+ by rewrite !mulf_neq0 // expf_eq0 lead_coef_eq0 (negPf dn0) andbF.
+ move/eqP=> Eq2; apply: (@eq_dvdp _ (c1 *: q2 - c2 *: q1) _ _ sn0).
+ rewrite mulrDl -scaleNr -!scalerAl -Eq1 -Eq2 !scalerA.
+ by rewrite mulNr mulrC scaleNr -scalerBr addrC addKr.
+have sn0 : c1 * c2 != 0.
+ by rewrite !mulf_neq0 // expf_eq0 lead_coef_eq0 (negPf dn0) andbF.
+move/eqP=> Eq2; apply: (@eq_dvdp _ (c1 *: q2 + c2 *: q1) _ _ sn0).
+by rewrite mulrDl -!scalerAl -Eq1 -Eq2 !scalerA mulrC addrC scalerDr.
+Qed.
+
+Lemma dvdp_addl n d m : d %| n -> (d %| m + n) = (d %| m).
+Proof. by rewrite addrC; exact: dvdp_addr. Qed.
+
+Lemma dvdp_add d m n : d %| m -> d %| n -> d %| m + n.
+Proof. by move/dvdp_addr->. Qed.
+
+Lemma dvdp_add_eq d m n : d %| m + n -> (d %| m) = (d %| n).
+Proof. by move=> ?; apply/idP/idP; [move/dvdp_addr <-| move/dvdp_addl <-]. Qed.
+
+Lemma dvdp_subr d m n : d %| m -> (d %| m - n) = (d %| n).
+Proof. by move=> ?; apply dvdp_add_eq; rewrite -addrA addNr simp. Qed.
+
+Lemma dvdp_subl d m n : d %| n -> (d %| m - n) = (d %| m).
+Proof. by move/dvdp_addl<-; rewrite subrK. Qed.
+
+Lemma dvdp_sub d m n : d %| m -> d %| n -> d %| m - n.
+Proof. by move=> *; rewrite dvdp_subl. Qed.
+
+Lemma dvdp_mod d n m : d %| n -> (d %| m) = (d %| m %% n).
+Proof.
+case: (altP (n =P 0)) => [-> | nn0]; first by rewrite modp0.
+case: (altP (d =P 0)) => [-> | dn0]; first by move/dvd0pP->; rewrite modp0.
+rewrite dvdp_eq; set c1 := _ ^+ _; set q1 := _ %/ _; move/eqP=> Eq1.
+apply/idP/idP; rewrite dvdp_eq; set c2 := _ ^+ _; set q2 := _ %/ _.
+ have sn0 : c1 * c2 != 0.
+ by rewrite !mulf_neq0 // expf_eq0 lead_coef_eq0 (negPf dn0) andbF.
+ pose quo := (c1 * lead_coef n ^+ scalp m n) *: q2 - c2 *: (m %/ n) * q1.
+ move/eqP=> Eq2; apply: (@eq_dvdp _ quo _ _ sn0).
+ rewrite mulrDl mulNr -!scalerAl -!mulrA -Eq1 -Eq2 -scalerAr !scalerA.
+ rewrite mulrC [_ * c2]mulrC mulrA -[((_ * _) * _) *: _]scalerA -scalerBr.
+ by rewrite divp_eq addrC addKr.
+have sn0 : c1 * c2 * lead_coef n ^+ scalp m n != 0.
+ rewrite !mulf_neq0 // expf_eq0 lead_coef_eq0 ?(negPf dn0) ?andbF //.
+ by rewrite (negPf nn0) andbF.
+move/eqP=> Eq2; apply: (@eq_dvdp _ (c2 *: (m %/ n) * q1 + c1 *: q2) _ _ sn0).
+rewrite -scalerA divp_eq scalerDr -!scalerA Eq2 scalerAl scalerAr Eq1.
+by rewrite scalerAl mulrDl mulrA.
+Qed.
+
+Lemma dvdp_trans : transitive (@dvdp R).
+Proof.
+move=> n d m.
+case: (altP (d =P 0)) => [-> | dn0]; first by move/dvd0pP->.
+case: (altP (n =P 0)) => [-> | nn0]; first by move=> _ /dvd0pP ->.
+rewrite dvdp_eq; set c1 := _ ^+ _; set q1 := _ %/ _; move/eqP=> Hq1.
+rewrite dvdp_eq; set c2 := _ ^+ _; set q2 := _ %/ _; move/eqP=> Hq2.
+have sn0 : c1 * c2 != 0 by rewrite mulf_neq0 // expf_neq0 // lead_coef_eq0.
+by apply: (@eq_dvdp _ (q2 * q1) _ _ sn0); rewrite -scalerA Hq2 scalerAr Hq1 mulrA.
+Qed.
+
+Lemma dvdp_mulIl p q : p %| p * q.
+Proof. by apply: dvdp_mulr; exact: dvdpp. Qed.
+
+Lemma dvdp_mulIr p q : q %| p * q.
+Proof. by apply: dvdp_mull; exact: dvdpp. Qed.
+
+Lemma dvdp_mul2r r p q : r != 0 -> (p * r %| q * r) = (p %| q).
+Proof.
+move => nzr.
+case: (eqVneq p 0) => [-> | pn0].
+ by rewrite mul0r !dvd0p mulf_eq0 (negPf nzr) orbF.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite mul0r !dvdp0.
+apply/idP/idP; last by move => ?; rewrite dvdp_mul ?dvdpp.
+rewrite dvdp_eq; set c := _ ^+ _; set x := _ %/ _; move/eqP=> Hx.
+apply: (@eq_dvdp c x).
+ by rewrite expf_neq0 // lead_coef_eq0 mulf_neq0.
+by apply: (GRing.mulIf nzr); rewrite -GRing.mulrA -GRing.scalerAl.
+Qed.
+
+Lemma dvdp_mul2l r p q: r != 0 -> (r * p %| r * q) = (p %| q).
+Proof. by rewrite ![r * _]GRing.mulrC; apply: dvdp_mul2r. Qed.
+
+Lemma ltn_divpr d p q :
+ d %| q -> (size p < size (q %/ d)) = (size (p * d) < size q).
+Proof. by move=> dv_d_q; rewrite !ltnNge leq_divpl. Qed.
+
+Lemma dvdp_exp d k p : 0 < k -> d %| p -> d %| (p ^+ k).
+Proof. by case: k => // k _ d_dv_m; rewrite exprS dvdp_mulr. Qed.
+
+Lemma dvdp_exp2l d k l : k <= l -> d ^+ k %| d ^+ l.
+Proof.
+by move/subnK <-; rewrite exprD dvdp_mull // ?lead_coef_exp ?unitrX.
+Qed.
+
+Lemma dvdp_Pexp2l d k l : 1 < size d -> (d ^+ k %| d ^+ l) = (k <= l).
+Proof.
+move=> sd; case: leqP => [|gt_n_m]; first exact: dvdp_exp2l.
+have dn0 : d != 0 by rewrite -size_poly_gt0; apply: ltn_trans sd.
+rewrite gtNdvdp ?expf_neq0 // polySpred ?expf_neq0 // size_exp /=.
+rewrite [size (d ^+ k)]polySpred ?expf_neq0 // size_exp ltnS ltn_mul2l.
+by move: sd; rewrite -subn_gt0 subn1; move->.
+Qed.
+
+Lemma dvdp_exp2r p q k : p %| q -> p ^+ k %| q ^+ k.
+Proof.
+case: (eqVneq p 0) => [-> | pn0]; first by move/dvd0pP->.
+rewrite dvdp_eq; set c := _ ^+ _; set t := _ %/ _; move/eqP=> e.
+apply: (@eq_dvdp (c ^+ k) (t ^+ k)); first by rewrite !expf_neq0 ?lead_coef_eq0.
+by rewrite -exprMn -exprZn; congr (_ ^+ k).
+Qed.
+
+Lemma dvdp_exp_sub p q k l: p != 0 ->
+ (p ^+ k %| q * p ^+ l) = (p ^+ (k - l) %| q).
+Proof.
+move=> pn0; case: (leqP k l)=> hkl.
+ move:(hkl); rewrite -subn_eq0; move/eqP->; rewrite expr0 dvd1p.
+ apply: dvdp_mull; case: (ltnP 1%N (size p)) => sp.
+ by rewrite dvdp_Pexp2l.
+ move: sp; case esp: (size p) => [|sp].
+ by move/eqP: esp; rewrite size_poly_eq0 (negPf pn0).
+ rewrite ltnS leqn0; move/eqP=> sp0; move/eqP: esp; rewrite sp0.
+ by case/size_poly1P => c cn0 ->; move/subnK: hkl<-; rewrite exprD dvdp_mulIr.
+rewrite -{1}[k](@subnK l) 1?ltnW// exprD dvdp_mul2r//.
+elim: l {hkl}=> [|l ihl]; first by rewrite expr0 oner_eq0.
+by rewrite exprS mulf_neq0.
+Qed.
+
+Lemma dvdp_XsubCl p x : ('X - x%:P) %| p = root p x.
+Proof. rewrite dvdpE; exact: Ring.rdvdp_XsubCl. Qed.
+
+Lemma polyXsubCP p x : reflect (p.[x] = 0) (('X - x%:P) %| p).
+Proof. rewrite dvdpE; exact: Ring.polyXsubCP. Qed.
+
+Lemma eqp_div_XsubC p c :
+ (p == (p %/ ('X - c%:P)) * ('X - c%:P)) = ('X - c%:P %| p).
+Proof. by rewrite dvdp_eq lead_coefXsubC expr1n scale1r. Qed.
+
+Lemma root_factor_theorem p x : root p x = (('X - x%:P) %| p).
+Proof. by rewrite dvdp_XsubCl. Qed.
+
+Lemma uniq_roots_dvdp p rs : all (root p) rs -> uniq_roots rs ->
+ (\prod_(z <- rs) ('X - z%:P)) %| p.
+Proof.
+move=> rrs; case/(uniq_roots_prod_XsubC rrs)=> q ->.
+by apply: dvdp_mull; rewrite // (eqP (monic_prod_XsubC _)) unitr1.
+Qed.
+
+
+Lemma root_bigmul : forall x (ps : seq {poly R}),
+ ~~root (\big[*%R/1]_(p <- ps) p) x = all (fun p => ~~ root p x) ps.
+Proof.
+move=> x; elim; first by rewrite big_nil root1.
+by move=> p ps ihp; rewrite big_cons /= rootM negb_or ihp.
+Qed.
+
+Lemma eqpP m n :
+ reflect (exists2 c12, (c12.1 != 0) && (c12.2 != 0) & c12.1 *: m = c12.2 *: n)
+ (m %= n).
+Proof.
+apply: (iffP idP) => [| [[c1 c2]/andP[nz_c1 nz_c2 eq_cmn]]]; last first.
+ rewrite /eqp (@eq_dvdp c2 c1%:P) -?eq_cmn ?mul_polyC // (@eq_dvdp c1 c2%:P) //.
+ by rewrite eq_cmn mul_polyC.
+case: (eqVneq m 0) => [-> | m_nz].
+ by case/andP => /dvd0pP -> _; exists (1, 1); rewrite ?scaler0 // oner_eq0.
+case: (eqVneq n 0) => [-> | n_nz].
+ by case/andP => _ /dvd0pP ->; exists (1, 1); rewrite ?scaler0 // oner_eq0.
+case/andP; rewrite !dvdp_eq; set c1 := _ ^+ _; set c2 := _ ^+ _.
+set q1 := _ %/ _; set q2 := _ %/ _; move/eqP => Hq1 /eqP Hq2;
+have Hc1 : c1 != 0 by rewrite expf_eq0 lead_coef_eq0 negb_and m_nz orbT.
+have Hc2 : c2 != 0 by rewrite expf_eq0 lead_coef_eq0 negb_and n_nz orbT.
+have def_q12: q1 * q2 = (c1 * c2)%:P.
+ apply: (mulIf m_nz); rewrite mulrAC mulrC -Hq1 -scalerAr -Hq2 scalerA.
+ by rewrite -mul_polyC.
+have: q1 * q2 != 0 by rewrite def_q12 -size_poly_eq0 size_polyC mulf_neq0.
+rewrite mulf_eq0; case/norP=> nz_q1 nz_q2.
+have: size q2 <= 1%N.
+ have:= size_mul nz_q1 nz_q2; rewrite def_q12 size_polyC mulf_neq0 //=.
+ by rewrite polySpred // => ->; rewrite leq_addl.
+rewrite leq_eqVlt ltnS leqn0 size_poly_eq0 (negPf nz_q2) orbF.
+case/size_poly1P=> c cn0 cqe; exists (c2, c); first by rewrite Hc2.
+by rewrite Hq2 -mul_polyC -cqe.
+Qed.
+
+Lemma eqp_eq p q: p %= q -> (lead_coef q) *: p = (lead_coef p) *: q.
+Proof.
+move=> /eqpP [[c1 c2] /= /andP [nz_c1 nz_c2]] eq.
+have/(congr1 lead_coef) := eq; rewrite !lead_coefZ.
+move=> eqC; apply/(@mulfI _ c2%:P); rewrite ?polyC_eq0 //.
+rewrite !mul_polyC scalerA -eqC mulrC -scalerA eq.
+by rewrite !scalerA mulrC.
+Qed.
+
+Lemma eqpxx : reflexive (@eqp R).
+Proof. by move=> p; rewrite /eqp dvdpp. Qed.
+
+Hint Resolve eqpxx.
+
+Lemma eqp_sym : symmetric (@eqp R).
+Proof. by move=> p q; rewrite /eqp andbC. Qed.
+
+Lemma eqp_trans : transitive (@eqp R).
+Proof.
+move=> p q r; case/andP=> Dp pD; case/andP=> Dq qD.
+by rewrite /eqp (dvdp_trans Dp) // (dvdp_trans qD).
+Qed.
+
+Lemma eqp_ltrans : left_transitive (@eqp R).
+Proof.
+move=> p q r pq.
+by apply/idP/idP=> e; apply: eqp_trans e; rewrite // eqp_sym.
+Qed.
+
+Lemma eqp_rtrans : right_transitive (@eqp R).
+Proof. by move=> x y xy z; rewrite eqp_sym (eqp_ltrans xy) eqp_sym. Qed.
+
+Lemma eqp0 : forall p, (p %= 0) = (p == 0).
+Proof.
+move=> p; case: eqP; move/eqP=> Ep; first by rewrite (eqP Ep) eqpxx.
+by apply/negP; case/andP=> _; rewrite /dvdp modp0 (negPf Ep).
+Qed.
+
+Lemma eqp01 : 0 %= (1 : {poly R}) = false.
+Proof.
+case abs : (0 %= 1) => //; case/eqpP: abs=> [[c1 c2]] /andP [c1n0 c2n0] /=.
+by rewrite scaler0 alg_polyC; move/eqP; rewrite eq_sym polyC_eq0 (negbTE c2n0).
+Qed.
+
+Lemma eqp_scale p c : c != 0 -> c *: p %= p.
+Proof.
+move=> c0; apply/eqpP; exists (1, c); first by rewrite c0 oner_eq0.
+by rewrite scale1r.
+Qed.
+
+Lemma eqp_size p q : p %= q -> size p = size q.
+Proof.
+case: (q =P 0); move/eqP => Eq; first by rewrite (eqP Eq) eqp0; move/eqP->.
+rewrite eqp_sym; case: (p =P 0); move/eqP => Ep.
+ by rewrite (eqP Ep) eqp0; move/eqP->.
+by case/andP => Dp Dq; apply: anti_leq; rewrite !dvdp_leq.
+Qed.
+
+Lemma size_poly_eq1 p : (size p == 1%N) = (p %= 1).
+Proof.
+apply/size_poly1P/idP=> [[c cn0 ep] |].
+ by apply/eqpP; exists (1, c); rewrite ?oner_eq0 // alg_polyC scale1r.
+by move/eqp_size; rewrite size_poly1; move/eqP; move/size_poly1P.
+Qed.
+
+Lemma polyXsubC_eqp1 (x : R) : ('X - x%:P %= 1) = false.
+Proof. by rewrite -size_poly_eq1 size_XsubC. Qed.
+
+Lemma dvdp_eqp1 p q : p %| q -> q %= 1 -> p %= 1.
+Proof.
+move=> dpq hq.
+have sizeq : size q == 1%N by rewrite size_poly_eq1.
+have n0q : q != 0.
+ by case abs: (q == 0) => //; move: hq; rewrite (eqP abs) eqp01.
+rewrite -size_poly_eq1 eqn_leq -{1}(eqP sizeq) dvdp_leq //=.
+case p0 : (size p == 0%N); last by rewrite neq0_lt0n.
+by move: dpq; rewrite size_poly_eq0 in p0; rewrite (eqP p0) dvd0p (negbTE n0q).
+Qed.
+
+Lemma eqp_dvdr q p d: p %= q -> d %| p = (d %| q).
+Proof.
+suff Hmn m n: m %= n -> (d %| m) -> (d %| n).
+ by move=> mn; apply/idP/idP; apply: Hmn=> //; rewrite eqp_sym.
+by rewrite /eqp; case/andP=> pq qp dp; apply: (dvdp_trans dp).
+Qed.
+
+Lemma eqp_dvdl d2 d1 p : d1 %= d2 -> d1 %| p = (d2 %| p).
+suff Hmn m n: m %= n -> (m %| p) -> (n %| p).
+ by move=> ?; apply/idP/idP; apply: Hmn; rewrite // eqp_sym.
+by rewrite /eqp; case/andP=> dd' d'd dp; apply: (dvdp_trans d'd).
+Qed.
+
+Lemma dvdp_scaler c m n : c != 0 -> m %| c *: n = (m %| n).
+Proof. move=> cn0; apply: eqp_dvdr; exact: eqp_scale. Qed.
+
+Lemma dvdp_scalel c m n : c != 0 -> (c *: m %| n) = (m %| n).
+Proof. move=> cn0; apply: eqp_dvdl; exact: eqp_scale. Qed.
+
+Lemma dvdp_opp d p : d %| (- p) = (d %| p).
+Proof.
+by apply: eqp_dvdr; rewrite -scaleN1r eqp_scale // oppr_eq0 oner_eq0.
+Qed.
+
+Lemma eqp_mul2r r p q : r != 0 -> (p * r %= q * r) = (p %= q).
+Proof. by move => nz_r; rewrite /eqp !dvdp_mul2r. Qed.
+
+Lemma eqp_mul2l r p q: r != 0 -> (r * p %= r * q) = (p %= q).
+Proof. by move => nz_r; rewrite /eqp !dvdp_mul2l. Qed.
+
+Lemma eqp_mull r p q: (q %= r) -> (p * q %= p * r).
+Proof.
+case/eqpP=> [[c d]] /andP [c0 d0 e]; apply/eqpP; exists (c, d); rewrite ?c0 //.
+by rewrite scalerAr e -scalerAr.
+Qed.
+
+Lemma eqp_mulr q p r : (p %= q) -> (p * r %= q * r).
+Proof. by move=> epq; rewrite ![_ * r]mulrC eqp_mull. Qed.
+
+Lemma eqp_exp p q k : p %= q -> p ^+ k %= q ^+ k.
+Proof.
+move=> pq; elim: k=> [|k ihk]; first by rewrite !expr0 eqpxx.
+by rewrite !exprS (@eqp_trans (q * p ^+ k)) // (eqp_mulr, eqp_mull).
+Qed.
+
+Lemma polyC_eqp1 (c : R) : (c%:P %= 1) = (c != 0).
+Proof.
+apply/eqpP/idP => [[[x y]] |nc0] /=.
+ case c0: (c == 0); rewrite // alg_polyC (eqP c0) scaler0.
+ by case/andP=> _ /=; move/negbTE<-; move/eqP; rewrite eq_sym polyC_eq0.
+exists (1, c); first by rewrite nc0 /= oner_neq0.
+by rewrite alg_polyC scale1r.
+Qed.
+
+Lemma dvdUp d p: d %= 1 -> d %| p.
+Proof. by move/eqp_dvdl->; rewrite dvd1p. Qed.
+
+Lemma dvdp_size_eqp p q : p %| q -> size p == size q = (p %= q).
+Proof.
+move=> pq; apply/idP/idP; last by move/eqp_size->.
+case (q =P 0)=> [->|]; [|move/eqP => Hq].
+ by rewrite size_poly0 size_poly_eq0; move/eqP->; rewrite eqpxx.
+case (p =P 0)=> [->|]; [|move/eqP => Hp].
+ by rewrite size_poly0 eq_sym size_poly_eq0; move/eqP->; rewrite eqpxx.
+move: pq; rewrite dvdp_eq; set c := _ ^+ _; set x := _ %/ _; move/eqP=> eqpq.
+move:(eqpq); move/(congr1 (size \o (@polyseq R)))=> /=.
+have cn0 : c != 0 by rewrite expf_neq0 // lead_coef_eq0.
+rewrite (@eqp_size _ q); last by exact: eqp_scale.
+rewrite size_mul ?p0 // => [-> HH|]; last first.
+ apply/eqP=> HH; move: eqpq; rewrite HH mul0r.
+ by move/eqP; rewrite scale_poly_eq0 (negPf Hq) (negPf cn0).
+suff: size x == 1%N.
+ case/size_poly1P=> y H1y H2y.
+ by apply/eqpP; exists (y, c); rewrite ?H1y // eqpq H2y mul_polyC.
+case: (size p) HH (size_poly_eq0 p)=> [|n]; first by case: eqP Hp.
+by rewrite addnS -add1n eqn_add2r;move/eqP->.
+Qed.
+
+Lemma eqp_root p q : p %= q -> root p =1 root q.
+Proof.
+move/eqpP=> [[c d]] /andP [c0 d0 e] x; move/negPf:c0=>c0; move/negPf:d0=>d0.
+rewrite rootE -[_==_]orFb -c0 -mulf_eq0 -hornerZ e hornerZ.
+by rewrite mulf_eq0 d0.
+Qed.
+
+Lemma eqp_rmod_mod p q : rmodp p q %= modp p q.
+Proof.
+rewrite modpE eqp_sym; case: ifP => ulcq //.
+apply: eqp_scale; rewrite invr_eq0 //.
+by apply: expf_neq0; apply: contraTneq ulcq => ->; rewrite unitr0.
+Qed.
+
+Lemma eqp_rdiv_div p q : rdivp p q %= divp p q.
+Proof.
+rewrite divpE eqp_sym; case: ifP=> ulcq //; apply: eqp_scale; rewrite invr_eq0 //.
+by apply: expf_neq0; apply: contraTneq ulcq => ->; rewrite unitr0.
+Qed.
+
+Lemma dvd_eqp_divl d p q (dvd_dp : d %| q) (eq_pq : p %= q) :
+ p %/ d %= q %/ d.
+Proof.
+case: (eqVneq q 0) eq_pq=> [->|q_neq0]; first by rewrite eqp0=> /eqP->.
+have d_neq0: d != 0 by apply: contraL dvd_dp=> /eqP->; rewrite dvd0p.
+move=> eq_pq; rewrite -(@eqp_mul2r d) // !divpK // ?(eqp_dvdr _ eq_pq) //.
+rewrite (eqp_ltrans (eqp_scale _ _)) ?lc_expn_scalp_neq0 //.
+by rewrite (eqp_rtrans (eqp_scale _ _)) ?lc_expn_scalp_neq0.
+Qed.
+
+Definition gcdp_rec p q :=
+ let: (p1, q1) := if size p < size q then (q, p) else (p, q) in
+ if p1 == 0 then q1 else
+ let fix loop (n : nat) (pp qq : {poly R}) {struct n} :=
+ let rr := modp pp qq in
+ if rr == 0 then qq else
+ if n is n1.+1 then loop n1 qq rr else rr in
+ loop (size p1) p1 q1.
+
+Definition gcdp := nosimpl gcdp_rec.
+
+Lemma gcd0p : left_id 0 gcdp.
+Proof.
+move=> p; rewrite /gcdp /gcdp_rec size_poly0 size_poly_gt0 if_neg.
+case: ifP => /= [_ | nzp]; first by rewrite eqxx.
+by rewrite polySpred !(modp0, nzp) //; case: _.-1 => [|m]; rewrite mod0p eqxx.
+Qed.
+
+Lemma gcdp0 : right_id 0 gcdp.
+Proof.
+move=> p; have:= gcd0p p; rewrite /gcdp /gcdp_rec size_poly0 size_poly_gt0.
+by rewrite if_neg; case: ifP => /= p0; rewrite ?(eqxx, p0) // (eqP p0).
+Qed.
+
+Lemma gcdpE p q :
+ gcdp p q = if size p < size q
+ then gcdp (modp q p) p else gcdp (modp p q) q.
+Proof.
+pose gcdpE_rec := fix gcdpE_rec (n : nat) (pp qq : {poly R}) {struct n} :=
+ let rr := modp pp qq in
+ if rr == 0 then qq else
+ if n is n1.+1 then gcdpE_rec n1 qq rr else rr.
+have Irec: forall k l p q, size q <= k -> size q <= l
+ -> size q < size p -> gcdpE_rec k p q = gcdpE_rec l p q.
++ elim=> [|m Hrec] [|n] //= p1 q1.
+ - rewrite leqn0 size_poly_eq0; move/eqP=> -> _.
+ rewrite size_poly0 size_poly_gt0 modp0 => nzp.
+ by rewrite (negPf nzp); case: n => [|n] /=; rewrite mod0p eqxx.
+ - rewrite leqn0 size_poly_eq0 => _; move/eqP=> ->.
+ rewrite size_poly0 size_poly_gt0 modp0 => nzp.
+ by rewrite (negPf nzp); case: m {Hrec} => [|m] /=; rewrite mod0p eqxx.
+ case: ifP => Epq Sm Sn Sq //; rewrite ?Epq //.
+ case: (eqVneq q1 0) => [->|nzq].
+ by case: n m {Sm Sn Hrec} => [|m] [|n] //=; rewrite mod0p eqxx.
+ apply: Hrec; last by rewrite ltn_modp.
+ by rewrite -ltnS (leq_trans _ Sm) // ltn_modp.
+ by rewrite -ltnS (leq_trans _ Sn) // ltn_modp.
+case: (eqVneq p 0) => [-> | nzp].
+ by rewrite mod0p modp0 gcd0p gcdp0 if_same.
+case: (eqVneq q 0) => [-> | nzq].
+ by rewrite mod0p modp0 gcd0p gcdp0 if_same.
+rewrite /gcdp /gcdp_rec.
+case: ltnP; rewrite (negPf nzp, negPf nzq) //=.
+ move=> ltpq; rewrite ltn_modp (negPf nzp) //=.
+ rewrite -(ltn_predK ltpq) /=; case: eqP => [->|].
+ by case: (size p) => [|[|s]]; rewrite /= modp0 (negPf nzp) // mod0p eqxx.
+ move/eqP=> nzqp; rewrite (negPf nzp).
+ apply: Irec => //; last by rewrite ltn_modp.
+ by rewrite -ltnS (ltn_predK ltpq) (leq_trans _ ltpq) ?leqW // ltn_modp.
+ by rewrite ltnW // ltn_modp.
+move=> leqp; rewrite ltn_modp (negPf nzq) //=.
+have p_gt0: size p > 0 by rewrite size_poly_gt0.
+rewrite -(prednK p_gt0) /=; case: eqP => [->|].
+ by case: (size q) => [|[|s]]; rewrite /= modp0 (negPf nzq) // mod0p eqxx.
+move/eqP=> nzpq; rewrite (negPf nzq); apply: Irec => //; rewrite ?ltn_modp //.
+ by rewrite -ltnS (prednK p_gt0) (leq_trans _ leqp) // ltn_modp.
+by rewrite ltnW // ltn_modp.
+Qed.
+
+Lemma size_gcd1p p : size (gcdp 1 p) = 1%N.
+Proof.
+rewrite gcdpE size_polyC oner_eq0 /= modp1; case: ltnP.
+ by rewrite gcd0p size_polyC oner_eq0.
+move/size1_polyC=> e; rewrite e.
+case p00: (p`_0 == 0); first by rewrite (eqP p00) modp0 gcdp0 size_poly1.
+by rewrite modpC ?p00 // gcd0p size_polyC p00.
+Qed.
+
+Lemma size_gcdp1 p : size (gcdp p 1) = 1%N.
+rewrite gcdpE size_polyC oner_eq0 /= modp1; case: ltnP; last first.
+ by rewrite gcd0p size_polyC oner_eq0.
+rewrite ltnS leqn0 size_poly_eq0; move/eqP->; rewrite gcdp0 modp0 size_polyC.
+by rewrite oner_eq0.
+Qed.
+
+Lemma gcdpp : idempotent gcdp.
+Proof. by move=> p; rewrite gcdpE ltnn modpp gcd0p. Qed.
+
+Lemma dvdp_gcdlr p q : (gcdp p q %| p) && (gcdp p q %| q).
+Proof.
+elim: {p q}minn {-2}p {-2}q (leqnn (minn (size q) (size p))) => [|r Hrec] p q.
+ rewrite geq_min !leqn0 !size_poly_eq0.
+ by case/pred2P=> ->; rewrite (gcdp0, gcd0p) dvdpp ?andbT /=.
+case: (eqVneq p 0) => [-> _|nz_p]; first by rewrite gcd0p dvdpp andbT.
+case: (eqVneq q 0) => [->|nz_q]; first by rewrite gcdp0 dvdpp /=.
+rewrite gcdpE minnC /minn; case: ltnP => [lt_pq | le_pq] le_qr.
+ suffices: minn (size p) (size (q %% p)) <= r.
+ by move/Hrec; case/andP => E1 E2; rewrite E2 (dvdp_mod _ E2).
+ by rewrite geq_min orbC -ltnS (leq_trans _ le_qr) ?ltn_modp.
+suffices: minn (size q) (size (p %% q)) <= r.
+ by move/Hrec; case/andP => E1 E2; rewrite E2 andbT (dvdp_mod _ E2).
+by rewrite geq_min orbC -ltnS (leq_trans _ le_qr) ?ltn_modp.
+Qed.
+
+Lemma dvdp_gcdl p q : gcdp p q %| p.
+Proof. by case/andP: (dvdp_gcdlr p q). Qed.
+
+Lemma dvdp_gcdr p q :gcdp p q %| q.
+Proof. by case/andP: (dvdp_gcdlr p q). Qed.
+
+Lemma leq_gcdpl p q : p != 0 -> size (gcdp p q) <= size p.
+Proof. by move=> pn0; move: (dvdp_gcdl p q); apply: dvdp_leq. Qed.
+
+Lemma leq_gcdpr p q : q != 0 -> size (gcdp p q) <= size q.
+Proof. by move=> qn0; move: (dvdp_gcdr p q); apply: dvdp_leq. Qed.
+
+Lemma dvdp_gcd p m n : p %| gcdp m n = (p %| m) && (p %| n).
+Proof.
+apply/idP/andP=> [dv_pmn | [dv_pm dv_pn]].
+ by rewrite ?(dvdp_trans dv_pmn) ?dvdp_gcdl ?dvdp_gcdr.
+move: (leqnn (minn (size n) (size m))) dv_pm dv_pn.
+elim: {m n}minn {-2}m {-2}n => [|r Hrec] m n.
+ rewrite geq_min !leqn0 !size_poly_eq0.
+ by case/pred2P=> ->; rewrite (gcdp0, gcd0p).
+case: (eqVneq m 0) => [-> _|nz_m]; first by rewrite gcd0p /=.
+case: (eqVneq n 0) => [->|nz_n]; first by rewrite gcdp0 /=.
+rewrite gcdpE minnC /minn; case: ltnP => Cnm le_r dv_m dv_n.
+ apply: Hrec => //; last by rewrite -(dvdp_mod _ dv_m).
+ by rewrite geq_min orbC -ltnS (leq_trans _ le_r) ?ltn_modp.
+apply: Hrec => //; last by rewrite -(dvdp_mod _ dv_n).
+by rewrite geq_min orbC -ltnS (leq_trans _ le_r) ?ltn_modp.
+Qed.
+
+
+Lemma gcdpC : forall p q, gcdp p q %= gcdp q p.
+Proof. by move=> p q; rewrite /eqp !dvdp_gcd !dvdp_gcdl !dvdp_gcdr. Qed.
+
+Lemma gcd1p p : gcdp 1 p %= 1.
+Proof.
+rewrite -size_poly_eq1 gcdpE size_poly1; case: ltnP.
+ by rewrite modp1 gcd0p size_poly1 eqxx.
+move/size1_polyC=> e; rewrite e.
+case p00: (p`_0 == 0); first by rewrite (eqP p00) modp0 gcdp0 size_poly1.
+by rewrite modpC ?p00 // gcd0p size_polyC p00.
+Qed.
+
+Lemma gcdp1 p : gcdp p 1 %= 1.
+Proof. by rewrite (eqp_ltrans (gcdpC _ _)) gcd1p. Qed.
+
+Lemma gcdp_addl_mul p q r: gcdp r (p * r + q) %= gcdp r q.
+Proof.
+suff h m n d : gcdp d n %| gcdp d (m * d + n).
+ apply/andP; split => //; rewrite {2}(_: q = (-p) * r + (p * r + q)) ?H //.
+ by rewrite GRing.mulNr GRing.addKr.
+by rewrite dvdp_gcd dvdp_gcdl /= dvdp_addr ?dvdp_gcdr ?dvdp_mull ?dvdp_gcdl.
+Qed.
+
+Lemma gcdp_addl m n : gcdp m (m + n) %= gcdp m n.
+Proof. by rewrite -{2}(mul1r m) gcdp_addl_mul. Qed.
+
+Lemma gcdp_addr m n : gcdp m (n + m) %= gcdp m n.
+Proof. by rewrite addrC gcdp_addl. Qed.
+
+Lemma gcdp_mull m n : gcdp n (m * n) %= n.
+Proof.
+case: (eqVneq n 0) => [-> | nn0]; first by rewrite gcd0p mulr0 eqpxx.
+case: (eqVneq m 0) => [-> | mn0]; first by rewrite mul0r gcdp0 eqpxx.
+rewrite gcdpE modp_mull gcd0p size_mul //; case: ifP; first by rewrite eqpxx.
+rewrite (polySpred mn0) addSn /= -{1}[size n]add0n ltn_add2r; move/negbT.
+rewrite -ltnNge prednK ?size_poly_gt0 // leq_eqVlt ltnS leqn0 size_poly_eq0.
+rewrite (negPf mn0) orbF; case/size_poly1P=> c cn0 -> {mn0 m}; rewrite mul_polyC.
+suff -> : n %% (c *: n) = 0 by rewrite gcd0p; exact: eqp_scale.
+by apply/modp_eq0P; rewrite dvdp_scalel.
+Qed.
+
+Lemma gcdp_mulr m n : gcdp n (n * m) %= n.
+Proof. by rewrite mulrC gcdp_mull. Qed.
+
+Lemma gcdp_scalel c m n : c != 0 -> gcdp (c *: m) n %= gcdp m n.
+Proof.
+move=> cn0; rewrite /eqp dvdp_gcd [gcdp m n %| _]dvdp_gcd !dvdp_gcdr !andbT.
+apply/andP; split; last first.
+ by apply: dvdp_trans (dvdp_gcdl _ _) _; rewrite dvdp_scaler.
+by apply: dvdp_trans (dvdp_gcdl _ _) _; rewrite dvdp_scalel.
+Qed.
+
+Lemma gcdp_scaler c m n : c != 0 -> gcdp m (c *: n) %= gcdp m n.
+Proof.
+move=> cn0; apply: eqp_trans (gcdpC _ _) _.
+apply: eqp_trans (gcdp_scalel _ _ _) _ => //; exact: gcdpC.
+Qed.
+
+Lemma dvdp_gcd_idl m n : m %| n -> gcdp m n %= m.
+Proof.
+case: (eqVneq m 0) => [-> | mn0].
+ by rewrite dvd0p => /eqP ->; rewrite gcdp0 eqpxx.
+rewrite dvdp_eq; move/eqP; move/(f_equal (gcdp m)) => h.
+apply: eqp_trans (gcdp_mull (n %/ m) _); rewrite -h eqp_sym gcdp_scaler //.
+by rewrite expf_neq0 // lead_coef_eq0.
+Qed.
+
+Lemma dvdp_gcd_idr m n : n %| m -> gcdp m n %= n.
+Proof. move/dvdp_gcd_idl => h; apply: eqp_trans h; exact: gcdpC. Qed.
+
+Lemma gcdp_exp p k l : gcdp (p ^+ k) (p ^+ l) %= p ^+ minn k l.
+Proof.
+wlog leqmn: k l / k <= l.
+ move=> hwlog; case: (leqP k l); first exact: hwlog.
+ move/ltnW; rewrite minnC; move/hwlog=> h; apply: eqp_trans h; exact: gcdpC.
+rewrite (minn_idPl leqmn); move/subnK: leqmn<-; rewrite exprD.
+apply: eqp_trans (gcdp_mull _ _) _; exact: eqpxx.
+Qed.
+
+Lemma gcdp_eq0 p q : gcdp p q == 0 = (p == 0) && (q == 0).
+Proof.
+apply/idP/idP; last by case/andP => /eqP -> /eqP ->; rewrite gcdp0.
+have h m n: gcdp m n == 0 -> (m == 0).
+ by rewrite -(dvd0p m); move/eqP<-; rewrite dvdp_gcdl.
+by move=> ?; rewrite (h _ q) // (h _ p) // -eqp0 (eqp_ltrans (gcdpC _ _)) eqp0.
+Qed.
+
+Lemma eqp_gcdr p q r : q %= r -> gcdp p q %= gcdp p r.
+Proof.
+move=> eqr; rewrite /eqp !(dvdp_gcd, dvdp_gcdl, andbT) /=.
+by rewrite -(eqp_dvdr _ eqr) dvdp_gcdr (eqp_dvdr _ eqr) dvdp_gcdr.
+Qed.
+
+Lemma eqp_gcdl r p q : p %= q -> gcdp p r %= gcdp q r.
+move=> eqr; rewrite /eqp !(dvdp_gcd, dvdp_gcdr, andbT) /=.
+by rewrite -(eqp_dvdr _ eqr) dvdp_gcdl (eqp_dvdr _ eqr) dvdp_gcdl.
+Qed.
+
+Lemma eqp_gcd p1 p2 q1 q2 : p1 %= p2 -> q1 %= q2 -> gcdp p1 q1 %= gcdp p2 q2.
+Proof.
+move=> e1 e2.
+by apply: eqp_trans (eqp_gcdr _ e2); apply: eqp_trans (eqp_gcdl _ e1).
+Qed.
+
+Lemma eqp_rgcd_gcd p q : rgcdp p q %= gcdp p q.
+Proof.
+move: (leqnn (minn (size p) (size q))); move: {2}(minn (size p) (size q)) => n.
+elim: n p q => [p q|n ihn p q hs].
+ rewrite leqn0 /minn; case: ltnP => _; rewrite size_poly_eq0; move/eqP->.
+ by rewrite gcd0p rgcd0p eqpxx.
+ by rewrite gcdp0 rgcdp0 eqpxx.
+case: (eqVneq p 0) => [-> | pn0]; first by rewrite gcd0p rgcd0p eqpxx.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite gcdp0 rgcdp0 eqpxx.
+rewrite gcdpE rgcdpE; case: ltnP => sp.
+ have e := (eqp_rmod_mod q p); move: (e); move/(eqp_gcdl p) => h.
+ apply: eqp_trans h; apply: ihn; rewrite (eqp_size e) geq_min.
+ by rewrite -ltnS (leq_trans _ hs) // (minn_idPl (ltnW _)) ?ltn_modp.
+have e := (eqp_rmod_mod p q); move: (e); move/(eqp_gcdl q) => h.
+apply: eqp_trans h; apply: ihn; rewrite (eqp_size e) geq_min.
+by rewrite -ltnS (leq_trans _ hs) // (minn_idPr _) ?ltn_modp.
+Qed.
+
+Lemma gcdp_modr m n : gcdp m (n %% m) %= gcdp m n.
+Proof.
+case: (eqVneq m 0) => [-> | mn0]; first by rewrite modp0 eqpxx.
+have : (lead_coef m) ^+ (scalp n m) != 0 by rewrite expf_neq0 // lead_coef_eq0.
+move/gcdp_scaler; move/(_ m n) => h; apply: eqp_trans h; rewrite divp_eq.
+by rewrite eqp_sym gcdp_addl_mul.
+Qed.
+
+Lemma gcdp_modl m n : gcdp (m %% n) n %= gcdp m n.
+Proof.
+apply: eqp_trans (gcdpC _ _) _; apply: eqp_trans (gcdp_modr _ _) _.
+exact: gcdpC.
+Qed.
+
+Lemma gcdp_def d m n :
+ d %| m -> d %| n -> (forall d', d' %| m -> d' %| n -> d' %| d) ->
+ gcdp m n %= d.
+Proof.
+move=> dm dn h; rewrite /eqp dvdp_gcd dm dn !andbT.
+apply: h; [exact: dvdp_gcdl | exact: dvdp_gcdr].
+Qed.
+
+Definition coprimep p q := size (gcdp p q) == 1%N.
+
+Lemma coprimep_size_gcd p q : coprimep p q -> size (gcdp p q) = 1%N.
+Proof. by rewrite /coprimep=> /eqP. Qed.
+
+Lemma coprimep_def p q : (coprimep p q) = (size (gcdp p q) == 1%N).
+Proof. done. Qed.
+
+Lemma coprimep_scalel c m n :
+ c != 0 -> coprimep (c *: m) n = coprimep m n.
+Proof. by move=> ?; rewrite !coprimep_def (eqp_size (gcdp_scalel _ _ _)). Qed.
+
+Lemma coprimep_scaler c m n:
+ c != 0 -> coprimep m (c *: n) = coprimep m n.
+Proof. by move=> ?; rewrite !coprimep_def (eqp_size (gcdp_scaler _ _ _)). Qed.
+
+Lemma coprimepp p : coprimep p p = (size p == 1%N).
+Proof. by rewrite coprimep_def gcdpp. Qed.
+
+Lemma gcdp_eqp1 p q : gcdp p q %= 1 = (coprimep p q).
+Proof. by rewrite coprimep_def size_poly_eq1. Qed.
+
+Lemma coprimep_sym p q : coprimep p q = coprimep q p.
+Proof.
+by rewrite -!gcdp_eqp1; apply: eqp_ltrans; rewrite gcdpC.
+Qed.
+
+Lemma coprime1p p : coprimep 1 p.
+Proof.
+rewrite /coprimep -[1%N](size_poly1 R); apply/eqP; apply: eqp_size.
+exact: gcd1p.
+Qed.
+
+Lemma coprimep1 p : coprimep p 1.
+Proof. by rewrite coprimep_sym; apply: coprime1p. Qed.
+
+Lemma coprimep0 p : coprimep p 0 = (p %= 1).
+Proof. by rewrite /coprimep gcdp0 size_poly_eq1. Qed.
+
+Lemma coprime0p p : coprimep 0 p = (p %= 1).
+Proof. by rewrite coprimep_sym coprimep0. Qed.
+
+(* This is different from coprimeP in div. shall we keep this? *)
+Lemma coprimepP p q :
+ reflect (forall d, d %| p -> d %| q -> d %= 1) (coprimep p q).
+Proof.
+apply: (iffP idP)=> [|h].
+ rewrite /coprimep; move/eqP=> hs d dvddp dvddq.
+ have dvddg: d %| gcdp p q by rewrite dvdp_gcd dvddp dvddq.
+ by apply: (dvdp_eqp1 dvddg); rewrite -size_poly_eq1; apply/eqP.
+case/andP: (dvdp_gcdlr p q)=> h1 h2.
+by rewrite /coprimep size_poly_eq1; apply: h.
+Qed.
+
+Lemma coprimepPn p q : p != 0 ->
+ reflect (exists d, (d %| gcdp p q) && ~~ (d %= 1)) (~~ coprimep p q).
+Proof.
+move=> p0; apply: (iffP idP).
+ by rewrite -gcdp_eqp1=> ng1; exists (gcdp p q); rewrite dvdpp /=.
+case=> d; case/andP=> dg; apply: contra; rewrite -gcdp_eqp1=> g1.
+by move: dg; rewrite (eqp_dvdr _ g1) dvdp1 size_poly_eq1.
+Qed.
+
+Lemma coprimep_dvdl q p r : r %| q -> coprimep p q -> coprimep p r.
+Proof.
+move=> rq cpq; apply/coprimepP=> d dp dr; move/coprimepP:cpq=> cpq'.
+by apply: cpq'; rewrite // (dvdp_trans dr).
+Qed.
+
+Lemma coprimep_dvdr p q r :
+ r %| p -> coprimep p q -> coprimep r q.
+Proof.
+move=> rp; rewrite ![coprimep _ q]coprimep_sym.
+by move/coprimep_dvdl; apply.
+Qed.
+
+
+Lemma coprimep_modl p q : coprimep (p %% q) q = coprimep p q.
+Proof.
+symmetry; rewrite !coprimep_def.
+case: (ltnP (size p) (size q))=> hpq; first by rewrite modp_small.
+by rewrite gcdpE ltnNge hpq.
+Qed.
+
+Lemma coprimep_modr q p : coprimep q (p %% q) = coprimep q p.
+Proof. by rewrite ![coprimep q _]coprimep_sym coprimep_modl. Qed.
+
+Lemma rcoprimep_coprimep q p : rcoprimep q p = coprimep q p.
+Proof.
+by rewrite /coprimep /rcoprimep; rewrite (eqp_size (eqp_rgcd_gcd _ _)).
+Qed.
+
+Lemma eqp_coprimepr p q r : q %= r -> coprimep p q = coprimep p r.
+Proof.
+by rewrite -!gcdp_eqp1; move/(eqp_gcdr p) => h1; apply: (eqp_ltrans h1).
+Qed.
+
+Lemma eqp_coprimepl p q r : q %= r -> coprimep q p = coprimep r p.
+Proof. rewrite !(coprimep_sym _ p); exact: eqp_coprimepr. Qed.
+
+(* This should be implemented with an extended remainder sequence *)
+Fixpoint egcdp_rec p q k {struct k} : {poly R} * {poly R} :=
+ if k is k'.+1 then
+ if q == 0 then (1, 0) else
+ let: (u, v) := egcdp_rec q (p %% q) k' in
+ (lead_coef q ^+ scalp p q *: v, (u - v * (p %/ q)))
+ else (1, 0).
+
+Definition egcdp p q :=
+ if size q <= size p then egcdp_rec p q (size q)
+ else let e := egcdp_rec q p (size p) in (e.2, e.1).
+
+(* No provable egcd0p *)
+Lemma egcdp0 p : egcdp p 0 = (1, 0).
+Proof. by rewrite /egcdp size_poly0. Qed.
+
+Lemma egcdp_recP : forall k p q, q != 0 -> size q <= k -> size q <= size p ->
+ let e := (egcdp_rec p q k) in
+ [/\ size e.1 <= size q, size e.2 <= size p & gcdp p q %= e.1 * p + e.2 * q].
+Proof.
+elim=> [|k ihk] p q /= qn0; first by rewrite leqn0 size_poly_eq0 (negPf qn0).
+move=> sqSn qsp; case: (eqVneq q 0)=> q0; first by rewrite q0 eqxx in qn0.
+rewrite (negPf qn0).
+have sp : size p > 0 by apply: leq_trans qsp; rewrite size_poly_gt0.
+case: (eqVneq (p %% q) 0) => [r0 | rn0] /=.
+ rewrite r0 /egcdp_rec; case: k ihk sqSn => [|n] ihn sqSn /=.
+ rewrite !scaler0 !mul0r subr0 add0r mul1r size_poly0 size_poly1.
+ by rewrite dvdp_gcd_idr /dvdp ?r0.
+ rewrite !eqxx mul0r scaler0 /= mul0r add0r subr0 mul1r size_poly0 size_poly1.
+ by rewrite dvdp_gcd_idr /dvdp ?r0 //.
+have h1 : size (p %% q) <= k.
+ by rewrite -ltnS; apply: leq_trans sqSn; rewrite ltn_modp.
+have h2 : size (p %% q) <= size q by rewrite ltnW // ltn_modp.
+have := (ihk q (p %% q) rn0 h1 h2).
+case: (egcdp_rec _ _)=> u v /= => [[ihn'1 ihn'2 ihn'3]].
+rewrite gcdpE ltnNge qsp //= (eqp_ltrans (gcdpC _ _)); split; last first.
+- apply: (eqp_trans ihn'3).
+ rewrite mulrBl addrCA -scalerAl scalerAr -mulrA -mulrBr.
+ by rewrite divp_eq addrAC subrr add0r eqpxx.
+- apply: (leq_trans (size_add _ _)).
+ case: (eqVneq v 0)=> [-> | vn0].
+ rewrite mul0r size_opp size_poly0 maxn0; apply: leq_trans ihn'1 _.
+ exact: leq_modp.
+ case: (eqVneq (p %/ q) 0)=> [-> | qqn0].
+ rewrite mulr0 size_opp size_poly0 maxn0; apply: leq_trans ihn'1 _.
+ exact: leq_modp.
+ rewrite geq_max (leq_trans ihn'1) ?leq_modp //= size_opp size_mul //.
+ move: (ihn'2); rewrite -(leq_add2r (size (p %/ q))).
+ have : size v + size (p %/ q) > 0 by rewrite addn_gt0 size_poly_gt0 vn0.
+ have : size q + size (p %/ q) > 0 by rewrite addn_gt0 size_poly_gt0 qn0.
+ do 2! move/prednK=> {1}<-; rewrite ltnS => h; apply: leq_trans h _.
+ rewrite size_divp // addnBA; last by apply: leq_trans qsp; exact: leq_pred.
+ rewrite addnC -addnBA ?leq_pred //; move: qn0; rewrite -size_poly_eq0 -lt0n.
+ by move/prednK=> {1}<-; rewrite subSnn addn1.
+- by rewrite size_scale // lc_expn_scalp_neq0.
+Qed.
+
+Lemma egcdpP p q : p != 0 -> q != 0 -> forall (e := egcdp p q),
+ [/\ size e.1 <= size q, size e.2 <= size p & gcdp p q %= e.1 * p + e.2 * q].
+Proof.
+move=> pn0 qn0; rewrite /egcdp; case: (leqP (size q) (size p)) => /= hp.
+ by apply: egcdp_recP.
+move/ltnW: hp => hp; case: (egcdp_recP pn0 (leqnn (size p)) hp) => h1 h2 h3.
+by split => //; rewrite (eqp_ltrans (gcdpC _ _)) addrC.
+Qed.
+
+Lemma egcdpE p q (e := egcdp p q) : gcdp p q %= e.1 * p + e.2 * q.
+Proof.
+rewrite {}/e; have [-> /= | qn0] := eqVneq q 0.
+ by rewrite gcdp0 egcdp0 mul1r mulr0 addr0.
+have [p0 | pn0] := eqVneq p 0; last by case: (egcdpP pn0 qn0).
+rewrite p0 gcd0p mulr0 add0r /egcdp size_poly0 leqn0 size_poly_eq0 (negPf qn0).
+by rewrite /= mul1r.
+Qed.
+
+Lemma Bezoutp p q : exists u, u.1 * p + u.2 * q %= (gcdp p q).
+Proof.
+case: (eqVneq p 0) => [-> | pn0].
+ by rewrite gcd0p; exists (0, 1); rewrite mul0r mul1r add0r.
+case: (eqVneq q 0) => [-> | qn0].
+ by rewrite gcdp0; exists (1, 0); rewrite mul0r mul1r addr0.
+pose e := egcdp p q; exists e; rewrite eqp_sym.
+by case: (egcdpP pn0 qn0).
+Qed.
+
+Lemma Bezout_coprimepP : forall p q,
+ reflect (exists u, u.1 * p + u.2 * q %= 1) (coprimep p q).
+Proof.
+move=> p q; rewrite -gcdp_eqp1; apply:(iffP idP)=> [g1|].
+ case: (Bezoutp p q) => [[u v] Puv]; exists (u, v); exact: eqp_trans g1.
+case=>[[u v]]; rewrite eqp_sym=> Puv; rewrite /eqp (eqp_dvdr _ Puv).
+by rewrite dvdp_addr dvdp_mull ?dvdp_gcdl ?dvdp_gcdr //= dvd1p.
+Qed.
+
+Lemma coprimep_root p q x : coprimep p q -> root p x -> q.[x] != 0.
+Proof.
+case/Bezout_coprimepP=> [[u v] euv] px0.
+move/eqpP: euv => [[c1 c2]] /andP /= [c1n0 c2n0 e].
+suffices: c1 * (v.[x] * q.[x]) != 0.
+ by rewrite !mulf_eq0 !negb_or c1n0 /=; case/andP.
+move/(f_equal (fun t => horner t x)): e; rewrite /= !hornerZ hornerD.
+by rewrite !hornerM (eqP px0) mulr0 add0r hornerC mulr1; move->.
+Qed.
+
+Lemma Gauss_dvdpl p q d: coprimep d q -> (d %| p * q) = (d %| p).
+Proof.
+move/Bezout_coprimepP=>[[u v] Puv]; apply/idP/idP; last exact: dvdp_mulr.
+move:Puv; move/(eqp_mull p); rewrite mulr1 mulrDr eqp_sym=> peq dpq.
+rewrite (eqp_dvdr _ peq) dvdp_addr; first by rewrite mulrA mulrAC dvdp_mulr.
+by rewrite mulrA dvdp_mull ?dvdpp.
+Qed.
+
+Lemma Gauss_dvdpr p q d: coprimep d q -> (d %| q * p) = (d %| p).
+Proof. rewrite mulrC; exact: Gauss_dvdpl. Qed.
+
+(* This could be simplified with the introduction of lcmp *)
+Lemma Gauss_dvdp m n p : coprimep m n -> (m * n %| p) = (m %| p) && (n %| p).
+Proof.
+case: (eqVneq m 0) => [-> | mn0].
+ by rewrite coprime0p; move/eqp_dvdl->; rewrite !mul0r dvd0p dvd1p andbT.
+case: (eqVneq n 0) => [-> | nn0].
+ by rewrite coprimep0; move/eqp_dvdl->; rewrite !mulr0 dvd1p.
+move=> hc; apply/idP/idP.
+ move/Gauss_dvdpl: hc => <- h; move/(dvdp_mull m): (h); rewrite dvdp_mul2l //.
+ move->; move/(dvdp_mulr n): (h); rewrite dvdp_mul2r // andbT.
+ exact: dvdp_mulr.
+case/andP => dmp dnp; move: (dnp); rewrite dvdp_eq.
+set c2 := _ ^+ _; set q2 := _ %/ _; move/eqP=> e2.
+have := (sym_eq (Gauss_dvdpl q2 hc)); rewrite -e2.
+have -> : m %| c2 *: p by rewrite -mul_polyC dvdp_mull.
+rewrite dvdp_eq; set c3 := _ ^+ _; set q3 := _ %/ _; move/eqP=> e3.
+apply: (@eq_dvdp (c3 * c2) q3).
+ by rewrite mulf_neq0 // expf_neq0 // lead_coef_eq0.
+by rewrite mulrA -e3 -scalerAl -e2 scalerA.
+Qed.
+
+Lemma Gauss_gcdpr p m n : coprimep p m -> gcdp p (m * n) %= gcdp p n.
+Proof.
+move=> co_pm; apply/eqP; rewrite /eqp !dvdp_gcd !dvdp_gcdl /= andbC.
+rewrite dvdp_mull ?dvdp_gcdr // -(@Gauss_dvdpl _ m).
+ by rewrite mulrC dvdp_gcdr.
+apply/coprimepP=> d; rewrite dvdp_gcd; case/andP=> hdp _ hdm.
+by move/coprimepP: co_pm; apply.
+Qed.
+
+Lemma Gauss_gcdpl p m n : coprimep p n -> gcdp p (m * n) %= gcdp p m.
+Proof. by move=> co_pn; rewrite mulrC Gauss_gcdpr. Qed.
+
+Lemma coprimep_mulr p q r : coprimep p (q * r) = (coprimep p q && coprimep p r).
+Proof.
+apply/coprimepP/andP=> [hp|[/coprimepP hq hr]].
+ split; apply/coprimepP=> d dp dq; rewrite hp //;
+ [exact: dvdp_mulr|exact: dvdp_mull].
+move=> d dp dqr; move/(_ _ dp) in hq.
+rewrite Gauss_dvdpl in dqr; first exact: hq.
+by move/coprimep_dvdr:hr; apply.
+Qed.
+
+Lemma coprimep_mull p q r: coprimep (q * r) p = (coprimep q p && coprimep r p).
+Proof. by rewrite ![coprimep _ p]coprimep_sym coprimep_mulr. Qed.
+
+Lemma modp_coprime k u n : k != 0 -> (k * u) %% n %= 1 -> coprimep k n.
+Proof.
+move=> kn0 hmod; apply/Bezout_coprimepP.
+exists (((lead_coef n)^+(scalp (k * u) n) *: u), (- (k * u %/ n))).
+rewrite -scalerAl mulrC (divp_eq (u * k) n) mulNr -addrAC subrr add0r.
+by rewrite mulrC.
+Qed.
+
+Lemma coprimep_pexpl k m n : 0 < k -> coprimep (m ^+ k) n = coprimep m n.
+Proof.
+case: k => // k _; elim: k => [|k IHk]; first by rewrite expr1.
+by rewrite exprS coprimep_mull -IHk andbb.
+Qed.
+
+Lemma coprimep_pexpr k m n : 0 < k -> coprimep m (n ^+ k) = coprimep m n.
+Proof. by move=> k_gt0; rewrite !(coprimep_sym m) coprimep_pexpl. Qed.
+
+Lemma coprimep_expl k m n : coprimep m n -> coprimep (m ^+ k) n.
+Proof. by case: k => [|k] co_pm; rewrite ?coprime1p // coprimep_pexpl. Qed.
+
+Lemma coprimep_expr k m n : coprimep m n -> coprimep m (n ^+ k).
+Proof. by rewrite !(coprimep_sym m); exact: coprimep_expl. Qed.
+
+Lemma gcdp_mul2l p q r : gcdp (p * q) (p * r) %= (p * gcdp q r).
+Proof.
+case: (eqVneq p 0)=> [->|hp]; first by rewrite !mul0r gcdp0 eqpxx.
+rewrite /eqp !dvdp_gcd !dvdp_mul2l // dvdp_gcdr dvdp_gcdl !andbT.
+move: (Bezoutp q r) => [[u v]] huv.
+rewrite eqp_sym in huv; rewrite (eqp_dvdr _ (eqp_mull _ huv)).
+rewrite mulrDr ![p * (_ * _)]mulrCA.
+by apply: dvdp_add; rewrite dvdp_mull// (dvdp_gcdr, dvdp_gcdl).
+Qed.
+
+Lemma gcdp_mul2r q r p : gcdp (q * p) (r * p) %= (gcdp q r * p).
+Proof. by rewrite ![_ * p]GRing.mulrC gcdp_mul2l. Qed.
+
+Lemma mulp_gcdr p q r : r * (gcdp p q) %= gcdp (r * p) (r * q).
+Proof. by rewrite eqp_sym gcdp_mul2l. Qed.
+
+Lemma mulp_gcdl p q r : (gcdp p q) * r %= gcdp (p * r) (q * r).
+Proof. by rewrite eqp_sym gcdp_mul2r. Qed.
+
+Lemma coprimep_div_gcd p q : (p != 0) || (q != 0) ->
+ coprimep (p %/ (gcdp p q)) (q %/ gcdp p q).
+Proof.
+move=> hpq.
+have gpq0: gcdp p q != 0 by rewrite gcdp_eq0 negb_and.
+rewrite -gcdp_eqp1 -(@eqp_mul2r (gcdp p q)) // mul1r.
+have: gcdp p q %| p by rewrite dvdp_gcdl.
+have: gcdp p q %| q by rewrite dvdp_gcdr.
+rewrite !dvdp_eq eq_sym; move/eqP=> hq; rewrite eq_sym; move/eqP=> hp.
+rewrite (eqp_ltrans (mulp_gcdl _ _ _)) hq hp.
+have lcn0 k : (lead_coef (gcdp p q)) ^+ k != 0.
+ by rewrite expf_neq0 ?lead_coef_eq0.
+by apply: eqp_gcd; rewrite ?eqp_scale.
+Qed.
+
+Lemma divp_eq0 p q : (p %/ q == 0) = [|| p == 0, q ==0 | size p < size q].
+Proof.
+apply/eqP/idP=> [d0|]; last first.
+ case/or3P; [by move/eqP->; rewrite div0p| by move/eqP->; rewrite divp0|].
+ by move/divp_small.
+case: (eqVneq p 0) => [->|pn0]; first by rewrite eqxx.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite eqxx orbT.
+move: (divp_eq p q); rewrite d0 mul0r add0r.
+move/(f_equal (fun x : {poly R} => size x)).
+by rewrite size_scale ?lc_expn_scalp_neq0 // => ->; rewrite ltn_modp qn0 !orbT.
+Qed.
+
+Lemma dvdp_div_eq0 p q : q %| p -> (p %/ q == 0) = (p == 0).
+Proof.
+move=> dvdp_qp; have [->|p_neq0] := altP (p =P 0); first by rewrite div0p eqxx.
+rewrite divp_eq0 ltnNge dvdp_leq // (negPf p_neq0) orbF /=.
+by apply: contraTF dvdp_qp=> /eqP ->; rewrite dvd0p.
+Qed.
+
+Lemma Bezout_coprimepPn p q : p != 0 -> q != 0 ->
+ reflect (exists2 uv : {poly R} * {poly R},
+ (0 < size uv.1 < size q) && (0 < size uv.2 < size p) &
+ uv.1 * p = uv.2 * q)
+ (~~ (coprimep p q)).
+move=> pn0 qn0; apply: (iffP idP); last first.
+ case=> [[u v] /= /andP [/andP [ps1 s1] /andP [ps2 s2]] e].
+ have: ~~(size (q * p) <= size (u * p)).
+ rewrite -ltnNge !size_mul // -?size_poly_gt0 // (polySpred pn0) !addnS.
+ by rewrite ltn_add2r.
+ apply: contra => ?; apply: dvdp_leq; rewrite ?mulf_neq0 // -?size_poly_gt0 //.
+ by rewrite mulrC Gauss_dvdp // dvdp_mull // e dvdp_mull.
+rewrite coprimep_def neq_ltn.
+case/orP; first by rewrite ltnS leqn0 size_poly_eq0 gcdp_eq0 -[p == 0]negbK pn0.
+case sg: (size (gcdp p q)) => [|n] //; case: n sg=> [|n] // sg _.
+move: (dvdp_gcdl p q); rewrite dvdp_eq; set c1 := _ ^+ _; move/eqP=> hu1.
+move: (dvdp_gcdr p q); rewrite dvdp_eq; set c2 := _ ^+ _; move/eqP=> hv1.
+exists (c1 *: (q %/ gcdp p q), c2 *: (p %/ gcdp p q)); last first.
+ by rewrite -!{1}scalerAl !scalerAr hu1 hv1 mulrCA.
+rewrite !{1}size_scale ?lc_expn_scalp_neq0 //= !size_poly_gt0 !divp_eq0.
+rewrite gcdp_eq0 !(negPf pn0) !(negPf qn0) /= -!leqNgt leq_gcdpl //.
+rewrite leq_gcdpr //= !ltn_divpl -?size_poly_eq0 ?sg //.
+rewrite !size_mul // -?size_poly_eq0 ?sg // ![(_ + n.+2)%N]addnS /=.
+by rewrite -{1}(addn0 (size p)) -{1}(addn0 (size q)) !ltn_add2l.
+Qed.
+
+Lemma dvdp_pexp2r m n k : k > 0 -> (m ^+ k %| n ^+ k) = (m %| n).
+Proof.
+move=> k_gt0; apply/idP/idP; last exact: dvdp_exp2r.
+case: (eqVneq n 0) => [-> | nn0] //; case: (eqVneq m 0) => [-> | mn0].
+ move/prednK: k_gt0=> {1}<-; rewrite exprS mul0r //= !dvd0p expf_eq0.
+ by case/andP=> _ ->.
+set d := gcdp m n; have := (dvdp_gcdr m n); rewrite -/d dvdp_eq.
+set c1 := _ ^+ _; set n' := _ %/ _; move/eqP=> def_n.
+have := (dvdp_gcdl m n); rewrite -/d dvdp_eq.
+set c2 := _ ^+ _; set m' := _ %/ _; move/eqP=> def_m.
+have dn0 : d != 0 by rewrite gcdp_eq0 negb_and nn0 orbT.
+have c1n0 : c1 != 0 by rewrite !expf_neq0 // lead_coef_eq0.
+have c2n0 : c2 != 0 by rewrite !expf_neq0 // lead_coef_eq0.
+rewrite -(@dvdp_scaler (c1 ^+ k)) ?expf_neq0 ?lead_coef_eq0 //.
+have c2k_n0 : c2 ^+ k != 0 by rewrite !expf_neq0 // lead_coef_eq0.
+rewrite -(@dvdp_scalel (c2 ^+k)) // -!exprZn def_m def_n !exprMn.
+rewrite dvdp_mul2r ?expf_neq0 //.
+have: coprimep (m' ^+ k) (n' ^+ k).
+ rewrite coprimep_pexpl // coprimep_pexpr //; apply: coprimep_div_gcd.
+ by rewrite nn0 orbT.
+move/coprimepP=> hc hd.
+have /size_poly1P [c cn0 em'] : size m' == 1%N.
+ case: (eqVneq m' 0) => [m'0 |m'_n0].
+ move/eqP: def_m; rewrite m'0 mul0r scale_poly_eq0.
+ by rewrite (negPf mn0) (negPf c2n0).
+ have := (hc _ (dvdpp _) hd); rewrite -size_poly_eq1.
+ rewrite polySpred; last by rewrite expf_eq0 negb_and m'_n0 orbT.
+ rewrite size_exp eqSS muln_eq0; move: k_gt0; rewrite lt0n; move/negPf->.
+ by rewrite orbF -{2}(@prednK (size m')) ?lt0n // size_poly_eq0.
+rewrite -(@dvdp_scalel c2) // def_m em' mul_polyC dvdp_scalel //.
+by rewrite -(@dvdp_scaler c1) // def_n dvdp_mull.
+Qed.
+
+Lemma root_gcd p q x : root (gcdp p q) x = root p x && root q x.
+Proof.
+rewrite /= !root_factor_theorem; apply/idP/andP=> [dg| [dp dq]].
+ by split; apply: dvdp_trans dg _; rewrite ?(dvdp_gcdl, dvdp_gcdr).
+have:= (Bezoutp p q)=> [[[u v]]]; rewrite eqp_sym=> e.
+by rewrite (eqp_dvdr _ e) dvdp_addl dvdp_mull.
+Qed.
+
+Lemma root_biggcd : forall x (ps : seq {poly R}),
+ root (\big[gcdp/0]_(p <- ps) p) x = all (fun p => root p x) ps.
+Proof.
+move=> x; elim; first by rewrite big_nil root0.
+by move=> p ps ihp; rewrite big_cons /= root_gcd ihp.
+Qed.
+
+(* "gdcop Q P" is the Greatest Divisor of P which is coprime to Q *)
+(* if P null, we pose that gdcop returns 1 if Q null, 0 otherwise*)
+Fixpoint gdcop_rec q p k :=
+ if k is m.+1 then
+ if coprimep p q then p
+ else gdcop_rec q (divp p (gcdp p q)) m
+ else (q == 0)%:R.
+
+Definition gdcop q p := gdcop_rec q p (size p).
+
+CoInductive gdcop_spec q p : {poly R} -> Type :=
+ GdcopSpec r of (dvdp r p) & ((coprimep r q) || (p == 0))
+ & (forall d, dvdp d p -> coprimep d q -> dvdp d r)
+ : gdcop_spec q p r.
+
+Lemma gdcop0 q : gdcop q 0 = (q == 0)%:R.
+Proof. by rewrite /gdcop size_poly0. Qed.
+
+Lemma gdcop_recP : forall q p k,
+ size p <= k -> gdcop_spec q p (gdcop_rec q p k).
+Proof.
+move=> q p k; elim: k p => [p | k ihk p] /=.
+ rewrite leqn0 size_poly_eq0; move/eqP->.
+ case q0: (_ == _); split; rewrite ?coprime1p // ?eqxx ?orbT //.
+ by move=> d _; rewrite (eqP q0) coprimep0 dvdp1 size_poly_eq1.
+move=> hs; case cop : (coprimep _ _); first by split; rewrite ?dvdpp ?cop.
+case (eqVneq p 0) => [-> | p0].
+ by rewrite div0p; apply: ihk; rewrite size_poly0 leq0n.
+case: (eqVneq q 0) => [-> | q0].
+ rewrite gcdp0 divpp ?p0 //= => {hs ihk}; case: k=> /=.
+ rewrite eqxx; split; rewrite ?dvd1p ?coprimep0 ?eqpxx //=.
+ by move=> d _; rewrite coprimep0 dvdp1 size_poly_eq1.
+ move=> n; rewrite coprimep0 polyC_eqp1 //; rewrite lc_expn_scalp_neq0.
+ split; first by rewrite (@eqp_dvdl 1) ?dvd1p // polyC_eqp1 lc_expn_scalp_neq0.
+ by rewrite coprimep0 polyC_eqp1 // ?lc_expn_scalp_neq0.
+ by move=> d _; rewrite coprimep0; move/eqp_dvdl->; rewrite dvd1p.
+move: (dvdp_gcdl p q); rewrite dvdp_eq; move/eqP=> e.
+have sgp : size (gcdp p q) <= size p.
+ by apply: dvdp_leq; rewrite ?gcdp_eq0 ?p0 ?q0 // dvdp_gcdl.
+have : p %/ gcdp p q != 0; last move/negPf=>p'n0.
+ move: (dvdp_mulIl (p %/ gcdp p q) (gcdp p q)); move/dvdpN0; apply; rewrite -e.
+ by rewrite scale_poly_eq0 negb_or lc_expn_scalp_neq0.
+have gn0 : gcdp p q != 0.
+ move: (dvdp_mulIr (p %/ gcdp p q) (gcdp p q)); move/dvdpN0; apply; rewrite -e.
+ by rewrite scale_poly_eq0 negb_or lc_expn_scalp_neq0.
+have sp' : size (p %/ (gcdp p q)) <= k.
+ rewrite size_divp ?sgp // leq_subLR (leq_trans hs)//.
+ rewrite -subn_gt0 addnK -subn1 ltn_subRL addn0 ltnNge leq_eqVlt.
+ by rewrite [_ == _]cop ltnS leqn0 size_poly_eq0 (negPf gn0).
+case (ihk _ sp')=> r' dr'p'; first rewrite p'n0 orbF=> cr'q maxr'.
+constructor=> //=; rewrite ?(negPf p0) ?orbF //.
+ apply: (dvdp_trans dr'p'); apply: divp_dvd; exact: dvdp_gcdl.
+move=> d dp cdq; apply: maxr'; last by rewrite cdq.
+case dpq: (d %| gcdp p q).
+ move: (dpq); rewrite dvdp_gcd dp /= => dq; apply: dvdUp; move: cdq.
+ apply: contraLR=> nd1; apply/coprimepPn; last first.
+ by exists d; rewrite dvdp_gcd dvdpp dq nd1.
+ move/negP: p0; move/negP; apply: contra=> d0; move:dp; rewrite (eqP d0).
+ by rewrite dvd0p.
+move: (dp); apply: contraLR=> ndp'.
+rewrite (@eqp_dvdr ((lead_coef (gcdp p q) ^+ scalp p (gcdp p q))*:p)).
+ by rewrite e; rewrite Gauss_dvdpl //; apply: (coprimep_dvdl (dvdp_gcdr _ _)).
+by rewrite eqp_sym eqp_scale // lc_expn_scalp_neq0.
+Qed.
+
+Lemma gdcopP q p : gdcop_spec q p (gdcop q p).
+Proof. by rewrite /gdcop; apply: gdcop_recP. Qed.
+
+Lemma coprimep_gdco p q : (q != 0)%B -> coprimep (gdcop p q) p.
+Proof. by move=> q_neq0; case: gdcopP=> d; rewrite (negPf q_neq0) orbF. Qed.
+
+Lemma size2_dvdp_gdco p q d : p != 0 -> size d = 2%N ->
+ (d %| (gdcop q p)) = (d %| p) && ~~(d %| q).
+Proof.
+case: (eqVneq d 0) => [-> | dn0]; first by rewrite size_poly0.
+move=> p0 sd; apply/idP/idP.
+ case: gdcopP=> r rp crq maxr dr; move/negPf: (p0)=> p0f.
+ rewrite (dvdp_trans dr) //=.
+ move: crq; apply: contraL=> dq; rewrite p0f orbF; apply/coprimepPn.
+ by move:p0; apply: contra=> r0; move: rp; rewrite (eqP r0) dvd0p.
+ by exists d; rewrite dvdp_gcd dr dq -size_poly_eq1 sd.
+case/andP=> dp dq; case: gdcopP=> r rp crq maxr; apply: maxr=> //.
+apply/coprimepP=> x xd xq.
+move: (dvdp_leq dn0 xd); rewrite leq_eqVlt sd; case/orP; last first.
+ rewrite ltnS leq_eqVlt; case/orP; first by rewrite -size_poly_eq1.
+ rewrite ltnS leqn0 size_poly_eq0; move/eqP=> x0; move: xd; rewrite x0 dvd0p.
+ by rewrite (negPf dn0).
+by rewrite -sd dvdp_size_eqp //; move/(eqp_dvdl q); rewrite xq (negPf dq).
+Qed.
+
+Lemma dvdp_gdco p q : (gdcop p q) %| q.
+Proof. by case: gdcopP. Qed.
+
+Lemma root_gdco p q x : p != 0 -> root (gdcop q p) x = root p x && ~~(root q x).
+Proof.
+move=> p0 /=; rewrite !root_factor_theorem.
+apply: size2_dvdp_gdco; rewrite ?p0 //.
+by rewrite size_addl size_polyX // size_opp size_polyC ltnS; case: (x != 0).
+Qed.
+
+Lemma dvdp_comp_poly r p q : (p %| q) -> (p \Po r) %| (q \Po r).
+Proof.
+case: (eqVneq p 0) => [-> | pn0].
+ by rewrite comp_poly0 !dvd0p; move/eqP->; rewrite comp_poly0.
+rewrite dvdp_eq; set c := _ ^+ _; set s := _ %/ _; move/eqP=> Hq.
+apply: (@eq_dvdp c (s \Po r)); first by rewrite expf_neq0 // lead_coef_eq0.
+by rewrite -comp_polyZ Hq comp_polyM.
+Qed.
+
+Lemma gcdp_comp_poly r p q : gcdp p q \Po r %= gcdp (p \Po r) (q \Po r).
+Proof.
+apply/andP; split.
+ by rewrite dvdp_gcd !dvdp_comp_poly ?dvdp_gcdl ?dvdp_gcdr.
+case: (Bezoutp p q) => [[u v]] /andP [].
+move/(dvdp_comp_poly r) => Huv _.
+rewrite (dvdp_trans _ Huv) // comp_polyD !comp_polyM.
+by rewrite dvdp_add // dvdp_mull // (dvdp_gcdl,dvdp_gcdr).
+Qed.
+
+Lemma coprimep_comp_poly r p q : coprimep p q -> coprimep (p \Po r) (q \Po r).
+Proof.
+rewrite -!gcdp_eqp1 -!size_poly_eq1 -!dvdp1; move/(dvdp_comp_poly r).
+rewrite comp_polyC => Hgcd.
+by apply: dvdp_trans Hgcd; case/andP: (gcdp_comp_poly r p q).
+Qed.
+
+Lemma coprimep_addl_mul p q r : coprimep r (p * r + q) = coprimep r q.
+Proof. by rewrite !coprimep_def (eqp_size (gcdp_addl_mul _ _ _)). Qed.
+
+Definition irreducible_poly p :=
+ (size p > 1) * (forall q, size q != 1%N -> q %| p -> q %= p) : Prop.
+
+Lemma irredp_neq0 p : irreducible_poly p -> p != 0.
+Proof. by rewrite -size_poly_eq0 -lt0n => [[/ltnW]]. Qed.
+
+Definition apply_irredp p (irr_p : irreducible_poly p) := irr_p.2.
+Coercion apply_irredp : irreducible_poly >-> Funclass.
+
+Lemma modp_XsubC p c : p %% ('X - c%:P) = p.[c]%:P.
+Proof.
+have: root (p - p.[c]%:P) c by rewrite /root !hornerE subrr.
+case/factor_theorem=> q /(canRL (subrK _)) Dp; rewrite modpE /= lead_coefXsubC.
+rewrite GRing.unitr1 expr1n invr1 scale1r {1}Dp.
+rewrite RingMonic.rmodp_addl_mul_small // ?monicXsubC // size_XsubC size_polyC.
+by case: (p.[c] == 0).
+Qed.
+
+Lemma coprimep_XsubC p c : coprimep p ('X - c%:P) = ~~ root p c.
+Proof.
+rewrite -coprimep_modl modp_XsubC /root -alg_polyC.
+have [-> | /coprimep_scalel->] := altP eqP; last exact: coprime1p.
+by rewrite scale0r /coprimep gcd0p size_XsubC.
+Qed.
+
+Lemma coprimepX p : coprimep p 'X = ~~ root p 0.
+Proof. by rewrite -['X]subr0 coprimep_XsubC. Qed.
+
+Lemma eqp_monic : {in monic &, forall p q, (p %= q) = (p == q)}.
+Proof.
+move=> p q monic_p monic_q; apply/idP/eqP=> [|-> //].
+case/eqpP=> [[a b] /= /andP[a_neq0 _] eq_pq].
+apply: (@mulfI _ a%:P); first by rewrite polyC_eq0.
+rewrite !mul_polyC eq_pq; congr (_ *: q); apply: (mulIf (oner_neq0 _)).
+by rewrite -{1}(monicP monic_q) -(monicP monic_p) -!lead_coefZ eq_pq.
+Qed.
+
+
+Lemma dvdp_mul_XsubC p q c :
+ (p %| ('X - c%:P) * q) = ((if root p c then p %/ ('X - c%:P) else p) %| q).
+Proof.
+case: ifPn => [| not_pc0]; last by rewrite Gauss_dvdpr ?coprimep_XsubC.
+rewrite root_factor_theorem -eqp_div_XsubC mulrC => /eqP{1}->.
+by rewrite dvdp_mul2l ?polyXsubC_eq0.
+Qed.
+
+Lemma dvdp_prod_XsubC (I : Type) (r : seq I) (F : I -> R) p :
+ p %| \prod_(i <- r) ('X - (F i)%:P) ->
+ {m | p %= \prod_(i <- mask m r) ('X - (F i)%:P)}.
+Proof.
+elim: r => [|i r IHr] in p *.
+ by rewrite big_nil dvdp1; exists nil; rewrite // big_nil -size_poly_eq1.
+rewrite big_cons dvdp_mul_XsubC root_factor_theorem -eqp_div_XsubC.
+case: eqP => [{2}-> | _] /IHr[m Dp]; last by exists (false :: m).
+by exists (true :: m); rewrite /= mulrC big_cons eqp_mul2l ?polyXsubC_eq0.
+Qed.
+
+Lemma irredp_XsubC (x : R) : irreducible_poly ('X - x%:P).
+Proof.
+split=> [|d size_d d_dv_Xx]; first by rewrite size_XsubC.
+have: ~ d %= 1 by apply/negP; rewrite -size_poly_eq1.
+have [|m /=] := @dvdp_prod_XsubC _ [:: x] id d; first by rewrite big_seq1.
+by case: m => [|[] [|_ _] /=]; rewrite (big_nil, big_seq1).
+Qed.
+
+Lemma irredp_XsubCP d p :
+ irreducible_poly p -> d %| p -> {d %= 1} + {d %= p}.
+Proof.
+move=> irred_p dvd_dp; have [] := boolP (_ %= 1); first by left.
+by rewrite -size_poly_eq1=> /irred_p /(_ dvd_dp); right.
+Qed.
+
+End IDomainPseudoDivision.
+
+Hint Resolve eqpxx divp0 divp1 mod0p modp0 modp1 dvdp_mull dvdp_mulr dvdpp.
+Hint Resolve dvdp0.
+
+End CommonIdomain.
+
+Module Idomain.
+
+Include IdomainDefs.
+Export IdomainDefs.
+Include WeakIdomain.
+Include CommonIdomain.
+
+End Idomain.
+
+Module IdomainMonic.
+
+Import Ring ComRing UnitRing IdomainDefs Idomain.
+
+Section MonicDivisor.
+
+Variable R : idomainType.
+Variable q : {poly R}.
+Hypothesis monq : q \is monic.
+
+Implicit Type p d r : {poly R}.
+
+Lemma divpE p : p %/ q = rdivp p q.
+Proof. by rewrite divpE (eqP monq) unitr1 expr1n invr1 scale1r. Qed.
+
+Lemma modpE p : p %% q = rmodp p q.
+Proof. by rewrite modpE (eqP monq) unitr1 expr1n invr1 scale1r. Qed.
+
+Lemma scalpE p : scalp p q = 0%N.
+Proof. by rewrite scalpE (eqP monq) unitr1. Qed.
+
+Lemma divp_eq p : p = (p %/ q) * q + (p %% q).
+Proof. by rewrite -divp_eq (eqP monq) expr1n scale1r. Qed.
+
+Lemma divpp p : q %/ q = 1.
+Proof. by rewrite divpp ?monic_neq0 // (eqP monq) expr1n. Qed.
+
+Lemma dvdp_eq p : (q %| p) = (p == (p %/ q) * q).
+Proof. by rewrite dvdp_eq (eqP monq) expr1n scale1r. Qed.
+
+Lemma dvdpP p : reflect (exists qq, p = qq * q) (q %| p).
+Proof.
+apply: (iffP idP); first by rewrite dvdp_eq; move/eqP=> e; exists (p %/ q).
+by case=> qq ->; rewrite dvdp_mull // dvdpp.
+Qed.
+
+Lemma mulpK p : p * q %/ q = p.
+Proof. by rewrite mulpK ?monic_neq0 // (eqP monq) expr1n scale1r. Qed.
+
+Lemma mulKp p : q * p %/ q = p.
+Proof. by rewrite mulrC; exact: mulpK. Qed.
+
+End MonicDivisor.
+
+End IdomainMonic.
+
+Module IdomainUnit.
+
+Import Ring ComRing UnitRing IdomainDefs Idomain.
+
+Section UnitDivisor.
+
+Variable R : idomainType.
+Variable d : {poly R}.
+
+Hypothesis ulcd : lead_coef d \in GRing.unit.
+
+Implicit Type p q r : {poly R}.
+
+Lemma divp_eq p : p = (p %/ d) * d + (p %% d).
+Proof. by have := (divp_eq p d); rewrite scalpE ulcd expr0 scale1r. Qed.
+
+Lemma edivpP p q r : p = q * d + r -> size r < size d ->
+ q = (p %/ d) /\ r = p %% d.
+Proof.
+move=> ep srd; have := (divp_eq p); rewrite {1}ep.
+move/eqP; rewrite -subr_eq -addrA addrC eq_sym -subr_eq -mulrBl; move/eqP.
+have lcdn0 : lead_coef d != 0 by apply: contraTneq ulcd => ->; rewrite unitr0.
+case abs: (p %/ d - q == 0).
+ move: abs; rewrite subr_eq0; move/eqP->; rewrite subrr mul0r; move/eqP.
+ by rewrite eq_sym subr_eq0; move/eqP->.
+have hleq : size d <= size ((p %/ d - q) * d).
+ rewrite size_proper_mul; last first.
+ by rewrite mulf_eq0 (negPf lcdn0) orbF lead_coef_eq0 abs.
+ move: abs; rewrite -size_poly_eq0; move/negbT; rewrite -lt0n; move/prednK<-.
+ by rewrite addSn /= leq_addl.
+have hlt : size (r - p %% d) < size d.
+ apply: leq_ltn_trans (size_add _ _) _; rewrite size_opp.
+ by rewrite gtn_max srd ltn_modp /= -lead_coef_eq0.
+by move=> e; have:= (leq_trans hlt hleq); rewrite e ltnn.
+Qed.
+
+Lemma divpP p q r : p = q * d + r -> size r < size d ->
+ q = (p %/ d).
+Proof. by move/edivpP=> h; case/h. Qed.
+
+Lemma modpP p q r : p = q * d + r -> size r < size d -> r = (p %% d).
+Proof. by move/edivpP=> h; case/h. Qed.
+
+Lemma ulc_eqpP p q : lead_coef q \is a GRing.unit ->
+ reflect (exists2 c : R, c != 0 & p = c *: q) (p %= q).
+Proof.
+ case: (altP (lead_coef q =P 0)) => [->|]; first by rewrite unitr0.
+ rewrite lead_coef_eq0 => nz_q ulcq; apply: (iffP idP).
+ case: (altP (p =P 0)) => [->|nz_p].
+ by rewrite eqp_sym eqp0 (negbTE nz_q).
+ move/eqp_eq=> eq; exists (lead_coef p / lead_coef q).
+ by rewrite mulf_neq0 // ?invr_eq0 lead_coef_eq0.
+ by apply/(scaler_injl ulcq); rewrite scalerA mulrCA divrr // mulr1.
+ by case=> c nz_c ->; apply/eqpP; exists (1, c); rewrite ?scale1r ?oner_eq0.
+Qed.
+
+Lemma dvdp_eq p : (d %| p) = (p == p %/ d * d).
+Proof.
+apply/eqP/eqP=> [modp0 | ->]; last exact: modp_mull.
+by rewrite {1}(divp_eq p) modp0 addr0.
+Qed.
+
+Lemma ucl_eqp_eq p q : lead_coef q \is a GRing.unit ->
+ p %= q -> p = (lead_coef p / lead_coef q) *: q.
+Proof.
+move=> ulcq /eqp_eq; move/(congr1 ( *:%R (lead_coef q)^-1 )).
+by rewrite !scalerA mulrC divrr // scale1r mulrC.
+Qed.
+
+Lemma modp_scalel c p : (c *: p) %% d = c *: (p %% d).
+Proof.
+case: (altP (c =P 0)) => [-> | cn0]; first by rewrite !scale0r mod0p.
+have e : (c *: p) = (c *: (p %/ d)) * d + c *: (p %% d).
+ by rewrite -scalerAl -scalerDr -divp_eq.
+have s: size (c *: (p %% d)) < size d.
+ rewrite -mul_polyC; apply: leq_ltn_trans (size_mul_leq _ _) _.
+ rewrite size_polyC cn0 addSn add0n /= ltn_modp.
+ by rewrite -lead_coef_eq0; apply: contraTneq ulcd => ->; rewrite unitr0.
+by case: (edivpP e s) => _ ->.
+Qed.
+
+Lemma divp_scalel c p : (c *: p) %/ d = c *: (p %/ d).
+Proof.
+case: (altP (c =P 0)) => [-> | cn0]; first by rewrite !scale0r div0p.
+have e : (c *: p) = (c *: (p %/ d)) * d + c *: (p %% d).
+ by rewrite -scalerAl -scalerDr -divp_eq.
+have s: size (c *: (p %% d)) < size d.
+ rewrite -mul_polyC; apply: leq_ltn_trans (size_mul_leq _ _) _.
+ rewrite size_polyC cn0 addSn add0n /= ltn_modp.
+ by rewrite -lead_coef_eq0; apply: contraTneq ulcd => ->; rewrite unitr0.
+by case: (edivpP e s) => ->.
+Qed.
+
+Lemma eqp_modpl p q : p %= q -> (p %% d) %= (q %% d).
+Proof.
+case/eqpP=> [[c1 c2]] /andP /= [c1n0 c2n0 e].
+by apply/eqpP; exists (c1, c2); rewrite ?c1n0 //= -!modp_scalel e.
+Qed.
+
+Lemma eqp_divl p q : p %= q -> (p %/ d) %= (q %/ d).
+Proof.
+case/eqpP=> [[c1 c2]] /andP /= [c1n0 c2n0 e].
+by apply/eqpP; exists (c1, c2); rewrite ?c1n0 // -!divp_scalel e.
+Qed.
+
+Lemma modp_opp p : (- p) %% d = - (p %% d).
+Proof.
+by rewrite -mulN1r -[- (_ %% _)]mulN1r -polyC_opp !mul_polyC modp_scalel.
+Qed.
+
+Lemma divp_opp p : (- p) %/ d = - (p %/ d).
+Proof.
+by rewrite -mulN1r -[- (_ %/ _)]mulN1r -polyC_opp !mul_polyC divp_scalel.
+Qed.
+
+Lemma modp_add p q : (p + q) %% d = p %% d + q %% d.
+Proof.
+have hs : size (p %% d + q %% d) < size d.
+ apply: leq_ltn_trans (size_add _ _) _.
+ rewrite gtn_max !ltn_modp andbb -lead_coef_eq0.
+ by apply: contraTneq ulcd => ->; rewrite unitr0.
+have he : (p + q) = (p %/ d + q %/ d) * d + (p %% d + q %% d).
+ rewrite {1}(divp_eq p) {1}(divp_eq q) addrAC addrA -mulrDl.
+ by rewrite [_ %% _ + _]addrC addrA.
+by case: (edivpP he hs).
+Qed.
+
+Lemma divp_add p q : (p + q) %/ d = p %/ d + q %/ d.
+Proof.
+have hs : size (p %% d + q %% d) < size d.
+ apply: leq_ltn_trans (size_add _ _) _.
+ rewrite gtn_max !ltn_modp andbb -lead_coef_eq0.
+ by apply: contraTneq ulcd => ->; rewrite unitr0.
+have he : (p + q) = (p %/ d + q %/ d) * d + (p %% d + q %% d).
+ rewrite {1}(divp_eq p) {1}(divp_eq q) addrAC addrA -mulrDl.
+ by rewrite [_ %% _ + _]addrC addrA.
+by case: (edivpP he hs).
+Qed.
+
+Lemma mulpK q : (q * d) %/ d = q.
+Proof.
+case/edivpP: (sym_eq (addr0 (q * d))); rewrite // size_poly0 size_poly_gt0.
+by rewrite -lead_coef_eq0; apply: contraTneq ulcd => ->; rewrite unitr0.
+Qed.
+
+Lemma mulKp q : (d * q) %/ d = q.
+Proof. rewrite mulrC; exact: mulpK. Qed.
+
+Lemma divp_addl_mul_small q r :
+ size r < size d -> (q * d + r) %/ d = q.
+Proof. by move=> srd; rewrite divp_add (divp_small srd) addr0 mulpK. Qed.
+
+Lemma modp_addl_mul_small q r :
+ size r < size d -> (q * d + r) %% d = r.
+Proof. by move=> srd; rewrite modp_add modp_mull add0r modp_small. Qed.
+
+Lemma divp_addl_mul q r : (q * d + r) %/ d = q + r %/ d.
+Proof. by rewrite divp_add mulpK. Qed.
+
+Lemma divpp : d %/ d = 1.
+Proof. by rewrite -{1}(mul1r d) mulpK. Qed.
+
+Lemma leq_trunc_divp m : size (m %/ d * d) <= size m.
+Proof.
+have dn0 : d != 0.
+ by rewrite -lead_coef_eq0; apply: contraTneq ulcd => ->; rewrite unitr0.
+case q0 : (m %/ d == 0); first by rewrite (eqP q0) mul0r size_poly0 leq0n.
+rewrite {2}(divp_eq m) size_addl // size_mul ?q0 //; move/negbT: q0.
+rewrite -size_poly_gt0; move/prednK<-; rewrite addSn /=.
+by move: dn0; rewrite -(ltn_modp m); move/ltn_addl->.
+Qed.
+
+Lemma dvdpP p : reflect (exists q, p = q * d) (d %| p).
+Proof.
+apply: (iffP idP) => [| [k ->]]; last by apply/eqP; rewrite modp_mull.
+by rewrite dvdp_eq; move/eqP->; exists (p %/ d).
+Qed.
+
+Lemma divpK p : d %| p -> p %/ d * d = p.
+Proof. by rewrite dvdp_eq; move/eqP. Qed.
+
+Lemma divpKC p : d %| p -> d * (p %/ d) = p.
+Proof. by move=> ?; rewrite mulrC divpK. Qed.
+
+Lemma dvdp_eq_div p q : d %| p -> (q == p %/ d) = (q * d == p).
+Proof.
+move/divpK=> {2}<-; apply/eqP/eqP; first by move->.
+suff dn0 : d != 0 by move/(mulIf dn0).
+by rewrite -lead_coef_eq0; apply: contraTneq ulcd => ->; rewrite unitr0.
+Qed.
+
+Lemma dvdp_eq_mul p q : d %| p -> (p == q * d) = (p %/ d == q).
+Proof. by move=>dv_d_p; rewrite eq_sym -dvdp_eq_div // eq_sym. Qed.
+
+Lemma divp_mulA p q : d %| q -> p * (q %/ d) = p * q %/ d.
+Proof.
+move=> hdm; apply/eqP; rewrite eq_sym -dvdp_eq_mul.
+ by rewrite -mulrA divpK.
+by move/divpK: hdm<-; rewrite mulrA dvdp_mull // dvdpp.
+Qed.
+
+Lemma divp_mulAC m n : d %| m -> m %/ d * n = m * n %/ d.
+Proof. by move=> hdm; rewrite mulrC (mulrC m); exact: divp_mulA. Qed.
+
+Lemma divp_mulCA p q : d %| p -> d %| q -> p * (q %/ d) = q * (p %/ d).
+Proof. by move=> hdp hdq; rewrite mulrC divp_mulAC // divp_mulA. Qed.
+
+Lemma modp_mul p q : (p * (q %% d)) %% d = (p * q) %% d.
+Proof.
+have -> : q %% d = q - q %/ d * d by rewrite {2}(divp_eq q) -addrA addrC subrK.
+rewrite mulrDr modp_add // -mulNr mulrA -{2}[_ %% _]addr0; congr (_ + _).
+by apply/eqP; apply: dvdp_mull; exact: dvdpp.
+Qed.
+
+End UnitDivisor.
+
+Section MoreUnitDivisor.
+
+Variable R : idomainType.
+Variable d : {poly R}.
+Hypothesis ulcd : lead_coef d \in GRing.unit.
+
+Implicit Types p q : {poly R}.
+
+Lemma expp_sub m n : n <= m -> (d ^+ (m - n))%N = d ^+ m %/ d ^+ n.
+Proof.
+by move/subnK=> {2}<-; rewrite exprD mulpK // lead_coef_exp unitrX.
+Qed.
+
+Lemma divp_pmul2l p q : lead_coef q \in GRing.unit -> d * p %/ (d * q) = p %/ q.
+Proof.
+move=> uq.
+have udq: lead_coef (d * q) \in GRing.unit.
+ by rewrite lead_coefM unitrM_comm ?ulcd //; red; rewrite mulrC.
+rewrite {1}(divp_eq uq p) mulrDr mulrCA divp_addl_mul //.
+have dn0 : d != 0.
+ by rewrite -lead_coef_eq0; apply: contraTneq ulcd => ->; rewrite unitr0.
+have qn0 : q != 0.
+ by rewrite -lead_coef_eq0; apply: contraTneq uq => ->; rewrite unitr0.
+have dqn0 : d * q != 0 by rewrite mulf_eq0 negb_or dn0.
+suff : size (d * (p %% q)) < size (d * q).
+ by rewrite ltnNge -divpN0 // negbK => /eqP ->; rewrite addr0.
+case: (altP ( (p %% q) =P 0)) => [-> | rn0].
+ by rewrite mulr0 size_poly0 size_poly_gt0.
+rewrite !size_mul //; move: dn0; rewrite -size_poly_gt0.
+by move/prednK<-; rewrite !addSn /= ltn_add2l ltn_modp.
+Qed.
+
+Lemma divp_pmul2r p q :
+ lead_coef p \in GRing.unit -> q * d %/ (p * d) = q %/ p.
+Proof. by move=> uq; rewrite -!(mulrC d) divp_pmul2l. Qed.
+
+Lemma divp_divl r p q :
+ lead_coef r \in GRing.unit -> lead_coef p \in GRing.unit ->
+ q %/ p %/ r = q %/ (p * r).
+Proof.
+move=> ulcr ulcp.
+have e : q = (q %/ p %/ r) * (p * r) + ((q %/ p) %% r * p + q %% p).
+ rewrite addrA (mulrC p) mulrA -mulrDl; rewrite -divp_eq //; exact: divp_eq.
+have pn0 : p != 0.
+ by rewrite -lead_coef_eq0; apply: contraTneq ulcp => ->; rewrite unitr0.
+have rn0 : r != 0.
+ by rewrite -lead_coef_eq0; apply: contraTneq ulcr => ->; rewrite unitr0.
+have s : size ((q %/ p) %% r * p + q %% p) < size (p * r).
+ case: (altP ((q %/ p) %% r =P 0)) => [-> | qn0].
+ rewrite mul0r add0r size_mul // (polySpred rn0) addnS /=.
+ by apply: leq_trans (leq_addr _ _); rewrite ltn_modp.
+ rewrite size_addl mulrC.
+ by rewrite !size_mul // (polySpred pn0) !addSn /= ltn_add2l ltn_modp.
+ rewrite size_mul // (polySpred qn0) addnS /=.
+ by apply: leq_trans (leq_addr _ _); rewrite ltn_modp.
+case: (edivpP _ e s) => //; rewrite lead_coefM unitrM_comm ?ulcp //.
+by red; rewrite mulrC.
+Qed.
+
+Lemma divpAC p q : lead_coef p \in GRing.unit -> q %/ d %/ p = q %/ p %/ d.
+Proof. by move=> ulcp; rewrite !divp_divl // mulrC. Qed.
+
+Lemma modp_scaler c p : c \in GRing.unit -> p %% (c *: d) = (p %% d).
+Proof.
+move=> cn0; case: (eqVneq d 0) => [-> | dn0]; first by rewrite scaler0 !modp0.
+have e : p = (c^-1 *: (p %/ d)) * (c *: d) + (p %% d).
+ by rewrite scalerCA scalerA mulVr // scale1r -(divp_eq ulcd).
+suff s : size (p %% d) < size (c *: d).
+ by rewrite (modpP _ e s) // -mul_polyC lead_coefM lead_coefC unitrM cn0.
+by rewrite size_scale ?ltn_modp //; apply: contraTneq cn0 => ->; rewrite unitr0.
+Qed.
+
+Lemma divp_scaler c p : c \in GRing.unit -> p %/ (c *: d) = c^-1 *: (p %/ d).
+Proof.
+move=> cn0; case: (eqVneq d 0) => [-> | dn0].
+ by rewrite scaler0 !divp0 scaler0.
+have e : p = (c^-1 *: (p %/ d)) * (c *: d) + (p %% d).
+ by rewrite scalerCA scalerA mulVr // scale1r -(divp_eq ulcd).
+suff s : size (p %% d) < size (c *: d).
+ by rewrite (divpP _ e s) // -mul_polyC lead_coefM lead_coefC unitrM cn0.
+by rewrite size_scale ?ltn_modp //; apply: contraTneq cn0 => ->; rewrite unitr0.
+Qed.
+
+End MoreUnitDivisor.
+
+End IdomainUnit.
+
+Module Field.
+
+Import Ring ComRing UnitRing.
+Include IdomainDefs.
+Export IdomainDefs.
+Include CommonIdomain.
+
+Section FieldDivision.
+
+Variable F : fieldType.
+
+Implicit Type p q r d : {poly F}.
+
+Lemma divp_eq p q : p = (p %/ q) * q + (p %% q).
+Proof.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite modp0 mulr0 add0r.
+by apply: IdomainUnit.divp_eq; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma divp_modpP p q d r : p = q * d + r -> size r < size d ->
+ q = (p %/ d) /\ r = p %% d.
+Proof.
+move=> he hs; apply: IdomainUnit.edivpP => //; rewrite unitfE lead_coef_eq0.
+by rewrite -size_poly_gt0; apply: leq_trans hs.
+Qed.
+
+Lemma divpP p q d r : p = q * d + r -> size r < size d ->
+ q = (p %/ d).
+Proof. by move/divp_modpP=> h; case/h. Qed.
+
+Lemma modpP p q d r : p = q * d + r -> size r < size d -> r = (p %% d).
+Proof. by move/divp_modpP=> h; case/h. Qed.
+
+Lemma eqpfP p q : p %= q -> p = (lead_coef p / lead_coef q) *: q.
+Proof.
+have [->|nz_q] := altP (q =P 0).
+ by rewrite eqp0 => /eqP ->; rewrite scaler0.
+move/IdomainUnit.ucl_eqp_eq; apply; rewrite unitfE.
+by move: nz_q; rewrite -lead_coef_eq0 => nz_qT.
+Qed.
+
+Lemma dvdp_eq q p : (q %| p) = (p == p %/ q * q).
+Proof.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite dvd0p mulr0 eq_sym.
+by apply: IdomainUnit.dvdp_eq; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma eqpf_eq p q : reflect (exists2 c, c != 0 & p = c *: q) (p %= q).
+Proof.
+apply: (iffP idP); last first.
+ case=> c nz_c ->; apply/eqpP.
+ by exists (1, c); rewrite ?scale1r ?oner_eq0.
+have [->|nz_q] := altP (q =P 0).
+ by rewrite eqp0=> /eqP ->; exists 1; rewrite ?scale1r ?oner_eq0.
+case/IdomainUnit.ulc_eqpP; first by rewrite unitfE lead_coef_eq0.
+by move=> c nz_c ->; exists c.
+Qed.
+
+Lemma modp_scalel c p q : (c *: p) %% q = c *: (p %% q).
+Proof.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite !modp0.
+by apply: IdomainUnit.modp_scalel; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma mulpK p q : q != 0 -> p * q %/ q = p.
+Proof. by move=> qn0; rewrite IdomainUnit.mulpK // unitfE lead_coef_eq0. Qed.
+
+Lemma mulKp p q : q != 0 -> q * p %/ q = p.
+Proof. by rewrite mulrC; exact: mulpK. Qed.
+
+Lemma divp_scalel c p q : (c *: p) %/ q = c *: (p %/ q).
+Proof.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite !divp0 scaler0.
+by apply: IdomainUnit.divp_scalel; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma modp_scaler c p d : c != 0 -> p %% (c *: d) = (p %% d).
+Proof.
+move=> cn0; case: (eqVneq d 0) => [-> | dn0]; first by rewrite scaler0 !modp0.
+have e : p = (c^-1 *: (p %/ d)) * (c *: d) + (p %% d).
+ by rewrite scalerCA scalerA mulVf // scale1r -divp_eq.
+suff s : size (p %% d) < size (c *: d) by rewrite (modpP e s).
+by rewrite size_scale ?ltn_modp.
+Qed.
+
+Lemma divp_scaler c p d : c != 0 -> p %/ (c *: d) = c^-1 *: (p %/ d).
+Proof.
+move=> cn0; case: (eqVneq d 0) => [-> | dn0].
+ by rewrite scaler0 !divp0 scaler0.
+have e : p = (c^-1 *: (p %/ d)) * (c *: d) + (p %% d).
+ by rewrite scalerCA scalerA mulVf // scale1r -divp_eq.
+suff s : size (p %% d) < size (c *: d) by rewrite (divpP e s).
+by rewrite size_scale ?ltn_modp.
+Qed.
+
+Lemma eqp_modpl d p q : p %= q -> (p %% d) %= (q %% d).
+Proof.
+case/eqpP=> [[c1 c2]] /andP /= [c1n0 c2n0 e].
+by apply/eqpP; exists (c1, c2); rewrite ?c1n0 // -!modp_scalel e.
+Qed.
+
+Lemma eqp_divl d p q : p %= q -> (p %/ d) %= (q %/ d).
+Proof.
+case/eqpP=> [[c1 c2]] /andP /= [c1n0 c2n0 e].
+by apply/eqpP; exists (c1, c2); rewrite ?c1n0 // -!divp_scalel e.
+Qed.
+
+Lemma eqp_modpr d p q : p %= q -> (d %% p) %= (d %% q).
+Proof.
+case/eqpP=> [[c1 c2]] /andP [c1n0 c2n0 e].
+have -> : p = (c1^-1 * c2) *: q by rewrite -scalerA -e scalerA mulVf // scale1r.
+by rewrite modp_scaler ?eqpxx // mulf_eq0 negb_or invr_eq0 c1n0.
+Qed.
+
+Lemma eqp_mod p1 p2 q1 q2 : p1 %= p2 -> q1 %= q2 -> p1 %% q1 %= p2 %% q2.
+Proof.
+move=> e1 e2; apply: eqp_trans (eqp_modpr _ e2).
+apply: eqp_trans (eqp_modpl _ e1); exact: eqpxx.
+Qed.
+
+Lemma eqp_divr (d m n : {poly F}) : m %= n -> (d %/ m) %= (d %/ n).
+Proof.
+case/eqpP=> [[c1 c2]] /andP [c1n0 c2n0 e].
+have -> : m = (c1^-1 * c2) *: n by rewrite -scalerA -e scalerA mulVf // scale1r.
+by rewrite divp_scaler ?eqp_scale // ?invr_eq0 mulf_eq0 negb_or invr_eq0 c1n0.
+Qed.
+
+Lemma eqp_div p1 p2 q1 q2 : p1 %= p2 -> q1 %= q2 -> p1 %/ q1 %= p2 %/ q2.
+Proof.
+move=> e1 e2; apply: eqp_trans (eqp_divr _ e2).
+apply: eqp_trans (eqp_divl _ e1); exact: eqpxx.
+Qed.
+
+Lemma eqp_gdcor p q r : q %= r -> gdcop p q %= gdcop p r.
+Proof.
+move=> eqr; rewrite /gdcop (eqp_size eqr).
+move: (size r)=> n; elim: n p q r eqr => [|n ihn] p q r; first by rewrite eqpxx.
+move=> eqr /=; rewrite (eqp_coprimepl p eqr); case: ifP => _ //; apply: ihn.
+apply: eqp_div => //; exact: eqp_gcdl.
+Qed.
+
+Lemma eqp_gdcol p q r : q %= r -> gdcop q p %= gdcop r p.
+Proof.
+move=> eqr; rewrite /gdcop; move: (size p)=> n.
+elim: n p q r eqr {1 3}p (eqpxx p) => [|n ihn] p q r eqr s esp /=.
+ move: eqr; case: (eqVneq q 0)=> [-> | nq0 eqr] /=.
+ by rewrite eqp_sym eqp0; move->; rewrite eqxx eqpxx.
+ suff rn0 : r != 0 by rewrite (negPf nq0) (negPf rn0) eqpxx.
+ by apply: contraTneq eqr => ->; rewrite eqp0.
+rewrite (eqp_coprimepr _ eqr) (eqp_coprimepl _ esp); case: ifP=> _ //.
+apply: ihn => //; apply: eqp_div => //; exact: eqp_gcd.
+Qed.
+
+Lemma eqp_rgdco_gdco q p : rgdcop q p %= gdcop q p.
+Proof.
+rewrite /rgdcop /gdcop; move: (size p)=> n.
+elim: n p q {1 3}p {1 3}q (eqpxx p) (eqpxx q) => [|n ihn] p q s t /= sp tq.
+ move: tq; case: (eqVneq t 0)=> [-> | nt0 etq].
+ by rewrite eqp_sym eqp0; move->; rewrite eqxx eqpxx.
+ suff qn0 : q != 0 by rewrite (negPf nt0) (negPf qn0) eqpxx.
+ by apply: contraTneq etq => ->; rewrite eqp0.
+rewrite rcoprimep_coprimep (eqp_coprimepl t sp) (eqp_coprimepr p tq).
+case: ifP=> // _; apply: ihn => //; apply: eqp_trans (eqp_rdiv_div _ _) _.
+by apply: eqp_div => //; apply: eqp_trans (eqp_rgcd_gcd _ _) _; apply: eqp_gcd.
+Qed.
+
+Lemma modp_opp p q : (- p) %% q = - (p %% q).
+Proof.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite !modp0.
+by apply: IdomainUnit.modp_opp; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma divp_opp p q : (- p) %/ q = - (p %/ q).
+Proof.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite !divp0 oppr0.
+by apply: IdomainUnit.divp_opp; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma modp_add d p q : (p + q) %% d = p %% d + q %% d.
+Proof.
+case: (eqVneq d 0) => [-> | dn0]; first by rewrite !modp0.
+by apply: IdomainUnit.modp_add; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma modNp p q : (- p) %% q = - (p %% q).
+Proof. by apply/eqP; rewrite -addr_eq0 -modp_add addNr mod0p. Qed.
+
+Lemma divp_add d p q : (p + q) %/ d = p %/ d + q %/ d.
+Proof.
+case: (eqVneq d 0) => [-> | dn0]; first by rewrite !divp0 addr0.
+by apply: IdomainUnit.divp_add; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma divp_addl_mul_small d q r :
+ size r < size d -> (q * d + r) %/ d = q.
+Proof.
+move=> srd; rewrite divp_add (divp_small srd) addr0 mulpK //.
+by rewrite -size_poly_gt0; apply: leq_trans srd.
+Qed.
+
+Lemma modp_addl_mul_small d q r :
+ size r < size d -> (q * d + r) %% d = r.
+Proof. by move=> srd; rewrite modp_add modp_mull add0r modp_small. Qed.
+
+Lemma divp_addl_mul d q r : d != 0 -> (q * d + r) %/ d = q + r %/ d.
+Proof. by move=> dn0; rewrite divp_add mulpK. Qed.
+
+Lemma divpp d : d != 0 -> d %/ d = 1.
+Proof.
+by move=> dn0; apply: IdomainUnit.divpp; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma leq_trunc_divp d m : size (m %/ d * d) <= size m.
+Proof.
+case: (eqVneq d 0) => [-> | dn0]; first by rewrite mulr0 size_poly0.
+by apply: IdomainUnit.leq_trunc_divp; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma divpK d p : d %| p -> p %/ d * d = p.
+Proof.
+case: (eqVneq d 0) => [-> | dn0]; first by move/dvd0pP->; rewrite mulr0.
+by apply: IdomainUnit.divpK; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma divpKC d p : d %| p -> d * (p %/ d) = p.
+Proof. by move=> ?; rewrite mulrC divpK. Qed.
+
+Lemma dvdp_eq_div d p q : d != 0 -> d %| p -> (q == p %/ d) = (q * d == p).
+Proof.
+by move=> dn0; apply: IdomainUnit.dvdp_eq_div; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma dvdp_eq_mul d p q : d != 0 -> d %| p -> (p == q * d) = (p %/ d == q).
+Proof. by move=> dn0 dv_d_p; rewrite eq_sym -dvdp_eq_div // eq_sym. Qed.
+
+Lemma divp_mulA d p q : d %| q -> p * (q %/ d) = p * q %/ d.
+Proof.
+case: (eqVneq d 0) => [-> | dn0]; first by move/dvd0pP->; rewrite !divp0 mulr0.
+by apply: IdomainUnit.divp_mulA; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma divp_mulAC d m n : d %| m -> m %/ d * n = m * n %/ d.
+Proof. by move=> hdm; rewrite mulrC (mulrC m); exact: divp_mulA. Qed.
+
+Lemma divp_mulCA d p q : d %| p -> d %| q -> p * (q %/ d) = q * (p %/ d).
+Proof. by move=> hdp hdq; rewrite mulrC divp_mulAC // divp_mulA. Qed.
+
+Lemma expp_sub d m n : d != 0 -> m >= n -> (d ^+ (m - n))%N = d ^+ m %/ d ^+ n.
+Proof. by move=> dn0 /subnK=> {2}<-; rewrite exprD mulpK // expf_neq0. Qed.
+
+Lemma divp_pmul2l d q p : d != 0 -> q != 0 -> d * p %/ (d * q) = p %/ q.
+Proof.
+by move=> dn0 qn0; apply: IdomainUnit.divp_pmul2l; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma divp_pmul2r d p q : d != 0 -> p != 0 -> q * d %/ (p * d) = q %/ p.
+Proof. by move=> dn0 qn0; rewrite -!(mulrC d) divp_pmul2l. Qed.
+
+Lemma divp_divl r p q : q %/ p %/ r = q %/ (p * r).
+Proof.
+case: (eqVneq r 0) => [-> | rn0]; first by rewrite mulr0 !divp0.
+case: (eqVneq p 0) => [-> | pn0]; first by rewrite mul0r !divp0 div0p.
+by apply: IdomainUnit.divp_divl; rewrite unitfE lead_coef_eq0.
+Qed.
+
+Lemma divpAC d p q : q %/ d %/ p = q %/ p %/ d.
+Proof. by rewrite !divp_divl // mulrC. Qed.
+
+Lemma edivp_def p q : edivp p q = (0%N, p %/ q, p %% q).
+Proof.
+rewrite Idomain.edivp_def; congr (_, _, _); rewrite /scalp 2!unlock /=.
+case (eqVneq q 0) => [-> | qn0]; first by rewrite eqxx lead_coef0 unitr0.
+rewrite (negPf qn0) /= unitfE lead_coef_eq0 qn0 /=.
+by case: (redivp_rec _ _ _ _) => [[]].
+Qed.
+
+Lemma divpE p q : p %/ q = (lead_coef q)^-(rscalp p q) *: (rdivp p q).
+Proof.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite rdivp0 divp0 scaler0.
+by rewrite Idomain.divpE unitfE lead_coef_eq0 qn0.
+Qed.
+
+Lemma modpE p q : p %% q = (lead_coef q)^-(rscalp p q) *: (rmodp p q).
+Proof.
+case: (eqVneq q 0) => [-> | qn0].
+ by rewrite rmodp0 modp0 /rscalp unlock eqxx lead_coef0 expr0 invr1 scale1r.
+by rewrite Idomain.modpE unitfE lead_coef_eq0 qn0.
+Qed.
+
+Lemma scalpE p q : scalp p q = 0%N.
+Proof.
+case: (eqVneq q 0) => [-> | qn0]; first by rewrite scalp0.
+by rewrite Idomain.scalpE unitfE lead_coef_eq0 qn0.
+Qed.
+
+(* Just to have it without importing the weak theory *)
+Lemma dvdpE p q : p %| q = rdvdp p q. Proof. exact: Idomain.dvdpE. Qed.
+
+CoInductive edivp_spec m d : nat * {poly F} * {poly F} -> Type :=
+ EdivpSpec n q r of
+ m = q * d + r & (d != 0) ==> (size r < size d) : edivp_spec m d (n, q, r).
+
+Lemma edivpP m d : edivp_spec m d (edivp m d).
+Proof.
+rewrite edivp_def; constructor; first exact: divp_eq.
+by apply/implyP=> dn0; rewrite ltn_modp.
+Qed.
+
+Lemma edivp_eq d q r : size r < size d -> edivp (q * d + r) d = (0%N, q, r).
+Proof.
+move=> srd; apply: Idomain.edivp_eq ; rewrite // unitfE lead_coef_eq0.
+rewrite -size_poly_gt0; exact: leq_trans srd.
+Qed.
+
+Lemma modp_mul p q m : (p * (q %% m)) %% m = (p * q) %% m.
+Proof.
+have ->: q %% m = q - q %/ m * m by rewrite {2}(divp_eq q m) -addrA addrC subrK.
+rewrite mulrDr modp_add // -mulNr mulrA -{2}[_ %% _]addr0; congr (_ + _).
+by apply/eqP; apply: dvdp_mull; exact: dvdpp.
+Qed.
+
+Lemma dvdpP p q : reflect (exists qq, p = qq * q) (q %| p).
+Proof.
+case: (eqVneq q 0)=> [-> | qn0]; last first.
+ by apply: IdomainUnit.dvdpP; rewrite unitfE lead_coef_eq0.
+rewrite dvd0p.
+by apply: (iffP idP) => [/eqP->| [? ->]]; [exists 1|]; rewrite mulr0.
+Qed.
+
+Lemma Bezout_eq1_coprimepP : forall p q,
+ reflect (exists u, u.1 * p + u.2 * q = 1) (coprimep p q).
+Proof.
+move=> p q; apply:(iffP idP)=> [hpq|]; last first.
+ by case=>[[u v]] /= e; apply/Bezout_coprimepP; exists (u, v); rewrite e eqpxx.
+case/Bezout_coprimepP: hpq => [[u v]] /=.
+case/eqpP=> [[c1 c2]] /andP /= [c1n0 c2n0] e.
+exists (c2^-1 *: (c1 *: u), c2^-1 *: (c1 *: v)); rewrite /= -!scalerAl.
+by rewrite -!scalerDr e scalerA mulVf // scale1r.
+Qed.
+
+Lemma dvdp_gdcor p q : q != 0 -> p %| (gdcop q p) * (q ^+ size p).
+Proof.
+move=> q_neq0; rewrite /gdcop.
+elim: (size p) {-2 5}p (leqnn (size p))=> {p} [|n ihn] p.
+ rewrite size_poly_leq0; move/eqP->.
+ by rewrite size_poly0 /= dvd0p expr0 mulr1 (negPf q_neq0).
+move=> hsp /=; have [->|p_neq0] := altP (p =P 0).
+ rewrite size_poly0 /= dvd0p expr0 mulr1 div0p /=.
+ case: ifP=> // _; have := (ihn 0).
+ by rewrite size_poly0 expr0 mulr1 dvd0p=> /(_ isT).
+have [|ncop_pq] := boolP (coprimep _ _); first by rewrite dvdp_mulr ?dvdpp.
+have g_gt1: (1 < size (gcdp p q))%N.
+ have [|//|/eqP] := ltngtP; last by rewrite -coprimep_def (negPf ncop_pq).
+ by rewrite ltnS leqn0 size_poly_eq0 gcdp_eq0 (negPf p_neq0).
+have sd : (size (p %/ gcdp p q) < size p)%N.
+ rewrite size_divp -?size_poly_eq0 -(subnKC g_gt1) // add2n /=.
+ by rewrite -[size _]prednK ?size_poly_gt0 // ltnS subSS leq_subr.
+rewrite -{1}[p](divpK (dvdp_gcdl _ q)) -(subnKC sd) addSnnS exprD mulrA.
+rewrite dvdp_mul ?ihn //; first by rewrite -ltnS (leq_trans sd).
+by rewrite exprS dvdp_mulr // dvdp_gcdr.
+Qed.
+
+Lemma reducible_cubic_root p q :
+ size p <= 4 -> 1 < size q < size p -> q %| p -> {r | root p r}.
+Proof.
+move=> p_le4 /andP[]; rewrite leq_eqVlt eq_sym.
+have [/poly2_root[x qx0] _ _ | _ /= q_gt2 p_gt_q] := size q =P 2.
+ by exists x; rewrite -!dvdp_XsubCl in qx0 *; apply: (dvdp_trans qx0).
+case/dvdpP/sig_eqW=> r def_p; rewrite def_p.
+suffices /poly2_root[x rx0]: size r = 2 by exists x; rewrite rootM rx0.
+have /norP[nz_r nz_q]: ~~ [|| r == 0 | q == 0].
+ by rewrite -mulf_eq0 -def_p -size_poly_gt0 (leq_ltn_trans _ p_gt_q).
+rewrite def_p size_mul // -subn1 leq_subLR ltn_subRL in p_gt_q p_le4.
+by apply/eqP; rewrite -(eqn_add2r (size q)) eqn_leq (leq_trans p_le4).
+Qed.
+
+Lemma cubic_irreducible p :
+ 1 < size p <= 4 -> (forall x, ~~ root p x) -> irreducible_poly p.
+Proof.
+move=> /andP[p_gt1 p_le4] root'p; split=> // q sz_q_neq1 q_dv_p.
+have nz_p: p != 0 by rewrite -size_poly_gt0 ltnW.
+have nz_q: q != 0 by apply: contraTneq q_dv_p => ->; rewrite dvd0p.
+have q_gt1: size q > 1 by rewrite ltn_neqAle eq_sym sz_q_neq1 size_poly_gt0.
+rewrite -dvdp_size_eqp // eqn_leq dvdp_leq //= leqNgt; apply/negP=> p_gt_q.
+by have [|x /idPn//] := reducible_cubic_root p_le4 _ q_dv_p; rewrite q_gt1.
+Qed.
+
+Section FieldRingMap.
+
+Variable rR : ringType.
+
+Variable f : {rmorphism F -> rR}.
+Local Notation "p ^f" := (map_poly f p) : ring_scope.
+
+Implicit Type a b : {poly F}.
+
+Lemma redivp_map a b :
+ redivp a^f b^f = (rscalp a b, (rdivp a b)^f, (rmodp a b)^f).
+Proof.
+rewrite /rdivp /rscalp /rmodp !unlock map_poly_eq0 size_map_poly.
+case: eqP; rewrite /= -(rmorph0 (map_poly_rmorphism f)) //; move/eqP=> q_nz.
+move: (size a) => m; elim: m 0%N 0 a => [|m IHm] qq r a /=.
+ rewrite -!mul_polyC !size_map_poly !lead_coef_map // -(map_polyXn f).
+ by rewrite -!(map_polyC f) -!rmorphM -rmorphB -rmorphD; case: (_ < _).
+rewrite -!mul_polyC !size_map_poly !lead_coef_map // -(map_polyXn f).
+by rewrite -!(map_polyC f) -!rmorphM -rmorphB -rmorphD /= IHm; case: (_ < _).
+Qed.
+
+End FieldRingMap.
+
+Section FieldMap.
+
+Variable rR : idomainType.
+
+Variable f : {rmorphism F -> rR}.
+Local Notation "p ^f" := (map_poly f p) : ring_scope.
+
+Implicit Type a b : {poly F}.
+
+Lemma edivp_map a b :
+ edivp a^f b^f = (0%N, (a %/ b)^f, (a %% b)^f).
+Proof.
+case: (eqVneq b 0) => [-> | bn0].
+ rewrite (rmorph0 (map_poly_rmorphism f)) WeakIdomain.edivp_def !modp0 !divp0.
+ by rewrite (rmorph0 (map_poly_rmorphism f)) scalp0.
+rewrite unlock redivp_map lead_coef_map rmorph_unit; last first.
+ by rewrite unitfE lead_coef_eq0.
+rewrite modpE divpE !map_polyZ !rmorphV ?rmorphX // unitfE.
+by rewrite expf_neq0 // lead_coef_eq0.
+Qed.
+
+Lemma scalp_map p q : scalp p^f q^f = scalp p q.
+Proof. by rewrite /scalp edivp_map edivp_def. Qed.
+
+Lemma map_divp p q : (p %/ q)^f = p^f %/ q^f.
+Proof. by rewrite /divp edivp_map edivp_def. Qed.
+
+Lemma map_modp p q : (p %% q)^f = p^f %% q^f.
+Proof. by rewrite /modp edivp_map edivp_def. Qed.
+
+Lemma egcdp_map p q :
+ egcdp (map_poly f p) (map_poly f q)
+ = (map_poly f (egcdp p q).1, map_poly f (egcdp p q).2).
+Proof.
+wlog le_qp: p q / size q <= size p.
+ move=> IH; have [/IH// | lt_qp] := leqP (size q) (size p).
+ have /IH := ltnW lt_qp; rewrite /egcdp !size_map_poly ltnW // leqNgt lt_qp /=.
+ by case: (egcdp_rec _ _ _) => u v [-> ->].
+rewrite /egcdp !size_map_poly {}le_qp; move: (size q) => n.
+elim: n => /= [|n IHn] in p q *; first by rewrite rmorph1 rmorph0.
+rewrite map_poly_eq0; have [_ | nz_q] := ifPn; first by rewrite rmorph1 rmorph0.
+rewrite -map_modp (IHn q (p %% q)); case: (egcdp_rec _ _ n) => u v /=.
+by rewrite map_polyZ lead_coef_map -rmorphX scalp_map rmorphB rmorphM -map_divp.
+Qed.
+
+Lemma dvdp_map p q : (p^f %| q^f) = (p %| q).
+Proof. by rewrite /dvdp -map_modp map_poly_eq0. Qed.
+
+Lemma eqp_map p q : (p^f %= q^f) = (p %= q).
+Proof. by rewrite /eqp !dvdp_map. Qed.
+
+Lemma gcdp_map p q : (gcdp p q)^f = gcdp p^f q^f.
+Proof.
+wlog lt_p_q: p q / size p < size q.
+ move=> IH; case: (ltnP (size p) (size q)) => [|le_q_p]; first exact: IH.
+ rewrite gcdpE (gcdpE p^f) !size_map_poly ltnNge le_q_p /= -map_modp.
+ case: (eqVneq q 0) => [-> | q_nz]; first by rewrite rmorph0 !gcdp0.
+ by rewrite IH ?ltn_modp.
+elim: {q}_.+1 p {-2}q (ltnSn (size q)) lt_p_q => // m IHm p q le_q_m lt_p_q.
+rewrite gcdpE (gcdpE p^f) !size_map_poly lt_p_q -map_modp.
+case: (eqVneq p 0) => [-> | q_nz]; first by rewrite rmorph0 !gcdp0.
+by rewrite IHm ?(leq_trans lt_p_q) ?ltn_modp.
+Qed.
+
+Lemma coprimep_map p q : coprimep p^f q^f = coprimep p q.
+Proof. by rewrite -!gcdp_eqp1 -eqp_map rmorph1 gcdp_map. Qed.
+
+Lemma gdcop_rec_map p q n : (gdcop_rec p q n)^f = (gdcop_rec p^f q^f n).
+Proof.
+elim: n p q => [|n IH] => /= p q.
+ by rewrite map_poly_eq0; case: eqP; rewrite ?rmorph1 ?rmorph0.
+rewrite /coprimep -gcdp_map size_map_poly.
+by case: eqP => Hq0 //; rewrite -map_divp -IH.
+Qed.
+
+Lemma gdcop_map p q : (gdcop p q)^f = (gdcop p^f q^f).
+Proof. by rewrite /gdcop gdcop_rec_map !size_map_poly. Qed.
+
+End FieldMap.
+
+End FieldDivision.
+
+End Field.
+
+Module ClosedField.
+
+Import Field.
+
+Section closed.
+
+Variable F : closedFieldType.
+
+Lemma root_coprimep (p q : {poly F}):
+ (forall x, root p x -> q.[x] != 0) -> coprimep p q.
+Proof.
+move=> Ncmn; rewrite -gcdp_eqp1 -size_poly_eq1; apply/closed_rootP.
+by case=> r; rewrite root_gcd !rootE=> /andP [/Ncmn/negbTE->].
+Qed.
+
+Lemma coprimepP (p q : {poly F}):
+ reflect (forall x, root p x -> q.[x] != 0) (coprimep p q).
+Proof.
+ by apply: (iffP idP)=> [/coprimep_root|/root_coprimep].
+Qed.
+
+End closed.
+
+End ClosedField.
+
+End Pdiv.
+
+Export Pdiv.Field.
diff --git a/mathcomp/algebra/rat.v b/mathcomp/algebra/rat.v
new file mode 100644
index 0000000..b2b88d9
--- /dev/null
+++ b/mathcomp/algebra/rat.v
@@ -0,0 +1,808 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype.
+Require Import bigop ssralg div ssrnum ssrint.
+
+(******************************************************************************)
+(* This file defines a datatype for rational numbers and equips it with a *)
+(* structure of archimedean, real field, with int and nat declared as closed *)
+(* subrings. *)
+(* rat == the type of rational number, with single constructor Rat *)
+(* Rat p h == the element of type rat build from p a pair of integers and*)
+(* h a proof of (0 < p.2) && coprime `|p.1| `|p.2| *)
+(* n%:Q == explicit cast from int to rat, postfix notation for the *)
+(* ratz constant *)
+(* numq r == numerator of (r : rat) *)
+(* denq r == denominator of (r : rat) *)
+(* x \is a Qint == x is an element of rat whose denominator is equal to 1 *)
+(* x \is a Qnat == x is a positive element of rat whose denominator is equal *)
+(* to 1 *)
+(* ratr x == generic embedding of (r : R) into an arbitrary unitring. *)
+(******************************************************************************)
+
+Import GRing.Theory.
+Import Num.Theory.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Local Open Scope ring_scope.
+Local Notation sgr := Num.sg.
+
+Record rat : Set := Rat {
+ valq : (int * int) ;
+ _ : (0 < valq.2) && coprime `|valq.1| `|valq.2|
+}.
+
+Bind Scope ring_scope with rat.
+Delimit Scope rat_scope with Q.
+
+Definition ratz (n : int) := @Rat (n, 1) (coprimen1 _).
+(* Coercion ratz (n : int) := @Rat (n, 1) (coprimen1 _). *)
+
+Canonical rat_subType := Eval hnf in [subType for valq].
+Definition rat_eqMixin := [eqMixin of rat by <:].
+Canonical rat_eqType := EqType rat rat_eqMixin.
+Definition rat_choiceMixin := [choiceMixin of rat by <:].
+Canonical rat_choiceType := ChoiceType rat rat_choiceMixin.
+Definition rat_countMixin := [countMixin of rat by <:].
+Canonical rat_countType := CountType rat rat_countMixin.
+Canonical rat_subCountType := [subCountType of rat].
+
+Definition numq x := nosimpl ((valq x).1).
+Definition denq x := nosimpl ((valq x).2).
+
+Lemma denq_gt0 x : 0 < denq x.
+Proof. by rewrite /denq; case: x=> [[a b] /= /andP []]. Qed.
+Hint Resolve denq_gt0.
+
+Definition denq_ge0 x := ltrW (denq_gt0 x).
+
+Lemma denq_lt0 x : (denq x < 0) = false. Proof. by rewrite ltr_gtF. Qed.
+
+Lemma denq_neq0 x : denq x != 0.
+Proof. by rewrite /denq gtr_eqF ?denq_gt0. Qed.
+Hint Resolve denq_neq0.
+
+Lemma denq_eq0 x : (denq x == 0) = false.
+Proof. exact: negPf (denq_neq0 _). Qed.
+
+Lemma coprime_num_den x : coprime `|numq x| `|denq x|.
+Proof. by rewrite /numq /denq; case: x=> [[a b] /= /andP []]. Qed.
+
+Fact RatK x P : @Rat (numq x, denq x) P = x.
+Proof. by move:x P => [[a b] P'] P; apply: val_inj. Qed.
+
+Fact fracq_subproof : forall x : int * int,
+ let n :=
+ if x.2 == 0 then 0 else
+ (-1) ^ ((x.2 < 0) (+) (x.1 < 0)) * (`|x.1| %/ gcdn `|x.1| `|x.2|)%:Z in
+ let d := if x.2 == 0 then 1 else (`|x.2| %/ gcdn `|x.1| `|x.2|)%:Z in
+ (0 < d) && (coprime `|n| `|d|).
+Proof.
+move=> [m n] /=; case: (altP (n =P 0))=> [//|n0].
+rewrite ltz_nat divn_gt0 ?gcdn_gt0 ?absz_gt0 ?n0 ?orbT //.
+rewrite dvdn_leq ?absz_gt0 ?dvdn_gcdr //= !abszM absz_sign mul1n.
+have [->|m0] := altP (m =P 0); first by rewrite div0n gcd0n divnn absz_gt0 n0.
+move: n0 m0; rewrite -!absz_gt0 absz_nat.
+move: `|_|%N `|_|%N => {m n} [|m] [|n] // _ _.
+rewrite /coprime -(@eqn_pmul2l (gcdn m.+1 n.+1)) ?gcdn_gt0 //.
+rewrite muln_gcdr; do 2!rewrite muln_divCA ?(dvdn_gcdl, dvdn_gcdr) ?divnn //.
+by rewrite ?gcdn_gt0 ?muln1.
+Qed.
+
+Definition fracq (x : int * int) := nosimpl (@Rat (_, _) (fracq_subproof x)).
+
+Fact ratz_frac n : ratz n = fracq (n, 1).
+Proof. by apply: val_inj; rewrite /= gcdn1 !divn1 abszE mulr_sign_norm. Qed.
+
+Fact valqK x : fracq (valq x) = x.
+Proof.
+move:x => [[n d] /= Pnd]; apply: val_inj=> /=.
+move: Pnd; rewrite /coprime /fracq /=; case/andP=> hd; move/eqP=> hnd.
+by rewrite ltr_gtF ?gtr_eqF //= hnd !divn1 mulz_sign_abs abszE gtr0_norm.
+Qed.
+
+Fact scalq_key : unit. Proof. by []. Qed.
+Definition scalq_def x := sgr x.2 * (gcdn `|x.1| `|x.2|)%:Z.
+Definition scalq := locked_with scalq_key scalq_def.
+Canonical scalq_unlockable := [unlockable fun scalq].
+
+Fact scalq_eq0 x : (scalq x == 0) = (x.2 == 0).
+Proof.
+case: x => n d; rewrite unlock /= mulf_eq0 sgr_eq0 /= eqz_nat.
+rewrite -[gcdn _ _ == 0%N]negbK -lt0n gcdn_gt0 ?absz_gt0 [X in ~~ X]orbC.
+by case: sgrP.
+Qed.
+
+Lemma sgr_scalq x : sgr (scalq x) = sgr x.2.
+Proof.
+rewrite unlock sgrM sgr_id -[(gcdn _ _)%:Z]intz sgr_nat.
+by rewrite -lt0n gcdn_gt0 ?absz_gt0 orbC; case: sgrP; rewrite // mul0r.
+Qed.
+
+Lemma signr_scalq x : (scalq x < 0) = (x.2 < 0).
+Proof. by rewrite -!sgr_cp0 sgr_scalq. Qed.
+
+Lemma scalqE x :
+ x.2 != 0 -> scalq x = (-1) ^+ (x.2 < 0)%R * (gcdn `|x.1| `|x.2|)%:Z.
+Proof. by rewrite unlock; case: sgrP. Qed.
+
+Fact valq_frac x :
+ x.2 != 0 -> x = (scalq x * numq (fracq x), scalq x * denq (fracq x)).
+Proof.
+case: x => [n d] /= d_neq0; rewrite /denq /numq scalqE //= (negPf d_neq0).
+rewrite mulr_signM -mulrA -!PoszM addKb.
+do 2!rewrite muln_divCA ?(dvdn_gcdl, dvdn_gcdr) // divnn.
+by rewrite gcdn_gt0 !absz_gt0 d_neq0 orbT !muln1 !mulz_sign_abs.
+Qed.
+
+Definition zeroq := fracq (0, 1).
+Definition oneq := fracq (1, 1).
+
+Fact frac0q x : fracq (0, x) = zeroq.
+Proof.
+apply: val_inj; rewrite //= div0n !gcd0n !mulr0 !divnn.
+by have [//|x_neq0] := altP eqP; rewrite absz_gt0 x_neq0.
+Qed.
+
+Fact fracq0 x : fracq (x, 0) = zeroq. Proof. exact/eqP. Qed.
+
+CoInductive fracq_spec (x : int * int) : int * int -> rat -> Type :=
+ | FracqSpecN of x.2 = 0 : fracq_spec x (x.1, 0) zeroq
+ | FracqSpecP k fx of k != 0 : fracq_spec x (k * numq fx, k * denq fx) fx.
+
+Fact fracqP x : fracq_spec x x (fracq x).
+Proof.
+case: x => n d /=; have [d_eq0 | d_neq0] := eqVneq d 0.
+ by rewrite d_eq0 fracq0; constructor.
+by rewrite {2}[(_, _)]valq_frac //; constructor; rewrite scalq_eq0.
+Qed.
+
+Lemma rat_eqE x y : (x == y) = (numq x == numq y) && (denq x == denq y).
+Proof.
+rewrite -val_eqE [val x]surjective_pairing [val y]surjective_pairing /=.
+by rewrite xpair_eqE.
+Qed.
+
+Lemma sgr_denq x : sgr (denq x) = 1. Proof. by apply/eqP; rewrite sgr_cp0. Qed.
+
+Lemma normr_denq x : `|denq x| = denq x. Proof. by rewrite gtr0_norm. Qed.
+
+Lemma absz_denq x : `|denq x|%N = denq x :> int.
+Proof. by rewrite abszE normr_denq. Qed.
+
+Lemma rat_eq x y : (x == y) = (numq x * denq y == numq y * denq x).
+Proof.
+symmetry; rewrite rat_eqE andbC.
+have [->|] /= := altP (denq _ =P _); first by rewrite (inj_eq (mulIf _)).
+apply: contraNF => /eqP hxy; rewrite -absz_denq -[X in _ == X]absz_denq.
+rewrite eqz_nat /= eqn_dvd.
+rewrite -(@Gauss_dvdr _ `|numq x|) 1?coprime_sym ?coprime_num_den // andbC.
+rewrite -(@Gauss_dvdr _ `|numq y|) 1?coprime_sym ?coprime_num_den //.
+by rewrite -!abszM hxy -{1}hxy !abszM !dvdn_mull ?dvdnn.
+Qed.
+
+Fact fracq_eq x y : x.2 != 0 -> y.2 != 0 ->
+ (fracq x == fracq y) = (x.1 * y.2 == y.1 * x.2).
+Proof.
+case: fracqP=> //= u fx u_neq0 _; case: fracqP=> //= v fy v_neq0 _; symmetry.
+rewrite [X in (_ == X)]mulrC mulrACA [X in (_ == X)]mulrACA.
+by rewrite [denq _ * _]mulrC (inj_eq (mulfI _)) ?mulf_neq0 // rat_eq.
+Qed.
+
+Fact fracq_eq0 x : (fracq x == zeroq) = (x.1 == 0) || (x.2 == 0).
+Proof.
+move: x=> [n d] /=; have [->|d0] := altP (d =P 0).
+ by rewrite fracq0 eqxx orbT.
+by rewrite orbF fracq_eq ?d0 //= mulr1 mul0r.
+Qed.
+
+Fact fracqMM x n d : x != 0 -> fracq (x * n, x * d) = fracq (n, d).
+Proof.
+move=> x_neq0; apply/eqP.
+have [->|d_neq0] := eqVneq d 0; first by rewrite mulr0 !fracq0.
+by rewrite fracq_eq ?mulf_neq0 //= mulrCA mulrA.
+Qed.
+
+Definition addq_subdef (x y : int * int) := (x.1 * y.2 + y.1 * x.2, x.2 * y.2).
+Definition addq (x y : rat) := nosimpl fracq (addq_subdef (valq x) (valq y)).
+
+Definition oppq_subdef (x : int * int) := (- x.1, x.2).
+Definition oppq (x : rat) := nosimpl fracq (oppq_subdef (valq x)).
+
+Fact addq_subdefC : commutative addq_subdef.
+Proof. by move=> x y; rewrite /addq_subdef addrC [_.2 * _]mulrC. Qed.
+
+Fact addq_subdefA : associative addq_subdef.
+Proof.
+move=> x y z; rewrite /addq_subdef.
+by rewrite !mulrA !mulrDl addrA ![_ * x.2]mulrC !mulrA.
+Qed.
+
+Fact addq_frac x y : x.2 != 0 -> y.2 != 0 ->
+ (addq (fracq x) (fracq y)) = fracq (addq_subdef x y).
+Proof.
+case: fracqP => // u fx u_neq0 _; case: fracqP => // v fy v_neq0 _.
+rewrite /addq_subdef /= ![(_ * numq _) * _]mulrACA [(_ * denq _) * _]mulrACA.
+by rewrite [v * _]mulrC -mulrDr fracqMM ?mulf_neq0.
+Qed.
+
+Fact ratzD : {morph ratz : x y / x + y >-> addq x y}.
+Proof.
+by move=> x y /=; rewrite !ratz_frac addq_frac // /addq_subdef /= !mulr1.
+Qed.
+
+Fact oppq_frac x : oppq (fracq x) = fracq (oppq_subdef x).
+Proof.
+rewrite /oppq_subdef; case: fracqP => /= [|u fx u_neq0].
+ by rewrite fracq0.
+by rewrite -mulrN fracqMM.
+Qed.
+
+Fact ratzN : {morph ratz : x / - x >-> oppq x}.
+Proof. by move=> x /=; rewrite !ratz_frac oppq_frac // /add /= !mulr1. Qed.
+
+Fact addqC : commutative addq.
+Proof. by move=> x y; rewrite /addq /=; rewrite addq_subdefC. Qed.
+
+Fact addqA : associative addq.
+Proof.
+move=> x y z; rewrite -[x]valqK -[y]valqK -[z]valqK.
+by rewrite !addq_frac ?mulf_neq0 ?denq_neq0 // addq_subdefA.
+Qed.
+
+Fact add0q : left_id zeroq addq.
+Proof.
+move=> x; rewrite -[x]valqK addq_frac ?denq_neq0 // /addq_subdef /=.
+by rewrite mul0r add0r mulr1 mul1r -surjective_pairing.
+Qed.
+
+Fact addNq : left_inverse (fracq (0, 1)) oppq addq.
+Proof.
+move=> x; rewrite -[x]valqK !(addq_frac, oppq_frac) ?denq_neq0 //.
+rewrite /addq_subdef /oppq_subdef //= mulNr addNr; apply/eqP.
+by rewrite fracq_eq ?mulf_neq0 ?denq_neq0 //= !mul0r.
+Qed.
+
+Definition rat_ZmodMixin := ZmodMixin addqA addqC add0q addNq.
+Canonical rat_ZmodType := ZmodType rat rat_ZmodMixin.
+
+Definition mulq_subdef (x y : int * int) := nosimpl (x.1 * y.1, x.2 * y.2).
+Definition mulq (x y : rat) := nosimpl fracq (mulq_subdef (valq x) (valq y)).
+
+Fact mulq_subdefC : commutative mulq_subdef.
+Proof. by move=> x y; rewrite /mulq_subdef mulrC [_ * x.2]mulrC. Qed.
+
+Fact mul_subdefA : associative mulq_subdef.
+Proof. by move=> x y z; rewrite /mulq_subdef !mulrA. Qed.
+
+Definition invq_subdef (x : int * int) := nosimpl (x.2, x.1).
+Definition invq (x : rat) := nosimpl fracq (invq_subdef (valq x)).
+
+Fact mulq_frac x y : (mulq (fracq x) (fracq y)) = fracq (mulq_subdef x y).
+Proof.
+rewrite /mulq_subdef; case: fracqP => /= [|u fx u_neq0].
+ by rewrite mul0r fracq0 /mulq /mulq_subdef /= mul0r frac0q.
+case: fracqP=> /= [|v fy v_neq0].
+ by rewrite mulr0 fracq0 /mulq /mulq_subdef /= mulr0 frac0q.
+by rewrite ![_ * (v * _)]mulrACA fracqMM ?mulf_neq0.
+Qed.
+
+Fact ratzM : {morph ratz : x y / x * y >-> mulq x y}.
+Proof. by move=> x y /=; rewrite !ratz_frac mulq_frac // /= !mulr1. Qed.
+
+Fact invq_frac x :
+ x.1 != 0 -> x.2 != 0 -> invq (fracq x) = fracq (invq_subdef x).
+Proof.
+by rewrite /invq_subdef; case: fracqP => // k {x} x k0; rewrite fracqMM.
+Qed.
+
+Fact mulqC : commutative mulq.
+Proof. by move=> x y; rewrite /mulq mulq_subdefC. Qed.
+
+Fact mulqA : associative mulq.
+Proof.
+by move=> x y z; rewrite -[x]valqK -[y]valqK -[z]valqK !mulq_frac mul_subdefA.
+Qed.
+
+Fact mul1q : left_id oneq mulq.
+Proof.
+move=> x; rewrite -[x]valqK; rewrite mulq_frac /mulq_subdef.
+by rewrite !mul1r -surjective_pairing.
+Qed.
+
+Fact mulq_addl : left_distributive mulq addq.
+Proof.
+move=> x y z; rewrite -[x]valqK -[y]valqK -[z]valqK /=.
+rewrite !(mulq_frac, addq_frac) ?mulf_neq0 ?denq_neq0 //=.
+apply/eqP; rewrite fracq_eq ?mulf_neq0 ?denq_neq0 //= !mulrDl; apply/eqP.
+by rewrite !mulrA ![_ * (valq z).1]mulrC !mulrA ![_ * (valq x).2]mulrC !mulrA.
+Qed.
+
+Fact nonzero1q : oneq != zeroq. Proof. by []. Qed.
+
+Definition rat_comRingMixin :=
+ ComRingMixin mulqA mulqC mul1q mulq_addl nonzero1q.
+Canonical rat_Ring := Eval hnf in RingType rat rat_comRingMixin.
+Canonical rat_comRing := Eval hnf in ComRingType rat mulqC.
+
+Fact mulVq x : x != 0 -> mulq (invq x) x = 1.
+Proof.
+rewrite -[x]valqK fracq_eq ?denq_neq0 //= mulr1 mul0r=> nx0.
+rewrite !(mulq_frac, invq_frac) ?denq_neq0 //.
+by apply/eqP; rewrite fracq_eq ?mulf_neq0 ?denq_neq0 //= mulr1 mul1r mulrC.
+Qed.
+
+Fact invq0 : invq 0 = 0. Proof. by apply/eqP. Qed.
+
+Definition RatFieldUnitMixin := FieldUnitMixin mulVq invq0.
+Canonical rat_unitRing :=
+ Eval hnf in UnitRingType rat RatFieldUnitMixin.
+Canonical rat_comUnitRing := Eval hnf in [comUnitRingType of rat].
+
+Fact rat_field_axiom : GRing.Field.mixin_of rat_unitRing. Proof. exact. Qed.
+
+Definition RatFieldIdomainMixin := (FieldIdomainMixin rat_field_axiom).
+Canonical rat_iDomain :=
+ Eval hnf in IdomainType rat (FieldIdomainMixin rat_field_axiom).
+Canonical rat_fieldType := FieldType rat rat_field_axiom.
+
+Lemma numq_eq0 x : (numq x == 0) = (x == 0).
+Proof.
+rewrite -[x]valqK fracq_eq0; case: fracqP=> /= [|k {x} x k0].
+ by rewrite eqxx orbT.
+by rewrite !mulf_eq0 (negPf k0) /= denq_eq0 orbF.
+Qed.
+
+Notation "n %:Q" := ((n : int)%:~R : rat)
+ (at level 2, left associativity, format "n %:Q") : ring_scope.
+
+Hint Resolve denq_neq0 denq_gt0 denq_ge0.
+
+Definition subq (x y : rat) : rat := (addq x (oppq y)).
+Definition divq (x y : rat) : rat := (mulq x (invq y)).
+
+Notation "0" := zeroq : rat_scope.
+Notation "1" := oneq : rat_scope.
+Infix "+" := addq : rat_scope.
+Notation "- x" := (oppq x) : rat_scope.
+Infix "*" := mulq : rat_scope.
+Notation "x ^-1" := (invq x) : rat_scope.
+Infix "-" := subq : rat_scope.
+Infix "/" := divq : rat_scope.
+
+(* ratz should not be used, %:Q should be used instead *)
+Lemma ratzE n : ratz n = n%:Q.
+Proof.
+elim: n=> [|n ihn|n ihn]; first by rewrite mulr0z ratz_frac.
+ by rewrite intS mulrzDl ratzD ihn.
+by rewrite intS opprD mulrzDl ratzD ihn.
+Qed.
+
+Lemma numq_int n : numq n%:Q = n. Proof. by rewrite -ratzE. Qed.
+Lemma denq_int n : denq n%:Q = 1. Proof. by rewrite -ratzE. Qed.
+
+Lemma rat0 : 0%:Q = 0. Proof. by []. Qed.
+Lemma rat1 : 1%:Q = 1. Proof. by []. Qed.
+
+Lemma numqN x : numq (- x) = - numq x.
+Proof.
+rewrite /numq; case: x=> [[a b] /= /andP [hb]]; rewrite /coprime=> /eqP hab.
+by rewrite ltr_gtF ?gtr_eqF // {2}abszN hab divn1 mulz_sign_abs.
+Qed.
+
+Lemma denqN x : denq (- x) = denq x.
+Proof.
+rewrite /denq; case: x=> [[a b] /= /andP [hb]]; rewrite /coprime=> /eqP hab.
+by rewrite gtr_eqF // abszN hab divn1 gtz0_abs.
+Qed.
+
+(* Will be subsumed by pnatr_eq0 *)
+Fact intq_eq0 n : (n%:~R == 0 :> rat) = (n == 0)%N.
+Proof. by rewrite -ratzE /ratz rat_eqE /numq /denq /= mulr0 eqxx andbT. Qed.
+
+(* fracq should never appear, its canonical form is _%:Q / _%:Q *)
+Lemma fracqE x : fracq x = x.1%:Q / x.2%:Q.
+Proof.
+move:x => [m n] /=.
+case n0: (n == 0); first by rewrite (eqP n0) fracq0 rat0 invr0 mulr0.
+rewrite -[m%:Q]valqK -[n%:Q]valqK.
+rewrite [_^-1]invq_frac ?(denq_neq0, numq_eq0, n0, intq_eq0) //.
+rewrite [_ / _]mulq_frac /= /invq_subdef /mulq_subdef /=.
+by rewrite -!/(numq _) -!/(denq _) !numq_int !denq_int mul1r mulr1.
+Qed.
+
+Lemma divq_num_den x : (numq x)%:Q / (denq x)%:Q = x.
+Proof. by rewrite -{3}[x]valqK [valq _]surjective_pairing /= fracqE. Qed.
+
+CoInductive divq_spec (n d : int) : int -> int -> rat -> Type :=
+| DivqSpecN of d = 0 : divq_spec n d n 0 0
+| DivqSpecP k x of k != 0 : divq_spec n d (k * numq x) (k * denq x) x.
+
+(* replaces fracqP *)
+Lemma divqP n d : divq_spec n d n d (n%:Q / d%:Q).
+Proof.
+set x := (n, d); rewrite -[n]/x.1 -[d]/x.2 -fracqE.
+by case: fracqP => [_|k fx k_neq0] /=; constructor.
+Qed.
+
+Lemma divq_eq (nx dx ny dy : rat) :
+ dx != 0 -> dy != 0 -> (nx / dx == ny / dy) = (nx * dy == ny * dx).
+Proof.
+move=> dx_neq0 dy_neq0; rewrite -(inj_eq (@mulIf _ (dx * dy) _)) ?mulf_neq0 //.
+by rewrite mulrA divfK // mulrCA divfK // [dx * _ ]mulrC.
+Qed.
+
+CoInductive rat_spec (* (x : rat) *) : rat -> int -> int -> Type :=
+ Rat_spec (n : int) (d : nat) & coprime `|n| d.+1
+ : rat_spec (* x *) (n%:Q / d.+1%:Q) n d.+1.
+
+Lemma ratP x : rat_spec x (numq x) (denq x).
+Proof.
+rewrite -{1}[x](divq_num_den); case hd: denq => [p|n].
+ have: 0 < p%:Z by rewrite -hd denq_gt0.
+ case: p hd=> //= n hd; constructor; rewrite -?hd ?divq_num_den //.
+ by rewrite -[n.+1]/`|n.+1|%N -hd coprime_num_den.
+by move: (denq_gt0 x); rewrite hd.
+Qed.
+
+Lemma coprimeq_num n d : coprime `|n| `|d| -> numq (n%:~R / d%:~R) = sgr d * n.
+Proof.
+move=> cnd /=; have <- := fracqE (n, d).
+rewrite /numq /= (eqP (cnd : _ == 1%N)) divn1.
+have [|d_gt0|d_lt0] := sgrP d;
+by rewrite (mul0r, mul1r, mulN1r) //= ?[_ ^ _]signrN ?mulNr mulz_sign_abs.
+Qed.
+
+Lemma coprimeq_den n d :
+ coprime `|n| `|d| -> denq (n%:~R / d%:~R) = (if d == 0 then 1 else `|d|).
+Proof.
+move=> cnd; have <- := fracqE (n, d).
+by rewrite /denq /= (eqP (cnd : _ == 1%N)) divn1; case: d {cnd}.
+Qed.
+
+Lemma denqVz (i : int) : i != 0 -> denq (i%:~R^-1) = `|i|.
+Proof.
+by move=> h; rewrite -div1r -[1]/(1%:~R) coprimeq_den /= ?coprime1n // (negPf h).
+Qed.
+
+Lemma numqE x : (numq x)%:~R = x * (denq x)%:~R.
+Proof. by rewrite -{2}[x]divq_num_den divfK // intq_eq0 denq_eq0. Qed.
+
+Lemma denqP x : {d | denq x = d.+1}.
+Proof. by rewrite /denq; case: x => [[_ [[|d]|]] //= _]; exists d. Qed.
+
+Definition normq (x : rat) : rat := `|numq x|%:~R / (denq x)%:~R.
+Definition le_rat (x y : rat) := numq x * denq y <= numq y * denq x.
+Definition lt_rat (x y : rat) := numq x * denq y < numq y * denq x.
+
+Lemma gt_rat0 x : lt_rat 0 x = (0 < numq x).
+Proof. by rewrite /lt_rat mul0r mulr1. Qed.
+
+Lemma lt_rat0 x : lt_rat x 0 = (numq x < 0).
+Proof. by rewrite /lt_rat mul0r mulr1. Qed.
+
+Lemma ge_rat0 x : le_rat 0 x = (0 <= numq x).
+Proof. by rewrite /le_rat mul0r mulr1. Qed.
+
+Lemma le_rat0 x : le_rat x 0 = (numq x <= 0).
+Proof. by rewrite /le_rat mul0r mulr1. Qed.
+
+Fact le_rat0D x y : le_rat 0 x -> le_rat 0 y -> le_rat 0 (x + y).
+Proof.
+rewrite !ge_rat0 => hnx hny.
+have hxy: (0 <= numq x * denq y + numq y * denq x).
+ by rewrite addr_ge0 ?mulr_ge0.
+by rewrite /numq /= -!/(denq _) ?mulf_eq0 ?denq_eq0 !ler_gtF ?mulr_ge0.
+Qed.
+
+Fact le_rat0M x y : le_rat 0 x -> le_rat 0 y -> le_rat 0 (x * y).
+Proof.
+rewrite !ge_rat0 => hnx hny.
+have hxy: (0 <= numq x * denq y + numq y * denq x).
+ by rewrite addr_ge0 ?mulr_ge0.
+by rewrite /numq /= -!/(denq _) ?mulf_eq0 ?denq_eq0 !ler_gtF ?mulr_ge0.
+Qed.
+
+Fact le_rat0_anti x : le_rat 0 x -> le_rat x 0 -> x = 0.
+Proof.
+by move=> hx hy; apply/eqP; rewrite -numq_eq0 eqr_le -ge_rat0 -le_rat0 hx hy.
+Qed.
+
+Lemma sgr_numq_div (n d : int) : sgr (numq (n%:Q / d%:Q)) = sgr n * sgr d.
+Proof.
+set x := (n, d); rewrite -[n]/x.1 -[d]/x.2 -fracqE.
+case: fracqP => [|k fx k_neq0] /=; first by rewrite mulr0.
+by rewrite !sgrM mulrACA -expr2 sqr_sg k_neq0 sgr_denq mulr1 mul1r.
+Qed.
+
+Fact subq_ge0 x y : le_rat 0 (y - x) = le_rat x y.
+Proof.
+symmetry; rewrite ge_rat0 /le_rat -subr_ge0.
+case: ratP => nx dx cndx; case: ratP => ny dy cndy.
+rewrite -!mulNr addf_div ?intq_eq0 // !mulNr -!rmorphM -rmorphB /=.
+symmetry; rewrite !lerNgt -sgr_cp0 sgr_numq_div mulrC gtr0_sg //.
+by rewrite mul1r sgr_cp0.
+Qed.
+
+Fact le_rat_total : total le_rat.
+Proof. by move=> x y; apply: ler_total. Qed.
+
+Fact numq_sign_mul (b : bool) x : numq ((-1) ^+ b * x) = (-1) ^+ b * numq x.
+Proof. by case: b; rewrite ?(mul1r, mulN1r) // numqN. Qed.
+
+Fact numq_div_lt0 n d : n != 0 -> d != 0 ->
+ (numq (n%:~R / d%:~R) < 0)%R = (n < 0)%R (+) (d < 0)%R.
+Proof.
+move=> n0 d0; rewrite -sgr_cp0 sgr_numq_div !sgr_def n0 d0.
+by rewrite !mulr1n -signr_addb; case: (_ (+) _).
+Qed.
+
+Lemma normr_num_div n d : `|numq (n%:~R / d%:~R)| = numq (`|n|%:~R / `|d|%:~R).
+Proof.
+rewrite (normrEsg n) (normrEsg d) !rmorphM /= invfM mulrACA !sgr_def.
+have [->|n_neq0] := altP eqP; first by rewrite mul0r mulr0.
+have [->|d_neq0] := altP eqP; first by rewrite invr0 !mulr0.
+rewrite !intr_sign invr_sign -signr_addb numq_sign_mul -numq_div_lt0 //.
+by apply: (canRL (signrMK _)); rewrite mulz_sign_abs.
+Qed.
+
+Fact norm_ratN x : normq (- x) = normq x.
+Proof. by rewrite /normq numqN denqN normrN. Qed.
+
+Fact ge_rat0_norm x : le_rat 0 x -> normq x = x.
+Proof.
+rewrite ge_rat0; case: ratP=> [] // n d cnd n_ge0.
+by rewrite /normq /= normr_num_div ?ger0_norm // divq_num_den.
+Qed.
+
+Fact lt_rat_def x y : (lt_rat x y) = (y != x) && (le_rat x y).
+Proof. by rewrite /lt_rat ltr_def rat_eq. Qed.
+
+Definition ratLeMixin := RealLeMixin le_rat0D le_rat0M le_rat0_anti
+ subq_ge0 (@le_rat_total 0) norm_ratN ge_rat0_norm lt_rat_def.
+
+Canonical rat_numDomainType := NumDomainType rat ratLeMixin.
+Canonical rat_numFieldType := [numFieldType of rat].
+Canonical rat_realDomainType := RealDomainType rat (@le_rat_total 0).
+Canonical rat_realFieldType := [realFieldType of rat].
+
+Lemma numq_ge0 x : (0 <= numq x) = (0 <= x).
+Proof.
+by case: ratP => n d cnd; rewrite ?pmulr_lge0 ?invr_gt0 (ler0z, ltr0z).
+Qed.
+
+Lemma numq_le0 x : (numq x <= 0) = (x <= 0).
+Proof. by rewrite -oppr_ge0 -numqN numq_ge0 oppr_ge0. Qed.
+
+Lemma numq_gt0 x : (0 < numq x) = (0 < x).
+Proof. by rewrite !ltrNge numq_le0. Qed.
+
+Lemma numq_lt0 x : (numq x < 0) = (x < 0).
+Proof. by rewrite !ltrNge numq_ge0. Qed.
+
+Lemma sgr_numq x : sgz (numq x) = sgz x.
+Proof.
+apply/eqP; case: (sgzP x); rewrite sgz_cp0 ?(numq_gt0, numq_lt0) //.
+by move->.
+Qed.
+
+Lemma denq_mulr_sign (b : bool) x : denq ((-1) ^+ b * x) = denq x.
+Proof. by case: b; rewrite ?(mul1r, mulN1r) // denqN. Qed.
+
+Lemma denq_norm x : denq `|x| = denq x.
+Proof. by rewrite normrEsign denq_mulr_sign. Qed.
+
+Fact rat_archimedean : Num.archimedean_axiom [numDomainType of rat].
+Proof.
+move=> x; exists `|numq x|.+1; rewrite mulrS ltr_spaddl //.
+rewrite pmulrn abszE intr_norm numqE normrM ler_pemulr ?norm_ge0 //.
+by rewrite -intr_norm ler1n absz_gt0 denq_eq0.
+Qed.
+
+Canonical archiType := ArchiFieldType rat rat_archimedean.
+
+Section QintPred.
+
+Definition Qint := [qualify a x : rat | denq x == 1].
+Fact Qint_key : pred_key Qint. Proof. by []. Qed.
+Canonical Qint_keyed := KeyedQualifier Qint_key.
+
+Lemma Qint_def x : (x \is a Qint) = (denq x == 1). Proof. by []. Qed.
+
+Lemma numqK : {in Qint, cancel (fun x => numq x) intr}.
+Proof. by move=> x /(_ =P 1 :> int) Zx; rewrite numqE Zx rmorph1 mulr1. Qed.
+
+Lemma QintP x : reflect (exists z, x = z%:~R) (x \in Qint).
+Proof.
+apply: (iffP idP) => [/numqK <- | [z ->]]; first by exists (numq x).
+by rewrite Qint_def denq_int.
+Qed.
+
+Fact Qint_subring_closed : subring_closed Qint.
+Proof.
+split=> // _ _ /QintP[x ->] /QintP[y ->]; apply/QintP.
+ by exists (x - y); rewrite -rmorphB.
+by exists (x * y); rewrite -rmorphM.
+Qed.
+
+Canonical Qint_opprPred := OpprPred Qint_subring_closed.
+Canonical Qint_addrPred := AddrPred Qint_subring_closed.
+Canonical Qint_mulrPred := MulrPred Qint_subring_closed.
+Canonical Qint_zmodPred := ZmodPred Qint_subring_closed.
+Canonical Qint_semiringPred := SemiringPred Qint_subring_closed.
+Canonical Qint_smulrPred := SmulrPred Qint_subring_closed.
+Canonical Qint_subringPred := SubringPred Qint_subring_closed.
+
+End QintPred.
+
+Section QnatPred.
+
+Definition Qnat := [qualify a x : rat | (x \is a Qint) && (0 <= x)].
+Fact Qnat_key : pred_key Qnat. Proof. by []. Qed.
+Canonical Qnat_keyed := KeyedQualifier Qnat_key.
+
+Lemma Qnat_def x : (x \is a Qnat) = (x \is a Qint) && (0 <= x).
+Proof. by []. Qed.
+
+Lemma QnatP x : reflect (exists n : nat, x = n%:R) (x \in Qnat).
+Proof.
+rewrite Qnat_def; apply: (iffP idP) => [/andP []|[n ->]]; last first.
+ by rewrite Qint_def pmulrn denq_int eqxx ler0z.
+by move=> /QintP [] [] n ->; rewrite ?ler0z // => _; exists n.
+Qed.
+
+Fact Qnat_semiring_closed : semiring_closed Qnat.
+Proof.
+do 2?split; move => // x y; rewrite !Qnat_def => /andP[xQ hx] /andP[yQ hy].
+ by rewrite rpredD // addr_ge0.
+by rewrite rpredM // mulr_ge0.
+Qed.
+
+Canonical Qnat_addrPred := AddrPred Qnat_semiring_closed.
+Canonical Qnat_mulrPred := MulrPred Qnat_semiring_closed.
+Canonical Qnat_semiringPred := SemiringPred Qnat_semiring_closed.
+
+End QnatPred.
+
+Lemma natq_div m n : n %| m -> (m %/ n)%:R = m%:R / n%:R :> rat.
+Proof. by apply: char0_natf_div; apply: char_num. Qed.
+
+Section InRing.
+
+Variable R : unitRingType.
+
+Definition ratr x : R := (numq x)%:~R / (denq x)%:~R.
+
+Lemma ratr_int z : ratr z%:~R = z%:~R.
+Proof. by rewrite /ratr numq_int denq_int divr1. Qed.
+
+Lemma ratr_nat n : ratr n%:R = n%:R.
+Proof. exact: (ratr_int n). Qed.
+
+Lemma rpred_rat S (ringS : @divringPred R S) (kS : keyed_pred ringS) a :
+ ratr a \in kS.
+Proof. by rewrite rpred_div ?rpred_int. Qed.
+
+End InRing.
+
+Section Fmorph.
+
+Implicit Type rR : unitRingType.
+
+Lemma fmorph_rat (aR : fieldType) rR (f : {rmorphism aR -> rR}) a :
+ f (ratr _ a) = ratr _ a.
+Proof. by rewrite fmorph_div !rmorph_int. Qed.
+
+Lemma fmorph_eq_rat rR (f : {rmorphism rat -> rR}) : f =1 ratr _.
+Proof. by move=> a; rewrite -{1}[a]divq_num_den fmorph_div !rmorph_int. Qed.
+
+End Fmorph.
+
+Section Linear.
+
+Implicit Types (U V : lmodType rat) (A B : lalgType rat).
+
+Lemma rat_linear U V (f : U -> V) : additive f -> linear f.
+Proof.
+move=> fB a u v; pose phi := Additive fB; rewrite [f _](raddfD phi).
+congr (_ + _); rewrite -{2}[a]divq_num_den mulrC -scalerA.
+apply: canRL (scalerK _) _; first by rewrite intr_eq0 denq_neq0.
+by rewrite !scaler_int -raddfMz scalerMzl -mulrzr -numqE scaler_int raddfMz.
+Qed.
+
+Lemma rat_lrmorphism A B (f : A -> B) : rmorphism f -> lrmorphism f.
+Proof. by case=> /rat_linear fZ fM; do ?split=> //; apply: fZ. Qed.
+
+End Linear.
+
+Section InPrealField.
+
+Variable F : numFieldType.
+
+Fact ratr_is_rmorphism : rmorphism (@ratr F).
+Proof.
+have injZtoQ: @injective rat int intr by exact: intr_inj.
+have nz_den x: (denq x)%:~R != 0 :> F by rewrite intr_eq0 denq_eq0.
+do 2?split; rewrite /ratr ?divr1 // => x y; last first.
+ rewrite mulrC mulrAC; apply: canLR (mulKf (nz_den _)) _; rewrite !mulrA.
+ do 2!apply: canRL (mulfK (nz_den _)) _; rewrite -!rmorphM; congr _%:~R.
+ apply: injZtoQ; rewrite !rmorphM [x * y]lock /= !numqE -lock.
+ by rewrite -!mulrA mulrA mulrCA -!mulrA (mulrCA y).
+apply: (canLR (mulfK (nz_den _))); apply: (mulIf (nz_den x)).
+rewrite mulrAC mulrBl divfK ?nz_den // mulrAC -!rmorphM.
+apply: (mulIf (nz_den y)); rewrite mulrAC mulrBl divfK ?nz_den //.
+rewrite -!(rmorphM, rmorphB); congr _%:~R; apply: injZtoQ.
+rewrite !(rmorphM, rmorphB) [_ - _]lock /= -lock !numqE.
+by rewrite (mulrAC y) -!mulrBl -mulrA mulrAC !mulrA.
+Qed.
+
+Canonical ratr_additive := Additive ratr_is_rmorphism.
+Canonical ratr_rmorphism := RMorphism ratr_is_rmorphism.
+
+Lemma ler_rat : {mono (@ratr F) : x y / x <= y}.
+Proof.
+move=> x y /=; case: (ratP x) => nx dx cndx; case: (ratP y) => ny dy cndy.
+rewrite !fmorph_div /= !ratr_int !ler_pdivl_mulr ?ltr0z //.
+by rewrite ![_ / _ * _]mulrAC !ler_pdivr_mulr ?ltr0z // -!rmorphM /= !ler_int.
+Qed.
+
+Lemma ltr_rat : {mono (@ratr F) : x y / x < y}.
+Proof. exact: lerW_mono ler_rat. Qed.
+
+Lemma ler0q x : (0 <= ratr F x) = (0 <= x).
+Proof. by rewrite (_ : 0 = ratr F 0) ?ler_rat ?rmorph0. Qed.
+
+Lemma lerq0 x : (ratr F x <= 0) = (x <= 0).
+Proof. by rewrite (_ : 0 = ratr F 0) ?ler_rat ?rmorph0. Qed.
+
+Lemma ltr0q x : (0 < ratr F x) = (0 < x).
+Proof. by rewrite (_ : 0 = ratr F 0) ?ltr_rat ?rmorph0. Qed.
+
+Lemma ltrq0 x : (ratr F x < 0) = (x < 0).
+Proof. by rewrite (_ : 0 = ratr F 0) ?ltr_rat ?rmorph0. Qed.
+
+Lemma ratr_sg x : ratr F (sgr x) = sgr (ratr F x).
+Proof. by rewrite !sgr_def fmorph_eq0 ltrq0 rmorphMn rmorph_sign. Qed.
+
+Lemma ratr_norm x : ratr F `|x| = `|ratr F x|.
+Proof.
+rewrite {2}[x]numEsign rmorphMsign normrMsign [`|ratr F _|]ger0_norm //.
+by rewrite ler0q ?normr_ge0.
+Qed.
+
+End InPrealField.
+
+Implicit Arguments ratr [[R]].
+
+(* Conntecting rationals to the ring an field tactics *)
+
+Ltac rat_to_ring :=
+ rewrite -?[0%Q]/(0 : rat)%R -?[1%Q]/(1 : rat)%R
+ -?[(_ - _)%Q]/(_ - _ : rat)%R -?[(_ / _)%Q]/(_ / _ : rat)%R
+ -?[(_ + _)%Q]/(_ + _ : rat)%R -?[(_ * _)%Q]/(_ * _ : rat)%R
+ -?[(- _)%Q]/(- _ : rat)%R -?[(_ ^-1)%Q]/(_ ^-1 : rat)%R /=.
+
+Ltac ring_to_rat :=
+ rewrite -?[0%R]/0%Q -?[1%R]/1%Q
+ -?[(_ - _)%R]/(_ - _)%Q -?[(_ / _)%R]/(_ / _)%Q
+ -?[(_ + _)%R]/(_ + _)%Q -?[(_ * _)%R]/(_ * _)%Q
+ -?[(- _)%R]/(- _)%Q -?[(_ ^-1)%R]/(_ ^-1)%Q /=.
+
+Lemma rat_ring_theory : (ring_theory 0%Q 1%Q addq mulq subq oppq eq).
+Proof.
+split => * //; rat_to_ring;
+by rewrite ?(add0r, addrA, mul1r, mulrA, mulrDl, subrr) // (addrC, mulrC).
+Qed.
+
+Require setoid_ring.Field_theory setoid_ring.Field_tac.
+
+Lemma rat_field_theory :
+ Field_theory.field_theory 0%Q 1%Q addq mulq subq oppq divq invq eq.
+Proof.
+split => //; first exact rat_ring_theory.
+by move=> p /eqP p_neq0; rat_to_ring; rewrite mulVf.
+Qed.
+
+Add Field rat_field : rat_field_theory.
diff --git a/mathcomp/algebra/ring_quotient.v b/mathcomp/algebra/ring_quotient.v
new file mode 100644
index 0000000..479508c
--- /dev/null
+++ b/mathcomp/algebra/ring_quotient.v
@@ -0,0 +1,650 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import eqtype choice ssreflect ssrbool ssrnat ssrfun seq.
+Require Import ssralg generic_quotient.
+
+
+(******************************************************************************)
+(* This file describes quotients of algebraic structures. *)
+(* *)
+(* It defines a join hierarchy mxing the structures defined in file ssralg *)
+(* (up to unit ring type) and the quotType quotient structure defined in *)
+(* file generic_quotient. Every structure in that (join) hierarchy is *)
+(* parametrized by a base type T and the constants and operations on the *)
+(* base type that will be used to confer its algebraic structure to the *)
+(* quotient. Note that T itself is in general not an instance of an *)
+(* algebraic structure. The canonical surjection from T onto its quotient *)
+(* should be compatible with the parameter operations. *)
+(* *)
+(* The second part of the file provides a definition of (non trivial) *)
+(* decidable ideals (resp. prime ideals) of an arbitrary instance of ring *)
+(* structure and a construction of the quotient of a ring by such an ideal. *)
+(* These definitions extend the hierarchy of sub-structures defined in file *)
+(* ssralg (see Module Pred in ssralg), following a similar methodology. *)
+(* Although the definition of the (structure of) quotient of a ring by an *)
+(* ideal is a general one, we do not provide infrastructure for the case of *)
+(* non commutative ring and left or two-sided ideals. *)
+(* *)
+(* The file defines the following Structures: *)
+(* zmodQuotType T e z n a == Z-module obtained by quotienting type T *)
+(* with the relation e and whose neutral, *)
+(* opposite and addition are the images in the *)
+(* quotient of the parameters z, n and a, *)
+(* respectively. *)
+(* ringQuotType T e z n a o m == ring obtained by quotienting type T with *)
+(* the relation e and whose zero opposite, *)
+(* addition, one, and multiplication are the *)
+(* images in the quotient of the parameters *)
+(* z, n, a, o, m, respectively. *)
+(* unitRingQuotType ... u i == As in the previous cases, instance of unit *)
+(* ring whose unit predicate is obtained from *)
+(* u and the inverse from i. *)
+(* idealr R S == (S : pred R) is a non-trivial, decidable, *)
+(* right ideal of the ring R. *)
+(* prime_idealr R S == (S : pred R) is a non-trivial, decidable, *)
+(* right, prime ideal of the ring R. *)
+(* *)
+(* The formalization of ideals features the following constructions: *)
+(* nontrivial_ideal S == the collective predicate (S : pred R) on the *)
+(* ring R is stable by the ring product and does *)
+(* contain R's one. *)
+(* prime_idealr_closed S := u * v \in S -> (u \in S) || (v \in S) *)
+(* idealr_closed S == the collective predicate (S : pred R) on the *)
+(* ring R represents a (right) ideal. This *)
+(* implies its being a nontrivial_ideal. *)
+(* *)
+(* MkIdeal idealS == packs idealS : nontrivial_ideal S into an *)
+(* idealr S interface structure associating the *)
+(* idealr_closed property to the canonical *)
+(* pred_key S (see ssrbool), which must already *)
+(* be an zmodPred (see ssralg). *)
+(* MkPrimeIdeal pidealS == packs pidealS : prime_idealr_closed S into a *)
+(* prime_idealr S interface structure associating *)
+(* the prime_idealr_closed property to the *)
+(* canonical pred_key S (see ssrbool), which must *)
+(* already be an idealr (see above). *)
+(* {ideal_quot kI} == quotient by the keyed (right) ideal predicate *)
+(* kI of a commutative ring R. Note that we indeed*)
+(* only provide canonical structures of ring *)
+(* quotients for the case of commutative rings, *)
+(* for which a right ideal is obviously a *)
+(* two-sided ideal. *)
+(* *)
+(* Note : *)
+(* if (I : pred R) is a predicate over a ring R and (ideal : idealr I) is an *)
+(* instance of (right) ideal, in order to quantify over an arbitrary (keyed) *)
+(* predicate describing ideal, use type (keyed_pred ideal), as in: *)
+(* forall (kI : keyed_pred ideal),... *)
+(******************************************************************************)
+
+
+Import GRing.Theory.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Local Open Scope ring_scope.
+Local Open Scope quotient_scope.
+
+Reserved Notation "{ideal_quot I }" (at level 0, format "{ideal_quot I }").
+Reserved Notation "m = n %[mod_ideal I ]" (at level 70, n at next level,
+ format "'[hv ' m '/' = n '/' %[mod_ideal I ] ']'").
+Reserved Notation "m == n %[mod_ideal I ]" (at level 70, n at next level,
+ format "'[hv ' m '/' == n '/' %[mod_ideal I ] ']'").
+Reserved Notation "m <> n %[mod_ideal I ]" (at level 70, n at next level,
+ format "'[hv ' m '/' <> n '/' %[mod_ideal I ] ']'").
+Reserved Notation "m != n %[mod_ideal I ]" (at level 70, n at next level,
+ format "'[hv ' m '/' != n '/' %[mod_ideal I ] ']'").
+
+
+Section ZmodQuot.
+
+Variable (T : Type).
+Variable eqT : rel T.
+Variables (zeroT : T) (oppT : T -> T) (addT : T -> T -> T).
+
+Record zmod_quot_mixin_of (Q : Type) (qc : quot_class_of T Q)
+ (zc : GRing.Zmodule.class_of Q) := ZmodQuotMixinPack {
+ zmod_eq_quot_mixin :> eq_quot_mixin_of eqT qc zc;
+ _ : \pi_(QuotTypePack qc Q) zeroT = 0 :> GRing.Zmodule.Pack zc Q;
+ _ : {morph \pi_(QuotTypePack qc Q) : x /
+ oppT x >-> @GRing.opp (GRing.Zmodule.Pack zc Q) x};
+ _ : {morph \pi_(QuotTypePack qc Q) : x y /
+ addT x y >-> @GRing.add (GRing.Zmodule.Pack zc Q) x y}
+}.
+
+Record zmod_quot_class_of (Q : Type) : Type := ZmodQuotClass {
+ zmod_quot_quot_class :> quot_class_of T Q;
+ zmod_quot_zmod_class :> GRing.Zmodule.class_of Q;
+ zmod_quot_mixin :> zmod_quot_mixin_of
+ zmod_quot_quot_class zmod_quot_zmod_class
+}.
+
+Structure zmodQuotType : Type := ZmodQuotTypePack {
+ zmod_quot_sort :> Type;
+ _ : zmod_quot_class_of zmod_quot_sort;
+ _ : Type
+}.
+
+Implicit Type zqT : zmodQuotType.
+
+Definition zmod_quot_class zqT : zmod_quot_class_of zqT :=
+ let: ZmodQuotTypePack _ cT _ as qT' := zqT return zmod_quot_class_of qT' in cT.
+
+Definition zmod_eq_quot_class zqT (zqc : zmod_quot_class_of zqT) :
+ eq_quot_class_of eqT zqT := EqQuotClass zqc.
+
+Canonical zmodQuotType_eqType zqT := Equality.Pack (zmod_quot_class zqT) zqT.
+Canonical zmodQuotType_choiceType zqT :=
+ Choice.Pack (zmod_quot_class zqT) zqT.
+Canonical zmodQuotType_zmodType zqT :=
+ GRing.Zmodule.Pack (zmod_quot_class zqT) zqT.
+Canonical zmodQuotType_quotType zqT := QuotTypePack (zmod_quot_class zqT) zqT.
+Canonical zmodQuotType_eqQuotType zqT := EqQuotTypePack
+ (zmod_eq_quot_class (zmod_quot_class zqT)) zqT.
+
+Coercion zmodQuotType_eqType : zmodQuotType >-> eqType.
+Coercion zmodQuotType_choiceType : zmodQuotType >-> choiceType.
+Coercion zmodQuotType_zmodType : zmodQuotType >-> zmodType.
+Coercion zmodQuotType_quotType : zmodQuotType >-> quotType.
+Coercion zmodQuotType_eqQuotType : zmodQuotType >-> eqQuotType.
+
+Definition ZmodQuotType_pack Q :=
+ fun (qT : quotType T) (zT : zmodType) qc zc
+ of phant_id (quot_class qT) qc & phant_id (GRing.Zmodule.class zT) zc =>
+ fun m => ZmodQuotTypePack (@ZmodQuotClass Q qc zc m) Q.
+
+Definition ZmodQuotMixin_pack Q :=
+ fun (qT : eqQuotType eqT) (qc : eq_quot_class_of eqT Q)
+ of phant_id (eq_quot_class qT) qc =>
+ fun (zT : zmodType) zc of phant_id (GRing.Zmodule.class zT) zc =>
+ fun e m0 mN mD => @ZmodQuotMixinPack Q qc zc e m0 mN mD.
+
+Definition ZmodQuotType_clone (Q : Type) qT cT
+ of phant_id (zmod_quot_class qT) cT := @ZmodQuotTypePack Q cT Q.
+
+Lemma zmod_quot_mixinP zqT :
+ zmod_quot_mixin_of (zmod_quot_class zqT) (zmod_quot_class zqT).
+Proof. by case: zqT => [] ? [] ? ? []. Qed.
+
+Lemma pi_zeror zqT : \pi_zqT zeroT = 0.
+Proof. by case: zqT => [] ? [] ? ? []. Qed.
+
+Lemma pi_oppr zqT : {morph \pi_zqT : x / oppT x >-> - x}.
+Proof. by case: zqT => [] ? [] ? ? []. Qed.
+
+Lemma pi_addr zqT : {morph \pi_zqT : x y / addT x y >-> x + y}.
+Proof. by case: zqT => [] ? [] ? ? []. Qed.
+
+Canonical pi_zero_quot_morph zqT := PiMorph (pi_zeror zqT).
+Canonical pi_opp_quot_morph zqT := PiMorph1 (pi_oppr zqT).
+Canonical pi_add_quot_morph zqT := PiMorph2 (pi_addr zqT).
+
+End ZmodQuot.
+
+Notation ZmodQuotType z o a Q m :=
+ (@ZmodQuotType_pack _ _ z o a Q _ _ _ _ id id m).
+Notation "[ 'zmodQuotType' z , o & a 'of' Q ]" :=
+ (@ZmodQuotType_clone _ _ z o a Q _ _ id)
+ (at level 0, format "[ 'zmodQuotType' z , o & a 'of' Q ]") : form_scope.
+Notation ZmodQuotMixin Q m0 mN mD :=
+ (@ZmodQuotMixin_pack _ _ _ _ _ Q _ _ id _ _ id (pi_eq_quot _) m0 mN mD).
+
+Section PiAdditive.
+
+Variables (V : zmodType) (equivV : rel V) (zeroV : V).
+Variable Q : @zmodQuotType V equivV zeroV -%R +%R.
+
+Lemma pi_is_additive : additive \pi_Q.
+Proof. by move=> x y /=; rewrite !piE. Qed.
+
+Canonical pi_additive := Additive pi_is_additive.
+
+End PiAdditive.
+
+Section RingQuot.
+
+Variable (T : Type).
+Variable eqT : rel T.
+Variables (zeroT : T) (oppT : T -> T) (addT : T -> T -> T).
+Variables (oneT : T) (mulT : T -> T -> T).
+
+Record ring_quot_mixin_of (Q : Type) (qc : quot_class_of T Q)
+ (rc : GRing.Ring.class_of Q) := RingQuotMixinPack {
+ ring_zmod_quot_mixin :> zmod_quot_mixin_of eqT zeroT oppT addT qc rc;
+ _ : \pi_(QuotTypePack qc Q) oneT = 1 :> GRing.Ring.Pack rc Q;
+ _ : {morph \pi_(QuotTypePack qc Q) : x y /
+ mulT x y >-> @GRing.mul (GRing.Ring.Pack rc Q) x y}
+}.
+
+Record ring_quot_class_of (Q : Type) : Type := RingQuotClass {
+ ring_quot_quot_class :> quot_class_of T Q;
+ ring_quot_ring_class :> GRing.Ring.class_of Q;
+ ring_quot_mixin :> ring_quot_mixin_of
+ ring_quot_quot_class ring_quot_ring_class
+}.
+
+Structure ringQuotType : Type := RingQuotTypePack {
+ ring_quot_sort :> Type;
+ _ : ring_quot_class_of ring_quot_sort;
+ _ : Type
+}.
+
+Implicit Type rqT : ringQuotType.
+
+Definition ring_quot_class rqT : ring_quot_class_of rqT :=
+ let: RingQuotTypePack _ cT _ as qT' := rqT return ring_quot_class_of qT' in cT.
+
+Definition ring_zmod_quot_class rqT (rqc : ring_quot_class_of rqT) :
+ zmod_quot_class_of eqT zeroT oppT addT rqT := ZmodQuotClass rqc.
+Definition ring_eq_quot_class rqT (rqc : ring_quot_class_of rqT) :
+ eq_quot_class_of eqT rqT := EqQuotClass rqc.
+
+Canonical ringQuotType_eqType rqT := Equality.Pack (ring_quot_class rqT) rqT.
+Canonical ringQuotType_choiceType rqT := Choice.Pack (ring_quot_class rqT) rqT.
+Canonical ringQuotType_zmodType rqT :=
+ GRing.Zmodule.Pack (ring_quot_class rqT) rqT.
+Canonical ringQuotType_ringType rqT :=
+ GRing.Ring.Pack (ring_quot_class rqT) rqT.
+Canonical ringQuotType_quotType rqT := QuotTypePack (ring_quot_class rqT) rqT.
+Canonical ringQuotType_eqQuotType rqT :=
+ EqQuotTypePack (ring_eq_quot_class (ring_quot_class rqT)) rqT.
+Canonical ringQuotType_zmodQuotType rqT :=
+ ZmodQuotTypePack (ring_zmod_quot_class (ring_quot_class rqT)) rqT.
+
+Coercion ringQuotType_eqType : ringQuotType >-> eqType.
+Coercion ringQuotType_choiceType : ringQuotType >-> choiceType.
+Coercion ringQuotType_zmodType : ringQuotType >-> zmodType.
+Coercion ringQuotType_ringType : ringQuotType >-> ringType.
+Coercion ringQuotType_quotType : ringQuotType >-> quotType.
+Coercion ringQuotType_eqQuotType : ringQuotType >-> eqQuotType.
+Coercion ringQuotType_zmodQuotType : ringQuotType >-> zmodQuotType.
+
+Definition RingQuotType_pack Q :=
+ fun (qT : quotType T) (zT : ringType) qc rc
+ of phant_id (quot_class qT) qc & phant_id (GRing.Ring.class zT) rc =>
+ fun m => RingQuotTypePack (@RingQuotClass Q qc rc m) Q.
+
+Definition RingQuotMixin_pack Q :=
+ fun (qT : zmodQuotType eqT zeroT oppT addT) =>
+ fun (qc : zmod_quot_class_of eqT zeroT oppT addT Q)
+ of phant_id (zmod_quot_class qT) qc =>
+ fun (rT : ringType) rc of phant_id (GRing.Ring.class rT) rc =>
+ fun mZ m1 mM => @RingQuotMixinPack Q qc rc mZ m1 mM.
+
+Definition RingQuotType_clone (Q : Type) qT cT
+ of phant_id (ring_quot_class qT) cT := @RingQuotTypePack Q cT Q.
+
+Lemma ring_quot_mixinP rqT :
+ ring_quot_mixin_of (ring_quot_class rqT) (ring_quot_class rqT).
+Proof. by case: rqT => [] ? [] ? ? []. Qed.
+
+Lemma pi_oner rqT : \pi_rqT oneT = 1.
+Proof. by case: rqT => [] ? [] ? ? []. Qed.
+
+Lemma pi_mulr rqT : {morph \pi_rqT : x y / mulT x y >-> x * y}.
+Proof. by case: rqT => [] ? [] ? ? []. Qed.
+
+Canonical pi_one_quot_morph rqT := PiMorph (pi_oner rqT).
+Canonical pi_mul_quot_morph rqT := PiMorph2 (pi_mulr rqT).
+
+End RingQuot.
+
+Notation RingQuotType o mul Q mix :=
+ (@RingQuotType_pack _ _ _ _ _ o mul Q _ _ _ _ id id mix).
+Notation "[ 'ringQuotType' o & m 'of' Q ]" :=
+ (@RingQuotType_clone _ _ _ _ _ o m Q _ _ id)
+ (at level 0, format "[ 'ringQuotType' o & m 'of' Q ]") : form_scope.
+Notation RingQuotMixin Q m1 mM :=
+ (@RingQuotMixin_pack _ _ _ _ _ _ _ Q _ _ id _ _ id (zmod_quot_mixinP _) m1 mM).
+
+Section PiRMorphism.
+
+Variables (R : ringType) (equivR : rel R) (zeroR : R).
+
+Variable Q : @ringQuotType R equivR zeroR -%R +%R 1 *%R.
+
+Lemma pi_is_multiplicative : multiplicative \pi_Q.
+Proof. by split; do ?move=> x y /=; rewrite !piE. Qed.
+
+Canonical pi_rmorphism := AddRMorphism pi_is_multiplicative.
+
+End PiRMorphism.
+
+Section UnitRingQuot.
+
+Variable (T : Type).
+Variable eqT : rel T.
+Variables (zeroT : T) (oppT : T -> T) (addT : T -> T -> T).
+Variables (oneT : T) (mulT : T -> T -> T).
+Variables (unitT : pred T) (invT : T -> T).
+
+Record unit_ring_quot_mixin_of (Q : Type) (qc : quot_class_of T Q)
+ (rc : GRing.UnitRing.class_of Q) := UnitRingQuotMixinPack {
+ unit_ring_zmod_quot_mixin :>
+ ring_quot_mixin_of eqT zeroT oppT addT oneT mulT qc rc;
+ _ : {mono \pi_(QuotTypePack qc Q) : x /
+ unitT x >-> x \in @GRing.unit (GRing.UnitRing.Pack rc Q)};
+ _ : {morph \pi_(QuotTypePack qc Q) : x /
+ invT x >-> @GRing.inv (GRing.UnitRing.Pack rc Q) x}
+}.
+
+Record unit_ring_quot_class_of (Q : Type) : Type := UnitRingQuotClass {
+ unit_ring_quot_quot_class :> quot_class_of T Q;
+ unit_ring_quot_ring_class :> GRing.UnitRing.class_of Q;
+ unit_ring_quot_mixin :> unit_ring_quot_mixin_of
+ unit_ring_quot_quot_class unit_ring_quot_ring_class
+}.
+
+Structure unitRingQuotType : Type := UnitRingQuotTypePack {
+ unit_ring_quot_sort :> Type;
+ _ : unit_ring_quot_class_of unit_ring_quot_sort;
+ _ : Type
+}.
+
+Implicit Type rqT : unitRingQuotType.
+
+Definition unit_ring_quot_class rqT : unit_ring_quot_class_of rqT :=
+ let: UnitRingQuotTypePack _ cT _ as qT' := rqT
+ return unit_ring_quot_class_of qT' in cT.
+
+Definition unit_ring_ring_quot_class rqT (rqc : unit_ring_quot_class_of rqT) :
+ ring_quot_class_of eqT zeroT oppT addT oneT mulT rqT := RingQuotClass rqc.
+Definition unit_ring_zmod_quot_class rqT (rqc : unit_ring_quot_class_of rqT) :
+ zmod_quot_class_of eqT zeroT oppT addT rqT := ZmodQuotClass rqc.
+Definition unit_ring_eq_quot_class rqT (rqc : unit_ring_quot_class_of rqT) :
+ eq_quot_class_of eqT rqT := EqQuotClass rqc.
+
+Canonical unitRingQuotType_eqType rqT :=
+ Equality.Pack (unit_ring_quot_class rqT) rqT.
+Canonical unitRingQuotType_choiceType rqT :=
+ Choice.Pack (unit_ring_quot_class rqT) rqT.
+Canonical unitRingQuotType_zmodType rqT :=
+ GRing.Zmodule.Pack (unit_ring_quot_class rqT) rqT.
+Canonical unitRingQuotType_ringType rqT :=
+ GRing.Ring.Pack (unit_ring_quot_class rqT) rqT.
+Canonical unitRingQuotType_unitRingType rqT :=
+ GRing.UnitRing.Pack (unit_ring_quot_class rqT) rqT.
+Canonical unitRingQuotType_quotType rqT :=
+ QuotTypePack (unit_ring_quot_class rqT) rqT.
+Canonical unitRingQuotType_eqQuotType rqT :=
+ EqQuotTypePack (unit_ring_eq_quot_class (unit_ring_quot_class rqT)) rqT.
+Canonical unitRingQuotType_zmodQuotType rqT :=
+ ZmodQuotTypePack (unit_ring_zmod_quot_class (unit_ring_quot_class rqT)) rqT.
+Canonical unitRingQuotType_ringQuotType rqT :=
+ RingQuotTypePack (unit_ring_ring_quot_class (unit_ring_quot_class rqT)) rqT.
+
+Coercion unitRingQuotType_eqType : unitRingQuotType >-> eqType.
+Coercion unitRingQuotType_choiceType : unitRingQuotType >-> choiceType.
+Coercion unitRingQuotType_zmodType : unitRingQuotType >-> zmodType.
+Coercion unitRingQuotType_ringType : unitRingQuotType >-> ringType.
+Coercion unitRingQuotType_unitRingType : unitRingQuotType >-> unitRingType.
+Coercion unitRingQuotType_quotType : unitRingQuotType >-> quotType.
+Coercion unitRingQuotType_eqQuotType : unitRingQuotType >-> eqQuotType.
+Coercion unitRingQuotType_zmodQuotType : unitRingQuotType >-> zmodQuotType.
+Coercion unitRingQuotType_ringQuotType : unitRingQuotType >-> ringQuotType.
+
+Definition UnitRingQuotType_pack Q :=
+ fun (qT : quotType T) (rT : unitRingType) qc rc
+ of phant_id (quot_class qT) qc & phant_id (GRing.UnitRing.class rT) rc =>
+ fun m => UnitRingQuotTypePack (@UnitRingQuotClass Q qc rc m) Q.
+
+Definition UnitRingQuotMixin_pack Q :=
+ fun (qT : ringQuotType eqT zeroT oppT addT oneT mulT) =>
+ fun (qc : ring_quot_class_of eqT zeroT oppT addT oneT mulT Q)
+ of phant_id (zmod_quot_class qT) qc =>
+ fun (rT : unitRingType) rc of phant_id (GRing.UnitRing.class rT) rc =>
+ fun mR mU mV => @UnitRingQuotMixinPack Q qc rc mR mU mV.
+
+Definition UnitRingQuotType_clone (Q : Type) qT cT
+ of phant_id (unit_ring_quot_class qT) cT := @UnitRingQuotTypePack Q cT Q.
+
+Lemma unit_ring_quot_mixinP rqT :
+ unit_ring_quot_mixin_of (unit_ring_quot_class rqT) (unit_ring_quot_class rqT).
+Proof. by case: rqT => [] ? [] ? ? []. Qed.
+
+Lemma pi_unitr rqT : {mono \pi_rqT : x / unitT x >-> x \in GRing.unit}.
+Proof. by case: rqT => [] ? [] ? ? []. Qed.
+
+Lemma pi_invr rqT : {morph \pi_rqT : x / invT x >-> x^-1}.
+Proof. by case: rqT => [] ? [] ? ? []. Qed.
+
+Canonical pi_unit_quot_morph rqT := PiMono1 (pi_unitr rqT).
+Canonical pi_inv_quot_morph rqT := PiMorph1 (pi_invr rqT).
+
+End UnitRingQuot.
+
+Notation UnitRingQuotType u i Q mix :=
+ (@UnitRingQuotType_pack _ _ _ _ _ _ _ u i Q _ _ _ _ id id mix).
+Notation "[ 'unitRingQuotType' u & i 'of' Q ]" :=
+ (@UnitRingQuotType_clone _ _ _ _ _ _ _ u i Q _ _ id)
+ (at level 0, format "[ 'unitRingQuotType' u & i 'of' Q ]") : form_scope.
+Notation UnitRingQuotMixin Q mU mV :=
+ (@UnitRingQuotMixin_pack _ _ _ _ _ _ _ _ _ Q
+ _ _ id _ _ id (zmod_quot_mixinP _) mU mV).
+
+Section IdealDef.
+
+Definition nontrivial_ideal (R : ringType) (S : predPredType R) : Prop :=
+ 1 \notin S /\ forall a, {in S, forall u, a * u \in S}.
+
+Definition prime_idealr_closed (R : ringType) (S : predPredType R) : Prop :=
+ forall u v, u * v \in S -> (u \in S) || (v \in S).
+
+Definition idealr_closed (R : ringType) (S : predPredType R) :=
+ [/\ 0 \in S, 1 \notin S & forall a, {in S &, forall u v, a * u + v \in S}].
+
+Lemma idealr_closed_nontrivial R S : @idealr_closed R S -> nontrivial_ideal S.
+Proof. by case=> S0 S1 hS; split => // a x xS; rewrite -[_ * _]addr0 hS. Qed.
+
+Lemma idealr_closedB R S : @idealr_closed R S -> zmod_closed S.
+Proof. by case=> S0 _ hS; split=> // x y xS yS; rewrite -mulN1r addrC hS. Qed.
+
+Coercion idealr_closedB : idealr_closed >-> zmod_closed.
+Coercion idealr_closed_nontrivial : idealr_closed >-> nontrivial_ideal.
+
+Structure idealr (R : ringType) (S : predPredType R) := MkIdeal {
+ idealr_zmod :> zmodPred S;
+ _ : nontrivial_ideal S
+}.
+
+Structure prime_idealr (R : ringType) (S : predPredType R) := MkPrimeIdeal {
+ prime_idealr_zmod :> idealr S;
+ _ : prime_idealr_closed S
+}.
+
+Definition Idealr (R : ringType) (I : predPredType R) (zmodI : zmodPred I)
+ (kI : keyed_pred zmodI) : nontrivial_ideal I -> idealr I.
+Proof. by move=> kI1; split => //. Qed.
+
+Section IdealTheory.
+Variables (R : ringType) (I : predPredType R)
+ (idealrI : idealr I) (kI : keyed_pred idealrI).
+
+Lemma idealr1 : 1 \in kI = false.
+Proof. by apply: negPf; case: idealrI kI => ? /= [? _] [] /= _ ->. Qed.
+
+Lemma idealMr a u : u \in kI -> a * u \in kI.
+Proof.
+by case: idealrI kI=> ? /= [? hI] [] /= ? hkI; rewrite !hkI; apply: hI.
+Qed.
+
+Lemma idealr0 : 0 \in kI. Proof. exact: rpred0. Qed.
+
+End IdealTheory.
+
+Section PrimeIdealTheory.
+
+Variables (R : comRingType) (I : predPredType R)
+ (pidealrI : prime_idealr I) (kI : keyed_pred pidealrI).
+
+Lemma prime_idealrM u v : (u * v \in kI) = (u \in kI) || (v \in kI).
+Proof.
+apply/idP/idP; last by case/orP => /idealMr hI; rewrite // mulrC.
+by case: pidealrI kI=> ? /= hI [] /= ? hkI; rewrite !hkI; apply: hI.
+Qed.
+
+End PrimeIdealTheory.
+
+End IdealDef.
+
+Module Quotient.
+Section ZmodQuotient.
+Variables (R : zmodType) (I : predPredType R)
+ (zmodI : zmodPred I) (kI : keyed_pred zmodI).
+
+Definition equiv (x y : R) := (x - y) \in kI.
+
+Lemma equivE x y : (equiv x y) = (x - y \in kI). Proof. by []. Qed.
+
+Lemma equiv_is_equiv : equiv_class_of equiv.
+Proof.
+split=> [x|x y|y x z]; rewrite !equivE ?subrr ?rpred0 //.
+ by rewrite -opprB rpredN.
+by move=> *; rewrite -[x](addrNK y) -addrA rpredD.
+Qed.
+
+Canonical equiv_equiv := EquivRelPack equiv_is_equiv.
+Canonical equiv_encModRel := defaultEncModRel equiv.
+
+Definition type := {eq_quot equiv}.
+Definition type_of of phant R := type.
+
+Canonical rquot_quotType := [quotType of type].
+Canonical rquot_eqType := [eqType of type].
+Canonical rquot_choiceType := [choiceType of type].
+Canonical rquot_eqQuotType := [eqQuotType equiv of type].
+
+Lemma idealrBE x y : (x - y) \in kI = (x == y %[mod type]).
+Proof. by rewrite piE equivE. Qed.
+
+Lemma idealrDE x y : (x + y) \in kI = (x == - y %[mod type]).
+Proof. by rewrite -idealrBE opprK. Qed.
+
+Definition zero : type := lift_cst type 0.
+Definition add := lift_op2 type +%R.
+Definition opp := lift_op1 type -%R.
+
+Canonical pi_zero_morph := PiConst zero.
+
+Lemma pi_opp : {morph \pi : x / - x >-> opp x}.
+Proof.
+move=> x; unlock opp; apply/eqP; rewrite piE equivE.
+by rewrite -opprD rpredN idealrDE opprK reprK.
+Qed.
+
+Canonical pi_opp_morph := PiMorph1 pi_opp.
+
+Lemma pi_add : {morph \pi : x y / x + y >-> add x y}.
+Proof.
+move=> x y /=; unlock add; apply/eqP; rewrite piE equivE.
+rewrite opprD addrAC addrA -addrA.
+by rewrite rpredD // (idealrBE, idealrDE) ?pi_opp ?reprK.
+Qed.
+Canonical pi_add_morph := PiMorph2 pi_add.
+
+Lemma addqA: associative add.
+Proof. by move=> x y z; rewrite -[x]reprK -[y]reprK -[z]reprK !piE addrA. Qed.
+
+Lemma addqC: commutative add.
+Proof. by move=> x y; rewrite -[x]reprK -[y]reprK !piE addrC. Qed.
+
+Lemma add0q: left_id zero add.
+Proof. by move=> x; rewrite -[x]reprK !piE add0r. Qed.
+
+Lemma addNq: left_inverse zero opp add.
+Proof. by move=> x; rewrite -[x]reprK !piE addNr. Qed.
+
+Definition rquot_zmodMixin := ZmodMixin addqA addqC add0q addNq.
+Canonical rquot_zmodType := Eval hnf in ZmodType type rquot_zmodMixin.
+
+Definition rquot_zmodQuotMixin := ZmodQuotMixin type (lock _) pi_opp pi_add.
+Canonical rquot_zmodQuotType := ZmodQuotType 0 -%R +%R type rquot_zmodQuotMixin.
+
+End ZmodQuotient.
+
+Notation "{quot I }" := (@type_of _ _ _ I (Phant _)).
+
+Section RingQuotient.
+
+Variables (R : comRingType) (I : predPredType R)
+ (idealI : idealr I) (kI : keyed_pred idealI).
+
+Local Notation type := {quot kI}.
+
+Definition one: type := lift_cst type 1.
+Definition mul := lift_op2 type *%R.
+
+Canonical pi_one_morph := PiConst one.
+
+Lemma pi_mul: {morph \pi : x y / x * y >-> mul x y}.
+Proof.
+move=> x y; unlock mul; apply/eqP; rewrite piE equivE.
+rewrite -[_ * _](addrNK (x * repr (\pi_type y))) -mulrBr.
+rewrite -addrA -mulrBl rpredD //.
+ by rewrite idealMr // idealrDE opprK reprK.
+by rewrite mulrC idealMr // idealrDE opprK reprK.
+Qed.
+Canonical pi_mul_morph := PiMorph2 pi_mul.
+
+Lemma mulqA: associative mul.
+Proof. by move=> x y z; rewrite -[x]reprK -[y]reprK -[z]reprK !piE mulrA. Qed.
+
+Lemma mulqC: commutative mul.
+Proof. by move=> x y; rewrite -[x]reprK -[y]reprK !piE mulrC. Qed.
+
+Lemma mul1q: left_id one mul.
+Proof. by move=> x; rewrite -[x]reprK !piE mul1r. Qed.
+
+Lemma mulq_addl: left_distributive mul +%R.
+Proof.
+move=> x y z; rewrite -[x]reprK -[y]reprK -[z]reprK.
+by apply/eqP; rewrite piE /= mulrDl equiv_refl.
+Qed.
+
+Lemma nonzero1q: one != 0.
+Proof. by rewrite piE equivE subr0 idealr1. Qed.
+
+Definition rquot_comRingMixin :=
+ ComRingMixin mulqA mulqC mul1q mulq_addl nonzero1q.
+
+Canonical rquot_ringType := Eval hnf in RingType type rquot_comRingMixin.
+Canonical rquot_comRingType := Eval hnf in ComRingType type mulqC.
+
+Definition rquot_ringQuotMixin := RingQuotMixin type (lock _) pi_mul.
+Canonical rquot_ringQuotType := RingQuotType 1 *%R type rquot_ringQuotMixin.
+
+End RingQuotient.
+
+Section IDomainQuotient.
+
+Variables (R : comRingType) (I : predPredType R)
+ (pidealI : prime_idealr I) (kI : keyed_pred pidealI).
+
+Lemma rquot_IdomainAxiom (x y : {quot kI}): x * y = 0 -> (x == 0) || (y == 0).
+Proof.
+by move=> /eqP; rewrite -[x]reprK -[y]reprK !piE !equivE !subr0 prime_idealrM.
+Qed.
+
+End IDomainQuotient.
+
+End Quotient.
+
+Notation "{ideal_quot I }" := (@Quotient.type_of _ _ _ I (Phant _)).
+Notation "x == y %[mod_ideal I ]" :=
+ (x == y %[mod {ideal_quot I}]) : quotient_scope.
+Notation "x = y %[mod_ideal I ]" :=
+ (x = y %[mod {ideal_quot I}]) : quotient_scope.
+Notation "x != y %[mod_ideal I ]" :=
+ (x != y %[mod {ideal_quot I}]) : quotient_scope.
+Notation "x <> y %[mod_ideal I ]" :=
+ (x <> y %[mod {ideal_quot I}]) : quotient_scope.
+
+Canonical Quotient.rquot_eqType.
+Canonical Quotient.rquot_choiceType.
+Canonical Quotient.rquot_zmodType.
+Canonical Quotient.rquot_ringType.
+Canonical Quotient.rquot_quotType.
+Canonical Quotient.rquot_eqQuotType.
+Canonical Quotient.rquot_zmodQuotType.
+Canonical Quotient.rquot_ringQuotType.
diff --git a/mathcomp/algebra/ssralg.v b/mathcomp/algebra/ssralg.v
new file mode 100644
index 0000000..5f4592b
--- /dev/null
+++ b/mathcomp/algebra/ssralg.v
@@ -0,0 +1,6230 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat div seq choice fintype.
+Require Import finfun bigop prime binomial.
+
+(******************************************************************************)
+(* The algebraic part of the Algebraic Hierarchy, as described in *)
+(* ``Packaging mathematical structures'', TPHOLs09, by *)
+(* Francois Garillot, Georges Gonthier, Assia Mahboubi, Laurence Rideau *)
+(* *)
+(* This file defines for each Structure (Zmodule, Ring, etc ...) its type, *)
+(* its packers and its canonical properties : *)
+(* *)
+(* * Zmodule (additive abelian groups): *)
+(* zmodType == interface type for Zmodule structure. *)
+(* ZmodMixin addA addC add0x addNx == builds the mixin for a Zmodule from the *)
+(* algebraic properties of its operations. *)
+(* ZmodType V m == packs the mixin m to build a Zmodule of type *)
+(* zmodType. The carrier type V must have a *)
+(* choiceType canonical structure. *)
+(* [zmodType of V for S] == V-clone of the zmodType structure S: a copy of S *)
+(* where the sort carrier has been replaced by V, *)
+(* and which is therefore a zmodType structure on V. *)
+(* The sort carrier for S must be convertible to V. *)
+(* [zmodType of V] == clone of a canonical zmodType structure on V. *)
+(* Similar to the above, except S is inferred, but *)
+(* possibly with a syntactically different carrier. *)
+(* 0 == the zero (additive identity) of a Zmodule. *)
+(* x + y == the sum of x and y (in a Zmodule). *)
+(* - x == the opposite (additive inverse) of x. *)
+(* x - y == the difference of x and y; this is only notation *)
+(* for x + (- y). *)
+(* x *+ n == n times x, with n in nat (non-negative), i.e., *)
+(* x + (x + .. (x + x)..) (n terms); x *+ 1 is thus *)
+(* convertible to x, and x *+ 2 to x + x. *)
+(* x *- n == notation for - (x *+ n), the opposite of x *+ n. *)
+(* \sum_<range> e == iterated sum for a Zmodule (cf bigop.v). *)
+(* e`_i == nth 0 e i, when e : seq M and M has a zmodType *)
+(* structure. *)
+(* support f == 0.-support f, i.e., [pred x | f x != 0]. *)
+(* oppr_closed S <-> collective predicate S is closed under opposite. *)
+(* addr_closed S <-> collective predicate S is closed under finite *)
+(* sums (0 and x + y in S, for x, y in S). *)
+(* zmod_closed S <-> collective predicate S is closed under zmodType *)
+(* operations (0 and x - y in S, for x, y in S). *)
+(* This property coerces to oppr_pred and addr_pred. *)
+(* OpprPred oppS == packs oppS : oppr_closed S into an opprPred S *)
+(* interface structure associating this property to *)
+(* the canonical pred_key S, i.e. the k for which S *)
+(* has a Canonical keyed_pred k structure (see file *)
+(* ssrbool.v). *)
+(* AddrPred addS == packs addS : addr_closed S into an addrPred S *)
+(* interface structure associating this property to *)
+(* the canonical pred_key S (see above). *)
+(* ZmodPred oppS == packs oppS : oppr_closed S into an zmodPred S *)
+(* interface structure associating the zmod_closed *)
+(* property to the canonical pred_key S (see above), *)
+(* which must already be an addrPred. *)
+(* [zmodMixin of M by <:] == zmodType mixin for a subType whose base type is *)
+(* a zmodType and whose predicate's canonical *)
+(* pred_key is a zmodPred. *)
+(* --> Coq can be made to behave as if all predicates had canonical zmodPred *)
+(* keys by executing Import DefaultKeying GRing.DefaultPred. The required *)
+(* oppr_closed and addr_closed assumptions will be either abstracted, *)
+(* resolved or issued as separate proof obligations by the ssreflect *)
+(* plugin abstraction and Prop-irrelevance functions. *)
+(* * Ring (non-commutative rings): *)
+(* ringType == interface type for a Ring structure. *)
+(* RingMixin mulA mul1x mulx1 mulDx mulxD == builds the mixin for a Ring from *)
+(* the algebraic properties of its multiplicative *)
+(* operators; the carrier type must have a zmodType *)
+(* structure. *)
+(* RingType R m == packs the ring mixin m into a ringType. *)
+(* R^c == the converse Ring for R: R^c is convertible to R *)
+(* but when R has a canonical ringType structure *)
+(* R^c has the converse one: if x y : R^c, then *)
+(* x * y = (y : R) * (x : R). *)
+(* [ringType of R for S] == R-clone of the ringType structure S. *)
+(* [ringType of R] == clone of a canonical ringType structure on R. *)
+(* 1 == the multiplicative identity element of a Ring. *)
+(* n%:R == the ring image of an n in nat; this is just *)
+(* notation for 1 *+ n, so 1%:R is convertible to 1 *)
+(* and 2%:R to 1 + 1. *)
+(* x * y == the ring product of x and y. *)
+(* \prod_<range> e == iterated product for a ring (cf bigop.v). *)
+(* x ^+ n == x to the nth power with n in nat (non-negative), *)
+(* i.e., x * (x * .. (x * x)..) (n factors); x ^+ 1 *)
+(* is thus convertible to x, and x ^+ 2 to x * x. *)
+(* GRing.sign R b := (-1) ^+ b in R : ringType, with b : bool. *)
+(* This is a parsing-only helper notation, to be *)
+(* used for defining more specific instances. *)
+(* GRing.comm x y <-> x and y commute, i.e., x * y = y * x. *)
+(* GRing.lreg x <-> x if left-regular, i.e., *%R x is injective. *)
+(* GRing.rreg x <-> x if right-regular, i.e., *%R x is injective. *)
+(* [char R] == the characteristic of R, defined as the set of *)
+(* prime numbers p such that p%:R = 0 in R. The set *)
+(* [char p] has a most one element, and is *)
+(* implemented as a pred_nat collective predicate *)
+(* (see prime.v); thus the statement p \in [char R] *)
+(* can be read as `R has characteristic p', while *)
+(* [char R] =i pred0 means `R has characteristic 0' *)
+(* when R is a field. *)
+(* Frobenius_aut chRp == the Frobenius automorphism mapping x in R to *)
+(* x ^+ p, where chRp : p \in [char R] is a proof *)
+(* that R has (non-zero) characteristic p. *)
+(* mulr_closed S <-> collective predicate S is closed under finite *)
+(* products (1 and x * y in S for x, y in S). *)
+(* smulr_closed S <-> collective predicate S is closed under products *)
+(* and opposite (-1 and x * y in S for x, y in S). *)
+(* semiring_closed S <-> collective predicate S is closed under semiring *)
+(* operations (0, 1, x + y and x * y in S). *)
+(* subring_closed S <-> collective predicate S is closed under ring *)
+(* operations (1, x - y and x * y in S). *)
+(* MulrPred mulS == packs mulS : mulr_closed S into a mulrPred S, *)
+(* SmulrPred mulS smulrPred S, semiringPred S, or subringPred S *)
+(* SemiringPred mulS interface structure, corresponding to the above *)
+(* SubRingPred mulS properties, respectively, provided S already has *)
+(* the supplementary zmodType closure properties. *)
+(* The properties above coerce to subproperties so, *)
+(* e.g., ringS : subring_closed S can be used for *)
+(* the proof obligations of all prerequisites. *)
+(* [ringMixin of R by <:] == ringType mixin for a subType whose base type is *)
+(* a ringType and whose predicate's canonical key *)
+(* is a SubringPred. *)
+(* --> As for zmodType predicates, Import DefaultKeying GRing.DefaultPred *)
+(* turns unresolved GRing.Pred unification constraints into proof *)
+(* obligations for basic closure assumptions. *)
+(* *)
+(* * ComRing (commutative Rings): *)
+(* comRingType == interface type for commutative ring structure. *)
+(* ComRingType R mulC == packs mulC into a comRingType; the carrier type *)
+(* R must have a ringType canonical structure. *)
+(* ComRingMixin mulA mulC mul1x mulDx == builds the mixin for a Ring (i.e., a *)
+(* *non commutative* ring), using the commutativity *)
+(* to reduce the number of proof obligations. *)
+(* [comRingType of R for S] == R-clone of the comRingType structure S. *)
+(* [comRingType of R] == clone of a canonical comRingType structure on R. *)
+(* [comRingMixin of R by <:] == comutativity mixin axiom for R when it is a *)
+(* subType of a commutative ring. *)
+(* *)
+(* * UnitRing (Rings whose units have computable inverses): *)
+(* unitRingType == interface type for the UnitRing structure. *)
+(* UnitRingMixin mulVr mulrV unitP inv0id == builds the mixin for a UnitRing *)
+(* from the properties of the inverse operation and *)
+(* the boolean test for being a unit (invertible). *)
+(* The inverse of a non-unit x is constrained to be *)
+(* x itself (property inv0id). The carrier type *)
+(* must have a ringType canonical structure. *)
+(* UnitRingType R m == packs the unit ring mixin m into a unitRingType. *)
+(* WARNING: while it is possible to omit R for most of the *)
+(* XxxType functions, R MUST be explicitly given *)
+(* when UnitRingType is used with a mixin produced *)
+(* by ComUnitRingMixin, otherwise the resulting *)
+(* structure will have the WRONG sort key and will *)
+(* NOT BE USED during type inference. *)
+(* [unitRingType of R for S] == R-clone of the unitRingType structure S. *)
+(* [unitRingType of R] == clones a canonical unitRingType structure on R. *)
+(* x \is a GRing.unit <=> x is a unit (i.e., has an inverse). *)
+(* x^-1 == the ring inverse of x, if x is a unit, else x. *)
+(* x / y == x divided by y (notation for x * y^-1). *)
+(* x ^- n := notation for (x ^+ n)^-1, the inverse of x ^+ n. *)
+(* invr_closed S <-> collective predicate S is closed under inverse. *)
+(* divr_closed S <-> collective predicate S is closed under division *)
+(* (1 and x / y in S). *)
+(* sdivr_closed S <-> collective predicate S is closed under division *)
+(* and opposite (-1 and x / y in S, for x, y in S). *)
+(* divring_closed S <-> collective predicate S is closed under unitRing *)
+(* operations (1, x - y and x / y in S). *)
+(* DivrPred invS == packs invS : mulr_closed S into a divrPred S, *)
+(* SdivrPred invS sdivrPred S or divringPred S interface structure, *)
+(* DivringPred invS corresponding to the above properties, resp., *)
+(* provided S already has the supplementary ringType *)
+(* closure properties. The properties above coerce *)
+(* to subproperties, as explained above. *)
+(* [unitRingMixin of R by <:] == unitRingType mixin for a subType whose base *)
+(* type is a unitRingType and whose predicate's *)
+(* canonical key is a divringPred and whose ring *)
+(* structure is compatible with the base type's. *)
+(* *)
+(* * ComUnitRing (commutative rings with computable inverses): *)
+(* comUnitRingType == interface type for ComUnitRing structure. *)
+(* ComUnitRingMixin mulVr unitP inv0id == builds the mixin for a UnitRing (a *)
+(* *non commutative* unit ring, using commutativity *)
+(* to simplify the proof obligations; the carrier *)
+(* type must have a comRingType structure. *)
+(* WARNING: ALWAYS give an explicit type argument *)
+(* to UnitRingType along with a mixin produced by *)
+(* ComUnitRingMixin (see above). *)
+(* [comUnitRingType of R] == a comUnitRingType structure for R created by *)
+(* merging canonical comRingType and unitRingType *)
+(* structures on R. *)
+(* *)
+(* * IntegralDomain (integral, commutative, ring with partial inverses): *)
+(* idomainType == interface type for the IntegralDomain structure. *)
+(* IdomainType R mulf_eq0 == packs the integrality property into an *)
+(* idomainType integral domain structure; R must *)
+(* have a comUnitRingType canonical structure. *)
+(* [idomainType of R for S] == R-clone of the idomainType structure S. *)
+(* [idomainType of R] == clone of a canonical idomainType structure on R. *)
+(* [idomainMixin of R by <:] == mixin axiom for a idomain subType. *)
+(* *)
+(* * Field (commutative fields): *)
+(* fieldType == interface type for fields. *)
+(* GRing.Field.axiom inv == the field axiom (x != 0 -> inv x * x = 1). *)
+(* FieldUnitMixin mulVx unitP inv0id == builds a *non commutative unit ring* *)
+(* mixin, using the field axiom to simplify proof *)
+(* obligations. The carrier type must have a *)
+(* comRingType canonical structure. *)
+(* FieldMixin mulVx == builds the field mixin from the field axiom. The *)
+(* carrier type must have a comRingType structure. *)
+(* FieldIdomainMixin m == builds an *idomain* mixin from a field mixin m. *)
+(* FieldType R m == packs the field mixin M into a fieldType. The *)
+(* carrier type R must be an idomainType. *)
+(* [fieldType of F for S] == F-clone of the fieldType structure S. *)
+(* [fieldType of F] == clone of a canonical fieldType structure on F. *)
+(* [fieldMixin of R by <:] == mixin axiom for a field subType. *)
+(* *)
+(* * DecidableField (fields with a decidable first order theory): *)
+(* decFieldType == interface type for DecidableField structure. *)
+(* DecFieldMixin satP == builds the mixin for a DecidableField from the *)
+(* correctness of its satisfiability predicate. The *)
+(* carrier type must have a unitRingType structure. *)
+(* DecFieldType F m == packs the decidable field mixin m into a *)
+(* decFieldType; the carrier type F must have a *)
+(* fieldType structure. *)
+(* [decFieldType of F for S] == F-clone of the decFieldType structure S. *)
+(* [decFieldType of F] == clone of a canonical decFieldType structure on F *)
+(* GRing.term R == the type of formal expressions in a unit ring R *)
+(* with formal variables 'X_k, k : nat, and *)
+(* manifest constants x%:T, x : R. The notation of *)
+(* all the ring operations is redefined for terms, *)
+(* in scope %T. *)
+(* GRing.formula R == the type of first order formulas over R; the %T *)
+(* scope binds the logical connectives /\, \/, ~, *)
+(* ==>, ==, and != to formulae; GRing.True/False *)
+(* and GRing.Bool b denote constant formulae, and *)
+(* quantifiers are written 'forall/'exists 'X_k, f. *)
+(* GRing.Unit x tests for ring units *)
+(* GRing.If p_f t_f e_f emulates if-then-else *)
+(* GRing.Pick p_f t_f e_f emulates fintype.pick *)
+(* foldr GRing.Exists/Forall q_f xs can be used *)
+(* to write iterated quantifiers. *)
+(* GRing.eval e t == the value of term t with valuation e : seq R *)
+(* (e maps 'X_i to e`_i). *)
+(* GRing.same_env e1 e2 <-> environments e1 and e2 are extensionally equal. *)
+(* GRing.qf_form f == f is quantifier-free. *)
+(* GRing.holds e f == the intuitionistic CiC interpretation of the *)
+(* formula f holds with valuation e. *)
+(* GRing.qf_eval e f == the value (in bool) of a quantifier-free f. *)
+(* GRing.sat e f == valuation e satisfies f (only in a decField). *)
+(* GRing.sol n f == a sequence e of size n such that e satisfies f, *)
+(* if one exists, or [::] if there is no such e. *)
+(* QEdecFieldMixin wfP okP == a decidable field Mixin built from a quantifier *)
+(* eliminator p and proofs wfP : GRing.wf_QE_proj p *)
+(* and okP : GRing.valid_QE_proj p that p returns *)
+(* well-formed and valid formulae, i.e., p i (u, v) *)
+(* is a quantifier-free formula equivalent to *)
+(* 'exists 'X_i, u1 == 0 /\ ... /\ u_m == 0 /\ v1 != 0 ... /\ v_n != 0 *)
+(* *)
+(* * ClosedField (algebraically closed fields): *)
+(* closedFieldType == interface type for the ClosedField structure. *)
+(* ClosedFieldType F m == packs the closed field mixin m into a *)
+(* closedFieldType. The carrier F must have a *)
+(* decFieldType structure. *)
+(* [closedFieldType of F on S] == F-clone of a closedFieldType structure S. *)
+(* [closedFieldType of F] == clone of a canonicalclosedFieldType structure *)
+(* on F. *)
+(* *)
+(* * Lmodule (module with left multiplication by external scalars). *)
+(* lmodType R == interface type for an Lmodule structure with *)
+(* scalars of type R; R must have a ringType *)
+(* structure. *)
+(* LmodMixin scalA scal1v scalxD scalDv == builds an Lmodule mixin from the *)
+(* algebraic properties of the scaling operation; *)
+(* the module carrier type must have a zmodType *)
+(* structure, and the scalar carrier must have a *)
+(* ringType structure. *)
+(* LmodType R V m == packs the mixin v to build an Lmodule of type *)
+(* lmodType R. The carrier type V must have a *)
+(* zmodType structure. *)
+(* [lmodType R of V for S] == V-clone of an lmodType R structure S. *)
+(* [lmodType R of V] == clone of a canonical lmodType R structure on V. *)
+(* a *: v == v scaled by a, when v is in an Lmodule V and a *)
+(* is in the scalar Ring of V. *)
+(* scaler_closed S <-> collective predicate S is closed under scaling. *)
+(* linear_closed S <-> collective predicate S is closed under linear *)
+(* combinations (a *: u + v in S when u, v in S). *)
+(* submod_closed S <-> collective predicate S is closed under lmodType *)
+(* operations (0 and a *: u + v in S). *)
+(* SubmodPred scaleS == packs scaleS : scaler_closed S in a submodPred S *)
+(* interface structure corresponding to the above *)
+(* property, provided S's key is a zmodPred; *)
+(* submod_closed coerces to all the prerequisites. *)
+(* [lmodMixin of V by <:] == mixin for a subType of an lmodType, whose *)
+(* predicate's key is a submodPred. *)
+(* *)
+(* * Lalgebra (left algebra, ring with scaling that associates on the left): *)
+(* lalgType R == interface type for Lalgebra structures with *)
+(* scalars in R; R must have ringType structure. *)
+(* LalgType R V scalAl == packs scalAl : k (x y) = (k x) y into an *)
+(* Lalgebra of type lalgType R. The carrier type V *)
+(* must have both lmodType R and ringType canonical *)
+(* structures. *)
+(* R^o == the regular algebra of R: R^o is convertible to *)
+(* R, but when R has a ringType structure then R^o *)
+(* extends it to an lalgType structure by letting R *)
+(* act on itself: if x : R and y : R^o then *)
+(* x *: y = x * (y : R). *)
+(* k%:A == the image of the scalar k in an L-algebra; this *)
+(* is simply notation for k *: 1. *)
+(* [lalgType R of V for S] == V-clone the lalgType R structure S. *)
+(* [lalgType R of V] == clone of a canonical lalgType R structure on V. *)
+(* subalg_closed S <-> collective predicate S is closed under lalgType *)
+(* operations (1, a *: u + v and u * v in S). *)
+(* SubalgPred scaleS == packs scaleS : scaler_closed S in a subalgPred S *)
+(* interface structure corresponding to the above *)
+(* property, provided S's key is a subringPred; *)
+(* subalg_closed coerces to all the prerequisites. *)
+(* [lalgMixin of V by <:] == mixin axiom for a subType of an lalgType. *)
+(* *)
+(* * Algebra (ring with scaling that associates both left and right): *)
+(* algType R == type for Algebra structure with scalars in R. *)
+(* R should be a commutative ring. *)
+(* AlgType R A scalAr == packs scalAr : k (x y) = x (k y) into an Algebra *)
+(* Structure of type algType R. The carrier type A *)
+(* must have an lalgType R structure. *)
+(* CommAlgType R A == creates an Algebra structure for an A that has *)
+(* both lalgType R and comRingType structures. *)
+(* [algType R of V for S] == V-clone of an algType R structure on S. *)
+(* [algType R of V] == clone of a canonical algType R structure on V. *)
+(* [algMixin of V by <:] == mixin axiom for a subType of an algType. *)
+(* *)
+(* * UnitAlgebra (algebra with computable inverses): *)
+(* unitAlgType R == interface type for UnitAlgebra structure with *)
+(* scalars in R; R should have a unitRingType *)
+(* structure. *)
+(* [unitAlgType R of V] == a unitAlgType R structure for V created by *)
+(* merging canonical algType and unitRingType on V. *)
+(* divalg_closed S <-> collective predicate S is closed under all *)
+(* unitAlgType operations (1, a *: u + v and u / v *)
+(* are in S fo u, v in S). *)
+(* DivalgPred scaleS == packs scaleS : scaler_closed S in a divalgPred S *)
+(* interface structure corresponding to the above *)
+(* property, provided S's key is a divringPred; *)
+(* divalg_closed coerces to all the prerequisites. *)
+(* *)
+(* In addition to this structure hierarchy, we also develop a separate, *)
+(* parallel hierarchy for morphisms linking these structures: *)
+(* *)
+(* * Additive (additive functions): *)
+(* additive f <-> f of type U -> V is additive, i.e., f maps the *)
+(* Zmodule structure of U to that of V, 0 to 0, *)
+(* - to - and + to + (equivalently, binary - to -). *)
+(* := {morph f : u v / u + v}. *)
+(* {additive U -> V} == the interface type for a Structure (keyed on *)
+(* a function f : U -> V) that encapsulates the *)
+(* additive property; both U and V must have *)
+(* zmodType canonical structures. *)
+(* Additive add_f == packs add_f : additive f into an additive *)
+(* function structure of type {additive U -> V}. *)
+(* [additive of f as g] == an f-clone of the additive structure on the *)
+(* function g -- f and g must be convertible. *)
+(* [additive of f] == a clone of an existing additive structure on f. *)
+(* *)
+(* * RMorphism (ring morphisms): *)
+(* multiplicative f <-> f of type R -> S is multiplicative, i.e., f *)
+(* maps 1 and * in R to 1 and * in S, respectively, *)
+(* R ans S must have canonical ringType structures. *)
+(* rmorphism f <-> f is a ring morphism, i.e., f is both additive *)
+(* and multiplicative. *)
+(* {rmorphism R -> S} == the interface type for ring morphisms, i.e., *)
+(* a Structure that encapsulates the rmorphism *)
+(* property for functions f : R -> S; both R and S *)
+(* must have ringType structures. *)
+(* RMorphism morph_f == packs morph_f : rmorphism f into a Ring morphism *)
+(* structure of type {rmorphism R -> S}. *)
+(* AddRMorphism mul_f == packs mul_f : multiplicative f into an rmorphism *)
+(* structure of type {rmorphism R -> S}; f must *)
+(* already have an {additive R -> S} structure. *)
+(* [rmorphism of f as g] == an f-clone of the rmorphism structure of g. *)
+(* [rmorphism of f] == a clone of an existing additive structure on f. *)
+(* -> If R and S are UnitRings the f also maps units to units and inverses *)
+(* of units to inverses; if R is a field then f if a field isomorphism *)
+(* between R and its image. *)
+(* -> As rmorphism coerces to both additive and multiplicative, all *)
+(* structures for f can be built from a single proof of rmorphism f. *)
+(* -> Additive properties (raddf_suffix, see below) are duplicated and *)
+(* specialised for RMorphism (as rmorph_suffix). This allows more *)
+(* precise rewriting and cleaner chaining: although raddf lemmas will *)
+(* recognize RMorphism functions, the converse will not hold (we cannot *)
+(* add reverse inheritance rules because of incomplete backtracking in *)
+(* the Canonical Projection unification), so one would have to insert a *)
+(* /= every time one switched from additive to multiplicative rules. *)
+(* -> The property duplication also means that it is not strictly necessary *)
+(* to declare all Additive instances. *)
+(* *)
+(* * Linear (linear functions): *)
+(* scalable f <-> f of type U -> V is scalable, i.e., f morphs *)
+(* scaling on U to scaling on V, a *: _ to a *: _. *)
+(* U and V must both have lmodType R structures, *)
+(* for the same ringType R. *)
+(* scalable_for s f <-> f is scalable for scaling operator s, i.e., *)
+(* f morphs a *: _ to s a _; the range of f only *)
+(* need to be a zmodType. The scaling operator s *)
+(* should be one of *:%R (see scalable, above), *%R *)
+(* or a combination nu \; *%R or nu \; *:%R with *)
+(* nu : {rmorphism _}; otherwise some of the theory *)
+(* (e.g., the linearZ rule) will not apply. *)
+(* linear f <-> f of type U -> V is linear, i.e., f morphs *)
+(* linear combinations a *: u + v in U to similar *)
+(* linear combinations in V; U and V must both have *)
+(* lmodType R structures, for the same ringType R. *)
+(* := forall a, {morph f: u v / a *: u + v}. *)
+(* scalar f <-> f of type U -> R is a scalar function, i.e., *)
+(* f (a *: u + v) = a * f u + f v. *)
+(* linear_for s f <-> f is linear for the scaling operator s, i.e., *)
+(* f (a *: u + v) = s a (f u) + f v. The range of f *)
+(* only needs to be a zmodType, but s MUST be of *)
+(* the form described in in scalable_for paragraph *)
+(* for this predicate to type check. *)
+(* lmorphism f <-> f is both additive and scalable. This is in *)
+(* fact equivalent to linear f, although somewhat *)
+(* less convenient to prove. *)
+(* lmorphism_for s f <-> f is both additive and scalable for s. *)
+(* {linear U -> V} == the interface type for linear functions, i.e., a *)
+(* Structure that encapsulates the linear property *)
+(* for functions f : U -> V; both U and V must have *)
+(* lmodType R structures, for the same R. *)
+(* {scalar U} == the interface type for scalar functions, of type *)
+(* U -> R where U has an lmodType R structure. *)
+(* {linear U -> V | s} == the interface type for functions linear for s. *)
+(* Linear lin_f == packs lin_f : lmorphism_for s f into a linear *)
+(* function structure of type {linear U -> V | s}. *)
+(* As linear_for s f coerces to lmorphism_for s f, *)
+(* Linear can be used with lin_f : linear_for s f *)
+(* (indeed, that is the recommended usage). Note *)
+(* that as linear f, scalar f, {linear U -> V} and *)
+(* {scalar U} are simply notation for corresponding *)
+(* generic "_for" forms, Linear can be used for any *)
+(* of these special cases, transparently. *)
+(* AddLinear scal_f == packs scal_f : scalable_for s f into a *)
+(* {linear U -> V | s} structure; f must already *)
+(* have an additive structure; as with Linear, *)
+(* AddLinear can be used with lin_f : linear f, etc *)
+(* [linear of f as g] == an f-clone of the linear structure of g. *)
+(* [linear of f] == a clone of an existing linear structure on f. *)
+(* (a *: u)%Rlin == transient forms that simplify to a *: u, a * u, *)
+(* (a * u)%Rlin nu a *: u, and nu a * u, respectively, and are *)
+(* (a *:^nu u)%Rlin created by rewriting with the linearZ lemma. The *)
+(* (a *^nu u)%Rlin forms allows the RHS of linearZ to be matched *)
+(* reliably, using the GRing.Scale.law structure. *)
+(* -> Similarly to Ring morphisms, additive properties are specialized for *)
+(* linear functions. *)
+(* -> Although {scalar U} is convertible to {linear U -> R^o}, it does not *)
+(* actually use R^o, so that rewriting preserves the canonical structure *)
+(* of the range of scalar functions. *)
+(* -> The generic linearZ lemma uses a set of bespoke interface structures to *)
+(* ensure that both left-to-right and right-to-left rewriting work even in *)
+(* the presence of scaling functions that simplify non-trivially (e.g., *)
+(* idfun \; *%R). Because most of the canonical instances and projections *)
+(* are coercions the machinery will be mostly invisible (with only the *)
+(* {linear ...} structure and %Rlin notations showing), but users should *)
+(* beware that in (a *: f u)%Rlin, a actually occurs in the f u subterm. *)
+(* -> The simpler linear_LR, or more specialized linearZZ and scalarZ rules *)
+(* should be used instead of linearZ if there are complexity issues, as *)
+(* well as for explicit forward and backward application, as the main *)
+(* parameter of linearZ is a proper sub-interface of {linear fUV | s}. *)
+(* *)
+(* * LRMorphism (linear ring morphisms, i.e., algebra morphisms): *)
+(* lrmorphism f <-> f of type A -> B is a linear Ring (Algebra) *)
+(* morphism: f is both additive, multiplicative and *)
+(* scalable. A and B must both have lalgType R *)
+(* canonical structures, for the same ringType R. *)
+(* lrmorphism_for s f <-> f a linear Ring morphism for the scaling *)
+(* operator s: f is additive, multiplicative and *)
+(* scalable for s. A must be an lalgType R, but B *)
+(* only needs to have a ringType structure. *)
+(* {lrmorphism A -> B} == the interface type for linear morphisms, i.e., a *)
+(* Structure that encapsulates the lrmorphism *)
+(* property for functions f : A -> B; both A and B *)
+(* must have lalgType R structures, for the same R. *)
+(* {lrmorphism A -> B | s} == the interface type for morphisms linear for s. *)
+(* LRmorphism lrmorph_f == packs lrmorph_f : lrmorphism_for s f into a *)
+(* linear morphism structure of type *)
+(* {lrmorphism A -> B | s}. Like Linear, LRmorphism *)
+(* can be used transparently for lrmorphism f. *)
+(* AddLRmorphism scal_f == packs scal_f : scalable_for s f into a linear *)
+(* morphism structure of type *)
+(* {lrmorphism A -> B | s}; f must already have an *)
+(* {rmorphism A -> B} structure, and AddLRmorphism *)
+(* can be applied to a linear_for s f, linear f, *)
+(* scalar f, etc argument, like AddLinear. *)
+(* [lrmorphism of f] == creates an lrmorphism structure from existing *)
+(* rmorphism and linear structures on f; this is *)
+(* the preferred way of creating lrmorphism *)
+(* structures. *)
+(* -> Linear and rmorphism properties do not need to be specialized for *)
+(* as we supply inheritance join instances in both directions. *)
+(* Finally we supply some helper notation for morphisms: *)
+(* x^f == the image of x under some morphism. This *)
+(* notation is only reserved (not defined) here; *)
+(* it is bound locally in sections where some *)
+(* morphism is used heavily (e.g., the container *)
+(* morphism in the parametricity sections of poly *)
+(* and matrix, or the Frobenius section here). *)
+(* \0 == the constant null function, which has a *)
+(* canonical linear structure, and simplifies on *)
+(* application (see ssrfun.v). *)
+(* f \+ g == the additive composition of f and g, i.e., the *)
+(* function x |-> f x + g x; f \+ g is canonically *)
+(* linear when f and g are, and simplifies on *)
+(* application (see ssrfun.v). *)
+(* f \- g == the function x |-> f x - g x, canonically *)
+(* linear when f and g are, and simplifies on *)
+(* application. *)
+(* k \*: f == the function x |-> k *: f x, which is *)
+(* canonically linear when f is and simplifies on *)
+(* application (this is a shorter alternative to *)
+(* *:%R k \o f). *)
+(* GRing.in_alg A == the ring morphism that injects R into A, where A *)
+(* has an lalgType R structure; GRing.in_alg A k *)
+(* simplifies to k%:A. *)
+(* a \*o f == the function x |-> a * f x, canonically linear *)
+(* linear when f is and its codomain is an algType *)
+(* and which simplifies on application. *)
+(* a \o* f == the function x |-> f x * a, canonically linear *)
+(* linear when f is and its codomain is an lalgType *)
+(* and which simplifies on application. *)
+(* The Lemmas about these structures are contained in both the GRing module *)
+(* and in the submodule GRing.Theory, which can be imported when unqualified *)
+(* access to the theory is needed (GRing.Theory also allows the unqualified *)
+(* use of additive, linear, Linear, etc). The main GRing module should NOT be *)
+(* imported. *)
+(* Notations are defined in scope ring_scope (delimiter %R), except term *)
+(* and formula notations, which are in term_scope (delimiter %T). *)
+(* This library also extends the conventional suffixes described in library *)
+(* ssrbool.v with the following: *)
+(* 0 -- ring 0, as in addr0 : x + 0 = x. *)
+(* 1 -- ring 1, as in mulr1 : x * 1 = x. *)
+(* D -- ring addition, as in linearD : f (u + v) = f u + f v. *)
+(* B -- ring subtraction, as in opprB : - (x - y) = y - x. *)
+(* M -- ring multiplication, as in invfM : (x * y)^-1 = x^-1 * y^-1. *)
+(* Mn -- ring by nat multiplication, as in raddfMn : f (x *+ n) = f x *+ n. *)
+(* N -- ring opposite, as in mulNr : (- x) * y = - (x * y). *)
+(* V -- ring inverse, as in mulVr : x^-1 * x = 1. *)
+(* X -- ring exponentiation, as in rmorphX : f (x ^+ n) = f x ^+ n. *)
+(* Z -- (left) module scaling, as in linearZ : f (a *: v) = s *: f v. *)
+(* The operator suffixes D, B, M and X are also used for the corresponding *)
+(* operations on nat, as in natrX : (m ^ n)%:R = m%:R ^+ n. For the binary *)
+(* power operator, a trailing "n" suffix is used to indicate the operator *)
+(* suffix applies to the left-hand ring argument, as in *)
+(* expr1n : 1 ^+ n = 1 vs. expr1 : x ^+ 1 = x. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Reserved Notation "+%R" (at level 0).
+Reserved Notation "-%R" (at level 0).
+Reserved Notation "*%R" (at level 0, format " *%R").
+Reserved Notation "*:%R" (at level 0, format " *:%R").
+Reserved Notation "n %:R" (at level 2, left associativity, format "n %:R").
+Reserved Notation "k %:A" (at level 2, left associativity, format "k %:A").
+Reserved Notation "[ 'char' F ]" (at level 0, format "[ 'char' F ]").
+
+Reserved Notation "x %:T" (at level 2, left associativity, format "x %:T").
+Reserved Notation "''X_' i" (at level 8, i at level 2, format "''X_' i").
+(* Patch for recurring Coq parser bug: Coq seg faults when a level 200 *)
+(* notation is used as a pattern. *)
+Reserved Notation "''exists' ''X_' i , f"
+ (at level 199, i at level 2, right associativity,
+ format "'[hv' ''exists' ''X_' i , '/ ' f ']'").
+Reserved Notation "''forall' ''X_' i , f"
+ (at level 199, i at level 2, right associativity,
+ format "'[hv' ''forall' ''X_' i , '/ ' f ']'").
+
+Reserved Notation "x ^f" (at level 2, left associativity, format "x ^f").
+
+Reserved Notation "\0" (at level 0).
+Reserved Notation "f \+ g" (at level 50, left associativity).
+Reserved Notation "f \- g" (at level 50, left associativity).
+Reserved Notation "a \*o f" (at level 40).
+Reserved Notation "a \o* f" (at level 40).
+Reserved Notation "a \*: f" (at level 40).
+
+Delimit Scope ring_scope with R.
+Delimit Scope term_scope with T.
+Local Open Scope ring_scope.
+
+Module Import GRing.
+
+Import Monoid.Theory.
+
+Module Zmodule.
+
+Record mixin_of (V : Type) : Type := Mixin {
+ zero : V;
+ opp : V -> V;
+ add : V -> V -> V;
+ _ : associative add;
+ _ : commutative add;
+ _ : left_id zero add;
+ _ : left_inverse zero opp add
+}.
+
+Section ClassDef.
+
+Record class_of T := Class { base : Choice.class_of T; mixin : mixin_of T }.
+Local Coercion base : class_of >-> Choice.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack m :=
+ fun bT b & phant_id (Choice.class bT) b => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> Choice.class_of.
+Coercion mixin : class_of >-> mixin_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Notation zmodType := type.
+Notation ZmodType T m := (@pack T m _ _ id).
+Notation ZmodMixin := Mixin.
+Notation "[ 'zmodType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'zmodType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'zmodType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'zmodType' 'of' T ]") : form_scope.
+End Exports.
+
+End Zmodule.
+Import Zmodule.Exports.
+
+Definition zero V := Zmodule.zero (Zmodule.class V).
+Definition opp V := Zmodule.opp (Zmodule.class V).
+Definition add V := Zmodule.add (Zmodule.class V).
+
+Local Notation "0" := (zero _) : ring_scope.
+Local Notation "-%R" := (@opp _) : ring_scope.
+Local Notation "- x" := (opp x) : ring_scope.
+Local Notation "+%R" := (@add _) : ring_scope.
+Local Notation "x + y" := (add x y) : ring_scope.
+Local Notation "x - y" := (x + - y) : ring_scope.
+
+Definition natmul V x n := nosimpl iterop _ n +%R x (zero V).
+
+Local Notation "x *+ n" := (natmul x n) : ring_scope.
+Local Notation "x *- n" := (- (x *+ n)) : ring_scope.
+
+Local Notation "\sum_ ( i <- r | P ) F" := (\big[+%R/0]_(i <- r | P) F).
+Local Notation "\sum_ ( m <= i < n ) F" := (\big[+%R/0]_(m <= i < n) F).
+Local Notation "\sum_ ( i < n ) F" := (\big[+%R/0]_(i < n) F).
+Local Notation "\sum_ ( i 'in' A ) F" := (\big[+%R/0]_(i in A) F).
+
+Local Notation "s `_ i" := (nth 0 s i) : ring_scope.
+
+Section ZmoduleTheory.
+
+Variable V : zmodType.
+Implicit Types x y : V.
+
+Lemma addrA : @associative V +%R. Proof. by case V => T [? []]. Qed.
+Lemma addrC : @commutative V V +%R. Proof. by case V => T [? []]. Qed.
+Lemma add0r : @left_id V V 0 +%R. Proof. by case V => T [? []]. Qed.
+Lemma addNr : @left_inverse V V V 0 -%R +%R. Proof. by case V => T [? []]. Qed.
+
+Lemma addr0 : @right_id V V 0 +%R.
+Proof. by move=> x; rewrite addrC add0r. Qed.
+Lemma addrN : @right_inverse V V V 0 -%R +%R.
+Proof. by move=> x; rewrite addrC addNr. Qed.
+Definition subrr := addrN.
+
+Canonical add_monoid := Monoid.Law addrA add0r addr0.
+Canonical add_comoid := Monoid.ComLaw addrC.
+
+Lemma addrCA : @left_commutative V V +%R. Proof. exact: mulmCA. Qed.
+Lemma addrAC : @right_commutative V V +%R. Proof. exact: mulmAC. Qed.
+Lemma addrACA : @interchange V +%R +%R. Proof. exact: mulmACA. Qed.
+
+Lemma addKr : @left_loop V V -%R +%R.
+Proof. by move=> x y; rewrite addrA addNr add0r. Qed.
+Lemma addNKr : @rev_left_loop V V -%R +%R.
+Proof. by move=> x y; rewrite addrA addrN add0r. Qed.
+Lemma addrK : @right_loop V V -%R +%R.
+Proof. by move=> x y; rewrite -addrA addrN addr0. Qed.
+Lemma addrNK : @rev_right_loop V V -%R +%R.
+Proof. by move=> x y; rewrite -addrA addNr addr0. Qed.
+Definition subrK := addrNK.
+Lemma addrI : @right_injective V V V +%R.
+Proof. move=> x; exact: can_inj (addKr x). Qed.
+Lemma addIr : @left_injective V V V +%R.
+Proof. move=> y; exact: can_inj (addrK y). Qed.
+Lemma opprK : @involutive V -%R.
+Proof. by move=> x; apply: (@addIr (- x)); rewrite addNr addrN. Qed.
+Lemma oppr_inj : @injective V V -%R.
+Proof. exact: inv_inj opprK. Qed.
+Lemma oppr0 : -0 = 0 :> V.
+Proof. by rewrite -[-0]add0r subrr. Qed.
+Lemma oppr_eq0 x : (- x == 0) = (x == 0).
+Proof. by rewrite (inv_eq opprK) oppr0. Qed.
+
+Lemma subr0 x : x - 0 = x. Proof. by rewrite oppr0 addr0. Qed.
+Lemma sub0r x : 0 - x = - x. Proof. by rewrite add0r. Qed.
+
+Lemma opprD : {morph -%R: x y / x + y : V}.
+Proof.
+by move=> x y; apply: (@addrI (x + y)); rewrite addrA subrr addrAC addrK subrr.
+Qed.
+
+Lemma opprB x y : - (x - y) = y - x.
+Proof. by rewrite opprD addrC opprK. Qed.
+
+Lemma subr_eq x y z : (x - z == y) = (x == y + z).
+Proof. exact: can2_eq (subrK z) (addrK z) x y. Qed.
+
+Lemma subr_eq0 x y : (x - y == 0) = (x == y).
+Proof. by rewrite subr_eq add0r. Qed.
+
+Lemma addr_eq0 x y : (x + y == 0) = (x == - y).
+Proof. by rewrite -[x == _]subr_eq0 opprK. Qed.
+
+Lemma eqr_opp x y : (- x == - y) = (x == y).
+Proof. exact: can_eq opprK x y. Qed.
+
+Lemma eqr_oppLR x y : (- x == y) = (x == - y).
+Proof. exact: inv_eq opprK x y. Qed.
+
+Lemma mulr0n x : x *+ 0 = 0. Proof. by []. Qed.
+Lemma mulr1n x : x *+ 1 = x. Proof. by []. Qed.
+Lemma mulr2n x : x *+ 2 = x + x. Proof. by []. Qed.
+
+Lemma mulrS x n : x *+ n.+1 = x + x *+ n.
+Proof. by case: n => //=; rewrite addr0. Qed.
+
+Lemma mulrSr x n : x *+ n.+1 = x *+ n + x.
+Proof. by rewrite addrC mulrS. Qed.
+
+Lemma mulrb x (b : bool) : x *+ b = (if b then x else 0).
+Proof. by case: b. Qed.
+
+Lemma mul0rn n : 0 *+ n = 0 :> V.
+Proof. by elim: n => // n IHn; rewrite mulrS add0r. Qed.
+
+Lemma mulNrn x n : (- x) *+ n = x *- n.
+Proof. by elim: n => [|n IHn]; rewrite ?oppr0 // !mulrS opprD IHn. Qed.
+
+Lemma mulrnDl n : {morph (fun x => x *+ n) : x y / x + y}.
+Proof.
+move=> x y; elim: n => [|n IHn]; rewrite ?addr0 // !mulrS.
+by rewrite addrCA -!addrA -IHn -addrCA.
+Qed.
+
+Lemma mulrnDr x m n : x *+ (m + n) = x *+ m + x *+ n.
+Proof.
+elim: m => [|m IHm]; first by rewrite add0r.
+by rewrite !mulrS IHm addrA.
+Qed.
+
+Lemma mulrnBl n : {morph (fun x => x *+ n) : x y / x - y}.
+Proof.
+move=> x y; elim: n => [|n IHn]; rewrite ?subr0 // !mulrS -!addrA; congr(_ + _).
+by rewrite addrC IHn -!addrA opprD [_ - y]addrC.
+Qed.
+
+Lemma mulrnBr x m n : n <= m -> x *+ (m - n) = x *+ m - x *+ n.
+Proof.
+elim: m n => [|m IHm] [|n le_n_m]; rewrite ?subr0 // {}IHm //.
+by rewrite mulrSr mulrS opprD addrA addrK.
+Qed.
+
+Lemma mulrnA x m n : x *+ (m * n) = x *+ m *+ n.
+Proof.
+by rewrite mulnC; elim: n => //= n IHn; rewrite mulrS mulrnDr IHn.
+Qed.
+
+Lemma mulrnAC x m n : x *+ m *+ n = x *+ n *+ m.
+Proof. by rewrite -!mulrnA mulnC. Qed.
+
+Lemma sumrN I r P (F : I -> V) :
+ (\sum_(i <- r | P i) - F i = - (\sum_(i <- r | P i) F i)).
+Proof. by rewrite (big_morph _ opprD oppr0). Qed.
+
+Lemma sumrB I r (P : pred I) (F1 F2 : I -> V) :
+ \sum_(i <- r | P i) (F1 i - F2 i)
+ = \sum_(i <- r | P i) F1 i - \sum_(i <- r | P i) F2 i.
+Proof. by rewrite -sumrN -big_split /=. Qed.
+
+Lemma sumrMnl I r P (F : I -> V) n :
+ \sum_(i <- r | P i) F i *+ n = (\sum_(i <- r | P i) F i) *+ n.
+Proof. by rewrite (big_morph _ (mulrnDl n) (mul0rn _)). Qed.
+
+Lemma sumrMnr x I r P (F : I -> nat) :
+ \sum_(i <- r | P i) x *+ F i = x *+ (\sum_(i <- r | P i) F i).
+Proof. by rewrite (big_morph _ (mulrnDr x) (erefl _)). Qed.
+
+Lemma sumr_const (I : finType) (A : pred I) (x : V) :
+ \sum_(i in A) x = x *+ #|A|.
+Proof. by rewrite big_const -iteropE. Qed.
+
+Lemma telescope_sumr n m (f : nat -> V) : n <= m ->
+ \sum_(n <= k < m) (f k.+1 - f k) = f m - f n.
+Proof.
+rewrite leq_eqVlt => /predU1P[-> | ]; first by rewrite subrr big_geq.
+case: m => // m lenm; rewrite sumrB big_nat_recr // big_nat_recl //=.
+by rewrite addrC opprD addrA subrK addrC.
+Qed.
+
+Section ClosedPredicates.
+
+Variable S : predPredType V.
+
+Definition addr_closed := 0 \in S /\ {in S &, forall u v, u + v \in S}.
+Definition oppr_closed := {in S, forall u, - u \in S}.
+Definition subr_2closed := {in S &, forall u v, u - v \in S}.
+Definition zmod_closed := 0 \in S /\ subr_2closed.
+
+Lemma zmod_closedN : zmod_closed -> oppr_closed.
+Proof. by case=> S0 SB y Sy; rewrite -sub0r !SB. Qed.
+
+Lemma zmod_closedD : zmod_closed -> addr_closed.
+Proof.
+by case=> S0 SB; split=> // y z Sy Sz; rewrite -[z]opprK -[- z]sub0r !SB.
+Qed.
+
+End ClosedPredicates.
+
+End ZmoduleTheory.
+
+Implicit Arguments addrI [[V] x1 x2].
+Implicit Arguments addIr [[V] x1 x2].
+Implicit Arguments oppr_inj [[V] x1 x2].
+
+Module Ring.
+
+Record mixin_of (R : zmodType) : Type := Mixin {
+ one : R;
+ mul : R -> R -> R;
+ _ : associative mul;
+ _ : left_id one mul;
+ _ : right_id one mul;
+ _ : left_distributive mul +%R;
+ _ : right_distributive mul +%R;
+ _ : one != 0
+}.
+
+Definition EtaMixin R one mul mulA mul1x mulx1 mul_addl mul_addr nz1 :=
+ let _ := @Mixin R one mul mulA mul1x mulx1 mul_addl mul_addr nz1 in
+ @Mixin (Zmodule.Pack (Zmodule.class R) R) _ _
+ mulA mul1x mulx1 mul_addl mul_addr nz1.
+
+Section ClassDef.
+
+Record class_of (R : Type) : Type := Class {
+ base : Zmodule.class_of R;
+ mixin : mixin_of (Zmodule.Pack base R)
+}.
+Local Coercion base : class_of >-> Zmodule.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+
+Definition pack b0 (m0 : mixin_of (@Zmodule.Pack T b0 T)) :=
+ fun bT b & phant_id (Zmodule.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> Zmodule.class_of.
+Coercion mixin : class_of >-> mixin_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Notation ringType := type.
+Notation RingType T m := (@pack T _ m _ _ id _ id).
+Notation RingMixin := Mixin.
+Notation "[ 'ringType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'ringType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'ringType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'ringType' 'of' T ]") : form_scope.
+End Exports.
+
+End Ring.
+Import Ring.Exports.
+
+Definition one (R : ringType) : R := Ring.one (Ring.class R).
+Definition mul (R : ringType) : R -> R -> R := Ring.mul (Ring.class R).
+Definition exp R x n := nosimpl iterop _ n (@mul R) x (one R).
+Notation sign R b := (exp (- one R) (nat_of_bool b)) (only parsing).
+Definition comm R x y := @mul R x y = mul y x.
+Definition lreg R x := injective (@mul R x).
+Definition rreg R x := injective ((@mul R)^~ x).
+
+Local Notation "1" := (one _) : ring_scope.
+Local Notation "- 1" := (- (1)) : ring_scope.
+Local Notation "n %:R" := (1 *+ n) : ring_scope.
+Local Notation "*%R" := (@mul _).
+Local Notation "x * y" := (mul x y) : ring_scope.
+Local Notation "x ^+ n" := (exp x n) : ring_scope.
+
+Local Notation "\prod_ ( i <- r | P ) F" := (\big[*%R/1]_(i <- r | P) F).
+Local Notation "\prod_ ( i | P ) F" := (\big[*%R/1]_(i | P) F).
+Local Notation "\prod_ ( i 'in' A ) F" := (\big[*%R/1]_(i in A) F).
+Local Notation "\prod_ ( m <= i < n ) F" := (\big[*%R/1%R]_(m <= i < n) F%R).
+
+(* The ``field'' characteristic; the definition, and many of the theorems, *)
+(* has to apply to rings as well; indeed, we need the Frobenius automorphism *)
+(* results for a non commutative ring in the proof of Gorenstein 2.6.3. *)
+Definition char (R : Ring.type) of phant R : nat_pred :=
+ [pred p | prime p & p%:R == 0 :> R].
+
+Local Notation "[ 'char' R ]" := (char (Phant R)) : ring_scope.
+
+(* Converse ring tag. *)
+Definition converse R : Type := R.
+Local Notation "R ^c" := (converse R) (at level 2, format "R ^c") : type_scope.
+
+Section RingTheory.
+
+Variable R : ringType.
+Implicit Types x y : R.
+
+Lemma mulrA : @associative R *%R. Proof. by case R => T [? []]. Qed.
+Lemma mul1r : @left_id R R 1 *%R. Proof. by case R => T [? []]. Qed.
+Lemma mulr1 : @right_id R R 1 *%R. Proof. by case R => T [? []]. Qed.
+Lemma mulrDl : @left_distributive R R *%R +%R.
+Proof. by case R => T [? []]. Qed.
+Lemma mulrDr : @right_distributive R R *%R +%R.
+Proof. by case R => T [? []]. Qed.
+Lemma oner_neq0 : 1 != 0 :> R. Proof. by case R => T [? []]. Qed.
+Lemma oner_eq0 : (1 == 0 :> R) = false. Proof. exact: negbTE oner_neq0. Qed.
+
+Lemma mul0r : @left_zero R R 0 *%R.
+Proof.
+by move=> x; apply: (addIr (1 * x)); rewrite -mulrDl !add0r mul1r.
+Qed.
+Lemma mulr0 : @right_zero R R 0 *%R.
+Proof.
+by move=> x; apply: (addIr (x * 1)); rewrite -mulrDr !add0r mulr1.
+Qed.
+Lemma mulrN x y : x * (- y) = - (x * y).
+Proof. by apply: (addrI (x * y)); rewrite -mulrDr !subrr mulr0. Qed.
+Lemma mulNr x y : (- x) * y = - (x * y).
+Proof. by apply: (addrI (x * y)); rewrite -mulrDl !subrr mul0r. Qed.
+Lemma mulrNN x y : (- x) * (- y) = x * y.
+Proof. by rewrite mulrN mulNr opprK. Qed.
+Lemma mulN1r x : -1 * x = - x.
+Proof. by rewrite mulNr mul1r. Qed.
+Lemma mulrN1 x : x * -1 = - x.
+Proof. by rewrite mulrN mulr1. Qed.
+
+Canonical mul_monoid := Monoid.Law mulrA mul1r mulr1.
+Canonical muloid := Monoid.MulLaw mul0r mulr0.
+Canonical addoid := Monoid.AddLaw mulrDl mulrDr.
+
+Lemma mulr_suml I r P (F : I -> R) x :
+ (\sum_(i <- r | P i) F i) * x = \sum_(i <- r | P i) F i * x.
+Proof. exact: big_distrl. Qed.
+
+Lemma mulr_sumr I r P (F : I -> R) x :
+ x * (\sum_(i <- r | P i) F i) = \sum_(i <- r | P i) x * F i.
+Proof. exact: big_distrr. Qed.
+
+Lemma mulrBl x y z : (y - z) * x = y * x - z * x.
+Proof. by rewrite mulrDl mulNr. Qed.
+
+Lemma mulrBr x y z : x * (y - z) = x * y - x * z.
+Proof. by rewrite mulrDr mulrN. Qed.
+
+Lemma mulrnAl x y n : (x *+ n) * y = (x * y) *+ n.
+Proof. by elim: n => [|n IHn]; rewrite ?mul0r // !mulrS mulrDl IHn. Qed.
+
+Lemma mulrnAr x y n : x * (y *+ n) = (x * y) *+ n.
+Proof. by elim: n => [|n IHn]; rewrite ?mulr0 // !mulrS mulrDr IHn. Qed.
+
+Lemma mulr_natl x n : n%:R * x = x *+ n.
+Proof. by rewrite mulrnAl mul1r. Qed.
+
+Lemma mulr_natr x n : x * n%:R = x *+ n.
+Proof. by rewrite mulrnAr mulr1. Qed.
+
+Lemma natrD m n : (m + n)%:R = m%:R + n%:R :> R.
+Proof. exact: mulrnDr. Qed.
+
+Lemma natrB m n : n <= m -> (m - n)%:R = m%:R - n%:R :> R.
+Proof. exact: mulrnBr. Qed.
+
+Definition natr_sum := big_morph (natmul 1) natrD (mulr0n 1).
+
+Lemma natrM m n : (m * n)%:R = m%:R * n%:R :> R.
+Proof. by rewrite mulrnA -mulr_natr. Qed.
+
+Lemma expr0 x : x ^+ 0 = 1. Proof. by []. Qed.
+Lemma expr1 x : x ^+ 1 = x. Proof. by []. Qed.
+Lemma expr2 x : x ^+ 2 = x * x. Proof. by []. Qed.
+
+Lemma exprS x n : x ^+ n.+1 = x * x ^+ n.
+Proof. by case: n => //; rewrite mulr1. Qed.
+
+Lemma expr0n n : 0 ^+ n = (n == 0%N)%:R :> R.
+Proof. by case: n => // n; rewrite exprS mul0r. Qed.
+
+Lemma expr1n n : 1 ^+ n = 1 :> R.
+Proof. by elim: n => // n IHn; rewrite exprS mul1r. Qed.
+
+Lemma exprD x m n : x ^+ (m + n) = x ^+ m * x ^+ n.
+Proof. by elim: m => [|m IHm]; rewrite ?mul1r // !exprS -mulrA -IHm. Qed.
+
+Lemma exprSr x n : x ^+ n.+1 = x ^+ n * x.
+Proof. by rewrite -addn1 exprD expr1. Qed.
+
+Lemma commr_sym x y : comm x y -> comm y x. Proof. by []. Qed.
+Lemma commr_refl x : comm x x. Proof. by []. Qed.
+
+Lemma commr0 x : comm x 0.
+Proof. by rewrite /comm mulr0 mul0r. Qed.
+
+Lemma commr1 x : comm x 1.
+Proof. by rewrite /comm mulr1 mul1r. Qed.
+
+Lemma commrN x y : comm x y -> comm x (- y).
+Proof. by move=> com_xy; rewrite /comm mulrN com_xy mulNr. Qed.
+
+Lemma commrN1 x : comm x (-1).
+Proof. apply: commrN; exact: commr1. Qed.
+
+Lemma commrD x y z : comm x y -> comm x z -> comm x (y + z).
+Proof. by rewrite /comm mulrDl mulrDr => -> ->. Qed.
+
+Lemma commrMn x y n : comm x y -> comm x (y *+ n).
+Proof.
+rewrite /comm => com_xy.
+by elim: n => [|n IHn]; rewrite ?commr0 // mulrS commrD.
+Qed.
+
+Lemma commrM x y z : comm x y -> comm x z -> comm x (y * z).
+Proof. by move=> com_xy; rewrite /comm mulrA com_xy -!mulrA => ->. Qed.
+
+Lemma commr_nat x n : comm x n%:R.
+Proof. by apply: commrMn; exact: commr1. Qed.
+
+Lemma commrX x y n : comm x y -> comm x (y ^+ n).
+Proof.
+rewrite /comm => com_xy.
+by elim: n => [|n IHn]; rewrite ?commr1 // exprS commrM.
+Qed.
+
+Lemma exprMn_comm x y n : comm x y -> (x * y) ^+ n = x ^+ n * y ^+ n.
+Proof.
+move=> com_xy; elim: n => /= [|n IHn]; first by rewrite mulr1.
+by rewrite !exprS IHn !mulrA; congr (_ * _); rewrite -!mulrA -commrX.
+Qed.
+
+Lemma commr_sign x n : comm x ((-1) ^+ n).
+Proof. exact: (commrX n (commrN1 x)). Qed.
+
+Lemma exprMn_n x m n : (x *+ m) ^+ n = x ^+ n *+ (m ^ n) :> R.
+Proof.
+elim: n => [|n IHn]; first by rewrite mulr1n.
+rewrite exprS IHn -mulr_natr -mulrA -commr_nat mulr_natr -mulrnA -expnSr.
+by rewrite -mulr_natr mulrA -exprS mulr_natr.
+Qed.
+
+Lemma exprM x m n : x ^+ (m * n) = x ^+ m ^+ n.
+Proof.
+elim: m => [|m IHm]; first by rewrite expr1n.
+by rewrite mulSn exprD IHm exprS exprMn_comm //; exact: commrX.
+Qed.
+
+Lemma exprAC x m n : (x ^+ m) ^+ n = (x ^+ n) ^+ m.
+Proof. by rewrite -!exprM mulnC. Qed.
+
+Lemma expr_mod n x i : x ^+ n = 1 -> x ^+ (i %% n) = x ^+ i.
+Proof.
+move=> xn1; rewrite {2}(divn_eq i n) exprD mulnC exprM xn1.
+by rewrite expr1n mul1r.
+Qed.
+
+Lemma expr_dvd n x i : x ^+ n = 1 -> n %| i -> x ^+ i = 1.
+Proof.
+by move=> xn1 dvd_n_i; rewrite -(expr_mod i xn1) (eqnP dvd_n_i).
+Qed.
+
+Lemma natrX n k : (n ^ k)%:R = n%:R ^+ k :> R.
+Proof. by rewrite exprMn_n expr1n. Qed.
+
+Lemma signr_odd n : (-1) ^+ (odd n) = (-1) ^+ n :> R.
+Proof.
+elim: n => //= n IHn; rewrite exprS -{}IHn.
+by case/odd: n; rewrite !mulN1r ?opprK.
+Qed.
+
+Lemma signr_eq0 n : ((-1) ^+ n == 0 :> R) = false.
+Proof. by rewrite -signr_odd; case: odd; rewrite ?oppr_eq0 oner_eq0. Qed.
+
+Lemma mulr_sign (b : bool) x : (-1) ^+ b * x = (if b then - x else x).
+Proof. by case: b; rewrite ?mulNr mul1r. Qed.
+
+Lemma signr_addb b1 b2 : (-1) ^+ (b1 (+) b2) = (-1) ^+ b1 * (-1) ^+ b2 :> R.
+Proof. by rewrite mulr_sign; case: b1 b2 => [] []; rewrite ?opprK. Qed.
+
+Lemma signrE (b : bool) : (-1) ^+ b = 1 - b.*2%:R :> R.
+Proof. by case: b; rewrite ?subr0 // opprD addNKr. Qed.
+
+Lemma signrN b : (-1) ^+ (~~ b) = - (-1) ^+ b :> R.
+Proof. by case: b; rewrite ?opprK. Qed.
+
+Lemma mulr_signM (b1 b2 : bool) x1 x2 :
+ ((-1) ^+ b1 * x1) * ((-1) ^+ b2 * x2) = (-1) ^+ (b1 (+) b2) * (x1 * x2).
+Proof.
+by rewrite signr_addb -!mulrA; congr (_ * _); rewrite !mulrA commr_sign.
+Qed.
+
+Lemma exprNn x n : (- x) ^+ n = (-1) ^+ n * x ^+ n :> R.
+Proof. by rewrite -mulN1r exprMn_comm // /comm mulN1r mulrN mulr1. Qed.
+
+Lemma sqrrN x : (- x) ^+ 2 = x ^+ 2.
+Proof. exact: mulrNN. Qed.
+
+Lemma sqrr_sign n : ((-1) ^+ n) ^+ 2 = 1 :> R.
+Proof. by rewrite exprAC sqrrN !expr1n. Qed.
+
+Lemma signrMK n : @involutive R ( *%R ((-1) ^+ n)).
+Proof. by move=> x; rewrite mulrA -expr2 sqrr_sign mul1r. Qed.
+
+Lemma mulrI_eq0 x y : lreg x -> (x * y == 0) = (y == 0).
+Proof. by move=> reg_x; rewrite -{1}(mulr0 x) (inj_eq reg_x). Qed.
+
+Lemma lreg_neq0 x : lreg x -> x != 0.
+Proof. by move=> reg_x; rewrite -[x]mulr1 mulrI_eq0 ?oner_eq0. Qed.
+
+Lemma mulrI0_lreg x : (forall y, x * y = 0 -> y = 0) -> lreg x.
+Proof.
+move=> reg_x y z eq_xy_xz; apply/eqP; rewrite -subr_eq0 [y - z]reg_x //.
+by rewrite mulrBr eq_xy_xz subrr.
+Qed.
+
+Lemma lregN x : lreg x -> lreg (- x).
+Proof. by move=> reg_x y z; rewrite !mulNr => /oppr_inj/reg_x. Qed.
+
+Lemma lreg1 : lreg (1 : R).
+Proof. by move=> x y; rewrite !mul1r. Qed.
+
+Lemma lregM x y : lreg x -> lreg y -> lreg (x * y).
+Proof. by move=> reg_x reg_y z t; rewrite -!mulrA => /reg_x/reg_y. Qed.
+
+Lemma lregX x n : lreg x -> lreg (x ^+ n).
+Proof.
+by move=> reg_x; elim: n => [|n]; [exact: lreg1 | rewrite exprS; exact: lregM].
+Qed.
+
+Lemma lreg_sign n : lreg ((-1) ^+ n : R).
+Proof. by apply: lregX; apply: lregN; apply: lreg1. Qed.
+
+Lemma prodr_const (I : finType) (A : pred I) (x : R) :
+ \prod_(i in A) x = x ^+ #|A|.
+Proof. by rewrite big_const -iteropE. Qed.
+
+Lemma prodrXr x I r P (F : I -> nat) :
+ \prod_(i <- r | P i) x ^+ F i = x ^+ (\sum_(i <- r | P i) F i).
+Proof. by rewrite (big_morph _ (exprD _) (erefl _)). Qed.
+
+Lemma prodrN (I : finType) (A : pred I) (F : I -> R) :
+ \prod_(i in A) - F i = (- 1) ^+ #|A| * \prod_(i in A) F i.
+Proof.
+rewrite -sum1_card; elim/big_rec3: _ => [|i x n _ _ ->]; first by rewrite mulr1.
+by rewrite exprS !mulrA mulN1r !mulNr commrX //; apply: commrN1.
+Qed.
+
+Lemma prodrMn n (I : finType) (A : pred I) (F : I -> R) :
+ \prod_(i in A) (F i *+ n) = \prod_(i in A) F i *+ n ^ #|A|.
+Proof.
+rewrite -sum1_card; elim/big_rec3: _ => // i x m _ _ ->.
+by rewrite mulrnAr mulrnAl expnS mulrnA.
+Qed.
+
+Lemma natr_prod I r P (F : I -> nat) :
+ (\prod_(i <- r | P i) F i)%:R = \prod_(i <- r | P i) (F i)%:R :> R.
+Proof. exact: (big_morph _ natrM). Qed.
+
+Lemma exprDn_comm x y n (cxy : comm x y) :
+ (x + y) ^+ n = \sum_(i < n.+1) (x ^+ (n - i) * y ^+ i) *+ 'C(n, i).
+Proof.
+elim: n => [|n IHn]; rewrite big_ord_recl mulr1 ?big_ord0 ?addr0 //=.
+rewrite exprS {}IHn /= mulrDl !big_distrr /= big_ord_recl mulr1 subn0.
+rewrite !big_ord_recr /= !binn !subnn !mul1r !subn0 bin0 !exprS -addrA.
+congr (_ + _); rewrite addrA -big_split /=; congr (_ + _).
+apply: eq_bigr => i _; rewrite !mulrnAr !mulrA -exprS -subSn ?(valP i) //.
+by rewrite subSS (commrX _ (commr_sym cxy)) -mulrA -exprS -mulrnDr.
+Qed.
+
+Lemma exprBn_comm x y n (cxy : comm x y) :
+ (x - y) ^+ n =
+ \sum_(i < n.+1) ((-1) ^+ i * x ^+ (n - i) * y ^+ i) *+ 'C(n, i).
+Proof.
+rewrite exprDn_comm; last exact: commrN.
+by apply: eq_bigr => i _; congr (_ *+ _); rewrite -commr_sign -mulrA -exprNn.
+Qed.
+
+Lemma subrXX_comm x y n (cxy : comm x y) :
+ x ^+ n - y ^+ n = (x - y) * (\sum_(i < n) x ^+ (n.-1 - i) * y ^+ i).
+Proof.
+case: n => [|n]; first by rewrite big_ord0 mulr0 subrr.
+rewrite mulrBl !big_distrr big_ord_recl big_ord_recr /= subnn mulr1 mul1r.
+rewrite subn0 -!exprS opprD -!addrA; congr (_ + _); rewrite addrA -sumrB.
+rewrite big1 ?add0r // => i _; rewrite !mulrA -exprS -subSn ?(valP i) //.
+by rewrite subSS (commrX _ (commr_sym cxy)) -mulrA -exprS subrr.
+Qed.
+
+Lemma exprD1n x n : (x + 1) ^+ n = \sum_(i < n.+1) x ^+ i *+ 'C(n, i).
+Proof.
+rewrite addrC (exprDn_comm n (commr_sym (commr1 x))).
+by apply: eq_bigr => i _; rewrite expr1n mul1r.
+Qed.
+
+Lemma subrX1 x n : x ^+ n - 1 = (x - 1) * (\sum_(i < n) x ^+ i).
+Proof.
+rewrite -!(opprB 1) mulNr -{1}(expr1n n).
+rewrite (subrXX_comm _ (commr_sym (commr1 x))); congr (- (_ * _)).
+by apply: eq_bigr => i _; rewrite expr1n mul1r.
+Qed.
+
+Lemma sqrrD1 x : (x + 1) ^+ 2 = x ^+ 2 + x *+ 2 + 1.
+Proof.
+rewrite exprD1n !big_ord_recr big_ord0 /= add0r.
+by rewrite addrC addrA addrAC.
+Qed.
+
+Lemma sqrrB1 x : (x - 1) ^+ 2 = x ^+ 2 - x *+ 2 + 1.
+Proof. by rewrite -sqrrN opprB addrC sqrrD1 sqrrN mulNrn. Qed.
+
+Lemma subr_sqr_1 x : x ^+ 2 - 1 = (x - 1) * (x + 1).
+Proof. by rewrite subrX1 !big_ord_recr big_ord0 /= addrAC add0r. Qed.
+
+Definition Frobenius_aut p of p \in [char R] := fun x => x ^+ p.
+
+Section FrobeniusAutomorphism.
+
+Variable p : nat.
+Hypothesis charFp : p \in [char R].
+
+Lemma charf0 : p%:R = 0 :> R. Proof. by apply/eqP; case/andP: charFp. Qed.
+Lemma charf_prime : prime p. Proof. by case/andP: charFp. Qed.
+Hint Resolve charf_prime.
+
+Lemma mulrn_char x : x *+ p = 0. Proof. by rewrite -mulr_natl charf0 mul0r. Qed.
+
+Lemma natr_mod_char n : (n %% p)%:R = n%:R :> R.
+Proof. by rewrite {2}(divn_eq n p) natrD mulrnA mulrn_char add0r. Qed.
+
+Lemma dvdn_charf n : (p %| n)%N = (n%:R == 0 :> R).
+Proof.
+apply/idP/eqP=> [/dvdnP[n' ->]|n0]; first by rewrite natrM charf0 mulr0.
+apply/idPn; rewrite -prime_coprime // => /eqnP pn1.
+have [a _ /dvdnP[b]] := Bezoutl n (prime_gt0 charf_prime).
+move/(congr1 (fun m => m%:R : R))/eqP.
+by rewrite natrD !natrM charf0 n0 !mulr0 pn1 addr0 oner_eq0.
+Qed.
+
+Lemma charf_eq : [char R] =i (p : nat_pred).
+Proof.
+move=> q; apply/andP/eqP=> [[q_pr q0] | ->]; last by rewrite charf0.
+by apply/eqP; rewrite eq_sym -dvdn_prime2 // dvdn_charf.
+Qed.
+
+Lemma bin_lt_charf_0 k : 0 < k < p -> 'C(p, k)%:R = 0 :> R.
+Proof. by move=> lt0kp; apply/eqP; rewrite -dvdn_charf prime_dvd_bin. Qed.
+
+Local Notation "x ^f" := (Frobenius_aut charFp x).
+
+Lemma Frobenius_autE x : x^f = x ^+ p. Proof. by []. Qed.
+Local Notation fE := Frobenius_autE.
+
+Lemma Frobenius_aut0 : 0^f = 0.
+Proof. by rewrite fE -(prednK (prime_gt0 charf_prime)) exprS mul0r. Qed.
+
+Lemma Frobenius_aut1 : 1^f = 1.
+Proof. by rewrite fE expr1n. Qed.
+
+Lemma Frobenius_autD_comm x y (cxy : comm x y) : (x + y)^f = x^f + y^f.
+Proof.
+have defp := prednK (prime_gt0 charf_prime).
+rewrite !fE exprDn_comm // big_ord_recr subnn -defp big_ord_recl /= defp.
+rewrite subn0 mulr1 mul1r bin0 binn big1 ?addr0 // => i _.
+by rewrite -mulr_natl bin_lt_charf_0 ?mul0r //= -{2}defp ltnS (valP i).
+Qed.
+
+Lemma Frobenius_autMn x n : (x *+ n)^f = x^f *+ n.
+Proof.
+elim: n => [|n IHn]; first exact: Frobenius_aut0.
+rewrite !mulrS Frobenius_autD_comm ?IHn //; exact: commrMn.
+Qed.
+
+Lemma Frobenius_aut_nat n : (n%:R)^f = n%:R.
+Proof. by rewrite Frobenius_autMn Frobenius_aut1. Qed.
+
+Lemma Frobenius_autM_comm x y : comm x y -> (x * y)^f = x^f * y^f.
+Proof. by exact: exprMn_comm. Qed.
+
+Lemma Frobenius_autX x n : (x ^+ n)^f = x^f ^+ n.
+Proof. by rewrite !fE -!exprM mulnC. Qed.
+
+Lemma Frobenius_autN x : (- x)^f = - x^f.
+Proof.
+apply/eqP; rewrite -subr_eq0 opprK addrC.
+by rewrite -(Frobenius_autD_comm (commrN _)) // subrr Frobenius_aut0.
+Qed.
+
+Lemma Frobenius_autB_comm x y : comm x y -> (x - y)^f = x^f - y^f.
+Proof.
+by move/commrN/Frobenius_autD_comm->; rewrite Frobenius_autN.
+Qed.
+
+End FrobeniusAutomorphism.
+
+Lemma exprNn_char x n : [char R].-nat n -> (- x) ^+ n = - (x ^+ n).
+Proof.
+pose p := pdiv n; have [|n_gt1 charRn] := leqP n 1; first by case: (n) => [|[]].
+have charRp: p \in [char R] by rewrite (pnatPpi charRn) // pi_pdiv.
+have /p_natP[e ->]: p.-nat n by rewrite -(eq_pnat _ (charf_eq charRp)).
+elim: e => // e IHe; rewrite expnSr !exprM {}IHe.
+by rewrite -Frobenius_autE Frobenius_autN.
+Qed.
+
+Section Char2.
+
+Hypothesis charR2 : 2 \in [char R].
+
+Lemma addrr_char2 x : x + x = 0. Proof. by rewrite -mulr2n mulrn_char. Qed.
+
+Lemma oppr_char2 x : - x = x.
+Proof. by apply/esym/eqP; rewrite -addr_eq0 addrr_char2. Qed.
+
+Lemma subr_char2 x y : x - y = x + y. Proof. by rewrite oppr_char2. Qed.
+
+Lemma addrK_char2 x : involutive (+%R^~ x).
+Proof. by move=> y; rewrite /= -subr_char2 addrK. Qed.
+
+Lemma addKr_char2 x : involutive (+%R x).
+Proof. by move=> y; rewrite -{1}[x]oppr_char2 addKr. Qed.
+
+End Char2.
+
+Canonical converse_eqType := [eqType of R^c].
+Canonical converse_choiceType := [choiceType of R^c].
+Canonical converse_zmodType := [zmodType of R^c].
+
+Definition converse_ringMixin :=
+ let mul' x y := y * x in
+ let mulrA' x y z := esym (mulrA z y x) in
+ let mulrDl' x y z := mulrDr z x y in
+ let mulrDr' x y z := mulrDl y z x in
+ @Ring.Mixin converse_zmodType
+ 1 mul' mulrA' mulr1 mul1r mulrDl' mulrDr' oner_neq0.
+Canonical converse_ringType := RingType R^c converse_ringMixin.
+
+Section ClosedPredicates.
+
+Variable S : predPredType R.
+
+Definition mulr_2closed := {in S &, forall u v, u * v \in S}.
+Definition mulr_closed := 1 \in S /\ mulr_2closed.
+Definition smulr_closed := -1 \in S /\ mulr_2closed.
+Definition semiring_closed := addr_closed S /\ mulr_closed.
+Definition subring_closed := [/\ 1 \in S, subr_2closed S & mulr_2closed].
+
+Lemma smulr_closedM : smulr_closed -> mulr_closed.
+Proof. by case=> SN1 SM; split=> //; rewrite -[1]mulr1 -mulrNN SM. Qed.
+
+Lemma smulr_closedN : smulr_closed -> oppr_closed S.
+Proof. by case=> SN1 SM x Sx; rewrite -mulN1r SM. Qed.
+
+Lemma semiring_closedD : semiring_closed -> addr_closed S. Proof. by case. Qed.
+
+Lemma semiring_closedM : semiring_closed -> mulr_closed. Proof. by case. Qed.
+
+Lemma subring_closedB : subring_closed -> zmod_closed S.
+Proof. by case=> S1 SB _; split; rewrite // -(subrr 1) SB. Qed.
+
+Lemma subring_closedM : subring_closed -> smulr_closed.
+Proof.
+by case=> S1 SB SM; split; rewrite ?(zmod_closedN (subring_closedB _)).
+Qed.
+
+Lemma subring_closed_semi : subring_closed -> semiring_closed.
+Proof.
+by move=> ringS; split; [apply/zmod_closedD/subring_closedB | case: ringS].
+Qed.
+
+End ClosedPredicates.
+
+End RingTheory.
+
+Section RightRegular.
+
+Variable R : ringType.
+Implicit Types x y : R.
+Let Rc := converse_ringType R.
+
+Lemma mulIr_eq0 x y : rreg x -> (y * x == 0) = (y == 0).
+Proof. exact: (@mulrI_eq0 Rc). Qed.
+
+Lemma mulIr0_rreg x : (forall y, y * x = 0 -> y = 0) -> rreg x.
+Proof. exact: (@mulrI0_lreg Rc). Qed.
+
+Lemma rreg_neq0 x : rreg x -> x != 0.
+Proof. exact: (@lreg_neq0 Rc). Qed.
+
+Lemma rregN x : rreg x -> rreg (- x).
+Proof. exact: (@lregN Rc). Qed.
+
+Lemma rreg1 : rreg (1 : R).
+Proof. exact: (@lreg1 Rc). Qed.
+
+Lemma rregM x y : rreg x -> rreg y -> rreg (x * y).
+Proof. by move=> reg_x reg_y; exact: (@lregM Rc). Qed.
+
+Lemma revrX x n : (x : Rc) ^+ n = (x : R) ^+ n.
+Proof. by elim: n => // n IHn; rewrite exprS exprSr IHn. Qed.
+
+Lemma rregX x n : rreg x -> rreg (x ^+ n).
+Proof. by move/(@lregX Rc x n); rewrite revrX. Qed.
+
+End RightRegular.
+
+Module Lmodule.
+
+Structure mixin_of (R : ringType) (V : zmodType) : Type := Mixin {
+ scale : R -> V -> V;
+ _ : forall a b v, scale a (scale b v) = scale (a * b) v;
+ _ : left_id 1 scale;
+ _ : right_distributive scale +%R;
+ _ : forall v, {morph scale^~ v: a b / a + b}
+}.
+
+Section ClassDef.
+
+Variable R : ringType.
+
+Structure class_of V := Class {
+ base : Zmodule.class_of V;
+ mixin : mixin_of R (Zmodule.Pack base V)
+}.
+Local Coercion base : class_of >-> Zmodule.class_of.
+
+Structure type (phR : phant R) := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variable (phR : phant R) (T : Type) (cT : type phR).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack phR T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+
+Definition pack b0 (m0 : mixin_of R (@Zmodule.Pack T b0 T)) :=
+ fun bT b & phant_id (Zmodule.class bT) b =>
+ fun m & phant_id m0 m => Pack phR (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Import Exports.
+Coercion base : class_of >-> Zmodule.class_of.
+Coercion mixin : class_of >-> mixin_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Notation lmodType R := (type (Phant R)).
+Notation LmodType R T m := (@pack _ (Phant R) T _ m _ _ id _ id).
+Notation LmodMixin := Mixin.
+Notation "[ 'lmodType' R 'of' T 'for' cT ]" := (@clone _ (Phant R) T cT _ idfun)
+ (at level 0, format "[ 'lmodType' R 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'lmodType' R 'of' T ]" := (@clone _ (Phant R) T _ _ id)
+ (at level 0, format "[ 'lmodType' R 'of' T ]") : form_scope.
+End Exports.
+
+End Lmodule.
+Import Lmodule.Exports.
+
+Definition scale (R : ringType) (V : lmodType R) :=
+ Lmodule.scale (Lmodule.class V).
+
+Local Notation "*:%R" := (@scale _ _).
+Local Notation "a *: v" := (scale a v) : ring_scope.
+
+Section LmoduleTheory.
+
+Variables (R : ringType) (V : lmodType R).
+Implicit Types (a b c : R) (u v : V).
+
+Local Notation "*:%R" := (@scale R V).
+
+Lemma scalerA a b v : a *: (b *: v) = a * b *: v.
+Proof. by case: V v => ? [] ? []. Qed.
+
+Lemma scale1r : @left_id R V 1 *:%R.
+Proof. by case: V => ? [] ? []. Qed.
+
+Lemma scalerDr a : {morph *:%R a : u v / u + v}.
+Proof. by case: V a => ? [] ? []. Qed.
+
+Lemma scalerDl v : {morph *:%R^~ v : a b / a + b}.
+Proof. by case: V v => ? [] ? []. Qed.
+
+Lemma scale0r v : 0 *: v = 0.
+Proof. by apply: (addIr (1 *: v)); rewrite -scalerDl !add0r. Qed.
+
+Lemma scaler0 a : a *: 0 = 0 :> V.
+Proof. by rewrite -{1}(scale0r 0) scalerA mulr0 scale0r. Qed.
+
+Lemma scaleNr a v : - a *: v = - (a *: v).
+Proof. by apply: (addIr (a *: v)); rewrite -scalerDl !addNr scale0r. Qed.
+
+Lemma scaleN1r v : (- 1) *: v = - v.
+Proof. by rewrite scaleNr scale1r. Qed.
+
+Lemma scalerN a v : a *: (- v) = - (a *: v).
+Proof. by apply: (addIr (a *: v)); rewrite -scalerDr !addNr scaler0. Qed.
+
+Lemma scalerBl a b v : (a - b) *: v = a *: v - b *: v.
+Proof. by rewrite scalerDl scaleNr. Qed.
+
+Lemma scalerBr a u v : a *: (u - v) = a *: u - a *: v.
+Proof. by rewrite scalerDr scalerN. Qed.
+
+Lemma scaler_nat n v : n%:R *: v = v *+ n.
+Proof.
+elim: n => /= [|n ]; first by rewrite scale0r.
+by rewrite !mulrS scalerDl ?scale1r => ->.
+Qed.
+
+Lemma scaler_sign (b : bool) v: (-1) ^+ b *: v = (if b then - v else v).
+Proof. by case: b; rewrite ?scaleNr scale1r. Qed.
+
+Lemma signrZK n : @involutive V ( *:%R ((-1) ^+ n)).
+Proof. by move=> u; rewrite scalerA -expr2 sqrr_sign scale1r. Qed.
+
+Lemma scalerMnl a v n : a *: v *+ n = (a *+ n) *: v.
+Proof.
+elim: n => [|n IHn]; first by rewrite !mulr0n scale0r.
+by rewrite !mulrSr IHn scalerDl.
+Qed.
+
+Lemma scalerMnr a v n : a *: v *+ n = a *: (v *+ n).
+Proof.
+elim: n => [|n IHn]; first by rewrite !mulr0n scaler0.
+by rewrite !mulrSr IHn scalerDr.
+Qed.
+
+Lemma scaler_suml v I r (P : pred I) F :
+ (\sum_(i <- r | P i) F i) *: v = \sum_(i <- r | P i) F i *: v.
+Proof. exact: (big_morph _ (scalerDl v) (scale0r v)). Qed.
+
+Lemma scaler_sumr a I r (P : pred I) (F : I -> V) :
+ a *: (\sum_(i <- r | P i) F i) = \sum_(i <- r | P i) a *: F i.
+Proof. exact: big_endo (scalerDr a) (scaler0 a) I r P F. Qed.
+
+Section ClosedPredicates.
+
+Variable S : predPredType V.
+
+Definition scaler_closed := forall a, {in S, forall v, a *: v \in S}.
+Definition linear_closed := forall a, {in S &, forall u v, a *: u + v \in S}.
+Definition submod_closed := 0 \in S /\ linear_closed.
+
+Lemma linear_closedB : linear_closed -> subr_2closed S.
+Proof. by move=> Slin u v Su Sv; rewrite addrC -scaleN1r Slin. Qed.
+
+Lemma submod_closedB : submod_closed -> zmod_closed S.
+Proof. by case=> S0 /linear_closedB. Qed.
+
+Lemma submod_closedZ : submod_closed -> scaler_closed.
+Proof. by case=> S0 Slin a v Sv; rewrite -[a *: v]addr0 Slin. Qed.
+
+End ClosedPredicates.
+
+End LmoduleTheory.
+
+Module Lalgebra.
+
+Definition axiom (R : ringType) (V : lmodType R) (mul : V -> V -> V) :=
+ forall a u v, a *: mul u v = mul (a *: u) v.
+
+Section ClassDef.
+
+Variable R : ringType.
+
+Record class_of (T : Type) : Type := Class {
+ base : Ring.class_of T;
+ mixin : Lmodule.mixin_of R (Zmodule.Pack base T);
+ ext : @axiom R (Lmodule.Pack _ (Lmodule.Class mixin) T) (Ring.mul base)
+}.
+Definition base2 R m := Lmodule.Class (@mixin R m).
+Local Coercion base : class_of >-> Ring.class_of.
+Local Coercion base2 : class_of >-> Lmodule.class_of.
+
+Structure type (phR : phant R) := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variable (phR : phant R) (T : Type) (cT : type phR).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack phR T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack T b0 mul0 (axT : @axiom R (@Lmodule.Pack R _ T b0 T) mul0) :=
+ fun bT b & phant_id (Ring.class bT) (b : Ring.class_of T) =>
+ fun mT m & phant_id (@Lmodule.class R phR mT) (@Lmodule.Class R T b m) =>
+ fun ax & phant_id axT ax =>
+ Pack (Phant R) (@Class T b m ax) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+Definition lmodType := @Lmodule.Pack R phR cT xclass xT.
+Definition lmod_ringType := @Lmodule.Pack R phR ringType xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> Ring.class_of.
+Coercion base2 : class_of >-> Lmodule.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Coercion lmodType : type >-> Lmodule.type.
+Canonical lmodType.
+Canonical lmod_ringType.
+Notation lalgType R := (type (Phant R)).
+Notation LalgType R T a := (@pack _ (Phant R) T _ _ a _ _ id _ _ id _ id).
+Notation "[ 'lalgType' R 'of' T 'for' cT ]" := (@clone _ (Phant R) T cT _ idfun)
+ (at level 0, format "[ 'lalgType' R 'of' T 'for' cT ]")
+ : form_scope.
+Notation "[ 'lalgType' R 'of' T ]" := (@clone _ (Phant R) T _ _ id)
+ (at level 0, format "[ 'lalgType' R 'of' T ]") : form_scope.
+End Exports.
+
+End Lalgebra.
+Import Lalgebra.Exports.
+
+(* Scalar injection (see the definition of in_alg A below). *)
+Local Notation "k %:A" := (k *: 1) : ring_scope.
+
+(* Regular ring algebra tag. *)
+Definition regular R : Type := R.
+Local Notation "R ^o" := (regular R) (at level 2, format "R ^o") : type_scope.
+
+Section LalgebraTheory.
+
+Variables (R : ringType) (A : lalgType R).
+Implicit Types x y : A.
+
+Lemma scalerAl k (x y : A) : k *: (x * y) = k *: x * y.
+Proof. by case: A k x y => ? []. Qed.
+
+Lemma mulr_algl a x : a%:A * x = a *: x.
+Proof. by rewrite -scalerAl mul1r. Qed.
+
+Canonical regular_eqType := [eqType of R^o].
+Canonical regular_choiceType := [choiceType of R^o].
+Canonical regular_zmodType := [zmodType of R^o].
+Canonical regular_ringType := [ringType of R^o].
+
+Definition regular_lmodMixin :=
+ let mkMixin := @Lmodule.Mixin R regular_zmodType (@mul R) in
+ mkMixin (@mulrA R) (@mul1r R) (@mulrDr R) (fun v a b => mulrDl a b v).
+
+Canonical regular_lmodType := LmodType R R^o regular_lmodMixin.
+Canonical regular_lalgType := LalgType R R^o (@mulrA regular_ringType).
+
+Section ClosedPredicates.
+
+Variable S : predPredType A.
+
+Definition subalg_closed := [/\ 1 \in S, linear_closed S & mulr_2closed S].
+
+Lemma subalg_closedZ : subalg_closed -> submod_closed S.
+Proof. by case=> S1 Slin _; split; rewrite // -(subrr 1) linear_closedB. Qed.
+
+Lemma subalg_closedBM : subalg_closed -> subring_closed S.
+Proof. by case=> S1 Slin SM; split=> //; apply: linear_closedB. Qed.
+
+End ClosedPredicates.
+
+End LalgebraTheory.
+
+(* Morphism hierarchy. *)
+
+Module Additive.
+
+Section ClassDef.
+
+Variables U V : zmodType.
+
+Definition axiom (f : U -> V) := {morph f : x y / x - y}.
+
+Structure map (phUV : phant (U -> V)) := Pack {apply; _ : axiom apply}.
+Local Coercion apply : map >-> Funclass.
+
+Variables (phUV : phant (U -> V)) (f g : U -> V) (cF : map phUV).
+Definition class := let: Pack _ c as cF' := cF return axiom cF' in c.
+Definition clone fA of phant_id g (apply cF) & phant_id fA class :=
+ @Pack phUV f fA.
+
+End ClassDef.
+
+Module Exports.
+Notation additive f := (axiom f).
+Coercion apply : map >-> Funclass.
+Notation Additive fA := (Pack (Phant _) fA).
+Notation "{ 'additive' fUV }" := (map (Phant fUV))
+ (at level 0, format "{ 'additive' fUV }") : ring_scope.
+Notation "[ 'additive' 'of' f 'as' g ]" := (@clone _ _ _ f g _ _ idfun id)
+ (at level 0, format "[ 'additive' 'of' f 'as' g ]") : form_scope.
+Notation "[ 'additive' 'of' f ]" := (@clone _ _ _ f f _ _ id id)
+ (at level 0, format "[ 'additive' 'of' f ]") : form_scope.
+End Exports.
+
+End Additive.
+Include Additive.Exports. (* Allows GRing.additive to resolve conflicts. *)
+
+(* Lifted additive operations. *)
+Section LiftedZmod.
+Variables (U : Type) (V : zmodType).
+Definition null_fun_head (phV : phant V) of U : V := let: Phant := phV in 0.
+Definition add_fun_head t (f g : U -> V) x := let: tt := t in f x + g x.
+Definition sub_fun_head t (f g : U -> V) x := let: tt := t in f x - g x.
+End LiftedZmod.
+
+(* Lifted multiplication. *)
+Section LiftedRing.
+Variables (R : ringType) (T : Type).
+Implicit Type f : T -> R.
+Definition mull_fun_head t a f x := let: tt := t in a * f x.
+Definition mulr_fun_head t a f x := let: tt := t in f x * a.
+End LiftedRing.
+
+(* Lifted linear operations. *)
+Section LiftedScale.
+Variables (R : ringType) (U : Type) (V : lmodType R) (A : lalgType R).
+Definition scale_fun_head t a (f : U -> V) x := let: tt := t in a *: f x.
+Definition in_alg_head (phA : phant A) k : A := let: Phant := phA in k%:A.
+End LiftedScale.
+
+Notation null_fun V := (null_fun_head (Phant V)) (only parsing).
+(* The real in_alg notation is declared after GRing.Theory so that at least *)
+(* in Coq 8.2 it gets precedence when GRing.Theory is not imported. *)
+Local Notation in_alg_loc A := (in_alg_head (Phant A)) (only parsing).
+
+Local Notation "\0" := (null_fun _) : ring_scope.
+Local Notation "f \+ g" := (add_fun_head tt f g) : ring_scope.
+Local Notation "f \- g" := (sub_fun_head tt f g) : ring_scope.
+Local Notation "a \*: f" := (scale_fun_head tt a f) : ring_scope.
+Local Notation "x \*o f" := (mull_fun_head tt x f) : ring_scope.
+Local Notation "x \o* f" := (mulr_fun_head tt x f) : ring_scope.
+
+Section AdditiveTheory.
+
+Section Properties.
+
+Variables (U V : zmodType) (k : unit) (f : {additive U -> V}).
+
+Lemma raddfB : {morph f : x y / x - y}. Proof. exact: Additive.class. Qed.
+
+Lemma raddf0 : f 0 = 0.
+Proof. by rewrite -[0]subr0 raddfB subrr. Qed.
+
+Lemma raddf_eq0 x : injective f -> (f x == 0) = (x == 0).
+Proof. by move=> /inj_eq <-; rewrite raddf0. Qed.
+
+Lemma raddfN : {morph f : x / - x}.
+Proof. by move=> x /=; rewrite -sub0r raddfB raddf0 sub0r. Qed.
+
+Lemma raddfD : {morph f : x y / x + y}.
+Proof. by move=> x y; rewrite -[y]opprK raddfB -raddfN. Qed.
+
+Lemma raddfMn n : {morph f : x / x *+ n}.
+Proof. by elim: n => [|n IHn] x /=; rewrite ?raddf0 // !mulrS raddfD IHn. Qed.
+
+Lemma raddfMNn n : {morph f : x / x *- n}.
+Proof. by move=> x /=; rewrite raddfN raddfMn. Qed.
+
+Lemma raddf_sum I r (P : pred I) E :
+ f (\sum_(i <- r | P i) E i) = \sum_(i <- r | P i) f (E i).
+Proof. exact: (big_morph f raddfD raddf0). Qed.
+
+Lemma can2_additive f' : cancel f f' -> cancel f' f -> additive f'.
+Proof. by move=> fK f'K x y /=; apply: (canLR fK); rewrite raddfB !f'K. Qed.
+
+Lemma bij_additive :
+ bijective f -> exists2 f' : {additive V -> U}, cancel f f' & cancel f' f.
+Proof. by case=> f' fK f'K; exists (Additive (can2_additive fK f'K)). Qed.
+
+Fact locked_is_additive : additive (locked_with k (f : U -> V)).
+Proof. by case: k f => [] []. Qed.
+Canonical locked_additive := Additive locked_is_additive.
+
+End Properties.
+
+Section RingProperties.
+
+Variables (R S : ringType) (f : {additive R -> S}).
+
+Lemma raddfMnat n x : f (n%:R * x) = n%:R * f x.
+Proof. by rewrite !mulr_natl raddfMn. Qed.
+
+Lemma raddfMsign n x : f ((-1) ^+ n * x) = (-1) ^+ n * f x.
+Proof. by rewrite !(mulr_sign, =^~ signr_odd) (fun_if f) raddfN. Qed.
+
+Variables (U : lmodType R) (V : lmodType S) (h : {additive U -> V}).
+
+Lemma raddfZnat n u : h (n%:R *: u) = n%:R *: h u.
+Proof. by rewrite !scaler_nat raddfMn. Qed.
+
+Lemma raddfZsign n u : h ((-1) ^+ n *: u) = (-1) ^+ n *: h u.
+Proof. by rewrite !(scaler_sign, =^~ signr_odd) (fun_if h) raddfN. Qed.
+
+End RingProperties.
+
+Section AddFun.
+
+Variables (U V W : zmodType) (f g : {additive V -> W}) (h : {additive U -> V}).
+
+Fact idfun_is_additive : additive (@idfun U).
+Proof. by []. Qed.
+Canonical idfun_additive := Additive idfun_is_additive.
+
+Fact comp_is_additive : additive (f \o h).
+Proof. by move=> x y /=; rewrite !raddfB. Qed.
+Canonical comp_additive := Additive comp_is_additive.
+
+Fact opp_is_additive : additive (-%R : U -> U).
+Proof. by move=> x y; rewrite /= opprD. Qed.
+Canonical opp_additive := Additive opp_is_additive.
+
+Fact null_fun_is_additive : additive (\0 : U -> V).
+Proof. by move=> /=; rewrite subr0. Qed.
+Canonical null_fun_additive := Additive null_fun_is_additive.
+
+Fact add_fun_is_additive : additive (f \+ g).
+Proof.
+by move=> x y /=; rewrite !raddfB addrCA -!addrA addrCA -opprD.
+Qed.
+Canonical add_fun_additive := Additive add_fun_is_additive.
+
+Fact sub_fun_is_additive : additive (f \- g).
+Proof.
+by move=> x y /=; rewrite !raddfB addrAC -!addrA -!opprD addrAC addrA.
+Qed.
+Canonical sub_fun_additive := Additive sub_fun_is_additive.
+
+End AddFun.
+
+Section MulFun.
+
+Variables (R : ringType) (U : zmodType).
+Variables (a : R) (f : {additive U -> R}).
+
+Fact mull_fun_is_additive : additive (a \*o f).
+Proof. by move=> x y /=; rewrite raddfB mulrBr. Qed.
+Canonical mull_fun_additive := Additive mull_fun_is_additive.
+
+Fact mulr_fun_is_additive : additive (a \o* f).
+Proof. by move=> x y /=; rewrite raddfB mulrBl. Qed.
+Canonical mulr_fun_additive := Additive mulr_fun_is_additive.
+
+End MulFun.
+
+Section ScaleFun.
+
+Variables (R : ringType) (U : zmodType) (V : lmodType R).
+Variables (a : R) (f : {additive U -> V}).
+
+Canonical scale_additive := Additive (@scalerBr R V a).
+Canonical scale_fun_additive := [additive of a \*: f as f \; *:%R a].
+
+End ScaleFun.
+
+End AdditiveTheory.
+
+Module RMorphism.
+
+Section ClassDef.
+
+Variables R S : ringType.
+
+Definition mixin_of (f : R -> S) :=
+ {morph f : x y / x * y}%R * (f 1 = 1) : Prop.
+
+Record class_of f : Prop := Class {base : additive f; mixin : mixin_of f}.
+Local Coercion base : class_of >-> additive.
+
+Structure map (phRS : phant (R -> S)) := Pack {apply; _ : class_of apply}.
+Local Coercion apply : map >-> Funclass.
+Variables (phRS : phant (R -> S)) (f g : R -> S) (cF : map phRS).
+
+Definition class := let: Pack _ c as cF' := cF return class_of cF' in c.
+
+Definition clone fM of phant_id g (apply cF) & phant_id fM class :=
+ @Pack phRS f fM.
+
+Definition pack (fM : mixin_of f) :=
+ fun (bF : Additive.map phRS) fA & phant_id (Additive.class bF) fA =>
+ Pack phRS (Class fA fM).
+
+Canonical additive := Additive.Pack phRS class.
+
+End ClassDef.
+
+Module Exports.
+Notation multiplicative f := (mixin_of f).
+Notation rmorphism f := (class_of f).
+Coercion base : rmorphism >-> Additive.axiom.
+Coercion mixin : rmorphism >-> multiplicative.
+Coercion apply : map >-> Funclass.
+Notation RMorphism fM := (Pack (Phant _) fM).
+Notation AddRMorphism fM := (pack fM id).
+Notation "{ 'rmorphism' fRS }" := (map (Phant fRS))
+ (at level 0, format "{ 'rmorphism' fRS }") : ring_scope.
+Notation "[ 'rmorphism' 'of' f 'as' g ]" := (@clone _ _ _ f g _ _ idfun id)
+ (at level 0, format "[ 'rmorphism' 'of' f 'as' g ]") : form_scope.
+Notation "[ 'rmorphism' 'of' f ]" := (@clone _ _ _ f f _ _ id id)
+ (at level 0, format "[ 'rmorphism' 'of' f ]") : form_scope.
+Coercion additive : map >-> Additive.map.
+Canonical additive.
+End Exports.
+
+End RMorphism.
+Include RMorphism.Exports.
+
+Section RmorphismTheory.
+
+Section Properties.
+
+Variables (R S : ringType) (k : unit) (f : {rmorphism R -> S}).
+
+Lemma rmorph0 : f 0 = 0. Proof. exact: raddf0. Qed.
+Lemma rmorphN : {morph f : x / - x}. Proof. exact: raddfN. Qed.
+Lemma rmorphD : {morph f : x y / x + y}. Proof. exact: raddfD. Qed.
+Lemma rmorphB : {morph f: x y / x - y}. Proof. exact: raddfB. Qed.
+Lemma rmorphMn n : {morph f : x / x *+ n}. Proof. exact: raddfMn. Qed.
+Lemma rmorphMNn n : {morph f : x / x *- n}. Proof. exact: raddfMNn. Qed.
+Lemma rmorph_sum I r (P : pred I) E :
+ f (\sum_(i <- r | P i) E i) = \sum_(i <- r | P i) f (E i).
+Proof. exact: raddf_sum. Qed.
+Lemma rmorphMsign n : {morph f : x / (- 1) ^+ n * x}.
+Proof. exact: raddfMsign. Qed.
+
+Lemma rmorphismP : rmorphism f. Proof. exact: RMorphism.class. Qed.
+Lemma rmorphismMP : multiplicative f. Proof. exact: rmorphismP. Qed.
+Lemma rmorph1 : f 1 = 1. Proof. by case: rmorphismMP. Qed.
+Lemma rmorphM : {morph f: x y / x * y}. Proof. by case: rmorphismMP. Qed.
+
+Lemma rmorph_prod I r (P : pred I) E :
+ f (\prod_(i <- r | P i) E i) = \prod_(i <- r | P i) f (E i).
+Proof. exact: (big_morph f rmorphM rmorph1). Qed.
+
+Lemma rmorphX n : {morph f: x / x ^+ n}.
+Proof. by elim: n => [|n IHn] x; rewrite ?rmorph1 // !exprS rmorphM IHn. Qed.
+
+Lemma rmorph_nat n : f n%:R = n%:R. Proof. by rewrite rmorphMn rmorph1. Qed.
+Lemma rmorphN1 : f (- 1) = (- 1). Proof. by rewrite rmorphN rmorph1. Qed.
+
+Lemma rmorph_sign n : f ((- 1) ^+ n) = (- 1) ^+ n.
+Proof. by rewrite rmorphX rmorphN1. Qed.
+
+Lemma rmorph_char p : p \in [char R] -> p \in [char S].
+Proof. by rewrite !inE -rmorph_nat => /andP[-> /= /eqP->]; rewrite rmorph0. Qed.
+
+Lemma rmorph_eq_nat x n : injective f -> (f x == n%:R) = (x == n%:R).
+Proof. by move/inj_eq <-; rewrite rmorph_nat. Qed.
+
+Lemma rmorph_eq1 x : injective f -> (f x == 1) = (x == 1).
+Proof. exact: rmorph_eq_nat 1%N. Qed.
+
+Lemma can2_rmorphism f' : cancel f f' -> cancel f' f -> rmorphism f'.
+Proof.
+move=> fK f'K; split; first exact: can2_additive fK f'K.
+by split=> [x y|]; apply: (canLR fK); rewrite /= (rmorphM, rmorph1) ?f'K.
+Qed.
+
+Lemma bij_rmorphism :
+ bijective f -> exists2 f' : {rmorphism S -> R}, cancel f f' & cancel f' f.
+Proof. by case=> f' fK f'K; exists (RMorphism (can2_rmorphism fK f'K)). Qed.
+
+Fact locked_is_multiplicative : multiplicative (locked_with k (f : R -> S)).
+Proof. by case: k f => [] [? []]. Qed.
+Canonical locked_rmorphism := AddRMorphism locked_is_multiplicative.
+
+End Properties.
+
+Section Projections.
+
+Variables (R S T : ringType) (f : {rmorphism S -> T}) (g : {rmorphism R -> S}).
+
+Fact idfun_is_multiplicative : multiplicative (@idfun R).
+Proof. by []. Qed.
+Canonical idfun_rmorphism := AddRMorphism idfun_is_multiplicative.
+
+Fact comp_is_multiplicative : multiplicative (f \o g).
+Proof. by split=> [x y|] /=; rewrite ?rmorph1 ?rmorphM. Qed.
+Canonical comp_rmorphism := AddRMorphism comp_is_multiplicative.
+
+End Projections.
+
+Section InAlgebra.
+
+Variables (R : ringType) (A : lalgType R).
+
+Fact in_alg_is_rmorphism : rmorphism (in_alg_loc A).
+Proof.
+split=> [x y|]; first exact: scalerBl.
+by split=> [x y|] /=; rewrite ?scale1r // -scalerAl mul1r scalerA.
+Qed.
+Canonical in_alg_additive := Additive in_alg_is_rmorphism.
+Canonical in_alg_rmorphism := RMorphism in_alg_is_rmorphism.
+
+Lemma in_algE a : in_alg_loc A a = a%:A. Proof. by []. Qed.
+
+End InAlgebra.
+
+End RmorphismTheory.
+
+Module Scale.
+
+Section ScaleLaw.
+
+Structure law (R : ringType) (V : zmodType) (s : R -> V -> V) := Law {
+ op : R -> V -> V;
+ _ : op = s;
+ _ : op (-1) =1 -%R;
+ _ : forall a, additive (op a)
+}.
+
+Definition mul_law R := Law (erefl *%R) (@mulN1r R) (@mulrBr R).
+Definition scale_law R U := Law (erefl *:%R) (@scaleN1r R U) (@scalerBr R U).
+
+Variables (R : ringType) (V : zmodType) (s : R -> V -> V) (s_law : law s).
+Local Notation s_op := (op s_law).
+
+Lemma opE : s_op = s. Proof. by case: s_law. Qed.
+Lemma N1op : s_op (-1) =1 -%R. Proof. by case: s_law. Qed.
+Fact opB a : additive (s_op a). Proof. by case: s_law. Qed.
+Definition op_additive a := Additive (opB a).
+
+Variables (aR : ringType) (nu : {rmorphism aR -> R}).
+Fact comp_opE : nu \; s_op = nu \; s. Proof. exact: congr1 opE. Qed.
+Fact compN1op : (nu \; s_op) (-1) =1 -%R.
+Proof. by move=> v; rewrite /= rmorphN1 N1op. Qed.
+Definition comp_law : law (nu \; s) := Law comp_opE compN1op (fun a => opB _).
+
+End ScaleLaw.
+
+End Scale.
+
+Module Linear.
+
+Section ClassDef.
+
+Variables (R : ringType) (U : lmodType R) (V : zmodType) (s : R -> V -> V).
+Implicit Type phUV : phant (U -> V).
+
+Local Coercion Scale.op : Scale.law >-> Funclass.
+Definition axiom (f : U -> V) (s_law : Scale.law s) of s = s_law :=
+ forall a, {morph f : u v / a *: u + v >-> s a u + v}.
+Definition mixin_of (f : U -> V) :=
+ forall a, {morph f : v / a *: v >-> s a v}.
+
+Record class_of f : Prop := Class {base : additive f; mixin : mixin_of f}.
+Local Coercion base : class_of >-> additive.
+
+Lemma class_of_axiom f s_law Ds : @axiom f s_law Ds -> class_of f.
+Proof.
+move=> fL; have fB: additive f.
+ by move=> x y /=; rewrite -scaleN1r addrC fL Ds Scale.N1op addrC.
+by split=> // a v /=; rewrite -[a *: v](addrK v) fB fL addrK Ds.
+Qed.
+
+Structure map (phUV : phant (U -> V)) := Pack {apply; _ : class_of apply}.
+Local Coercion apply : map >-> Funclass.
+
+Variables (phUV : phant (U -> V)) (f g : U -> V) (cF : map phUV).
+Definition class := let: Pack _ c as cF' := cF return class_of cF' in c.
+Definition clone fL of phant_id g (apply cF) & phant_id fL class :=
+ @Pack phUV f fL.
+
+Definition pack (fZ : mixin_of f) :=
+ fun (bF : Additive.map phUV) fA & phant_id (Additive.class bF) fA =>
+ Pack phUV (Class fA fZ).
+
+Canonical additive := Additive.Pack phUV class.
+
+(* Support for right-to-left rewriting with the generic linearZ rule. *)
+Notation mapUV := (map (Phant (U -> V))).
+Definition map_class := mapUV.
+Definition map_at (a : R) := mapUV.
+Structure map_for a s_a := MapFor {map_for_map : mapUV; _ : s a = s_a}.
+Definition unify_map_at a (f : map_at a) := MapFor f (erefl (s a)).
+Structure wrapped := Wrap {unwrap : mapUV}.
+Definition wrap (f : map_class) := Wrap f.
+
+End ClassDef.
+
+Module Exports.
+Canonical Scale.mul_law.
+Canonical Scale.scale_law.
+Canonical Scale.comp_law.
+Canonical Scale.op_additive.
+Delimit Scope linear_ring_scope with linR.
+Notation "a *: u" := (@Scale.op _ _ *:%R _ a u) : linear_ring_scope.
+Notation "a * u" := (@Scale.op _ _ *%R _ a u) : linear_ring_scope.
+Notation "a *:^ nu u" := (@Scale.op _ _ (nu \; *:%R) _ a u)
+ (at level 40, nu at level 1, format "a *:^ nu u") : linear_ring_scope.
+Notation "a *^ nu u" := (@Scale.op _ _ (nu \; *%R) _ a u)
+ (at level 40, nu at level 1, format "a *^ nu u") : linear_ring_scope.
+Notation scalable_for s f := (mixin_of s f).
+Notation scalable f := (scalable_for *:%R f).
+Notation linear_for s f := (axiom f (erefl s)).
+Notation linear f := (linear_for *:%R f).
+Notation scalar f := (linear_for *%R f).
+Notation lmorphism_for s f := (class_of s f).
+Notation lmorphism f := (lmorphism_for *:%R f).
+Coercion class_of_axiom : axiom >-> lmorphism_for.
+Coercion base : lmorphism_for >-> Additive.axiom.
+Coercion mixin : lmorphism_for >-> scalable.
+Coercion apply : map >-> Funclass.
+Notation Linear fL := (Pack (Phant _) fL).
+Notation AddLinear fZ := (pack fZ id).
+Notation "{ 'linear' fUV | s }" := (map s (Phant fUV))
+ (at level 0, format "{ 'linear' fUV | s }") : ring_scope.
+Notation "{ 'linear' fUV }" := {linear fUV | *:%R}
+ (at level 0, format "{ 'linear' fUV }") : ring_scope.
+Notation "{ 'scalar' U }" := {linear U -> _ | *%R}
+ (at level 0, format "{ 'scalar' U }") : ring_scope.
+Notation "[ 'linear' 'of' f 'as' g ]" := (@clone _ _ _ _ _ f g _ _ idfun id)
+ (at level 0, format "[ 'linear' 'of' f 'as' g ]") : form_scope.
+Notation "[ 'linear' 'of' f ]" := (@clone _ _ _ _ _ f f _ _ id id)
+ (at level 0, format "[ 'linear' 'of' f ]") : form_scope.
+Coercion additive : map >-> Additive.map.
+Canonical additive.
+(* Support for right-to-left rewriting with the generic linearZ rule. *)
+Coercion map_for_map : map_for >-> map.
+Coercion unify_map_at : map_at >-> map_for.
+Canonical unify_map_at.
+Coercion unwrap : wrapped >-> map.
+Coercion wrap : map_class >-> wrapped.
+Canonical wrap.
+End Exports.
+
+End Linear.
+Include Linear.Exports.
+
+Section LinearTheory.
+
+Variable R : ringType.
+
+Section GenericProperties.
+
+Variables (U : lmodType R) (V : zmodType) (s : R -> V -> V) (k : unit).
+Variable f : {linear U -> V | s}.
+
+Lemma linear0 : f 0 = 0. Proof. exact: raddf0. Qed.
+Lemma linearN : {morph f : x / - x}. Proof. exact: raddfN. Qed.
+Lemma linearD : {morph f : x y / x + y}. Proof. exact: raddfD. Qed.
+Lemma linearB : {morph f : x y / x - y}. Proof. exact: raddfB. Qed.
+Lemma linearMn n : {morph f : x / x *+ n}. Proof. exact: raddfMn. Qed.
+Lemma linearMNn n : {morph f : x / x *- n}. Proof. exact: raddfMNn. Qed.
+Lemma linear_sum I r (P : pred I) E :
+ f (\sum_(i <- r | P i) E i) = \sum_(i <- r | P i) f (E i).
+Proof. exact: raddf_sum. Qed.
+
+Lemma linearZ_LR : scalable_for s f. Proof. by case: f => ? []. Qed.
+Lemma linearP a : {morph f : u v / a *: u + v >-> s a u + v}.
+Proof. by move=> u v /=; rewrite linearD linearZ_LR. Qed.
+
+Fact locked_is_scalable : scalable_for s (locked_with k (f : U -> V)).
+Proof. by case: k f => [] [? []]. Qed.
+Canonical locked_linear := AddLinear locked_is_scalable.
+
+End GenericProperties.
+
+Section BidirectionalLinearZ.
+
+Variables (U : lmodType R) (V : zmodType) (s : R -> V -> V).
+
+(* The general form of the linearZ lemma uses some bespoke interfaces to *)
+(* allow right-to-left rewriting when a composite scaling operation such as *)
+(* conjC \; *%R has been expanded, say in a^* * f u. This redex is matched *)
+(* by using the Scale.law interface to recognize a "head" scaling operation *)
+(* h (here *%R), stow away its "scalar" c, then reconcile h c and s a, once *)
+(* s is known, that is, once the Linear.map structure for f has been found. *)
+(* In general, s and a need not be equal to h and c; indeed they need not *)
+(* have the same type! The unification is performed by the unify_map_at *)
+(* default instance for the Linear.map_for U s a h_c sub-interface of *)
+(* Linear.map; the h_c pattern uses the Scale.law structure to insure it is *)
+(* inferred when rewriting right-to-left. *)
+(* The wrap on the rhs allows rewriting f (a *: b *: u) into a *: b *: f u *)
+(* with rewrite !linearZ /= instead of rewrite linearZ /= linearZ /=. *)
+(* Without it, the first rewrite linearZ would produce *)
+(* (a *: apply (map_for_map (@check_map_at .. a f)) (b *: u)%R)%Rlin *)
+(* and matching the second rewrite LHS would bypass the unify_map_at default *)
+(* instance for b, reuse the one for a, and subsequently fail to match the *)
+(* b *: u argument. The extra wrap / unwrap ensures that this can't happen. *)
+(* In the RL direction, the wrap / unwrap will be inserted on the redex side *)
+(* as needed, without causing unnecessary delta-expansion: using an explicit *)
+(* identity function would have Coq normalize the redex to head normal, then *)
+(* reduce the identity to expose the map_for_map projection, and the *)
+(* expanded Linear.map structure would then be exposed in the result. *)
+(* Most of this machinery will be invisible to a casual user, because all *)
+(* the projections and default instances involved are declared as coercions. *)
+
+Variables (S : ringType) (h : S -> V -> V) (h_law : Scale.law h).
+
+Lemma linearZ c a (h_c := Scale.op h_law c) (f : Linear.map_for U s a h_c) u :
+ f (a *: u) = h_c (Linear.wrap f u).
+Proof. by rewrite linearZ_LR; case: f => f /= ->. Qed.
+
+End BidirectionalLinearZ.
+
+Section LmodProperties.
+
+Variables (U V : lmodType R) (f : {linear U -> V}).
+
+Lemma linearZZ : scalable f. Proof. exact: linearZ_LR. Qed.
+Lemma linearPZ : linear f. Proof. exact: linearP. Qed.
+
+Lemma can2_linear f' : cancel f f' -> cancel f' f -> linear f'.
+Proof. by move=> fK f'K a x y /=; apply: (canLR fK); rewrite linearP !f'K. Qed.
+
+Lemma bij_linear :
+ bijective f -> exists2 f' : {linear V -> U}, cancel f f' & cancel f' f.
+Proof. by case=> f' fK f'K; exists (Linear (can2_linear fK f'K)). Qed.
+
+End LmodProperties.
+
+Section ScalarProperties.
+
+Variable (U : lmodType R) (f : {scalar U}).
+
+Lemma scalarZ : scalable_for *%R f. Proof. exact: linearZ_LR. Qed.
+Lemma scalarP : scalar f. Proof. exact: linearP. Qed.
+
+End ScalarProperties.
+
+Section LinearLmod.
+
+Variables (W U : lmodType R) (V : zmodType) (s : R -> V -> V).
+Variables (f : {linear U -> V | s}) (h : {linear W -> U}).
+
+Lemma idfun_is_scalable : scalable (@idfun U). Proof. by []. Qed.
+Canonical idfun_linear := AddLinear idfun_is_scalable.
+
+Lemma opp_is_scalable : scalable (-%R : U -> U).
+Proof. by move=> a v /=; rewrite scalerN. Qed.
+Canonical opp_linear := AddLinear opp_is_scalable.
+
+Lemma comp_is_scalable : scalable_for s (f \o h).
+Proof. by move=> a v /=; rewrite !linearZ_LR. Qed.
+Canonical comp_linear := AddLinear comp_is_scalable.
+
+Variables (s_law : Scale.law s) (g : {linear U -> V | Scale.op s_law}).
+Let Ds : s =1 Scale.op s_law. Proof. by rewrite Scale.opE. Qed.
+
+Lemma null_fun_is_scalable : scalable_for (Scale.op s_law) (\0 : U -> V).
+Proof. by move=> a v /=; rewrite raddf0. Qed.
+Canonical null_fun_linear := AddLinear null_fun_is_scalable.
+
+Lemma add_fun_is_scalable : scalable_for s (f \+ g).
+Proof. by move=> a u; rewrite /= !linearZ_LR !Ds raddfD. Qed.
+Canonical add_fun_linear := AddLinear add_fun_is_scalable.
+
+Lemma sub_fun_is_scalable : scalable_for s (f \- g).
+Proof. by move=> a u; rewrite /= !linearZ_LR !Ds raddfB. Qed.
+Canonical sub_fun_linear := AddLinear sub_fun_is_scalable.
+
+End LinearLmod.
+
+Section LinearLalg.
+
+Variables (A : lalgType R) (U : lmodType R).
+
+Variables (a : A) (f : {linear U -> A}).
+
+Fact mulr_fun_is_scalable : scalable (a \o* f).
+Proof. by move=> k x /=; rewrite linearZ scalerAl. Qed.
+Canonical mulr_fun_linear := AddLinear mulr_fun_is_scalable.
+
+End LinearLalg.
+
+End LinearTheory.
+
+Module LRMorphism.
+
+Section ClassDef.
+
+Variables (R : ringType) (A : lalgType R) (B : ringType) (s : R -> B -> B).
+
+Record class_of (f : A -> B) : Prop :=
+ Class {base : rmorphism f; mixin : scalable_for s f}.
+Local Coercion base : class_of >-> rmorphism.
+Definition base2 f (fLM : class_of f) := Linear.Class fLM (mixin fLM).
+Local Coercion base2 : class_of >-> lmorphism.
+
+Structure map (phAB : phant (A -> B)) := Pack {apply; _ : class_of apply}.
+Local Coercion apply : map >-> Funclass.
+
+Variables (phAB : phant (A -> B)) (f : A -> B) (cF : map phAB).
+Definition class := let: Pack _ c as cF' := cF return class_of cF' in c.
+
+Definition clone :=
+ fun (g : RMorphism.map phAB) fM & phant_id (RMorphism.class g) fM =>
+ fun (h : Linear.map s phAB) fZ &
+ phant_id (Linear.mixin (Linear.class h)) fZ =>
+ Pack phAB (@Class f fM fZ).
+
+Definition pack (fZ : scalable_for s f) :=
+ fun (g : RMorphism.map phAB) fM & phant_id (RMorphism.class g) fM =>
+ Pack phAB (Class fM fZ).
+
+Canonical additive := Additive.Pack phAB class.
+Canonical rmorphism := RMorphism.Pack phAB class.
+Canonical linear := Linear.Pack phAB class.
+Canonical join_rmorphism := @RMorphism.Pack _ _ phAB linear class.
+Canonical join_linear := @Linear.Pack R A B s phAB rmorphism class.
+
+End ClassDef.
+
+Module Exports.
+Notation lrmorphism_for s f := (class_of s f).
+Notation lrmorphism f := (lrmorphism_for *:%R f).
+Coercion base : lrmorphism_for >-> RMorphism.class_of.
+Coercion base2 : lrmorphism_for >-> lmorphism_for.
+Coercion apply : map >-> Funclass.
+Notation LRMorphism f_lrM := (Pack (Phant _) (Class f_lrM f_lrM)).
+Notation AddLRMorphism fZ := (pack fZ id).
+Notation "{ 'lrmorphism' fAB | s }" := (map s (Phant fAB))
+ (at level 0, format "{ 'lrmorphism' fAB | s }") : ring_scope.
+Notation "{ 'lrmorphism' fAB }" := {lrmorphism fAB | *:%R}
+ (at level 0, format "{ 'lrmorphism' fAB }") : ring_scope.
+Notation "[ 'lrmorphism' 'of' f ]" := (@clone _ _ _ _ _ f _ _ id _ _ id)
+ (at level 0, format "[ 'lrmorphism' 'of' f ]") : form_scope.
+Coercion additive : map >-> Additive.map.
+Canonical additive.
+Coercion rmorphism : map >-> RMorphism.map.
+Canonical rmorphism.
+Coercion linear : map >-> Linear.map.
+Canonical linear.
+Canonical join_rmorphism.
+Canonical join_linear.
+End Exports.
+
+End LRMorphism.
+Include LRMorphism.Exports.
+
+Section LRMorphismTheory.
+
+Variables (R : ringType) (A B : lalgType R) (C : ringType) (s : R -> C -> C).
+Variables (k : unit) (f : {lrmorphism A -> B}) (g : {lrmorphism B -> C | s}).
+
+Definition idfun_lrmorphism := [lrmorphism of @idfun A].
+Definition comp_lrmorphism := [lrmorphism of g \o f].
+Definition locked_lrmorphism := [lrmorphism of locked_with k (f : A -> B)].
+
+Lemma rmorph_alg a : f a%:A = a%:A.
+Proof. by rewrite linearZ rmorph1. Qed.
+
+Lemma lrmorphismP : lrmorphism f. Proof. exact: LRMorphism.class. Qed.
+
+Lemma can2_lrmorphism f' : cancel f f' -> cancel f' f -> lrmorphism f'.
+Proof.
+move=> fK f'K; split; [exact: (can2_rmorphism fK) | exact: (can2_linear fK)].
+Qed.
+
+Lemma bij_lrmorphism :
+ bijective f -> exists2 f' : {lrmorphism B -> A}, cancel f f' & cancel f' f.
+Proof.
+by case/bij_rmorphism=> f' fK f'K; exists (AddLRMorphism (can2_linear fK f'K)).
+Qed.
+
+End LRMorphismTheory.
+
+Module ComRing.
+
+Definition RingMixin R one mul mulA mulC mul1x mul_addl :=
+ let mulx1 := Monoid.mulC_id mulC mul1x in
+ let mul_addr := Monoid.mulC_dist mulC mul_addl in
+ @Ring.EtaMixin R one mul mulA mul1x mulx1 mul_addl mul_addr.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class {base : Ring.class_of R; mixin : commutative (Ring.mul base)}.
+Local Coercion base : class_of >-> Ring.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variable (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack mul0 (m0 : @commutative T T mul0) :=
+ fun bT b & phant_id (Ring.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> Ring.class_of.
+Implicit Arguments mixin [R].
+Coercion mixin : class_of >-> commutative.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Notation comRingType := type.
+Notation ComRingType T m := (@pack T _ m _ _ id _ id).
+Notation ComRingMixin := RingMixin.
+Notation "[ 'comRingType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'comRingType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'comRingType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'comRingType' 'of' T ]") : form_scope.
+End Exports.
+
+End ComRing.
+Import ComRing.Exports.
+
+Section ComRingTheory.
+
+Variable R : comRingType.
+Implicit Types x y : R.
+
+Lemma mulrC : @commutative R R *%R. Proof. by case: R => T []. Qed.
+Canonical mul_comoid := Monoid.ComLaw mulrC.
+Lemma mulrCA : @left_commutative R R *%R. Proof. exact: mulmCA. Qed.
+Lemma mulrAC : @right_commutative R R *%R. Proof. exact: mulmAC. Qed.
+Lemma mulrACA : @interchange R *%R *%R. Proof. exact: mulmACA. Qed.
+
+Lemma exprMn n : {morph (fun x => x ^+ n) : x y / x * y}.
+Proof. move=> x y; apply: exprMn_comm; exact: mulrC. Qed.
+
+Lemma prodrXl n I r (P : pred I) (F : I -> R) :
+ \prod_(i <- r | P i) F i ^+ n = (\prod_(i <- r | P i) F i) ^+ n.
+Proof. by rewrite (big_morph _ (exprMn n) (expr1n _ n)). Qed.
+
+Lemma prodr_undup_exp_count (I : eqType) r (P : pred I) (F : I -> R) :
+ \prod_(i <- undup r | P i) F i ^+ count_mem i r = \prod_(i <- r | P i) F i.
+Proof. exact: big_undup_iterop_count. Qed.
+
+Lemma exprDn x y n :
+ (x + y) ^+ n = \sum_(i < n.+1) (x ^+ (n - i) * y ^+ i) *+ 'C(n, i).
+Proof. by rewrite exprDn_comm //; exact: mulrC. Qed.
+
+Lemma exprBn x y n :
+ (x - y) ^+ n =
+ \sum_(i < n.+1) ((-1) ^+ i * x ^+ (n - i) * y ^+ i) *+ 'C(n, i).
+Proof. by rewrite exprBn_comm //; exact: mulrC. Qed.
+
+Lemma subrXX x y n :
+ x ^+ n - y ^+ n = (x - y) * (\sum_(i < n) x ^+ (n.-1 - i) * y ^+ i).
+Proof. by rewrite -subrXX_comm //; exact: mulrC. Qed.
+
+Lemma sqrrD x y : (x + y) ^+ 2 = x ^+ 2 + x * y *+ 2 + y ^+ 2.
+Proof. by rewrite exprDn !big_ord_recr big_ord0 /= add0r mulr1 mul1r. Qed.
+
+Lemma sqrrB x y : (x - y) ^+ 2 = x ^+ 2 - x * y *+ 2 + y ^+ 2.
+Proof. by rewrite sqrrD mulrN mulNrn sqrrN. Qed.
+
+Lemma subr_sqr x y : x ^+ 2 - y ^+ 2 = (x - y) * (x + y).
+Proof. by rewrite subrXX !big_ord_recr big_ord0 /= add0r mulr1 mul1r. Qed.
+
+Lemma subr_sqrDB x y : (x + y) ^+ 2 - (x - y) ^+ 2 = x * y *+ 4.
+Proof.
+rewrite sqrrD sqrrB -!(addrAC _ (y ^+ 2)) opprB.
+by rewrite addrC addrA subrK -mulrnDr.
+Qed.
+
+Section FrobeniusAutomorphism.
+
+Variables (p : nat) (charRp : p \in [char R]).
+
+Lemma Frobenius_aut_is_rmorphism : rmorphism (Frobenius_aut charRp).
+Proof.
+split=> [x y|]; first exact: Frobenius_autB_comm (mulrC _ _).
+split=> [x y|]; first exact: Frobenius_autM_comm (mulrC _ _).
+exact: Frobenius_aut1.
+Qed.
+
+Canonical Frobenius_aut_additive := Additive Frobenius_aut_is_rmorphism.
+Canonical Frobenius_aut_rmorphism := RMorphism Frobenius_aut_is_rmorphism.
+
+End FrobeniusAutomorphism.
+
+Lemma exprDn_char x y n : [char R].-nat n -> (x + y) ^+ n = x ^+ n + y ^+ n.
+Proof.
+pose p := pdiv n; have [|n_gt1 charRn] := leqP n 1; first by case: (n) => [|[]].
+have charRp: p \in [char R] by rewrite (pnatPpi charRn) ?pi_pdiv.
+have{charRn} /p_natP[e ->]: p.-nat n by rewrite -(eq_pnat _ (charf_eq charRp)).
+by elim: e => // e IHe; rewrite !expnSr !exprM IHe -Frobenius_autE rmorphD.
+Qed.
+
+Lemma rmorph_comm (S : ringType) (f : {rmorphism R -> S}) x y :
+ comm (f x) (f y).
+Proof. by red; rewrite -!rmorphM mulrC. Qed.
+
+Section ScaleLinear.
+
+Variables (U V : lmodType R) (b : R) (f : {linear U -> V}).
+
+Lemma scale_is_scalable : scalable ( *:%R b : V -> V).
+Proof. by move=> a v /=; rewrite !scalerA mulrC. Qed.
+Canonical scale_linear := AddLinear scale_is_scalable.
+
+Lemma scale_fun_is_scalable : scalable (b \*: f).
+Proof. by move=> a v /=; rewrite !linearZ. Qed.
+Canonical scale_fun_linear := AddLinear scale_fun_is_scalable.
+
+End ScaleLinear.
+
+End ComRingTheory.
+
+Module Algebra.
+
+Section Mixin.
+
+Variables (R : ringType) (A : lalgType R).
+
+Definition axiom := forall k (x y : A), k *: (x * y) = x * (k *: y).
+
+Lemma comm_axiom : phant A -> commutative (@mul A) -> axiom.
+Proof. by move=> _ commA k x y; rewrite commA scalerAl commA. Qed.
+
+End Mixin.
+
+Section ClassDef.
+
+Variable R : ringType.
+
+Record class_of (T : Type) : Type := Class {
+ base : Lalgebra.class_of R T;
+ mixin : axiom (Lalgebra.Pack _ base T)
+}.
+Local Coercion base : class_of >-> Lalgebra.class_of.
+
+Structure type (phR : phant R) := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variable (phR : phant R) (T : Type) (cT : type phR).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack phR T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack b0 (ax0 : @axiom R b0) :=
+ fun bT b & phant_id (@Lalgebra.class R phR bT) b =>
+ fun ax & phant_id ax0 ax => Pack phR (@Class T b ax) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+Definition lmodType := @Lmodule.Pack R phR cT xclass xT.
+Definition lalgType := @Lalgebra.Pack R phR cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> Lalgebra.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Coercion lmodType : type >-> Lmodule.type.
+Canonical lmodType.
+Coercion lalgType : type >-> Lalgebra.type.
+Canonical lalgType.
+Notation algType R := (type (Phant R)).
+Notation AlgType R A ax := (@pack _ (Phant R) A _ ax _ _ id _ id).
+Notation CommAlgType R A := (AlgType R A (comm_axiom (Phant A) (@mulrC _))).
+Notation "[ 'algType' R 'of' T 'for' cT ]" := (@clone _ (Phant R) T cT _ idfun)
+ (at level 0, format "[ 'algType' R 'of' T 'for' cT ]")
+ : form_scope.
+Notation "[ 'algType' R 'of' T ]" := (@clone _ (Phant R) T _ _ id)
+ (at level 0, format "[ 'algType' R 'of' T ]") : form_scope.
+End Exports.
+
+End Algebra.
+Import Algebra.Exports.
+
+Section AlgebraTheory.
+
+Variables (R : comRingType) (A : algType R).
+Implicit Types (k : R) (x y : A).
+
+Lemma scalerAr k x y : k *: (x * y) = x * (k *: y).
+Proof. by case: A k x y => T []. Qed.
+
+Lemma scalerCA k x y : k *: x * y = x * (k *: y).
+Proof. by rewrite -scalerAl scalerAr. Qed.
+
+Lemma mulr_algr a x : x * a%:A = a *: x.
+Proof. by rewrite -scalerAr mulr1. Qed.
+
+Lemma exprZn k x n : (k *: x) ^+ n = k ^+ n *: x ^+ n.
+Proof.
+elim: n => [|n IHn]; first by rewrite !expr0 scale1r.
+by rewrite !exprS IHn -scalerA scalerAr scalerAl.
+Qed.
+
+Lemma scaler_prod I r (P : pred I) (F : I -> R) (G : I -> A) :
+ \prod_(i <- r | P i) (F i *: G i) =
+ \prod_(i <- r | P i) F i *: \prod_(i <- r | P i) G i.
+Proof.
+elim/big_rec3: _ => [|i x a _ _ ->]; first by rewrite scale1r.
+by rewrite -scalerAl -scalerAr scalerA.
+Qed.
+
+Lemma scaler_prodl (I : finType) (S : pred I) (F : I -> A) k :
+ \prod_(i in S) (k *: F i) = k ^+ #|S| *: \prod_(i in S) F i.
+Proof. by rewrite scaler_prod prodr_const. Qed.
+
+Lemma scaler_prodr (I : finType) (S : pred I) (F : I -> R) x :
+ \prod_(i in S) (F i *: x) = \prod_(i in S) F i *: x ^+ #|S|.
+Proof. by rewrite scaler_prod prodr_const. Qed.
+
+Canonical regular_comRingType := [comRingType of R^o].
+Canonical regular_algType := CommAlgType R R^o.
+
+Variables (U : lmodType R) (a : A) (f : {linear U -> A}).
+
+Lemma mull_fun_is_scalable : scalable (a \*o f).
+Proof. by move=> k x /=; rewrite linearZ scalerAr. Qed.
+Canonical mull_fun_linear := AddLinear mull_fun_is_scalable.
+
+End AlgebraTheory.
+
+Module UnitRing.
+
+Record mixin_of (R : ringType) : Type := Mixin {
+ unit : pred R;
+ inv : R -> R;
+ _ : {in unit, left_inverse 1 inv *%R};
+ _ : {in unit, right_inverse 1 inv *%R};
+ _ : forall x y, y * x = 1 /\ x * y = 1 -> unit x;
+ _ : {in [predC unit], inv =1 id}
+}.
+
+Definition EtaMixin R unit inv mulVr mulrV unitP inv_out :=
+ let _ := @Mixin R unit inv mulVr mulrV unitP inv_out in
+ @Mixin (Ring.Pack (Ring.class R) R) unit inv mulVr mulrV unitP inv_out.
+
+Section ClassDef.
+
+Record class_of (R : Type) : Type := Class {
+ base : Ring.class_of R;
+ mixin : mixin_of (Ring.Pack base R)
+}.
+Local Coercion base : class_of >-> Ring.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack b0 (m0 : mixin_of (@Ring.Pack T b0 T)) :=
+ fun bT b & phant_id (Ring.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> Ring.class_of.
+Coercion mixin : class_of >-> mixin_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Notation unitRingType := type.
+Notation UnitRingType T m := (@pack T _ m _ _ id _ id).
+Notation UnitRingMixin := EtaMixin.
+Notation "[ 'unitRingType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'unitRingType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'unitRingType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'unitRingType' 'of' T ]") : form_scope.
+End Exports.
+
+End UnitRing.
+Import UnitRing.Exports.
+
+Definition unit {R : unitRingType} :=
+ [qualify a u : R | UnitRing.unit (UnitRing.class R) u].
+Fact unit_key R : pred_key (@unit R). Proof. by []. Qed.
+Canonical unit_keyed R := KeyedQualifier (@unit_key R).
+Definition inv {R : unitRingType} : R -> R := UnitRing.inv (UnitRing.class R).
+
+Local Notation "x ^-1" := (inv x).
+Local Notation "x / y" := (x * y^-1).
+Local Notation "x ^- n" := ((x ^+ n)^-1).
+
+Section UnitRingTheory.
+
+Variable R : unitRingType.
+Implicit Types x y : R.
+
+Lemma divrr : {in unit, right_inverse 1 (@inv R) *%R}.
+Proof. by case: R => T [? []]. Qed.
+Definition mulrV := divrr.
+
+Lemma mulVr : {in unit, left_inverse 1 (@inv R) *%R}.
+Proof. by case: R => T [? []]. Qed.
+
+Lemma invr_out x : x \isn't a unit -> x^-1 = x.
+Proof. by case: R x => T [? []]. Qed.
+
+Lemma unitrP x : reflect (exists y, y * x = 1 /\ x * y = 1) (x \is a unit).
+Proof.
+apply: (iffP idP) => [Ux | []]; last by case: R x => T [? []].
+by exists x^-1; rewrite divrr ?mulVr.
+Qed.
+
+Lemma mulKr : {in unit, left_loop (@inv R) *%R}.
+Proof. by move=> x Ux y; rewrite mulrA mulVr ?mul1r. Qed.
+
+Lemma mulVKr : {in unit, rev_left_loop (@inv R) *%R}.
+Proof. by move=> x Ux y; rewrite mulrA mulrV ?mul1r. Qed.
+
+Lemma mulrK : {in unit, right_loop (@inv R) *%R}.
+Proof. by move=> x Ux y; rewrite -mulrA divrr ?mulr1. Qed.
+
+Lemma mulrVK : {in unit, rev_right_loop (@inv R) *%R}.
+Proof. by move=> x Ux y; rewrite -mulrA mulVr ?mulr1. Qed.
+Definition divrK := mulrVK.
+
+Lemma mulrI : {in @unit R, right_injective *%R}.
+Proof. by move=> x Ux; exact: can_inj (mulKr Ux). Qed.
+
+Lemma mulIr : {in @unit R, left_injective *%R}.
+Proof. by move=> x Ux; exact: can_inj (mulrK Ux). Qed.
+
+(* Due to noncommutativity, fractions are inverted. *)
+Lemma telescope_prodr n m (f : nat -> R) :
+ (forall k, n < k < m -> f k \is a unit) -> n < m ->
+ \prod_(n <= k < m) (f k / f k.+1) = f n / f m.
+Proof.
+move=> Uf /subnK-Dm; do [rewrite -{}Dm; move: {m}(m - _)%N => m] in Uf *.
+rewrite unlock /index_iota -addSnnS addnK /= -mulrA; congr (_ * _).
+have{Uf}: all [preim f of unit] (iota n.+1 m).
+ by apply/allP=> k; rewrite mem_iota addnC => /Uf.
+elim: m n => [|m IHm] n /=; first by rewrite mulr1.
+by rewrite -mulrA addSnnS => /andP[/mulKr-> /IHm].
+Qed.
+
+Lemma commrV x y : comm x y -> comm x y^-1.
+Proof.
+have [Uy cxy | /invr_out-> //] := boolP (y \in unit).
+by apply: (canLR (mulrK Uy)); rewrite -mulrA cxy mulKr.
+Qed.
+
+Lemma unitrE x : (x \is a unit) = (x / x == 1).
+Proof.
+apply/idP/eqP=> [Ux | xx1]; first exact: divrr.
+by apply/unitrP; exists x^-1; rewrite -commrV.
+Qed.
+
+Lemma invrK : involutive (@inv R).
+Proof.
+move=> x; case Ux: (x \in unit); last by rewrite !invr_out ?Ux.
+rewrite -(mulrK Ux _^-1) -mulrA commrV ?mulKr //.
+by apply/unitrP; exists x; rewrite divrr ?mulVr.
+Qed.
+
+Lemma invr_inj : injective (@inv R).
+Proof. exact: inv_inj invrK. Qed.
+
+Lemma unitrV x : (x^-1 \in unit) = (x \in unit).
+Proof. by rewrite !unitrE invrK commrV. Qed.
+
+Lemma unitr1 : 1 \in @unit R.
+Proof. by apply/unitrP; exists 1; rewrite mulr1. Qed.
+
+Lemma invr1 : 1^-1 = 1 :> R.
+Proof. by rewrite -{2}(mulVr unitr1) mulr1. Qed.
+
+Lemma div1r x : 1 / x = x^-1. Proof. by rewrite mul1r. Qed.
+Lemma divr1 x : x / 1 = x. Proof. by rewrite invr1 mulr1. Qed.
+
+Lemma natr_div m d :
+ d %| m -> d%:R \is a @unit R -> (m %/ d)%:R = m%:R / d%:R :> R.
+Proof.
+by rewrite dvdn_eq => /eqP def_m unit_d; rewrite -{2}def_m natrM mulrK.
+Qed.
+
+Lemma unitr0 : (0 \is a @unit R) = false.
+Proof.
+by apply/unitrP=> [[x [_]]]; apply/eqP; rewrite mul0r eq_sym oner_neq0.
+Qed.
+
+Lemma invr0 : 0^-1 = 0 :> R.
+Proof. by rewrite invr_out ?unitr0. Qed.
+
+Lemma unitrN1 : -1 \is a @unit R.
+Proof. by apply/unitrP; exists (-1); rewrite mulrNN mulr1. Qed.
+
+Lemma invrN1 : (-1)^-1 = -1 :> R.
+Proof. by rewrite -{2}(divrr unitrN1) mulN1r opprK. Qed.
+
+Lemma invr_sign n : ((-1) ^- n) = (-1) ^+ n :> R.
+Proof. by rewrite -signr_odd; case: (odd n); rewrite (invr1, invrN1). Qed.
+
+Lemma unitrMl x y : y \is a unit -> (x * y \is a unit) = (x \is a unit).
+Proof.
+move=> Uy; wlog Ux: x y Uy / x \is a unit => [WHxy|].
+ by apply/idP/idP=> Ux; first rewrite -(mulrK Uy x); rewrite WHxy ?unitrV.
+rewrite Ux; apply/unitrP; exists (y^-1 * x^-1).
+by rewrite -!mulrA mulKr ?mulrA ?mulrK ?divrr ?mulVr.
+Qed.
+
+Lemma unitrMr x y : x \is a unit -> (x * y \is a unit) = (y \is a unit).
+Proof.
+move=> Ux; apply/idP/idP=> [Uxy | Uy]; last by rewrite unitrMl.
+by rewrite -(mulKr Ux y) unitrMl ?unitrV.
+Qed.
+
+Lemma invrM : {in unit &, forall x y, (x * y)^-1 = y^-1 * x^-1}.
+Proof.
+move=> x y Ux Uy; have Uxy: (x * y \in unit) by rewrite unitrMl.
+by apply: (mulrI Uxy); rewrite divrr ?mulrA ?mulrK ?divrr.
+Qed.
+
+Lemma unitrM_comm x y :
+ comm x y -> (x * y \is a unit) = (x \is a unit) && (y \is a unit).
+Proof.
+move=> cxy; apply/idP/andP=> [Uxy | [Ux Uy]]; last by rewrite unitrMl.
+suffices Ux: x \in unit by rewrite unitrMr in Uxy.
+apply/unitrP; case/unitrP: Uxy => z [zxy xyz]; exists (y * z).
+rewrite mulrA xyz -{1}[y]mul1r -{1}zxy cxy -!mulrA (mulrA x) (mulrA _ z) xyz.
+by rewrite mul1r -cxy.
+Qed.
+
+Lemma unitrX x n : x \is a unit -> x ^+ n \is a unit.
+Proof.
+by move=> Ux; elim: n => [|n IHn]; rewrite ?unitr1 // exprS unitrMl.
+Qed.
+
+Lemma unitrX_pos x n : n > 0 -> (x ^+ n \in unit) = (x \in unit).
+Proof.
+case: n => // n _; rewrite exprS unitrM_comm; last exact: commrX.
+by case Ux: (x \is a unit); rewrite // unitrX.
+Qed.
+
+Lemma exprVn x n : x^-1 ^+ n = x ^- n.
+Proof.
+elim: n => [|n IHn]; first by rewrite !expr0 ?invr1.
+case Ux: (x \is a unit); first by rewrite exprSr exprS IHn -invrM // unitrX.
+by rewrite !invr_out ?unitrX_pos ?Ux.
+Qed.
+
+Lemma exprB m n x : n <= m -> x \is a unit -> x ^+ (m - n) = x ^+ m / x ^+ n.
+Proof. by move/subnK=> {2}<- Ux; rewrite exprD mulrK ?unitrX. Qed.
+
+Lemma invr_neq0 x : x != 0 -> x^-1 != 0.
+Proof.
+move=> nx0; case Ux: (x \is a unit); last by rewrite invr_out ?Ux.
+by apply/eqP=> x'0; rewrite -unitrV x'0 unitr0 in Ux.
+Qed.
+
+Lemma invr_eq0 x : (x^-1 == 0) = (x == 0).
+Proof. by apply: negb_inj; apply/idP/idP; move/invr_neq0; rewrite ?invrK. Qed.
+
+Lemma invr_eq1 x : (x^-1 == 1) = (x == 1).
+Proof. by rewrite (inv_eq invrK) invr1. Qed.
+
+Lemma rev_unitrP (x y : R^c) : y * x = 1 /\ x * y = 1 -> x \is a unit.
+Proof. by case=> [yx1 xy1]; apply/unitrP; exists y. Qed.
+
+Definition converse_unitRingMixin :=
+ @UnitRing.Mixin _ ((unit : pred_class) : pred R^c) _
+ mulrV mulVr rev_unitrP invr_out.
+Canonical converse_unitRingType := UnitRingType R^c converse_unitRingMixin.
+Canonical regular_unitRingType := [unitRingType of R^o].
+
+Section ClosedPredicates.
+
+Variables S : predPredType R.
+
+Definition invr_closed := {in S, forall x, x^-1 \in S}.
+Definition divr_2closed := {in S &, forall x y, x / y \in S}.
+Definition divr_closed := 1 \in S /\ divr_2closed.
+Definition sdivr_closed := -1 \in S /\ divr_2closed.
+Definition divring_closed := [/\ 1 \in S, subr_2closed S & divr_2closed].
+
+Lemma divr_closedV : divr_closed -> invr_closed.
+Proof. by case=> S1 Sdiv x Sx; rewrite -[x^-1]mul1r Sdiv. Qed.
+
+Lemma divr_closedM : divr_closed -> mulr_closed S.
+Proof.
+by case=> S1 Sdiv; split=> // x y Sx Sy; rewrite -[y]invrK -[y^-1]mul1r !Sdiv.
+Qed.
+
+Lemma sdivr_closed_div : sdivr_closed -> divr_closed.
+Proof. by case=> SN1 Sdiv; split; rewrite // -(divrr unitrN1) Sdiv. Qed.
+
+Lemma sdivr_closedM : sdivr_closed -> smulr_closed S.
+Proof.
+by move=> Sdiv; have [_ SM] := divr_closedM (sdivr_closed_div Sdiv); case: Sdiv.
+Qed.
+
+Lemma divring_closedBM : divring_closed -> subring_closed S.
+Proof. by case=> S1 SB Sdiv; split=> //; case: divr_closedM. Qed.
+
+Lemma divring_closed_div : divring_closed -> sdivr_closed.
+Proof.
+case=> S1 SB Sdiv; split; rewrite ?zmod_closedN //.
+exact/subring_closedB/divring_closedBM.
+Qed.
+
+End ClosedPredicates.
+
+End UnitRingTheory.
+
+Implicit Arguments invr_inj [[R] x1 x2].
+
+Section UnitRingMorphism.
+
+Variables (R S : unitRingType) (f : {rmorphism R -> S}).
+
+Lemma rmorph_unit x : x \in unit -> f x \in unit.
+Proof.
+case/unitrP=> y [yx1 xy1]; apply/unitrP.
+by exists (f y); rewrite -!rmorphM // yx1 xy1 rmorph1.
+Qed.
+
+Lemma rmorphV : {in unit, {morph f: x / x^-1}}.
+Proof.
+move=> x Ux; rewrite /= -[(f x)^-1]mul1r.
+by apply: (canRL (mulrK (rmorph_unit Ux))); rewrite -rmorphM mulVr ?rmorph1.
+Qed.
+
+Lemma rmorph_div x y : y \in unit -> f (x / y) = f x / f y.
+Proof. by move=> Uy; rewrite rmorphM rmorphV. Qed.
+
+End UnitRingMorphism.
+
+Module ComUnitRing.
+
+Section Mixin.
+
+Variables (R : comRingType) (unit : pred R) (inv : R -> R).
+Hypothesis mulVx : {in unit, left_inverse 1 inv *%R}.
+Hypothesis unitPl : forall x y, y * x = 1 -> unit x.
+
+Fact mulC_mulrV : {in unit, right_inverse 1 inv *%R}.
+Proof. by move=> x Ux /=; rewrite mulrC mulVx. Qed.
+
+Fact mulC_unitP x y : y * x = 1 /\ x * y = 1 -> unit x.
+Proof. case=> yx _; exact: unitPl yx. Qed.
+
+Definition Mixin := UnitRingMixin mulVx mulC_mulrV mulC_unitP.
+
+End Mixin.
+
+Section ClassDef.
+
+Record class_of (R : Type) : Type := Class {
+ base : ComRing.class_of R;
+ mixin : UnitRing.mixin_of (Ring.Pack base R)
+}.
+Local Coercion base : class_of >-> ComRing.class_of.
+Definition base2 R m := UnitRing.Class (@mixin R m).
+Local Coercion base2 : class_of >-> UnitRing.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack :=
+ fun bT b & phant_id (ComRing.class bT) (b : ComRing.class_of T) =>
+ fun mT m & phant_id (UnitRing.class mT) (@UnitRing.Class T b m) =>
+ Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+Definition comRingType := @ComRing.Pack cT xclass xT.
+Definition unitRingType := @UnitRing.Pack cT xclass xT.
+Definition com_unitRingType := @UnitRing.Pack comRingType xclass xT.
+
+End ClassDef.
+
+Module Import Exports.
+Coercion base : class_of >-> ComRing.class_of.
+Coercion mixin : class_of >-> UnitRing.mixin_of.
+Coercion base2 : class_of >-> UnitRing.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> UnitRing.type.
+Canonical unitRingType.
+Canonical com_unitRingType.
+Notation comUnitRingType := type.
+Notation ComUnitRingMixin := Mixin.
+Notation "[ 'comUnitRingType' 'of' T ]" := (@pack T _ _ id _ _ id)
+ (at level 0, format "[ 'comUnitRingType' 'of' T ]") : form_scope.
+End Exports.
+
+End ComUnitRing.
+Import ComUnitRing.Exports.
+
+Module UnitAlgebra.
+
+Section ClassDef.
+
+Variable R : ringType.
+
+Record class_of (T : Type) : Type := Class {
+ base : Algebra.class_of R T;
+ mixin : GRing.UnitRing.mixin_of (Ring.Pack base T)
+}.
+Definition base2 R m := UnitRing.Class (@mixin R m).
+Local Coercion base : class_of >-> Algebra.class_of.
+Local Coercion base2 : class_of >-> UnitRing.class_of.
+
+Structure type (phR : phant R) := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variable (phR : phant R) (T : Type) (cT : type phR).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack :=
+ fun bT b & phant_id (@Algebra.class R phR bT) (b : Algebra.class_of R T) =>
+ fun mT m & phant_id (UnitRing.mixin (UnitRing.class mT)) m =>
+ Pack (Phant R) (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+Definition unitRingType := @UnitRing.Pack cT xclass xT.
+Definition lmodType := @Lmodule.Pack R phR cT xclass xT.
+Definition lalgType := @Lalgebra.Pack R phR cT xclass xT.
+Definition algType := @Algebra.Pack R phR cT xclass xT.
+Definition lmod_unitRingType := @Lmodule.Pack R phR unitRingType xclass xT.
+Definition lalg_unitRingType := @Lalgebra.Pack R phR unitRingType xclass xT.
+Definition alg_unitRingType := @Algebra.Pack R phR unitRingType xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> Algebra.class_of.
+Coercion base2 : class_of >-> UnitRing.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Coercion unitRingType : type >-> UnitRing.type.
+Canonical unitRingType.
+Coercion lmodType : type >-> Lmodule.type.
+Canonical lmodType.
+Coercion lalgType : type >-> Lalgebra.type.
+Canonical lalgType.
+Coercion algType : type >-> Algebra.type.
+Canonical algType.
+Canonical lmod_unitRingType.
+Canonical lalg_unitRingType.
+Canonical alg_unitRingType.
+Notation unitAlgType R := (type (Phant R)).
+Notation "[ 'unitAlgType' R 'of' T ]" := (@pack _ (Phant R) T _ _ id _ _ id)
+ (at level 0, format "[ 'unitAlgType' R 'of' T ]") : form_scope.
+End Exports.
+
+End UnitAlgebra.
+Import UnitAlgebra.Exports.
+
+Section ComUnitRingTheory.
+
+Variable R : comUnitRingType.
+Implicit Types x y : R.
+
+Lemma unitrM x y : (x * y \in unit) = (x \in unit) && (y \in unit).
+Proof. by apply: unitrM_comm; exact: mulrC. Qed.
+
+Lemma unitrPr x : reflect (exists y, x * y = 1) (x \in unit).
+Proof.
+by apply: (iffP (unitrP x)) => [[y []] | [y]]; exists y; rewrite // mulrC.
+Qed.
+
+Lemma expr_div_n x y n : (x / y) ^+ n = x ^+ n / y ^+ n.
+Proof. by rewrite exprMn exprVn. Qed.
+
+Canonical regular_comUnitRingType := [comUnitRingType of R^o].
+Canonical regular_unitAlgType := [unitAlgType R of R^o].
+
+End ComUnitRingTheory.
+
+Section UnitAlgebraTheory.
+
+Variable (R : comUnitRingType) (A : unitAlgType R).
+Implicit Types (k : R) (x y : A).
+
+Lemma scaler_injl : {in unit, @right_injective R A A *:%R}.
+Proof.
+move=> k Uk x1 x2 Hx1x2.
+by rewrite -[x1]scale1r -(mulVr Uk) -scalerA Hx1x2 scalerA mulVr // scale1r.
+Qed.
+
+Lemma scaler_unit k x : k \in unit -> (k *: x \in unit) = (x \in unit).
+Proof.
+move=> Uk; apply/idP/idP=> [Ukx | Ux]; apply/unitrP; last first.
+ exists (k^-1 *: x^-1).
+ by rewrite -!scalerAl -!scalerAr !scalerA !mulVr // !mulrV // scale1r.
+exists (k *: (k *: x)^-1); split.
+ apply: (mulrI Ukx).
+ by rewrite mulr1 mulrA -scalerAr mulrV // -scalerAl mul1r.
+apply: (mulIr Ukx).
+by rewrite mul1r -mulrA -scalerAl mulVr // -scalerAr mulr1.
+Qed.
+
+Lemma invrZ k x : k \in unit -> x \in unit -> (k *: x)^-1 = k^-1 *: x^-1.
+Proof.
+move=> Uk Ux; have Ukx: (k *: x \in unit) by rewrite scaler_unit.
+apply: (mulIr Ukx).
+by rewrite mulVr // -scalerAl -scalerAr scalerA !mulVr // scale1r.
+Qed.
+
+Section ClosedPredicates.
+
+Variables S : predPredType A.
+
+Definition divalg_closed := [/\ 1 \in S, linear_closed S & divr_2closed S].
+
+Lemma divalg_closedBdiv : divalg_closed -> divring_closed S.
+Proof. by case=> S1 /linear_closedB. Qed.
+
+Lemma divalg_closedZ : divalg_closed -> subalg_closed S.
+Proof. by case=> S1 Slin Sdiv; split=> //; have [] := @divr_closedM A S. Qed.
+
+End ClosedPredicates.
+
+End UnitAlgebraTheory.
+
+(* Interface structures for algebraically closed predicates. *)
+Module Pred.
+
+Structure opp V S := Opp {opp_key : pred_key S; _ : @oppr_closed V S}.
+Structure add V S := Add {add_key : pred_key S; _ : @addr_closed V S}.
+Structure mul R S := Mul {mul_key : pred_key S; _ : @mulr_closed R S}.
+Structure zmod V S := Zmod {zmod_add : add S; _ : @oppr_closed V S}.
+Structure semiring R S := Semiring {semiring_add : add S; _ : @mulr_closed R S}.
+Structure smul R S := Smul {smul_opp : opp S; _ : @mulr_closed R S}.
+Structure div R S := Div {div_mul : mul S; _ : @invr_closed R S}.
+Structure submod R V S :=
+ Submod {submod_zmod : zmod S; _ : @scaler_closed R V S}.
+Structure subring R S := Subring {subring_zmod : zmod S; _ : @mulr_closed R S}.
+Structure sdiv R S := Sdiv {sdiv_smul : smul S; _ : @invr_closed R S}.
+Structure subalg (R : ringType) (A : lalgType R) S :=
+ Subalg {subalg_ring : subring S; _ : @scaler_closed R A S}.
+Structure divring R S :=
+ Divring {divring_ring : subring S; _ : @invr_closed R S}.
+Structure divalg (R : ringType) (A : unitAlgType R) S :=
+ Divalg {divalg_ring : divring S; _ : @scaler_closed R A S}.
+
+Section Subtyping.
+
+Ltac done := case=> *; assumption.
+Fact zmod_oppr R S : @zmod R S -> oppr_closed S. Proof. by []. Qed.
+Fact semiring_mulr R S : @semiring R S -> mulr_closed S. Proof. by []. Qed.
+Fact smul_mulr R S : @smul R S -> mulr_closed S. Proof. by []. Qed.
+Fact submod_scaler R V S : @submod R V S -> scaler_closed S. Proof. by []. Qed.
+Fact subring_mulr R S : @subring R S -> mulr_closed S. Proof. by []. Qed.
+Fact sdiv_invr R S : @sdiv R S -> invr_closed S. Proof. by []. Qed.
+Fact subalg_scaler R A S : @subalg R A S -> scaler_closed S. Proof. by []. Qed.
+Fact divring_invr R S : @divring R S -> invr_closed S. Proof. by []. Qed.
+Fact divalg_scaler R A S : @divalg R A S -> scaler_closed S. Proof. by []. Qed.
+
+Definition zmod_opp R S (addS : @zmod R S) :=
+ Opp (add_key (zmod_add addS)) (zmod_oppr addS).
+Definition semiring_mul R S (ringS : @semiring R S) :=
+ Mul (add_key (semiring_add ringS)) (semiring_mulr ringS).
+Definition smul_mul R S (mulS : @smul R S) :=
+ Mul (opp_key (smul_opp mulS)) (smul_mulr mulS).
+Definition subring_semi R S (ringS : @subring R S) :=
+ Semiring (zmod_add (subring_zmod ringS)) (subring_mulr ringS).
+Definition subring_smul R S (ringS : @subring R S) :=
+ Smul (zmod_opp (subring_zmod ringS)) (subring_mulr ringS).
+Definition sdiv_div R S (divS : @sdiv R S) :=
+ Div (smul_mul (sdiv_smul divS)) (sdiv_invr divS).
+Definition subalg_submod R A S (algS : @subalg R A S) :=
+ Submod (subring_zmod (subalg_ring algS)) (subalg_scaler algS).
+Definition divring_sdiv R S (ringS : @divring R S) :=
+ Sdiv (subring_smul (divring_ring ringS)) (divring_invr ringS).
+Definition divalg_alg R A S (algS : @divalg R A S) :=
+ Subalg (divring_ring (divalg_ring algS)) (divalg_scaler algS).
+
+End Subtyping.
+
+Section Extensionality.
+(* This could be avoided by exploiting the Coq 8.4 eta-convertibility. *)
+
+Lemma opp_ext (U : zmodType) S k (kS : @keyed_pred U S k) :
+ oppr_closed kS -> oppr_closed S.
+Proof. by move=> oppS x; rewrite -!(keyed_predE kS); apply: oppS. Qed.
+
+Lemma add_ext (U : zmodType) S k (kS : @keyed_pred U S k) :
+ addr_closed kS -> addr_closed S.
+Proof.
+by case=> S0 addS; split=> [|x y]; rewrite -!(keyed_predE kS) //; apply: addS.
+Qed.
+
+Lemma mul_ext (R : ringType) S k (kS : @keyed_pred R S k) :
+ mulr_closed kS -> mulr_closed S.
+Proof.
+by case=> S1 mulS; split=> [|x y]; rewrite -!(keyed_predE kS) //; apply: mulS.
+Qed.
+
+Lemma scale_ext (R : ringType) (U : lmodType R) S k (kS : @keyed_pred U S k) :
+ scaler_closed kS -> scaler_closed S.
+Proof. by move=> linS a x; rewrite -!(keyed_predE kS); apply: linS. Qed.
+
+Lemma inv_ext (R : unitRingType) S k (kS : @keyed_pred R S k) :
+ invr_closed kS -> invr_closed S.
+Proof. by move=> invS x; rewrite -!(keyed_predE kS); apply: invS. Qed.
+
+End Extensionality.
+
+Module Default.
+Definition opp V S oppS := @Opp V S (DefaultPredKey S) oppS.
+Definition add V S addS := @Add V S (DefaultPredKey S) addS.
+Definition mul R S mulS := @Mul R S (DefaultPredKey S) mulS.
+Definition zmod V S addS oppS := @Zmod V S (add addS) oppS.
+Definition semiring R S addS mulS := @Semiring R S (add addS) mulS.
+Definition smul R S oppS mulS := @Smul R S (opp oppS) mulS.
+Definition div R S mulS invS := @Div R S (mul mulS) invS.
+Definition submod R V S addS oppS linS := @Submod R V S (zmod addS oppS) linS.
+Definition subring R S addS oppS mulS := @Subring R S (zmod addS oppS) mulS.
+Definition sdiv R S oppS mulS invS := @Sdiv R S (smul oppS mulS) invS.
+Definition subalg R A S addS oppS mulS linS :=
+ @Subalg R A S (subring addS oppS mulS) linS.
+Definition divring R S addS oppS mulS invS :=
+ @Divring R S (subring addS oppS mulS) invS.
+Definition divalg R A S addS oppS mulS invS linS :=
+ @Divalg R A S (divring addS oppS mulS invS) linS.
+End Default.
+
+Module Exports.
+
+Notation oppr_closed := oppr_closed.
+Notation addr_closed := addr_closed.
+Notation mulr_closed := mulr_closed.
+Notation zmod_closed := zmod_closed.
+Notation smulr_closed := smulr_closed.
+Notation invr_closed := invr_closed.
+Notation divr_closed := divr_closed.
+Notation linear_closed := linear_closed.
+Notation submod_closed := submod_closed.
+Notation semiring_closed := semiring_closed.
+Notation subring_closed := subring_closed.
+Notation sdivr_closed := sdivr_closed.
+Notation subalg_closed := subalg_closed.
+Notation divring_closed := divring_closed.
+Notation divalg_closed := divalg_closed.
+
+Coercion zmod_closedD : zmod_closed >-> addr_closed.
+Coercion zmod_closedN : zmod_closed >-> oppr_closed.
+Coercion smulr_closedN : smulr_closed >-> oppr_closed.
+Coercion smulr_closedM : smulr_closed >-> mulr_closed.
+Coercion divr_closedV : divr_closed >-> invr_closed.
+Coercion divr_closedM : divr_closed >-> mulr_closed.
+Coercion submod_closedZ : submod_closed >-> scaler_closed.
+Coercion submod_closedB : submod_closed >-> zmod_closed.
+Coercion semiring_closedD : semiring_closed >-> addr_closed.
+Coercion semiring_closedM : semiring_closed >-> mulr_closed.
+Coercion subring_closedB : subring_closed >-> zmod_closed.
+Coercion subring_closedM : subring_closed >-> smulr_closed.
+Coercion subring_closed_semi : subring_closed >-> semiring_closed.
+Coercion sdivr_closedM : sdivr_closed >-> smulr_closed.
+Coercion sdivr_closed_div : sdivr_closed >-> divr_closed.
+Coercion subalg_closedZ : subalg_closed >-> submod_closed.
+Coercion subalg_closedBM : subalg_closed >-> subring_closed.
+Coercion divring_closedBM : divring_closed >-> subring_closed.
+Coercion divring_closed_div : divring_closed >-> sdivr_closed.
+Coercion divalg_closedZ : divalg_closed >-> subalg_closed.
+Coercion divalg_closedBdiv : divalg_closed >-> divring_closed.
+
+Coercion opp_key : opp >-> pred_key.
+Coercion add_key : add >-> pred_key.
+Coercion mul_key : mul >-> pred_key.
+Coercion zmod_opp : zmod >-> opp.
+Canonical zmod_opp.
+Coercion zmod_add : zmod >-> add.
+Coercion semiring_add : semiring >-> add.
+Coercion semiring_mul : semiring >-> mul.
+Canonical semiring_mul.
+Coercion smul_opp : smul >-> opp.
+Coercion smul_mul : smul >-> mul.
+Canonical smul_mul.
+Coercion div_mul : div >-> mul.
+Coercion submod_zmod : submod >-> zmod.
+Coercion subring_zmod : subring >-> zmod.
+Coercion subring_semi : subring >-> semiring.
+Canonical subring_semi.
+Coercion subring_smul : subring >-> smul.
+Canonical subring_smul.
+Coercion sdiv_smul : sdiv >-> smul.
+Coercion sdiv_div : sdiv >-> div.
+Canonical sdiv_div.
+Coercion subalg_submod : subalg >-> submod.
+Canonical subalg_submod.
+Coercion subalg_ring : subalg >-> subring.
+Coercion divring_ring : divring >-> subring.
+Coercion divring_sdiv : divring >-> sdiv.
+Canonical divring_sdiv.
+Coercion divalg_alg : divalg >-> subalg.
+Canonical divalg_alg.
+Coercion divalg_ring : divalg >-> divring.
+
+Notation opprPred := opp.
+Notation addrPred := add.
+Notation mulrPred := mul.
+Notation zmodPred := zmod.
+Notation semiringPred := semiring.
+Notation smulrPred := smul.
+Notation divrPred := div.
+Notation submodPred := submod.
+Notation subringPred := subring.
+Notation sdivrPred := sdiv.
+Notation subalgPred := subalg.
+Notation divringPred := divring.
+Notation divalgPred := divalg.
+
+Definition OpprPred U S k kS NkS := Opp k (@opp_ext U S k kS NkS).
+Definition AddrPred U S k kS DkS := Add k (@add_ext U S k kS DkS).
+Definition MulrPred R S k kS MkS := Mul k (@mul_ext R S k kS MkS).
+Definition ZmodPred U S k kS NkS := Zmod k (@opp_ext U S k kS NkS).
+Definition SemiringPred R S k kS MkS := Semiring k (@mul_ext R S k kS MkS).
+Definition SmulrPred R S k kS MkS := Smul k (@mul_ext R S k kS MkS).
+Definition DivrPred R S k kS VkS := Div k (@inv_ext R S k kS VkS).
+Definition SubmodPred R U S k kS ZkS := Submod k (@scale_ext R U S k kS ZkS).
+Definition SubringPred R S k kS MkS := Subring k (@mul_ext R S k kS MkS).
+Definition SdivrPred R S k kS VkS := Sdiv k (@inv_ext R S k kS VkS).
+Definition SubalgPred (R : ringType) (A : lalgType R) S k kS ZkS :=
+ Subalg k (@scale_ext R A S k kS ZkS).
+Definition DivringPred R S k kS VkS := Divring k (@inv_ext R S k kS VkS).
+Definition DivalgPred (R : ringType) (A : unitAlgType R) S k kS ZkS :=
+ Divalg k (@scale_ext R A S k kS ZkS).
+
+End Exports.
+
+End Pred.
+Import Pred.Exports.
+
+Module DefaultPred.
+
+Canonical Pred.Default.opp.
+Canonical Pred.Default.add.
+Canonical Pred.Default.mul.
+Canonical Pred.Default.zmod.
+Canonical Pred.Default.semiring.
+Canonical Pred.Default.smul.
+Canonical Pred.Default.div.
+Canonical Pred.Default.submod.
+Canonical Pred.Default.subring.
+Canonical Pred.Default.sdiv.
+Canonical Pred.Default.subalg.
+Canonical Pred.Default.divring.
+Canonical Pred.Default.divalg.
+
+End DefaultPred.
+
+Section ZmodulePred.
+
+Variables (V : zmodType) (S : predPredType V).
+
+Section Add.
+
+Variables (addS : addrPred S) (kS : keyed_pred addS).
+
+Lemma rpred0D : addr_closed kS.
+Proof.
+by split=> [|x y]; rewrite !keyed_predE; case: addS => _ [_]//; apply.
+Qed.
+
+Lemma rpred0 : 0 \in kS.
+Proof. by case: rpred0D. Qed.
+
+Lemma rpredD : {in kS &, forall u v, u + v \in kS}.
+Proof. by case: rpred0D. Qed.
+
+Lemma rpred_sum I r (P : pred I) F :
+ (forall i, P i -> F i \in kS) -> \sum_(i <- r | P i) F i \in kS.
+Proof. by move=> IH; elim/big_ind: _; [exact: rpred0 | exact: rpredD |]. Qed.
+
+Lemma rpredMn n : {in kS, forall u, u *+ n \in kS}.
+Proof. by move=> u Su; rewrite -(card_ord n) -sumr_const rpred_sum. Qed.
+
+End Add.
+
+Section Opp.
+
+Variables (oppS : opprPred S) (kS : keyed_pred oppS).
+
+Lemma rpredNr : oppr_closed kS.
+Proof. by move=> x; rewrite !keyed_predE; case: oppS => _; apply. Qed.
+
+Lemma rpredN : {mono -%R: u / u \in kS}.
+Proof. by move=> u; apply/idP/idP=> /rpredNr; rewrite ?opprK; apply. Qed.
+
+End Opp.
+
+Section Sub.
+
+Variables (subS : zmodPred S) (kS : keyed_pred subS).
+
+Lemma rpredB : {in kS &, forall u v, u - v \in kS}.
+Proof. by move=> u v Su Sv; rewrite /= rpredD ?rpredN. Qed.
+
+Lemma rpredMNn n : {in kS, forall u, u *- n \in kS}.
+Proof. by move=> u Su; rewrite /= rpredN rpredMn. Qed.
+
+Lemma rpredDr x y : x \in kS -> (y + x \in kS) = (y \in kS).
+Proof.
+move=> Sx; apply/idP/idP=> [Sxy | /rpredD-> //].
+by rewrite -(addrK x y) rpredB.
+Qed.
+
+Lemma rpredDl x y : x \in kS -> (x + y \in kS) = (y \in kS).
+Proof. by rewrite addrC; apply: rpredDr. Qed.
+
+Lemma rpredBr x y : x \in kS -> (y - x \in kS) = (y \in kS).
+Proof. by rewrite -rpredN; apply: rpredDr. Qed.
+
+Lemma rpredBl x y : x \in kS -> (x - y \in kS) = (y \in kS).
+Proof. by rewrite -(rpredN _ y); apply: rpredDl. Qed.
+
+End Sub.
+
+End ZmodulePred.
+
+Section RingPred.
+
+Variables (R : ringType) (S : predPredType R).
+
+Lemma rpredMsign (oppS : opprPred S) (kS : keyed_pred oppS) n x :
+ ((-1) ^+ n * x \in kS) = (x \in kS).
+Proof. by rewrite -signr_odd mulr_sign; case: ifP => // _; rewrite rpredN. Qed.
+
+Section Mul.
+
+Variables (mulS : mulrPred S) (kS : keyed_pred mulS).
+
+Lemma rpred1M : mulr_closed kS.
+Proof.
+by split=> [|x y]; rewrite !keyed_predE; case: mulS => _ [_] //; apply.
+Qed.
+
+Lemma rpred1 : 1 \in kS.
+Proof. by case: rpred1M. Qed.
+
+Lemma rpredM : {in kS &, forall u v, u * v \in kS}.
+Proof. by case: rpred1M. Qed.
+
+Lemma rpred_prod I r (P : pred I) F :
+ (forall i, P i -> F i \in kS) -> \prod_(i <- r | P i) F i \in kS.
+Proof. by move=> IH; elim/big_ind: _; [exact: rpred1 | exact: rpredM |]. Qed.
+
+Lemma rpredX n : {in kS, forall u, u ^+ n \in kS}.
+Proof. by move=> u Su; rewrite -(card_ord n) -prodr_const rpred_prod. Qed.
+
+End Mul.
+
+Lemma rpred_nat (rngS : semiringPred S) (kS : keyed_pred rngS) n : n%:R \in kS.
+Proof. by rewrite rpredMn ?rpred1. Qed.
+
+Lemma rpredN1 (mulS : smulrPred S) (kS : keyed_pred mulS) : -1 \in kS.
+Proof. by rewrite rpredN rpred1. Qed.
+
+Lemma rpred_sign (mulS : smulrPred S) (kS : keyed_pred mulS) n :
+ (-1) ^+ n \in kS.
+Proof. by rewrite rpredX ?rpredN1. Qed.
+
+End RingPred.
+
+Section LmodPred.
+
+Variables (R : ringType) (V : lmodType R) (S : predPredType V).
+
+Lemma rpredZsign (oppS : opprPred S) (kS : keyed_pred oppS) n u :
+ ((-1) ^+ n *: u \in kS) = (u \in kS).
+Proof. by rewrite -signr_odd scaler_sign fun_if if_arg rpredN if_same. Qed.
+
+Lemma rpredZnat (addS : addrPred S) (kS : keyed_pred addS) n :
+ {in kS, forall u, n%:R *: u \in kS}.
+Proof. by move=> u Su; rewrite /= scaler_nat rpredMn. Qed.
+
+Lemma rpredZ (linS : submodPred S) (kS : keyed_pred linS) : scaler_closed kS.
+Proof. by move=> a u; rewrite !keyed_predE; case: {kS}linS => _; apply. Qed.
+
+End LmodPred.
+
+Section UnitRingPred.
+
+Variable R : unitRingType.
+
+Section Div.
+
+Variables (S : predPredType R) (divS : divrPred S) (kS : keyed_pred divS).
+
+Lemma rpredVr x : x \in kS -> x^-1 \in kS.
+Proof. by rewrite !keyed_predE; case: divS x. Qed.
+
+Lemma rpredV x : (x^-1 \in kS) = (x \in kS).
+Proof. by apply/idP/idP=> /rpredVr; rewrite ?invrK. Qed.
+
+Lemma rpred_div : {in kS &, forall x y, x / y \in kS}.
+Proof. by move=> x y Sx Sy; rewrite /= rpredM ?rpredV. Qed.
+
+Lemma rpredXN n : {in kS, forall x, x ^- n \in kS}.
+Proof. by move=> x Sx; rewrite /= rpredV rpredX. Qed.
+
+Lemma rpredMl x y : x \in kS -> x \is a unit-> (x * y \in kS) = (y \in kS).
+Proof.
+move=> Sx Ux; apply/idP/idP=> [Sxy | /(rpredM Sx)-> //].
+by rewrite -(mulKr Ux y); rewrite rpredM ?rpredV.
+Qed.
+
+Lemma rpredMr x y : x \in kS -> x \is a unit -> (y * x \in kS) = (y \in kS).
+Proof.
+move=> Sx Ux; apply/idP/idP=> [Sxy | /rpredM-> //].
+by rewrite -(mulrK Ux y); rewrite rpred_div.
+Qed.
+
+Lemma rpred_divr x y : x \in kS -> x \is a unit -> (y / x \in kS) = (y \in kS).
+Proof. by rewrite -rpredV -unitrV; apply: rpredMr. Qed.
+
+Lemma rpred_divl x y : x \in kS -> x \is a unit -> (x / y \in kS) = (y \in kS).
+Proof. by rewrite -(rpredV y); apply: rpredMl. Qed.
+
+End Div.
+
+Fact unitr_sdivr_closed : @sdivr_closed R unit.
+Proof. by split=> [|x y Ux Uy]; rewrite ?unitrN1 // unitrMl ?unitrV. Qed.
+
+Canonical unit_opprPred := OpprPred unitr_sdivr_closed.
+Canonical unit_mulrPred := MulrPred unitr_sdivr_closed.
+Canonical unit_divrPred := DivrPred unitr_sdivr_closed.
+Canonical unit_smulrPred := SmulrPred unitr_sdivr_closed.
+Canonical unit_sdivrPred := SdivrPred unitr_sdivr_closed.
+
+Implicit Type x : R.
+
+Lemma unitrN x : (- x \is a unit) = (x \is a unit). Proof. exact: rpredN. Qed.
+
+Lemma invrN x : (- x)^-1 = - x^-1.
+Proof.
+have [Ux | U'x] := boolP (x \is a unit); last by rewrite !invr_out ?unitrN.
+by rewrite -mulN1r invrM ?unitrN1 // invrN1 mulrN1.
+Qed.
+
+Lemma invr_signM n x : ((-1) ^+ n * x)^-1 = (-1) ^+ n * x^-1.
+Proof. by rewrite -signr_odd !mulr_sign; case: ifP => // _; rewrite invrN. Qed.
+
+Lemma divr_signM (b1 b2 : bool) x1 x2:
+ ((-1) ^+ b1 * x1) / ((-1) ^+ b2 * x2) = (-1) ^+ (b1 (+) b2) * (x1 / x2).
+Proof. by rewrite invr_signM mulr_signM. Qed.
+
+End UnitRingPred.
+
+(* Reification of the theory of rings with units, in named style *)
+Section TermDef.
+
+Variable R : Type.
+
+Inductive term : Type :=
+| Var of nat
+| Const of R
+| NatConst of nat
+| Add of term & term
+| Opp of term
+| NatMul of term & nat
+| Mul of term & term
+| Inv of term
+| Exp of term & nat.
+
+Inductive formula : Type :=
+| Bool of bool
+| Equal of term & term
+| Unit of term
+| And of formula & formula
+| Or of formula & formula
+| Implies of formula & formula
+| Not of formula
+| Exists of nat & formula
+| Forall of nat & formula.
+
+End TermDef.
+
+Bind Scope term_scope with term.
+Bind Scope term_scope with formula.
+Arguments Scope Add [_ term_scope term_scope].
+Arguments Scope Opp [_ term_scope].
+Arguments Scope NatMul [_ term_scope nat_scope].
+Arguments Scope Mul [_ term_scope term_scope].
+Arguments Scope Mul [_ term_scope term_scope].
+Arguments Scope Inv [_ term_scope].
+Arguments Scope Exp [_ term_scope nat_scope].
+Arguments Scope Equal [_ term_scope term_scope].
+Arguments Scope Unit [_ term_scope].
+Arguments Scope And [_ term_scope term_scope].
+Arguments Scope Or [_ term_scope term_scope].
+Arguments Scope Implies [_ term_scope term_scope].
+Arguments Scope Not [_ term_scope].
+Arguments Scope Exists [_ nat_scope term_scope].
+Arguments Scope Forall [_ nat_scope term_scope].
+
+Implicit Arguments Bool [R].
+Prenex Implicits Const Add Opp NatMul Mul Exp Bool Unit And Or Implies Not.
+Prenex Implicits Exists Forall.
+
+Notation True := (Bool true).
+Notation False := (Bool false).
+
+Local Notation "''X_' i" := (Var _ i) : term_scope.
+Local Notation "n %:R" := (NatConst _ n) : term_scope.
+Local Notation "x %:T" := (Const x) : term_scope.
+Local Notation "0" := 0%:R%T : term_scope.
+Local Notation "1" := 1%:R%T : term_scope.
+Local Infix "+" := Add : term_scope.
+Local Notation "- t" := (Opp t) : term_scope.
+Local Notation "t - u" := (Add t (- u)) : term_scope.
+Local Infix "*" := Mul : term_scope.
+Local Infix "*+" := NatMul : term_scope.
+Local Notation "t ^-1" := (Inv t) : term_scope.
+Local Notation "t / u" := (Mul t u^-1) : term_scope.
+Local Infix "^+" := Exp : term_scope.
+Local Infix "==" := Equal : term_scope.
+Local Infix "/\" := And : term_scope.
+Local Infix "\/" := Or : term_scope.
+Local Infix "==>" := Implies : term_scope.
+Local Notation "~ f" := (Not f) : term_scope.
+Local Notation "x != y" := (Not (x == y)) : term_scope.
+Local Notation "''exists' ''X_' i , f" := (Exists i f) : term_scope.
+Local Notation "''forall' ''X_' i , f" := (Forall i f) : term_scope.
+
+Section Substitution.
+
+Variable R : Type.
+
+Fixpoint tsubst (t : term R) (s : nat * term R) :=
+ match t with
+ | 'X_i => if i == s.1 then s.2 else t
+ | _%:T | _%:R => t
+ | t1 + t2 => tsubst t1 s + tsubst t2 s
+ | - t1 => - tsubst t1 s
+ | t1 *+ n => tsubst t1 s *+ n
+ | t1 * t2 => tsubst t1 s * tsubst t2 s
+ | t1^-1 => (tsubst t1 s)^-1
+ | t1 ^+ n => tsubst t1 s ^+ n
+ end%T.
+
+Fixpoint fsubst (f : formula R) (s : nat * term R) :=
+ match f with
+ | Bool _ => f
+ | t1 == t2 => tsubst t1 s == tsubst t2 s
+ | Unit t1 => Unit (tsubst t1 s)
+ | f1 /\ f2 => fsubst f1 s /\ fsubst f2 s
+ | f1 \/ f2 => fsubst f1 s \/ fsubst f2 s
+ | f1 ==> f2 => fsubst f1 s ==> fsubst f2 s
+ | ~ f1 => ~ fsubst f1 s
+ | ('exists 'X_i, f1) => 'exists 'X_i, if i == s.1 then f1 else fsubst f1 s
+ | ('forall 'X_i, f1) => 'forall 'X_i, if i == s.1 then f1 else fsubst f1 s
+ end%T.
+
+End Substitution.
+
+Section EvalTerm.
+
+Variable R : unitRingType.
+
+(* Evaluation of a reified term into R a ring with units *)
+Fixpoint eval (e : seq R) (t : term R) {struct t} : R :=
+ match t with
+ | ('X_i)%T => e`_i
+ | (x%:T)%T => x
+ | (n%:R)%T => n%:R
+ | (t1 + t2)%T => eval e t1 + eval e t2
+ | (- t1)%T => - eval e t1
+ | (t1 *+ n)%T => eval e t1 *+ n
+ | (t1 * t2)%T => eval e t1 * eval e t2
+ | t1^-1%T => (eval e t1)^-1
+ | (t1 ^+ n)%T => eval e t1 ^+ n
+ end.
+
+Definition same_env (e e' : seq R) := nth 0 e =1 nth 0 e'.
+
+Lemma eq_eval e e' t : same_env e e' -> eval e t = eval e' t.
+Proof. by move=> eq_e; elim: t => //= t1 -> // t2 ->. Qed.
+
+Lemma eval_tsubst e t s :
+ eval e (tsubst t s) = eval (set_nth 0 e s.1 (eval e s.2)) t.
+Proof.
+case: s => i u; elim: t => //=; do 2?[move=> ? -> //] => j.
+by rewrite nth_set_nth /=; case: (_ == _).
+Qed.
+
+(* Evaluation of a reified formula *)
+Fixpoint holds (e : seq R) (f : formula R) {struct f} : Prop :=
+ match f with
+ | Bool b => b
+ | (t1 == t2)%T => eval e t1 = eval e t2
+ | Unit t1 => eval e t1 \in unit
+ | (f1 /\ f2)%T => holds e f1 /\ holds e f2
+ | (f1 \/ f2)%T => holds e f1 \/ holds e f2
+ | (f1 ==> f2)%T => holds e f1 -> holds e f2
+ | (~ f1)%T => ~ holds e f1
+ | ('exists 'X_i, f1)%T => exists x, holds (set_nth 0 e i x) f1
+ | ('forall 'X_i, f1)%T => forall x, holds (set_nth 0 e i x) f1
+ end.
+
+Lemma same_env_sym e e' : same_env e e' -> same_env e' e.
+Proof. exact: fsym. Qed.
+
+(* Extensionality of formula evaluation *)
+Lemma eq_holds e e' f : same_env e e' -> holds e f -> holds e' f.
+Proof.
+pose sv := set_nth (0 : R).
+have eq_i i v e1 e2: same_env e1 e2 -> same_env (sv e1 i v) (sv e2 i v).
+ by move=> eq_e j; rewrite !nth_set_nth /= eq_e.
+elim: f e e' => //=.
+- by move=> t1 t2 e e' eq_e; rewrite !(eq_eval _ eq_e).
+- by move=> t e e' eq_e; rewrite (eq_eval _ eq_e).
+- by move=> f1 IH1 f2 IH2 e e' eq_e; move/IH2: (eq_e); move/IH1: eq_e; tauto.
+- by move=> f1 IH1 f2 IH2 e e' eq_e; move/IH2: (eq_e); move/IH1: eq_e; tauto.
+- by move=> f1 IH1 f2 IH2 e e' eq_e f12; move/IH1: (same_env_sym eq_e); eauto.
+- by move=> f1 IH1 e e'; move/same_env_sym; move/IH1; tauto.
+- by move=> i f1 IH1 e e'; move/(eq_i i)=> eq_e [x f_ex]; exists x; eauto.
+by move=> i f1 IH1 e e'; move/(eq_i i); eauto.
+Qed.
+
+(* Evaluation and substitution by a constant *)
+Lemma holds_fsubst e f i v :
+ holds e (fsubst f (i, v%:T)%T) <-> holds (set_nth 0 e i v) f.
+Proof.
+elim: f e => //=; do [
+ by move=> *; rewrite !eval_tsubst
+| move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto
+| move=> f IHf e; move: (IHf e); tauto
+| move=> j f IHf e].
+- case eq_ji: (j == i); first rewrite (eqP eq_ji).
+ by split=> [] [x f_x]; exists x; rewrite set_set_nth eqxx in f_x *.
+ split=> [] [x f_x]; exists x; move: f_x; rewrite set_set_nth eq_sym eq_ji;
+ have:= IHf (set_nth 0 e j x); tauto.
+case eq_ji: (j == i); first rewrite (eqP eq_ji).
+ by split=> [] f_ x; move: (f_ x); rewrite set_set_nth eqxx.
+split=> [] f_ x; move: (IHf (set_nth 0 e j x)) (f_ x);
+ by rewrite set_set_nth eq_sym eq_ji; tauto.
+Qed.
+
+(* Boolean test selecting terms in the language of rings *)
+Fixpoint rterm (t : term R) :=
+ match t with
+ | _^-1 => false
+ | t1 + t2 | t1 * t2 => rterm t1 && rterm t2
+ | - t1 | t1 *+ _ | t1 ^+ _ => rterm t1
+ | _ => true
+ end%T.
+
+(* Boolean test selecting formulas in the theory of rings *)
+Fixpoint rformula (f : formula R) :=
+ match f with
+ | Bool _ => true
+ | t1 == t2 => rterm t1 && rterm t2
+ | Unit t1 => false
+ | f1 /\ f2 | f1 \/ f2 | f1 ==> f2 => rformula f1 && rformula f2
+ | ~ f1 | ('exists 'X__, f1) | ('forall 'X__, f1) => rformula f1
+ end%T.
+
+(* Upper bound of the names used in a term *)
+Fixpoint ub_var (t : term R) :=
+ match t with
+ | 'X_i => i.+1
+ | t1 + t2 | t1 * t2 => maxn (ub_var t1) (ub_var t2)
+ | - t1 | t1 *+ _ | t1 ^+ _ | t1^-1 => ub_var t1
+ | _ => 0%N
+ end%T.
+
+(* Replaces inverses in the term t by fresh variables, accumulating the *)
+(* substitution. *)
+Fixpoint to_rterm (t : term R) (r : seq (term R)) (n : nat) {struct t} :=
+ match t with
+ | t1^-1 =>
+ let: (t1', r1) := to_rterm t1 r n in
+ ('X_(n + size r1), rcons r1 t1')
+ | t1 + t2 =>
+ let: (t1', r1) := to_rterm t1 r n in
+ let: (t2', r2) := to_rterm t2 r1 n in
+ (t1' + t2', r2)
+ | - t1 =>
+ let: (t1', r1) := to_rterm t1 r n in
+ (- t1', r1)
+ | t1 *+ m =>
+ let: (t1', r1) := to_rterm t1 r n in
+ (t1' *+ m, r1)
+ | t1 * t2 =>
+ let: (t1', r1) := to_rterm t1 r n in
+ let: (t2', r2) := to_rterm t2 r1 n in
+ (Mul t1' t2', r2)
+ | t1 ^+ m =>
+ let: (t1', r1) := to_rterm t1 r n in
+ (t1' ^+ m, r1)
+ | _ => (t, r)
+ end%T.
+
+Lemma to_rterm_id t r n : rterm t -> to_rterm t r n = (t, r).
+Proof.
+elim: t r n => //.
+- by move=> t1 IHt1 t2 IHt2 r n /= /andP[rt1 rt2]; rewrite {}IHt1 // IHt2.
+- by move=> t IHt r n /= rt; rewrite {}IHt.
+- by move=> t IHt r n m /= rt; rewrite {}IHt.
+- by move=> t1 IHt1 t2 IHt2 r n /= /andP[rt1 rt2]; rewrite {}IHt1 // IHt2.
+- by move=> t IHt r n m /= rt; rewrite {}IHt.
+Qed.
+
+(* A ring formula stating that t1 is equal to 0 in the ring theory. *)
+(* Also applies to non commutative rings. *)
+Definition eq0_rform t1 :=
+ let m := ub_var t1 in
+ let: (t1', r1) := to_rterm t1 [::] m in
+ let fix loop r i := match r with
+ | [::] => t1' == 0
+ | t :: r' =>
+ let f := 'X_i * t == 1 /\ t * 'X_i == 1 in
+ 'forall 'X_i, (f \/ 'X_i == t /\ ~ ('exists 'X_i, f)) ==> loop r' i.+1
+ end%T
+ in loop r1 m.
+
+(* Transformation of a formula in the theory of rings with units into an *)
+(* equivalent formula in the sub-theory of rings. *)
+Fixpoint to_rform f :=
+ match f with
+ | Bool b => f
+ | t1 == t2 => eq0_rform (t1 - t2)
+ | Unit t1 => eq0_rform (t1 * t1^-1 - 1)
+ | f1 /\ f2 => to_rform f1 /\ to_rform f2
+ | f1 \/ f2 => to_rform f1 \/ to_rform f2
+ | f1 ==> f2 => to_rform f1 ==> to_rform f2
+ | ~ f1 => ~ to_rform f1
+ | ('exists 'X_i, f1) => 'exists 'X_i, to_rform f1
+ | ('forall 'X_i, f1) => 'forall 'X_i, to_rform f1
+ end%T.
+
+(* The transformation gives a ring formula. *)
+Lemma to_rform_rformula f : rformula (to_rform f).
+Proof.
+suffices eq0_ring t1: rformula (eq0_rform t1) by elim: f => //= => f1 ->.
+rewrite /eq0_rform; move: (ub_var t1) => m; set tr := _ m.
+suffices: all rterm (tr.1 :: tr.2).
+ case: tr => {t1} t1 r /= /andP[t1_r].
+ by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; exact: IHr.
+have: all rterm [::] by [].
+rewrite {}/tr; elim: t1 [::] => //=.
+- move=> t1 IHt1 t2 IHt2 r.
+ move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /= /andP[t1_r].
+ move/IHt2; case: to_rterm => {t2 r IHt2} t2 r /= /andP[t2_r].
+ by rewrite t1_r t2_r.
+- by move=> t1 IHt1 r /IHt1; case: to_rterm.
+- by move=> t1 IHt1 n r /IHt1; case: to_rterm.
+- move=> t1 IHt1 t2 IHt2 r.
+ move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /= /andP[t1_r].
+ move/IHt2; case: to_rterm => {t2 r IHt2} t2 r /= /andP[t2_r].
+ by rewrite t1_r t2_r.
+- move=> t1 IHt1 r.
+ by move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /=; rewrite all_rcons.
+- by move=> t1 IHt1 n r /IHt1; case: to_rterm.
+Qed.
+
+(* Correctness of the transformation. *)
+Lemma to_rformP e f : holds e (to_rform f) <-> holds e f.
+Proof.
+suffices{e f} equal0_equiv e t1 t2:
+ holds e (eq0_rform (t1 - t2)) <-> (eval e t1 == eval e t2).
+- elim: f e => /=; try tauto.
+ + move=> t1 t2 e.
+ by split; [move/equal0_equiv/eqP | move/eqP/equal0_equiv].
+ + move=> t1 e; rewrite unitrE; exact: equal0_equiv.
+ + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto.
+ + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto.
+ + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto.
+ + move=> f1 IHf1 e; move: (IHf1 e); tauto.
+ + by move=> n f1 IHf1 e; split=> [] [x] /IHf1; exists x.
+ + by move=> n f1 IHf1 e; split=> Hx x; apply/IHf1.
+rewrite -(add0r (eval e t2)) -(can2_eq (subrK _) (addrK _)).
+rewrite -/(eval e (t1 - t2)); move: (t1 - t2)%T => {t1 t2} t.
+have sub_var_tsubst s t0: s.1 >= ub_var t0 -> tsubst t0 s = t0.
+ elim: t0 {t} => //=.
+ - by move=> n; case: ltngtP.
+ - by move=> t1 IHt1 t2 IHt2; rewrite geq_max => /andP[/IHt1-> /IHt2->].
+ - by move=> t1 IHt1 /IHt1->.
+ - by move=> t1 IHt1 n /IHt1->.
+ - by move=> t1 IHt1 t2 IHt2; rewrite geq_max => /andP[/IHt1-> /IHt2->].
+ - by move=> t1 IHt1 /IHt1->.
+ - by move=> t1 IHt1 n /IHt1->.
+pose fix rsub t' m r : term R :=
+ if r is u :: r' then tsubst (rsub t' m.+1 r') (m, u^-1)%T else t'.
+pose fix ub_sub m r : Prop :=
+ if r is u :: r' then ub_var u <= m /\ ub_sub m.+1 r' else true.
+suffices{t} rsub_to_r t r0 m: m >= ub_var t -> ub_sub m r0 ->
+ let: (t', r) := to_rterm t r0 m in
+ [/\ take (size r0) r = r0,
+ ub_var t' <= m + size r, ub_sub m r & rsub t' m r = t].
+- have:= rsub_to_r t [::] _ (leqnn _); rewrite /eq0_rform.
+ case: (to_rterm _ _ _) => [t1' r1] [//|_ _ ub_r1 def_t].
+ rewrite -{2}def_t {def_t}.
+ elim: r1 (ub_var t) e ub_r1 => [|u r1 IHr1] m e /= => [_|[ub_u ub_r1]].
+ by split=> /eqP.
+ rewrite eval_tsubst /=; set y := eval e u; split=> t_eq0.
+ apply/IHr1=> //; apply: t_eq0.
+ rewrite nth_set_nth /= eqxx -(eval_tsubst e u (m, Const _)).
+ rewrite sub_var_tsubst //= -/y.
+ case Uy: (y \in unit); [left | right]; first by rewrite mulVr ?divrr.
+ split=> [|[z]]; first by rewrite invr_out ?Uy.
+ rewrite nth_set_nth /= eqxx.
+ rewrite -!(eval_tsubst _ _ (m, Const _)) !sub_var_tsubst // -/y => yz1.
+ by case/unitrP: Uy; exists z.
+ move=> x def_x; apply/IHr1=> //; suff ->: x = y^-1 by []; move: def_x.
+ rewrite nth_set_nth /= eqxx -(eval_tsubst e u (m, Const _)).
+ rewrite sub_var_tsubst //= -/y; case=> [[xy1 yx1] | [xy nUy]].
+ by rewrite -[y^-1]mul1r -[1]xy1 mulrK //; apply/unitrP; exists x.
+ rewrite invr_out //; apply/unitrP=> [[z yz1]]; case: nUy; exists z.
+ rewrite nth_set_nth /= eqxx -!(eval_tsubst _ _ (m, _%:T)%T).
+ by rewrite !sub_var_tsubst.
+have rsub_id r t0 n: ub_var t0 <= n -> rsub t0 n r = t0.
+ by elim: r n => //= t1 r IHr n let0n; rewrite IHr ?sub_var_tsubst ?leqW.
+have rsub_acc r s t1 m1:
+ ub_var t1 <= m1 + size r -> rsub t1 m1 (r ++ s) = rsub t1 m1 r.
+ elim: r t1 m1 => [|t1 r IHr] t2 m1 /=; first by rewrite addn0; apply: rsub_id.
+ by move=> letmr; rewrite IHr ?addSnnS.
+elim: t r0 m => /=; try do [
+ by move=> n r m hlt hub; rewrite take_size (ltn_addr _ hlt) rsub_id
+| by move=> n r m hlt hub; rewrite leq0n take_size rsub_id
+| move=> t1 IHt1 t2 IHt2 r m; rewrite geq_max; case/andP=> hub1 hub2 hmr;
+ case: to_rterm {IHt1 hub1 hmr}(IHt1 r m hub1 hmr) => t1' r1;
+ case=> htake1 hub1' hsub1 <-;
+ case: to_rterm {IHt2 hub2 hsub1}(IHt2 r1 m hub2 hsub1) => t2' r2 /=;
+ rewrite geq_max; case=> htake2 -> hsub2 /= <-;
+ rewrite -{1 2}(cat_take_drop (size r1) r2) htake2; set r3 := drop _ _;
+ rewrite size_cat addnA (leq_trans _ (leq_addr _ _)) //;
+ split=> {hsub2}//;
+ first by [rewrite takel_cat // -htake1 size_take geq_min leqnn orbT];
+ rewrite -(rsub_acc r1 r3 t1') {hub1'}// -{htake1}htake2 {r3}cat_take_drop;
+ by elim: r2 m => //= u r2 IHr2 m; rewrite IHr2
+| do [ move=> t1 IHt1 r m; do 2!move/IHt1=> {IHt1}IHt1
+ | move=> t1 IHt1 n r m; do 2!move/IHt1=> {IHt1}IHt1];
+ case: to_rterm IHt1 => t1' r1 [-> -> hsub1 <-]; split=> {hsub1}//;
+ by elim: r1 m => //= u r1 IHr1 m; rewrite IHr1].
+move=> t1 IH r m letm /IH {IH} /(_ letm) {letm}.
+case: to_rterm => t1' r1 /= [def_r ub_t1' ub_r1 <-].
+rewrite size_rcons addnS leqnn -{1}cats1 takel_cat ?def_r; last first.
+ by rewrite -def_r size_take geq_min leqnn orbT.
+elim: r1 m ub_r1 ub_t1' {def_r} => /= [|u r1 IHr1] m => [_|[->]].
+ by rewrite addn0 eqxx.
+by rewrite -addSnnS => /IHr1 IH /IH[_ _ ub_r1 ->].
+Qed.
+
+(* Boolean test selecting formulas which describe a constructible set, *)
+(* i.e. formulas without quantifiers. *)
+
+(* The quantifier elimination check. *)
+Fixpoint qf_form (f : formula R) :=
+ match f with
+ | Bool _ | _ == _ | Unit _ => true
+ | f1 /\ f2 | f1 \/ f2 | f1 ==> f2 => qf_form f1 && qf_form f2
+ | ~ f1 => qf_form f1
+ | _ => false
+ end%T.
+
+(* Boolean holds predicate for quantifier free formulas *)
+Definition qf_eval e := fix loop (f : formula R) : bool :=
+ match f with
+ | Bool b => b
+ | t1 == t2 => (eval e t1 == eval e t2)%bool
+ | Unit t1 => eval e t1 \in unit
+ | f1 /\ f2 => loop f1 && loop f2
+ | f1 \/ f2 => loop f1 || loop f2
+ | f1 ==> f2 => (loop f1 ==> loop f2)%bool
+ | ~ f1 => ~~ loop f1
+ |_ => false
+ end%T.
+
+(* qf_eval is equivalent to holds *)
+Lemma qf_evalP e f : qf_form f -> reflect (holds e f) (qf_eval e f).
+Proof.
+elim: f => //=; try by move=> *; exact: idP.
+- move=> t1 t2 _; exact: eqP.
+- move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1T]; last by right; case.
+ by case/IHf2; [left | right; case].
+- move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1F]; first by do 2 left.
+ by case/IHf2; [left; right | right; case].
+- move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1T]; last by left.
+ by case/IHf2; [left | right; move/(_ f1T)].
+by move=> f1 IHf1 /IHf1[]; [right | left].
+Qed.
+
+Implicit Type bc : seq (term R) * seq (term R).
+
+(* Quantifier-free formula are normalized into DNF. A DNF is *)
+(* represented by the type seq (seq (term R) * seq (term R)), where we *)
+(* separate positive and negative literals *)
+
+(* DNF preserving conjunction *)
+Definition and_dnf bcs1 bcs2 :=
+ \big[cat/nil]_(bc1 <- bcs1)
+ map (fun bc2 => (bc1.1 ++ bc2.1, bc1.2 ++ bc2.2)) bcs2.
+
+(* Computes a DNF from a qf ring formula *)
+Fixpoint qf_to_dnf (f : formula R) (neg : bool) {struct f} :=
+ match f with
+ | Bool b => if b (+) neg then [:: ([::], [::])] else [::]
+ | t1 == t2 => [:: if neg then ([::], [:: t1 - t2]) else ([:: t1 - t2], [::])]
+ | f1 /\ f2 => (if neg then cat else and_dnf) [rec f1, neg] [rec f2, neg]
+ | f1 \/ f2 => (if neg then and_dnf else cat) [rec f1, neg] [rec f2, neg]
+ | f1 ==> f2 => (if neg then and_dnf else cat) [rec f1, ~~ neg] [rec f2, neg]
+ | ~ f1 => [rec f1, ~~ neg]
+ | _ => if neg then [:: ([::], [::])] else [::]
+ end%T where "[ 'rec' f , neg ]" := (qf_to_dnf f neg).
+
+(* Conversely, transforms a DNF into a formula *)
+Definition dnf_to_form :=
+ let pos_lit t := And (t == 0) in let neg_lit t := And (t != 0) in
+ let cls bc := Or (foldr pos_lit True bc.1 /\ foldr neg_lit True bc.2) in
+ foldr cls False.
+
+(* Catenation of dnf is the Or of formulas *)
+Lemma cat_dnfP e bcs1 bcs2 :
+ qf_eval e (dnf_to_form (bcs1 ++ bcs2))
+ = qf_eval e (dnf_to_form bcs1 \/ dnf_to_form bcs2).
+Proof.
+by elim: bcs1 => //= bc1 bcs1 IH1; rewrite -orbA; congr orb; rewrite IH1.
+Qed.
+
+(* and_dnf is the And of formulas *)
+Lemma and_dnfP e bcs1 bcs2 :
+ qf_eval e (dnf_to_form (and_dnf bcs1 bcs2))
+ = qf_eval e (dnf_to_form bcs1 /\ dnf_to_form bcs2).
+Proof.
+elim: bcs1 => [|bc1 bcs1 IH1] /=; first by rewrite /and_dnf big_nil.
+rewrite /and_dnf big_cons -/(and_dnf bcs1 bcs2) cat_dnfP /=.
+rewrite {}IH1 /= andb_orl; congr orb.
+elim: bcs2 bc1 {bcs1} => [|bc2 bcs2 IH] bc1 /=; first by rewrite andbF.
+rewrite {}IH /= andb_orr; congr orb => {bcs2}.
+suffices aux (l1 l2 : seq (term R)) g : let redg := foldr (And \o g) True in
+ qf_eval e (redg (l1 ++ l2)) = qf_eval e (redg l1 /\ redg l2)%T.
++ by rewrite 2!aux /= 2!andbA -andbA -andbCA andbA andbCA andbA.
+by elim: l1 => [| t1 l1 IHl1] //=; rewrite -andbA IHl1.
+Qed.
+
+Lemma qf_to_dnfP e :
+ let qev f b := qf_eval e (dnf_to_form (qf_to_dnf f b)) in
+ forall f, qf_form f && rformula f -> qev f false = qf_eval e f.
+Proof.
+move=> qev; have qevT f: qev f true = ~~ qev f false.
+ rewrite {}/qev; elim: f => //=; do [by case | move=> f1 IH1 f2 IH2 | ].
+ - by move=> t1 t2; rewrite !andbT !orbF.
+ - by rewrite and_dnfP cat_dnfP negb_and -IH1 -IH2.
+ - by rewrite and_dnfP cat_dnfP negb_or -IH1 -IH2.
+ - by rewrite and_dnfP cat_dnfP /= negb_or IH1 -IH2 negbK.
+ by move=> t1 ->; rewrite negbK.
+rewrite /qev; elim=> //=; first by case.
+- by move=> t1 t2 _; rewrite subr_eq0 !andbT orbF.
+- move=> f1 IH1 f2 IH2; rewrite andbCA -andbA andbCA andbA; case/andP.
+ by rewrite and_dnfP /= => /IH1-> /IH2->.
+- move=> f1 IH1 f2 IH2; rewrite andbCA -andbA andbCA andbA; case/andP.
+ by rewrite cat_dnfP /= => /IH1-> => /IH2->.
+- move=> f1 IH1 f2 IH2; rewrite andbCA -andbA andbCA andbA; case/andP.
+ by rewrite cat_dnfP /= [qf_eval _ _]qevT -implybE => /IH1 <- /IH2->.
+by move=> f1 IH1 /IH1 <-; rewrite -qevT.
+Qed.
+
+Lemma dnf_to_form_qf bcs : qf_form (dnf_to_form bcs).
+Proof.
+by elim: bcs => //= [[clT clF] _ ->] /=; elim: clT => //=; elim: clF.
+Qed.
+
+Definition dnf_rterm cl := all rterm cl.1 && all rterm cl.2.
+
+Lemma qf_to_dnf_rterm f b : rformula f -> all dnf_rterm (qf_to_dnf f b).
+Proof.
+set ok := all dnf_rterm.
+have cat_ok bcs1 bcs2: ok bcs1 -> ok bcs2 -> ok (bcs1 ++ bcs2).
+ by move=> ok1 ok2; rewrite [ok _]all_cat; exact/andP.
+have and_ok bcs1 bcs2: ok bcs1 -> ok bcs2 -> ok (and_dnf bcs1 bcs2).
+ rewrite /and_dnf unlock; elim: bcs1 => //= cl1 bcs1 IH1; rewrite -andbA.
+ case/and3P=> ok11 ok12 ok1 ok2; rewrite cat_ok ?{}IH1 {bcs1 ok1}//.
+ elim: bcs2 ok2 => //= cl2 bcs2 IH2 /andP[ok2 /IH2->].
+ by rewrite /dnf_rterm !all_cat ok11 ok12 /= !andbT.
+elim: f b => //=; [ by do 2!case | | | | | by auto | | ];
+ try by repeat case/andP || intro; case: ifP; auto.
+by rewrite /dnf_rterm => ?? [] /= ->.
+Qed.
+
+Lemma dnf_to_rform bcs : rformula (dnf_to_form bcs) = all dnf_rterm bcs.
+Proof.
+elim: bcs => //= [[cl1 cl2] bcs ->]; rewrite {2}/dnf_rterm /=; congr (_ && _).
+by congr andb; [elim: cl1 | elim: cl2] => //= t cl ->; rewrite andbT.
+Qed.
+
+Section If.
+
+Variables (pred_f then_f else_f : formula R).
+
+Definition If := (pred_f /\ then_f \/ ~ pred_f /\ else_f)%T.
+
+Lemma If_form_qf :
+ qf_form pred_f -> qf_form then_f -> qf_form else_f -> qf_form If.
+Proof. by move=> /= -> -> ->. Qed.
+
+Lemma If_form_rf :
+ rformula pred_f -> rformula then_f -> rformula else_f -> rformula If.
+Proof. by move=> /= -> -> ->. Qed.
+
+Lemma eval_If e :
+ let ev := qf_eval e in ev If = (if ev pred_f then ev then_f else ev else_f).
+Proof. by rewrite /=; case: ifP => _; rewrite ?orbF. Qed.
+
+End If.
+
+Section Pick.
+
+Variables (I : finType) (pred_f then_f : I -> formula R) (else_f : formula R).
+
+Definition Pick :=
+ \big[Or/False]_(p : {ffun pred I})
+ ((\big[And/True]_i (if p i then pred_f i else ~ pred_f i))
+ /\ (if pick p is Some i then then_f i else else_f))%T.
+
+Lemma Pick_form_qf :
+ (forall i, qf_form (pred_f i)) ->
+ (forall i, qf_form (then_f i)) ->
+ qf_form else_f ->
+ qf_form Pick.
+Proof.
+move=> qfp qft qfe; have mA := (big_morph qf_form) true andb.
+rewrite mA // big1 //= => p _.
+rewrite mA // big1 => [|i _]; first by case: pick.
+by rewrite fun_if if_same /= qfp.
+Qed.
+
+Lemma eval_Pick e (qev := qf_eval e) :
+ let P i := qev (pred_f i) in
+ qev Pick = (if pick P is Some i then qev (then_f i) else qev else_f).
+Proof.
+move=> P; rewrite ((big_morph qev) false orb) //= big_orE /=.
+apply/existsP/idP=> [[p] | true_at_P].
+ rewrite ((big_morph qev) true andb) //= big_andE /=.
+ case/andP=> /forallP-eq_p_P.
+ rewrite (@eq_pick _ _ P) => [|i]; first by case: pick.
+ by move/(_ i): eq_p_P => /=; case: (p i) => //=; move/negbTE.
+exists [ffun i => P i] => /=; apply/andP; split.
+ rewrite ((big_morph qev) true andb) //= big_andE /=.
+ by apply/forallP=> i; rewrite /= ffunE; case Pi: (P i) => //=; apply: negbT.
+rewrite (@eq_pick _ _ P) => [|i]; first by case: pick true_at_P.
+by rewrite ffunE.
+Qed.
+
+End Pick.
+
+Section MultiQuant.
+
+Variable f : formula R.
+Implicit Types (I : seq nat) (e : seq R).
+
+Lemma foldExistsP I e :
+ (exists2 e', {in [predC I], same_env e e'} & holds e' f)
+ <-> holds e (foldr Exists f I).
+Proof.
+elim: I e => /= [|i I IHi] e.
+ by split=> [[e' eq_e] |]; [apply: eq_holds => i; rewrite eq_e | exists e].
+split=> [[e' eq_e f_e'] | [x]]; last set e_x := set_nth 0 e i x.
+ exists e'`_i; apply/IHi; exists e' => // j.
+ by have:= eq_e j; rewrite nth_set_nth /= !inE; case: eqP => // ->.
+case/IHi=> e' eq_e f_e'; exists e' => // j.
+by have:= eq_e j; rewrite nth_set_nth /= !inE; case: eqP.
+Qed.
+
+Lemma foldForallP I e :
+ (forall e', {in [predC I], same_env e e'} -> holds e' f)
+ <-> holds e (foldr Forall f I).
+Proof.
+elim: I e => /= [|i I IHi] e.
+ by split=> [|f_e e' eq_e]; [exact | apply: eq_holds f_e => i; rewrite eq_e].
+split=> [f_e' x | f_e e' eq_e]; first set e_x := set_nth 0 e i x.
+ apply/IHi=> e' eq_e; apply: f_e' => j.
+ by have:= eq_e j; rewrite nth_set_nth /= !inE; case: eqP.
+move/IHi: (f_e e'`_i); apply=> j.
+by have:= eq_e j; rewrite nth_set_nth /= !inE; case: eqP => // ->.
+Qed.
+
+End MultiQuant.
+
+End EvalTerm.
+
+Prenex Implicits dnf_rterm.
+
+Module IntegralDomain.
+
+Definition axiom (R : ringType) :=
+ forall x y : R, x * y = 0 -> (x == 0) || (y == 0).
+
+Section ClassDef.
+
+Record class_of (R : Type) : Type :=
+ Class {base : ComUnitRing.class_of R; mixin : axiom (Ring.Pack base R)}.
+Local Coercion base : class_of >-> ComUnitRing.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variable (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack b0 (m0 : axiom (@Ring.Pack T b0 T)) :=
+ fun bT b & phant_id (ComUnitRing.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+Definition comRingType := @ComRing.Pack cT xclass xT.
+Definition unitRingType := @UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @ComUnitRing.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> ComUnitRing.class_of.
+Implicit Arguments mixin [R x y].
+Coercion mixin : class_of >-> axiom.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> ComUnitRing.type.
+Canonical comUnitRingType.
+Notation idomainType := type.
+Notation IdomainType T m := (@pack T _ m _ _ id _ id).
+Notation "[ 'idomainType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'idomainType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'idomainType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'idomainType' 'of' T ]") : form_scope.
+End Exports.
+
+End IntegralDomain.
+Import IntegralDomain.Exports.
+
+Section IntegralDomainTheory.
+
+Variable R : idomainType.
+Implicit Types x y : R.
+
+Lemma mulf_eq0 x y : (x * y == 0) = (x == 0) || (y == 0).
+Proof.
+apply/eqP/idP; first by case: R x y => T [].
+by case/pred2P=> ->; rewrite (mulr0, mul0r).
+Qed.
+
+Lemma prodf_eq0 (I : finType) (P : pred I) (F : I -> R) :
+ reflect (exists2 i, P i & (F i == 0)) (\prod_(i | P i) F i == 0).
+Proof.
+apply: (iffP idP) => [|[i Pi /eqP Fi0]]; last first.
+ by rewrite (bigD1 i) //= Fi0 mul0r.
+elim: (index_enum _) => [|i r IHr]; first by rewrite big_nil oner_eq0.
+rewrite big_cons /=; have [Pi | _] := ifP; last exact: IHr.
+by rewrite mulf_eq0; case/orP=> // Fi0; exists i.
+Qed.
+
+Lemma prodf_seq_eq0 I r (P : pred I) (F : I -> R) :
+ (\prod_(i <- r | P i) F i == 0) = has (fun i => P i && (F i == 0)) r.
+Proof. by rewrite (big_morph _ mulf_eq0 (oner_eq0 _)) big_has_cond. Qed.
+
+Lemma mulf_neq0 x y : x != 0 -> y != 0 -> x * y != 0.
+Proof. move=> x0 y0; rewrite mulf_eq0; exact/norP. Qed.
+
+Lemma prodf_neq0 (I : finType) (P : pred I) (F : I -> R) :
+ reflect (forall i, P i -> (F i != 0)) (\prod_(i | P i) F i != 0).
+Proof.
+by rewrite (sameP (prodf_eq0 _ _) exists_inP) negb_exists_in; exact: forall_inP.
+Qed.
+
+Lemma prodf_seq_neq0 I r (P : pred I) (F : I -> R) :
+ (\prod_(i <- r | P i) F i != 0) = all (fun i => P i ==> (F i != 0)) r.
+Proof.
+rewrite prodf_seq_eq0 -all_predC; apply: eq_all => i /=.
+by rewrite implybE negb_and.
+Qed.
+
+Lemma expf_eq0 x n : (x ^+ n == 0) = (n > 0) && (x == 0).
+Proof.
+elim: n => [|n IHn]; first by rewrite oner_eq0.
+by rewrite exprS mulf_eq0 IHn andKb.
+Qed.
+
+Lemma sqrf_eq0 x : (x ^+ 2 == 0) = (x == 0). Proof. exact: expf_eq0. Qed.
+
+Lemma expf_neq0 x m : x != 0 -> x ^+ m != 0.
+Proof. by move=> x_nz; rewrite expf_eq0; apply/nandP; right. Qed.
+
+Lemma natf_neq0 n : (n%:R != 0 :> R) = [char R]^'.-nat n.
+Proof.
+have [-> | /prod_prime_decomp->] := posnP n; first by rewrite eqxx.
+rewrite !big_seq; elim/big_rec: _ => [|[p e] s /=]; first by rewrite oner_eq0.
+case/mem_prime_decomp=> p_pr _ _; rewrite pnat_mul pnat_exp eqn0Ngt orbC => <-.
+by rewrite natrM natrX mulf_eq0 expf_eq0 negb_or negb_and pnatE ?inE p_pr.
+Qed.
+
+Lemma natf0_char n : n > 0 -> n%:R == 0 :> R -> exists p, p \in [char R].
+Proof.
+move=> n_gt0 nR_0; exists (pdiv n`_[char R]).
+apply: pnatP (pdiv_dvd _); rewrite ?part_pnat // ?pdiv_prime //.
+by rewrite ltn_neqAle eq_sym partn_eq1 // -natf_neq0 nR_0 /=.
+Qed.
+
+Lemma charf'_nat n : [char R]^'.-nat n = (n%:R != 0 :> R).
+Proof.
+have [-> | n_gt0] := posnP n; first by rewrite eqxx.
+apply/idP/idP => [|nz_n]; last first.
+ by apply/pnatP=> // p p_pr p_dvd_n; apply: contra nz_n => /dvdn_charf <-.
+apply: contraL => n0; have [// | p charRp] := natf0_char _ n0.
+have [p_pr _] := andP charRp; rewrite (eq_pnat _ (eq_negn (charf_eq charRp))).
+by rewrite p'natE // (dvdn_charf charRp) n0.
+Qed.
+
+Lemma charf0P : [char R] =i pred0 <-> (forall n, (n%:R == 0 :> R) = (n == 0)%N).
+Proof.
+split=> charF0 n; last by rewrite !inE charF0 andbC; case: eqP => // ->.
+have [-> | n_gt0] := posnP; first exact: eqxx.
+by apply/negP; case/natf0_char=> // p; rewrite charF0.
+Qed.
+
+Lemma eqf_sqr x y : (x ^+ 2 == y ^+ 2) = (x == y) || (x == - y).
+Proof. by rewrite -subr_eq0 subr_sqr mulf_eq0 subr_eq0 addr_eq0. Qed.
+
+Lemma mulfI x : x != 0 -> injective ( *%R x).
+Proof.
+move=> nz_x y z; rewrite -[x * z]add0r; move/(canLR (addrK _))/eqP.
+rewrite -mulrN -mulrDr mulf_eq0 (negbTE nz_x) /=.
+by move/eqP/(canRL (subrK _)); rewrite add0r.
+Qed.
+
+Lemma mulIf x : x != 0 -> injective ( *%R^~ x).
+Proof. by move=> nz_x y z; rewrite -!(mulrC x); exact: mulfI. Qed.
+
+Lemma sqrf_eq1 x : (x ^+ 2 == 1) = (x == 1) || (x == -1).
+Proof. by rewrite -subr_eq0 subr_sqr_1 mulf_eq0 subr_eq0 addr_eq0. Qed.
+
+Lemma expfS_eq1 x n :
+ (x ^+ n.+1 == 1) = (x == 1) || (\sum_(i < n.+1) x ^+ i == 0).
+Proof. by rewrite -![_ == 1]subr_eq0 subrX1 mulf_eq0. Qed.
+
+Lemma lregP x : reflect (lreg x) (x != 0).
+Proof. by apply: (iffP idP) => [/mulfI | /lreg_neq0]. Qed.
+
+Lemma rregP x : reflect (rreg x) (x != 0).
+Proof. by apply: (iffP idP) => [/mulIf | /rreg_neq0]. Qed.
+
+Canonical regular_idomainType := [idomainType of R^o].
+
+End IntegralDomainTheory.
+
+Implicit Arguments lregP [[R] [x]].
+Implicit Arguments rregP [[R] [x]].
+
+Module Field.
+
+Definition mixin_of (F : unitRingType) := forall x : F, x != 0 -> x \in unit.
+
+Lemma IdomainMixin R : mixin_of R -> IntegralDomain.axiom R.
+Proof.
+move=> m x y xy0; apply/norP=> [[]] /m Ux /m.
+by rewrite -(unitrMr _ Ux) xy0 unitr0.
+Qed.
+
+Section Mixins.
+
+Variables (R : comRingType) (inv : R -> R).
+
+Definition axiom := forall x, x != 0 -> inv x * x = 1.
+Hypothesis mulVx : axiom.
+Hypothesis inv0 : inv 0 = 0.
+
+Fact intro_unit (x y : R) : y * x = 1 -> x != 0.
+Proof.
+by move=> yx1; apply: contraNneq (oner_neq0 R) => x0; rewrite -yx1 x0 mulr0.
+Qed.
+
+Fact inv_out : {in predC (predC1 0), inv =1 id}.
+Proof. by move=> x /negbNE/eqP->. Qed.
+
+Definition UnitMixin := ComUnitRing.Mixin mulVx intro_unit inv_out.
+
+Lemma Mixin : mixin_of (UnitRing.Pack (UnitRing.Class UnitMixin) R).
+Proof. by []. Qed.
+
+End Mixins.
+
+Section ClassDef.
+
+Record class_of (F : Type) : Type := Class {
+ base : IntegralDomain.class_of F;
+ mixin : mixin_of (UnitRing.Pack base F)
+}.
+Local Coercion base : class_of >-> IntegralDomain.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variable (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack b0 (m0 : mixin_of (@UnitRing.Pack T b0 T)) :=
+ fun bT b & phant_id (IntegralDomain.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+Definition comRingType := @ComRing.Pack cT xclass xT.
+Definition unitRingType := @UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @IntegralDomain.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> IntegralDomain.class_of.
+Implicit Arguments mixin [F x].
+Coercion mixin : class_of >-> mixin_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> IntegralDomain.type.
+Canonical idomainType.
+Notation fieldType := type.
+Notation FieldType T m := (@pack T _ m _ _ id _ id).
+Notation FieldUnitMixin := UnitMixin.
+Notation FieldIdomainMixin := IdomainMixin.
+Notation FieldMixin := Mixin.
+Notation "[ 'fieldType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'fieldType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'fieldType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'fieldType' 'of' T ]") : form_scope.
+End Exports.
+
+End Field.
+Import Field.Exports.
+
+Section FieldTheory.
+
+Variable F : fieldType.
+Implicit Types x y : F.
+
+Lemma fieldP : Field.mixin_of F. Proof. by case: F => T []. Qed.
+
+Lemma unitfE x : (x \in unit) = (x != 0).
+Proof. by apply/idP/idP=> [/(memPn _)-> | /fieldP]; rewrite ?unitr0. Qed.
+
+Lemma mulVf x : x != 0 -> x^-1 * x = 1.
+Proof. by rewrite -unitfE; exact: mulVr. Qed.
+Lemma divff x : x != 0 -> x / x = 1.
+Proof. by rewrite -unitfE; exact: divrr. Qed.
+Definition mulfV := divff.
+Lemma mulKf x : x != 0 -> cancel ( *%R x) ( *%R x^-1).
+Proof. by rewrite -unitfE; exact: mulKr. Qed.
+Lemma mulVKf x : x != 0 -> cancel ( *%R x^-1) ( *%R x).
+Proof. by rewrite -unitfE; exact: mulVKr. Qed.
+Lemma mulfK x : x != 0 -> cancel ( *%R^~ x) ( *%R^~ x^-1).
+Proof. by rewrite -unitfE; exact: mulrK. Qed.
+Lemma mulfVK x : x != 0 -> cancel ( *%R^~ x^-1) ( *%R^~ x).
+Proof. by rewrite -unitfE; exact: divrK. Qed.
+Definition divfK := mulfVK.
+
+Lemma invfM : {morph @inv F : x y / x * y}.
+Proof.
+move=> x y; case: (eqVneq x 0) => [-> |nzx]; first by rewrite !(mul0r, invr0).
+case: (eqVneq y 0) => [-> |nzy]; first by rewrite !(mulr0, invr0).
+by rewrite mulrC invrM ?unitfE.
+Qed.
+
+Lemma invf_div x y : (x / y)^-1 = y / x.
+Proof. by rewrite invfM invrK mulrC. Qed.
+
+Lemma expfB_cond m n x : (x == 0) + n <= m -> x ^+ (m - n) = x ^+ m / x ^+ n.
+Proof.
+move/subnK=> <-; rewrite addnA addnK !exprD.
+have [-> | nz_x] := altP eqP; first by rewrite !mulr0 !mul0r.
+by rewrite mulfK ?expf_neq0.
+Qed.
+
+Lemma expfB m n x : n < m -> x ^+ (m - n) = x ^+ m / x ^+ n.
+Proof. by move=> lt_n_m; apply: expfB_cond; case: eqP => // _; apply: ltnW. Qed.
+
+Lemma prodfV I r (P : pred I) (E : I -> F) :
+ \prod_(i <- r | P i) (E i)^-1 = (\prod_(i <- r | P i) E i)^-1.
+Proof. by rewrite (big_morph _ invfM (invr1 F)). Qed.
+
+Lemma prodf_div I r (P : pred I) (E D : I -> F) :
+ \prod_(i <- r | P i) (E i / D i) =
+ \prod_(i <- r | P i) E i / \prod_(i <- r | P i) D i.
+Proof. by rewrite big_split prodfV. Qed.
+
+Lemma telescope_prodf n m (f : nat -> F) :
+ (forall k, n < k < m -> f k != 0) -> n < m ->
+ \prod_(n <= k < m) (f k.+1 / f k) = f m / f n.
+Proof.
+move=> nz_f ltnm; apply: invr_inj; rewrite prodf_div !invf_div -prodf_div.
+by apply: telescope_prodr => // k /nz_f; rewrite unitfE.
+Qed.
+
+Lemma addf_div x1 y1 x2 y2 :
+ y1 != 0 -> y2 != 0 -> x1 / y1 + x2 / y2 = (x1 * y2 + x2 * y1) / (y1 * y2).
+Proof. by move=> nzy1 nzy2; rewrite invfM mulrDl !mulrA mulrAC !mulfK. Qed.
+
+Lemma mulf_div x1 y1 x2 y2 : (x1 / y1) * (x2 / y2) = (x1 * x2) / (y1 * y2).
+Proof. by rewrite mulrACA -invfM. Qed.
+
+
+Lemma char0_natf_div :
+ [char F] =i pred0 -> forall m d, d %| m -> (m %/ d)%:R = m%:R / d%:R :> F.
+Proof.
+move/charf0P=> char0F m [|d] d_dv_m; first by rewrite divn0 invr0 mulr0.
+by rewrite natr_div // unitfE char0F.
+Qed.
+
+Section FieldMorphismInj.
+
+Variables (R : ringType) (f : {rmorphism F -> R}).
+
+Lemma fmorph_eq0 x : (f x == 0) = (x == 0).
+Proof.
+have [-> | nz_x] := altP (x =P _); first by rewrite rmorph0 eqxx.
+apply/eqP; move/(congr1 ( *%R (f x^-1)))/eqP.
+by rewrite -rmorphM mulVf // mulr0 rmorph1 ?oner_eq0.
+Qed.
+
+Lemma fmorph_inj : injective f.
+Proof.
+move=> x y eqfxy; apply/eqP; rewrite -subr_eq0 -fmorph_eq0 rmorphB //.
+by rewrite eqfxy subrr.
+Qed.
+
+Lemma fmorph_eq1 x : (f x == 1) = (x == 1).
+Proof. by rewrite -(inj_eq fmorph_inj) rmorph1. Qed.
+
+Lemma fmorph_char : [char R] =i [char F].
+Proof. by move=> p; rewrite !inE -fmorph_eq0 rmorph_nat. Qed.
+
+End FieldMorphismInj.
+
+Section FieldMorphismInv.
+
+Variables (R : unitRingType) (f : {rmorphism F -> R}).
+
+Lemma fmorph_unit x : (f x \in unit) = (x != 0).
+Proof.
+have [-> |] := altP (x =P _); first by rewrite rmorph0 unitr0.
+by rewrite -unitfE; exact: rmorph_unit.
+Qed.
+
+Lemma fmorphV : {morph f: x / x^-1}.
+Proof.
+move=> x; have [-> | nz_x] := eqVneq x 0; first by rewrite !(invr0, rmorph0).
+by rewrite rmorphV ?unitfE.
+Qed.
+
+Lemma fmorph_div : {morph f : x y / x / y}.
+Proof. by move=> x y; rewrite rmorphM fmorphV. Qed.
+
+End FieldMorphismInv.
+
+Canonical regular_fieldType := [fieldType of F^o].
+
+Section ModuleTheory.
+
+Variable V : lmodType F.
+Implicit Types (a : F) (v : V).
+
+Lemma scalerK a : a != 0 -> cancel ( *:%R a : V -> V) ( *:%R a^-1).
+Proof. by move=> nz_a v; rewrite scalerA mulVf // scale1r. Qed.
+
+Lemma scalerKV a : a != 0 -> cancel ( *:%R a^-1 : V -> V) ( *:%R a).
+Proof. by rewrite -invr_eq0 -{3}[a]invrK; exact: scalerK. Qed.
+
+Lemma scalerI a : a != 0 -> injective ( *:%R a : V -> V).
+Proof. move=> nz_a; exact: can_inj (scalerK nz_a). Qed.
+
+Lemma scaler_eq0 a v : (a *: v == 0) = (a == 0) || (v == 0).
+Proof.
+have [-> | nz_a] := altP (a =P _); first by rewrite scale0r eqxx.
+by rewrite (can2_eq (scalerK nz_a) (scalerKV nz_a)) scaler0.
+Qed.
+
+Lemma rpredZeq S (modS : submodPred S) (kS : keyed_pred modS) a v :
+ (a *: v \in kS) = (a == 0) || (v \in kS).
+Proof.
+have [-> | nz_a] := altP eqP; first by rewrite scale0r rpred0.
+by apply/idP/idP; first rewrite -{2}(scalerK nz_a v); apply: rpredZ.
+Qed.
+
+End ModuleTheory.
+
+Lemma char_lalg (A : lalgType F) : [char A] =i [char F].
+Proof. by move=> p; rewrite inE -scaler_nat scaler_eq0 oner_eq0 orbF. Qed.
+
+Section Predicates.
+
+Context (S : pred_class) (divS : @divrPred F S) (kS : keyed_pred divS).
+
+Lemma fpredMl x y : x \in kS -> x != 0 -> (x * y \in kS) = (y \in kS).
+Proof. by rewrite -!unitfE; exact: rpredMl. Qed.
+
+Lemma fpredMr x y : x \in kS -> x != 0 -> (y * x \in kS) = (y \in kS).
+Proof. by rewrite -!unitfE; exact: rpredMr. Qed.
+
+Lemma fpred_divl x y : x \in kS -> x != 0 -> (x / y \in kS) = (y \in kS).
+Proof. by rewrite -!unitfE; exact: rpred_divl. Qed.
+
+Lemma fpred_divr x y : x \in kS -> x != 0 -> (y / x \in kS) = (y \in kS).
+Proof. by rewrite -!unitfE; exact: rpred_divr. Qed.
+
+End Predicates.
+
+End FieldTheory.
+
+Implicit Arguments fmorph_inj [[F] [R] x1 x2].
+
+Module DecidableField.
+
+Definition axiom (R : unitRingType) (s : seq R -> pred (formula R)) :=
+ forall e f, reflect (holds e f) (s e f).
+
+Record mixin_of (R : unitRingType) : Type :=
+ Mixin { sat : seq R -> pred (formula R); satP : axiom sat}.
+
+Section ClassDef.
+
+Record class_of (F : Type) : Type :=
+ Class {base : Field.class_of F; mixin : mixin_of (UnitRing.Pack base F)}.
+Local Coercion base : class_of >-> Field.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variable (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack b0 (m0 : mixin_of (@UnitRing.Pack T b0 T)) :=
+ fun bT b & phant_id (Field.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+Definition comRingType := @ComRing.Pack cT xclass xT.
+Definition unitRingType := @UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @IntegralDomain.Pack cT xclass xT.
+Definition fieldType := @Field.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> Field.class_of.
+Coercion mixin : class_of >-> mixin_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> IntegralDomain.type.
+Canonical idomainType.
+Coercion fieldType : type >-> Field.type.
+Canonical fieldType.
+Notation decFieldType := type.
+Notation DecFieldType T m := (@pack T _ m _ _ id _ id).
+Notation DecFieldMixin := Mixin.
+Notation "[ 'decFieldType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'decFieldType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'decFieldType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'decFieldType' 'of' T ]") : form_scope.
+End Exports.
+
+End DecidableField.
+Import DecidableField.Exports.
+
+Section DecidableFieldTheory.
+
+Variable F : decFieldType.
+
+Definition sat := DecidableField.sat (DecidableField.class F).
+
+Lemma satP : DecidableField.axiom sat.
+Proof. exact: DecidableField.satP. Qed.
+
+Fact sol_subproof n f :
+ reflect (exists s, (size s == n) && sat s f)
+ (sat [::] (foldr Exists f (iota 0 n))).
+Proof.
+apply: (iffP (satP _ _)) => [|[s]]; last first.
+ case/andP=> /eqP sz_s /satP f_s; apply/foldExistsP.
+ exists s => // i; rewrite !inE mem_iota -leqNgt add0n => le_n_i.
+ by rewrite !nth_default ?sz_s.
+case/foldExistsP=> e e0 f_e; set s := take n (set_nth 0 e n 0).
+have sz_s: size s = n by rewrite size_take size_set_nth leq_max leqnn.
+exists s; rewrite sz_s eqxx; apply/satP; apply: eq_holds f_e => i.
+case: (leqP n i) => [le_n_i | lt_i_n].
+ by rewrite -e0 ?nth_default ?sz_s // !inE mem_iota -leqNgt.
+by rewrite nth_take // nth_set_nth /= eq_sym eqn_leq leqNgt lt_i_n.
+Qed.
+
+Definition sol n f :=
+ if sol_subproof n f is ReflectT sP then xchoose sP else nseq n 0.
+
+Lemma size_sol n f : size (sol n f) = n.
+Proof.
+rewrite /sol; case: sol_subproof => [sP | _]; last exact: size_nseq.
+by case/andP: (xchooseP sP) => /eqP.
+Qed.
+
+Lemma solP n f : reflect (exists2 s, size s = n & holds s f) (sat (sol n f) f).
+Proof.
+rewrite /sol; case: sol_subproof => [sP | sPn].
+ case/andP: (xchooseP sP) => _ ->; left.
+ by case: sP => s; case/andP; move/eqP=> <-; move/satP; exists s.
+apply: (iffP (satP _ _)); first by exists (nseq n 0); rewrite ?size_nseq.
+by case=> s sz_s; move/satP=> f_s; case: sPn; exists s; rewrite sz_s eqxx.
+Qed.
+
+Lemma eq_sat f1 f2 :
+ (forall e, holds e f1 <-> holds e f2) -> sat^~ f1 =1 sat^~ f2.
+Proof. by move=> eqf12 e; apply/satP/satP; case: (eqf12 e). Qed.
+
+Lemma eq_sol f1 f2 :
+ (forall e, holds e f1 <-> holds e f2) -> sol^~ f1 =1 sol^~ f2.
+Proof.
+rewrite /sol => /eq_sat eqf12 n.
+do 2![case: sol_subproof] => //= [f1s f2s | ns1 [s f2s] | [s f1s] []].
+- by apply: eq_xchoose => s; rewrite eqf12.
+- by case: ns1; exists s; rewrite -eqf12.
+by exists s; rewrite eqf12.
+Qed.
+
+End DecidableFieldTheory.
+
+Implicit Arguments satP [[F] [e] [f]].
+Implicit Arguments solP [[F] [n] [f]].
+
+Section QE_Mixin.
+
+Variable F : Field.type.
+Implicit Type f : formula F.
+
+Variable proj : nat -> seq (term F) * seq (term F) -> formula F.
+(* proj is the elimination of a single existential quantifier *)
+
+(* The elimination projector is well_formed. *)
+Definition wf_QE_proj :=
+ forall i bc (bc_i := proj i bc),
+ dnf_rterm bc -> qf_form bc_i && rformula bc_i.
+
+(* The elimination projector is valid *)
+Definition valid_QE_proj :=
+ forall i bc (ex_i_bc := ('exists 'X_i, dnf_to_form [:: bc])%T) e,
+ dnf_rterm bc -> reflect (holds e ex_i_bc) (qf_eval e (proj i bc)).
+
+Hypotheses (wf_proj : wf_QE_proj) (ok_proj : valid_QE_proj).
+
+Let elim_aux f n := foldr Or False (map (proj n) (qf_to_dnf f false)).
+
+Fixpoint quantifier_elim f :=
+ match f with
+ | f1 /\ f2 => (quantifier_elim f1) /\ (quantifier_elim f2)
+ | f1 \/ f2 => (quantifier_elim f1) \/ (quantifier_elim f2)
+ | f1 ==> f2 => (~ quantifier_elim f1) \/ (quantifier_elim f2)
+ | ~ f => ~ quantifier_elim f
+ | ('exists 'X_n, f) => elim_aux (quantifier_elim f) n
+ | ('forall 'X_n, f) => ~ elim_aux (~ quantifier_elim f) n
+ | _ => f
+ end%T.
+
+Lemma quantifier_elim_wf f :
+ let qf := quantifier_elim f in rformula f -> qf_form qf && rformula qf.
+Proof.
+suffices aux_wf f0 n : let qf := elim_aux f0 n in
+ rformula f0 -> qf_form qf && rformula qf.
+- by elim: f => //=; do ?[ move=> f1 IH1 f2 IH2;
+ case/andP=> rf1 rf2;
+ case/andP:(IH1 rf1)=> -> ->;
+ case/andP:(IH2 rf2)=> -> -> //
+ | move=> n f1 IH rf1;
+ case/andP: (IH rf1)=> qff rf;
+ rewrite aux_wf ].
+rewrite /elim_aux => rf.
+suffices or_wf fs : let ofs := foldr Or False fs in
+ all (@qf_form F) fs && all (@rformula F) fs -> qf_form ofs && rformula ofs.
+- apply: or_wf.
+ suffices map_proj_wf bcs: let mbcs := map (proj n) bcs in
+ all dnf_rterm bcs -> all (@qf_form _) mbcs && all (@rformula _) mbcs.
+ by apply: map_proj_wf; exact: qf_to_dnf_rterm.
+ elim: bcs => [|bc bcs ihb] bcsr //= /andP[rbc rbcs].
+ by rewrite andbAC andbA wf_proj //= andbC ihb.
+elim: fs => //= g gs ihg; rewrite -andbA => /and4P[-> qgs -> rgs] /=.
+by apply: ihg; rewrite qgs rgs.
+Qed.
+
+Lemma quantifier_elim_rformP e f :
+ rformula f -> reflect (holds e f) (qf_eval e (quantifier_elim f)).
+Proof.
+pose rc e n f := exists x, qf_eval (set_nth 0 e n x) f.
+have auxP f0 e0 n0: qf_form f0 && rformula f0 ->
+ reflect (rc e0 n0 f0) (qf_eval e0 (elim_aux f0 n0)).
++ rewrite /elim_aux => cf; set bcs := qf_to_dnf f0 false.
+ apply: (@iffP (rc e0 n0 (dnf_to_form bcs))); last first.
+ - by case=> x; rewrite -qf_to_dnfP //; exists x.
+ - by case=> x; rewrite qf_to_dnfP //; exists x.
+ have: all dnf_rterm bcs by case/andP: cf => _; exact: qf_to_dnf_rterm.
+ elim: {f0 cf}bcs => [|bc bcs IHbcs] /=; first by right; case.
+ case/andP=> r_bc /IHbcs {IHbcs}bcsP.
+ have f_qf := dnf_to_form_qf [:: bc].
+ case: ok_proj => //= [ex_x|no_x].
+ left; case: ex_x => x /(qf_evalP _ f_qf); rewrite /= orbF => bc_x.
+ by exists x; rewrite /= bc_x.
+ apply: (iffP bcsP) => [[x bcs_x] | [x]] /=.
+ by exists x; rewrite /= bcs_x orbT.
+ case/orP => [bc_x|]; last by exists x.
+ by case: no_x; exists x; apply/(qf_evalP _ f_qf); rewrite /= bc_x.
+elim: f e => //.
+- move=> b e _; exact: idP.
+- move=> t1 t2 e _; exact: eqP.
+- move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; last by right; case.
+ by case/IH2; [left | right; case].
+- move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; first by do 2!left.
+ by case/IH2; [left; right | right; case].
+- move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; last by left.
+ by case/IH2; [left | right; move/(_ f1e)].
+- by move=> f IHf e /= /IHf[]; [right | left].
+- move=> n f IHf e /= rf; have rqf := quantifier_elim_wf rf.
+ by apply: (iffP (auxP _ _ _ rqf)) => [] [x]; exists x; exact/IHf.
+move=> n f IHf e /= rf; have rqf := quantifier_elim_wf rf.
+case: auxP => // [f_x|no_x]; first by right=> no_x; case: f_x => x /IHf[].
+by left=> x; apply/IHf=> //; apply/idPn=> f_x; case: no_x; exists x.
+Qed.
+
+Definition proj_sat e f := qf_eval e (quantifier_elim (to_rform f)).
+
+Lemma proj_satP : DecidableField.axiom proj_sat.
+Proof.
+move=> e f; have fP := quantifier_elim_rformP e (to_rform_rformula f).
+by apply: (iffP fP); move/to_rformP.
+Qed.
+
+Definition QEdecFieldMixin := DecidableField.Mixin proj_satP.
+
+End QE_Mixin.
+
+Module ClosedField.
+
+(* Axiom == all non-constant monic polynomials have a root *)
+Definition axiom (R : ringType) :=
+ forall n (P : nat -> R), n > 0 ->
+ exists x : R, x ^+ n = \sum_(i < n) P i * (x ^+ i).
+
+Section ClassDef.
+
+Record class_of (F : Type) : Type :=
+ Class {base : DecidableField.class_of F; _ : axiom (Ring.Pack base F)}.
+Local Coercion base : class_of >-> DecidableField.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variable (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+
+Definition pack b0 (m0 : axiom (@Ring.Pack T b0 T)) :=
+ fun bT b & phant_id (DecidableField.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+(* There should eventually be a constructor from polynomial resolution *)
+(* that builds the DecidableField mixin using QE. *)
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @Zmodule.Pack cT xclass xT.
+Definition ringType := @Ring.Pack cT xclass xT.
+Definition comRingType := @ComRing.Pack cT xclass xT.
+Definition unitRingType := @UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @IntegralDomain.Pack cT xclass xT.
+Definition fieldType := @Field.Pack cT xclass xT.
+Definition decFieldType := @DecidableField.Pack cT class xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> DecidableField.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> IntegralDomain.type.
+Canonical idomainType.
+Coercion fieldType : type >-> Field.type.
+Canonical fieldType.
+Coercion decFieldType : type >-> DecidableField.type.
+Canonical decFieldType.
+Notation closedFieldType := type.
+Notation ClosedFieldType T m := (@pack T _ m _ _ id _ id).
+Notation "[ 'closedFieldType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'closedFieldType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'closedFieldType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'closedFieldType' 'of' T ]") : form_scope.
+End Exports.
+
+End ClosedField.
+Import ClosedField.Exports.
+
+Section ClosedFieldTheory.
+
+Variable F : closedFieldType.
+
+Lemma solve_monicpoly : ClosedField.axiom F.
+Proof. by case: F => ? []. Qed.
+
+End ClosedFieldTheory.
+
+Module SubType.
+
+Section Zmodule.
+
+Variables (V : zmodType) (S : predPredType V).
+Variables (subS : zmodPred S) (kS : keyed_pred subS).
+Variable U : subType (mem kS).
+
+Let inU v Sv : U := Sub v Sv.
+Let zeroU := inU (rpred0 kS).
+
+Let oppU (u : U) := inU (rpredNr (valP u)).
+Let addU (u1 u2 : U) := inU (rpredD (valP u1) (valP u2)).
+
+Fact addA : associative addU.
+Proof. by move=> u1 u2 u3; apply: val_inj; rewrite !SubK addrA. Qed.
+Fact addC : commutative addU.
+Proof. by move=> u1 u2; apply: val_inj; rewrite !SubK addrC. Qed.
+Fact add0 : left_id zeroU addU.
+Proof. by move=> u; apply: val_inj; rewrite !SubK add0r. Qed.
+Fact addN : left_inverse zeroU oppU addU.
+Proof. by move=> u; apply: val_inj; rewrite !SubK addNr. Qed.
+
+Definition zmodMixin of phant U := ZmodMixin addA addC add0 addN.
+
+End Zmodule.
+
+Section Ring.
+
+Variables (R : ringType) (S : predPredType R).
+Variables (ringS : subringPred S) (kS : keyed_pred ringS).
+
+Definition cast_zmodType (V : zmodType) T (VeqT : V = T :> Type) :=
+ let cast mV := let: erefl in _ = T := VeqT return Zmodule.class_of T in mV in
+ Zmodule.Pack (cast (Zmodule.class V)) T.
+
+Variable (T : subType (mem kS)) (V : zmodType) (VeqT: V = T :> Type).
+
+Let inT x Sx : T := Sub x Sx.
+Let oneT := inT (rpred1 kS).
+Let mulT (u1 u2 : T) := inT (rpredM (valP u1) (valP u2)).
+Let T' := cast_zmodType VeqT.
+
+Hypothesis valM : {morph (val : T' -> R) : x y / x - y}.
+
+Let val0 : val (0 : T') = 0.
+Proof. by rewrite -(subrr (0 : T')) valM subrr. Qed.
+Let valD : {morph (val : T' -> R): x y / x + y}.
+Proof.
+by move=> u v; rewrite -{1}[v]opprK -[- v]sub0r !valM val0 sub0r opprK.
+Qed.
+
+Fact mulA : @associative T' mulT.
+Proof. by move=> u1 u2 u3; apply: val_inj; rewrite !SubK mulrA. Qed.
+Fact mul1l : left_id oneT mulT.
+Proof. by move=> u; apply: val_inj; rewrite !SubK mul1r. Qed.
+Fact mul1r : right_id oneT mulT.
+Proof. by move=> u; apply: val_inj; rewrite !SubK mulr1. Qed.
+Fact mulDl : @left_distributive T' T' mulT +%R.
+Proof. by move=> u1 u2 u3; apply: val_inj; rewrite !(SubK, valD) mulrDl. Qed.
+Fact mulDr : @right_distributive T' T' mulT +%R.
+Proof. by move=> u1 u2 u3; apply: val_inj; rewrite !(SubK, valD) mulrDr. Qed.
+Fact nz1 : oneT != 0 :> T'.
+Proof.
+by apply: contraNneq (oner_neq0 R) => eq10; rewrite -val0 -eq10 SubK.
+Qed.
+
+Definition ringMixin := RingMixin mulA mul1l mul1r mulDl mulDr nz1.
+
+End Ring.
+
+Section Lmodule.
+
+Variables (R : ringType) (V : lmodType R) (S : predPredType V).
+Variables (linS : submodPred S) (kS : keyed_pred linS).
+Variables (W : subType (mem kS)) (Z : zmodType) (ZeqW : Z = W :> Type).
+
+Let scaleW a (w : W) := (Sub _ : _ -> W) (rpredZ a (valP w)).
+Let W' := cast_zmodType ZeqW.
+
+Hypothesis valD : {morph (val : W' -> V) : x y / x + y}.
+
+Fact scaleA a b (w : W') : scaleW a (scaleW b w) = scaleW (a * b) w.
+Proof. by apply: val_inj; rewrite !SubK scalerA. Qed.
+Fact scale1 : left_id 1 scaleW.
+Proof. by move=> w; apply: val_inj; rewrite !SubK scale1r. Qed.
+Fact scaleDr : @right_distributive R W' scaleW +%R.
+Proof. by move=> a w w2; apply: val_inj; rewrite !(SubK, valD) scalerDr. Qed.
+Fact scaleDl w : {morph (scaleW^~ w : R -> W') : a b / a + b}.
+Proof. by move=> a b; apply: val_inj; rewrite !(SubK, valD) scalerDl. Qed.
+
+Definition lmodMixin := LmodMixin scaleA scale1 scaleDr scaleDl.
+
+End Lmodule.
+
+Lemma lalgMixin (R : ringType) (A : lalgType R) (B : lmodType R) (f : B -> A) :
+ phant B -> injective f -> scalable f ->
+ forall mulB, {morph f : x y / mulB x y >-> x * y} -> Lalgebra.axiom mulB.
+Proof.
+by move=> _ injf fZ mulB fM a x y; apply: injf; rewrite !(fZ, fM) scalerAl.
+Qed.
+
+Lemma comRingMixin (R : comRingType) (T : ringType) (f : T -> R) :
+ phant T -> injective f -> {morph f : x y / x * y} -> commutative (@mul T).
+Proof. by move=> _ inj_f fM x y; apply: inj_f; rewrite !fM mulrC. Qed.
+
+Lemma algMixin (R : comRingType) (A : algType R) (B : lalgType R) (f : B -> A) :
+ phant B -> injective f -> {morph f : x y / x * y} -> scalable f ->
+ @Algebra.axiom R B.
+Proof.
+by move=> _ inj_f fM fZ a x y; apply: inj_f; rewrite !(fM, fZ) scalerAr.
+Qed.
+
+Section UnitRing.
+
+Definition cast_ringType (Q : ringType) T (QeqT : Q = T :> Type) :=
+ let cast rQ := let: erefl in _ = T := QeqT return Ring.class_of T in rQ in
+ Ring.Pack (cast (Ring.class Q)) T.
+
+Variables (R : unitRingType) (S : predPredType R).
+Variables (ringS : divringPred S) (kS : keyed_pred ringS).
+
+Variables (T : subType (mem kS)) (Q : ringType) (QeqT : Q = T :> Type).
+
+Let inT x Sx : T := Sub x Sx.
+Let invT (u : T) := inT (rpredVr (valP u)).
+Let unitT := [qualify a u : T | val u \is a unit].
+Let T' := cast_ringType QeqT.
+
+Hypothesis val1 : val (1 : T') = 1.
+Hypothesis valM : {morph (val : T' -> R) : x y / x * y}.
+
+Fact mulVr :
+ {in (unitT : predPredType T'), left_inverse (1 : T') invT (@mul T')}.
+Proof. by move=> u Uu; apply: val_inj; rewrite val1 valM SubK mulVr. Qed.
+
+Fact mulrV : {in unitT, right_inverse (1 : T') invT (@mul T')}.
+Proof. by move=> u Uu; apply: val_inj; rewrite val1 valM SubK mulrV. Qed.
+
+Fact unitP (u v : T') : v * u = 1 /\ u * v = 1 -> u \in unitT.
+Proof.
+by case=> vu1 uv1; apply/unitrP; exists (val v); rewrite -!valM vu1 uv1.
+Qed.
+
+Fact unit_id : {in [predC unitT], invT =1 id}.
+Proof. by move=> u /invr_out def_u1; apply: val_inj; rewrite SubK. Qed.
+
+Definition unitRingMixin := UnitRingMixin mulVr mulrV unitP unit_id.
+
+End UnitRing.
+
+Lemma idomainMixin (R : idomainType) (T : ringType) (f : T -> R) :
+ phant T -> injective f -> f 0 = 0 -> {morph f : u v / u * v} ->
+ @IntegralDomain.axiom T.
+Proof.
+move=> _ injf f0 fM u v uv0.
+by rewrite -!(inj_eq injf) !f0 -mulf_eq0 -fM uv0 f0.
+Qed.
+
+Lemma fieldMixin (F : fieldType) (K : unitRingType) (f : K -> F) :
+ phant K -> injective f -> f 0 = 0 -> {mono f : u / u \in unit} ->
+ @Field.mixin_of K.
+Proof. by move=> _ injf f0 fU u; rewrite -fU unitfE -f0 inj_eq. Qed.
+
+Module Exports.
+
+Notation "[ 'zmodMixin' 'of' U 'by' <: ]" := (zmodMixin (Phant U))
+ (at level 0, format "[ 'zmodMixin' 'of' U 'by' <: ]") : form_scope.
+Notation "[ 'ringMixin' 'of' R 'by' <: ]" :=
+ (@ringMixin _ _ _ _ _ _ (@erefl Type R%type) (rrefl _))
+ (at level 0, format "[ 'ringMixin' 'of' R 'by' <: ]") : form_scope.
+Notation "[ 'lmodMixin' 'of' U 'by' <: ]" :=
+ (@lmodMixin _ _ _ _ _ _ _ (@erefl Type U%type) (rrefl _))
+ (at level 0, format "[ 'lmodMixin' 'of' U 'by' <: ]") : form_scope.
+Notation "[ 'lalgMixin' 'of' A 'by' <: ]" :=
+ ((lalgMixin (Phant A) val_inj (rrefl _)) *%R (rrefl _))
+ (at level 0, format "[ 'lalgMixin' 'of' A 'by' <: ]") : form_scope.
+Notation "[ 'comRingMixin' 'of' R 'by' <: ]" :=
+ (comRingMixin (Phant R) val_inj (rrefl _))
+ (at level 0, format "[ 'comRingMixin' 'of' R 'by' <: ]") : form_scope.
+Notation "[ 'algMixin' 'of' A 'by' <: ]" :=
+ (algMixin (Phant A) val_inj (rrefl _) (rrefl _))
+ (at level 0, format "[ 'algMixin' 'of' A 'by' <: ]") : form_scope.
+Notation "[ 'unitRingMixin' 'of' R 'by' <: ]" :=
+ (@unitRingMixin _ _ _ _ _ _ (@erefl Type R%type) (erefl _) (rrefl _))
+ (at level 0, format "[ 'unitRingMixin' 'of' R 'by' <: ]") : form_scope.
+Notation "[ 'idomainMixin' 'of' R 'by' <: ]" :=
+ (idomainMixin (Phant R) val_inj (erefl _) (rrefl _))
+ (at level 0, format "[ 'idomainMixin' 'of' R 'by' <: ]") : form_scope.
+Notation "[ 'fieldMixin' 'of' F 'by' <: ]" :=
+ (fieldMixin (Phant F) val_inj (erefl _) (frefl _))
+ (at level 0, format "[ 'fieldMixin' 'of' F 'by' <: ]") : form_scope.
+
+End Exports.
+
+End SubType.
+
+Module Theory.
+
+Definition addrA := addrA.
+Definition addrC := addrC.
+Definition add0r := add0r.
+Definition addNr := addNr.
+Definition addr0 := addr0.
+Definition addrN := addrN.
+Definition subrr := subrr.
+Definition addrCA := addrCA.
+Definition addrAC := addrAC.
+Definition addrACA := addrACA.
+Definition addKr := addKr.
+Definition addNKr := addNKr.
+Definition addrK := addrK.
+Definition addrNK := addrNK.
+Definition subrK := subrK.
+Definition addrI := @addrI.
+Definition addIr := @addIr.
+Implicit Arguments addrI [[V] x1 x2].
+Implicit Arguments addIr [[V] x1 x2].
+Definition opprK := opprK.
+Definition oppr_inj := @oppr_inj.
+Implicit Arguments oppr_inj [[V] x1 x2].
+Definition oppr0 := oppr0.
+Definition oppr_eq0 := oppr_eq0.
+Definition opprD := opprD.
+Definition opprB := opprB.
+Definition subr0 := subr0.
+Definition sub0r := sub0r.
+Definition subr_eq := subr_eq.
+Definition subr_eq0 := subr_eq0.
+Definition addr_eq0 := addr_eq0.
+Definition eqr_opp := eqr_opp.
+Definition eqr_oppLR := eqr_oppLR.
+Definition sumrN := sumrN.
+Definition sumrB := sumrB.
+Definition sumrMnl := sumrMnl.
+Definition sumrMnr := sumrMnr.
+Definition sumr_const := sumr_const.
+Definition telescope_sumr := telescope_sumr.
+Definition mulr0n := mulr0n.
+Definition mulr1n := mulr1n.
+Definition mulr2n := mulr2n.
+Definition mulrS := mulrS.
+Definition mulrSr := mulrSr.
+Definition mulrb := mulrb.
+Definition mul0rn := mul0rn.
+Definition mulNrn := mulNrn.
+Definition mulrnDl := mulrnDl.
+Definition mulrnDr := mulrnDr.
+Definition mulrnBl := mulrnBl.
+Definition mulrnBr := mulrnBr.
+Definition mulrnA := mulrnA.
+Definition mulrnAC := mulrnAC.
+Definition mulrA := mulrA.
+Definition mul1r := mul1r.
+Definition mulr1 := mulr1.
+Definition mulrDl := mulrDl.
+Definition mulrDr := mulrDr.
+Definition oner_neq0 := oner_neq0.
+Definition oner_eq0 := oner_eq0.
+Definition mul0r := mul0r.
+Definition mulr0 := mulr0.
+Definition mulrN := mulrN.
+Definition mulNr := mulNr.
+Definition mulrNN := mulrNN.
+Definition mulN1r := mulN1r.
+Definition mulrN1 := mulrN1.
+Definition mulr_suml := mulr_suml.
+Definition mulr_sumr := mulr_sumr.
+Definition mulrBl := mulrBl.
+Definition mulrBr := mulrBr.
+Definition mulrnAl := mulrnAl.
+Definition mulrnAr := mulrnAr.
+Definition mulr_natl := mulr_natl.
+Definition mulr_natr := mulr_natr.
+Definition natrD := natrD.
+Definition natrB := natrB.
+Definition natr_sum := natr_sum.
+Definition natrM := natrM.
+Definition natrX := natrX.
+Definition expr0 := expr0.
+Definition exprS := exprS.
+Definition expr1 := expr1.
+Definition expr2 := expr2.
+Definition expr0n := expr0n.
+Definition expr1n := expr1n.
+Definition exprD := exprD.
+Definition exprSr := exprSr.
+Definition commr_sym := commr_sym.
+Definition commr_refl := commr_refl.
+Definition commr0 := commr0.
+Definition commr1 := commr1.
+Definition commrN := commrN.
+Definition commrN1 := commrN1.
+Definition commrD := commrD.
+Definition commrMn := commrMn.
+Definition commrM := commrM.
+Definition commr_nat := commr_nat.
+Definition commrX := commrX.
+Definition exprMn_comm := exprMn_comm.
+Definition commr_sign := commr_sign.
+Definition exprMn_n := exprMn_n.
+Definition exprM := exprM.
+Definition exprAC := exprAC.
+Definition expr_mod := expr_mod.
+Definition expr_dvd := expr_dvd.
+Definition signr_odd := signr_odd.
+Definition signr_eq0 := signr_eq0.
+Definition mulr_sign := mulr_sign.
+Definition signr_addb := signr_addb.
+Definition signrN := signrN.
+Definition signrE := signrE.
+Definition mulr_signM := mulr_signM.
+Definition exprNn := exprNn.
+Definition sqrrN := sqrrN.
+Definition sqrr_sign := sqrr_sign.
+Definition signrMK := signrMK.
+Definition mulrI_eq0 := mulrI_eq0.
+Definition lreg_neq0 := lreg_neq0.
+Definition mulrI0_lreg := mulrI0_lreg.
+Definition lregN := lregN.
+Definition lreg1 := lreg1.
+Definition lregM := lregM.
+Definition lregX := lregX.
+Definition lreg_sign := lreg_sign.
+Definition lregP {R x} := @lregP R x.
+Definition mulIr_eq0 := mulIr_eq0.
+Definition mulIr0_rreg := mulIr0_rreg.
+Definition rreg_neq0 := rreg_neq0.
+Definition rregN := rregN.
+Definition rreg1 := rreg1.
+Definition rregM := rregM.
+Definition revrX := revrX.
+Definition rregX := rregX.
+Definition rregP {R x} := @rregP R x.
+Definition exprDn_comm := exprDn_comm.
+Definition exprBn_comm := exprBn_comm.
+Definition subrXX_comm := subrXX_comm.
+Definition exprD1n := exprD1n.
+Definition subrX1 := subrX1.
+Definition sqrrD1 := sqrrD1.
+Definition sqrrB1 := sqrrB1.
+Definition subr_sqr_1 := subr_sqr_1.
+Definition charf0 := charf0.
+Definition charf_prime := charf_prime.
+Definition mulrn_char := mulrn_char.
+Definition dvdn_charf := dvdn_charf.
+Definition charf_eq := charf_eq.
+Definition bin_lt_charf_0 := bin_lt_charf_0.
+Definition Frobenius_autE := Frobenius_autE.
+Definition Frobenius_aut0 := Frobenius_aut0.
+Definition Frobenius_aut1 := Frobenius_aut1.
+Definition Frobenius_autD_comm := Frobenius_autD_comm.
+Definition Frobenius_autMn := Frobenius_autMn.
+Definition Frobenius_aut_nat := Frobenius_aut_nat.
+Definition Frobenius_autM_comm := Frobenius_autM_comm.
+Definition Frobenius_autX := Frobenius_autX.
+Definition Frobenius_autN := Frobenius_autN.
+Definition Frobenius_autB_comm := Frobenius_autB_comm.
+Definition exprNn_char := exprNn_char.
+Definition addrr_char2 := addrr_char2.
+Definition oppr_char2 := oppr_char2.
+Definition addrK_char2 := addrK_char2.
+Definition addKr_char2 := addKr_char2.
+Definition prodr_const := prodr_const.
+Definition mulrC := mulrC.
+Definition mulrCA := mulrCA.
+Definition mulrAC := mulrAC.
+Definition mulrACA := mulrACA.
+Definition exprMn := exprMn.
+Definition prodrXl := prodrXl.
+Definition prodrXr := prodrXr.
+Definition prodrN := prodrN.
+Definition prodrMn := prodrMn.
+Definition natr_prod := natr_prod.
+Definition prodr_undup_exp_count := prodr_undup_exp_count.
+Definition exprDn := exprDn.
+Definition exprBn := exprBn.
+Definition subrXX := subrXX.
+Definition sqrrD := sqrrD.
+Definition sqrrB := sqrrB.
+Definition subr_sqr := subr_sqr.
+Definition subr_sqrDB := subr_sqrDB.
+Definition exprDn_char := exprDn_char.
+Definition mulrV := mulrV.
+Definition divrr := divrr.
+Definition mulVr := mulVr.
+Definition invr_out := invr_out.
+Definition unitrP {R x} := @unitrP R x.
+Definition mulKr := mulKr.
+Definition mulVKr := mulVKr.
+Definition mulrK := mulrK.
+Definition mulrVK := mulrVK.
+Definition divrK := divrK.
+Definition mulrI := mulrI.
+Definition mulIr := mulIr.
+Definition telescope_prodr := telescope_prodr.
+Definition commrV := commrV.
+Definition unitrE := unitrE.
+Definition invrK := invrK.
+Definition invr_inj := @invr_inj.
+Implicit Arguments invr_inj [[R] x1 x2].
+Definition unitrV := unitrV.
+Definition unitr1 := unitr1.
+Definition invr1 := invr1.
+Definition divr1 := divr1.
+Definition div1r := div1r.
+Definition natr_div := natr_div.
+Definition unitr0 := unitr0.
+Definition invr0 := invr0.
+Definition unitrN1 := unitrN1.
+Definition unitrN := unitrN.
+Definition invrN1 := invrN1.
+Definition invrN := invrN.
+Definition invr_sign := invr_sign.
+Definition unitrMl := unitrMl.
+Definition unitrMr := unitrMr.
+Definition invrM := invrM.
+Definition invr_eq0 := invr_eq0.
+Definition invr_eq1 := invr_eq1.
+Definition invr_neq0 := invr_neq0.
+Definition unitrM_comm := unitrM_comm.
+Definition unitrX := unitrX.
+Definition unitrX_pos := unitrX_pos.
+Definition exprVn := exprVn.
+Definition exprB := exprB.
+Definition invr_signM := invr_signM.
+Definition divr_signM := divr_signM.
+Definition rpred0D := rpred0D.
+Definition rpred0 := rpred0.
+Definition rpredD := rpredD.
+Definition rpredNr := rpredNr.
+Definition rpred_sum := rpred_sum.
+Definition rpredMn := rpredMn.
+Definition rpredN := rpredN.
+Definition rpredB := rpredB.
+Definition rpredMNn := rpredMNn.
+Definition rpredDr := rpredDr.
+Definition rpredDl := rpredDl.
+Definition rpredBr := rpredBr.
+Definition rpredBl := rpredBl.
+Definition rpredMsign := rpredMsign.
+Definition rpred1M := rpred1M.
+Definition rpred1 := rpred1.
+Definition rpredM := rpredM.
+Definition rpred_prod := rpred_prod.
+Definition rpredX := rpredX.
+Definition rpred_nat := rpred_nat.
+Definition rpredN1 := rpredN1.
+Definition rpred_sign := rpred_sign.
+Definition rpredZsign := rpredZsign.
+Definition rpredZnat := rpredZnat.
+Definition rpredZ := rpredZ.
+Definition rpredVr := rpredVr.
+Definition rpredV := rpredV.
+Definition rpred_div := rpred_div.
+Definition rpredXN := rpredXN.
+Definition rpredZeq := rpredZeq.
+Definition char_lalg := char_lalg.
+Definition rpredMr := rpredMr.
+Definition rpredMl := rpredMl.
+Definition rpred_divr := rpred_divr.
+Definition rpred_divl := rpred_divl.
+Definition eq_eval := eq_eval.
+Definition eval_tsubst := eval_tsubst.
+Definition eq_holds := eq_holds.
+Definition holds_fsubst := holds_fsubst.
+Definition unitrM := unitrM.
+Definition unitrPr {R x} := @unitrPr R x.
+Definition expr_div_n := expr_div_n.
+Definition mulf_eq0 := mulf_eq0.
+Definition prodf_eq0 := prodf_eq0.
+Definition prodf_seq_eq0 := prodf_seq_eq0.
+Definition mulf_neq0 := mulf_neq0.
+Definition prodf_neq0 := prodf_neq0.
+Definition prodf_seq_neq0 := prodf_seq_neq0.
+Definition expf_eq0 := expf_eq0.
+Definition sqrf_eq0 := sqrf_eq0.
+Definition expf_neq0 := expf_neq0.
+Definition natf_neq0 := natf_neq0.
+Definition natf0_char := natf0_char.
+Definition charf'_nat := charf'_nat.
+Definition charf0P := charf0P.
+Definition eqf_sqr := eqf_sqr.
+Definition mulfI := mulfI.
+Definition mulIf := mulIf.
+Definition sqrf_eq1 := sqrf_eq1.
+Definition expfS_eq1 := expfS_eq1.
+Definition fieldP := fieldP.
+Definition unitfE := unitfE.
+Definition mulVf := mulVf.
+Definition mulfV := mulfV.
+Definition divff := divff.
+Definition mulKf := mulKf.
+Definition mulVKf := mulVKf.
+Definition mulfK := mulfK.
+Definition mulfVK := mulfVK.
+Definition divfK := divfK.
+Definition invfM := invfM.
+Definition invf_div := invf_div.
+Definition expfB_cond := expfB_cond.
+Definition expfB := expfB.
+Definition prodfV := prodfV.
+Definition prodf_div := prodf_div.
+Definition telescope_prodf := telescope_prodf.
+Definition addf_div := addf_div.
+Definition mulf_div := mulf_div.
+Definition char0_natf_div := char0_natf_div.
+Definition fpredMr := fpredMr.
+Definition fpredMl := fpredMl.
+Definition fpred_divr := fpred_divr.
+Definition fpred_divl := fpred_divl.
+Definition satP {F e f} := @satP F e f.
+Definition eq_sat := eq_sat.
+Definition solP {F n f} := @solP F n f.
+Definition eq_sol := eq_sol.
+Definition size_sol := size_sol.
+Definition solve_monicpoly := solve_monicpoly.
+Definition raddf0 := raddf0.
+Definition raddf_eq0 := raddf_eq0.
+Definition raddfN := raddfN.
+Definition raddfD := raddfD.
+Definition raddfB := raddfB.
+Definition raddf_sum := raddf_sum.
+Definition raddfMn := raddfMn.
+Definition raddfMNn := raddfMNn.
+Definition raddfMnat := raddfMnat.
+Definition raddfMsign := raddfMsign.
+Definition can2_additive := can2_additive.
+Definition bij_additive := bij_additive.
+Definition rmorph0 := rmorph0.
+Definition rmorphN := rmorphN.
+Definition rmorphD := rmorphD.
+Definition rmorphB := rmorphB.
+Definition rmorph_sum := rmorph_sum.
+Definition rmorphMn := rmorphMn.
+Definition rmorphMNn := rmorphMNn.
+Definition rmorphismP := rmorphismP.
+Definition rmorphismMP := rmorphismMP.
+Definition rmorph1 := rmorph1.
+Definition rmorph_eq1 := rmorph_eq1.
+Definition rmorphM := rmorphM.
+Definition rmorphMsign := rmorphMsign.
+Definition rmorph_nat := rmorph_nat.
+Definition rmorph_eq_nat := rmorph_eq_nat.
+Definition rmorph_prod := rmorph_prod.
+Definition rmorphX := rmorphX.
+Definition rmorphN1 := rmorphN1.
+Definition rmorph_sign := rmorph_sign.
+Definition rmorph_char := rmorph_char.
+Definition can2_rmorphism := can2_rmorphism.
+Definition bij_rmorphism := bij_rmorphism.
+Definition rmorph_comm := rmorph_comm.
+Definition rmorph_unit := rmorph_unit.
+Definition rmorphV := rmorphV.
+Definition rmorph_div := rmorph_div.
+Definition fmorph_eq0 := fmorph_eq0.
+Definition fmorph_inj := @fmorph_inj.
+Implicit Arguments fmorph_inj [[F] [R] x1 x2].
+Definition fmorph_eq1 := fmorph_eq1.
+Definition fmorph_char := fmorph_char.
+Definition fmorph_unit := fmorph_unit.
+Definition fmorphV := fmorphV.
+Definition fmorph_div := fmorph_div.
+Definition scalerA := scalerA.
+Definition scale1r := scale1r.
+Definition scalerDr := scalerDr.
+Definition scalerDl := scalerDl.
+Definition scaler0 := scaler0.
+Definition scale0r := scale0r.
+Definition scaleNr := scaleNr.
+Definition scaleN1r := scaleN1r.
+Definition scalerN := scalerN.
+Definition scalerBl := scalerBl.
+Definition scalerBr := scalerBr.
+Definition scaler_nat := scaler_nat.
+Definition scalerMnl := scalerMnl.
+Definition scalerMnr := scalerMnr.
+Definition scaler_suml := scaler_suml.
+Definition scaler_sumr := scaler_sumr.
+Definition scaler_eq0 := scaler_eq0.
+Definition scalerK := scalerK.
+Definition scalerKV := scalerKV.
+Definition scalerI := scalerI.
+Definition scalerAl := scalerAl.
+Definition mulr_algl := mulr_algl.
+Definition scaler_sign := scaler_sign.
+Definition signrZK := signrZK.
+Definition scalerCA := scalerCA.
+Definition scalerAr := scalerAr.
+Definition mulr_algr := mulr_algr.
+Definition exprZn := exprZn.
+Definition scaler_prodl := scaler_prodl.
+Definition scaler_prodr := scaler_prodr.
+Definition scaler_prod := scaler_prod.
+Definition scaler_injl := scaler_injl.
+Definition scaler_unit := scaler_unit.
+Definition invrZ := invrZ.
+Definition raddfZnat := raddfZnat.
+Definition raddfZsign := raddfZsign.
+Definition in_algE := in_algE.
+Definition linear0 := linear0.
+Definition linearN := linearN.
+Definition linearD := linearD.
+Definition linearB := linearB.
+Definition linear_sum := linear_sum.
+Definition linearMn := linearMn.
+Definition linearMNn := linearMNn.
+Definition linearP := linearP.
+Definition linearZ_LR := linearZ_LR.
+Definition linearZ := linearZ.
+Definition linearPZ := linearPZ.
+Definition linearZZ := linearZZ.
+Definition scalarP := scalarP.
+Definition scalarZ := scalarZ.
+Definition can2_linear := can2_linear.
+Definition bij_linear := bij_linear.
+Definition rmorph_alg := rmorph_alg.
+Definition lrmorphismP := lrmorphismP.
+Definition can2_lrmorphism := can2_lrmorphism.
+Definition bij_lrmorphism := bij_lrmorphism.
+
+Notation null_fun V := (null_fun V) (only parsing).
+Notation in_alg A := (in_alg_loc A).
+
+End Theory.
+
+Notation in_alg A := (in_alg_loc A).
+
+End GRing.
+
+Export Zmodule.Exports Ring.Exports Lmodule.Exports Lalgebra.Exports.
+Export Additive.Exports RMorphism.Exports Linear.Exports LRMorphism.Exports.
+Export ComRing.Exports Algebra.Exports UnitRing.Exports UnitAlgebra.Exports.
+Export ComUnitRing.Exports IntegralDomain.Exports Field.Exports.
+Export DecidableField.Exports ClosedField.Exports.
+Export Pred.Exports SubType.Exports.
+Notation QEdecFieldMixin := QEdecFieldMixin.
+
+Notation "0" := (zero _) : ring_scope.
+Notation "-%R" := (@opp _) : ring_scope.
+Notation "- x" := (opp x) : ring_scope.
+Notation "+%R" := (@add _).
+Notation "x + y" := (add x y) : ring_scope.
+Notation "x - y" := (add x (- y)) : ring_scope.
+Notation "x *+ n" := (natmul x n) : ring_scope.
+Notation "x *- n" := (opp (x *+ n)) : ring_scope.
+Notation "s `_ i" := (seq.nth 0%R s%R i) : ring_scope.
+Notation support := 0.-support.
+
+Notation "1" := (one _) : ring_scope.
+Notation "- 1" := (opp 1) : ring_scope.
+
+Notation "n %:R" := (natmul 1 n) : ring_scope.
+Notation "[ 'char' R ]" := (char (Phant R)) : ring_scope.
+Notation Frobenius_aut chRp := (Frobenius_aut chRp).
+Notation "*%R" := (@mul _).
+Notation "x * y" := (mul x y) : ring_scope.
+Notation "x ^+ n" := (exp x n) : ring_scope.
+Notation "x ^-1" := (inv x) : ring_scope.
+Notation "x ^- n" := (inv (x ^+ n)) : ring_scope.
+Notation "x / y" := (mul x y^-1) : ring_scope.
+
+Notation "*:%R" := (@scale _ _).
+Notation "a *: m" := (scale a m) : ring_scope.
+Notation "k %:A" := (scale k 1) : ring_scope.
+Notation "\0" := (null_fun _) : ring_scope.
+Notation "f \+ g" := (add_fun_head tt f g) : ring_scope.
+Notation "f \- g" := (sub_fun_head tt f g) : ring_scope.
+Notation "a \*: f" := (scale_fun_head tt a f) : ring_scope.
+Notation "x \*o f" := (mull_fun_head tt x f) : ring_scope.
+Notation "x \o* f" := (mulr_fun_head tt x f) : ring_scope.
+
+Notation "\sum_ ( i <- r | P ) F" :=
+ (\big[+%R/0%R]_(i <- r | P%B) F%R) : ring_scope.
+Notation "\sum_ ( i <- r ) F" :=
+ (\big[+%R/0%R]_(i <- r) F%R) : ring_scope.
+Notation "\sum_ ( m <= i < n | P ) F" :=
+ (\big[+%R/0%R]_(m <= i < n | P%B) F%R) : ring_scope.
+Notation "\sum_ ( m <= i < n ) F" :=
+ (\big[+%R/0%R]_(m <= i < n) F%R) : ring_scope.
+Notation "\sum_ ( i | P ) F" :=
+ (\big[+%R/0%R]_(i | P%B) F%R) : ring_scope.
+Notation "\sum_ i F" :=
+ (\big[+%R/0%R]_i F%R) : ring_scope.
+Notation "\sum_ ( i : t | P ) F" :=
+ (\big[+%R/0%R]_(i : t | P%B) F%R) (only parsing) : ring_scope.
+Notation "\sum_ ( i : t ) F" :=
+ (\big[+%R/0%R]_(i : t) F%R) (only parsing) : ring_scope.
+Notation "\sum_ ( i < n | P ) F" :=
+ (\big[+%R/0%R]_(i < n | P%B) F%R) : ring_scope.
+Notation "\sum_ ( i < n ) F" :=
+ (\big[+%R/0%R]_(i < n) F%R) : ring_scope.
+Notation "\sum_ ( i 'in' A | P ) F" :=
+ (\big[+%R/0%R]_(i in A | P%B) F%R) : ring_scope.
+Notation "\sum_ ( i 'in' A ) F" :=
+ (\big[+%R/0%R]_(i in A) F%R) : ring_scope.
+
+Notation "\prod_ ( i <- r | P ) F" :=
+ (\big[*%R/1%R]_(i <- r | P%B) F%R) : ring_scope.
+Notation "\prod_ ( i <- r ) F" :=
+ (\big[*%R/1%R]_(i <- r) F%R) : ring_scope.
+Notation "\prod_ ( m <= i < n | P ) F" :=
+ (\big[*%R/1%R]_(m <= i < n | P%B) F%R) : ring_scope.
+Notation "\prod_ ( m <= i < n ) F" :=
+ (\big[*%R/1%R]_(m <= i < n) F%R) : ring_scope.
+Notation "\prod_ ( i | P ) F" :=
+ (\big[*%R/1%R]_(i | P%B) F%R) : ring_scope.
+Notation "\prod_ i F" :=
+ (\big[*%R/1%R]_i F%R) : ring_scope.
+Notation "\prod_ ( i : t | P ) F" :=
+ (\big[*%R/1%R]_(i : t | P%B) F%R) (only parsing) : ring_scope.
+Notation "\prod_ ( i : t ) F" :=
+ (\big[*%R/1%R]_(i : t) F%R) (only parsing) : ring_scope.
+Notation "\prod_ ( i < n | P ) F" :=
+ (\big[*%R/1%R]_(i < n | P%B) F%R) : ring_scope.
+Notation "\prod_ ( i < n ) F" :=
+ (\big[*%R/1%R]_(i < n) F%R) : ring_scope.
+Notation "\prod_ ( i 'in' A | P ) F" :=
+ (\big[*%R/1%R]_(i in A | P%B) F%R) : ring_scope.
+Notation "\prod_ ( i 'in' A ) F" :=
+ (\big[*%R/1%R]_(i in A) F%R) : ring_scope.
+
+Canonical add_monoid.
+Canonical add_comoid.
+Canonical mul_monoid.
+Canonical mul_comoid.
+Canonical muloid.
+Canonical addoid.
+
+Canonical locked_additive.
+Canonical locked_rmorphism.
+Canonical locked_linear.
+Canonical locked_lrmorphism.
+Canonical idfun_additive.
+Canonical idfun_rmorphism.
+Canonical idfun_linear.
+Canonical idfun_lrmorphism.
+Canonical comp_additive.
+Canonical comp_rmorphism.
+Canonical comp_linear.
+Canonical comp_lrmorphism.
+Canonical opp_additive.
+Canonical opp_linear.
+Canonical scale_additive.
+Canonical scale_linear.
+Canonical null_fun_additive.
+Canonical null_fun_linear.
+Canonical scale_fun_additive.
+Canonical scale_fun_linear.
+Canonical add_fun_additive.
+Canonical add_fun_linear.
+Canonical sub_fun_additive.
+Canonical sub_fun_linear.
+Canonical mull_fun_additive.
+Canonical mull_fun_linear.
+Canonical mulr_fun_additive.
+Canonical mulr_fun_linear.
+Canonical Frobenius_aut_additive.
+Canonical Frobenius_aut_rmorphism.
+Canonical in_alg_additive.
+Canonical in_alg_rmorphism.
+
+Notation "R ^c" := (converse R) (at level 2, format "R ^c") : type_scope.
+Canonical converse_eqType.
+Canonical converse_choiceType.
+Canonical converse_zmodType.
+Canonical converse_ringType.
+Canonical converse_unitRingType.
+
+Notation "R ^o" := (regular R) (at level 2, format "R ^o") : type_scope.
+Canonical regular_eqType.
+Canonical regular_choiceType.
+Canonical regular_zmodType.
+Canonical regular_ringType.
+Canonical regular_lmodType.
+Canonical regular_lalgType.
+Canonical regular_comRingType.
+Canonical regular_algType.
+Canonical regular_unitRingType.
+Canonical regular_comUnitRingType.
+Canonical regular_unitAlgType.
+Canonical regular_idomainType.
+Canonical regular_fieldType.
+
+Canonical unit_keyed.
+Canonical unit_opprPred.
+Canonical unit_mulrPred.
+Canonical unit_smulrPred.
+Canonical unit_divrPred.
+Canonical unit_sdivrPred.
+
+Bind Scope term_scope with term.
+Bind Scope term_scope with formula.
+
+Notation "''X_' i" := (Var _ i) : term_scope.
+Notation "n %:R" := (NatConst _ n) : term_scope.
+Notation "0" := 0%:R%T : term_scope.
+Notation "1" := 1%:R%T : term_scope.
+Notation "x %:T" := (Const x) : term_scope.
+Infix "+" := Add : term_scope.
+Notation "- t" := (Opp t) : term_scope.
+Notation "t - u" := (Add t (- u)) : term_scope.
+Infix "*" := Mul : term_scope.
+Infix "*+" := NatMul : term_scope.
+Notation "t ^-1" := (Inv t) : term_scope.
+Notation "t / u" := (Mul t u^-1) : term_scope.
+Infix "^+" := Exp : term_scope.
+Infix "==" := Equal : term_scope.
+Notation "x != y" := (GRing.Not (x == y)) : term_scope.
+Infix "/\" := And : term_scope.
+Infix "\/" := Or : term_scope.
+Infix "==>" := Implies : term_scope.
+Notation "~ f" := (Not f) : term_scope.
+Notation "''exists' ''X_' i , f" := (Exists i f) : term_scope.
+Notation "''forall' ''X_' i , f" := (Forall i f) : term_scope.
+
+(* Lifting Structure from the codomain of finfuns. *)
+Section FinFunZmod.
+
+Variable (aT : finType) (rT : zmodType).
+Implicit Types f g : {ffun aT -> rT}.
+
+Definition ffun_zero := [ffun a : aT => (0 : rT)].
+Definition ffun_opp f := [ffun a => - f a].
+Definition ffun_add f g := [ffun a => f a + g a].
+
+Fact ffun_addA : associative ffun_add.
+Proof. by move=> f1 f2 f3; apply/ffunP=> a; rewrite !ffunE addrA. Qed.
+Fact ffun_addC : commutative ffun_add.
+Proof. by move=> f1 f2; apply/ffunP=> a; rewrite !ffunE addrC. Qed.
+Fact ffun_add0 : left_id ffun_zero ffun_add.
+Proof. by move=> f; apply/ffunP=> a; rewrite !ffunE add0r. Qed.
+Fact ffun_addN : left_inverse ffun_zero ffun_opp ffun_add.
+Proof. by move=> f; apply/ffunP=> a; rewrite !ffunE addNr. Qed.
+
+Definition ffun_zmodMixin :=
+ Zmodule.Mixin ffun_addA ffun_addC ffun_add0 ffun_addN.
+Canonical ffun_zmodType := Eval hnf in ZmodType _ ffun_zmodMixin.
+
+Section Sum.
+
+Variables (I : Type) (r : seq I) (P : pred I) (F : I -> {ffun aT -> rT}).
+
+Lemma sum_ffunE x : (\sum_(i <- r | P i) F i) x = \sum_(i <- r | P i) F i x.
+Proof. by elim/big_rec2: _ => // [|i _ y _ <-]; rewrite !ffunE. Qed.
+
+Lemma sum_ffun :
+ \sum_(i <- r | P i) F i = [ffun x => \sum_(i <- r | P i) F i x].
+Proof. by apply/ffunP=> i; rewrite sum_ffunE ffunE. Qed.
+
+End Sum.
+
+Lemma ffunMnE f n x : (f *+ n) x = f x *+ n.
+Proof. by rewrite -[n]card_ord -!sumr_const sum_ffunE. Qed.
+
+End FinFunZmod.
+Canonical exp_zmodType (M : zmodType) n := [zmodType of M ^ n].
+
+Section FinFunRing.
+
+(* As rings require 1 != 0 in order to lift a ring structure over finfuns *)
+(* we need evidence that the domain is non-empty. *)
+
+Variable (aT : finType) (R : ringType) (a : aT).
+
+Definition ffun_one : {ffun aT -> R} := [ffun => 1].
+Definition ffun_mul (f g : {ffun aT -> R}) := [ffun x => f x * g x].
+
+Fact ffun_mulA : associative ffun_mul.
+Proof. by move=> f1 f2 f3; apply/ffunP=> i; rewrite !ffunE mulrA. Qed.
+Fact ffun_mul_1l : left_id ffun_one ffun_mul.
+Proof. by move=> f; apply/ffunP=> i; rewrite !ffunE mul1r. Qed.
+Fact ffun_mul_1r : right_id ffun_one ffun_mul.
+Proof. by move=> f; apply/ffunP=> i; rewrite !ffunE mulr1. Qed.
+Fact ffun_mul_addl : left_distributive ffun_mul (@ffun_add _ _).
+Proof. by move=> f1 f2 f3; apply/ffunP=> i; rewrite !ffunE mulrDl. Qed.
+Fact ffun_mul_addr : right_distributive ffun_mul (@ffun_add _ _).
+Proof. by move=> f1 f2 f3; apply/ffunP=> i; rewrite !ffunE mulrDr. Qed.
+Fact ffun1_nonzero : ffun_one != 0.
+Proof. by apply/eqP => /ffunP/(_ a)/eqP; rewrite !ffunE oner_eq0. Qed.
+
+Definition ffun_ringMixin :=
+ RingMixin ffun_mulA ffun_mul_1l ffun_mul_1r ffun_mul_addl ffun_mul_addr
+ ffun1_nonzero.
+Definition ffun_ringType :=
+ Eval hnf in RingType {ffun aT -> R} ffun_ringMixin.
+
+End FinFunRing.
+
+Section FinFunComRing.
+
+Variable (aT : finType) (R : comRingType) (a : aT).
+
+Fact ffun_mulC : commutative (@ffun_mul aT R).
+Proof. by move=> f1 f2; apply/ffunP=> i; rewrite !ffunE mulrC. Qed.
+
+Definition ffun_comRingType :=
+ Eval hnf in ComRingType (ffun_ringType R a) ffun_mulC.
+
+End FinFunComRing.
+
+Section FinFunLmod.
+
+Variable (R : ringType) (aT : finType) (rT : lmodType R).
+
+Implicit Types f g : {ffun aT -> rT}.
+
+Definition ffun_scale k f := [ffun a => k *: f a].
+
+Fact ffun_scaleA k1 k2 f :
+ ffun_scale k1 (ffun_scale k2 f) = ffun_scale (k1 * k2) f.
+Proof. by apply/ffunP=> a; rewrite !ffunE scalerA. Qed.
+Fact ffun_scale1 : left_id 1 ffun_scale.
+Proof. by move=> f; apply/ffunP=> a; rewrite !ffunE scale1r. Qed.
+Fact ffun_scale_addr k : {morph (ffun_scale k) : x y / x + y}.
+Proof. by move=> f g; apply/ffunP=> a; rewrite !ffunE scalerDr. Qed.
+Fact ffun_scale_addl u : {morph (ffun_scale)^~ u : k1 k2 / k1 + k2}.
+Proof. by move=> k1 k2; apply/ffunP=> a; rewrite !ffunE scalerDl. Qed.
+
+Definition ffun_lmodMixin :=
+ LmodMixin ffun_scaleA ffun_scale1 ffun_scale_addr ffun_scale_addl.
+Canonical ffun_lmodType :=
+ Eval hnf in LmodType R {ffun aT -> rT} ffun_lmodMixin.
+
+End FinFunLmod.
+Canonical exp_lmodType (R : ringType) (M : lmodType R) n :=
+ [lmodType R of M ^ n].
+
+(* External direct product. *)
+Section PairZmod.
+
+Variables M1 M2 : zmodType.
+
+Definition opp_pair (x : M1 * M2) := (- x.1, - x.2).
+Definition add_pair (x y : M1 * M2) := (x.1 + y.1, x.2 + y.2).
+
+Fact pair_addA : associative add_pair.
+Proof. by move=> x y z; congr (_, _); apply: addrA. Qed.
+
+Fact pair_addC : commutative add_pair.
+Proof. by move=> x y; congr (_, _); apply: addrC. Qed.
+
+Fact pair_add0 : left_id (0, 0) add_pair.
+Proof. by case=> x1 x2; congr (_, _); apply: add0r. Qed.
+
+Fact pair_addN : left_inverse (0, 0) opp_pair add_pair.
+Proof. by move=> x; congr (_, _); apply: addNr. Qed.
+
+Definition pair_zmodMixin := ZmodMixin pair_addA pair_addC pair_add0 pair_addN.
+Canonical pair_zmodType := Eval hnf in ZmodType (M1 * M2) pair_zmodMixin.
+
+End PairZmod.
+
+Section PairRing.
+
+Variables R1 R2 : ringType.
+
+Definition mul_pair (x y : R1 * R2) := (x.1 * y.1, x.2 * y.2).
+
+Fact pair_mulA : associative mul_pair.
+Proof. by move=> x y z; congr (_, _); apply: mulrA. Qed.
+
+Fact pair_mul1l : left_id (1, 1) mul_pair.
+Proof. by case=> x1 x2; congr (_, _); apply: mul1r. Qed.
+
+Fact pair_mul1r : right_id (1, 1) mul_pair.
+Proof. by case=> x1 x2; congr (_, _); apply: mulr1. Qed.
+
+Fact pair_mulDl : left_distributive mul_pair +%R.
+Proof. by move=> x y z; congr (_, _); apply: mulrDl. Qed.
+
+Fact pair_mulDr : right_distributive mul_pair +%R.
+Proof. by move=> x y z; congr (_, _); apply: mulrDr. Qed.
+
+Fact pair_one_neq0 : (1, 1) != 0 :> R1 * R2.
+Proof. by rewrite xpair_eqE oner_eq0. Qed.
+
+Definition pair_ringMixin :=
+ RingMixin pair_mulA pair_mul1l pair_mul1r pair_mulDl pair_mulDr pair_one_neq0.
+Canonical pair_ringType := Eval hnf in RingType (R1 * R2) pair_ringMixin.
+
+End PairRing.
+
+Section PairComRing.
+
+Variables R1 R2 : comRingType.
+
+Fact pair_mulC : commutative (@mul_pair R1 R2).
+Proof. by move=> x y; congr (_, _); apply: mulrC. Qed.
+
+Canonical pair_comRingType := Eval hnf in ComRingType (R1 * R2) pair_mulC.
+
+End PairComRing.
+
+Section PairLmod.
+
+Variables (R : ringType) (V1 V2 : lmodType R).
+
+Definition scale_pair a (v : V1 * V2) : V1 * V2 := (a *: v.1, a *: v.2).
+
+Fact pair_scaleA a b u : scale_pair a (scale_pair b u) = scale_pair (a * b) u.
+Proof. by congr (_, _); apply: scalerA. Qed.
+
+Fact pair_scale1 u : scale_pair 1 u = u.
+Proof. by case: u => u1 u2; congr (_, _); apply: scale1r. Qed.
+
+Fact pair_scaleDr : right_distributive scale_pair +%R.
+Proof. by move=> a u v; congr (_, _); apply: scalerDr. Qed.
+
+Fact pair_scaleDl u : {morph scale_pair^~ u: a b / a + b}.
+Proof. by move=> a b; congr (_, _); apply: scalerDl. Qed.
+
+Definition pair_lmodMixin :=
+ LmodMixin pair_scaleA pair_scale1 pair_scaleDr pair_scaleDl.
+Canonical pair_lmodType := Eval hnf in LmodType R (V1 * V2) pair_lmodMixin.
+
+End PairLmod.
+
+Section PairLalg.
+
+Variables (R : ringType) (A1 A2 : lalgType R).
+
+Fact pair_scaleAl a (u v : A1 * A2) : a *: (u * v) = (a *: u) * v.
+Proof. by congr (_, _); apply: scalerAl. Qed.
+Canonical pair_lalgType := Eval hnf in LalgType R (A1 * A2) pair_scaleAl.
+
+End PairLalg.
+
+Section PairAlg.
+
+Variables (R : comRingType) (A1 A2 : algType R).
+
+Fact pair_scaleAr a (u v : A1 * A2) : a *: (u * v) = u * (a *: v).
+Proof. by congr (_, _); apply: scalerAr. Qed.
+Canonical pair_algType := Eval hnf in AlgType R (A1 * A2) pair_scaleAr.
+
+End PairAlg.
+
+Section PairUnitRing.
+
+Variables R1 R2 : unitRingType.
+
+Definition pair_unitr :=
+ [qualify a x : R1 * R2 | (x.1 \is a GRing.unit) && (x.2 \is a GRing.unit)].
+Definition pair_invr x :=
+ if x \is a pair_unitr then (x.1^-1, x.2^-1) else x.
+
+Lemma pair_mulVl : {in pair_unitr, left_inverse 1 pair_invr *%R}.
+Proof.
+rewrite /pair_invr=> x; case: ifP => // /andP[Ux1 Ux2] _.
+by congr (_, _); apply: mulVr.
+Qed.
+
+Lemma pair_mulVr : {in pair_unitr, right_inverse 1 pair_invr *%R}.
+Proof.
+rewrite /pair_invr=> x; case: ifP => // /andP[Ux1 Ux2] _.
+by congr (_, _); apply: mulrV.
+Qed.
+
+Lemma pair_unitP x y : y * x = 1 /\ x * y = 1 -> x \is a pair_unitr.
+Proof.
+case=> [[y1x y2x] [x1y x2y]]; apply/andP.
+by split; apply/unitrP; [exists y.1 | exists y.2].
+Qed.
+
+Lemma pair_invr_out : {in [predC pair_unitr], pair_invr =1 id}.
+Proof. by rewrite /pair_invr => x /negPf/= ->. Qed.
+
+Definition pair_unitRingMixin :=
+ UnitRingMixin pair_mulVl pair_mulVr pair_unitP pair_invr_out.
+Canonical pair_unitRingType :=
+ Eval hnf in UnitRingType (R1 * R2) pair_unitRingMixin.
+
+End PairUnitRing.
+
+Canonical pair_comUnitRingType (R1 R2 : comUnitRingType) :=
+ Eval hnf in [comUnitRingType of R1 * R2].
+
+Canonical pair_unitAlgType (R : comUnitRingType) (A1 A2 : unitAlgType R) :=
+ Eval hnf in [unitAlgType R of A1 * A2].
+
+(* begin hide *)
+
+(* Testing subtype hierarchy
+Section Test0.
+
+Variables (T : choiceType) (S : predPredType T).
+
+Inductive B := mkB x & x \in S.
+Definition vB u := let: mkB x _ := u in x.
+
+Canonical B_subType := [subType for vB].
+Definition B_eqMixin := [eqMixin of B by <:].
+Canonical B_eqType := EqType B B_eqMixin.
+Definition B_choiceMixin := [choiceMixin of B by <:].
+Canonical B_choiceType := ChoiceType B B_choiceMixin.
+
+End Test0.
+
+Section Test1.
+
+Variables (R : unitRingType) (S : pred R).
+Variables (ringS : divringPred S) (kS : keyed_pred ringS).
+
+Definition B_zmodMixin := [zmodMixin of B kS by <:].
+Canonical B_zmodType := ZmodType (B kS) B_zmodMixin.
+Definition B_ringMixin := [ringMixin of B kS by <:].
+Canonical B_ringType := RingType (B kS) B_ringMixin.
+Definition B_unitRingMixin := [unitRingMixin of B kS by <:].
+Canonical B_unitRingType := UnitRingType (B kS) B_unitRingMixin.
+
+End Test1.
+
+Section Test2.
+
+Variables (R : comUnitRingType) (A : unitAlgType R) (S : pred A).
+Variables (algS : divalgPred S) (kS : keyed_pred algS).
+
+Definition B_lmodMixin := [lmodMixin of B kS by <:].
+Canonical B_lmodType := LmodType R (B kS) B_lmodMixin.
+Definition B_lalgMixin := [lalgMixin of B kS by <:].
+Canonical B_lalgType := LalgType R (B kS) B_lalgMixin.
+Definition B_algMixin := [algMixin of B kS by <:].
+Canonical B_algType := AlgType R (B kS) B_algMixin.
+Canonical B_unitAlgType := [unitAlgType R of B kS].
+
+End Test2.
+
+Section Test3.
+
+Variables (F : fieldType) (S : pred F).
+Variables (ringS : divringPred S) (kS : keyed_pred ringS).
+
+Definition B_comRingMixin := [comRingMixin of B kS by <:].
+Canonical B_comRingType := ComRingType (B kS) B_comRingMixin.
+Canonical B_comUnitRingType := [comUnitRingType of B kS].
+Definition B_idomainMixin := [idomainMixin of B kS by <:].
+Canonical B_idomainType := IdomainType (B kS) B_idomainMixin.
+Definition B_fieldMixin := [fieldMixin of B kS by <:].
+Canonical B_fieldType := FieldType (B kS) B_fieldMixin.
+
+End Test3.
+
+*)
+
+(* end hide *)
diff --git a/mathcomp/algebra/ssrint.v b/mathcomp/algebra/ssrint.v
new file mode 100644
index 0000000..cbf726e
--- /dev/null
+++ b/mathcomp/algebra/ssrint.v
@@ -0,0 +1,1782 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat choice seq.
+Require Import fintype finfun bigop ssralg ssrnum poly.
+Import GRing.Theory Num.Theory.
+
+(******************************************************************************)
+(* This file develops a basic theory of signed integers, defining: *)
+(* int == the type of signed integers, with two constructors Posz for *)
+(* non-negative integers and Negz for negative integers. It *)
+(* supports the realDomainType interface (and its parents). *)
+(* n%:Z == explicit cast from nat to int (:= Posz n); displayed as n. *)
+(* However (Posz m = Posz n) is displayed as (m = n :> int) *)
+(* (and so are ==, != and <>) *)
+(* Lemma NegzE : turns (Negz n) into - n.+1%:Z. *)
+(* x *~ m == m times x, with m : int; *)
+(* convertible to x *+ n if m is Posz n *)
+(* convertible to x *- n.+1 if m is Negz n. *)
+(* m%:~R == the image of m : int in a generic ring (:= 1 *~ m). *)
+(* x ^ m == x to the m, with m : int; *)
+(* convertible to x ^+ n if m is Posz n *)
+(* convertible to x ^- n.+1 if m is Negz n. *)
+(* sgz x == sign of x : R, *)
+(* equals (0 : int) if and only x == 0, *)
+(* equals (1 : int) if x is positive *)
+(* and (-1 : int) otherwise. *)
+(* `|m|%N == the n : nat such that `|m|%R = n%:Z, for m : int. *)
+(* `|m - n|%N == the distance between m and n; the '-' is specialized to *)
+(* the int type, so m and n can be either of type nat or int *)
+(* thanks to the Posz coercion; m and n are however parsed in *)
+(* the %N scope. The IntDist submodule provides this notation *)
+(* and the corresponding theory independently of the rest of *)
+(* of the int and ssralg libraries (and notations). *)
+(* Warning: due to the declaration of Posz as a coercion, two terms might be *)
+(* displayed the same while not being convertible, for instance: *)
+(* (Posz (x - y)) and (Posz x) - (Posz y) for x, y : nat. *)
+(******************************************************************************)
+
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Delimit Scope int_scope with Z.
+Local Open Scope int_scope.
+
+(* Defining int *)
+CoInductive int : Set := Posz of nat | Negz of nat.
+(* This must be deferred to module DistInt to work around the design flaws of *)
+(* the Coq module system. *)
+(* Coercion Posz : nat >-> int. *)
+
+Notation "n %:Z" := (Posz n)
+ (at level 2, left associativity, format "n %:Z", only parsing) : int_scope.
+Notation "n %:Z" := (Posz n)
+ (at level 2, left associativity, format "n %:Z", only parsing) : ring_scope.
+
+Notation "n = m :> 'in' 't'" := (Posz n = Posz m)
+ (at level 70, m at next level, format "n = m :> 'in' 't'") : ring_scope.
+Notation "n == m :> 'in' 't'" := (Posz n == Posz m)
+ (at level 70, m at next level, format "n == m :> 'in' 't'") : ring_scope.
+Notation "n != m :> 'in' 't'" := (Posz n != Posz m)
+ (at level 70, m at next level, format "n != m :> 'in' 't'") : ring_scope.
+Notation "n <> m :> 'in' 't'" := (Posz n <> Posz m)
+ (at level 70, m at next level, format "n <> m :> 'in' 't'") : ring_scope.
+
+Definition natsum_of_int (m : int) : nat + nat :=
+ match m with Posz p => inl _ p | Negz n => inr _ n end.
+
+Definition int_of_natsum (m : nat + nat) :=
+ match m with inl p => Posz p | inr n => Negz n end.
+
+Lemma natsum_of_intK : cancel natsum_of_int int_of_natsum.
+Proof. by case. Qed.
+
+Definition int_eqMixin := CanEqMixin natsum_of_intK.
+Definition int_countMixin := CanCountMixin natsum_of_intK.
+Definition int_choiceMixin := CountChoiceMixin int_countMixin.
+Canonical int_eqType := Eval hnf in EqType int int_eqMixin.
+Canonical int_choiceType := Eval hnf in ChoiceType int int_choiceMixin.
+Canonical int_countType := Eval hnf in CountType int int_countMixin.
+
+Lemma eqz_nat (m n : nat) : (m%:Z == n%:Z) = (m == n). Proof. by []. Qed.
+
+Module intZmod.
+Section intZmod.
+
+Definition addz (m n : int) :=
+ match m, n with
+ | Posz m', Posz n' => Posz (m' + n')
+ | Negz m', Negz n' => Negz (m' + n').+1
+ | Posz m', Negz n' => if n' < m' then Posz (m' - n'.+1) else Negz (n' - m')
+ | Negz n', Posz m' => if n' < m' then Posz (m' - n'.+1) else Negz (n' - m')
+ end.
+
+Definition oppz m := nosimpl
+ match m with
+ | Posz n => if n is (n'.+1)%N then Negz n' else Posz 0
+ | Negz n => Posz (n.+1)%N
+ end.
+
+Local Notation "0" := (Posz 0) : int_scope.
+Local Notation "-%Z" := (@oppz) : int_scope.
+Local Notation "- x" := (oppz x) : int_scope.
+Local Notation "+%Z" := (@addz) : int_scope.
+Local Notation "x + y" := (addz x y) : int_scope.
+Local Notation "x - y" := (x + - y) : int_scope.
+
+Lemma PoszD : {morph Posz : m n / (m + n)%N >-> m + n}. Proof. by []. Qed.
+
+Local Coercion Posz : nat >-> int.
+
+Lemma NegzE (n : nat) : Negz n = - n.+1. Proof. by []. Qed.
+
+Lemma int_rect (P : int -> Type) :
+ P 0 -> (forall n : nat, P n -> P (n.+1))
+ -> (forall n : nat, P (- n) -> P (- (n.+1)))
+ -> forall n : int, P n.
+Proof.
+by move=> P0 hPp hPn []; elim=> [|n ihn]//; do? [apply: hPn | apply: hPp].
+Qed.
+
+Definition int_rec := int_rect.
+Definition int_ind := int_rect.
+
+CoInductive int_spec (x : int) : int -> Type :=
+| ZintNull of x = 0 : int_spec x 0
+| ZintPos n of x = n.+1 : int_spec x n.+1
+| ZintNeg n of x = - (n.+1)%:Z : int_spec x (- n.+1).
+
+Lemma intP x : int_spec x x. Proof. by move: x=> [] []; constructor. Qed.
+
+Lemma addzC : commutative addz.
+Proof. by move=> [] m [] n //=; rewrite addnC. Qed.
+
+Lemma add0z : left_id 0 addz. Proof. by move=> [] [|]. Qed.
+
+Lemma oppzK : involutive oppz. Proof. by do 2?case. Qed.
+
+Lemma oppz_add : {morph oppz : m n / m + n}.
+Proof.
+move=> [[|n]|n] [[|m]|m] /=; rewrite ?NegzE ?oppzK ?addnS ?addn0 ?subn0 //;
+ rewrite ?ltnS[m <= n]leqNgt [n <= m]leqNgt; case: ltngtP=> hmn /=;
+ by rewrite ?hmn ?subnn // ?oppzK ?subSS ?subnS ?prednK // ?subn_gt0.
+Qed.
+
+Lemma add1Pz (n : int) : 1 + (n - 1) = n.
+Proof. by case: (intP n)=> // n' /= _; rewrite ?(subn1, addn0). Qed.
+
+Lemma subSz1 (n : int) : 1 + n - 1 = n.
+Proof.
+by apply: (inv_inj oppzK); rewrite addzC !oppz_add oppzK [_ - n]addzC add1Pz.
+Qed.
+
+Lemma addSnz (m : nat) (n : int) : (m.+1%N) + n = 1 + (m + n).
+Proof.
+move: m n=> [|m] [] [|n] //=; rewrite ?add1n ?subn1 // !(ltnS, subSS).
+rewrite [n <= m]leqNgt; case: ltngtP=> hmn /=; rewrite ?hmn ?subnn //.
+ by rewrite subnS add1n prednK ?subn_gt0.
+by rewrite ltnS leqn0 subn_eq0 leqNgt hmn /= subnS subn1.
+Qed.
+
+Lemma addSz (m n : int) : (1 + m) + n = 1 + (m + n).
+Proof.
+case: m => [] m; first by rewrite -PoszD add1n addSnz.
+rewrite !NegzE; apply: (inv_inj oppzK).
+rewrite !oppz_add !oppzK addSnz [-1%:Z + _]addzC addSnz add1Pz.
+by rewrite [-1%:Z + _]addzC subSz1.
+Qed.
+
+Lemma addPz (m n : int) : (m - 1) + n = (m + n) - 1.
+Proof.
+by apply: (inv_inj oppzK); rewrite !oppz_add oppzK [_ + 1]addzC addSz addzC.
+Qed.
+
+Lemma addzA : associative addz.
+Proof.
+elim=> [|m ihm|m ihm] n p; first by rewrite !add0z.
+ by rewrite -add1n PoszD !addSz ihm.
+by rewrite -add1n addnC PoszD oppz_add !addPz ihm.
+Qed.
+
+Lemma addNz : left_inverse (0:int) oppz addz. Proof. by do 3?elim. Qed.
+
+Lemma predn_int (n : nat) : 0 < n -> n.-1%:Z = n - 1.
+Proof. by case: n=> // n _ /=; rewrite subn1. Qed.
+
+Definition Mixin := ZmodMixin addzA addzC add0z addNz.
+
+End intZmod.
+End intZmod.
+
+Canonical int_ZmodType := ZmodType int intZmod.Mixin.
+
+Local Open Scope ring_scope.
+
+Section intZmoduleTheory.
+
+Local Coercion Posz : nat >-> int.
+
+Lemma PoszD : {morph Posz : n m / (n + m)%N >-> n + m}. Proof. by []. Qed.
+
+Lemma NegzE (n : nat) : Negz n = -(n.+1)%:Z. Proof. by []. Qed.
+
+Lemma int_rect (P : int -> Type) :
+ P 0 -> (forall n : nat, P n -> P (n.+1)%N)
+ -> (forall n : nat, P (- (n%:Z)) -> P (- (n.+1%N%:Z)))
+ -> forall n : int, P n.
+Proof.
+by move=> P0 hPp hPn []; elim=> [|n ihn]//; do? [apply: hPn | apply: hPp].
+Qed.
+
+Definition int_rec := int_rect.
+Definition int_ind := int_rect.
+
+CoInductive int_spec (x : int) : int -> Type :=
+| ZintNull : int_spec x 0
+| ZintPos n : int_spec x n.+1
+| ZintNeg n : int_spec x (- (n.+1)%:Z).
+
+Lemma intP x : int_spec x x.
+Proof. by move: x=> [] [] *; rewrite ?NegzE; constructor. Qed.
+
+Definition oppz_add := (@opprD [zmodType of int]).
+
+Lemma subzn (m n : nat) : (n <= m)%N -> m%:Z - n%:Z = (m - n)%N.
+Proof.
+elim: n=> //= [|n ihn] hmn; first by rewrite subr0 subn0.
+rewrite subnS -addn1 !PoszD opprD addrA ihn 1?ltnW //.
+by rewrite intZmod.predn_int // subn_gt0.
+Qed.
+
+Lemma subzSS (m n : nat) : m.+1%:Z - n.+1%:Z = m%:Z - n%:Z.
+Proof. by elim: n m=> [|n ihn] m //; rewrite !subzn. Qed.
+
+End intZmoduleTheory.
+
+Module intRing.
+Section intRing.
+
+Local Coercion Posz : nat >-> int.
+
+Definition mulz (m n : int) :=
+ match m, n with
+ | Posz m', Posz n' => (m' * n')%N%:Z
+ | Negz m', Negz n' => (m'.+1%N * n'.+1%N)%N%:Z
+ | Posz m', Negz n' => - (m' * (n'.+1%N))%N%:Z
+ | Negz n', Posz m' => - (m' * (n'.+1%N))%N%:Z
+ end.
+
+Local Notation "1" := (1%N:int) : int_scope.
+Local Notation "*%Z" := (@mulz) : int_scope.
+Local Notation "x * y" := (mulz x y) : int_scope.
+
+Lemma mul0z : left_zero 0 *%Z.
+Proof. by case=> [n|[|n]] //=; rewrite muln0. Qed.
+
+Lemma mulzC : commutative mulz.
+Proof. by move=> [] m [] n //=; rewrite mulnC. Qed.
+
+Lemma mulz0 : right_zero 0 *%Z.
+Proof. by move=> x; rewrite mulzC mul0z. Qed.
+
+Lemma mulzN (m n : int) : (m * (- n))%Z = - (m * n)%Z.
+Proof.
+by case: (intP m)=> {m} [|m|m]; rewrite ?mul0z //;
+case: (intP n)=> {n} [|n|n]; rewrite ?mulz0 //= mulnC.
+Qed.
+
+Lemma mulNz (m n : int) : ((- m) * n)%Z = - (m * n)%Z.
+Proof. by rewrite mulzC mulzN mulzC. Qed.
+
+Lemma mulzA : associative mulz.
+Proof.
+by move=> [] m [] n [] p; rewrite ?NegzE ?(mulnA,mulNz,mulzN,opprK) //= ?mulnA.
+Qed.
+
+Lemma mul1z : left_id 1%Z mulz.
+Proof. by case=> [[|n]|n] //=; rewrite ?mul1n// plusE addn0. Qed.
+
+Lemma mulzS (x : int) (n : nat) : (x * n.+1%:Z)%Z = x + (x * n)%Z.
+Proof.
+by case: (intP x)=> [|m'|m'] //=; [rewrite mulnS|rewrite mulSn -opprD].
+Qed.
+
+Lemma mulz_addl : left_distributive mulz (+%R).
+Proof.
+move=> x y z; elim: z=> [|n|n]; first by rewrite !(mul0z,mulzC).
+ by rewrite !mulzS=> ->; rewrite !addrA [X in X + _]addrAC.
+rewrite !mulzN !mulzS -!opprD=> /(inv_inj (@opprK _))->.
+by rewrite !addrA [X in X + _]addrAC.
+Qed.
+
+Lemma nonzero1z : 1%Z != 0. Proof. by []. Qed.
+
+Definition comMixin := ComRingMixin mulzA mulzC mul1z mulz_addl nonzero1z.
+
+End intRing.
+End intRing.
+
+Canonical int_Ring := Eval hnf in RingType int intRing.comMixin.
+Canonical int_comRing := Eval hnf in ComRingType int intRing.mulzC.
+
+Section intRingTheory.
+
+Implicit Types m n : int.
+Local Coercion Posz : nat >-> int.
+
+Lemma PoszM : {morph Posz : n m / (n * m)%N >-> n * m}. Proof. by []. Qed.
+
+Lemma intS (n : nat) : n.+1%:Z = 1 + n%:Z. Proof. by rewrite -PoszD. Qed.
+
+Lemma predn_int (n : nat) : (0 < n)%N -> n.-1%:Z = n%:Z - 1.
+Proof. exact: intZmod.predn_int. Qed.
+
+End intRingTheory.
+
+Module intUnitRing.
+Section intUnitRing.
+Implicit Types m n : int.
+Local Coercion Posz : nat >-> int.
+
+Definition unitz := [qualify a n : int | (n == 1) || (n == -1)].
+Definition invz n : int := n.
+
+Lemma mulVz : {in unitz, left_inverse 1%R invz *%R}.
+Proof. by move=> n /pred2P[] ->. Qed.
+
+Lemma mulzn_eq1 m (n : nat) : (m * n == 1) = (m == 1) && (n == 1%N).
+Proof. by case: m=> m /=; [rewrite -PoszM [_==_]muln_eq1 | case: n]. Qed.
+
+Lemma unitzPl m n : n * m = 1 -> m \is a unitz.
+Proof.
+case: m => m; move/eqP; rewrite qualifE.
+* by rewrite mulzn_eq1; case/andP=> _; move/eqP->.
+* by rewrite NegzE intS mulrN -mulNr mulzn_eq1; case/andP=> _.
+Qed.
+
+Lemma invz_out : {in [predC unitz], invz =1 id}.
+Proof. exact. Qed.
+
+Lemma idomain_axiomz m n : m * n = 0 -> (m == 0) || (n == 0).
+Proof.
+by case: m n => m [] n //=; move/eqP; rewrite ?(NegzE,mulrN,mulNr);
+ rewrite ?(inv_eq (@opprK _)) -PoszM [_==_]muln_eq0.
+Qed.
+
+Definition comMixin := ComUnitRingMixin mulVz unitzPl invz_out.
+
+End intUnitRing.
+End intUnitRing.
+
+Canonical int_unitRingType :=
+ Eval hnf in UnitRingType int intUnitRing.comMixin.
+Canonical int_comUnitRing := Eval hnf in [comUnitRingType of int].
+Canonical int_iDomain :=
+ Eval hnf in IdomainType int intUnitRing.idomain_axiomz.
+
+Definition absz m := match m with Posz p => p | Negz n => n.+1 end.
+Notation "m - n" :=
+ (@GRing.add int_ZmodType m%N (@GRing.opp int_ZmodType n%N)) : distn_scope.
+Arguments Scope absz [distn_scope].
+Local Notation "`| m |" := (absz m) : nat_scope.
+
+Module intOrdered.
+Section intOrdered.
+Implicit Types m n p : int.
+Local Coercion Posz : nat >-> int.
+
+Local Notation normz m := (absz m)%:Z.
+
+Definition lez m n :=
+ match m, n with
+ | Posz m', Posz n' => (m' <= n')%N
+ | Posz m', Negz n' => false
+ | Negz m', Posz n' => true
+ | Negz m', Negz n' => (n' <= m')%N
+ end.
+
+Definition ltz m n :=
+ match m, n with
+ | Posz m', Posz n' => (m' < n')%N
+ | Posz m', Negz n' => false
+ | Negz m', Posz n' => true
+ | Negz m', Negz n' => (n' < m')%N
+ end.
+
+Fact lez_norm_add x y : lez (normz (x + y)) (normz x + normz y).
+Proof.
+move: x y=> [] m [] n; rewrite /= ?addnS //=;
+rewrite /GRing.add /GRing.Zmodule.add /=; case: ltnP=> //=;
+rewrite ?addSn ?ltnS ?leq_subLR ?(addnS, addSn) ?(leq_trans _ (leqnSn _)) //;
+by rewrite 1?addnCA ?leq_addr ?addnA ?leq_addl.
+Qed.
+
+Fact ltz_add x y : ltz 0 x -> ltz 0 y -> ltz 0 (x + y).
+Proof. by move: x y => [] x [] y //= hx hy; rewrite ltn_addr. Qed.
+
+Fact eq0_normz x : normz x = 0 -> x = 0. Proof. by case: x. Qed.
+
+Fact lez_total x y : lez x y || lez y x.
+Proof. by move: x y => [] x [] y //=; apply: leq_total. Qed.
+
+Lemma abszN (n : nat) : absz (- n%:Z) = n. Proof. by case: n. Qed.
+
+Fact normzM : {morph (fun n => normz n) : x y / x * y}.
+Proof. by move=> [] x [] y; rewrite // abszN // mulnC. Qed.
+
+Lemma subz_ge0 m n : lez 0 (n - m) = lez m n.
+Proof.
+case: (intP m); case: (intP n)=> // {m n} m n /=;
+rewrite ?ltnS -?opprD ?opprB ?subzSS; case: leqP=> // hmn;
+by [ rewrite subzn //
+ | rewrite -opprB subzn ?(ltnW hmn) //;
+ move: hmn; rewrite -subn_gt0; case: (_ - _)%N].
+Qed.
+
+Fact lez_def x y : (lez x y) = (normz (y - x) == y - x).
+Proof. by rewrite -subz_ge0; move: (_ - _) => [] n //=; rewrite eqxx. Qed.
+
+Fact ltz_def x y : (ltz x y) = (y != x) && (lez x y).
+Proof.
+by move: x y=> [] x [] y //=; rewrite (ltn_neqAle, leq_eqVlt) // eq_sym.
+Qed.
+
+Definition Mixin :=
+ NumMixin lez_norm_add ltz_add eq0_normz (in2W lez_total) normzM
+ lez_def ltz_def.
+
+End intOrdered.
+End intOrdered.
+
+Canonical int_numDomainType := NumDomainType int intOrdered.Mixin.
+Canonical int_realDomainType := RealDomainType int (intOrdered.lez_total 0).
+
+Section intOrderedTheory.
+
+Local Coercion Posz : nat >-> int.
+Implicit Types m n p : nat.
+Implicit Types x y z : int.
+
+Lemma lez_nat m n : (m <= n :> int) = (m <= n)%N.
+Proof. by []. Qed.
+
+Lemma ltz_nat m n : (m < n :> int) = (m < n)%N.
+Proof. by rewrite ltnNge ltrNge lez_nat. Qed.
+
+Definition ltez_nat := (lez_nat, ltz_nat).
+
+Lemma leNz_nat m n : (- m%:Z <= n). Proof. by case: m. Qed.
+
+Lemma ltNz_nat m n : (- m%:Z < n) = (m != 0%N) || (n != 0%N).
+Proof. by move: m n=> [|?] []. Qed.
+
+Definition lteNz_nat := (leNz_nat, ltNz_nat).
+
+Lemma lezN_nat m n : (m%:Z <= - n%:Z) = (m == 0%N) && (n == 0%N).
+Proof. by move: m n=> [|?] []. Qed.
+
+Lemma ltzN_nat m n : (m%:Z < - n%:Z) = false.
+Proof. by move: m n=> [|?] []. Qed.
+
+Lemma le0z_nat n : 0 <= n :> int. Proof. by []. Qed.
+
+Lemma lez0_nat n : n <= 0 :> int = (n == 0%N :> nat). Proof. by elim: n. Qed.
+
+Definition ltezN_nat := (lezN_nat, ltzN_nat).
+Definition ltez_natE := (ltez_nat, lteNz_nat, ltezN_nat, le0z_nat, lez0_nat).
+
+Lemma gtz0_ge1 x : (0 < x) = (1 <= x). Proof. by case: (intP x). Qed.
+
+Lemma lez_add1r x y : (1 + x <= y) = (x < y).
+Proof. by rewrite -subr_gt0 gtz0_ge1 lter_sub_addr. Qed.
+
+Lemma lez_addr1 x y : (x + 1 <= y) = (x < y).
+Proof. by rewrite addrC lez_add1r. Qed.
+
+Lemma ltz_add1r x y : (x < 1 + y) = (x <= y).
+Proof. by rewrite -lez_add1r ler_add2l. Qed.
+
+Lemma ltz_addr1 x y : (x < y + 1) = (x <= y).
+Proof. by rewrite -lez_addr1 ler_add2r. Qed.
+
+End intOrderedTheory.
+
+Bind Scope ring_scope with int.
+
+(* definition of intmul *)
+Definition intmul (R : zmodType) (x : R) (n : int) := nosimpl
+ match n with
+ | Posz n => (x *+ n)%R
+ | Negz n => (x *- (n.+1))%R
+ end.
+
+Notation "*~%R" := (@intmul _) (at level 0, format " *~%R") : ring_scope.
+Notation "x *~ n" := (intmul x n)
+ (at level 40, left associativity, format "x *~ n") : ring_scope.
+Notation intr := ( *~%R 1).
+Notation "n %:~R" := (1 *~ n)%R
+ (at level 2, left associativity, format "n %:~R") : ring_scope.
+
+Lemma pmulrn (R : zmodType) (x : R) (n : nat) : x *+ n = x *~ n%:Z.
+Proof. by []. Qed.
+
+Lemma nmulrn (R : zmodType) (x : R) (n : nat) : x *- n = x *~ - n%:Z.
+Proof. by case: n=> [] //; rewrite ?oppr0. Qed.
+
+Section ZintLmod.
+
+Definition zmodule (M : Type) : Type := M.
+Local Notation "M ^z" := (zmodule M) (at level 2, format "M ^z") : type_scope.
+Local Coercion Posz : nat >-> int.
+
+Variable M : zmodType.
+
+Implicit Types m n : int.
+Implicit Types x y z : M.
+
+Fact mulrzA_C m n x : (x *~ n) *~ m = x *~ (m * n).
+Proof.
+elim: m=> [|m _|m _]; elim: n=> [|n _|n _]; rewrite /intmul //=;
+rewrite ?(muln0, mulr0n, mul0rn, oppr0, mulNrn, opprK) //;
+ do ?by rewrite mulnC mulrnA.
+* by rewrite -mulrnA mulnC.
+* by rewrite -mulrnA.
+Qed.
+
+Fact mulrzAC m n x : (x *~ n) *~ m = (x *~ m) *~ n.
+Proof. by rewrite !mulrzA_C mulrC. Qed.
+
+Fact mulr1z (x : M) : x *~ 1 = x. Proof. by []. Qed.
+
+Fact mulrzDr m : {morph ( *~%R^~ m : M -> M) : x y / x + y}.
+Proof.
+by elim: m=> [|m _|m _] x y;
+ rewrite ?addr0 /intmul //= ?mulrnDl // opprD.
+Qed.
+
+Lemma mulrzBl_nat (m n : nat) x : x *~ (m%:Z - n%:Z) = x *~ m - x *~ n.
+Proof.
+case: (leqP m n)=> hmn; rewrite /intmul //=.
+ rewrite addrC -{1}[m:int]opprK -opprD subzn //.
+ rewrite -{2}[n](@subnKC m)// mulrnDr opprD addrA subrr sub0r.
+ by case hdmn: (_ - _)%N=> [|dmn] /=; first by rewrite mulr0n oppr0.
+have hnm := ltnW hmn.
+rewrite -{2}[m](@subnKC n)// mulrnDr addrAC subrr add0r.
+by rewrite subzn.
+Qed.
+
+Fact mulrzDl x : {morph *~%R x : m n / m + n}.
+Proof.
+elim=> [|m _|m _]; elim=> [|n _|n _]; rewrite /intmul //=;
+rewrite -?(opprD) ?(add0r, addr0, mulrnDr, subn0) //.
+* by rewrite -/(intmul _ _) mulrzBl_nat.
+* by rewrite -/(intmul _ _) addrC mulrzBl_nat addrC.
+* by rewrite -addnS -addSn mulrnDr.
+Qed.
+
+Definition Mint_LmodMixin :=
+ @LmodMixin _ [zmodType of M] (fun n x => x *~ n)
+ mulrzA_C mulr1z mulrzDr mulrzDl.
+Canonical Mint_LmodType := LmodType int M^z Mint_LmodMixin.
+
+Lemma scalezrE n x : n *: (x : M^z) = x *~ n. Proof. by []. Qed.
+
+Lemma mulrzA x m n : x *~ (m * n) = x *~ m *~ n.
+Proof. by rewrite -!scalezrE scalerA mulrC. Qed.
+
+Lemma mulr0z x : x *~ 0 = 0. Proof. by []. Qed.
+
+Lemma mul0rz n : 0 *~ n = 0 :> M.
+Proof. by rewrite -scalezrE scaler0. Qed.
+
+Lemma mulrNz x n : x *~ (- n) = - (x *~ n).
+Proof. by rewrite -scalezrE scaleNr. Qed.
+
+Lemma mulrN1z x : x *~ (- 1) = - x. Proof. by rewrite -scalezrE scaleN1r. Qed.
+
+Lemma mulNrz x n : (- x) *~ n = - (x *~ n).
+Proof. by rewrite -scalezrE scalerN. Qed.
+
+Lemma mulrzBr x m n : x *~ (m - n) = x *~ m - x *~ n.
+Proof. by rewrite -scalezrE scalerBl. Qed.
+
+Lemma mulrzBl x y n : (x - y) *~ n = x *~ n - y *~ n.
+Proof. by rewrite -scalezrE scalerBr. Qed.
+
+Lemma mulrz_nat (n : nat) x : x *~ n%:R = x *+ n.
+Proof. by rewrite -scalezrE scaler_nat. Qed.
+
+Lemma mulrz_sumr : forall x I r (P : pred I) F,
+ x *~ (\sum_(i <- r | P i) F i) = \sum_(i <- r | P i) x *~ F i.
+Proof. by rewrite -/M^z; exact: scaler_suml. Qed.
+
+Lemma mulrz_suml : forall n I r (P : pred I) (F : I -> M),
+ (\sum_(i <- r | P i) F i) *~ n= \sum_(i <- r | P i) F i *~ n.
+Proof. by rewrite -/M^z; exact: scaler_sumr. Qed.
+
+Canonical intmul_additive x := Additive (@mulrzBr x).
+
+End ZintLmod.
+
+Lemma ffunMzE (I : finType) (M : zmodType) (f : {ffun I -> M}) z x :
+ (f *~ z) x = f x *~ z.
+Proof. by case: z => n; rewrite ?ffunE ffunMnE. Qed.
+
+Lemma intz (n : int) : n%:~R = n.
+Proof.
+elim: n=> //= n ihn; rewrite /intmul /=.
+ by rewrite -addn1 mulrnDr /= PoszD -ihn.
+by rewrite nmulrn intS opprD mulrzDl ihn.
+Qed.
+
+Lemma natz (n : nat) : n%:R = n%:Z :> int.
+Proof. by rewrite pmulrn intz. Qed.
+
+Section RintMod.
+
+Local Coercion Posz : nat >-> int.
+Variable R : ringType.
+
+Implicit Types m n : int.
+Implicit Types x y z : R.
+
+Lemma mulrzAl n x y : (x *~ n) * y = (x * y) *~ n.
+Proof.
+by elim: n=> //= *; rewrite ?mul0r ?mulr0z // /intmul /= -mulrnAl -?mulNr.
+Qed.
+
+Lemma mulrzAr n x y : x * (y *~ n) = (x * y) *~ n.
+Proof.
+by elim: n=> //= *; rewrite ?mulr0 ?mulr0z // /intmul /= -mulrnAr -?mulrN.
+Qed.
+
+Lemma mulrzl x n : n%:~R * x = x *~ n.
+Proof. by rewrite mulrzAl mul1r. Qed.
+
+Lemma mulrzr x n : x * n%:~R = x *~ n.
+Proof. by rewrite mulrzAr mulr1. Qed.
+
+Lemma mulNrNz n x : (-x) *~ (-n) = x *~ n.
+Proof. by rewrite mulNrz mulrNz opprK. Qed.
+
+Lemma mulrbz x (b : bool) : x *~ b = (if b then x else 0).
+Proof. by case: b. Qed.
+
+Lemma intrD m n : (m + n)%:~R = m%:~R + n%:~R :> R.
+Proof. exact: mulrzDl. Qed.
+
+Lemma intrM m n : (m * n)%:~R = m%:~R * n%:~R :> R.
+Proof. by rewrite mulrzA -mulrzr. Qed.
+
+Lemma intmul1_is_rmorphism : rmorphism ( *~%R (1 : R)).
+Proof.
+by do ?split; move=> // x y /=; rewrite ?intrD ?mulrNz ?intrM.
+Qed.
+Canonical intmul1_rmorphism := RMorphism intmul1_is_rmorphism.
+
+Lemma mulr2z n : n *~ 2 = n + n. Proof. exact: mulr2n. Qed.
+
+End RintMod.
+
+Lemma mulrzz m n : m *~ n = m * n. Proof. by rewrite -mulrzr intz. Qed.
+
+Lemma mulz2 n : n * 2%:Z = n + n. Proof. by rewrite -mulrzz. Qed.
+
+Lemma mul2z n : 2%:Z * n = n + n. Proof. by rewrite mulrC -mulrzz. Qed.
+
+Section LMod.
+
+Variable R : ringType.
+Variable V : (lmodType R).
+Local Coercion Posz : nat >-> int.
+
+Implicit Types m n : int.
+Implicit Types x y z : R.
+Implicit Types u v w : V.
+
+Lemma scaler_int n v : n%:~R *: v = v *~ n.
+Proof.
+elim: n=> [|n ihn|n ihn]; first by rewrite scale0r.
+ by rewrite intS !mulrzDl scalerDl ihn scale1r.
+by rewrite intS opprD !mulrzDl scalerDl ihn scaleN1r.
+Qed.
+
+Lemma scalerMzl a v n : (a *: v) *~ n = (a *~ n) *: v.
+Proof. by rewrite -mulrzl -scaler_int scalerA. Qed.
+
+Lemma scalerMzr a v n : (a *: v) *~ n = a *: (v *~ n).
+Proof. by rewrite -!scaler_int !scalerA mulrzr mulrzl. Qed.
+
+End LMod.
+
+Lemma mulrz_int (M : zmodType) (n : int) (x : M) : x *~ n%:~R = x *~ n.
+Proof. by rewrite -scalezrE scaler_int. Qed.
+
+Section MorphTheory.
+Local Coercion Posz : nat >-> int.
+Section Additive.
+Variables (U V : zmodType) (f : {additive U -> V}).
+
+Lemma raddfMz n : {morph f : x / x *~ n}.
+Proof.
+case: n=> n x /=; first exact: raddfMn.
+by rewrite NegzE !mulrNz; apply: raddfMNn.
+Qed.
+
+End Additive.
+
+Section Multiplicative.
+
+Variables (R S : ringType) (f : {rmorphism R -> S}).
+
+Lemma rmorphMz : forall n, {morph f : x / x *~ n}. Proof. exact: raddfMz. Qed.
+
+Lemma rmorph_int : forall n, f n%:~R = n%:~R.
+Proof. by move=> n; rewrite rmorphMz rmorph1. Qed.
+
+End Multiplicative.
+
+Section Linear.
+
+Variable R : ringType.
+Variables (U V : lmodType R) (f : {linear U -> V}).
+
+Lemma linearMn : forall n, {morph f : x / x *~ n}. Proof. exact: raddfMz. Qed.
+
+End Linear.
+
+Lemma raddf_int_scalable (aV rV : lmodType int) (f : {additive aV -> rV}) :
+ scalable f.
+Proof. by move=> z u; rewrite -[z]intz !scaler_int raddfMz. Qed.
+
+Section Zintmul1rMorph.
+
+Variable R : ringType.
+
+Lemma commrMz (x y : R) n : GRing.comm x y -> GRing.comm x (y *~ n).
+Proof. by rewrite /GRing.comm=> com_xy; rewrite mulrzAr mulrzAl com_xy. Qed.
+
+Lemma commr_int (x : R) n : GRing.comm x n%:~R.
+Proof. by apply: commrMz; apply: commr1. Qed.
+
+End Zintmul1rMorph.
+
+Section ZintBigMorphism.
+
+Variable R : ringType.
+
+Lemma sumMz : forall I r (P : pred I) F,
+ (\sum_(i <- r | P i) F i)%N%:~R = \sum_(i <- r | P i) ((F i)%:~R) :> R.
+Proof. by apply: big_morph=> // x y; rewrite !pmulrn -rmorphD. Qed.
+
+Lemma prodMz : forall I r (P : pred I) F,
+ (\prod_(i <- r | P i) F i)%N%:~R = \prod_(i <- r | P i) ((F i)%:~R) :> R.
+Proof. by apply: big_morph=> // x y; rewrite !pmulrn PoszM -rmorphM. Qed.
+
+End ZintBigMorphism.
+
+Section Frobenius.
+
+Variable R : ringType.
+Implicit Types x y : R.
+
+Variable p : nat.
+Hypothesis charFp : p \in [char R].
+
+Local Notation "x ^f" := (Frobenius_aut charFp x).
+
+Lemma Frobenius_autMz x n : (x *~ n)^f = x^f *~ n.
+Proof.
+case: n=> n /=; first exact: Frobenius_autMn.
+by rewrite !NegzE !mulrNz Frobenius_autN Frobenius_autMn.
+Qed.
+
+Lemma Frobenius_aut_int n : (n%:~R)^f = n%:~R.
+Proof. by rewrite Frobenius_autMz Frobenius_aut1. Qed.
+
+End Frobenius.
+
+Section NumMorphism.
+
+Section PO.
+
+Variables (R : numDomainType).
+
+Implicit Types n m : int.
+Implicit Types x y : R.
+
+Lemma rmorphzP (f : {rmorphism int -> R}) : f =1 ( *~%R 1).
+Proof.
+move=> n; wlog : n / 0 <= n; case: n=> [] n //; do ?exact.
+ by rewrite NegzE !rmorphN=>->.
+move=> _; elim: n=> [|n ihn]; first by rewrite rmorph0.
+by rewrite intS !rmorphD !rmorph1 ihn.
+Qed.
+
+(* intmul and ler/ltr *)
+Lemma ler_pmulz2r n (hn : 0 < n) : {mono *~%R^~ n :x y / x <= y :> R}.
+Proof. by move=> x y; case: n hn=> [[]|] // n _; rewrite ler_pmuln2r. Qed.
+
+Lemma ltr_pmulz2r n (hn : 0 < n) : {mono *~%R^~ n : x y / x < y :> R}.
+Proof. exact: lerW_mono (ler_pmulz2r _). Qed.
+
+Lemma ler_nmulz2r n (hn : n < 0) : {mono *~%R^~ n : x y /~ x <= y :> R}.
+Proof.
+move=> x y /=; rewrite -![_ *~ n]mulNrNz.
+by rewrite ler_pmulz2r (oppr_cp0, ler_opp2).
+Qed.
+
+Lemma ltr_nmulz2r n (hn : n < 0) : {mono *~%R^~ n : x y /~ x < y :> R}.
+Proof. exact: lerW_nmono (ler_nmulz2r _). Qed.
+
+Lemma ler_wpmulz2r n (hn : 0 <= n) : {homo *~%R^~ n : x y / x <= y :> R}.
+Proof. by move=> x y xy; case: n hn=> [] // n _; rewrite ler_wmuln2r. Qed.
+
+Lemma ler_wnmulz2r n (hn : n <= 0) : {homo *~%R^~ n : x y /~ x <= y :> R}.
+Proof.
+by move=> x y xy /=; rewrite -ler_opp2 -!mulrNz ler_wpmulz2r // oppr_ge0.
+Qed.
+
+Lemma mulrz_ge0 x n (x0 : 0 <= x) (n0 : 0 <= n) : 0 <= x *~ n.
+Proof. by rewrite -(mul0rz _ n) ler_wpmulz2r. Qed.
+
+Lemma mulrz_le0 x n (x0 : x <= 0) (n0 : n <= 0) : 0 <= x *~ n.
+Proof. by rewrite -(mul0rz _ n) ler_wnmulz2r. Qed.
+
+Lemma mulrz_ge0_le0 x n (x0 : 0 <= x) (n0 : n <= 0) : x *~ n <= 0.
+Proof. by rewrite -(mul0rz _ n) ler_wnmulz2r. Qed.
+
+Lemma mulrz_le0_ge0 x n (x0 : x <= 0) (n0 : 0 <= n) : x *~ n <= 0.
+Proof. by rewrite -(mul0rz _ n) ler_wpmulz2r. Qed.
+
+Lemma pmulrz_lgt0 x n (n0 : 0 < n) : 0 < x *~ n = (0 < x).
+Proof. by rewrite -(mul0rz _ n) ltr_pmulz2r // mul0rz. Qed.
+
+Lemma nmulrz_lgt0 x n (n0 : n < 0) : 0 < x *~ n = (x < 0).
+Proof. by rewrite -(mul0rz _ n) ltr_nmulz2r // mul0rz. Qed.
+
+Lemma pmulrz_llt0 x n (n0 : 0 < n) : x *~ n < 0 = (x < 0).
+Proof. by rewrite -(mul0rz _ n) ltr_pmulz2r // mul0rz. Qed.
+
+Lemma nmulrz_llt0 x n (n0 : n < 0) : x *~ n < 0 = (0 < x).
+Proof. by rewrite -(mul0rz _ n) ltr_nmulz2r // mul0rz. Qed.
+
+Lemma pmulrz_lge0 x n (n0 : 0 < n) : 0 <= x *~ n = (0 <= x).
+Proof. by rewrite -(mul0rz _ n) ler_pmulz2r // mul0rz. Qed.
+
+Lemma nmulrz_lge0 x n (n0 : n < 0) : 0 <= x *~ n = (x <= 0).
+Proof. by rewrite -(mul0rz _ n) ler_nmulz2r // mul0rz. Qed.
+
+Lemma pmulrz_lle0 x n (n0 : 0 < n) : x *~ n <= 0 = (x <= 0).
+Proof. by rewrite -(mul0rz _ n) ler_pmulz2r // mul0rz. Qed.
+
+Lemma nmulrz_lle0 x n (n0 : n < 0) : x *~ n <= 0 = (0 <= x).
+Proof. by rewrite -(mul0rz _ n) ler_nmulz2r // mul0rz. Qed.
+
+Lemma ler_wpmulz2l x (hx : 0 <= x) : {homo *~%R x : x y / x <= y}.
+Proof.
+by move=> m n /= hmn; rewrite -subr_ge0 -mulrzBr mulrz_ge0 // subr_ge0.
+Qed.
+
+Lemma ler_wnmulz2l x (hx : x <= 0) : {homo *~%R x : x y /~ x <= y}.
+Proof.
+by move=> m n /= hmn; rewrite -subr_ge0 -mulrzBr mulrz_le0 // subr_le0.
+Qed.
+
+Lemma ler_pmulz2l x (hx : 0 < x) : {mono *~%R x : x y / x <= y}.
+Proof.
+move=> m n /=; rewrite real_mono ?num_real // => {m n}.
+by move=> m n /= hmn; rewrite -subr_gt0 -mulrzBr pmulrz_lgt0 // subr_gt0.
+Qed.
+
+Lemma ler_nmulz2l x (hx : x < 0) : {mono *~%R x : x y /~ x <= y}.
+Proof.
+move=> m n /=; rewrite real_nmono ?num_real // => {m n}.
+by move=> m n /= hmn; rewrite -subr_gt0 -mulrzBr nmulrz_lgt0 // subr_lt0.
+Qed.
+
+Lemma ltr_pmulz2l x (hx : 0 < x) : {mono *~%R x : x y / x < y}.
+Proof. exact: lerW_mono (ler_pmulz2l _). Qed.
+
+Lemma ltr_nmulz2l x (hx : x < 0) : {mono *~%R x : x y /~ x < y}.
+Proof. exact: lerW_nmono (ler_nmulz2l _). Qed.
+
+Lemma pmulrz_rgt0 x n (x0 : 0 < x) : 0 < x *~ n = (0 < n).
+Proof. by rewrite -(mulr0z x) ltr_pmulz2l. Qed.
+
+Lemma nmulrz_rgt0 x n (x0 : x < 0) : 0 < x *~ n = (n < 0).
+Proof. by rewrite -(mulr0z x) ltr_nmulz2l. Qed.
+
+Lemma pmulrz_rlt0 x n (x0 : 0 < x) : x *~ n < 0 = (n < 0).
+Proof. by rewrite -(mulr0z x) ltr_pmulz2l. Qed.
+
+Lemma nmulrz_rlt0 x n (x0 : x < 0) : x *~ n < 0 = (0 < n).
+Proof. by rewrite -(mulr0z x) ltr_nmulz2l. Qed.
+
+Lemma pmulrz_rge0 x n (x0 : 0 < x) : 0 <= x *~ n = (0 <= n).
+Proof. by rewrite -(mulr0z x) ler_pmulz2l. Qed.
+
+Lemma nmulrz_rge0 x n (x0 : x < 0) : 0 <= x *~ n = (n <= 0).
+Proof. by rewrite -(mulr0z x) ler_nmulz2l. Qed.
+
+Lemma pmulrz_rle0 x n (x0 : 0 < x) : x *~ n <= 0 = (n <= 0).
+Proof. by rewrite -(mulr0z x) ler_pmulz2l. Qed.
+
+Lemma nmulrz_rle0 x n (x0 : x < 0) : x *~ n <= 0 = (0 <= n).
+Proof. by rewrite -(mulr0z x) ler_nmulz2l. Qed.
+
+Lemma mulrIz x (hx : x != 0) : injective ( *~%R x).
+Proof.
+move=> y z; rewrite -![x *~ _]mulrzr => /(mulfI hx).
+by apply: mono_inj y z; apply: ler_pmulz2l.
+Qed.
+
+Lemma ler_int m n : (m%:~R <= n%:~R :> R) = (m <= n).
+Proof. by rewrite ler_pmulz2l. Qed.
+
+Lemma ltr_int m n : (m%:~R < n%:~R :> R) = (m < n).
+Proof. by rewrite ltr_pmulz2l. Qed.
+
+Lemma eqr_int m n : (m%:~R == n%:~R :> R) = (m == n).
+Proof. by rewrite (inj_eq (mulrIz _)) ?oner_eq0. Qed.
+
+Lemma ler0z n : (0 <= n%:~R :> R) = (0 <= n).
+Proof. by rewrite pmulrz_rge0. Qed.
+
+Lemma ltr0z n : (0 < n%:~R :> R) = (0 < n).
+Proof. by rewrite pmulrz_rgt0. Qed.
+
+Lemma lerz0 n : (n%:~R <= 0 :> R) = (n <= 0).
+Proof. by rewrite pmulrz_rle0. Qed.
+
+Lemma ltrz0 n : (n%:~R < 0 :> R) = (n < 0).
+Proof. by rewrite pmulrz_rlt0. Qed.
+
+Lemma ler1z (n : int) : (1 <= n%:~R :> R) = (1 <= n).
+Proof. by rewrite -[1]/(1%:~R) ler_int. Qed.
+
+Lemma ltr1z (n : int) : (1 < n%:~R :> R) = (1 < n).
+Proof. by rewrite -[1]/(1%:~R) ltr_int. Qed.
+
+Lemma lerz1 n : (n%:~R <= 1 :> R) = (n <= 1).
+Proof. by rewrite -[1]/(1%:~R) ler_int. Qed.
+
+Lemma ltrz1 n : (n%:~R < 1 :> R) = (n < 1).
+Proof. by rewrite -[1]/(1%:~R) ltr_int. Qed.
+
+Lemma intr_eq0 n : (n%:~R == 0 :> R) = (n == 0).
+Proof. by rewrite -(mulr0z 1) (inj_eq (mulrIz _)) // oner_eq0. Qed.
+
+Lemma mulrz_eq0 x n : (x *~ n == 0) = ((n == 0) || (x == 0)).
+Proof. by rewrite -mulrzl mulf_eq0 intr_eq0. Qed.
+
+Lemma mulrz_neq0 x n : x *~ n != 0 = ((n != 0) && (x != 0)).
+Proof. by rewrite mulrz_eq0 negb_or. Qed.
+
+Lemma realz n : (n%:~R : R) \in Num.real.
+Proof. by rewrite -topredE /Num.real /= ler0z lerz0 ler_total. Qed.
+Hint Resolve realz.
+
+Definition intr_inj := @mulrIz 1 (oner_neq0 R).
+
+End PO.
+
+End NumMorphism.
+
+End MorphTheory.
+
+Implicit Arguments intr_inj [[R] x1 x2].
+
+Definition exprz (R : unitRingType) (x : R) (n : int) := nosimpl
+ match n with
+ | Posz n => x ^+ n
+ | Negz n => x ^- (n.+1)
+ end.
+
+Notation "x ^ n" := (exprz x n) : ring_scope.
+
+Section ExprzUnitRing.
+
+Variable R : unitRingType.
+Implicit Types x y : R.
+Implicit Types m n : int.
+Local Coercion Posz : nat >-> int.
+
+Lemma exprnP x (n : nat) : x ^+ n = x ^ n. Proof. by []. Qed.
+
+Lemma exprnN x (n : nat) : x ^- n = x ^ -n%:Z.
+Proof. by case: n=> //; rewrite oppr0 expr0 invr1. Qed.
+
+Lemma expr0z x : x ^ 0 = 1. Proof. by []. Qed.
+
+Lemma expr1z x : x ^ 1 = x. Proof. by []. Qed.
+
+Lemma exprN1 x : x ^ (-1) = x^-1. Proof. by []. Qed.
+
+Lemma invr_expz x n : (x ^ n)^-1 = x ^ (- n).
+Proof.
+by case: (intP n)=> // [|m]; rewrite ?opprK ?expr0z ?invr1 // invrK.
+Qed.
+
+Lemma exprz_inv x n : (x^-1) ^ n = x ^ (- n).
+Proof.
+by case: (intP n)=> // m; rewrite -[_ ^ (- _)]exprVn ?opprK ?invrK.
+Qed.
+
+Lemma exp1rz n : 1 ^ n = 1 :> R.
+Proof.
+by case: (intP n)=> // m; rewrite -?exprz_inv ?invr1; apply: expr1n.
+Qed.
+
+Lemma exprSz x (n : nat) : x ^ n.+1 = x * x ^ n. Proof. exact: exprS. Qed.
+
+Lemma exprSzr x (n : nat) : x ^ n.+1 = x ^ n * x.
+Proof. exact: exprSr. Qed.
+
+Fact exprzD_nat x (m n : nat) : x ^ (m%:Z + n) = x ^ m * x ^ n.
+Proof. exact: exprD. Qed.
+
+Fact exprzD_Nnat x (m n : nat) : x ^ (-m%:Z + -n%:Z) = x ^ (-m%:Z) * x ^ (-n%:Z).
+Proof. by rewrite -opprD -!exprz_inv exprzD_nat. Qed.
+
+Lemma exprzD_ss x m n : (0 <= m) && (0 <= n) || (m <= 0) && (n <= 0)
+ -> x ^ (m + n) = x ^ m * x ^ n.
+Proof.
+case: (intP m)=> {m} [|m|m]; case: (intP n)=> {n} [|n|n] //= _;
+by rewrite ?expr0z ?mul1r ?exprzD_nat ?exprzD_Nnat ?sub0r ?addr0 ?mulr1.
+Qed.
+
+Lemma exp0rz n : 0 ^ n = (n == 0)%:~R :> R.
+Proof. by case: (intP n)=> // m; rewrite -?exprz_inv ?invr0 exprSz mul0r. Qed.
+
+Lemma commrXz x y n : GRing.comm x y -> GRing.comm x (y ^ n).
+Proof.
+rewrite /GRing.comm; elim: n x y=> [|n ihn|n ihn] x y com_xy //=.
+* by rewrite expr0z mul1r mulr1.
+* by rewrite -exprnP commrX //.
+rewrite -exprz_inv -exprnP commrX //.
+case: (boolP (y \is a GRing.unit))=> uy; last by rewrite invr_out.
+by apply/eqP; rewrite (can2_eq (mulrVK _) (mulrK _)) // -mulrA com_xy mulKr.
+Qed.
+
+Lemma exprMz_comm x y n : x \is a GRing.unit -> y \is a GRing.unit ->
+ GRing.comm x y -> (x * y) ^ n = x ^ n * y ^ n.
+Proof.
+move=> ux uy com_xy; elim: n => [|n _|n _]; first by rewrite expr0z mulr1.
+ by rewrite -!exprnP exprMn_comm.
+rewrite -!exprnN -!exprVn com_xy -exprMn_comm ?invrM//.
+exact/commrV/commr_sym/commrV.
+Qed.
+
+Lemma commrXz_wmulls x y n :
+ 0 <= n -> GRing.comm x y -> (x * y) ^ n = x ^ n * y ^ n.
+Proof.
+move=> n0 com_xy; elim: n n0 => [|n _|n _] //; first by rewrite expr0z mulr1.
+by rewrite -!exprnP exprMn_comm.
+Qed.
+
+Lemma unitrXz x n (ux : x \is a GRing.unit) : x ^ n \is a GRing.unit.
+Proof.
+case: (intP n)=> {n} [|n|n]; rewrite ?expr0z ?unitr1 ?unitrX //.
+by rewrite -invr_expz unitrV unitrX.
+Qed.
+
+Lemma exprzDr x (ux : x \is a GRing.unit) m n : x ^ (m + n) = x ^ m * x ^ n.
+Proof.
+move: n m; apply: wlog_ler=> n m hnm.
+ by rewrite addrC hnm commrXz //; apply: commr_sym; apply: commrXz.
+case: (intP m) hnm=> {m} [|m|m]; rewrite ?mul1r ?add0r //;
+ case: (intP n)=> {n} [|n|n _]; rewrite ?mulr1 ?addr0 //;
+ do ?by rewrite exprzD_ss.
+rewrite -invr_expz subzSS !exprSzr invrM ?unitrX // -mulrA mulVKr //.
+case: (leqP n m)=> [|/ltnW] hmn; rewrite -{2}(subnK hmn) exprzD_nat -subzn //.
+ by rewrite mulrK ?unitrX.
+by rewrite invrM ?unitrXz // mulVKr ?unitrXz // -opprB -invr_expz.
+Qed.
+
+Lemma exprz_exp x m n : (x ^ m) ^ n = (x ^ (m * n)).
+Proof.
+wlog: n / 0 <= n.
+ by case: n=> [n -> //|n]; rewrite ?NegzE mulrN -?invr_expz=> -> /=.
+elim: n x m=> [|n ihn|n ihn] x m // _; first by rewrite mulr0 !expr0z.
+rewrite exprSz ihn // intS mulrDr mulr1 exprzD_ss //.
+by case: (intP m)=> // m'; rewrite ?oppr_le0 //.
+Qed.
+
+Lemma exprzAC x m n : (x ^ m) ^ n = (x ^ n) ^ m.
+Proof. by rewrite !exprz_exp mulrC. Qed.
+
+Lemma exprz_out x n (nux : x \isn't a GRing.unit) (hn : 0 <= n) :
+ x ^ (- n) = x ^ n.
+Proof. by case: (intP n) hn=> //= m; rewrite -exprnN -exprVn invr_out. Qed.
+
+End ExprzUnitRing.
+
+Section Exprz_Zint_UnitRing.
+
+Variable R : unitRingType.
+Implicit Types x y : R.
+Implicit Types m n : int.
+Local Coercion Posz : nat >-> int.
+
+Lemma exprz_pmulzl x m n : 0 <= n -> (x *~ m) ^ n = x ^ n *~ (m ^ n).
+Proof.
+by elim: n=> [|n ihn|n _] // _; rewrite !exprSz ihn // mulrzAr mulrzAl -mulrzA.
+Qed.
+
+Lemma exprz_pintl m n (hn : 0 <= n) : m%:~R ^ n = (m ^ n)%:~R :> R.
+Proof. by rewrite exprz_pmulzl // exp1rz. Qed.
+
+Lemma exprzMzl x m n (ux : x \is a GRing.unit) (um : m%:~R \is a @GRing.unit R):
+ (x *~ m) ^ n = (m%:~R ^ n) * x ^ n :> R.
+Proof.
+rewrite -[x *~ _]mulrzl exprMz_comm //.
+by apply: commr_sym; apply: commr_int.
+Qed.
+
+Lemma expNrz x n : (- x) ^ n = (-1) ^ n * x ^ n :> R.
+Proof.
+case: n=> [] n; rewrite ?NegzE; first by apply: exprNn.
+by rewrite -!exprz_inv !invrN invr1; apply: exprNn.
+Qed.
+
+Lemma unitr_n0expz x n :
+ n != 0 -> (x ^ n \is a GRing.unit) = (x \is a GRing.unit).
+Proof.
+by case: n => *; rewrite ?NegzE -?exprz_inv ?unitrX_pos ?unitrV ?lt0n.
+Qed.
+
+Lemma intrV (n : int) :
+ n \in [:: 0; 1; -1] -> n%:~R ^-1 = n%:~R :> R.
+Proof.
+by case: (intP n)=> // [|[]|[]] //; rewrite ?rmorphN ?invrN (invr0, invr1).
+Qed.
+
+Lemma rmorphXz (R' : unitRingType) (f : {rmorphism R -> R'}) n :
+ {in GRing.unit, {morph f : x / x ^ n}}.
+Proof. by case: n => n x Ux; rewrite ?rmorphV ?rpredX ?rmorphX. Qed.
+
+End Exprz_Zint_UnitRing.
+
+Section ExprzIdomain.
+
+Variable R : idomainType.
+Implicit Types x y : R.
+Implicit Types m n : int.
+Local Coercion Posz : nat >-> int.
+
+Lemma expfz_eq0 x n : (x ^ n == 0) = (n != 0) && (x == 0).
+Proof.
+by case: n=> n; rewrite ?NegzE -?exprz_inv ?expf_eq0 ?lt0n ?invr_eq0.
+Qed.
+
+Lemma expfz_neq0 x n : x != 0 -> x ^ n != 0.
+Proof. by move=> x_nz; rewrite expfz_eq0; apply/nandP; right. Qed.
+
+Lemma exprzMl x y n (ux : x \is a GRing.unit) (uy : y \is a GRing.unit) :
+ (x * y) ^ n = x ^ n * y ^ n.
+Proof. by rewrite exprMz_comm //; apply: mulrC. Qed.
+
+Lemma expfV (x : R) (i : int) : (x ^ i) ^-1 = (x ^-1) ^ i.
+Proof. by rewrite invr_expz exprz_inv. Qed.
+
+End ExprzIdomain.
+
+Section ExprzField.
+
+Variable F : fieldType.
+Implicit Types x y : F.
+Implicit Types m n : int.
+Local Coercion Posz : nat >-> int.
+
+Lemma expfzDr x m n : x != 0 -> x ^ (m + n) = x ^ m * x ^ n.
+Proof. by move=> hx; rewrite exprzDr ?unitfE. Qed.
+
+Lemma expfz_n0addr x m n : m + n != 0 -> x ^ (m + n) = x ^ m * x ^ n.
+Proof.
+have [-> hmn|nx0 _] := eqVneq x 0; last exact: expfzDr.
+rewrite !exp0rz (negPf hmn).
+case: (altP (m =P 0)) hmn=> [->|]; rewrite (mul0r, mul1r) //.
+by rewrite add0r=> /negPf->.
+Qed.
+
+Lemma expfzMl x y n : (x * y) ^ n = x ^ n * y ^ n.
+Proof.
+have [->|/negPf n0] := eqVneq n 0; first by rewrite !expr0z mulr1.
+case: (boolP ((x * y) == 0)); rewrite ?mulf_eq0.
+ by case/orP=> /eqP->; rewrite ?(mul0r, mulr0, exp0rz, n0).
+by case/norP=> x0 y0; rewrite exprzMl ?unitfE.
+Qed.
+
+Lemma fmorphXz (R : unitRingType) (f : {rmorphism F -> R}) n :
+ {morph f : x / x ^ n}.
+Proof. by case: n => n x; rewrite ?fmorphV rmorphX. Qed.
+
+End ExprzField.
+
+Section ExprzOrder.
+
+Variable R : realFieldType.
+Implicit Types x y : R.
+Implicit Types m n : int.
+Local Coercion Posz : nat >-> int.
+
+(* ler and exprz *)
+Lemma exprz_ge0 n x (hx : 0 <= x) : (0 <= x ^ n).
+Proof. by case: n=> n; rewrite ?NegzE -?invr_expz ?invr_ge0 ?exprn_ge0. Qed.
+
+Lemma exprz_gt0 n x (hx : 0 < x) : (0 < x ^ n).
+Proof. by case: n=> n; rewrite ?NegzE -?invr_expz ?invr_gt0 ?exprn_gt0. Qed.
+
+Definition exprz_gte0 := (exprz_ge0, exprz_gt0).
+
+Lemma ler_wpiexpz2l x (x0 : 0 <= x) (x1 : x <= 1) :
+ {in >= 0 &, {homo (exprz x) : x y /~ x <= y}}.
+Proof.
+move=> [] m [] n; rewrite -!topredE /= ?oppr_cp0 ?ltz_nat // => _ _.
+by rewrite lez_nat -?exprnP=> /ler_wiexpn2l; apply.
+Qed.
+
+Lemma ler_wniexpz2l x (x0 : 0 <= x) (x1 : x <= 1) :
+ {in < 0 &, {homo (exprz x) : x y /~ x <= y}}.
+Proof.
+move=> [] m [] n; rewrite ?NegzE -!topredE /= ?oppr_cp0 ?ltz_nat // => _ _.
+rewrite ler_opp2 lez_nat -?invr_expz=> hmn; move: (x0).
+rewrite le0r=> /orP [/eqP->|lx0]; first by rewrite !exp0rz invr0.
+by rewrite lef_pinv -?topredE /= ?exprz_gt0 // ler_wiexpn2l.
+Qed.
+
+Fact ler_wpeexpz2l x (x1 : 1 <= x) :
+ {in >= 0 &, {homo (exprz x) : x y / x <= y}}.
+Proof.
+move=> [] m [] n; rewrite -!topredE /= ?oppr_cp0 ?ltz_nat // => _ _.
+by rewrite lez_nat -?exprnP=> /ler_weexpn2l; apply.
+Qed.
+
+Fact ler_wneexpz2l x (x1 : 1 <= x) :
+ {in <= 0 &, {homo (exprz x) : x y / x <= y}}.
+Proof.
+move=> m n hm hn /= hmn.
+rewrite -lef_pinv -?topredE /= ?exprz_gt0 ?(ltr_le_trans ltr01) //.
+by rewrite !invr_expz ler_wpeexpz2l ?ler_opp2 -?topredE //= oppr_cp0.
+Qed.
+
+Lemma ler_weexpz2l x (x1 : 1 <= x) : {homo (exprz x) : x y / x <= y}.
+Proof.
+move=> m n /= hmn; case: (lerP 0 m)=> [|/ltrW] hm.
+ by rewrite ler_wpeexpz2l // [_ \in _](ler_trans hm).
+case: (lerP n 0)=> [|/ltrW] hn.
+ by rewrite ler_wneexpz2l // [_ \in _](ler_trans hmn).
+apply: (@ler_trans _ (x ^ 0)); first by rewrite ler_wneexpz2l.
+by rewrite ler_wpeexpz2l.
+Qed.
+
+Lemma pexprz_eq1 x n (x0 : 0 <= x) : (x ^ n == 1) = ((n == 0) || (x == 1)).
+Proof.
+case: n=> n; rewrite ?NegzE -?exprz_inv ?oppr_eq0 pexprn_eq1 // ?invr_eq1 //.
+by rewrite invr_ge0.
+Qed.
+
+Lemma ieexprIz x (x0 : 0 < x) (nx1 : x != 1) : injective (exprz x).
+Proof.
+apply: wlog_ltr=> // m n hmn; first by move=> hmn'; rewrite hmn.
+move=> /(f_equal ( *%R^~ (x ^ (- n)))).
+rewrite -!expfzDr ?gtr_eqF // subrr expr0z=> /eqP.
+by rewrite pexprz_eq1 ?(ltrW x0) // (negPf nx1) subr_eq0 orbF=> /eqP.
+Qed.
+
+Lemma ler_piexpz2l x (x0 : 0 < x) (x1 : x < 1) :
+ {in >= 0 &, {mono (exprz x) : x y /~ x <= y}}.
+Proof.
+apply: (nhomo_mono_in (nhomo_inj_in_lt _ _)).
+ by move=> n m hn hm /=; apply: ieexprIz; rewrite // ltr_eqF.
+by apply: ler_wpiexpz2l; rewrite ?ltrW.
+Qed.
+
+Lemma ltr_piexpz2l x (x0 : 0 < x) (x1 : x < 1) :
+ {in >= 0 &, {mono (exprz x) : x y /~ x < y}}.
+Proof. exact: (lerW_nmono_in (ler_piexpz2l _ _)). Qed.
+
+Lemma ler_niexpz2l x (x0 : 0 < x) (x1 : x < 1) :
+ {in < 0 &, {mono (exprz x) : x y /~ x <= y}}.
+Proof.
+apply: (nhomo_mono_in (nhomo_inj_in_lt _ _)).
+ by move=> n m hn hm /=; apply: ieexprIz; rewrite // ltr_eqF.
+by apply: ler_wniexpz2l; rewrite ?ltrW.
+Qed.
+
+Lemma ltr_niexpz2l x (x0 : 0 < x) (x1 : x < 1) :
+ {in < 0 &, {mono (exprz x) : x y /~ x < y}}.
+Proof. exact: (lerW_nmono_in (ler_niexpz2l _ _)). Qed.
+
+Lemma ler_eexpz2l x (x1 : 1 < x) : {mono (exprz x) : x y / x <= y}.
+Proof.
+apply: (homo_mono (homo_inj_lt _ _)).
+ by apply: ieexprIz; rewrite ?(ltr_trans ltr01) // gtr_eqF.
+by apply: ler_weexpz2l; rewrite ?ltrW.
+Qed.
+
+Lemma ltr_eexpz2l x (x1 : 1 < x) : {mono (exprz x) : x y / x < y}.
+Proof. exact: (lerW_mono (ler_eexpz2l _)). Qed.
+
+Lemma ler_wpexpz2r n (hn : 0 <= n) :
+{in >= 0 & , {homo ((@exprz R)^~ n) : x y / x <= y}}.
+Proof. by case: n hn=> // n _; apply: ler_expn2r. Qed.
+
+Lemma ler_wnexpz2r n (hn : n <= 0) :
+{in > 0 & , {homo ((@exprz R)^~ n) : x y /~ x <= y}}.
+Proof.
+move=> x y /= hx hy hxy; rewrite -lef_pinv ?[_ \in _]exprz_gt0 //.
+by rewrite !invr_expz ler_wpexpz2r ?[_ \in _]ltrW // oppr_cp0.
+Qed.
+
+Lemma pexpIrz n (n0 : n != 0) : {in >= 0 &, injective ((@exprz R)^~ n)}.
+Proof.
+move=> x y; rewrite ![_ \in _]le0r=> /orP [/eqP-> _ /eqP|hx].
+ by rewrite exp0rz ?(negPf n0) eq_sym expfz_eq0=> /andP [_ /eqP->].
+case/orP=> [/eqP-> /eqP|hy].
+ by rewrite exp0rz ?(negPf n0) expfz_eq0=> /andP [_ /eqP].
+move=> /(f_equal ( *%R^~ (y ^ (- n)))) /eqP.
+rewrite -expfzDr ?(gtr_eqF hy) // subrr expr0z -exprz_inv -expfzMl.
+rewrite pexprz_eq1 ?(negPf n0) /= ?mulr_ge0 ?invr_ge0 ?ltrW //.
+by rewrite (can2_eq (mulrVK _) (mulrK _)) ?unitfE ?(gtr_eqF hy) // mul1r=> /eqP.
+Qed.
+
+Lemma nexpIrz n (n0 : n != 0) : {in <= 0 &, injective ((@exprz R)^~ n)}.
+Proof.
+move=> x y; rewrite ![_ \in _]ler_eqVlt => /orP [/eqP -> _ /eqP|hx].
+ by rewrite exp0rz ?(negPf n0) eq_sym expfz_eq0=> /andP [_ /eqP->].
+case/orP=> [/eqP -> /eqP|hy].
+ by rewrite exp0rz ?(negPf n0) expfz_eq0=> /andP [_ /eqP].
+move=> /(f_equal ( *%R^~ (y ^ (- n)))) /eqP.
+rewrite -expfzDr ?(ltr_eqF hy) // subrr expr0z -exprz_inv -expfzMl.
+rewrite pexprz_eq1 ?(negPf n0) /= ?mulr_le0 ?invr_le0 ?ltrW //.
+by rewrite (can2_eq (mulrVK _) (mulrK _)) ?unitfE ?(ltr_eqF hy) // mul1r=> /eqP.
+Qed.
+
+Lemma ler_pexpz2r n (hn : 0 < n) :
+ {in >= 0 & , {mono ((@exprz R)^~ n) : x y / x <= y}}.
+Proof.
+apply: homo_mono_in (homo_inj_in_lt _ _).
+ by move=> x y hx hy /=; apply: pexpIrz; rewrite // gtr_eqF.
+by apply: ler_wpexpz2r; rewrite ltrW.
+Qed.
+
+Lemma ltr_pexpz2r n (hn : 0 < n) :
+ {in >= 0 & , {mono ((@exprz R)^~ n) : x y / x < y}}.
+Proof. exact: lerW_mono_in (ler_pexpz2r _). Qed.
+
+Lemma ler_nexpz2r n (hn : n < 0) :
+ {in > 0 & , {mono ((@exprz R)^~ n) : x y /~ x <= y}}.
+Proof.
+apply: nhomo_mono_in (nhomo_inj_in_lt _ _); last first.
+ by apply: ler_wnexpz2r; rewrite ltrW.
+by move=> x y hx hy /=; apply: pexpIrz; rewrite ?[_ \in _]ltrW ?ltr_eqF.
+Qed.
+
+Lemma ltr_nexpz2r n (hn : n < 0) :
+ {in > 0 & , {mono ((@exprz R)^~ n) : x y /~ x < y}}.
+Proof. exact: lerW_nmono_in (ler_nexpz2r _). Qed.
+
+Lemma eqr_expz2 n x y : n != 0 -> 0 <= x -> 0 <= y ->
+ (x ^ n == y ^ n) = (x == y).
+Proof. by move=> *; rewrite (inj_in_eq (pexpIrz _)). Qed.
+
+End ExprzOrder.
+
+Local Notation sgr := Num.sg.
+
+Section Sgz.
+
+Variable R : numDomainType.
+Implicit Types x y z : R.
+Implicit Types m n p : int.
+Local Coercion Posz : nat >-> int.
+
+Definition sgz x : int := if x == 0 then 0 else if x < 0 then -1 else 1.
+
+Lemma sgz_def x : sgz x = (-1) ^+ (x < 0)%R *+ (x != 0).
+Proof. by rewrite /sgz; case: (_ == _); case: (_ < _). Qed.
+
+Lemma sgrEz x : sgr x = (sgz x)%:~R. Proof. by rewrite !(fun_if intr). Qed.
+
+Lemma gtr0_sgz x : 0 < x -> sgz x = 1.
+Proof. by move=> x_gt0; rewrite /sgz ltr_neqAle andbC eqr_le ltr_geF //. Qed.
+
+Lemma ltr0_sgz x : x < 0 -> sgz x = -1.
+Proof. by move=> x_lt0; rewrite /sgz eq_sym eqr_le x_lt0 ltr_geF. Qed.
+
+Lemma sgz0 : sgz (0 : R) = 0. Proof. by rewrite /sgz eqxx. Qed.
+Lemma sgz1 : sgz (1 : R) = 1. Proof. by rewrite gtr0_sgz // ltr01. Qed.
+Lemma sgzN1 : sgz (-1 : R) = -1. Proof. by rewrite ltr0_sgz // ltrN10. Qed.
+
+Definition sgzE := (sgz0, sgz1, sgzN1).
+
+Lemma sgz_sgr x : sgz (sgr x) = sgz x.
+Proof. by rewrite !(fun_if sgz) !sgzE. Qed.
+
+Lemma normr_sgz x : `|sgz x| = (x != 0).
+Proof. by rewrite sgz_def -mulr_natr normrMsign normr_nat natz. Qed.
+
+Lemma normr_sg x : `|sgr x| = (x != 0)%:~R.
+Proof. by rewrite sgr_def -mulr_natr normrMsign normr_nat. Qed.
+
+End Sgz.
+
+Section MoreSgz.
+
+Variable R : numDomainType.
+
+Lemma sgz_int m : sgz (m%:~R : R) = sgz m.
+Proof. by rewrite /sgz intr_eq0 ltrz0. Qed.
+
+Lemma sgrz (n : int) : sgr n = sgz n. Proof. by rewrite sgrEz intz. Qed.
+
+Lemma intr_sg m : (sgr m)%:~R = sgr (m%:~R) :> R.
+Proof. by rewrite sgrz -sgz_int -sgrEz. Qed.
+
+Lemma sgz_id (x : R) : sgz (sgz x) = sgz x.
+Proof. by rewrite !(fun_if (@sgz _)). Qed.
+
+End MoreSgz.
+
+Section SgzReal.
+
+Variable R : realDomainType.
+Implicit Types x y z : R.
+Implicit Types m n p : int.
+Local Coercion Posz : nat >-> int.
+
+Lemma sgz_cp0 x :
+ ((sgz x == 1) = (0 < x)) *
+ ((sgz x == -1) = (x < 0)) *
+ ((sgz x == 0) = (x == 0)).
+Proof. by rewrite /sgz; case: ltrgtP. Qed.
+
+CoInductive sgz_val x : bool -> bool -> bool -> bool -> bool -> bool
+ -> bool -> bool -> bool -> bool -> bool -> bool
+ -> bool -> bool -> bool -> bool -> bool -> bool
+ -> R -> R -> int -> Set :=
+ | SgzNull of x = 0 : sgz_val x true true true true false false
+ true false false true false false true false false true false false 0 0 0
+ | SgzPos of x > 0 : sgz_val x false false true false false true
+ false false true false false true false false true false false true x 1 1
+ | SgzNeg of x < 0 : sgz_val x false true false false true false
+ false true false false true false false true false false true false (-x) (-1) (-1).
+
+Lemma sgzP x :
+ sgz_val x (0 == x) (x <= 0) (0 <= x) (x == 0) (x < 0) (0 < x)
+ (0 == sgr x) (-1 == sgr x) (1 == sgr x)
+ (sgr x == 0) (sgr x == -1) (sgr x == 1)
+ (0 == sgz x) (-1 == sgz x) (1 == sgz x)
+ (sgz x == 0) (sgz x == -1) (sgz x == 1) `|x| (sgr x) (sgz x).
+Proof.
+rewrite ![_ == sgz _]eq_sym ![_ == sgr _]eq_sym !sgr_cp0 !sgz_cp0.
+by rewrite /sgr /sgz !lerNgt; case: ltrgt0P; constructor.
+Qed.
+
+Lemma sgzN x : sgz (- x) = - sgz x.
+Proof. by rewrite /sgz oppr_eq0 oppr_lt0; case: ltrgtP. Qed.
+
+Lemma mulz_sg x : sgz x * sgz x = (x != 0)%:~R.
+Proof. by case: sgzP; rewrite ?(mulr0, mulr1, mulrNN). Qed.
+
+Lemma mulz_sg_eq1 x y : (sgz x * sgz y == 1) = (x != 0) && (sgz x == sgz y).
+Proof.
+do 2?case: sgzP=> _; rewrite ?(mulr0, mulr1, mulrN1, opprK, oppr0, eqxx);
+ by rewrite ?[0 == 1]eq_sym ?oner_eq0 //= eqr_oppLR oppr0 oner_eq0.
+Qed.
+
+Lemma mulz_sg_eqN1 x y : (sgz x * sgz y == -1) = (x != 0) && (sgz x == - sgz y).
+Proof. by rewrite -eqr_oppLR -mulrN -sgzN mulz_sg_eq1. Qed.
+
+(* Lemma muls_eqA x y z : sgr x != 0 -> *)
+(* (sgr y * sgr z == sgr x) = ((sgr y * sgr x == sgr z) && (sgr z != 0)). *)
+(* Proof. by do 3!case: sgrP=> _. Qed. *)
+
+Lemma sgzM x y : sgz (x * y) = sgz x * sgz y.
+Proof.
+case: (sgzP x)=> hx; first by rewrite hx ?mul0r sgz0.
+ case: (sgzP y)=> hy; first by rewrite hy !mulr0 sgz0.
+ by apply/eqP; rewrite mul1r sgz_cp0 pmulr_rgt0.
+ by apply/eqP; rewrite mul1r sgz_cp0 nmulr_llt0.
+case: (sgzP y)=> hy; first by rewrite hy !mulr0 sgz0.
+ by apply/eqP; rewrite mulr1 sgz_cp0 nmulr_rlt0.
+by apply/eqP; rewrite mulN1r opprK sgz_cp0 nmulr_rgt0.
+Qed.
+
+Lemma sgzX (n : nat) x : sgz (x ^+ n) = (sgz x) ^+ n.
+Proof. by elim: n => [|n IHn]; rewrite ?sgz1 // !exprS sgzM IHn. Qed.
+
+Lemma sgz_eq0 x : (sgz x == 0) = (x == 0).
+Proof. by rewrite sgz_cp0. Qed.
+
+Lemma sgz_odd (n : nat) x : x != 0 -> (sgz x) ^+ n = (sgz x) ^+ (odd n).
+Proof. by case: sgzP => //=; rewrite ?expr1n // signr_odd. Qed.
+
+Lemma sgz_gt0 x : (sgz x > 0) = (x > 0).
+Proof. by case: sgzP. Qed.
+
+Lemma sgz_lt0 x : (sgz x < 0) = (x < 0).
+Proof. by case: sgzP. Qed.
+
+Lemma sgz_ge0 x : (sgz x >= 0) = (x >= 0).
+Proof. by case: sgzP. Qed.
+
+Lemma sgz_le0 x : (sgz x <= 0) = (x <= 0).
+Proof. by case: sgzP. Qed.
+
+Lemma sgz_smul x y : sgz (y *~ (sgz x)) = (sgz x) * (sgz y).
+Proof. by rewrite -mulrzl sgzM -sgrEz sgz_sgr. Qed.
+
+Lemma sgrMz m x : sgr (x *~ m) = sgr x *~ sgr m.
+Proof. by rewrite -mulrzr sgrM -intr_sg mulrzr. Qed.
+
+End SgzReal.
+
+Lemma sgz_eq (R R' : realDomainType) (x : R) (y : R') :
+ (sgz x == sgz y) = ((x == 0) == (y == 0)) && ((0 < x) == (0 < y)).
+Proof. by do 2!case: sgzP. Qed.
+
+Lemma intr_sign (R : ringType) s : ((-1) ^+ s)%:~R = (-1) ^+ s :> R.
+Proof. exact: rmorph_sign. Qed.
+
+Section Absz.
+
+Implicit Types m n p : int.
+Open Scope nat_scope.
+Local Coercion Posz : nat >-> int.
+
+Lemma absz_nat (n : nat) : `|n| = n. Proof. by []. Qed.
+
+Lemma abszE (m : int) : `|m| = `|m|%R :> int. Proof. by []. Qed.
+
+Lemma absz0 : `|0%R| = 0. Proof. by []. Qed.
+
+Lemma abszN m : `|- m| = `|m|. Proof. by case: (normrN m). Qed.
+
+Lemma absz_eq0 m : (`|m| == 0) = (m == 0%R). Proof. by case: (intP m). Qed.
+
+Lemma absz_gt0 m : (`|m| > 0) = (m != 0%R). Proof. by case: (intP m). Qed.
+
+Lemma absz1 : `|1%R| = 1. Proof. by []. Qed.
+
+Lemma abszN1 : `|-1%R| = 1. Proof. by []. Qed.
+
+Lemma absz_id m : `|(`|m|)| = `|m|. Proof. by []. Qed.
+
+Lemma abszM m1 m2 : `|(m1 * m2)%R| = `|m1| * `|m2|.
+Proof. by case: m1 m2 => [[|m1]|m1] [[|m2]|m2]; rewrite //= mulnS mulnC. Qed.
+
+Lemma abszX (n : nat) m : `|m ^+ n| = `|m| ^ n.
+Proof. by elim: n => // n ihn; rewrite exprS expnS abszM ihn. Qed.
+
+Lemma absz_sg m : `|sgr m| = (m != 0%R). Proof. by case: (intP m). Qed.
+
+Lemma gez0_abs m : (0 <= m)%R -> `|m| = m :> int.
+Proof. by case: (intP m). Qed.
+
+Lemma gtz0_abs m : (0 < m)%R -> `|m| = m :> int.
+Proof. by case: (intP m). Qed.
+
+Lemma lez0_abs m : (m <= 0)%R -> `|m| = - m :> int.
+Proof. by case: (intP m). Qed.
+
+Lemma ltz0_abs m : (m < 0)%R -> `|m| = - m :> int.
+Proof. by case: (intP m). Qed.
+
+Lemma absz_sign s : `|(-1) ^+ s| = 1.
+Proof. by rewrite abszX exp1n. Qed.
+
+Lemma abszMsign s m : `|((-1) ^+ s * m)%R| = `|m|.
+Proof. by rewrite abszM absz_sign mul1n. Qed.
+
+Lemma mulz_sign_abs m : ((-1) ^+ (m < 0)%R * `|m|%:Z)%R = m.
+Proof. by rewrite abszE mulr_sign_norm. Qed.
+
+Lemma mulz_Nsign_abs m : ((-1) ^+ (0 < m)%R * `|m|%:Z)%R = - m.
+Proof. by rewrite abszE mulr_Nsign_norm. Qed.
+
+Lemma intEsign m : m = ((-1) ^+ (m < 0)%R * `|m|%:Z)%R.
+Proof. exact: numEsign. Qed.
+
+Lemma abszEsign m : `|m|%:Z = ((-1) ^+ (m < 0)%R * m)%R.
+Proof. exact: normrEsign. Qed.
+
+Lemma intEsg m : m = (sgz m * `|m|%:Z)%R.
+Proof. by rewrite -sgrz -numEsg. Qed.
+
+Lemma abszEsg m : (`|m|%:Z = sgz m * m)%R.
+Proof. by rewrite -sgrz -normrEsg. Qed.
+
+End Absz.
+
+Module Export IntDist.
+
+Notation "m - n" :=
+ (@GRing.add int_ZmodType m%N (@GRing.opp int_ZmodType n%N)) : distn_scope.
+Arguments Scope absz [distn_scope].
+Notation "`| m |" := (absz m) : nat_scope.
+Coercion Posz : nat >-> int.
+
+Section Distn.
+
+Open Scope nat_scope.
+Implicit Type m : int.
+Implicit Types n d : nat.
+
+Lemma distnC m1 m2 : `|m1 - m2| = `|m2 - m1|.
+Proof. by rewrite -opprB abszN. Qed.
+
+Lemma distnDl d n1 n2 : `|d + n1 - (d + n2)| = `|n1 - n2|.
+Proof. by rewrite !PoszD opprD addrCA -addrA addKr. Qed.
+
+Lemma distnDr d n1 n2 : `|n1 + d - (n2 + d)| = `|n1 - n2|.
+Proof. by rewrite -!(addnC d) distnDl. Qed.
+
+Lemma distnEr n1 n2 : n1 <= n2 -> `|n1 - n2| = n2 - n1.
+Proof. by move/subnK=> {1}<-; rewrite distnC PoszD addrK absz_nat. Qed.
+
+Lemma distnEl n1 n2 : n2 <= n1 -> `|n1 - n2| = n1 - n2.
+Proof. by move/distnEr <-; rewrite distnC. Qed.
+
+Lemma distn0 n : `|n - 0| = n.
+Proof. by rewrite subr0 absz_nat. Qed.
+
+Lemma dist0n n : `|0 - n| = n.
+Proof. by rewrite distnC distn0. Qed.
+
+Lemma distnn m : `|m - m| = 0.
+Proof. by rewrite subrr. Qed.
+
+Lemma distn_eq0 n1 n2 : (`|n1 - n2| == 0) = (n1 == n2).
+Proof. by rewrite absz_eq0 subr_eq0. Qed.
+
+Lemma distnS n : `|n - n.+1| = 1.
+Proof. exact: distnDr n 0 1. Qed.
+
+Lemma distSn n : `|n.+1 - n| = 1.
+Proof. exact: distnDr n 1 0. Qed.
+
+Lemma distn_eq1 n1 n2 :
+ (`|n1 - n2| == 1) = (if n1 < n2 then n1.+1 == n2 else n1 == n2.+1).
+Proof.
+case: ltnP => [lt_n12 | le_n21].
+ by rewrite eq_sym -(eqn_add2r n1) distnEr ?subnK // ltnW.
+by rewrite -(eqn_add2r n2) distnEl ?subnK.
+Qed.
+
+Lemma leq_add_dist m1 m2 m3 : `|m1 - m3| <= `|m1 - m2| + `|m2 - m3|.
+Proof. by rewrite -lez_nat PoszD !abszE ler_dist_add. Qed.
+
+(* Most of this proof generalizes to all real-ordered rings. *)
+Lemma leqif_add_distz m1 m2 m3 :
+ `|m1 - m3| <= `|m1 - m2| + `|m2 - m3|
+ ?= iff (m1 <= m2 <= m3)%R || (m3 <= m2 <= m1)%R.
+Proof.
+apply/leqifP; rewrite -ltz_nat -eqz_nat PoszD !abszE; apply/lerifP.
+wlog le_m31 : m1 m3 / (m3 <= m1)%R.
+ move=> IH; case/orP: (ler_total m1 m3) => /IH //.
+ by rewrite (addrC `|_|)%R orbC !(distrC m1) !(distrC m3).
+rewrite ger0_norm ?subr_ge0 // orb_idl => [|/andP[le_m12 le_m23]]; last first.
+ by have /eqP->: m2 == m3; rewrite ?lerr // eqr_le le_m23 (ler_trans le_m31).
+rewrite -{1}(subrK m2 m1) -addrA -subr_ge0 andbC -subr_ge0.
+by apply: lerif_add; apply/real_lerif_norm/num_real.
+Qed.
+
+Lemma leqif_add_dist n1 n2 n3 :
+ `|n1 - n3| <= `|n1 - n2| + `|n2 - n3|
+ ?= iff (n1 <= n2 <= n3) || (n3 <= n2 <= n1).
+Proof. exact: leqif_add_distz. Qed.
+
+Lemma sqrn_dist n1 n2 : `|n1 - n2| ^ 2 + 2 * (n1 * n2) = n1 ^ 2 + n2 ^ 2.
+Proof.
+wlog le_n21: n1 n2 / n2 <= n1.
+ move=> IH; case/orP: (leq_total n2 n1) => /IH //.
+ by rewrite (addnC (n2 ^ 2)) (mulnC n2) distnC.
+by rewrite distnEl ?sqrn_sub ?subnK ?nat_Cauchy.
+Qed.
+
+End Distn.
+
+End IntDist.
+
+Section NormInt.
+
+Variable R : numDomainType.
+
+Lemma intr_norm m : `|m|%:~R = `|m%:~R| :> R.
+Proof. by rewrite {2}[m]intEsign rmorphMsign normrMsign abszE normr_nat. Qed.
+
+Lemma normrMz m (x : R) : `|x *~ m| = `|x| *~ `|m|.
+Proof. by rewrite -mulrzl normrM -intr_norm mulrzl. Qed.
+
+Lemma expN1r (i : int) : (-1 : R) ^ i = (-1) ^+ `|i|.
+Proof.
+case: i => n; first by rewrite exprnP absz_nat.
+by rewrite NegzE abszN absz_nat -invr_expz expfV invrN1.
+Qed.
+
+End NormInt.
+
+Section PolyZintRing.
+
+Variable R : ringType.
+Implicit Types x y z: R.
+Implicit Types m n : int.
+Implicit Types i j k : nat.
+Implicit Types p q r : {poly R}.
+
+Lemma coefMrz : forall p n i, (p *~ n)`_i = (p`_i *~ n).
+Proof. by move=> p [] n i; rewrite ?NegzE (coefMNn, coefMn). Qed.
+
+Lemma polyC_mulrz : forall n, {morph (@polyC R) : c / c *~ n}.
+Proof.
+move=> [] n c; rewrite ?NegzE -?pmulrn ?polyC_muln //.
+by rewrite polyC_opp mulrNz polyC_muln nmulrn.
+Qed.
+
+Lemma hornerMz : forall n (p : {poly R}) x, (p *~ n).[x] = p.[x] *~ n.
+Proof. by case=> *; rewrite ?NegzE ?mulNzr ?(hornerN, hornerMn). Qed.
+
+Lemma horner_int : forall n x, (n%:~R : {poly R}).[x] = n%:~R.
+Proof. by move=> n x; rewrite hornerMz hornerC. Qed.
+
+Lemma derivMz : forall n p, (p *~ n)^`() = p^`() *~ n.
+Proof. by move=> [] n p; rewrite ?NegzE -?pmulrn (derivMn, derivMNn). Qed.
+
+End PolyZintRing.
+
+Section PolyZintOIdom.
+
+Variable R : realDomainType.
+
+Lemma mulpz (p : {poly R}) (n : int) : p *~ n = n%:~R *: p.
+Proof. by rewrite -[p *~ n]mulrzl -mul_polyC polyC_mulrz polyC1. Qed.
+
+End PolyZintOIdom.
+
+Section ZnatPred.
+
+Definition Znat := [qualify a n : int | 0 <= n].
+Fact Znat_key : pred_key Znat. by []. Qed.
+Canonical Znat_keyd := KeyedQualifier Znat_key.
+
+Lemma Znat_def n : (n \is a Znat) = (0 <= n). Proof. by []. Qed.
+
+Lemma Znat_semiring_closed : semiring_closed Znat.
+Proof. by do 2?split => //; [exact: addr_ge0 | exact: mulr_ge0]. Qed.
+Canonical Znat_addrPred := AddrPred Znat_semiring_closed.
+Canonical Znat_mulrPred := MulrPred Znat_semiring_closed.
+Canonical Znat_semiringPred := SemiringPred Znat_semiring_closed.
+
+Lemma ZnatP (m : int) : reflect (exists n : nat, m = n) (m \is a Znat).
+Proof. by apply: (iffP idP) => [|[n -> //]]; case: m => // n; exists n. Qed.
+
+End ZnatPred.
+
+Section rpred.
+
+Lemma rpredMz M S (addS : @zmodPred M S) (kS : keyed_pred addS) m :
+ {in kS, forall u, u *~ m \in kS}.
+Proof. by case: m => n u Su; rewrite ?rpredN ?rpredMn. Qed.
+
+Lemma rpred_int R S (ringS : @subringPred R S) (kS : keyed_pred ringS) m :
+ m%:~R \in kS.
+Proof. by rewrite rpredMz ?rpred1. Qed.
+
+Lemma rpredZint (R : ringType) (M : lmodType R) S
+ (addS : @zmodPred M S) (kS : keyed_pred addS) m :
+ {in kS, forall u, m%:~R *: u \in kS}.
+Proof. by move=> u Su; rewrite /= scaler_int rpredMz. Qed.
+
+Lemma rpredXz R S (divS : @divrPred R S) (kS : keyed_pred divS) m :
+ {in kS, forall x, x ^ m \in kS}.
+Proof. by case: m => n x Sx; rewrite ?rpredV rpredX. Qed.
+
+Lemma rpredXsign R S (divS : @divrPred R S) (kS : keyed_pred divS) n x :
+ (x ^ ((-1) ^+ n) \in kS) = (x \in kS).
+Proof. by rewrite -signr_odd; case: (odd n); rewrite ?rpredV. Qed.
+
+End rpred. \ No newline at end of file
diff --git a/mathcomp/algebra/ssrnum.v b/mathcomp/algebra/ssrnum.v
new file mode 100644
index 0000000..ab7afd0
--- /dev/null
+++ b/mathcomp/algebra/ssrnum.v
@@ -0,0 +1,4219 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div choice fintype.
+Require Import bigop ssralg finset fingroup zmodp poly.
+
+(******************************************************************************)
+(* *)
+(* This file defines some classes to manipulate number structures, i.e *)
+(* structures with an order and a norm *)
+(* *)
+(* * NumDomain (Integral domain with an order and a norm) *)
+(* NumMixin == the mixin that provides an order and a norm over *)
+(* a ring and their characteristic properties. *)
+(* numDomainType == interface for a num integral domain. *)
+(* NumDomainType T m *)
+(* == packs the num mixin into a numberDomainType. The *)
+(* carrier T must have a integral domain structure. *)
+(* [numDomainType of T for S ] *)
+(* == T-clone of the numDomainType structure S. *)
+(* [numDomainType of T] *)
+(* == clone of a canonical numDomainType structure on T. *)
+(* *)
+(* * NumField (Field with an order and a norm) *)
+(* numFieldType == interface for a num field. *)
+(* [numFieldType of T] *)
+(* == clone of a canonical numFieldType structure on T *)
+(* *)
+(* * NumClosedField (Closed Field with an order and a norm) *)
+(* numClosedFieldType *)
+(* == interface for a num closed field. *)
+(* [numClosedFieldType of T] *)
+(* == clone of a canonical numClosedFieldType structure on T *)
+(* *)
+(* * RealDomain (Num domain where all elements are positive or negative) *)
+(* realDomainType == interface for a real integral domain. *)
+(* RealDomainType T r *)
+(* == packs the real axiom r into a realDomainType. The *)
+(* carrier T must have a num domain structure. *)
+(* [realDomainType of T for S ] *)
+(* == T-clone of the realDomainType structure S. *)
+(* [realDomainType of T] *)
+(* == clone of a canonical realDomainType structure on T. *)
+(* *)
+(* * RealField (Num Field where all elements are positive or negative) *)
+(* realFieldType == interface for a real field. *)
+(* [realFieldType of T] *)
+(* == clone of a canonical realFieldType structure on T *)
+(* *)
+(* * ArchiField (A Real Field with the archimedean axiom) *)
+(* archiFieldType == interface for an archimedean field. *)
+(* ArchiFieldType T r *)
+(* == packs the archimeadean axiom r into an archiFieldType. *)
+(* The carrier T must have a real field type structure. *)
+(* [archiFieldType of T for S ] *)
+(* == T-clone of the archiFieldType structure S. *)
+(* [archiFieldType of T] *)
+(* == clone of a canonical archiFieldType structure on T *)
+(* *)
+(* * RealClosedField (Real Field with the real closed axiom) *)
+(* realClosedFieldType *)
+(* == interface for a real closed field. *)
+(* RealClosedFieldType T r *)
+(* == packs the real closed axiom r into a *)
+(* realClodedFieldType. The carrier T must have a real *)
+(* field type structure. *)
+(* [realClosedFieldType of T for S ] *)
+(* == T-clone of the realClosedFieldType structure S. *)
+(* [realClosedFieldype of T] *)
+(* == clone of a canonical realClosedFieldType structure on *)
+(* T. *)
+(* *)
+(* Over these structures, we have the following operations *)
+(* `|x| == norm of x. *)
+(* x <= y <=> x is less than or equal to y (:= '|y - x| == y - x). *)
+(* x < y <=> x is less than y (:= (x <= y) && (x != y)). *)
+(* x <= y ?= iff C <-> x is less than y, or equal iff C is true. *)
+(* Num.sg x == sign of x: equal to 0 iff x = 0, to 1 iff x > 0, and *)
+(* to -1 in all other cases (including x < 0). *)
+(* x \is a Num.pos <=> x is positive (:= x > 0). *)
+(* x \is a Num.neg <=> x is negative (:= x < 0). *)
+(* x \is a Num.nneg <=> x is positive or 0 (:= x >= 0). *)
+(* x \is a Num.real <=> x is real (:= x >= 0 or x < 0). *)
+(* Num.min x y == minimum of x y *)
+(* Num.max x y == maximum of x y *)
+(* Num.bound x == in archimedean fields, and upper bound for x, i.e., *)
+(* and n such that `|x| < n%:R. *)
+(* Num.sqrt x == in a real-closed field, a positive square root of x if *)
+(* x >= 0, or 0 otherwise. *)
+(* *)
+(* There are now three distinct uses of the symbols <, <=, > and >=: *)
+(* 0-ary, unary (prefix) and binary (infix). *)
+(* 0. <%R, <=%R, >%R, >=%R stand respectively for lt, le, gt and ge. *)
+(* 1. (< x), (<= x), (> x), (>= x) stand respectively for *)
+(* (gt x), (ge x), (lt x), (le x). *)
+(* So (< x) is a predicate characterizing elements smaller than x. *)
+(* 2. (x < y), (x <= y), ... mean what they are expected to. *)
+(* These convention are compatible with haskell's, *)
+(* where ((< y) x) = (x < y) = ((<) x y), *)
+(* except that we write <%R instead of (<). *)
+(* *)
+(* - list of prefixes : *)
+(* p : positive *)
+(* n : negative *)
+(* sp : strictly positive *)
+(* sn : strictly negative *)
+(* i : interior = in [0, 1] or ]0, 1[ *)
+(* e : exterior = in [1, +oo[ or ]1; +oo[ *)
+(* w : non strict (weak) monotony *)
+(* *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Local Open Scope ring_scope.
+Import GRing.Theory.
+
+Reserved Notation "<= y" (at level 35).
+Reserved Notation ">= y" (at level 35).
+Reserved Notation "< y" (at level 35).
+Reserved Notation "> y" (at level 35).
+Reserved Notation "<= y :> T" (at level 35, y at next level).
+Reserved Notation ">= y :> T" (at level 35, y at next level).
+Reserved Notation "< y :> T" (at level 35, y at next level).
+Reserved Notation "> y :> T" (at level 35, y at next level).
+
+Module Num.
+
+(* Principal mixin; further classes add axioms rather than operations. *)
+Record mixin_of (R : ringType) := Mixin {
+ norm_op : R -> R;
+ le_op : rel R;
+ lt_op : rel R;
+ _ : forall x y, le_op (norm_op (x + y)) (norm_op x + norm_op y);
+ _ : forall x y, lt_op 0 x -> lt_op 0 y -> lt_op 0 (x + y);
+ _ : forall x, norm_op x = 0 -> x = 0;
+ _ : forall x y, le_op 0 x -> le_op 0 y -> le_op x y || le_op y x;
+ _ : {morph norm_op : x y / x * y};
+ _ : forall x y, (le_op x y) = (norm_op (y - x) == y - x);
+ _ : forall x y, (lt_op x y) = (y != x) && (le_op x y)
+}.
+
+Local Notation ring_for T b := (@GRing.Ring.Pack T b T).
+
+(* Base interface. *)
+Module NumDomain.
+
+Section ClassDef.
+
+Record class_of T := Class {
+ base : GRing.IntegralDomain.class_of T;
+ mixin : mixin_of (ring_for T base)
+}.
+Local Coercion base : class_of >-> GRing.IntegralDomain.class_of.
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+Definition clone c of phant_id class c := @Pack T c T.
+Definition pack b0 (m0 : mixin_of (ring_for T b0)) :=
+ fun bT b & phant_id (GRing.IntegralDomain.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.IntegralDomain.class_of.
+Coercion mixin : class_of >-> mixin_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> GRing.IntegralDomain.type.
+Canonical idomainType.
+Notation numDomainType := type.
+Notation NumMixin := Mixin.
+Notation NumDomainType T m := (@pack T _ m _ _ id _ id).
+Notation "[ 'numDomainType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'numDomainType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'numDomainType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'numDomainType' 'of' T ]") : form_scope.
+End Exports.
+
+End NumDomain.
+Import NumDomain.Exports.
+
+Module Import Def. Section Def.
+Import NumDomain.
+Context {R : type}.
+Implicit Types (x y : R) (C : bool).
+
+Definition normr : R -> R := norm_op (class R).
+Definition ler : rel R := le_op (class R).
+Definition ltr : rel R := lt_op (class R).
+Local Notation "x <= y" := (ler x y) : ring_scope.
+Local Notation "x < y" := (ltr x y) : ring_scope.
+
+Definition ger : simpl_rel R := [rel x y | y <= x].
+Definition gtr : simpl_rel R := [rel x y | y < x].
+Definition lerif x y C : Prop := ((x <= y) * ((x == y) = C))%type.
+Definition sgr x : R := if x == 0 then 0 else if x < 0 then -1 else 1.
+Definition minr x y : R := if x <= y then x else y.
+Definition maxr x y : R := if y <= x then x else y.
+
+Definition Rpos : qualifier 0 R := [qualify x : R | 0 < x].
+Definition Rneg : qualifier 0 R := [qualify x : R | x < 0].
+Definition Rnneg : qualifier 0 R := [qualify x : R | 0 <= x].
+Definition Rreal : qualifier 0 R := [qualify x : R | (0 <= x) || (x <= 0)].
+End Def. End Def.
+
+(* Shorter qualified names, when Num.Def is not imported. *)
+Notation norm := normr.
+Notation le := ler.
+Notation lt := ltr.
+Notation ge := ger.
+Notation gt := gtr.
+Notation sg := sgr.
+Notation max := maxr.
+Notation min := minr.
+Notation pos := Rpos.
+Notation neg := Rneg.
+Notation nneg := Rnneg.
+Notation real := Rreal.
+
+Module Keys. Section Keys.
+Variable R : numDomainType.
+Fact Rpos_key : pred_key (@pos R). Proof. by []. Qed.
+Definition Rpos_keyed := KeyedQualifier Rpos_key.
+Fact Rneg_key : pred_key (@real R). Proof. by []. Qed.
+Definition Rneg_keyed := KeyedQualifier Rneg_key.
+Fact Rnneg_key : pred_key (@nneg R). Proof. by []. Qed.
+Definition Rnneg_keyed := KeyedQualifier Rnneg_key.
+Fact Rreal_key : pred_key (@real R). Proof. by []. Qed.
+Definition Rreal_keyed := KeyedQualifier Rreal_key.
+Definition ler_of_leif x y C (le_xy : @lerif R x y C) := le_xy.1 : le x y.
+End Keys. End Keys.
+
+(* (Exported) symbolic syntax. *)
+Module Import Syntax.
+Import Def Keys.
+
+Notation "`| x |" := (norm x) : ring_scope.
+
+Notation "<%R" := lt : ring_scope.
+Notation ">%R" := gt : ring_scope.
+Notation "<=%R" := le : ring_scope.
+Notation ">=%R" := ge : ring_scope.
+Notation "<?=%R" := lerif : ring_scope.
+
+Notation "< y" := (gt y) : ring_scope.
+Notation "< y :> T" := (< (y : T)) : ring_scope.
+Notation "> y" := (lt y) : ring_scope.
+Notation "> y :> T" := (> (y : T)) : ring_scope.
+
+Notation "<= y" := (ge y) : ring_scope.
+Notation "<= y :> T" := (<= (y : T)) : ring_scope.
+Notation ">= y" := (le y) : ring_scope.
+Notation ">= y :> T" := (>= (y : T)) : ring_scope.
+
+Notation "x < y" := (lt x y) : ring_scope.
+Notation "x < y :> T" := ((x : T) < (y : T)) : ring_scope.
+Notation "x > y" := (y < x) (only parsing) : ring_scope.
+Notation "x > y :> T" := ((x : T) > (y : T)) (only parsing) : ring_scope.
+
+Notation "x <= y" := (le x y) : ring_scope.
+Notation "x <= y :> T" := ((x : T) <= (y : T)) : ring_scope.
+Notation "x >= y" := (y <= x) (only parsing) : ring_scope.
+Notation "x >= y :> T" := ((x : T) >= (y : T)) (only parsing) : ring_scope.
+
+Notation "x <= y <= z" := ((x <= y) && (y <= z)) : ring_scope.
+Notation "x < y <= z" := ((x < y) && (y <= z)) : ring_scope.
+Notation "x <= y < z" := ((x <= y) && (y < z)) : ring_scope.
+Notation "x < y < z" := ((x < y) && (y < z)) : ring_scope.
+
+Notation "x <= y ?= 'iff' C" := (lerif x y C) : ring_scope.
+Notation "x <= y ?= 'iff' C :> R" := ((x : R) <= (y : R) ?= iff C)
+ (only parsing) : ring_scope.
+
+Coercion ler_of_leif : lerif >-> is_true.
+
+Canonical Rpos_keyed.
+Canonical Rneg_keyed.
+Canonical Rnneg_keyed.
+Canonical Rreal_keyed.
+
+End Syntax.
+
+Section ExtensionAxioms.
+
+Variable R : numDomainType.
+
+Definition real_axiom : Prop := forall x : R, x \is real.
+
+Definition archimedean_axiom : Prop := forall x : R, exists ub, `|x| < ub%:R.
+
+Definition real_closed_axiom : Prop :=
+ forall (p : {poly R}) (a b : R),
+ a <= b -> p.[a] <= 0 <= p.[b] -> exists2 x, a <= x <= b & root p x.
+
+End ExtensionAxioms.
+
+Local Notation num_for T b := (@NumDomain.Pack T b T).
+
+(* The rest of the numbers interface hierarchy. *)
+Module NumField.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : GRing.Field.class_of R; mixin : mixin_of (ring_for R base) }.
+Definition base2 R (c : class_of R) := NumDomain.Class (mixin c).
+Local Coercion base : class_of >-> GRing.Field.class_of.
+Local Coercion base2 : class_of >-> NumDomain.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+Definition pack :=
+ fun bT b & phant_id (GRing.Field.class bT) (b : GRing.Field.class_of T) =>
+ fun mT m & phant_id (NumDomain.class mT) (@NumDomain.Class T b m) =>
+ Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT.
+Definition numDomainType := @NumDomain.Pack cT xclass xT.
+Definition fieldType := @GRing.Field.Pack cT xclass xT.
+Definition join_numDomainType := @NumDomain.Pack fieldType xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.Field.class_of.
+Coercion base2 : class_of >-> NumDomain.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> GRing.IntegralDomain.type.
+Canonical idomainType.
+Coercion numDomainType : type >-> NumDomain.type.
+Canonical numDomainType.
+Coercion fieldType : type >-> GRing.Field.type.
+Canonical fieldType.
+Notation numFieldType := type.
+Notation "[ 'numFieldType' 'of' T ]" := (@pack T _ _ id _ _ id)
+ (at level 0, format "[ 'numFieldType' 'of' T ]") : form_scope.
+End Exports.
+
+End NumField.
+Import NumField.Exports.
+
+Module ClosedField.
+
+Section ClassDef.
+
+Record class_of R := Class {
+ base : GRing.ClosedField.class_of R;
+ mixin : mixin_of (ring_for R base)
+}.
+Definition base2 R (c : class_of R) := NumField.Class (mixin c).
+Local Coercion base : class_of >-> GRing.ClosedField.class_of.
+Local Coercion base2 : class_of >-> NumField.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+Definition pack :=
+ fun bT b & phant_id (GRing.ClosedField.class bT)
+ (b : GRing.ClosedField.class_of T) =>
+ fun mT m & phant_id (NumField.class mT) (@NumField.Class T b m) =>
+ Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT.
+Definition numDomainType := @NumDomain.Pack cT xclass xT.
+Definition fieldType := @GRing.Field.Pack cT xclass xT.
+Definition decFieldType := @GRing.DecidableField.Pack cT xclass xT.
+Definition closedFieldType := @GRing.ClosedField.Pack cT xclass xT.
+Definition join_dec_numDomainType := @NumDomain.Pack decFieldType xclass xT.
+Definition join_dec_numFieldType := @NumField.Pack decFieldType xclass xT.
+Definition join_numDomainType := @NumDomain.Pack closedFieldType xclass xT.
+Definition join_numFieldType := @NumField.Pack closedFieldType xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> GRing.ClosedField.class_of.
+Coercion base2 : class_of >-> NumField.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> GRing.IntegralDomain.type.
+Canonical idomainType.
+Coercion numDomainType : type >-> NumDomain.type.
+Canonical numDomainType.
+Coercion fieldType : type >-> GRing.Field.type.
+Canonical fieldType.
+Coercion decFieldType : type >-> GRing.DecidableField.type.
+Canonical decFieldType.
+Coercion closedFieldType : type >-> GRing.ClosedField.type.
+Canonical closedFieldType.
+Canonical join_dec_numDomainType.
+Canonical join_dec_numFieldType.
+Canonical join_numDomainType.
+Canonical join_numFieldType.
+Notation numClosedFieldType := type.
+Notation "[ 'numClosedFieldType' 'of' T ]" := (@pack T _ _ id _ _ id)
+ (at level 0, format "[ 'numClosedFieldType' 'of' T ]") : form_scope.
+End Exports.
+
+End ClosedField.
+Import ClosedField.Exports.
+
+Module RealDomain.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class {base : NumDomain.class_of R; _ : @real_axiom (num_for R base)}.
+Local Coercion base : class_of >-> NumDomain.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+Definition clone c of phant_id class c := @Pack T c T.
+Definition pack b0 (m0 : real_axiom (num_for T b0)) :=
+ fun bT b & phant_id (NumDomain.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT.
+Definition numDomainType := @NumDomain.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> NumDomain.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> GRing.IntegralDomain.type.
+Canonical idomainType.
+Coercion numDomainType : type >-> NumDomain.type.
+Canonical numDomainType.
+Notation realDomainType := type.
+Notation RealDomainType T m := (@pack T _ m _ _ id _ id).
+Notation "[ 'realDomainType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'realDomainType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'realDomainType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'realDomainType' 'of' T ]") : form_scope.
+End Exports.
+
+End RealDomain.
+Import RealDomain.Exports.
+
+Module RealField.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : NumField.class_of R; mixin : real_axiom (num_for R base) }.
+Definition base2 R (c : class_of R) := RealDomain.Class (@mixin R c).
+Local Coercion base : class_of >-> NumField.class_of.
+Local Coercion base2 : class_of >-> RealDomain.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+Definition pack :=
+ fun bT b & phant_id (NumField.class bT) (b : NumField.class_of T) =>
+ fun mT m & phant_id (RealDomain.class mT) (@RealDomain.Class T b m) =>
+ Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT.
+Definition numDomainType := @NumDomain.Pack cT xclass xT.
+Definition realDomainType := @RealDomain.Pack cT xclass xT.
+Definition fieldType := @GRing.Field.Pack cT xclass xT.
+Definition numFieldType := @NumField.Pack cT xclass xT.
+Definition join_realDomainType := @RealDomain.Pack numFieldType xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> NumField.class_of.
+Coercion base2 : class_of >-> RealDomain.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> GRing.IntegralDomain.type.
+Canonical idomainType.
+Coercion numDomainType : type >-> NumDomain.type.
+Canonical numDomainType.
+Coercion realDomainType : type >-> RealDomain.type.
+Canonical realDomainType.
+Coercion fieldType : type >-> GRing.Field.type.
+Canonical fieldType.
+Coercion numFieldType : type >-> NumField.type.
+Canonical numFieldType.
+Canonical join_realDomainType.
+Notation realFieldType := type.
+Notation "[ 'realFieldType' 'of' T ]" := (@pack T _ _ id _ _ id)
+ (at level 0, format "[ 'realFieldType' 'of' T ]") : form_scope.
+End Exports.
+
+End RealField.
+Import RealField.Exports.
+
+Module ArchimedeanField.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : RealField.class_of R; _ : archimedean_axiom (num_for R base) }.
+Local Coercion base : class_of >-> RealField.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+Definition clone c of phant_id class c := @Pack T c T.
+Definition pack b0 (m0 : archimedean_axiom (num_for T b0)) :=
+ fun bT b & phant_id (RealField.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT.
+Definition numDomainType := @NumDomain.Pack cT xclass xT.
+Definition realDomainType := @RealDomain.Pack cT xclass xT.
+Definition fieldType := @GRing.Field.Pack cT xclass xT.
+Definition numFieldType := @NumField.Pack cT xclass xT.
+Definition realFieldType := @RealField.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> RealField.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> GRing.IntegralDomain.type.
+Canonical idomainType.
+Coercion numDomainType : type >-> NumDomain.type.
+Canonical numDomainType.
+Coercion realDomainType : type >-> RealDomain.type.
+Canonical realDomainType.
+Coercion fieldType : type >-> GRing.Field.type.
+Canonical fieldType.
+Coercion numFieldType : type >-> NumField.type.
+Canonical numFieldType.
+Coercion realFieldType : type >-> RealField.type.
+Canonical realFieldType.
+Notation archiFieldType := type.
+Notation ArchiFieldType T m := (@pack T _ m _ _ id _ id).
+Notation "[ 'archiFieldType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'archiFieldType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'archiFieldType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'archiFieldType' 'of' T ]") : form_scope.
+End Exports.
+
+End ArchimedeanField.
+Import ArchimedeanField.Exports.
+
+Module RealClosedField.
+
+Section ClassDef.
+
+Record class_of R :=
+ Class { base : RealField.class_of R; _ : real_closed_axiom (num_for R base) }.
+Local Coercion base : class_of >-> RealField.class_of.
+
+Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (T : Type) (cT : type).
+Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+Definition clone c of phant_id class c := @Pack T c T.
+Definition pack b0 (m0 : real_closed_axiom (num_for T b0)) :=
+ fun bT b & phant_id (RealField.class bT) b =>
+ fun m & phant_id m0 m => Pack (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition ringType := @GRing.Ring.Pack cT xclass xT.
+Definition comRingType := @GRing.ComRing.Pack cT xclass xT.
+Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT.
+Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT.
+Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT.
+Definition numDomainType := @NumDomain.Pack cT xclass xT.
+Definition realDomainType := @RealDomain.Pack cT xclass xT.
+Definition fieldType := @GRing.Field.Pack cT xclass xT.
+Definition numFieldType := @NumField.Pack cT xclass xT.
+Definition realFieldType := @RealField.Pack cT xclass xT.
+
+End ClassDef.
+
+Module Exports.
+Coercion base : class_of >-> RealField.class_of.
+Coercion sort : type >-> Sortclass.
+Bind Scope ring_scope with sort.
+Coercion eqType : type >-> Equality.type.
+Canonical eqType.
+Coercion choiceType : type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType : type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion ringType : type >-> GRing.Ring.type.
+Canonical ringType.
+Coercion comRingType : type >-> GRing.ComRing.type.
+Canonical comRingType.
+Coercion unitRingType : type >-> GRing.UnitRing.type.
+Canonical unitRingType.
+Coercion comUnitRingType : type >-> GRing.ComUnitRing.type.
+Canonical comUnitRingType.
+Coercion idomainType : type >-> GRing.IntegralDomain.type.
+Canonical idomainType.
+Coercion numDomainType : type >-> NumDomain.type.
+Canonical numDomainType.
+Coercion realDomainType : type >-> RealDomain.type.
+Canonical realDomainType.
+Coercion fieldType : type >-> GRing.Field.type.
+Canonical fieldType.
+Coercion numFieldType : type >-> NumField.type.
+Canonical numFieldType.
+Coercion realFieldType : type >-> RealField.type.
+Canonical realFieldType.
+Notation rcfType := Num.RealClosedField.type.
+Notation RcfType T m := (@pack T _ m _ _ id _ id).
+Notation "[ 'rcfType' 'of' T 'for' cT ]" := (@clone T cT _ idfun)
+ (at level 0, format "[ 'rcfType' 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'rcfType' 'of' T ]" := (@clone T _ _ id)
+ (at level 0, format "[ 'rcfType' 'of' T ]") : form_scope.
+End Exports.
+
+End RealClosedField.
+Import RealClosedField.Exports.
+
+(* The elementary theory needed to support the definition of the derived *)
+(* operations for the extensions described above. *)
+Module Import Internals.
+
+Section Domain.
+Variable R : numDomainType.
+Implicit Types x y : R.
+
+(* Lemmas from the signature *)
+
+Lemma normr0_eq0 x : `|x| = 0 -> x = 0.
+Proof. by case: R x => ? [? []]. Qed.
+
+Lemma ler_norm_add x y : `|x + y| <= `|x| + `|y|.
+Proof. by case: R x y => ? [? []]. Qed.
+
+Lemma addr_gt0 x y : 0 < x -> 0 < y -> 0 < x + y.
+Proof. by case: R x y => ? [? []]. Qed.
+
+Lemma ger_leVge x y : 0 <= x -> 0 <= y -> (x <= y) || (y <= x).
+Proof. by case: R x y => ? [? []]. Qed.
+
+Lemma normrM : {morph norm : x y / x * y : R}.
+Proof. by case: R => ? [? []]. Qed.
+
+Lemma ler_def x y : (x <= y) = (`|y - x| == y - x).
+Proof. by case: R x y => ? [? []]. Qed.
+
+Lemma ltr_def x y : (x < y) = (y != x) && (x <= y).
+Proof. by case: R x y => ? [? []]. Qed.
+
+(* Basic consequences (just enough to get predicate closure properties). *)
+
+Lemma ger0_def x : (0 <= x) = (`|x| == x).
+Proof. by rewrite ler_def subr0. Qed.
+
+Lemma subr_ge0 x y : (0 <= x - y) = (y <= x).
+Proof. by rewrite ger0_def -ler_def. Qed.
+
+Lemma oppr_ge0 x : (0 <= - x) = (x <= 0).
+Proof. by rewrite -sub0r subr_ge0. Qed.
+
+Lemma ler01 : 0 <= 1 :> R.
+Proof.
+have n1_nz: `|1| != 0 :> R by apply: contraNneq (@oner_neq0 R) => /normr0_eq0->.
+by rewrite ger0_def -(inj_eq (mulfI n1_nz)) -normrM !mulr1.
+Qed.
+
+Lemma ltr01 : 0 < 1 :> R. Proof. by rewrite ltr_def oner_neq0 ler01. Qed.
+
+Lemma ltrW x y : x < y -> x <= y. Proof. by rewrite ltr_def => /andP[]. Qed.
+
+Lemma lerr x : x <= x.
+Proof.
+have n2: `|2%:R| == 2%:R :> R by rewrite -ger0_def ltrW ?addr_gt0 ?ltr01.
+rewrite ler_def subrr -(inj_eq (addrI `|0|)) addr0 -mulr2n -mulr_natr.
+by rewrite -(eqP n2) -normrM mul0r.
+Qed.
+
+Lemma le0r x : (0 <= x) = (x == 0) || (0 < x).
+Proof. by rewrite ltr_def; case: eqP => // ->; rewrite lerr. Qed.
+
+Lemma addr_ge0 x y : 0 <= x -> 0 <= y -> 0 <= x + y.
+Proof.
+rewrite le0r; case/predU1P=> [-> | x_pos]; rewrite ?add0r // le0r.
+by case/predU1P=> [-> | y_pos]; rewrite ltrW ?addr0 ?addr_gt0.
+Qed.
+
+Lemma pmulr_rgt0 x y : 0 < x -> (0 < x * y) = (0 < y).
+Proof.
+rewrite !ltr_def !ger0_def normrM mulf_eq0 negb_or => /andP[x_neq0 /eqP->].
+by rewrite x_neq0 (inj_eq (mulfI x_neq0)).
+Qed.
+
+(* Closure properties of the real predicates. *)
+
+Lemma posrE x : (x \is pos) = (0 < x). Proof. by []. Qed.
+Lemma nnegrE x : (x \is nneg) = (0 <= x). Proof. by []. Qed.
+Lemma realE x : (x \is real) = (0 <= x) || (x <= 0). Proof. by []. Qed.
+
+Fact pos_divr_closed : divr_closed (@pos R).
+Proof.
+split=> [|x y x_gt0 y_gt0]; rewrite posrE ?ltr01 //.
+have [Uy|/invr_out->] := boolP (y \is a GRing.unit); last by rewrite pmulr_rgt0.
+by rewrite -(pmulr_rgt0 _ y_gt0) mulrC divrK.
+Qed.
+Canonical pos_mulrPred := MulrPred pos_divr_closed.
+Canonical pos_divrPred := DivrPred pos_divr_closed.
+
+Fact nneg_divr_closed : divr_closed (@nneg R).
+Proof.
+split=> [|x y]; rewrite !nnegrE ?ler01 ?le0r // -!posrE.
+case/predU1P=> [-> _ | x_gt0]; first by rewrite mul0r eqxx.
+by case/predU1P=> [-> | y_gt0]; rewrite ?invr0 ?mulr0 ?eqxx // orbC rpred_div.
+Qed.
+Canonical nneg_mulrPred := MulrPred nneg_divr_closed.
+Canonical nneg_divrPred := DivrPred nneg_divr_closed.
+
+Fact nneg_addr_closed : addr_closed (@nneg R).
+Proof. by split; [exact: lerr | exact: addr_ge0]. Qed.
+Canonical nneg_addrPred := AddrPred nneg_addr_closed.
+Canonical nneg_semiringPred := SemiringPred nneg_divr_closed.
+
+Fact real_oppr_closed : oppr_closed (@real R).
+Proof. by move=> x; rewrite /= !realE oppr_ge0 orbC -!oppr_ge0 opprK. Qed.
+Canonical real_opprPred := OpprPred real_oppr_closed.
+
+Fact real_addr_closed : addr_closed (@real R).
+Proof.
+split=> [|x y Rx Ry]; first by rewrite realE lerr.
+without loss{Rx} x_ge0: x y Ry / 0 <= x.
+ case/orP: Rx => [? | x_le0]; first exact.
+ by rewrite -rpredN opprD; apply; rewrite ?rpredN ?oppr_ge0.
+case/orP: Ry => [y_ge0 | y_le0]; first by rewrite realE -nnegrE rpredD.
+by rewrite realE -[y]opprK orbC -oppr_ge0 opprB !subr_ge0 ger_leVge ?oppr_ge0.
+Qed.
+Canonical real_addrPred := AddrPred real_addr_closed.
+Canonical real_zmodPred := ZmodPred real_oppr_closed.
+
+Fact real_divr_closed : divr_closed (@real R).
+Proof.
+split=> [|x y Rx Ry]; first by rewrite realE ler01.
+without loss{Rx} x_ge0: x / 0 <= x.
+ case/orP: Rx => [? | x_le0]; first exact.
+ by rewrite -rpredN -mulNr; apply; rewrite ?oppr_ge0.
+without loss{Ry} y_ge0: y / 0 <= y; last by rewrite realE -nnegrE rpred_div.
+case/orP: Ry => [? | y_le0]; first exact.
+by rewrite -rpredN -mulrN -invrN; apply; rewrite ?oppr_ge0.
+Qed.
+Canonical real_mulrPred := MulrPred real_divr_closed.
+Canonical real_smulrPred := SmulrPred real_divr_closed.
+Canonical real_divrPred := DivrPred real_divr_closed.
+Canonical real_sdivrPred := SdivrPred real_divr_closed.
+Canonical real_semiringPred := SemiringPred real_divr_closed.
+Canonical real_subringPred := SubringPred real_divr_closed.
+Canonical real_divringPred := DivringPred real_divr_closed.
+
+End Domain.
+
+Lemma num_real (R : realDomainType) (x : R) : x \is real.
+Proof. by case: R x => T []. Qed.
+
+Fact archi_bound_subproof (R : archiFieldType) : archimedean_axiom R.
+Proof. by case: R => ? []. Qed.
+
+Section RealClosed.
+Variable R : rcfType.
+
+Lemma poly_ivt : real_closed_axiom R. Proof. by case: R => ? []. Qed.
+
+Fact sqrtr_subproof (x : R) :
+ exists2 y, 0 <= y & if 0 <= x return bool then y ^+ 2 == x else y == 0.
+Proof.
+case x_ge0: (0 <= x); last by exists 0; rewrite ?lerr.
+have le0x1: 0 <= x + 1 by rewrite -nnegrE rpredD ?rpred1.
+have [|y /andP[y_ge0 _]] := @poly_ivt ('X^2 - x%:P) _ _ le0x1.
+ rewrite !hornerE -subr_ge0 add0r opprK x_ge0 -expr2 sqrrD mulr1.
+ by rewrite addrAC !addrA addrK -nnegrE !rpredD ?rpredX ?rpred1.
+by rewrite rootE !hornerE subr_eq0; exists y.
+Qed.
+
+End RealClosed.
+
+End Internals.
+
+Module PredInstances.
+
+Canonical pos_mulrPred.
+Canonical pos_divrPred.
+
+Canonical nneg_addrPred.
+Canonical nneg_mulrPred.
+Canonical nneg_divrPred.
+Canonical nneg_semiringPred.
+
+Canonical real_addrPred.
+Canonical real_opprPred.
+Canonical real_zmodPred.
+Canonical real_mulrPred.
+Canonical real_smulrPred.
+Canonical real_divrPred.
+Canonical real_sdivrPred.
+Canonical real_semiringPred.
+Canonical real_subringPred.
+Canonical real_divringPred.
+
+End PredInstances.
+
+Module Import ExtraDef.
+
+Definition archi_bound {R} x := sval (sigW (@archi_bound_subproof R x)).
+
+Definition sqrtr {R} x := s2val (sig2W (@sqrtr_subproof R x)).
+
+End ExtraDef.
+
+Notation bound := archi_bound.
+Notation sqrt := sqrtr.
+
+Module Theory.
+
+Section NumIntegralDomainTheory.
+
+Variable R : numDomainType.
+Implicit Types x y z t : R.
+
+(* Lemmas from the signature (reexported from internals). *)
+
+Definition ler_norm_add x y : `|x + y| <= `|x| + `|y| := ler_norm_add x y.
+Definition addr_gt0 x y : 0 < x -> 0 < y -> 0 < x + y := @addr_gt0 R x y.
+Definition normr0_eq0 x : `|x| = 0 -> x = 0 := @normr0_eq0 R x.
+Definition ger_leVge x y : 0 <= x -> 0 <= y -> (x <= y) || (y <= x) :=
+ @ger_leVge R x y.
+Definition normrM : {morph normr : x y / x * y : R} := @normrM R.
+Definition ler_def x y : (x <= y) = (`|y - x| == y - x) := @ler_def R x y.
+Definition ltr_def x y : (x < y) = (y != x) && (x <= y) := @ltr_def R x y.
+
+(* Predicate and relation definitions. *)
+
+Lemma gerE x y : ge x y = (y <= x). Proof. by []. Qed.
+Lemma gtrE x y : gt x y = (y < x). Proof. by []. Qed.
+Lemma posrE x : (x \is pos) = (0 < x). Proof. by []. Qed.
+Lemma negrE x : (x \is neg) = (x < 0). Proof. by []. Qed.
+Lemma nnegrE x : (x \is nneg) = (0 <= x). Proof. by []. Qed.
+Lemma realE x : (x \is real) = (0 <= x) || (x <= 0). Proof. by []. Qed.
+
+(* General properties of <= and < *)
+
+Lemma lerr x : x <= x. Proof. exact: lerr. Qed.
+Lemma ltrr x : x < x = false. Proof. by rewrite ltr_def eqxx. Qed.
+Lemma ltrW x y : x < y -> x <= y. Proof. exact: ltrW. Qed.
+Hint Resolve lerr ltrr ltrW.
+
+Lemma ltr_neqAle x y : (x < y) = (x != y) && (x <= y).
+Proof. by rewrite ltr_def eq_sym. Qed.
+
+Lemma ler_eqVlt x y : (x <= y) = (x == y) || (x < y).
+Proof. by rewrite ltr_neqAle; case: eqP => // ->; rewrite lerr. Qed.
+
+Lemma lt0r x : (0 < x) = (x != 0) && (0 <= x). Proof. by rewrite ltr_def. Qed.
+Lemma le0r x : (0 <= x) = (x == 0) || (0 < x). Proof. exact: le0r. Qed.
+
+Lemma lt0r_neq0 (x : R) : 0 < x -> x != 0.
+Proof. by rewrite lt0r; case/andP. Qed.
+
+Lemma ltr0_neq0 (x : R) : 0 < x -> x != 0.
+Proof. by rewrite lt0r; case/andP. Qed.
+
+Lemma gtr_eqF x y : y < x -> x == y = false.
+Proof. by rewrite ltr_def; case/andP; move/negPf=> ->. Qed.
+
+Lemma ltr_eqF x y : x < y -> x == y = false.
+Proof. by move=> hyx; rewrite eq_sym gtr_eqF. Qed.
+
+Lemma pmulr_rgt0 x y : 0 < x -> (0 < x * y) = (0 < y).
+Proof. exact: pmulr_rgt0. Qed.
+
+Lemma pmulr_rge0 x y : 0 < x -> (0 <= x * y) = (0 <= y).
+Proof.
+by rewrite !le0r mulf_eq0; case: eqP => // [-> /negPf[] | _ /pmulr_rgt0->].
+Qed.
+
+(* Integer comparisons and characteristic 0. *)
+Lemma ler01 : 0 <= 1 :> R. Proof. exact: ler01. Qed.
+Lemma ltr01 : 0 < 1 :> R. Proof. exact: ltr01. Qed.
+Lemma ler0n n : 0 <= n%:R :> R. Proof. by rewrite -nnegrE rpred_nat. Qed.
+Hint Resolve ler01 ltr01 ler0n.
+Lemma ltr0Sn n : 0 < n.+1%:R :> R.
+Proof. by elim: n => // n; apply: addr_gt0. Qed.
+Lemma ltr0n n : (0 < n%:R :> R) = (0 < n)%N.
+Proof. by case: n => //= n; apply: ltr0Sn. Qed.
+Hint Resolve ltr0Sn.
+
+Lemma pnatr_eq0 n : (n%:R == 0 :> R) = (n == 0)%N.
+Proof. by case: n => [|n]; rewrite ?mulr0n ?eqxx // gtr_eqF. Qed.
+
+Lemma char_num : [char R] =i pred0.
+Proof. by case=> // p /=; rewrite !inE pnatr_eq0 andbF. Qed.
+
+(* Properties of the norm. *)
+
+Lemma ger0_def x : (0 <= x) = (`|x| == x). Proof. exact: ger0_def. Qed.
+Lemma normr_idP {x} : reflect (`|x| = x) (0 <= x).
+Proof. by rewrite ger0_def; apply: eqP. Qed.
+Lemma ger0_norm x : 0 <= x -> `|x| = x. Proof. exact: normr_idP. Qed.
+
+Lemma normr0 : `|0| = 0 :> R. Proof. exact: ger0_norm. Qed.
+Lemma normr1 : `|1| = 1 :> R. Proof. exact: ger0_norm. Qed.
+Lemma normr_nat n : `|n%:R| = n%:R :> R. Proof. exact: ger0_norm. Qed.
+Lemma normrMn x n : `|x *+ n| = `|x| *+ n.
+Proof. by rewrite -mulr_natl normrM normr_nat mulr_natl. Qed.
+
+Lemma normr_prod I r (P : pred I) (F : I -> R) :
+ `|\prod_(i <- r | P i) F i| = \prod_(i <- r | P i) `|F i|.
+Proof. exact: (big_morph norm normrM normr1). Qed.
+
+Lemma normrX n x : `|x ^+ n| = `|x| ^+ n.
+Proof. by rewrite -(card_ord n) -!prodr_const normr_prod. Qed.
+
+Lemma normr_unit : {homo (@norm R) : x / x \is a GRing.unit}.
+Proof.
+move=> x /= /unitrP [y [yx xy]]; apply/unitrP; exists `|y|.
+by rewrite -!normrM xy yx normr1.
+Qed.
+
+Lemma normrV : {in GRing.unit, {morph (@normr R) : x / x ^-1}}.
+Proof.
+move=> x ux; apply: (mulrI (normr_unit ux)).
+by rewrite -normrM !divrr ?normr1 ?normr_unit.
+Qed.
+
+Lemma normr0P {x} : reflect (`|x| = 0) (x == 0).
+Proof. by apply: (iffP eqP)=> [->|/normr0_eq0 //]; apply: normr0. Qed.
+
+Definition normr_eq0 x := sameP (`|x| =P 0) normr0P.
+
+Lemma normrN1 : `|-1| = 1 :> R.
+Proof.
+have: `|-1| ^+ 2 == 1 :> R by rewrite -normrX -signr_odd normr1.
+rewrite sqrf_eq1 => /orP[/eqP //|]; rewrite -ger0_def le0r oppr_eq0 oner_eq0.
+by move/(addr_gt0 ltr01); rewrite subrr ltrr.
+Qed.
+
+Lemma normrN x : `|- x| = `|x|.
+Proof. by rewrite -mulN1r normrM normrN1 mul1r. Qed.
+
+Lemma distrC x y : `|x - y| = `|y - x|.
+Proof. by rewrite -opprB normrN. Qed.
+
+Lemma ler0_def x : (x <= 0) = (`|x| == - x).
+Proof. by rewrite ler_def sub0r normrN. Qed.
+
+Lemma normr_id x : `|`|x| | = `|x|.
+Proof.
+have nz2: 2%:R != 0 :> R by rewrite pnatr_eq0.
+apply: (mulfI nz2); rewrite -{1}normr_nat -normrM mulr_natl mulr2n ger0_norm //.
+by rewrite -{2}normrN -normr0 -(subrr x) ler_norm_add.
+Qed.
+
+Lemma normr_ge0 x : 0 <= `|x|. Proof. by rewrite ger0_def normr_id. Qed.
+Hint Resolve normr_ge0.
+
+Lemma ler0_norm x : x <= 0 -> `|x| = - x.
+Proof. by move=> x_le0; rewrite -[r in _ = r]ger0_norm ?normrN ?oppr_ge0. Qed.
+
+Definition gtr0_norm x (hx : 0 < x) := ger0_norm (ltrW hx).
+Definition ltr0_norm x (hx : x < 0) := ler0_norm (ltrW hx).
+
+(* Comparision to 0 of a difference *)
+
+Lemma subr_ge0 x y : (0 <= y - x) = (x <= y). Proof. exact: subr_ge0. Qed.
+Lemma subr_gt0 x y : (0 < y - x) = (x < y).
+Proof. by rewrite !ltr_def subr_eq0 subr_ge0. Qed.
+Lemma subr_le0 x y : (y - x <= 0) = (y <= x).
+Proof. by rewrite -subr_ge0 opprB add0r subr_ge0. Qed.
+Lemma subr_lt0 x y : (y - x < 0) = (y < x).
+Proof. by rewrite -subr_gt0 opprB add0r subr_gt0. Qed.
+
+Definition subr_lte0 := (subr_le0, subr_lt0).
+Definition subr_gte0 := (subr_ge0, subr_gt0).
+Definition subr_cp0 := (subr_lte0, subr_gte0).
+
+(* Ordered ring properties. *)
+
+Lemma ler_asym : antisymmetric (<=%R : rel R).
+Proof.
+move=> x y; rewrite !ler_def distrC -opprB -addr_eq0 => /andP[/eqP->].
+by rewrite -mulr2n -mulr_natl mulf_eq0 subr_eq0 pnatr_eq0 => /eqP.
+Qed.
+
+Lemma eqr_le x y : (x == y) = (x <= y <= x).
+Proof. by apply/eqP/idP=> [->|/ler_asym]; rewrite ?lerr. Qed.
+
+Lemma ltr_trans : transitive (@ltr R).
+Proof.
+move=> y x z le_xy le_yz.
+by rewrite -subr_gt0 -(subrK y z) -addrA addr_gt0 ?subr_gt0.
+Qed.
+
+Lemma ler_lt_trans y x z : x <= y -> y < z -> x < z.
+Proof. by rewrite !ler_eqVlt => /orP[/eqP -> //|/ltr_trans]; apply. Qed.
+
+Lemma ltr_le_trans y x z : x < y -> y <= z -> x < z.
+Proof. by rewrite !ler_eqVlt => lxy /orP[/eqP <- //|/(ltr_trans lxy)]. Qed.
+
+Lemma ler_trans : transitive (@ler R).
+Proof.
+move=> y x z; rewrite !ler_eqVlt => /orP [/eqP -> //|lxy].
+by move=> /orP [/eqP <-|/(ltr_trans lxy) ->]; rewrite ?lxy orbT.
+Qed.
+
+Definition lter01 := (ler01, ltr01).
+Definition lterr := (lerr, ltrr).
+
+Lemma addr_ge0 x y : 0 <= x -> 0 <= y -> 0 <= x + y.
+Proof. exact: addr_ge0. Qed.
+
+Lemma lerifP x y C : reflect (x <= y ?= iff C) (if C then x == y else x < y).
+Proof.
+rewrite /lerif ler_eqVlt; apply: (iffP idP)=> [|[]].
+ by case: C => [/eqP->|lxy]; rewrite ?eqxx // lxy ltr_eqF.
+by move=> /orP[/eqP->|lxy] <-; rewrite ?eqxx // ltr_eqF.
+Qed.
+
+Lemma ltr_asym x y : x < y < x = false.
+Proof. by apply/negP=> /andP [/ltr_trans hyx /hyx]; rewrite ltrr. Qed.
+
+Lemma ler_anti : antisymmetric (@ler R).
+Proof. by move=> x y; rewrite -eqr_le=> /eqP. Qed.
+
+Lemma ltr_le_asym x y : x < y <= x = false.
+Proof. by rewrite ltr_neqAle -andbA -eqr_le eq_sym; case: (_ == _). Qed.
+
+Lemma ler_lt_asym x y : x <= y < x = false.
+Proof. by rewrite andbC ltr_le_asym. Qed.
+
+Definition lter_anti := (=^~ eqr_le, ltr_asym, ltr_le_asym, ler_lt_asym).
+
+Lemma ltr_geF x y : x < y -> (y <= x = false).
+Proof.
+by move=> xy; apply: contraTF isT=> /(ltr_le_trans xy); rewrite ltrr.
+Qed.
+
+Lemma ler_gtF x y : x <= y -> (y < x = false).
+Proof. by apply: contraTF=> /ltr_geF->. Qed.
+
+Definition ltr_gtF x y hxy := ler_gtF (@ltrW x y hxy).
+
+(* Norm and order properties. *)
+
+Lemma normr_le0 x : (`|x| <= 0) = (x == 0).
+Proof. by rewrite -normr_eq0 eqr_le normr_ge0 andbT. Qed.
+
+Lemma normr_lt0 x : `|x| < 0 = false.
+Proof. by rewrite ltr_neqAle normr_le0 normr_eq0 andNb. Qed.
+
+Lemma normr_gt0 x : (`|x| > 0) = (x != 0).
+Proof. by rewrite ltr_def normr_eq0 normr_ge0 andbT. Qed.
+
+Definition normrE x := (normr_id, normr0, normr1, normrN1, normr_ge0, normr_eq0,
+ normr_lt0, normr_le0, normr_gt0, normrN).
+
+End NumIntegralDomainTheory.
+
+Implicit Arguments ler01 [R].
+Implicit Arguments ltr01 [R].
+Implicit Arguments normr_idP [R x].
+Implicit Arguments normr0P [R x].
+Implicit Arguments lerifP [R x y C].
+Hint Resolve @ler01 @ltr01 lerr ltrr ltrW ltr_eqF ltr0Sn ler0n normr_ge0.
+
+Section NumIntegralDomainMonotonyTheory.
+
+Variables R R' : numDomainType.
+Implicit Types m n p : nat.
+Implicit Types x y z : R.
+Implicit Types u v w : R'.
+
+Section AcrossTypes.
+
+Variable D D' : pred R.
+Variable (f : R -> R').
+
+Lemma ltrW_homo : {homo f : x y / x < y} -> {homo f : x y / x <= y}.
+Proof. by move=> mf x y /=; rewrite ler_eqVlt => /orP[/eqP->|/mf/ltrW]. Qed.
+
+Lemma ltrW_nhomo : {homo f : x y /~ x < y} -> {homo f : x y /~ x <= y}.
+Proof. by move=> mf x y /=; rewrite ler_eqVlt => /orP[/eqP->|/mf/ltrW]. Qed.
+
+Lemma homo_inj_lt :
+ injective f -> {homo f : x y / x <= y} -> {homo f : x y / x < y}.
+Proof.
+by move=> fI mf x y /= hxy; rewrite ltr_neqAle (inj_eq fI) mf (ltr_eqF, ltrW).
+Qed.
+
+Lemma nhomo_inj_lt :
+ injective f -> {homo f : x y /~ x <= y} -> {homo f : x y /~ x < y}.
+Proof.
+by move=> fI mf x y /= hxy; rewrite ltr_neqAle (inj_eq fI) mf (gtr_eqF, ltrW).
+Qed.
+
+Lemma mono_inj : {mono f : x y / x <= y} -> injective f.
+Proof. by move=> mf x y /eqP; rewrite eqr_le !mf -eqr_le=> /eqP. Qed.
+
+Lemma nmono_inj : {mono f : x y /~ x <= y} -> injective f.
+Proof. by move=> mf x y /eqP; rewrite eqr_le !mf -eqr_le=> /eqP. Qed.
+
+Lemma lerW_mono : {mono f : x y / x <= y} -> {mono f : x y / x < y}.
+Proof.
+by move=> mf x y /=; rewrite !ltr_neqAle mf inj_eq //; apply: mono_inj.
+Qed.
+
+Lemma lerW_nmono : {mono f : x y /~ x <= y} -> {mono f : x y /~ x < y}.
+Proof.
+by move=> mf x y /=; rewrite !ltr_neqAle mf eq_sym inj_eq //; apply: nmono_inj.
+Qed.
+
+(* Monotony in D D' *)
+Lemma ltrW_homo_in :
+ {in D & D', {homo f : x y / x < y}} -> {in D & D', {homo f : x y / x <= y}}.
+Proof.
+by move=> mf x y hx hy /=; rewrite ler_eqVlt => /orP[/eqP->|/mf/ltrW] //; apply.
+Qed.
+
+Lemma ltrW_nhomo_in :
+ {in D & D', {homo f : x y /~ x < y}} -> {in D & D', {homo f : x y /~ x <= y}}.
+Proof.
+by move=> mf x y hx hy /=; rewrite ler_eqVlt => /orP[/eqP->|/mf/ltrW] //; apply.
+Qed.
+
+Lemma homo_inj_in_lt :
+ {in D & D', injective f} -> {in D & D', {homo f : x y / x <= y}} ->
+ {in D & D', {homo f : x y / x < y}}.
+Proof.
+move=> fI mf x y hx hy /= hxy; rewrite ltr_neqAle; apply/andP; split.
+ by apply: contraTN hxy => /eqP /fI -> //; rewrite ltrr.
+by rewrite mf // (ltr_eqF, ltrW).
+Qed.
+
+Lemma nhomo_inj_in_lt :
+ {in D & D', injective f} -> {in D & D', {homo f : x y /~ x <= y}} ->
+ {in D & D', {homo f : x y /~ x < y}}.
+Proof.
+move=> fI mf x y hx hy /= hxy; rewrite ltr_neqAle; apply/andP; split.
+ by apply: contraTN hxy => /eqP /fI -> //; rewrite ltrr.
+by rewrite mf // (gtr_eqF, ltrW).
+Qed.
+
+Lemma mono_inj_in : {in D &, {mono f : x y / x <= y}} -> {in D &, injective f}.
+Proof.
+by move=> mf x y hx hy /= /eqP; rewrite eqr_le !mf // -eqr_le => /eqP.
+Qed.
+
+Lemma nmono_inj_in :
+ {in D &, {mono f : x y /~ x <= y}} -> {in D &, injective f}.
+Proof.
+by move=> mf x y hx hy /= /eqP; rewrite eqr_le !mf // -eqr_le => /eqP.
+Qed.
+
+Lemma lerW_mono_in :
+ {in D &, {mono f : x y / x <= y}} -> {in D &, {mono f : x y / x < y}}.
+Proof.
+move=> mf x y hx hy /=; rewrite !ltr_neqAle mf // (@inj_in_eq _ _ D) //.
+exact: mono_inj_in.
+Qed.
+
+Lemma lerW_nmono_in :
+ {in D &, {mono f : x y /~ x <= y}} -> {in D &, {mono f : x y /~ x < y}}.
+Proof.
+move=> mf x y hx hy /=; rewrite !ltr_neqAle mf // eq_sym (@inj_in_eq _ _ D) //.
+exact: nmono_inj_in.
+Qed.
+
+End AcrossTypes.
+
+Section NatToR.
+
+Variable (f : nat -> R).
+
+Lemma ltn_ltrW_homo :
+ {homo f : m n / (m < n)%N >-> m < n} ->
+ {homo f : m n / (m <= n)%N >-> m <= n}.
+Proof. by move=> mf m n /=; rewrite leq_eqVlt => /orP[/eqP->|/mf/ltrW //]. Qed.
+
+Lemma ltn_ltrW_nhomo :
+ {homo f : m n / (n < m)%N >-> m < n} ->
+ {homo f : m n / (n <= m)%N >-> m <= n}.
+Proof. by move=> mf m n /=; rewrite leq_eqVlt => /orP[/eqP->|/mf/ltrW//]. Qed.
+
+Lemma homo_inj_ltn_lt :
+ injective f -> {homo f : m n / (m <= n)%N >-> m <= n} ->
+ {homo f : m n / (m < n)%N >-> m < n}.
+Proof.
+move=> fI mf m n /= hmn.
+by rewrite ltr_neqAle (inj_eq fI) mf ?neq_ltn ?hmn ?orbT // ltnW.
+Qed.
+
+Lemma nhomo_inj_ltn_lt :
+ injective f -> {homo f : m n / (n <= m)%N >-> m <= n} ->
+ {homo f : m n / (n < m)%N >-> m < n}.
+Proof.
+move=> fI mf m n /= hmn; rewrite ltr_def (inj_eq fI).
+by rewrite mf ?neq_ltn ?hmn // ltnW.
+Qed.
+
+Lemma leq_mono_inj : {mono f : m n / (m <= n)%N >-> m <= n} -> injective f.
+Proof. by move=> mf m n /eqP; rewrite eqr_le !mf -eqn_leq => /eqP. Qed.
+
+Lemma leq_nmono_inj : {mono f : m n / (n <= m)%N >-> m <= n} -> injective f.
+Proof. by move=> mf m n /eqP; rewrite eqr_le !mf -eqn_leq => /eqP. Qed.
+
+Lemma leq_lerW_mono :
+ {mono f : m n / (m <= n)%N >-> m <= n} ->
+ {mono f : m n / (m < n)%N >-> m < n}.
+Proof.
+move=> mf m n /=; rewrite !ltr_neqAle mf inj_eq ?ltn_neqAle 1?eq_sym //.
+exact: leq_mono_inj.
+Qed.
+
+Lemma leq_lerW_nmono :
+ {mono f : m n / (n <= m)%N >-> m <= n} ->
+ {mono f : m n / (n < m)%N >-> m < n}.
+Proof.
+move=> mf x y /=; rewrite ltr_neqAle mf eq_sym inj_eq ?ltn_neqAle 1?eq_sym //.
+exact: leq_nmono_inj.
+Qed.
+
+Lemma homo_leq_mono :
+ {homo f : m n / (m < n)%N >-> m < n} ->
+ {mono f : m n / (m <= n)%N >-> m <= n}.
+Proof.
+move=> mf m n /=; case: leqP; last by move=> /mf /ltr_geF.
+by rewrite leq_eqVlt => /orP[/eqP->|/mf/ltrW //]; rewrite lerr.
+Qed.
+
+Lemma nhomo_leq_mono :
+ {homo f : m n / (n < m)%N >-> m < n} ->
+ {mono f : m n / (n <= m)%N >-> m <= n}.
+Proof.
+move=> mf m n /=; case: leqP; last by move=> /mf /ltr_geF.
+by rewrite leq_eqVlt => /orP[/eqP->|/mf/ltrW //]; rewrite lerr.
+Qed.
+
+End NatToR.
+
+End NumIntegralDomainMonotonyTheory.
+
+Section NumDomainOperationTheory.
+
+Variable R : numDomainType.
+Implicit Types x y z t : R.
+
+(* Comparision and opposite. *)
+
+Lemma ler_opp2 : {mono -%R : x y /~ x <= y :> R}.
+Proof. by move=> x y /=; rewrite -subr_ge0 opprK addrC subr_ge0. Qed.
+Hint Resolve ler_opp2.
+Lemma ltr_opp2 : {mono -%R : x y /~ x < y :> R}.
+Proof. by move=> x y /=; rewrite lerW_nmono. Qed.
+Hint Resolve ltr_opp2.
+Definition lter_opp2 := (ler_opp2, ltr_opp2).
+
+Lemma ler_oppr x y : (x <= - y) = (y <= - x).
+Proof. by rewrite (monoRL (@opprK _) ler_opp2). Qed.
+
+Lemma ltr_oppr x y : (x < - y) = (y < - x).
+Proof. by rewrite (monoRL (@opprK _) (lerW_nmono _)). Qed.
+
+Definition lter_oppr := (ler_oppr, ltr_oppr).
+
+Lemma ler_oppl x y : (- x <= y) = (- y <= x).
+Proof. by rewrite (monoLR (@opprK _) ler_opp2). Qed.
+
+Lemma ltr_oppl x y : (- x < y) = (- y < x).
+Proof. by rewrite (monoLR (@opprK _) (lerW_nmono _)). Qed.
+
+Definition lter_oppl := (ler_oppl, ltr_oppl).
+
+Lemma oppr_ge0 x : (0 <= - x) = (x <= 0).
+Proof. by rewrite lter_oppr oppr0. Qed.
+
+Lemma oppr_gt0 x : (0 < - x) = (x < 0).
+Proof. by rewrite lter_oppr oppr0. Qed.
+
+Definition oppr_gte0 := (oppr_ge0, oppr_gt0).
+
+Lemma oppr_le0 x : (- x <= 0) = (0 <= x).
+Proof. by rewrite lter_oppl oppr0. Qed.
+
+Lemma oppr_lt0 x : (- x < 0) = (0 < x).
+Proof. by rewrite lter_oppl oppr0. Qed.
+
+Definition oppr_lte0 := (oppr_le0, oppr_lt0).
+Definition oppr_cp0 := (oppr_gte0, oppr_lte0).
+Definition lter_oppE := (oppr_cp0, lter_opp2).
+
+Lemma ge0_cp x : 0 <= x -> (- x <= 0) * (- x <= x).
+Proof. by move=> hx; rewrite oppr_cp0 hx (@ler_trans _ 0) ?oppr_cp0. Qed.
+
+Lemma gt0_cp x : 0 < x ->
+ (0 <= x) * (- x <= 0) * (- x <= x) * (- x < 0) * (- x < x).
+Proof.
+move=> hx; move: (ltrW hx) => hx'; rewrite !ge0_cp hx' //.
+by rewrite oppr_cp0 hx // (@ltr_trans _ 0) ?oppr_cp0.
+Qed.
+
+Lemma le0_cp x : x <= 0 -> (0 <= - x) * (x <= - x).
+Proof. by move=> hx; rewrite oppr_cp0 hx (@ler_trans _ 0) ?oppr_cp0. Qed.
+
+Lemma lt0_cp x :
+ x < 0 -> (x <= 0) * (0 <= - x) * (x <= - x) * (0 < - x) * (x < - x).
+Proof.
+move=> hx; move: (ltrW hx) => hx'; rewrite !le0_cp // hx'.
+by rewrite oppr_cp0 hx // (@ltr_trans _ 0) ?oppr_cp0.
+Qed.
+
+(* Properties of the real subset. *)
+
+Lemma ger0_real x : 0 <= x -> x \is real.
+Proof. by rewrite realE => ->. Qed.
+
+Lemma ler0_real x : x <= 0 -> x \is real.
+Proof. by rewrite realE orbC => ->. Qed.
+
+Lemma gtr0_real x : 0 < x -> x \is real.
+Proof. by move=> /ltrW/ger0_real. Qed.
+
+Lemma ltr0_real x : x < 0 -> x \is real.
+Proof. by move=> /ltrW/ler0_real. Qed.
+
+Lemma real0 : 0 \is @real R. Proof. by rewrite ger0_real. Qed.
+Hint Resolve real0.
+
+Lemma real1 : 1 \is @real R. Proof. by rewrite ger0_real. Qed.
+Hint Resolve real1.
+
+Lemma realn n : n%:R \is @real R. Proof. by rewrite ger0_real. Qed.
+
+Lemma ler_leVge x y : x <= 0 -> y <= 0 -> (x <= y) || (y <= x).
+Proof. by rewrite -!oppr_ge0 => /(ger_leVge _) h /h; rewrite !ler_opp2. Qed.
+
+Lemma real_leVge x y : x \is real -> y \is real -> (x <= y) || (y <= x).
+Proof.
+rewrite !realE; have [x_ge0 _|x_nge0 /= x_le0] := boolP (_ <= _); last first.
+ by have [/(ler_trans x_le0)->|_ /(ler_leVge x_le0) //] := boolP (0 <= _).
+by have [/(ger_leVge x_ge0)|_ /ler_trans->] := boolP (0 <= _); rewrite ?orbT.
+Qed.
+
+Lemma realB : {in real &, forall x y, x - y \is real}.
+Proof. exact: rpredB. Qed.
+
+Lemma realN : {mono (@GRing.opp R) : x / x \is real}.
+Proof. exact: rpredN. Qed.
+
+(* :TODO: add a rpredBC in ssralg *)
+Lemma realBC x y : (x - y \is real) = (y - x \is real).
+Proof. by rewrite -realN opprB. Qed.
+
+Lemma realD : {in real &, forall x y, x + y \is real}.
+Proof. exact: rpredD. Qed.
+
+(* dichotomy and trichotomy *)
+
+CoInductive ler_xor_gt (x y : R) : R -> R -> bool -> bool -> Set :=
+ | LerNotGt of x <= y : ler_xor_gt x y (y - x) (y - x) true false
+ | GtrNotLe of y < x : ler_xor_gt x y (x - y) (x - y) false true.
+
+CoInductive ltr_xor_ge (x y : R) : R -> R -> bool -> bool -> Set :=
+ | LtrNotGe of x < y : ltr_xor_ge x y (y - x) (y - x) false true
+ | GerNotLt of y <= x : ltr_xor_ge x y (x - y) (x - y) true false.
+
+CoInductive comparer x y : R -> R ->
+ bool -> bool -> bool -> bool -> bool -> bool -> Set :=
+ | ComparerLt of x < y : comparer x y (y - x) (y - x)
+ false false true false true false
+ | ComparerGt of x > y : comparer x y (x - y) (x - y)
+ false false false true false true
+ | ComparerEq of x = y : comparer x y 0 0
+ true true true true false false.
+
+Lemma real_lerP x y :
+ x \is real -> y \is real ->
+ ler_xor_gt x y `|x - y| `|y - x| (x <= y) (y < x).
+Proof.
+move=> xR /(real_leVge xR); have [le_xy _|Nle_xy /= le_yx] := boolP (_ <= _).
+ have [/(ler_lt_trans le_xy)|] := boolP (_ < _); first by rewrite ltrr.
+ by rewrite ler0_norm ?ger0_norm ?subr_cp0 ?opprB //; constructor.
+have [lt_yx|] := boolP (_ < _).
+ by rewrite ger0_norm ?ler0_norm ?subr_cp0 ?opprB //; constructor.
+by rewrite ltr_def le_yx andbT negbK=> /eqP exy; rewrite exy lerr in Nle_xy.
+Qed.
+
+Lemma real_ltrP x y :
+ x \is real -> y \is real ->
+ ltr_xor_ge x y `|x - y| `|y - x| (y <= x) (x < y).
+Proof. by move=> xR yR; case: real_lerP=> //; constructor. Qed.
+
+Lemma real_ltrNge : {in real &, forall x y, (x < y) = ~~ (y <= x)}.
+Proof. by move=> x y xR yR /=; case: real_lerP. Qed.
+
+Lemma real_lerNgt : {in real &, forall x y, (x <= y) = ~~ (y < x)}.
+Proof. by move=> x y xR yR /=; case: real_lerP. Qed.
+
+Lemma real_ltrgtP x y :
+ x \is real -> y \is real ->
+ comparer x y `|x - y| `|y - x|
+ (y == x) (x == y) (x <= y) (y <= x) (x < y) (x > y).
+Proof.
+move=> xR yR; case: real_lerP => // [le_yx|lt_xy]; last first.
+ by rewrite gtr_eqF // ltr_eqF // ler_gtF ?ltrW //; constructor.
+case: real_lerP => // [le_xy|lt_yx]; last first.
+ by rewrite ltr_eqF // gtr_eqF //; constructor.
+have /eqP ->: x == y by rewrite eqr_le le_yx le_xy.
+by rewrite subrr eqxx; constructor.
+Qed.
+
+CoInductive ger0_xor_lt0 (x : R) : R -> bool -> bool -> Set :=
+ | Ger0NotLt0 of 0 <= x : ger0_xor_lt0 x x false true
+ | Ltr0NotGe0 of x < 0 : ger0_xor_lt0 x (- x) true false.
+
+CoInductive ler0_xor_gt0 (x : R) : R -> bool -> bool -> Set :=
+ | Ler0NotLe0 of x <= 0 : ler0_xor_gt0 x (- x) false true
+ | Gtr0NotGt0 of 0 < x : ler0_xor_gt0 x x true false.
+
+CoInductive comparer0 x :
+ R -> bool -> bool -> bool -> bool -> bool -> bool -> Set :=
+ | ComparerGt0 of 0 < x : comparer0 x x false false false true false true
+ | ComparerLt0 of x < 0 : comparer0 x (- x) false false true false true false
+ | ComparerEq0 of x = 0 : comparer0 x 0 true true true true false false.
+
+Lemma real_ger0P x : x \is real -> ger0_xor_lt0 x `|x| (x < 0) (0 <= x).
+Proof.
+move=> hx; rewrite -{2}[x]subr0; case: real_ltrP;
+by rewrite ?subr0 ?sub0r //; constructor.
+Qed.
+
+Lemma real_ler0P x : x \is real -> ler0_xor_gt0 x `|x| (0 < x) (x <= 0).
+Proof.
+move=> hx; rewrite -{2}[x]subr0; case: real_ltrP;
+by rewrite ?subr0 ?sub0r //; constructor.
+Qed.
+
+Lemma real_ltrgt0P x :
+ x \is real ->
+ comparer0 x `|x| (0 == x) (x == 0) (x <= 0) (0 <= x) (x < 0) (x > 0).
+Proof.
+move=> hx; rewrite -{2}[x]subr0; case: real_ltrgtP;
+by rewrite ?subr0 ?sub0r //; constructor.
+Qed.
+
+Lemma real_neqr_lt : {in real &, forall x y, (x != y) = (x < y) || (y < x)}.
+Proof. by move=> * /=; case: real_ltrgtP. Qed.
+
+Lemma ler_sub_real x y : x <= y -> y - x \is real.
+Proof. by move=> le_xy; rewrite ger0_real // subr_ge0. Qed.
+
+Lemma ger_sub_real x y : x <= y -> x - y \is real.
+Proof. by move=> le_xy; rewrite ler0_real // subr_le0. Qed.
+
+Lemma ler_real y x : x <= y -> (x \is real) = (y \is real).
+Proof. by move=> le_xy; rewrite -(addrNK x y) rpredDl ?ler_sub_real. Qed.
+
+Lemma ger_real x y : y <= x -> (x \is real) = (y \is real).
+Proof. by move=> le_yx; rewrite -(ler_real le_yx). Qed.
+
+Lemma ger1_real x : 1 <= x -> x \is real. Proof. by move=> /ger_real->. Qed.
+Lemma ler1_real x : x <= 1 -> x \is real. Proof. by move=> /ler_real->. Qed.
+
+Lemma Nreal_leF x y : y \is real -> x \notin real -> (x <= y) = false.
+Proof. by move=> yR; apply: contraNF=> /ler_real->. Qed.
+
+Lemma Nreal_geF x y : y \is real -> x \notin real -> (y <= x) = false.
+Proof. by move=> yR; apply: contraNF=> /ger_real->. Qed.
+
+Lemma Nreal_ltF x y : y \is real -> x \notin real -> (x < y) = false.
+Proof. by move=> yR xNR; rewrite ltr_def Nreal_leF ?andbF. Qed.
+
+Lemma Nreal_gtF x y : y \is real -> x \notin real -> (y < x) = false.
+Proof. by move=> yR xNR; rewrite ltr_def Nreal_geF ?andbF. Qed.
+
+(* real wlog *)
+
+Lemma real_wlog_ler P :
+ (forall a b, P b a -> P a b) -> (forall a b, a <= b -> P a b) ->
+ forall a b : R, a \is real -> b \is real -> P a b.
+Proof.
+move=> sP hP a b ha hb; wlog: a b ha hb / a <= b => [hwlog|]; last exact: hP.
+by case: (real_lerP ha hb)=> [/hP //|/ltrW hba]; apply: sP; apply: hP.
+Qed.
+
+Lemma real_wlog_ltr P :
+ (forall a, P a a) -> (forall a b, (P b a -> P a b)) ->
+ (forall a b, a < b -> P a b) ->
+ forall a b : R, a \is real -> b \is real -> P a b.
+Proof.
+move=> rP sP hP; apply: real_wlog_ler=> // a b.
+by rewrite ler_eqVlt; case: (altP (_ =P _))=> [->|] //= _ lab; apply: hP.
+Qed.
+
+(* Monotony of addition *)
+Lemma ler_add2l x : {mono +%R x : y z / y <= z}.
+Proof.
+by move=> y z /=; rewrite -subr_ge0 opprD addrAC addNKr addrC subr_ge0.
+Qed.
+
+Lemma ler_add2r x : {mono +%R^~ x : y z / y <= z}.
+Proof. by move=> y z /=; rewrite ![_ + x]addrC ler_add2l. Qed.
+
+Lemma ltr_add2r z x y : (x + z < y + z) = (x < y).
+Proof. by rewrite (lerW_mono (ler_add2r _)). Qed.
+
+Lemma ltr_add2l z x y : (z + x < z + y) = (x < y).
+Proof. by rewrite (lerW_mono (ler_add2l _)). Qed.
+
+Definition ler_add2 := (ler_add2l, ler_add2r).
+Definition ltr_add2 := (ltr_add2l, ltr_add2r).
+Definition lter_add2 := (ler_add2, ltr_add2).
+
+(* Addition, subtraction and transitivity *)
+Lemma ler_add x y z t : x <= y -> z <= t -> x + z <= y + t.
+Proof. by move=> lxy lzt; rewrite (@ler_trans _ (y + z)) ?lter_add2. Qed.
+
+Lemma ler_lt_add x y z t : x <= y -> z < t -> x + z < y + t.
+Proof. by move=> lxy lzt; rewrite (@ler_lt_trans _ (y + z)) ?lter_add2. Qed.
+
+Lemma ltr_le_add x y z t : x < y -> z <= t -> x + z < y + t.
+Proof. by move=> lxy lzt; rewrite (@ltr_le_trans _ (y + z)) ?lter_add2. Qed.
+
+Lemma ltr_add x y z t : x < y -> z < t -> x + z < y + t.
+Proof. by move=> lxy lzt; rewrite ltr_le_add // ltrW. Qed.
+
+Lemma ler_sub x y z t : x <= y -> t <= z -> x - z <= y - t.
+Proof. by move=> lxy ltz; rewrite ler_add // lter_opp2. Qed.
+
+Lemma ler_lt_sub x y z t : x <= y -> t < z -> x - z < y - t.
+Proof. by move=> lxy lzt; rewrite ler_lt_add // lter_opp2. Qed.
+
+Lemma ltr_le_sub x y z t : x < y -> t <= z -> x - z < y - t.
+Proof. by move=> lxy lzt; rewrite ltr_le_add // lter_opp2. Qed.
+
+Lemma ltr_sub x y z t : x < y -> t < z -> x - z < y - t.
+Proof. by move=> lxy lzt; rewrite ltr_add // lter_opp2. Qed.
+
+Lemma ler_subl_addr x y z : (x - y <= z) = (x <= z + y).
+Proof. by rewrite (monoLR (addrK _) (ler_add2r _)). Qed.
+
+Lemma ltr_subl_addr x y z : (x - y < z) = (x < z + y).
+Proof. by rewrite (monoLR (addrK _) (ltr_add2r _)). Qed.
+
+Lemma ler_subr_addr x y z : (x <= y - z) = (x + z <= y).
+Proof. by rewrite (monoLR (addrNK _) (ler_add2r _)). Qed.
+
+Lemma ltr_subr_addr x y z : (x < y - z) = (x + z < y).
+Proof. by rewrite (monoLR (addrNK _) (ltr_add2r _)). Qed.
+
+Definition ler_sub_addr := (ler_subl_addr, ler_subr_addr).
+Definition ltr_sub_addr := (ltr_subl_addr, ltr_subr_addr).
+Definition lter_sub_addr := (ler_sub_addr, ltr_sub_addr).
+
+Lemma ler_subl_addl x y z : (x - y <= z) = (x <= y + z).
+Proof. by rewrite lter_sub_addr addrC. Qed.
+
+Lemma ltr_subl_addl x y z : (x - y < z) = (x < y + z).
+Proof. by rewrite lter_sub_addr addrC. Qed.
+
+Lemma ler_subr_addl x y z : (x <= y - z) = (z + x <= y).
+Proof. by rewrite lter_sub_addr addrC. Qed.
+
+Lemma ltr_subr_addl x y z : (x < y - z) = (z + x < y).
+Proof. by rewrite lter_sub_addr addrC. Qed.
+
+Definition ler_sub_addl := (ler_subl_addl, ler_subr_addl).
+Definition ltr_sub_addl := (ltr_subl_addl, ltr_subr_addl).
+Definition lter_sub_addl := (ler_sub_addl, ltr_sub_addl).
+
+Lemma ler_addl x y : (x <= x + y) = (0 <= y).
+Proof. by rewrite -{1}[x]addr0 lter_add2. Qed.
+
+Lemma ltr_addl x y : (x < x + y) = (0 < y).
+Proof. by rewrite -{1}[x]addr0 lter_add2. Qed.
+
+Lemma ler_addr x y : (x <= y + x) = (0 <= y).
+Proof. by rewrite -{1}[x]add0r lter_add2. Qed.
+
+Lemma ltr_addr x y : (x < y + x) = (0 < y).
+Proof. by rewrite -{1}[x]add0r lter_add2. Qed.
+
+Lemma ger_addl x y : (x + y <= x) = (y <= 0).
+Proof. by rewrite -{2}[x]addr0 lter_add2. Qed.
+
+Lemma gtr_addl x y : (x + y < x) = (y < 0).
+Proof. by rewrite -{2}[x]addr0 lter_add2. Qed.
+
+Lemma ger_addr x y : (y + x <= x) = (y <= 0).
+Proof. by rewrite -{2}[x]add0r lter_add2. Qed.
+
+Lemma gtr_addr x y : (y + x < x) = (y < 0).
+Proof. by rewrite -{2}[x]add0r lter_add2. Qed.
+
+Definition cpr_add := (ler_addl, ler_addr, ger_addl, ger_addl,
+ ltr_addl, ltr_addr, gtr_addl, gtr_addl).
+
+(* Addition with left member knwon to be positive/negative *)
+Lemma ler_paddl y x z : 0 <= x -> y <= z -> y <= x + z.
+Proof. by move=> *; rewrite -[y]add0r ler_add. Qed.
+
+Lemma ltr_paddl y x z : 0 <= x -> y < z -> y < x + z.
+Proof. by move=> *; rewrite -[y]add0r ler_lt_add. Qed.
+
+Lemma ltr_spaddl y x z : 0 < x -> y <= z -> y < x + z.
+Proof. by move=> *; rewrite -[y]add0r ltr_le_add. Qed.
+
+Lemma ltr_spsaddl y x z : 0 < x -> y < z -> y < x + z.
+Proof. by move=> *; rewrite -[y]add0r ltr_add. Qed.
+
+Lemma ler_naddl y x z : x <= 0 -> y <= z -> x + y <= z.
+Proof. by move=> *; rewrite -[z]add0r ler_add. Qed.
+
+Lemma ltr_naddl y x z : x <= 0 -> y < z -> x + y < z.
+Proof. by move=> *; rewrite -[z]add0r ler_lt_add. Qed.
+
+Lemma ltr_snaddl y x z : x < 0 -> y <= z -> x + y < z.
+Proof. by move=> *; rewrite -[z]add0r ltr_le_add. Qed.
+
+Lemma ltr_snsaddl y x z : x < 0 -> y < z -> x + y < z.
+Proof. by move=> *; rewrite -[z]add0r ltr_add. Qed.
+
+(* Addition with right member we know positive/negative *)
+Lemma ler_paddr y x z : 0 <= x -> y <= z -> y <= z + x.
+Proof. by move=> *; rewrite [_ + x]addrC ler_paddl. Qed.
+
+Lemma ltr_paddr y x z : 0 <= x -> y < z -> y < z + x.
+Proof. by move=> *; rewrite [_ + x]addrC ltr_paddl. Qed.
+
+Lemma ltr_spaddr y x z : 0 < x -> y <= z -> y < z + x.
+Proof. by move=> *; rewrite [_ + x]addrC ltr_spaddl. Qed.
+
+Lemma ltr_spsaddr y x z : 0 < x -> y < z -> y < z + x.
+Proof. by move=> *; rewrite [_ + x]addrC ltr_spsaddl. Qed.
+
+Lemma ler_naddr y x z : x <= 0 -> y <= z -> y + x <= z.
+Proof. by move=> *; rewrite [_ + x]addrC ler_naddl. Qed.
+
+Lemma ltr_naddr y x z : x <= 0 -> y < z -> y + x < z.
+Proof. by move=> *; rewrite [_ + x]addrC ltr_naddl. Qed.
+
+Lemma ltr_snaddr y x z : x < 0 -> y <= z -> y + x < z.
+Proof. by move=> *; rewrite [_ + x]addrC ltr_snaddl. Qed.
+
+Lemma ltr_snsaddr y x z : x < 0 -> y < z -> y + x < z.
+Proof. by move=> *; rewrite [_ + x]addrC ltr_snsaddl. Qed.
+
+(* x and y have the same sign and their sum is null *)
+Lemma paddr_eq0 (x y : R) :
+ 0 <= x -> 0 <= y -> (x + y == 0) = (x == 0) && (y == 0).
+Proof.
+rewrite le0r; case/orP=> [/eqP->|hx]; first by rewrite add0r eqxx.
+by rewrite (gtr_eqF hx) /= => hy; rewrite gtr_eqF // ltr_spaddl.
+Qed.
+
+Lemma naddr_eq0 (x y : R) :
+ x <= 0 -> y <= 0 -> (x + y == 0) = (x == 0) && (y == 0).
+Proof.
+by move=> lex0 ley0; rewrite -oppr_eq0 opprD paddr_eq0 ?oppr_cp0 // !oppr_eq0.
+Qed.
+
+Lemma addr_ss_eq0 (x y : R) :
+ (0 <= x) && (0 <= y) || (x <= 0) && (y <= 0) ->
+ (x + y == 0) = (x == 0) && (y == 0).
+Proof. by case/orP=> /andP []; [apply: paddr_eq0 | apply: naddr_eq0]. Qed.
+
+(* big sum and ler *)
+Lemma sumr_ge0 I (r : seq I) (P : pred I) (F : I -> R) :
+ (forall i, P i -> (0 <= F i)) -> 0 <= \sum_(i <- r | P i) (F i).
+Proof. exact: (big_ind _ _ (@ler_paddl 0)). Qed.
+
+Lemma ler_sum I (r : seq I) (P : pred I) (F G : I -> R) :
+ (forall i, P i -> F i <= G i) ->
+ \sum_(i <- r | P i) F i <= \sum_(i <- r | P i) G i.
+Proof. exact: (big_ind2 _ (lerr _) ler_add). Qed.
+
+Lemma psumr_eq0 (I : eqType) (r : seq I) (P : pred I) (F : I -> R) :
+ (forall i, P i -> 0 <= F i) ->
+ (\sum_(i <- r | P i) (F i) == 0) = (all (fun i => (P i) ==> (F i == 0)) r).
+Proof.
+elim: r=> [|a r ihr hr] /=; rewrite (big_nil, big_cons); first by rewrite eqxx.
+by case: ifP=> pa /=; rewrite ?paddr_eq0 ?ihr ?hr // sumr_ge0.
+Qed.
+
+(* :TODO: Cyril : See which form to keep *)
+Lemma psumr_eq0P (I : finType) (P : pred I) (F : I -> R) :
+ (forall i, P i -> 0 <= F i) -> \sum_(i | P i) F i = 0 ->
+ (forall i, P i -> F i = 0).
+Proof.
+move=> F_ge0 /eqP; rewrite psumr_eq0 // -big_all big_andE => /forallP hF i Pi.
+by move: (hF i); rewrite implyTb Pi /= => /eqP.
+Qed.
+
+(* mulr and ler/ltr *)
+
+Lemma ler_pmul2l x : 0 < x -> {mono *%R x : x y / x <= y}.
+Proof.
+by move=> x_gt0 y z /=; rewrite -subr_ge0 -mulrBr pmulr_rge0 // subr_ge0.
+Qed.
+
+Lemma ltr_pmul2l x : 0 < x -> {mono *%R x : x y / x < y}.
+Proof. by move=> x_gt0; apply: lerW_mono (ler_pmul2l _). Qed.
+
+Definition lter_pmul2l := (ler_pmul2l, ltr_pmul2l).
+
+Lemma ler_pmul2r x : 0 < x -> {mono *%R^~ x : x y / x <= y}.
+Proof. by move=> x_gt0 y z /=; rewrite ![_ * x]mulrC ler_pmul2l. Qed.
+
+Lemma ltr_pmul2r x : 0 < x -> {mono *%R^~ x : x y / x < y}.
+Proof. by move=> x_gt0; apply: lerW_mono (ler_pmul2r _). Qed.
+
+Definition lter_pmul2r := (ler_pmul2r, ltr_pmul2r).
+
+Lemma ler_nmul2l x : x < 0 -> {mono *%R x : x y /~ x <= y}.
+Proof.
+by move=> x_lt0 y z /=; rewrite -ler_opp2 -!mulNr ler_pmul2l ?oppr_gt0.
+Qed.
+
+Lemma ltr_nmul2l x : x < 0 -> {mono *%R x : x y /~ x < y}.
+Proof. by move=> x_lt0; apply: lerW_nmono (ler_nmul2l _). Qed.
+
+Definition lter_nmul2l := (ler_nmul2l, ltr_nmul2l).
+
+Lemma ler_nmul2r x : x < 0 -> {mono *%R^~ x : x y /~ x <= y}.
+Proof. by move=> x_lt0 y z /=; rewrite ![_ * x]mulrC ler_nmul2l. Qed.
+
+Lemma ltr_nmul2r x : x < 0 -> {mono *%R^~ x : x y /~ x < y}.
+Proof. by move=> x_lt0; apply: lerW_nmono (ler_nmul2r _). Qed.
+
+Definition lter_nmul2r := (ler_nmul2r, ltr_nmul2r).
+
+Lemma ler_wpmul2l x : 0 <= x -> {homo *%R x : y z / y <= z}.
+Proof.
+by rewrite le0r => /orP[/eqP-> y z | /ler_pmul2l/mono2W//]; rewrite !mul0r.
+Qed.
+
+Lemma ler_wpmul2r x : 0 <= x -> {homo *%R^~ x : y z / y <= z}.
+Proof. by move=> x_ge0 y z leyz; rewrite ![_ * x]mulrC ler_wpmul2l. Qed.
+
+Lemma ler_wnmul2l x : x <= 0 -> {homo *%R x : y z /~ y <= z}.
+Proof.
+by move=> x_le0 y z leyz; rewrite -![x * _]mulrNN ler_wpmul2l ?lter_oppE.
+Qed.
+
+Lemma ler_wnmul2r x : x <= 0 -> {homo *%R^~ x : y z /~ y <= z}.
+Proof.
+by move=> x_le0 y z leyz; rewrite -![_ * x]mulrNN ler_wpmul2r ?lter_oppE.
+Qed.
+
+(* Binary forms, for backchaining. *)
+
+Lemma ler_pmul x1 y1 x2 y2 :
+ 0 <= x1 -> 0 <= x2 -> x1 <= y1 -> x2 <= y2 -> x1 * x2 <= y1 * y2.
+Proof.
+move=> x1ge0 x2ge0 le_xy1 le_xy2; have y1ge0 := ler_trans x1ge0 le_xy1.
+exact: ler_trans (ler_wpmul2r x2ge0 le_xy1) (ler_wpmul2l y1ge0 le_xy2).
+Qed.
+
+Lemma ltr_pmul x1 y1 x2 y2 :
+ 0 <= x1 -> 0 <= x2 -> x1 < y1 -> x2 < y2 -> x1 * x2 < y1 * y2.
+Proof.
+move=> x1ge0 x2ge0 lt_xy1 lt_xy2; have y1gt0 := ler_lt_trans x1ge0 lt_xy1.
+by rewrite (ler_lt_trans (ler_wpmul2r x2ge0 (ltrW lt_xy1))) ?ltr_pmul2l.
+Qed.
+
+(* complement for x *+ n and <= or < *)
+
+Lemma ler_pmuln2r n : (0 < n)%N -> {mono (@GRing.natmul R)^~ n : x y / x <= y}.
+Proof.
+by case: n => // n _ x y /=; rewrite -mulr_natl -[y *+ _]mulr_natl ler_pmul2l.
+Qed.
+
+Lemma ltr_pmuln2r n : (0 < n)%N -> {mono (@GRing.natmul R)^~ n : x y / x < y}.
+Proof. by move/ler_pmuln2r/lerW_mono. Qed.
+
+Lemma pmulrnI n : (0 < n)%N -> injective ((@GRing.natmul R)^~ n).
+Proof. by move/ler_pmuln2r/mono_inj. Qed.
+
+Lemma eqr_pmuln2r n : (0 < n)%N -> {mono (@GRing.natmul R)^~ n : x y / x == y}.
+Proof. by move/pmulrnI/inj_eq. Qed.
+
+Lemma pmulrn_lgt0 x n : (0 < n)%N -> (0 < x *+ n) = (0 < x).
+Proof. by move=> n_gt0; rewrite -(mul0rn _ n) ltr_pmuln2r // mul0rn. Qed.
+
+Lemma pmulrn_llt0 x n : (0 < n)%N -> (x *+ n < 0) = (x < 0).
+Proof. by move=> n_gt0; rewrite -(mul0rn _ n) ltr_pmuln2r // mul0rn. Qed.
+
+Lemma pmulrn_lge0 x n : (0 < n)%N -> (0 <= x *+ n) = (0 <= x).
+Proof. by move=> n_gt0; rewrite -(mul0rn _ n) ler_pmuln2r // mul0rn. Qed.
+
+Lemma pmulrn_lle0 x n : (0 < n)%N -> (x *+ n <= 0) = (x <= 0).
+Proof. by move=> n_gt0; rewrite -(mul0rn _ n) ler_pmuln2r // mul0rn. Qed.
+
+Lemma ltr_wmuln2r x y n : x < y -> (x *+ n < y *+ n) = (0 < n)%N.
+Proof. by move=> ltxy; case: n=> // n; rewrite ltr_pmuln2r. Qed.
+
+Lemma ltr_wpmuln2r n : (0 < n)%N -> {homo (@GRing.natmul R)^~ n : x y / x < y}.
+Proof. by move=> n_gt0 x y /= / ltr_wmuln2r ->. Qed.
+
+Lemma ler_wmuln2r n : {homo (@GRing.natmul R)^~ n : x y / x <= y}.
+Proof. by move=> x y hxy /=; case: n=> // n; rewrite ler_pmuln2r. Qed.
+
+Lemma mulrn_wge0 x n : 0 <= x -> 0 <= x *+ n.
+Proof. by move=> /(ler_wmuln2r n); rewrite mul0rn. Qed.
+
+Lemma mulrn_wle0 x n : x <= 0 -> x *+ n <= 0.
+Proof. by move=> /(ler_wmuln2r n); rewrite mul0rn. Qed.
+
+Lemma ler_muln2r n x y : (x *+ n <= y *+ n) = ((n == 0%N) || (x <= y)).
+Proof. by case: n => [|n]; rewrite ?lerr ?eqxx // ler_pmuln2r. Qed.
+
+Lemma ltr_muln2r n x y : (x *+ n < y *+ n) = ((0 < n)%N && (x < y)).
+Proof. by case: n => [|n]; rewrite ?lerr ?eqxx // ltr_pmuln2r. Qed.
+
+Lemma eqr_muln2r n x y : (x *+ n == y *+ n) = (n == 0)%N || (x == y).
+Proof. by rewrite !eqr_le !ler_muln2r -orb_andr. Qed.
+
+(* More characteristic zero properties. *)
+
+Lemma mulrn_eq0 x n : (x *+ n == 0) = ((n == 0)%N || (x == 0)).
+Proof. by rewrite -mulr_natl mulf_eq0 pnatr_eq0. Qed.
+
+Lemma mulrIn x : x != 0 -> injective (GRing.natmul x).
+Proof.
+move=> x_neq0 m n; without loss /subnK <-: m n / (n <= m)%N.
+ by move=> IH eq_xmn; case/orP: (leq_total m n) => /IH->.
+by move/eqP; rewrite mulrnDr -subr_eq0 addrK mulrn_eq0 => /predU1P[-> | /idPn].
+Qed.
+
+Lemma ler_wpmuln2l x :
+ 0 <= x -> {homo (@GRing.natmul R x) : m n / (m <= n)%N >-> m <= n}.
+Proof. by move=> xge0 m n /subnK <-; rewrite mulrnDr ler_paddl ?mulrn_wge0. Qed.
+
+Lemma ler_wnmuln2l x :
+ x <= 0 -> {homo (@GRing.natmul R x) : m n / (n <= m)%N >-> m <= n}.
+Proof.
+by move=> xle0 m n hmn /=; rewrite -ler_opp2 -!mulNrn ler_wpmuln2l // oppr_cp0.
+Qed.
+
+Lemma mulrn_wgt0 x n : 0 < x -> 0 < x *+ n = (0 < n)%N.
+Proof. by case: n => // n hx; rewrite pmulrn_lgt0. Qed.
+
+Lemma mulrn_wlt0 x n : x < 0 -> x *+ n < 0 = (0 < n)%N.
+Proof. by case: n => // n hx; rewrite pmulrn_llt0. Qed.
+
+Lemma ler_pmuln2l x :
+ 0 < x -> {mono (@GRing.natmul R x) : m n / (m <= n)%N >-> m <= n}.
+Proof.
+move=> x_gt0 m n /=; case: leqP => hmn; first by rewrite ler_wpmuln2l // ltrW.
+rewrite -(subnK (ltnW hmn)) mulrnDr ger_addr ltr_geF //.
+by rewrite mulrn_wgt0 // subn_gt0.
+Qed.
+
+Lemma ltr_pmuln2l x :
+ 0 < x -> {mono (@GRing.natmul R x) : m n / (m < n)%N >-> m < n}.
+Proof. by move=> x_gt0; apply: leq_lerW_mono (ler_pmuln2l _). Qed.
+
+Lemma ler_nmuln2l x :
+ x < 0 -> {mono (@GRing.natmul R x) : m n / (n <= m)%N >-> m <= n}.
+Proof.
+by move=> x_lt0 m n /=; rewrite -ler_opp2 -!mulNrn ler_pmuln2l // oppr_gt0.
+Qed.
+
+Lemma ltr_nmuln2l x :
+ x < 0 -> {mono (@GRing.natmul R x) : m n / (n < m)%N >-> m < n}.
+Proof. by move=> x_lt0; apply: leq_lerW_nmono (ler_nmuln2l _). Qed.
+
+Lemma ler_nat m n : (m%:R <= n%:R :> R) = (m <= n)%N.
+Proof. by rewrite ler_pmuln2l. Qed.
+
+Lemma ltr_nat m n : (m%:R < n%:R :> R) = (m < n)%N.
+Proof. by rewrite ltr_pmuln2l. Qed.
+
+Lemma eqr_nat m n : (m%:R == n%:R :> R) = (m == n)%N.
+Proof. by rewrite (inj_eq (mulrIn _)) ?oner_eq0. Qed.
+
+Lemma pnatr_eq1 n : (n%:R == 1 :> R) = (n == 1)%N.
+Proof. exact: eqr_nat 1%N. Qed.
+
+Lemma lern0 n : (n%:R <= 0 :> R) = (n == 0%N).
+Proof. by rewrite -[0]/0%:R ler_nat leqn0. Qed.
+
+Lemma ltrn0 n : (n%:R < 0 :> R) = false.
+Proof. by rewrite -[0]/0%:R ltr_nat ltn0. Qed.
+
+Lemma ler1n n : 1 <= n%:R :> R = (1 <= n)%N. Proof. by rewrite -ler_nat. Qed.
+Lemma ltr1n n : 1 < n%:R :> R = (1 < n)%N. Proof. by rewrite -ltr_nat. Qed.
+Lemma lern1 n : n%:R <= 1 :> R = (n <= 1)%N. Proof. by rewrite -ler_nat. Qed.
+Lemma ltrn1 n : n%:R < 1 :> R = (n < 1)%N. Proof. by rewrite -ltr_nat. Qed.
+
+Lemma ltrN10 : -1 < 0 :> R. Proof. by rewrite oppr_lt0. Qed.
+Lemma lerN10 : -1 <= 0 :> R. Proof. by rewrite oppr_le0. Qed.
+Lemma ltr10 : 1 < 0 :> R = false. Proof. by rewrite ler_gtF. Qed.
+Lemma ler10 : 1 <= 0 :> R = false. Proof. by rewrite ltr_geF. Qed.
+Lemma ltr0N1 : 0 < -1 :> R = false. Proof. by rewrite ler_gtF // lerN10. Qed.
+Lemma ler0N1 : 0 <= -1 :> R = false. Proof. by rewrite ltr_geF // ltrN10. Qed.
+
+Lemma pmulrn_rgt0 x n : 0 < x -> 0 < x *+ n = (0 < n)%N.
+Proof. by move=> x_gt0; rewrite -(mulr0n x) ltr_pmuln2l. Qed.
+
+Lemma pmulrn_rlt0 x n : 0 < x -> x *+ n < 0 = false.
+Proof. by move=> x_gt0; rewrite -(mulr0n x) ltr_pmuln2l. Qed.
+
+Lemma pmulrn_rge0 x n : 0 < x -> 0 <= x *+ n.
+Proof. by move=> x_gt0; rewrite -(mulr0n x) ler_pmuln2l. Qed.
+
+Lemma pmulrn_rle0 x n : 0 < x -> x *+ n <= 0 = (n == 0)%N.
+Proof. by move=> x_gt0; rewrite -(mulr0n x) ler_pmuln2l ?leqn0. Qed.
+
+Lemma nmulrn_rgt0 x n : x < 0 -> 0 < x *+ n = false.
+Proof. by move=> x_lt0; rewrite -(mulr0n x) ltr_nmuln2l. Qed.
+
+Lemma nmulrn_rge0 x n : x < 0 -> 0 <= x *+ n = (n == 0)%N.
+Proof. by move=> x_lt0; rewrite -(mulr0n x) ler_nmuln2l ?leqn0. Qed.
+
+Lemma nmulrn_rle0 x n : x < 0 -> x *+ n <= 0.
+Proof. by move=> x_lt0; rewrite -(mulr0n x) ler_nmuln2l. Qed.
+
+(* (x * y) compared to 0 *)
+(* Remark : pmulr_rgt0 and pmulr_rge0 are defined above *)
+
+(* x positive and y right *)
+Lemma pmulr_rlt0 x y : 0 < x -> (x * y < 0) = (y < 0).
+Proof. by move=> x_gt0; rewrite -oppr_gt0 -mulrN pmulr_rgt0 // oppr_gt0. Qed.
+
+Lemma pmulr_rle0 x y : 0 < x -> (x * y <= 0) = (y <= 0).
+Proof. by move=> x_gt0; rewrite -oppr_ge0 -mulrN pmulr_rge0 // oppr_ge0. Qed.
+
+(* x positive and y left *)
+Lemma pmulr_lgt0 x y : 0 < x -> (0 < y * x) = (0 < y).
+Proof. by move=> x_gt0; rewrite mulrC pmulr_rgt0. Qed.
+
+Lemma pmulr_lge0 x y : 0 < x -> (0 <= y * x) = (0 <= y).
+Proof. by move=> x_gt0; rewrite mulrC pmulr_rge0. Qed.
+
+Lemma pmulr_llt0 x y : 0 < x -> (y * x < 0) = (y < 0).
+Proof. by move=> x_gt0; rewrite mulrC pmulr_rlt0. Qed.
+
+Lemma pmulr_lle0 x y : 0 < x -> (y * x <= 0) = (y <= 0).
+Proof. by move=> x_gt0; rewrite mulrC pmulr_rle0. Qed.
+
+(* x negative and y right *)
+Lemma nmulr_rgt0 x y : x < 0 -> (0 < x * y) = (y < 0).
+Proof. by move=> x_lt0; rewrite -mulrNN pmulr_rgt0 lter_oppE. Qed.
+
+Lemma nmulr_rge0 x y : x < 0 -> (0 <= x * y) = (y <= 0).
+Proof. by move=> x_lt0; rewrite -mulrNN pmulr_rge0 lter_oppE. Qed.
+
+Lemma nmulr_rlt0 x y : x < 0 -> (x * y < 0) = (0 < y).
+Proof. by move=> x_lt0; rewrite -mulrNN pmulr_rlt0 lter_oppE. Qed.
+
+Lemma nmulr_rle0 x y : x < 0 -> (x * y <= 0) = (0 <= y).
+Proof. by move=> x_lt0; rewrite -mulrNN pmulr_rle0 lter_oppE. Qed.
+
+(* x negative and y left *)
+Lemma nmulr_lgt0 x y : x < 0 -> (0 < y * x) = (y < 0).
+Proof. by move=> x_lt0; rewrite mulrC nmulr_rgt0. Qed.
+
+Lemma nmulr_lge0 x y : x < 0 -> (0 <= y * x) = (y <= 0).
+Proof. by move=> x_lt0; rewrite mulrC nmulr_rge0. Qed.
+
+Lemma nmulr_llt0 x y : x < 0 -> (y * x < 0) = (0 < y).
+Proof. by move=> x_lt0; rewrite mulrC nmulr_rlt0. Qed.
+
+Lemma nmulr_lle0 x y : x < 0 -> (y * x <= 0) = (0 <= y).
+Proof. by move=> x_lt0; rewrite mulrC nmulr_rle0. Qed.
+
+(* weak and symmetric lemmas *)
+Lemma mulr_ge0 x y : 0 <= x -> 0 <= y -> 0 <= x * y.
+Proof. by move=> x_ge0 y_ge0; rewrite -(mulr0 x) ler_wpmul2l. Qed.
+
+Lemma mulr_le0 x y : x <= 0 -> y <= 0 -> 0 <= x * y.
+Proof. by move=> x_le0 y_le0; rewrite -(mulr0 x) ler_wnmul2l. Qed.
+
+Lemma mulr_ge0_le0 x y : 0 <= x -> y <= 0 -> x * y <= 0.
+Proof. by move=> x_le0 y_le0; rewrite -(mulr0 x) ler_wpmul2l. Qed.
+
+Lemma mulr_le0_ge0 x y : x <= 0 -> 0 <= y -> x * y <= 0.
+Proof. by move=> x_le0 y_le0; rewrite -(mulr0 x) ler_wnmul2l. Qed.
+
+(* mulr_gt0 with only one case *)
+
+Lemma mulr_gt0 x y : 0 < x -> 0 < y -> 0 < x * y.
+Proof. by move=> x_gt0 y_gt0; rewrite pmulr_rgt0. Qed.
+
+(* Iterated products *)
+
+Lemma prodr_ge0 I r (P : pred I) (E : I -> R) :
+ (forall i, P i -> 0 <= E i) -> 0 <= \prod_(i <- r | P i) E i.
+Proof. by move=> Ege0; rewrite -nnegrE rpred_prod. Qed.
+
+Lemma prodr_gt0 I r (P : pred I) (E : I -> R) :
+ (forall i, P i -> 0 < E i) -> 0 < \prod_(i <- r | P i) E i.
+Proof. by move=> Ege0; rewrite -posrE rpred_prod. Qed.
+
+Lemma ler_prod I r (P : pred I) (E1 E2 : I -> R) :
+ (forall i, P i -> 0 <= E1 i <= E2 i) ->
+ \prod_(i <- r | P i) E1 i <= \prod_(i <- r | P i) E2 i.
+Proof.
+move=> leE12; elim/(big_load (fun x => 0 <= x)): _.
+elim/big_rec2: _ => // i x2 x1 /leE12/andP[le0Ei leEi12] [x1ge0 le_x12].
+by rewrite mulr_ge0 // ler_pmul.
+Qed.
+
+Lemma ltr_prod (E1 E2 : nat -> R) (n m : nat) :
+ (m < n)%N -> (forall i, (m <= i < n)%N -> 0 <= E1 i < E2 i) ->
+ \prod_(m <= i < n) E1 i < \prod_(m <= i < n) E2 i.
+Proof.
+elim: n m => // n ihn m; rewrite ltnS leq_eqVlt; case/orP => [/eqP -> | ltnm hE].
+ by move/(_ n) => /andb_idr; rewrite !big_nat1 leqnn ltnSn /=; case/andP.
+rewrite big_nat_recr ?[X in _ < X]big_nat_recr ?(ltnW ltnm) //=.
+move/andb_idr: (hE n); rewrite leqnn ltnW //=; case/andP => h1n h12n.
+rewrite big_nat_cond [X in _ < X * _]big_nat_cond; apply: ltr_pmul => //=.
+- apply: prodr_ge0 => i; rewrite andbT; case/andP=> hm hn.
+ by move/andb_idr: (hE i); rewrite hm /= ltnS ltnW //=; case/andP.
+rewrite -!big_nat_cond; apply: ihn => // i /andP [hm hn]; apply: hE.
+by rewrite hm ltnW.
+Qed.
+
+(* real of mul *)
+
+Lemma realMr x y : x != 0 -> x \is real -> (x * y \is real) = (y \is real).
+Proof.
+move=> x_neq0 xR; case: real_ltrgtP x_neq0 => // hx _; rewrite !realE.
+ by rewrite nmulr_rge0 // nmulr_rle0 // orbC.
+by rewrite pmulr_rge0 // pmulr_rle0 // orbC.
+Qed.
+
+Lemma realrM x y : y != 0 -> y \is real -> (x * y \is real) = (x \is real).
+Proof. by move=> y_neq0 yR; rewrite mulrC realMr. Qed.
+
+Lemma realM : {in real &, forall x y, x * y \is real}.
+Proof. exact: rpredM. Qed.
+
+Lemma realrMn x n : (n != 0)%N -> (x *+ n \is real) = (x \is real).
+Proof. by move=> n_neq0; rewrite -mulr_natl realMr ?realn ?pnatr_eq0. Qed.
+
+(* ler/ltr and multiplication between a positive/negative *)
+
+Lemma ger_pmull x y : 0 < y -> (x * y <= y) = (x <= 1).
+Proof. by move=> hy; rewrite -{2}[y]mul1r ler_pmul2r. Qed.
+
+Lemma gtr_pmull x y : 0 < y -> (x * y < y) = (x < 1).
+Proof. by move=> hy; rewrite -{2}[y]mul1r ltr_pmul2r. Qed.
+
+Lemma ger_pmulr x y : 0 < y -> (y * x <= y) = (x <= 1).
+Proof. by move=> hy; rewrite -{2}[y]mulr1 ler_pmul2l. Qed.
+
+Lemma gtr_pmulr x y : 0 < y -> (y * x < y) = (x < 1).
+Proof. by move=> hy; rewrite -{2}[y]mulr1 ltr_pmul2l. Qed.
+
+Lemma ler_pmull x y : 0 < y -> (y <= x * y) = (1 <= x).
+Proof. by move=> hy; rewrite -{1}[y]mul1r ler_pmul2r. Qed.
+
+Lemma ltr_pmull x y : 0 < y -> (y < x * y) = (1 < x).
+Proof. by move=> hy; rewrite -{1}[y]mul1r ltr_pmul2r. Qed.
+
+Lemma ler_pmulr x y : 0 < y -> (y <= y * x) = (1 <= x).
+Proof. by move=> hy; rewrite -{1}[y]mulr1 ler_pmul2l. Qed.
+
+Lemma ltr_pmulr x y : 0 < y -> (y < y * x) = (1 < x).
+Proof. by move=> hy; rewrite -{1}[y]mulr1 ltr_pmul2l. Qed.
+
+Lemma ger_nmull x y : y < 0 -> (x * y <= y) = (1 <= x).
+Proof. by move=> hy; rewrite -{2}[y]mul1r ler_nmul2r. Qed.
+
+Lemma gtr_nmull x y : y < 0 -> (x * y < y) = (1 < x).
+Proof. by move=> hy; rewrite -{2}[y]mul1r ltr_nmul2r. Qed.
+
+Lemma ger_nmulr x y : y < 0 -> (y * x <= y) = (1 <= x).
+Proof. by move=> hy; rewrite -{2}[y]mulr1 ler_nmul2l. Qed.
+
+Lemma gtr_nmulr x y : y < 0 -> (y * x < y) = (1 < x).
+Proof. by move=> hy; rewrite -{2}[y]mulr1 ltr_nmul2l. Qed.
+
+Lemma ler_nmull x y : y < 0 -> (y <= x * y) = (x <= 1).
+Proof. by move=> hy; rewrite -{1}[y]mul1r ler_nmul2r. Qed.
+
+Lemma ltr_nmull x y : y < 0 -> (y < x * y) = (x < 1).
+Proof. by move=> hy; rewrite -{1}[y]mul1r ltr_nmul2r. Qed.
+
+Lemma ler_nmulr x y : y < 0 -> (y <= y * x) = (x <= 1).
+Proof. by move=> hy; rewrite -{1}[y]mulr1 ler_nmul2l. Qed.
+
+Lemma ltr_nmulr x y : y < 0 -> (y < y * x) = (x < 1).
+Proof. by move=> hy; rewrite -{1}[y]mulr1 ltr_nmul2l. Qed.
+
+(* ler/ltr and multiplication between a positive/negative
+ and a exterior (1 <= _) or interior (0 <= _ <= 1) *)
+
+Lemma ler_pemull x y : 0 <= y -> 1 <= x -> y <= x * y.
+Proof. by move=> hy hx; rewrite -{1}[y]mul1r ler_wpmul2r. Qed.
+
+Lemma ler_nemull x y : y <= 0 -> 1 <= x -> x * y <= y.
+Proof. by move=> hy hx; rewrite -{2}[y]mul1r ler_wnmul2r. Qed.
+
+Lemma ler_pemulr x y : 0 <= y -> 1 <= x -> y <= y * x.
+Proof. by move=> hy hx; rewrite -{1}[y]mulr1 ler_wpmul2l. Qed.
+
+Lemma ler_nemulr x y : y <= 0 -> 1 <= x -> y * x <= y.
+Proof. by move=> hy hx; rewrite -{2}[y]mulr1 ler_wnmul2l. Qed.
+
+Lemma ler_pimull x y : 0 <= y -> x <= 1 -> x * y <= y.
+Proof. by move=> hy hx; rewrite -{2}[y]mul1r ler_wpmul2r. Qed.
+
+Lemma ler_nimull x y : y <= 0 -> x <= 1 -> y <= x * y.
+Proof. by move=> hy hx; rewrite -{1}[y]mul1r ler_wnmul2r. Qed.
+
+Lemma ler_pimulr x y : 0 <= y -> x <= 1 -> y * x <= y.
+Proof. by move=> hy hx; rewrite -{2}[y]mulr1 ler_wpmul2l. Qed.
+
+Lemma ler_nimulr x y : y <= 0 -> x <= 1 -> y <= y * x.
+Proof. by move=> hx hy; rewrite -{1}[y]mulr1 ler_wnmul2l. Qed.
+
+Lemma mulr_ile1 x y : 0 <= x -> 0 <= y -> x <= 1 -> y <= 1 -> x * y <= 1.
+Proof. by move=> *; rewrite (@ler_trans _ y) ?ler_pimull. Qed.
+
+Lemma mulr_ilt1 x y : 0 <= x -> 0 <= y -> x < 1 -> y < 1 -> x * y < 1.
+Proof. by move=> *; rewrite (@ler_lt_trans _ y) ?ler_pimull // ltrW. Qed.
+
+Definition mulr_ilte1 := (mulr_ile1, mulr_ilt1).
+
+Lemma mulr_ege1 x y : 1 <= x -> 1 <= y -> 1 <= x * y.
+Proof.
+by move=> le1x le1y; rewrite (@ler_trans _ y) ?ler_pemull // (ler_trans ler01).
+Qed.
+
+Lemma mulr_egt1 x y : 1 < x -> 1 < y -> 1 < x * y.
+Proof.
+by move=> le1x lt1y; rewrite (@ltr_trans _ y) // ltr_pmull // (ltr_trans ltr01).
+Qed.
+Definition mulr_egte1 := (mulr_ege1, mulr_egt1).
+Definition mulr_cp1 := (mulr_ilte1, mulr_egte1).
+
+(* ler and ^-1 *)
+
+Lemma invr_gt0 x : (0 < x^-1) = (0 < x).
+Proof.
+have [ux | nux] := boolP (x \is a GRing.unit); last by rewrite invr_out.
+by apply/idP/idP=> /ltr_pmul2r<-; rewrite mul0r (mulrV, mulVr) ?ltr01.
+Qed.
+
+Lemma invr_ge0 x : (0 <= x^-1) = (0 <= x).
+Proof. by rewrite !le0r invr_gt0 invr_eq0. Qed.
+
+Lemma invr_lt0 x : (x^-1 < 0) = (x < 0).
+Proof. by rewrite -oppr_cp0 -invrN invr_gt0 oppr_cp0. Qed.
+
+Lemma invr_le0 x : (x^-1 <= 0) = (x <= 0).
+Proof. by rewrite -oppr_cp0 -invrN invr_ge0 oppr_cp0. Qed.
+
+Definition invr_gte0 := (invr_ge0, invr_gt0).
+Definition invr_lte0 := (invr_le0, invr_lt0).
+
+Lemma divr_ge0 x y : 0 <= x -> 0 <= y -> 0 <= x / y.
+Proof. by move=> x_ge0 y_ge0; rewrite mulr_ge0 ?invr_ge0. Qed.
+
+Lemma divr_gt0 x y : 0 < x -> 0 < y -> 0 < x / y.
+Proof. by move=> x_gt0 y_gt0; rewrite pmulr_rgt0 ?invr_gt0. Qed.
+
+Lemma realV : {mono (@GRing.inv R) : x / x \is real}.
+Proof. exact: rpredV. Qed.
+
+(* ler and exprn *)
+Lemma exprn_ge0 n x : 0 <= x -> 0 <= x ^+ n.
+Proof. by move=> xge0; rewrite -nnegrE rpredX. Qed.
+
+Lemma realX n : {in real, forall x, x ^+ n \is real}.
+Proof. exact: rpredX. Qed.
+
+Lemma exprn_gt0 n x : 0 < x -> 0 < x ^+ n.
+Proof.
+by rewrite !lt0r expf_eq0 => /andP[/negPf-> /exprn_ge0->]; rewrite andbF.
+Qed.
+
+Definition exprn_gte0 := (exprn_ge0, exprn_gt0).
+
+Lemma exprn_ile1 n x : 0 <= x -> x <= 1 -> x ^+ n <= 1.
+Proof.
+move=> xge0 xle1; elim: n=> [|*]; rewrite ?expr0 // exprS.
+by rewrite mulr_ile1 ?exprn_ge0.
+Qed.
+
+Lemma exprn_ilt1 n x : 0 <= x -> x < 1 -> x ^+ n < 1 = (n != 0%N).
+Proof.
+move=> xge0 xlt1.
+case: n; [by rewrite eqxx ltrr | elim=> [|n ihn]; first by rewrite expr1].
+by rewrite exprS mulr_ilt1 // exprn_ge0.
+Qed.
+
+Definition exprn_ilte1 := (exprn_ile1, exprn_ilt1).
+
+Lemma exprn_ege1 n x : 1 <= x -> 1 <= x ^+ n.
+Proof.
+by move=> x_ge1; elim: n=> [|n ihn]; rewrite ?expr0 // exprS mulr_ege1.
+Qed.
+
+Lemma exprn_egt1 n x : 1 < x -> 1 < x ^+ n = (n != 0%N).
+Proof.
+move=> xgt1; case: n; first by rewrite eqxx ltrr.
+elim=> [|n ihn]; first by rewrite expr1.
+by rewrite exprS mulr_egt1 // exprn_ge0.
+Qed.
+
+Definition exprn_egte1 := (exprn_ege1, exprn_egt1).
+Definition exprn_cp1 := (exprn_ilte1, exprn_egte1).
+
+Lemma ler_iexpr x n : (0 < n)%N -> 0 <= x -> x <= 1 -> x ^+ n <= x.
+Proof. by case: n => n // *; rewrite exprS ler_pimulr // exprn_ile1. Qed.
+
+Lemma ltr_iexpr x n : 0 < x -> x < 1 -> (x ^+ n < x) = (1 < n)%N.
+Proof.
+case: n=> [|[|n]] //; first by rewrite expr0 => _ /ltr_gtF ->.
+by move=> x0 x1; rewrite exprS gtr_pmulr // ?exprn_ilt1 // ltrW.
+Qed.
+
+Definition lter_iexpr := (ler_iexpr, ltr_iexpr).
+
+Lemma ler_eexpr x n : (0 < n)%N -> 1 <= x -> x <= x ^+ n.
+Proof.
+case: n => // n _ x_ge1.
+by rewrite exprS ler_pemulr ?(ler_trans _ x_ge1) // exprn_ege1.
+Qed.
+
+Lemma ltr_eexpr x n : 1 < x -> (x < x ^+ n) = (1 < n)%N.
+Proof.
+move=> x_ge1; case: n=> [|[|n]] //; first by rewrite expr0 ltr_gtF.
+by rewrite exprS ltr_pmulr ?(ltr_trans _ x_ge1) ?exprn_egt1.
+Qed.
+
+Definition lter_eexpr := (ler_eexpr, ltr_eexpr).
+Definition lter_expr := (lter_iexpr, lter_eexpr).
+
+Lemma ler_wiexpn2l x :
+ 0 <= x -> x <= 1 -> {homo (GRing.exp x) : m n / (n <= m)%N >-> m <= n}.
+Proof.
+move=> xge0 xle1 m n /= hmn.
+by rewrite -(subnK hmn) exprD ler_pimull ?(exprn_ge0, exprn_ile1).
+Qed.
+
+Lemma ler_weexpn2l x :
+ 1 <= x -> {homo (GRing.exp x) : m n / (m <= n)%N >-> m <= n}.
+Proof.
+move=> xge1 m n /= hmn; rewrite -(subnK hmn) exprD.
+by rewrite ler_pemull ?(exprn_ge0, exprn_ege1) // (ler_trans _ xge1) ?ler01.
+Qed.
+
+Lemma ieexprn_weq1 x n : 0 <= x -> (x ^+ n == 1) = ((n == 0%N) || (x == 1)).
+Proof.
+move=> xle0; case: n => [|n]; first by rewrite expr0 eqxx.
+case: (@real_ltrgtP x 1); do ?by rewrite ?ger0_real.
++ by move=> x_lt1; rewrite ?ltr_eqF // exprn_ilt1.
++ by move=> x_lt1; rewrite ?gtr_eqF // exprn_egt1.
+by move->; rewrite expr1n eqxx.
+Qed.
+
+Lemma ieexprIn x : 0 < x -> x != 1 -> injective (GRing.exp x).
+Proof.
+move=> x_gt0 x_neq1 m n; without loss /subnK <-: m n / (n <= m)%N.
+ by move=> IH eq_xmn; case/orP: (leq_total m n) => /IH->.
+case: {m}(m - n)%N => // m /eqP/idPn[]; rewrite -[x ^+ n]mul1r exprD.
+by rewrite (inj_eq (mulIf _)) ?ieexprn_weq1 ?ltrW // expf_neq0 ?gtr_eqF.
+Qed.
+
+Lemma ler_iexpn2l x :
+ 0 < x -> x < 1 -> {mono (GRing.exp x) : m n / (n <= m)%N >-> m <= n}.
+Proof.
+move=> xgt0 xlt1; apply: (nhomo_leq_mono (nhomo_inj_ltn_lt _ _)); last first.
+ by apply: ler_wiexpn2l; rewrite ltrW.
+by apply: ieexprIn; rewrite ?ltr_eqF ?ltr_cpable.
+Qed.
+
+Lemma ltr_iexpn2l x :
+ 0 < x -> x < 1 -> {mono (GRing.exp x) : m n / (n < m)%N >-> m < n}.
+Proof. by move=> xgt0 xlt1; apply: (leq_lerW_nmono (ler_iexpn2l _ _)). Qed.
+
+Definition lter_iexpn2l := (ler_iexpn2l, ltr_iexpn2l).
+
+Lemma ler_eexpn2l x :
+ 1 < x -> {mono (GRing.exp x) : m n / (m <= n)%N >-> m <= n}.
+Proof.
+move=> xgt1; apply: (homo_leq_mono (homo_inj_ltn_lt _ _)); last first.
+ by apply: ler_weexpn2l; rewrite ltrW.
+by apply: ieexprIn; rewrite ?gtr_eqF ?gtr_cpable //; apply: ltr_trans xgt1.
+Qed.
+
+Lemma ltr_eexpn2l x :
+ 1 < x -> {mono (GRing.exp x) : m n / (m < n)%N >-> m < n}.
+Proof. by move=> xgt1; apply: (leq_lerW_mono (ler_eexpn2l _)). Qed.
+
+Definition lter_eexpn2l := (ler_eexpn2l, ltr_eexpn2l).
+
+Lemma ltr_expn2r n x y : 0 <= x -> x < y -> x ^+ n < y ^+ n = (n != 0%N).
+Proof.
+move=> xge0 xlty; case: n; first by rewrite ltrr.
+elim=> [|n IHn]; rewrite ?[_ ^+ _.+2]exprS //.
+rewrite (@ler_lt_trans _ (x * y ^+ n.+1)) ?ler_wpmul2l ?ltr_pmul2r ?IHn //.
+ by rewrite ltrW // ihn.
+by rewrite exprn_gt0 // (ler_lt_trans xge0).
+Qed.
+
+Lemma ler_expn2r n : {in nneg & , {homo ((@GRing.exp R)^~ n) : x y / x <= y}}.
+Proof.
+move=> x y /= x0 y0 xy; elim: n => [|n IHn]; rewrite !(expr0, exprS) //.
+by rewrite (@ler_trans _ (x * y ^+ n)) ?ler_wpmul2l ?ler_wpmul2r ?exprn_ge0.
+Qed.
+
+Definition lter_expn2r := (ler_expn2r, ltr_expn2r).
+
+Lemma ltr_wpexpn2r n :
+ (0 < n)%N -> {in nneg & , {homo ((@GRing.exp R)^~ n) : x y / x < y}}.
+Proof. by move=> ngt0 x y /= x0 y0 hxy; rewrite ltr_expn2r // -lt0n. Qed.
+
+Lemma ler_pexpn2r n :
+ (0 < n)%N -> {in nneg & , {mono ((@GRing.exp R)^~ n) : x y / x <= y}}.
+Proof.
+case: n => // n _ x y; rewrite !qualifE /= => x_ge0 y_ge0.
+have [-> | nzx] := eqVneq x 0; first by rewrite exprS mul0r exprn_ge0.
+rewrite -subr_ge0 subrXX pmulr_lge0 ?subr_ge0 //= big_ord_recr /=.
+rewrite subnn expr0 mul1r /= ltr_spaddr // ?exprn_gt0 ?lt0r ?nzx //.
+by rewrite sumr_ge0 // => i _; rewrite mulr_ge0 ?exprn_ge0.
+Qed.
+
+Lemma ltr_pexpn2r n :
+ (0 < n)%N -> {in nneg & , {mono ((@GRing.exp R)^~ n) : x y / x < y}}.
+Proof.
+by move=> n_gt0 x y x_ge0 y_ge0; rewrite !ltr_neqAle !eqr_le !ler_pexpn2r.
+Qed.
+
+Definition lter_pexpn2r := (ler_pexpn2r, ltr_pexpn2r).
+
+Lemma pexpIrn n : (0 < n)%N -> {in nneg &, injective ((@GRing.exp R)^~ n)}.
+Proof. by move=> n_gt0; apply: mono_inj_in (ler_pexpn2r _). Qed.
+
+(* expr and ler/ltr *)
+Lemma expr_le1 n x : (0 < n)%N -> 0 <= x -> (x ^+ n <= 1) = (x <= 1).
+Proof.
+by move=> ngt0 xge0; rewrite -{1}[1](expr1n _ n) ler_pexpn2r // [_ \in _]ler01.
+Qed.
+
+Lemma expr_lt1 n x : (0 < n)%N -> 0 <= x -> (x ^+ n < 1) = (x < 1).
+Proof.
+by move=> ngt0 xge0; rewrite -{1}[1](expr1n _ n) ltr_pexpn2r // [_ \in _]ler01.
+Qed.
+
+Definition expr_lte1 := (expr_le1, expr_lt1).
+
+Lemma expr_ge1 n x : (0 < n)%N -> 0 <= x -> (1 <= x ^+ n) = (1 <= x).
+Proof.
+by move=> ngt0 xge0; rewrite -{1}[1](expr1n _ n) ler_pexpn2r // [_ \in _]ler01.
+Qed.
+
+Lemma expr_gt1 n x : (0 < n)%N -> 0 <= x -> (1 < x ^+ n) = (1 < x).
+Proof.
+by move=> ngt0 xge0; rewrite -{1}[1](expr1n _ n) ltr_pexpn2r // [_ \in _]ler01.
+Qed.
+
+Definition expr_gte1 := (expr_ge1, expr_gt1).
+
+Lemma pexpr_eq1 x n : (0 < n)%N -> 0 <= x -> (x ^+ n == 1) = (x == 1).
+Proof. by move=> ngt0 xge0; rewrite !eqr_le expr_le1 // expr_ge1. Qed.
+
+Lemma pexprn_eq1 x n : 0 <= x -> (x ^+ n == 1) = (n == 0%N) || (x == 1).
+Proof. by case: n => [|n] xge0; rewrite ?eqxx // pexpr_eq1 ?gtn_eqF. Qed.
+
+Lemma eqr_expn2 n x y :
+ (0 < n)%N -> 0 <= x -> 0 <= y -> (x ^+ n == y ^+ n) = (x == y).
+Proof. by move=> ngt0 xge0 yge0; rewrite (inj_in_eq (pexpIrn _)). Qed.
+
+Lemma sqrp_eq1 x : 0 <= x -> (x ^+ 2 == 1) = (x == 1).
+Proof. by move/pexpr_eq1->. Qed.
+
+Lemma sqrn_eq1 x : x <= 0 -> (x ^+ 2 == 1) = (x == -1).
+Proof. by rewrite -sqrrN -oppr_ge0 -eqr_oppLR => /sqrp_eq1. Qed.
+
+Lemma ler_sqr : {in nneg &, {mono (fun x => x ^+ 2) : x y / x <= y}}.
+Proof. exact: ler_pexpn2r. Qed.
+
+Lemma ltr_sqr : {in nneg &, {mono (fun x => x ^+ 2) : x y / x < y}}.
+Proof. exact: ltr_pexpn2r. Qed.
+
+Lemma ler_pinv :
+ {in [pred x in GRing.unit | 0 < x] &, {mono (@GRing.inv R) : x y /~ x <= y}}.
+Proof.
+move=> x y /andP [ux hx] /andP [uy hy] /=.
+rewrite -(ler_pmul2l hx) -(ler_pmul2r hy).
+by rewrite !(divrr, mulrVK) ?unitf_gt0 // mul1r.
+Qed.
+
+Lemma ler_ninv :
+ {in [pred x in GRing.unit | x < 0] &, {mono (@GRing.inv R) : x y /~ x <= y}}.
+Proof.
+move=> x y /andP [ux hx] /andP [uy hy] /=.
+rewrite -(ler_nmul2l hx) -(ler_nmul2r hy).
+by rewrite !(divrr, mulrVK) ?unitf_lt0 // mul1r.
+Qed.
+
+Lemma ltr_pinv :
+ {in [pred x in GRing.unit | 0 < x] &, {mono (@GRing.inv R) : x y /~ x < y}}.
+Proof. exact: lerW_nmono_in ler_pinv. Qed.
+
+Lemma ltr_ninv :
+ {in [pred x in GRing.unit | x < 0] &, {mono (@GRing.inv R) : x y /~ x < y}}.
+Proof. exact: lerW_nmono_in ler_ninv. Qed.
+
+Lemma invr_gt1 x : x \is a GRing.unit -> 0 < x -> (1 < x^-1) = (x < 1).
+Proof.
+by move=> Ux xgt0; rewrite -{1}[1]invr1 ltr_pinv ?inE ?unitr1 ?ltr01 ?Ux.
+Qed.
+
+Lemma invr_ge1 x : x \is a GRing.unit -> 0 < x -> (1 <= x^-1) = (x <= 1).
+Proof.
+by move=> Ux xgt0; rewrite -{1}[1]invr1 ler_pinv ?inE ?unitr1 ?ltr01 // Ux.
+Qed.
+
+Definition invr_gte1 := (invr_ge1, invr_gt1).
+
+Lemma invr_le1 x (ux : x \is a GRing.unit) (hx : 0 < x) :
+ (x^-1 <= 1) = (1 <= x).
+Proof. by rewrite -invr_ge1 ?invr_gt0 ?unitrV // invrK. Qed.
+
+Lemma invr_lt1 x (ux : x \is a GRing.unit) (hx : 0 < x) : (x^-1 < 1) = (1 < x).
+Proof. by rewrite -invr_gt1 ?invr_gt0 ?unitrV // invrK. Qed.
+
+Definition invr_lte1 := (invr_le1, invr_lt1).
+Definition invr_cp1 := (invr_gte1, invr_lte1).
+
+(* norm *)
+
+Lemma real_ler_norm x : x \is real -> x <= `|x|.
+Proof.
+by case/real_ger0P=> hx //; rewrite (ler_trans (ltrW hx)) // oppr_ge0 ltrW.
+Qed.
+
+(* norm + add *)
+
+Lemma normr_real x : `|x| \is real. Proof. by rewrite ger0_real. Qed.
+Hint Resolve normr_real.
+
+Lemma ler_norm_sum I r (G : I -> R) (P : pred I):
+ `|\sum_(i <- r | P i) G i| <= \sum_(i <- r | P i) `|G i|.
+Proof.
+elim/big_rec2: _ => [|i y x _]; first by rewrite normr0.
+by rewrite -(ler_add2l `|G i|); apply: ler_trans; apply: ler_norm_add.
+Qed.
+
+Lemma ler_norm_sub x y : `|x - y| <= `|x| + `|y|.
+Proof. by rewrite (ler_trans (ler_norm_add _ _)) ?normrN. Qed.
+
+Lemma ler_dist_add z x y : `|x - y| <= `|x - z| + `|z - y|.
+Proof. by rewrite (ler_trans _ (ler_norm_add _ _)) // addrA addrNK. Qed.
+
+Lemma ler_sub_norm_add x y : `|x| - `|y| <= `|x + y|.
+Proof.
+rewrite -{1}[x](addrK y) lter_sub_addl.
+by rewrite (ler_trans (ler_norm_add _ _)) // addrC normrN.
+Qed.
+
+Lemma ler_sub_dist x y : `|x| - `|y| <= `|x - y|.
+Proof. by rewrite -[`|y|]normrN ler_sub_norm_add. Qed.
+
+Lemma ler_dist_dist x y : `|`|x| - `|y| | <= `|x - y|.
+Proof.
+have [||_|_] // := @real_lerP `|x| `|y|; last by rewrite ler_sub_dist.
+by rewrite distrC ler_sub_dist.
+Qed.
+
+Lemma ler_dist_norm_add x y : `| `|x| - `|y| | <= `| x + y |.
+Proof. by rewrite -[y]opprK normrN ler_dist_dist. Qed.
+
+Lemma real_ler_norml x y : x \is real -> (`|x| <= y) = (- y <= x <= y).
+Proof.
+move=> xR; wlog x_ge0 : x xR / 0 <= x => [hwlog|].
+ move: (xR) => /(@real_leVge 0) /orP [|/hwlog->|hx] //.
+ by rewrite -[x]opprK normrN ler_opp2 andbC ler_oppl hwlog ?realN ?oppr_ge0.
+rewrite ger0_norm //; have [le_xy|] := boolP (x <= y); last by rewrite andbF.
+by rewrite (ler_trans _ x_ge0) // oppr_le0 (ler_trans x_ge0).
+Qed.
+
+Lemma real_ler_normlP x y :
+ x \is real -> reflect ((-x <= y) * (x <= y)) (`|x| <= y).
+Proof.
+by move=> Rx; rewrite real_ler_norml // ler_oppl; apply: (iffP andP) => [] [].
+Qed.
+Implicit Arguments real_ler_normlP [x y].
+
+Lemma real_eqr_norml x y :
+ x \is real -> (`|x| == y) = ((x == y) || (x == -y)) && (0 <= y).
+Proof.
+move=> Rx.
+apply/idP/idP=> [|/andP[/pred2P[]-> /ger0_norm/eqP]]; rewrite ?normrE //.
+case: real_ler0P => // hx; rewrite 1?eqr_oppLR => /eqP exy.
+ by move: hx; rewrite exy ?oppr_le0 eqxx orbT //.
+by move: hx=> /ltrW; rewrite exy eqxx.
+Qed.
+
+Lemma real_eqr_norm2 x y :
+ x \is real -> y \is real -> (`|x| == `|y|) = (x == y) || (x == -y).
+Proof.
+move=> Rx Ry; rewrite real_eqr_norml // normrE andbT.
+by case: real_ler0P; rewrite // opprK orbC.
+Qed.
+
+Lemma real_ltr_norml x y : x \is real -> (`|x| < y) = (- y < x < y).
+Proof.
+move=> Rx; wlog x_ge0 : x Rx / 0 <= x => [hwlog|].
+ move: (Rx) => /(@real_leVge 0) /orP [|/hwlog->|hx] //.
+ by rewrite -[x]opprK normrN ltr_opp2 andbC ltr_oppl hwlog ?realN ?oppr_ge0.
+rewrite ger0_norm //; have [le_xy|] := boolP (x < y); last by rewrite andbF.
+by rewrite (ltr_le_trans _ x_ge0) // oppr_lt0 (ler_lt_trans x_ge0).
+Qed.
+
+Definition real_lter_norml := (real_ler_norml, real_ltr_norml).
+
+Lemma real_ltr_normlP x y :
+ x \is real -> reflect ((-x < y) * (x < y)) (`|x| < y).
+Proof.
+move=> Rx; rewrite real_ltr_norml // ltr_oppl.
+by apply: (iffP (@andP _ _)); case.
+Qed.
+Implicit Arguments real_ltr_normlP [x y].
+
+Lemma real_ler_normr x y : y \is real -> (x <= `|y|) = (x <= y) || (x <= - y).
+Proof.
+move=> Ry.
+have [xR|xNR] := boolP (x \is real); last by rewrite ?Nreal_leF ?realN.
+rewrite real_lerNgt ?real_ltr_norml // negb_and -?real_lerNgt ?realN //.
+by rewrite orbC ler_oppr.
+Qed.
+
+Lemma real_ltr_normr x y : y \is real -> (x < `|y|) = (x < y) || (x < - y).
+Proof.
+move=> Ry.
+have [xR|xNR] := boolP (x \is real); last by rewrite ?Nreal_ltF ?realN.
+rewrite real_ltrNge ?real_ler_norml // negb_and -?real_ltrNge ?realN //.
+by rewrite orbC ltr_oppr.
+Qed.
+
+Definition real_lter_normr := (real_ler_normr, real_ltr_normr).
+
+Lemma ler_nnorml x y : y < 0 -> `|x| <= y = false.
+Proof. by move=> y_lt0; rewrite ltr_geF // (ltr_le_trans y_lt0). Qed.
+
+Lemma ltr_nnorml x y : y <= 0 -> `|x| < y = false.
+Proof. by move=> y_le0; rewrite ler_gtF // (ler_trans y_le0). Qed.
+
+Definition lter_nnormr := (ler_nnorml, ltr_nnorml).
+
+Lemma real_ler_distl x y e :
+ x - y \is real -> (`|x - y| <= e) = (y - e <= x <= y + e).
+Proof. by move=> Rxy; rewrite real_lter_norml // !lter_sub_addl. Qed.
+
+Lemma real_ltr_distl x y e :
+ x - y \is real -> (`|x - y| < e) = (y - e < x < y + e).
+Proof. by move=> Rxy; rewrite real_lter_norml // !lter_sub_addl. Qed.
+
+Definition real_lter_distl := (real_ler_distl, real_ltr_distl).
+
+(* GG: pointless duplication }-( *)
+Lemma eqr_norm_id x : (`|x| == x) = (0 <= x). Proof. by rewrite ger0_def. Qed.
+Lemma eqr_normN x : (`|x| == - x) = (x <= 0). Proof. by rewrite ler0_def. Qed.
+Definition eqr_norm_idVN := =^~ (ger0_def, ler0_def).
+
+Lemma real_exprn_even_ge0 n x : x \is real -> ~~ odd n -> 0 <= x ^+ n.
+Proof.
+move=> xR even_n; have [/exprn_ge0 -> //|x_lt0] := real_ger0P xR.
+rewrite -[x]opprK -mulN1r exprMn -signr_odd (negPf even_n) expr0 mul1r.
+by rewrite exprn_ge0 ?oppr_ge0 ?ltrW.
+Qed.
+
+Lemma real_exprn_even_gt0 n x :
+ x \is real -> ~~ odd n -> (0 < x ^+ n) = (n == 0)%N || (x != 0).
+Proof.
+move=> xR n_even; rewrite lt0r real_exprn_even_ge0 ?expf_eq0 //.
+by rewrite andbT negb_and lt0n negbK.
+Qed.
+
+Lemma real_exprn_even_le0 n x :
+ x \is real -> ~~ odd n -> (x ^+ n <= 0) = (n != 0%N) && (x == 0).
+Proof.
+move=> xR n_even; rewrite !real_lerNgt ?rpred0 ?rpredX //.
+by rewrite real_exprn_even_gt0 // negb_or negbK.
+Qed.
+
+Lemma real_exprn_even_lt0 n x :
+ x \is real -> ~~ odd n -> (x ^+ n < 0) = false.
+Proof. by move=> xR n_even; rewrite ler_gtF // real_exprn_even_ge0. Qed.
+
+Lemma real_exprn_odd_ge0 n x :
+ x \is real -> odd n -> (0 <= x ^+ n) = (0 <= x).
+Proof.
+case/real_ger0P => [x_ge0|x_lt0] n_odd; first by rewrite exprn_ge0.
+apply: negbTE; rewrite ltr_geF //.
+case: n n_odd => // n /= n_even; rewrite exprS pmulr_llt0 //.
+by rewrite real_exprn_even_gt0 ?ler0_real ?ltrW // ltr_eqF ?orbT.
+Qed.
+
+Lemma real_exprn_odd_gt0 n x : x \is real -> odd n -> (0 < x ^+ n) = (0 < x).
+Proof.
+by move=> xR n_odd; rewrite !lt0r expf_eq0 real_exprn_odd_ge0; case: n n_odd.
+Qed.
+
+Lemma real_exprn_odd_le0 n x : x \is real -> odd n -> (x ^+ n <= 0) = (x <= 0).
+Proof.
+by move=> xR n_odd; rewrite !real_lerNgt ?rpred0 ?rpredX // real_exprn_odd_gt0.
+Qed.
+
+Lemma real_exprn_odd_lt0 n x : x \is real -> odd n -> (x ^+ n < 0) = (x < 0).
+Proof.
+by move=> xR n_odd; rewrite !real_ltrNge ?rpred0 ?rpredX // real_exprn_odd_ge0.
+Qed.
+
+(* GG: Could this be a better definition of "real" ? *)
+Lemma realEsqr x : (x \is real) = (0 <= x ^+ 2).
+Proof. by rewrite ger0_def normrX eqf_sqr -ger0_def -ler0_def. Qed.
+
+Lemma real_normK x : x \is real -> `|x| ^+ 2 = x ^+ 2.
+Proof. by move=> Rx; rewrite -normrX ger0_norm -?realEsqr. Qed.
+
+(* Binary sign ((-1) ^+ s). *)
+
+Lemma normr_sign s : `|(-1) ^+ s| = 1 :> R.
+Proof. by rewrite normrX normrN1 expr1n. Qed.
+
+Lemma normrMsign s x : `|(-1) ^+ s * x| = `|x|.
+Proof. by rewrite normrM normr_sign mul1r. Qed.
+
+Lemma signr_gt0 (b : bool) : (0 < (-1) ^+ b :> R) = ~~ b.
+Proof. by case: b; rewrite (ltr01, ltr0N1). Qed.
+
+Lemma signr_lt0 (b : bool) : ((-1) ^+ b < 0 :> R) = b.
+Proof. by case: b; rewrite // ?(ltrN10, ltr10). Qed.
+
+Lemma signr_ge0 (b : bool) : (0 <= (-1) ^+ b :> R) = ~~ b.
+Proof. by rewrite le0r signr_eq0 signr_gt0. Qed.
+
+Lemma signr_le0 (b : bool) : ((-1) ^+ b <= 0 :> R) = b.
+Proof. by rewrite ler_eqVlt signr_eq0 signr_lt0. Qed.
+
+(* This actually holds for char R != 2. *)
+Lemma signr_inj : injective (fun b : bool => (-1) ^+ b : R).
+Proof. exact: can_inj (fun x => 0 >= x) signr_le0. Qed.
+
+(* Ternary sign (sg). *)
+
+Lemma sgr_def x : sg x = (-1) ^+ (x < 0)%R *+ (x != 0).
+Proof. by rewrite /sg; do 2!case: ifP => //. Qed.
+
+Lemma neqr0_sign x : x != 0 -> (-1) ^+ (x < 0)%R = sgr x.
+Proof. by rewrite sgr_def => ->. Qed.
+
+Lemma gtr0_sg x : 0 < x -> sg x = 1.
+Proof. by move=> x_gt0; rewrite /sg gtr_eqF // ltr_gtF. Qed.
+
+Lemma ltr0_sg x : x < 0 -> sg x = -1.
+Proof. by move=> x_lt0; rewrite /sg x_lt0 ltr_eqF. Qed.
+
+Lemma sgr0 : sg 0 = 0 :> R. Proof. by rewrite /sgr eqxx. Qed.
+Lemma sgr1 : sg 1 = 1 :> R. Proof. by rewrite gtr0_sg // ltr01. Qed.
+Lemma sgrN1 : sg (-1) = -1 :> R. Proof. by rewrite ltr0_sg // ltrN10. Qed.
+Definition sgrE := (sgr0, sgr1, sgrN1).
+
+Lemma sqr_sg x : sg x ^+ 2 = (x != 0)%:R.
+Proof. by rewrite sgr_def exprMn_n sqrr_sign -mulnn mulnb andbb. Qed.
+
+Lemma mulr_sg_eq1 x y : (sg x * y == 1) = (x != 0) && (sg x == y).
+Proof.
+rewrite /sg eq_sym; case: ifP => _; first by rewrite mul0r oner_eq0.
+by case: ifP => _; rewrite ?mul1r // mulN1r eqr_oppLR.
+Qed.
+
+Lemma mulr_sg_eqN1 x y : (sg x * sg y == -1) = (x != 0) && (sg x == - sg y).
+Proof.
+move/sg: y => y; rewrite /sg eq_sym eqr_oppLR.
+case: ifP => _; first by rewrite mul0r oppr0 oner_eq0.
+by case: ifP => _; rewrite ?mul1r // mulN1r eqr_oppLR.
+Qed.
+
+Lemma sgr_eq0 x : (sg x == 0) = (x == 0).
+Proof. by rewrite -sqrf_eq0 sqr_sg pnatr_eq0; case: (x == 0). Qed.
+
+Lemma sgr_odd n x : x != 0 -> (sg x) ^+ n = (sg x) ^+ (odd n).
+Proof. by rewrite /sg; do 2!case: ifP => // _; rewrite ?expr1n ?signr_odd. Qed.
+
+Lemma sgrMn x n : sg (x *+ n) = (n != 0%N)%:R * sg x.
+Proof.
+case: n => [|n]; first by rewrite mulr0n sgr0 mul0r.
+by rewrite !sgr_def mulrn_eq0 mul1r pmulrn_llt0.
+Qed.
+
+Lemma sgr_nat n : sg n%:R = (n != 0%N)%:R :> R.
+Proof. by rewrite sgrMn sgr1 mulr1. Qed.
+
+Lemma sgr_id x : sg (sg x) = sg x.
+Proof. by rewrite !(fun_if sg) !sgrE. Qed.
+
+Lemma sgr_lt0 x : (sg x < 0) = (x < 0).
+Proof.
+rewrite /sg; case: eqP => [-> // | _].
+by case: ifP => _; rewrite ?ltrN10 // ltr_gtF.
+Qed.
+
+Lemma sgr_le0 x : (sgr x <= 0) = (x <= 0).
+Proof. by rewrite !ler_eqVlt sgr_eq0 sgr_lt0. Qed.
+
+(* sign and norm *)
+
+Lemma realEsign x : x \is real -> x = (-1) ^+ (x < 0)%R * `|x|.
+Proof. by case/real_ger0P; rewrite (mul1r, mulN1r) ?opprK. Qed.
+
+Lemma realNEsign x : x \is real -> - x = (-1) ^+ (0 < x)%R * `|x|.
+Proof. by move=> Rx; rewrite -normrN -oppr_lt0 -realEsign ?rpredN. Qed.
+
+Lemma real_normrEsign (x : R) (xR : x \is real) : `|x| = (-1) ^+ (x < 0)%R * x.
+Proof. by rewrite {3}[x]realEsign // signrMK. Qed.
+
+(* GG: pointless duplication... *)
+Lemma real_mulr_sign_norm x : x \is real -> (-1) ^+ (x < 0)%R * `|x| = x.
+Proof. by move/realEsign. Qed.
+
+Lemma real_mulr_Nsign_norm x : x \is real -> (-1) ^+ (0 < x)%R * `|x| = - x.
+Proof. by move/realNEsign. Qed.
+
+Lemma realEsg x : x \is real -> x = sgr x * `|x|.
+Proof.
+move=> xR; have [-> | ] := eqVneq x 0; first by rewrite normr0 mulr0.
+by move=> /neqr0_sign <-; rewrite -realEsign.
+Qed.
+
+Lemma normr_sg x : `|sg x| = (x != 0)%:R.
+Proof. by rewrite sgr_def -mulr_natr normrMsign normr_nat. Qed.
+
+Lemma sgr_norm x : sg `|x| = (x != 0)%:R.
+Proof. by rewrite /sg ler_gtF ?normr_ge0 // normr_eq0 mulrb if_neg. Qed.
+
+(* lerif *)
+
+Lemma lerif_refl x C : reflect (x <= x ?= iff C) C.
+Proof. by apply: (iffP idP) => [-> | <-] //; split; rewrite ?eqxx. Qed.
+
+Lemma lerif_trans x1 x2 x3 C12 C23 :
+ x1 <= x2 ?= iff C12 -> x2 <= x3 ?= iff C23 -> x1 <= x3 ?= iff C12 && C23.
+Proof.
+move=> ltx12 ltx23; apply/lerifP; rewrite -ltx12.
+case eqx12: (x1 == x2).
+ by rewrite (eqP eqx12) ltr_neqAle !ltx23 andbT; case C23.
+by rewrite (@ltr_le_trans _ x2) ?ltx23 // ltr_neqAle eqx12 ltx12.
+Qed.
+
+Lemma lerif_le x y : x <= y -> x <= y ?= iff (x >= y).
+Proof. by move=> lexy; split=> //; rewrite eqr_le lexy. Qed.
+
+Lemma lerif_eq x y : x <= y -> x <= y ?= iff (x == y).
+Proof. by []. Qed.
+
+Lemma ger_lerif x y C : x <= y ?= iff C -> (y <= x) = C.
+Proof. by case=> le_xy; rewrite eqr_le le_xy. Qed.
+
+Lemma ltr_lerif x y C : x <= y ?= iff C -> (x < y) = ~~ C.
+Proof. by move=> le_xy; rewrite ltr_neqAle !le_xy andbT. Qed.
+
+Lemma lerif_nat m n C : (m%:R <= n%:R ?= iff C :> R) = (m <= n ?= iff C)%N.
+Proof. by rewrite /lerif !ler_nat eqr_nat. Qed.
+
+Lemma mono_in_lerif (A : pred R) (f : R -> R) C :
+ {in A &, {mono f : x y / x <= y}} ->
+ {in A &, forall x y, (f x <= f y ?= iff C) = (x <= y ?= iff C)}.
+Proof.
+by move=> mf x y Ax Ay; rewrite /lerif mf ?(inj_in_eq (mono_inj_in mf)).
+Qed.
+
+Lemma mono_lerif (f : R -> R) C :
+ {mono f : x y / x <= y} ->
+ forall x y, (f x <= f y ?= iff C) = (x <= y ?= iff C).
+Proof. by move=> mf x y; rewrite /lerif mf (inj_eq (mono_inj _)). Qed.
+
+Lemma nmono_in_lerif (A : pred R) (f : R -> R) C :
+ {in A &, {mono f : x y /~ x <= y}} ->
+ {in A &, forall x y, (f x <= f y ?= iff C) = (y <= x ?= iff C)}.
+Proof.
+by move=> mf x y Ax Ay; rewrite /lerif eq_sym mf ?(inj_in_eq (nmono_inj_in mf)).
+Qed.
+
+Lemma nmono_lerif (f : R -> R) C :
+ {mono f : x y /~ x <= y} ->
+ forall x y, (f x <= f y ?= iff C) = (y <= x ?= iff C).
+Proof. by move=> mf x y; rewrite /lerif eq_sym mf ?(inj_eq (nmono_inj mf)). Qed.
+
+Lemma lerif_subLR x y z C : (x - y <= z ?= iff C) = (x <= z + y ?= iff C).
+Proof. by rewrite /lerif !eqr_le ler_subr_addr ler_subl_addr. Qed.
+
+Lemma lerif_subRL x y z C : (x <= y - z ?= iff C) = (x + z <= y ?= iff C).
+Proof. by rewrite -lerif_subLR opprK. Qed.
+
+Lemma lerif_add x1 y1 C1 x2 y2 C2 :
+ x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 ->
+ x1 + x2 <= y1 + y2 ?= iff C1 && C2.
+Proof.
+rewrite -(mono_lerif _ (ler_add2r x2)) -(mono_lerif C2 (ler_add2l y1)).
+exact: lerif_trans.
+Qed.
+
+Lemma lerif_sum (I : finType) (P C : pred I) (E1 E2 : I -> R) :
+ (forall i, P i -> E1 i <= E2 i ?= iff C i) ->
+ \sum_(i | P i) E1 i <= \sum_(i | P i) E2 i ?= iff [forall (i | P i), C i].
+Proof.
+move=> leE12; rewrite -big_andE.
+elim/big_rec3: _ => [|i Ci m2 m1 /leE12]; first by rewrite /lerif lerr eqxx.
+exact: lerif_add.
+Qed.
+
+Lemma lerif_0_sum (I : finType) (P C : pred I) (E : I -> R) :
+ (forall i, P i -> 0 <= E i ?= iff C i) ->
+ 0 <= \sum_(i | P i) E i ?= iff [forall (i | P i), C i].
+Proof. by move/lerif_sum; rewrite big1_eq. Qed.
+
+Lemma real_lerif_norm x : x \is real -> x <= `|x| ?= iff (0 <= x).
+Proof.
+by move=> xR; rewrite ger0_def eq_sym; apply: lerif_eq; rewrite real_ler_norm.
+Qed.
+
+Lemma lerif_pmul x1 x2 y1 y2 C1 C2 :
+ 0 <= x1 -> 0 <= x2 -> x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 ->
+ x1 * x2 <= y1 * y2 ?= iff (y1 * y2 == 0) || C1 && C2.
+Proof.
+move=> x1_ge0 x2_ge0 le_xy1 le_xy2; have [y_0 | ] := altP (_ =P 0).
+ apply/lerifP; rewrite y_0 /= mulf_eq0 !eqr_le x1_ge0 x2_ge0 !andbT.
+ move/eqP: y_0; rewrite mulf_eq0.
+ by case/pred2P=> <-; rewrite (le_xy1, le_xy2) ?orbT.
+rewrite /= mulf_eq0 => /norP[y1nz y2nz].
+have y1_gt0: 0 < y1 by rewrite ltr_def y1nz (ler_trans _ le_xy1).
+have [x2_0 | x2nz] := eqVneq x2 0.
+ apply/lerifP; rewrite -le_xy2 x2_0 eq_sym (negPf y2nz) andbF mulr0.
+ by rewrite mulr_gt0 // ltr_def y2nz -x2_0 le_xy2.
+have:= le_xy2; rewrite -(mono_lerif _ (ler_pmul2l y1_gt0)).
+by apply: lerif_trans; rewrite (mono_lerif _ (ler_pmul2r _)) // ltr_def x2nz.
+Qed.
+
+Lemma lerif_nmul x1 x2 y1 y2 C1 C2 :
+ y1 <= 0 -> y2 <= 0 -> x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 ->
+ y1 * y2 <= x1 * x2 ?= iff (x1 * x2 == 0) || C1 && C2.
+Proof.
+rewrite -!oppr_ge0 -mulrNN -[x1 * x2]mulrNN => y1le0 y2le0 le_xy1 le_xy2.
+by apply: lerif_pmul => //; rewrite (nmono_lerif _ ler_opp2).
+Qed.
+
+Lemma lerif_pprod (I : finType) (P C : pred I) (E1 E2 : I -> R) :
+ (forall i, P i -> 0 <= E1 i) ->
+ (forall i, P i -> E1 i <= E2 i ?= iff C i) ->
+ let pi E := \prod_(i | P i) E i in
+ pi E1 <= pi E2 ?= iff (pi E2 == 0) || [forall (i | P i), C i].
+Proof.
+move=> E1_ge0 leE12 /=; rewrite -big_andE; elim/(big_load (fun x => 0 <= x)): _.
+elim/big_rec3: _ => [|i Ci m2 m1 Pi [m1ge0 le_m12]].
+ by split=> //; apply/lerifP; rewrite orbT.
+have Ei_ge0 := E1_ge0 i Pi; split; first by rewrite mulr_ge0.
+congr (lerif _ _ _): (lerif_pmul Ei_ge0 m1ge0 (leE12 i Pi) le_m12).
+by rewrite mulf_eq0 -!orbA; congr (_ || _); rewrite !orb_andr orbA orbb.
+Qed.
+
+(* Mean inequalities. *)
+
+Lemma real_lerif_mean_square_scaled x y :
+ x \is real -> y \is real -> x * y *+ 2 <= x ^+ 2 + y ^+ 2 ?= iff (x == y).
+Proof.
+move=> Rx Ry; rewrite -[_ *+ 2]add0r -lerif_subRL addrAC -sqrrB -subr_eq0.
+by rewrite -sqrf_eq0 eq_sym; apply: lerif_eq; rewrite -realEsqr rpredB.
+Qed.
+
+Lemma real_lerif_AGM2_scaled x y :
+ x \is real -> y \is real -> x * y *+ 4 <= (x + y) ^+ 2 ?= iff (x == y).
+Proof.
+move=> Rx Ry; rewrite sqrrD addrAC (mulrnDr _ 2) -lerif_subLR addrK.
+exact: real_lerif_mean_square_scaled.
+Qed.
+
+Lemma lerif_AGM_scaled (I : finType) (A : pred I) (E : I -> R) (n := #|A|) :
+ {in A, forall i, 0 <= E i *+ n} ->
+ \prod_(i in A) (E i *+ n) <= (\sum_(i in A) E i) ^+ n
+ ?= iff [forall i in A, forall j in A, E i == E j].
+Proof.
+elim: {A}_.+1 {-2}A (ltnSn #|A|) => // m IHm A leAm in E n * => Ege0.
+apply/lerifP; case: ifPn => [/forall_inP-Econstant | Enonconstant].
+ have [i /= Ai | A0] := pickP (mem A); last by rewrite [n]eq_card0 ?big_pred0.
+ have /eqfun_inP-E_i := Econstant i Ai; rewrite -(eq_bigr _ E_i) sumr_const.
+ by rewrite exprMn_n prodrMn -(eq_bigr _ E_i) prodr_const.
+set mu := \sum_(i in A) E i; pose En i := E i *+ n.
+pose cmp_mu s := [pred i | s * mu < s * En i].
+have{Enonconstant} has_cmp_mu e (s := (-1) ^+ e): {i | i \in A & cmp_mu s i}.
+ apply/sig2W/exists_inP; apply: contraR Enonconstant.
+ rewrite negb_exists_in => /forall_inP-mu_s_A.
+ have n_gt0 i: i \in A -> (0 < n)%N by rewrite [n](cardD1 i) => ->.
+ have{mu_s_A} mu_s_A i: i \in A -> s * En i <= s * mu.
+ move=> Ai; rewrite real_lerNgt ?mu_s_A ?rpredMsign ?ger0_real ?Ege0 //.
+ by rewrite -(pmulrn_lge0 _ (n_gt0 i Ai)) -sumrMnl sumr_ge0.
+ have [_ /esym/eqfun_inP] := lerif_sum (fun i Ai => lerif_eq (mu_s_A i Ai)).
+ rewrite sumr_const -/n -mulr_sumr sumrMnl -/mu mulrnAr eqxx => A_mu.
+ apply/forall_inP=> i Ai; apply/eqfun_inP=> j Aj.
+ by apply: (pmulrnI (n_gt0 i Ai)); apply: (can_inj (signrMK e)); rewrite !A_mu.
+have [[i Ai Ei_lt_mu] [j Aj Ej_gt_mu]] := (has_cmp_mu 1, has_cmp_mu 0)%N.
+rewrite {cmp_mu has_cmp_mu}/= !mul1r !mulN1r ltr_opp2 in Ei_lt_mu Ej_gt_mu.
+pose A' := [predD1 A & i]; pose n' := #|A'|.
+have [Dn n_gt0]: n = n'.+1 /\ (n > 0)%N by rewrite [n](cardD1 i) Ai.
+have i'j: j != i by apply: contraTneq Ej_gt_mu => ->; rewrite ltr_gtF.
+have{i'j} A'j: j \in A' by rewrite !inE Aj i'j.
+have mu_gt0: 0 < mu := ler_lt_trans (Ege0 i Ai) Ei_lt_mu.
+rewrite (bigD1 i) // big_andbC (bigD1 j) //= mulrA; set pi := \prod_(k | _) _.
+have [-> | nz_pi] := eqVneq pi 0; first by rewrite !mulr0 exprn_gt0.
+have{nz_pi} pi_gt0: 0 < pi.
+ by rewrite ltr_def nz_pi prodr_ge0 // => k /andP[/andP[_ /Ege0]].
+rewrite -/(En i) -/(En j); pose E' := [eta En with j |-> En i + En j - mu].
+have E'ge0 k: k \in A' -> E' k *+ n' >= 0.
+ case/andP=> /= _ Ak; apply: mulrn_wge0; case: ifP => _; last exact: Ege0.
+ by rewrite subr_ge0 ler_paddl ?Ege0 // ltrW.
+rewrite -/n Dn in leAm; have{leAm IHm E'ge0}: _ <= _ := IHm _ leAm _ E'ge0.
+have ->: \sum_(k in A') E' k = mu *+ n'.
+ apply: (addrI mu); rewrite -mulrS -Dn -sumrMnl (bigD1 i Ai) big_andbC /=.
+ rewrite !(bigD1 j A'j) /= addrCA eqxx !addrA subrK; congr (_ + _).
+ by apply: eq_bigr => k /andP[_ /negPf->].
+rewrite prodrMn exprMn_n -/n' ler_pmuln2r ?expn_gt0; last by case: (n').
+have ->: \prod_(k in A') E' k = E' j * pi.
+ by rewrite (bigD1 j) //=; congr *%R; apply: eq_bigr => k /andP[_ /negPf->].
+rewrite -(ler_pmul2l mu_gt0) -exprS -Dn mulrA; apply: ltr_le_trans.
+rewrite ltr_pmul2r //= eqxx -addrA mulrDr mulrC -ltr_subl_addl -mulrBl.
+by rewrite mulrC ltr_pmul2r ?subr_gt0.
+Qed.
+
+(* Polynomial bound. *)
+
+Implicit Type p : {poly R}.
+
+Lemma poly_disk_bound p b : {ub | forall x, `|x| <= b -> `|p.[x]| <= ub}.
+Proof.
+exists (\sum_(j < size p) `|p`_j| * b ^+ j) => x le_x_b.
+rewrite horner_coef (ler_trans (ler_norm_sum _ _ _)) ?ler_sum // => j _.
+rewrite normrM normrX ler_wpmul2l ?ler_expn2r ?unfold_in ?normr_ge0 //.
+exact: ler_trans (normr_ge0 x) le_x_b.
+Qed.
+
+End NumDomainOperationTheory.
+
+Hint Resolve ler_opp2 ltr_opp2 real0 real1 normr_real.
+Implicit Arguments ler_sqr [[R] x y].
+Implicit Arguments ltr_sqr [[R] x y].
+Implicit Arguments signr_inj [[R] x1 x2].
+Implicit Arguments real_ler_normlP [R x y].
+Implicit Arguments real_ltr_normlP [R x y].
+Implicit Arguments lerif_refl [R x C].
+Implicit Arguments mono_in_lerif [R A f C].
+Implicit Arguments nmono_in_lerif [R A f C].
+Implicit Arguments mono_lerif [R f C].
+Implicit Arguments nmono_lerif [R f C].
+
+Section NumDomainMonotonyTheoryForReals.
+
+Variables (R R' : numDomainType) (D : pred R) (f : R -> R').
+Implicit Types (m n p : nat) (x y z : R) (u v w : R').
+
+Lemma real_mono :
+ {homo f : x y / x < y} -> {in real &, {mono f : x y / x <= y}}.
+Proof.
+move=> mf x y xR yR /=; have [lt_xy | le_yx] := real_lerP xR yR.
+ by rewrite ltrW_homo.
+by rewrite ltr_geF ?mf.
+Qed.
+
+Lemma real_nmono :
+ {homo f : x y /~ x < y} -> {in real &, {mono f : x y /~ x <= y}}.
+Proof.
+move=> mf x y xR yR /=; have [lt_xy|le_yx] := real_ltrP xR yR.
+ by rewrite ltr_geF ?mf.
+by rewrite ltrW_nhomo.
+Qed.
+
+(* GG: Domain should precede condition. *)
+Lemma real_mono_in :
+ {in D &, {homo f : x y / x < y}} ->
+ {in [pred x in D | x \is real] &, {mono f : x y / x <= y}}.
+Proof.
+move=> Dmf x y /andP[hx xR] /andP[hy yR] /=.
+have [lt_xy|le_yx] := real_lerP xR yR; first by rewrite (ltrW_homo_in Dmf).
+by rewrite ltr_geF ?Dmf.
+Qed.
+
+Lemma real_nmono_in :
+ {in D &, {homo f : x y /~ x < y}} ->
+ {in [pred x in D | x \is real] &, {mono f : x y /~ x <= y}}.
+Proof.
+move=> Dmf x y /andP[hx xR] /andP[hy yR] /=.
+have [lt_xy|le_yx] := real_ltrP xR yR; last by rewrite (ltrW_nhomo_in Dmf).
+by rewrite ltr_geF ?Dmf.
+Qed.
+
+End NumDomainMonotonyTheoryForReals.
+
+Section FinGroup.
+
+Import GroupScope.
+
+Variables (R : numDomainType) (gT : finGroupType).
+Implicit Types G : {group gT}.
+
+Lemma natrG_gt0 G : #|G|%:R > 0 :> R.
+Proof. by rewrite ltr0n cardG_gt0. Qed.
+
+Lemma natrG_neq0 G : #|G|%:R != 0 :> R.
+Proof. by rewrite gtr_eqF // natrG_gt0. Qed.
+
+Lemma natr_indexg_gt0 G B : #|G : B|%:R > 0 :> R.
+Proof. by rewrite ltr0n indexg_gt0. Qed.
+
+Lemma natr_indexg_neq0 G B : #|G : B|%:R != 0 :> R.
+Proof. by rewrite gtr_eqF // natr_indexg_gt0. Qed.
+
+End FinGroup.
+
+Section NumFieldTheory.
+
+Variable F : numFieldType.
+Implicit Types x y z t : F.
+
+Lemma unitf_gt0 x : 0 < x -> x \is a GRing.unit.
+Proof. by move=> hx; rewrite unitfE eq_sym ltr_eqF. Qed.
+
+Lemma unitf_lt0 x : x < 0 -> x \is a GRing.unit.
+Proof. by move=> hx; rewrite unitfE ltr_eqF. Qed.
+
+Lemma lef_pinv : {in pos &, {mono (@GRing.inv F) : x y /~ x <= y}}.
+Proof. by move=> x y hx hy /=; rewrite ler_pinv ?inE ?unitf_gt0. Qed.
+
+Lemma lef_ninv : {in neg &, {mono (@GRing.inv F) : x y /~ x <= y}}.
+Proof. by move=> x y hx hy /=; rewrite ler_ninv ?inE ?unitf_lt0. Qed.
+
+Lemma ltf_pinv : {in pos &, {mono (@GRing.inv F) : x y /~ x < y}}.
+Proof. exact: lerW_nmono_in lef_pinv. Qed.
+
+Lemma ltf_ninv: {in neg &, {mono (@GRing.inv F) : x y /~ x < y}}.
+Proof. exact: lerW_nmono_in lef_ninv. Qed.
+
+Definition ltef_pinv := (lef_pinv, ltf_pinv).
+Definition ltef_ninv := (lef_ninv, ltf_ninv).
+
+Lemma invf_gt1 x : 0 < x -> (1 < x^-1) = (x < 1).
+Proof. by move=> x_gt0; rewrite -{1}[1]invr1 ltf_pinv ?posrE ?ltr01. Qed.
+
+Lemma invf_ge1 x : 0 < x -> (1 <= x^-1) = (x <= 1).
+Proof. by move=> x_lt0; rewrite -{1}[1]invr1 lef_pinv ?posrE ?ltr01. Qed.
+
+Definition invf_gte1 := (invf_ge1, invf_gt1).
+
+Lemma invf_le1 x : 0 < x -> (x^-1 <= 1) = (1 <= x).
+Proof. by move=> x_gt0; rewrite -invf_ge1 ?invr_gt0 // invrK. Qed.
+
+Lemma invf_lt1 x : 0 < x -> (x^-1 < 1) = (1 < x).
+Proof. by move=> x_lt0; rewrite -invf_gt1 ?invr_gt0 // invrK. Qed.
+
+Definition invf_lte1 := (invf_le1, invf_lt1).
+Definition invf_cp1 := (invf_gte1, invf_lte1).
+
+(* These lemma are all combinations of mono(LR|RL) with ler_[pn]mul2[rl]. *)
+Lemma ler_pdivl_mulr z x y : 0 < z -> (x <= y / z) = (x * z <= y).
+Proof. by move=> z_gt0; rewrite -(@ler_pmul2r _ z) ?mulfVK ?gtr_eqF. Qed.
+
+Lemma ltr_pdivl_mulr z x y : 0 < z -> (x < y / z) = (x * z < y).
+Proof. by move=> z_gt0; rewrite -(@ltr_pmul2r _ z) ?mulfVK ?gtr_eqF. Qed.
+
+Definition lter_pdivl_mulr := (ler_pdivl_mulr, ltr_pdivl_mulr).
+
+Lemma ler_pdivr_mulr z x y : 0 < z -> (y / z <= x) = (y <= x * z).
+Proof. by move=> z_gt0; rewrite -(@ler_pmul2r _ z) ?mulfVK ?gtr_eqF. Qed.
+
+Lemma ltr_pdivr_mulr z x y : 0 < z -> (y / z < x) = (y < x * z).
+Proof. by move=> z_gt0; rewrite -(@ltr_pmul2r _ z) ?mulfVK ?gtr_eqF. Qed.
+
+Definition lter_pdivr_mulr := (ler_pdivr_mulr, ltr_pdivr_mulr).
+
+Lemma ler_pdivl_mull z x y : 0 < z -> (x <= z^-1 * y) = (z * x <= y).
+Proof. by move=> z_gt0; rewrite mulrC ler_pdivl_mulr ?[z * _]mulrC. Qed.
+
+Lemma ltr_pdivl_mull z x y : 0 < z -> (x < z^-1 * y) = (z * x < y).
+Proof. by move=> z_gt0; rewrite mulrC ltr_pdivl_mulr ?[z * _]mulrC. Qed.
+
+Definition lter_pdivl_mull := (ler_pdivl_mull, ltr_pdivl_mull).
+
+Lemma ler_pdivr_mull z x y : 0 < z -> (z^-1 * y <= x) = (y <= z * x).
+Proof. by move=> z_gt0; rewrite mulrC ler_pdivr_mulr ?[z * _]mulrC. Qed.
+
+Lemma ltr_pdivr_mull z x y : 0 < z -> (z^-1 * y < x) = (y < z * x).
+Proof. by move=> z_gt0; rewrite mulrC ltr_pdivr_mulr ?[z * _]mulrC. Qed.
+
+Definition lter_pdivr_mull := (ler_pdivr_mull, ltr_pdivr_mull).
+
+Lemma ler_ndivl_mulr z x y : z < 0 -> (x <= y / z) = (y <= x * z).
+Proof. by move=> z_lt0; rewrite -(@ler_nmul2r _ z) ?mulfVK ?ltr_eqF. Qed.
+
+Lemma ltr_ndivl_mulr z x y : z < 0 -> (x < y / z) = (y < x * z).
+Proof. by move=> z_lt0; rewrite -(@ltr_nmul2r _ z) ?mulfVK ?ltr_eqF. Qed.
+
+Definition lter_ndivl_mulr := (ler_ndivl_mulr, ltr_ndivl_mulr).
+
+Lemma ler_ndivr_mulr z x y : z < 0 -> (y / z <= x) = (x * z <= y).
+Proof. by move=> z_lt0; rewrite -(@ler_nmul2r _ z) ?mulfVK ?ltr_eqF. Qed.
+
+Lemma ltr_ndivr_mulr z x y : z < 0 -> (y / z < x) = (x * z < y).
+Proof. by move=> z_lt0; rewrite -(@ltr_nmul2r _ z) ?mulfVK ?ltr_eqF. Qed.
+
+Definition lter_ndivr_mulr := (ler_ndivr_mulr, ltr_ndivr_mulr).
+
+Lemma ler_ndivl_mull z x y : z < 0 -> (x <= z^-1 * y) = (y <= z * x).
+Proof. by move=> z_lt0; rewrite mulrC ler_ndivl_mulr ?[z * _]mulrC. Qed.
+
+Lemma ltr_ndivl_mull z x y : z < 0 -> (x < z^-1 * y) = (y < z * x).
+Proof. by move=> z_lt0; rewrite mulrC ltr_ndivl_mulr ?[z * _]mulrC. Qed.
+
+Definition lter_ndivl_mull := (ler_ndivl_mull, ltr_ndivl_mull).
+
+Lemma ler_ndivr_mull z x y : z < 0 -> (z^-1 * y <= x) = (z * x <= y).
+Proof. by move=> z_lt0; rewrite mulrC ler_ndivr_mulr ?[z * _]mulrC. Qed.
+
+Lemma ltr_ndivr_mull z x y : z < 0 -> (z^-1 * y < x) = (z * x < y).
+Proof. by move=> z_lt0; rewrite mulrC ltr_ndivr_mulr ?[z * _]mulrC. Qed.
+
+Definition lter_ndivr_mull := (ler_ndivr_mull, ltr_ndivr_mull).
+
+Lemma natf_div m d : (d %| m)%N -> (m %/ d)%:R = m%:R / d%:R :> F.
+Proof. by apply: char0_natf_div; apply: (@char_num F). Qed.
+
+Lemma normfV : {morph (@norm F) : x / x ^-1}.
+Proof.
+move=> x /=; have [/normrV //|Nux] := boolP (x \is a GRing.unit).
+by rewrite !invr_out // unitfE normr_eq0 -unitfE.
+Qed.
+
+Lemma normf_div : {morph (@norm F) : x y / x / y}.
+Proof. by move=> x y /=; rewrite normrM normfV. Qed.
+
+Lemma invr_sg x : (sg x)^-1 = sgr x.
+Proof. by rewrite !(fun_if GRing.inv) !(invr0, invrN, invr1). Qed.
+
+Lemma sgrV x : sgr x^-1 = sgr x.
+Proof. by rewrite /sgr invr_eq0 invr_lt0. Qed.
+
+(* Interval midpoint. *)
+
+Local Notation mid x y := ((x + y) / 2%:R).
+
+Lemma midf_le x y : x <= y -> (x <= mid x y) * (mid x y <= y).
+Proof.
+move=> lexy; rewrite ler_pdivl_mulr ?ler_pdivr_mulr ?ltr0Sn //.
+by rewrite !mulrDr !mulr1 ler_add2r ler_add2l.
+Qed.
+
+Lemma midf_lt x y : x < y -> (x < mid x y) * (mid x y < y).
+Proof.
+move=> ltxy; rewrite ltr_pdivl_mulr ?ltr_pdivr_mulr ?ltr0Sn //.
+by rewrite !mulrDr !mulr1 ltr_add2r ltr_add2l.
+Qed.
+
+Definition midf_lte := (midf_le, midf_lt).
+
+(* The AGM, unscaled but without the nth root. *)
+
+Lemma real_lerif_mean_square x y :
+ x \is real -> y \is real -> x * y <= mid (x ^+ 2) (y ^+ 2) ?= iff (x == y).
+Proof.
+move=> Rx Ry; rewrite -(mono_lerif (ler_pmul2r (ltr_nat F 0 2))).
+by rewrite divfK ?pnatr_eq0 // mulr_natr; apply: real_lerif_mean_square_scaled.
+Qed.
+
+Lemma real_lerif_AGM2 x y :
+ x \is real -> y \is real -> x * y <= mid x y ^+ 2 ?= iff (x == y).
+Proof.
+move=> Rx Ry; rewrite -(mono_lerif (ler_pmul2r (ltr_nat F 0 4))).
+rewrite mulr_natr (natrX F 2 2) -exprMn divfK ?pnatr_eq0 //.
+exact: real_lerif_AGM2_scaled.
+Qed.
+
+Lemma lerif_AGM (I : finType) (A : pred I) (E : I -> F) :
+ let n := #|A| in let mu := (\sum_(i in A) E i) / n%:R in
+ {in A, forall i, 0 <= E i} ->
+ \prod_(i in A) E i <= mu ^+ n
+ ?= iff [forall i in A, forall j in A, E i == E j].
+Proof.
+move=> n mu Ege0; have [n0 | n_gt0] := posnP n.
+ by rewrite n0 -big_andE !(big_pred0 _ _ _ _ (card0_eq n0)); apply/lerifP.
+pose E' i := E i / n%:R.
+have defE' i: E' i *+ n = E i by rewrite -mulr_natr divfK ?pnatr_eq0 -?lt0n.
+have /lerif_AGM_scaled (i): i \in A -> 0 <= E' i *+ n by rewrite defE' => /Ege0.
+rewrite -/n -mulr_suml (eq_bigr _ (in1W defE')); congr (_ <= _ ?= iff _).
+by do 2![apply: eq_forallb_in => ? _]; rewrite -(eqr_pmuln2r n_gt0) !defE'.
+Qed.
+
+Implicit Type p : {poly F}.
+Lemma Cauchy_root_bound p : p != 0 -> {b | forall x, root p x -> `|x| <= b}.
+Proof.
+move=> nz_p; set a := lead_coef p; set n := (size p).-1.
+have [q Dp]: {q | forall x, x != 0 -> p.[x] = (a - q.[x^-1] / x) * x ^+ n}.
+ exists (- \poly_(i < n) p`_(n - i.+1)) => x nz_x.
+ rewrite hornerN mulNr opprK horner_poly mulrDl !mulr_suml addrC.
+ rewrite horner_coef polySpred // big_ord_recr (reindex_inj rev_ord_inj) /=.
+ rewrite -/n -lead_coefE; congr (_ + _); apply: eq_bigr=> i _.
+ by rewrite exprB ?unitfE // -exprVn mulrA mulrAC exprSr mulrA.
+have [b ub_q] := poly_disk_bound q 1; exists (b / `|a| + 1) => x px0.
+have b_ge0: 0 <= b by rewrite (ler_trans (normr_ge0 q.[1])) ?ub_q ?normr1.
+have{b_ge0} ba_ge0: 0 <= b / `|a| by rewrite divr_ge0 ?normr_ge0.
+rewrite real_lerNgt ?rpredD ?rpred1 ?ger0_real ?normr_ge0 //.
+apply: contraL px0 => lb_x; rewrite rootE.
+have x_ge1: 1 <= `|x| by rewrite (ler_trans _ (ltrW lb_x)) // ler_paddl.
+have nz_x: x != 0 by rewrite -normr_gt0 (ltr_le_trans ltr01).
+rewrite {}Dp // mulf_neq0 ?expf_neq0 // subr_eq0 eq_sym.
+have: (b / `|a|) < `|x| by rewrite (ltr_trans _ lb_x) // ltr_spaddr ?ltr01.
+apply: contraTneq => /(canRL (divfK nz_x))Dax.
+rewrite ltr_pdivr_mulr ?normr_gt0 ?lead_coef_eq0 // mulrC -normrM -{}Dax.
+by rewrite ler_gtF // ub_q // normfV invf_le1 ?normr_gt0.
+Qed.
+
+Import GroupScope.
+
+Lemma natf_indexg (gT : finGroupType) (G H : {group gT}) :
+ H \subset G -> #|G : H|%:R = (#|G|%:R / #|H|%:R)%R :> F.
+Proof. by move=> sHG; rewrite -divgS // natf_div ?cardSg. Qed.
+
+End NumFieldTheory.
+
+Section RealDomainTheory.
+
+Hint Resolve lerr.
+
+Variable R : realDomainType.
+Implicit Types x y z t : R.
+
+Lemma num_real x : x \is real. Proof. exact: num_real. Qed.
+Hint Resolve num_real.
+
+Lemma ler_total : total (@le R). Proof. by move=> x y; apply: real_leVge. Qed.
+
+Lemma ltr_total x y : x != y -> (x < y) || (y < x).
+Proof. by rewrite !ltr_def [_ == y]eq_sym => ->; apply: ler_total. Qed.
+
+Lemma wlog_ler P :
+ (forall a b, P b a -> P a b) -> (forall a b, a <= b -> P a b) ->
+ forall a b : R, P a b.
+Proof. by move=> sP hP a b; apply: real_wlog_ler. Qed.
+
+Lemma wlog_ltr P :
+ (forall a, P a a) ->
+ (forall a b, (P b a -> P a b)) -> (forall a b, a < b -> P a b) ->
+ forall a b : R, P a b.
+Proof. by move=> rP sP hP a b; apply: real_wlog_ltr. Qed.
+
+Lemma ltrNge x y : (x < y) = ~~ (y <= x). Proof. exact: real_ltrNge. Qed.
+
+Lemma lerNgt x y : (x <= y) = ~~ (y < x). Proof. exact: real_lerNgt. Qed.
+
+Lemma lerP x y : ler_xor_gt x y `|x - y| `|y - x| (x <= y) (y < x).
+Proof. exact: real_lerP. Qed.
+
+Lemma ltrP x y : ltr_xor_ge x y `|x - y| `|y - x| (y <= x) (x < y).
+Proof. exact: real_ltrP. Qed.
+
+Lemma ltrgtP x y :
+ comparer x y `|x - y| `|y - x| (y == x) (x == y)
+ (x <= y) (y <= x) (x < y) (x > y) .
+Proof. exact: real_ltrgtP. Qed.
+
+Lemma ger0P x : ger0_xor_lt0 x `|x| (x < 0) (0 <= x).
+Proof. exact: real_ger0P. Qed.
+
+Lemma ler0P x : ler0_xor_gt0 x `|x| (0 < x) (x <= 0).
+Proof. exact: real_ler0P. Qed.
+
+Lemma ltrgt0P x :
+ comparer0 x `|x| (0 == x) (x == 0) (x <= 0) (0 <= x) (x < 0) (x > 0).
+Proof. exact: real_ltrgt0P. Qed.
+
+Lemma neqr_lt x y : (x != y) = (x < y) || (y < x).
+Proof. exact: real_neqr_lt. Qed.
+
+Lemma eqr_leLR x y z t :
+ (x <= y -> z <= t) -> (y < x -> t < z) -> (x <= y) = (z <= t).
+Proof. by move=> *; apply/idP/idP; rewrite // !lerNgt; apply: contra. Qed.
+
+Lemma eqr_leRL x y z t :
+ (x <= y -> z <= t) -> (y < x -> t < z) -> (z <= t) = (x <= y).
+Proof. by move=> *; symmetry; apply: eqr_leLR. Qed.
+
+Lemma eqr_ltLR x y z t :
+ (x < y -> z < t) -> (y <= x -> t <= z) -> (x < y) = (z < t).
+Proof. by move=> *; rewrite !ltrNge; congr negb; apply: eqr_leLR. Qed.
+
+Lemma eqr_ltRL x y z t :
+ (x < y -> z < t) -> (y <= x -> t <= z) -> (z < t) = (x < y).
+Proof. by move=> *; symmetry; apply: eqr_ltLR. Qed.
+
+(* sign *)
+
+Lemma mulr_lt0 x y :
+ (x * y < 0) = [&& x != 0, y != 0 & (x < 0) (+) (y < 0)].
+Proof.
+have [x_gt0|x_lt0|->] /= := ltrgt0P x; last by rewrite mul0r.
+ by rewrite pmulr_rlt0 //; case: ltrgt0P.
+by rewrite nmulr_rlt0 //; case: ltrgt0P.
+Qed.
+
+Lemma neq0_mulr_lt0 x y :
+ x != 0 -> y != 0 -> (x * y < 0) = (x < 0) (+) (y < 0).
+Proof. by move=> x_neq0 y_neq0; rewrite mulr_lt0 x_neq0 y_neq0. Qed.
+
+Lemma mulr_sign_lt0 (b : bool) x :
+ ((-1) ^+ b * x < 0) = (x != 0) && (b (+) (x < 0)%R).
+Proof. by rewrite mulr_lt0 signr_lt0 signr_eq0. Qed.
+
+(* sign & norm*)
+
+Lemma mulr_sign_norm x : (-1) ^+ (x < 0)%R * `|x| = x.
+Proof. by rewrite real_mulr_sign_norm. Qed.
+
+Lemma mulr_Nsign_norm x : (-1) ^+ (0 < x)%R * `|x| = - x.
+Proof. by rewrite real_mulr_Nsign_norm. Qed.
+
+Lemma numEsign x : x = (-1) ^+ (x < 0)%R * `|x|.
+Proof. by rewrite -realEsign. Qed.
+
+Lemma numNEsign x : -x = (-1) ^+ (0 < x)%R * `|x|.
+Proof. by rewrite -realNEsign. Qed.
+
+Lemma normrEsign x : `|x| = (-1) ^+ (x < 0)%R * x.
+Proof. by rewrite -real_normrEsign. Qed.
+
+End RealDomainTheory.
+
+Hint Resolve num_real.
+
+Section RealDomainMonotony.
+
+Variables (R : realDomainType) (R' : numDomainType) (D : pred R) (f : R -> R').
+Implicit Types (m n p : nat) (x y z : R) (u v w : R').
+
+Hint Resolve (@num_real R).
+
+Lemma homo_mono : {homo f : x y / x < y} -> {mono f : x y / x <= y}.
+Proof. by move=> mf x y; apply: real_mono. Qed.
+
+Lemma nhomo_mono : {homo f : x y /~ x < y} -> {mono f : x y /~ x <= y}.
+Proof. by move=> mf x y; apply: real_nmono. Qed.
+
+Lemma homo_mono_in :
+ {in D &, {homo f : x y / x < y}} -> {in D &, {mono f : x y / x <= y}}.
+Proof.
+by move=> mf x y Dx Dy; apply: (real_mono_in mf); rewrite ?inE ?Dx ?Dy /=.
+Qed.
+
+Lemma nhomo_mono_in :
+ {in D &, {homo f : x y /~ x < y}} -> {in D &, {mono f : x y /~ x <= y}}.
+Proof.
+by move=> mf x y Dx Dy; apply: (real_nmono_in mf); rewrite ?inE ?Dx ?Dy /=.
+Qed.
+
+End RealDomainMonotony.
+
+Section RealDomainOperations.
+
+(* sgr section *)
+
+Variable R : realDomainType.
+Implicit Types x y z t : R.
+Hint Resolve (@num_real R).
+
+Lemma sgr_cp0 x :
+ ((sg x == 1) = (0 < x)) *
+ ((sg x == -1) = (x < 0)) *
+ ((sg x == 0) = (x == 0)).
+Proof.
+rewrite -[1]/((-1) ^+ false) -signrN lt0r lerNgt sgr_def.
+case: (x =P 0) => [-> | _]; first by rewrite !(eq_sym 0) !signr_eq0 ltrr eqxx.
+by rewrite !(inj_eq signr_inj) eqb_id eqbF_neg signr_eq0 //.
+Qed.
+
+CoInductive sgr_val x : R -> bool -> bool -> bool -> bool -> bool -> bool
+ -> bool -> bool -> bool -> bool -> bool -> bool -> R -> Set :=
+ | SgrNull of x = 0 : sgr_val x 0 true true true true false false
+ true false false true false false 0
+ | SgrPos of x > 0 : sgr_val x x false false true false false true
+ false false true false false true 1
+ | SgrNeg of x < 0 : sgr_val x (- x) false true false false true false
+ false true false false true false (-1).
+
+Lemma sgrP x :
+ sgr_val x `|x| (0 == x) (x <= 0) (0 <= x) (x == 0) (x < 0) (0 < x)
+ (0 == sg x) (-1 == sg x) (1 == sg x)
+ (sg x == 0) (sg x == -1) (sg x == 1) (sg x).
+Proof.
+by rewrite ![_ == sg _]eq_sym !sgr_cp0 /sg; case: ltrgt0P; constructor.
+Qed.
+
+Lemma normrEsg x : `|x| = sg x * x.
+Proof. by case: sgrP; rewrite ?(mul0r, mul1r, mulN1r). Qed.
+
+Lemma numEsg x : x = sg x * `|x|.
+Proof. by case: sgrP; rewrite !(mul1r, mul0r, mulrNN). Qed.
+
+(* GG: duplicate! *)
+Lemma mulr_sg_norm x : sg x * `|x| = x. Proof. by rewrite -numEsg. Qed.
+
+Lemma sgrM x y : sg (x * y) = sg x * sg y.
+Proof.
+rewrite !sgr_def mulr_lt0 andbA mulrnAr mulrnAl -mulrnA mulnb -negb_or mulf_eq0.
+by case: (~~ _) => //; rewrite signr_addb.
+Qed.
+
+Lemma sgrN x : sg (- x) = - sg x.
+Proof. by rewrite -mulrN1 sgrM sgrN1 mulrN1. Qed.
+
+Lemma sgrX n x : sg (x ^+ n) = (sg x) ^+ n.
+Proof. by elim: n => [|n IHn]; rewrite ?sgr1 // !exprS sgrM IHn. Qed.
+
+Lemma sgr_smul x y : sg (sg x * y) = sg x * sg y.
+Proof. by rewrite sgrM sgr_id. Qed.
+
+Lemma sgr_gt0 x : (sg x > 0) = (x > 0).
+Proof. by rewrite -sgr_cp0 sgr_id sgr_cp0. Qed.
+
+Lemma sgr_ge0 x : (sgr x >= 0) = (x >= 0).
+Proof. by rewrite !lerNgt sgr_lt0. Qed.
+
+(* norm section *)
+
+Lemma ler_norm x : (x <= `|x|).
+Proof. exact: real_ler_norm. Qed.
+
+Lemma ler_norml x y : (`|x| <= y) = (- y <= x <= y).
+Proof. exact: real_ler_norml. Qed.
+
+Lemma ler_normlP x y : reflect ((- x <= y) * (x <= y)) (`|x| <= y).
+Proof. exact: real_ler_normlP. Qed.
+Implicit Arguments ler_normlP [x y].
+
+Lemma eqr_norml x y : (`|x| == y) = ((x == y) || (x == -y)) && (0 <= y).
+Proof. exact: real_eqr_norml. Qed.
+
+Lemma eqr_norm2 x y : (`|x| == `|y|) = (x == y) || (x == -y).
+Proof. exact: real_eqr_norm2. Qed.
+
+Lemma ltr_norml x y : (`|x| < y) = (- y < x < y).
+Proof. exact: real_ltr_norml. Qed.
+
+Definition lter_norml := (ler_norml, ltr_norml).
+
+Lemma ltr_normlP x y : reflect ((-x < y) * (x < y)) (`|x| < y).
+Proof. exact: real_ltr_normlP. Qed.
+Implicit Arguments ltr_normlP [x y].
+
+Lemma ler_normr x y : (x <= `|y|) = (x <= y) || (x <= - y).
+Proof. by rewrite lerNgt ltr_norml negb_and -!lerNgt orbC ler_oppr. Qed.
+
+Lemma ltr_normr x y : (x < `|y|) = (x < y) || (x < - y).
+Proof. by rewrite ltrNge ler_norml negb_and -!ltrNge orbC ltr_oppr. Qed.
+
+Definition lter_normr := (ler_normr, ltr_normr).
+
+Lemma ler_distl x y e : (`|x - y| <= e) = (y - e <= x <= y + e).
+Proof. by rewrite lter_norml !lter_sub_addl. Qed.
+
+Lemma ltr_distl x y e : (`|x - y| < e) = (y - e < x < y + e).
+Proof. by rewrite lter_norml !lter_sub_addl. Qed.
+
+Definition lter_distl := (ler_distl, ltr_distl).
+
+Lemma exprn_even_ge0 n x : ~~ odd n -> 0 <= x ^+ n.
+Proof. by move=> even_n; rewrite real_exprn_even_ge0 ?num_real. Qed.
+
+Lemma exprn_even_gt0 n x : ~~ odd n -> (0 < x ^+ n) = (n == 0)%N || (x != 0).
+Proof. by move=> even_n; rewrite real_exprn_even_gt0 ?num_real. Qed.
+
+Lemma exprn_even_le0 n x : ~~ odd n -> (x ^+ n <= 0) = (n != 0%N) && (x == 0).
+Proof. by move=> even_n; rewrite real_exprn_even_le0 ?num_real. Qed.
+
+Lemma exprn_even_lt0 n x : ~~ odd n -> (x ^+ n < 0) = false.
+Proof. by move=> even_n; rewrite real_exprn_even_lt0 ?num_real. Qed.
+
+Lemma exprn_odd_ge0 n x : odd n -> (0 <= x ^+ n) = (0 <= x).
+Proof. by move=> even_n; rewrite real_exprn_odd_ge0 ?num_real. Qed.
+
+Lemma exprn_odd_gt0 n x : odd n -> (0 < x ^+ n) = (0 < x).
+Proof. by move=> even_n; rewrite real_exprn_odd_gt0 ?num_real. Qed.
+
+Lemma exprn_odd_le0 n x : odd n -> (x ^+ n <= 0) = (x <= 0).
+Proof. by move=> even_n; rewrite real_exprn_odd_le0 ?num_real. Qed.
+
+Lemma exprn_odd_lt0 n x : odd n -> (x ^+ n < 0) = (x < 0).
+Proof. by move=> even_n; rewrite real_exprn_odd_lt0 ?num_real. Qed.
+
+(* Special lemmas for squares. *)
+
+Lemma sqr_ge0 x : 0 <= x ^+ 2. Proof. by rewrite exprn_even_ge0. Qed.
+
+Lemma sqr_norm_eq1 x : (x ^+ 2 == 1) = (`|x| == 1).
+Proof. by rewrite sqrf_eq1 eqr_norml ler01 andbT. Qed.
+
+Lemma lerif_mean_square_scaled x y :
+ x * y *+ 2 <= x ^+ 2 + y ^+ 2 ?= iff (x == y).
+Proof. exact: real_lerif_mean_square_scaled. Qed.
+
+Lemma lerif_AGM2_scaled x y : x * y *+ 4 <= (x + y) ^+ 2 ?= iff (x == y).
+Proof. exact: real_lerif_AGM2_scaled. Qed.
+
+Section MinMax.
+
+(* GG: Many of the first lemmas hold unconditionally, and others hold for *)
+(* the real subset of a general domain. *)
+Lemma minrC : @commutative R R min.
+Proof. by move=> x y; rewrite /min; case: ltrgtP. Qed.
+
+Lemma minrr : @idempotent R min.
+Proof. by move=> x; rewrite /min if_same. Qed.
+
+Lemma minr_l x y : x <= y -> min x y = x.
+Proof. by rewrite /minr => ->. Qed.
+
+Lemma minr_r x y : y <= x -> min x y = y.
+Proof. by move/minr_l; rewrite minrC. Qed.
+
+Lemma maxrC : @commutative R R max.
+Proof. by move=> x y; rewrite /maxr; case: ltrgtP. Qed.
+
+Lemma maxrr : @idempotent R max.
+Proof. by move=> x; rewrite /max if_same. Qed.
+
+Lemma maxr_l x y : y <= x -> max x y = x.
+Proof. by move=> hxy; rewrite /max hxy. Qed.
+
+Lemma maxr_r x y : x <= y -> max x y = y.
+Proof. by move=> hxy; rewrite maxrC maxr_l. Qed.
+
+Lemma addr_min_max x y : min x y + max x y = x + y.
+Proof.
+case: (lerP x y)=> hxy; first by rewrite maxr_r ?minr_l.
+by rewrite maxr_l ?minr_r ?ltrW // addrC.
+Qed.
+
+Lemma addr_max_min x y : max x y + min x y = x + y.
+Proof. by rewrite addrC addr_min_max. Qed.
+
+Lemma minr_to_max x y : min x y = x + y - max x y.
+Proof. by rewrite -[x + y]addr_min_max addrK. Qed.
+
+Lemma maxr_to_min x y : max x y = x + y - min x y.
+Proof. by rewrite -[x + y]addr_max_min addrK. Qed.
+
+Lemma minrA x y z : min x (min y z) = min (min x y) z.
+Proof.
+rewrite /min; case: (lerP y z) => [hyz | /ltrW hyz].
+ by case: lerP => hxy; rewrite ?hyz // (@ler_trans _ y).
+case: lerP=> hxz; first by rewrite !(ler_trans hxz).
+case: (lerP x y)=> hxy; first by rewrite lerNgt hxz.
+by case: ltrgtP hyz.
+Qed.
+
+Lemma minrCA : @left_commutative R R min.
+Proof. by move=> x y z; rewrite !minrA [minr x y]minrC. Qed.
+
+Lemma minrAC : @right_commutative R R min.
+Proof. by move=> x y z; rewrite -!minrA [minr y z]minrC. Qed.
+
+CoInductive minr_spec x y : bool -> bool -> R -> Type :=
+| Minr_r of x <= y : minr_spec x y true false x
+| Minr_l of y < x : minr_spec x y false true y.
+
+Lemma minrP x y : minr_spec x y (x <= y) (y < x) (min x y).
+Proof.
+case: lerP=> hxy; first by rewrite minr_l //; constructor.
+by rewrite minr_r 1?ltrW //; constructor.
+Qed.
+
+Lemma oppr_max x y : - max x y = min (- x) (- y).
+Proof.
+case: minrP; rewrite lter_opp2 => hxy; first by rewrite maxr_l.
+by rewrite maxr_r // ltrW.
+Qed.
+
+Lemma oppr_min x y : - min x y = max (- x) (- y).
+Proof. by rewrite -[maxr _ _]opprK oppr_max !opprK. Qed.
+
+Lemma maxrA x y z : max x (max y z) = max (max x y) z.
+Proof. by apply/eqP; rewrite -eqr_opp !oppr_max minrA. Qed.
+
+Lemma maxrCA : @left_commutative R R max.
+Proof. by move=> x y z; rewrite !maxrA [maxr x y]maxrC. Qed.
+
+Lemma maxrAC : @right_commutative R R max.
+Proof. by move=> x y z; rewrite -!maxrA [maxr y z]maxrC. Qed.
+
+CoInductive maxr_spec x y : bool -> bool -> R -> Type :=
+| Maxr_r of y <= x : maxr_spec x y true false x
+| Maxr_l of x < y : maxr_spec x y false true y.
+
+Lemma maxrP x y : maxr_spec x y (y <= x) (x < y) (maxr x y).
+Proof.
+case: lerP => hxy; first by rewrite maxr_l //; constructor.
+by rewrite maxr_r 1?ltrW //; constructor.
+Qed.
+
+Lemma eqr_minl x y : (min x y == x) = (x <= y).
+Proof. by case: minrP=> hxy; rewrite ?eqxx // ltr_eqF. Qed.
+
+Lemma eqr_minr x y : (min x y == y) = (y <= x).
+Proof. by rewrite minrC eqr_minl. Qed.
+
+Lemma eqr_maxl x y : (max x y == x) = (y <= x).
+Proof. by case: maxrP=> hxy; rewrite ?eqxx // eq_sym ltr_eqF. Qed.
+
+Lemma eqr_maxr x y : (max x y == y) = (x <= y).
+Proof. by rewrite maxrC eqr_maxl. Qed.
+
+Lemma ler_minr x y z : (x <= min y z) = (x <= y) && (x <= z).
+Proof.
+case: minrP=> hyz.
+ by case: lerP=> hxy //; rewrite (ler_trans _ hyz).
+by case: lerP=> hxz; rewrite andbC // (ler_trans hxz) // ltrW.
+Qed.
+
+Lemma ler_minl x y z : (min y z <= x) = (y <= x) || (z <= x).
+Proof.
+case: minrP => hyz.
+ case: lerP => hyx //=; symmetry; apply: negbTE.
+ by rewrite -ltrNge (@ltr_le_trans _ y).
+case: lerP => hzx; rewrite orbC //=; symmetry; apply: negbTE.
+by rewrite -ltrNge (@ltr_trans _ z).
+Qed.
+
+Lemma ler_maxr x y z : (x <= max y z) = (x <= y) || (x <= z).
+Proof. by rewrite -lter_opp2 oppr_max ler_minl !ler_opp2. Qed.
+
+Lemma ler_maxl x y z : (max y z <= x) = (y <= x) && (z <= x).
+Proof. by rewrite -lter_opp2 oppr_max ler_minr !ler_opp2. Qed.
+
+Lemma ltr_minr x y z : (x < min y z) = (x < y) && (x < z).
+Proof. by rewrite !ltrNge ler_minl negb_or. Qed.
+
+Lemma ltr_minl x y z : (min y z < x) = (y < x) || (z < x).
+Proof. by rewrite !ltrNge ler_minr negb_and. Qed.
+
+Lemma ltr_maxr x y z : (x < max y z) = (x < y) || (x < z).
+Proof. by rewrite !ltrNge ler_maxl negb_and. Qed.
+
+Lemma ltr_maxl x y z : (max y z < x) = (y < x) && (z < x).
+Proof. by rewrite !ltrNge ler_maxr negb_or. Qed.
+
+Definition lter_minr := (ler_minr, ltr_minr).
+Definition lter_minl := (ler_minl, ltr_minl).
+Definition lter_maxr := (ler_maxr, ltr_maxr).
+Definition lter_maxl := (ler_maxl, ltr_maxl).
+
+Lemma addr_minl : @left_distributive R R +%R min.
+Proof.
+move=> x y z; case: minrP=> hxy; first by rewrite minr_l // ler_add2r.
+by rewrite minr_r // ltrW // ltr_add2r.
+Qed.
+
+Lemma addr_minr : @right_distributive R R +%R min.
+Proof.
+move=> x y z; case: minrP=> hxy; first by rewrite minr_l // ler_add2l.
+by rewrite minr_r // ltrW // ltr_add2l.
+Qed.
+
+Lemma addr_maxl : @left_distributive R R +%R max.
+Proof.
+move=> x y z; rewrite -[_ + _]opprK opprD oppr_max.
+by rewrite addr_minl -!opprD oppr_min !opprK.
+Qed.
+
+Lemma addr_maxr : @right_distributive R R +%R max.
+Proof.
+move=> x y z; rewrite -[_ + _]opprK opprD oppr_max.
+by rewrite addr_minr -!opprD oppr_min !opprK.
+Qed.
+
+Lemma minrK x y : max (min x y) x = x.
+Proof. by case: minrP => hxy; rewrite ?maxrr ?maxr_r // ltrW. Qed.
+
+Lemma minKr x y : min y (max x y) = y.
+Proof. by case: maxrP => hxy; rewrite ?minrr ?minr_l. Qed.
+
+Lemma maxr_minl : @left_distributive R R max min.
+Proof.
+move=> x y z; case: minrP => hxy.
+ by case: maxrP => hm; rewrite minr_l // ler_maxr (hxy, lerr) ?orbT.
+by case: maxrP => hyz; rewrite minr_r // ler_maxr (ltrW hxy, lerr) ?orbT.
+Qed.
+
+Lemma maxr_minr : @right_distributive R R max min.
+Proof. by move=> x y z; rewrite maxrC maxr_minl ![_ _ x]maxrC. Qed.
+
+Lemma minr_maxl : @left_distributive R R min max.
+Proof.
+move=> x y z; rewrite -[min _ _]opprK !oppr_min [- max x y]oppr_max.
+by rewrite maxr_minl !(oppr_max, oppr_min, opprK).
+Qed.
+
+Lemma minr_maxr : @right_distributive R R min max.
+Proof. by move=> x y z; rewrite minrC minr_maxl ![_ _ x]minrC. Qed.
+
+Lemma minr_pmulr x y z : 0 <= x -> x * min y z = min (x * y) (x * z).
+Proof.
+case: sgrP=> // hx _; first by rewrite hx !mul0r minrr.
+case: minrP=> hyz; first by rewrite minr_l // ler_pmul2l.
+by rewrite minr_r // ltrW // ltr_pmul2l.
+Qed.
+
+Lemma minr_nmulr x y z : x <= 0 -> x * min y z = max (x * y) (x * z).
+Proof.
+move=> hx; rewrite -[_ * _]opprK -mulNr minr_pmulr ?oppr_cp0 //.
+by rewrite oppr_min !mulNr !opprK.
+Qed.
+
+Lemma maxr_pmulr x y z : 0 <= x -> x * max y z = max (x * y) (x * z).
+Proof.
+move=> hx; rewrite -[_ * _]opprK -mulrN oppr_max minr_pmulr //.
+by rewrite oppr_min !mulrN !opprK.
+Qed.
+
+Lemma maxr_nmulr x y z : x <= 0 -> x * max y z = min (x * y) (x * z).
+Proof.
+move=> hx; rewrite -[_ * _]opprK -mulrN oppr_max minr_nmulr //.
+by rewrite oppr_max !mulrN !opprK.
+Qed.
+
+Lemma minr_pmull x y z : 0 <= x -> min y z * x = min (y * x) (z * x).
+Proof. by move=> *; rewrite mulrC minr_pmulr // ![_ * x]mulrC. Qed.
+
+Lemma minr_nmull x y z : x <= 0 -> min y z * x = max (y * x) (z * x).
+Proof. by move=> *; rewrite mulrC minr_nmulr // ![_ * x]mulrC. Qed.
+
+Lemma maxr_pmull x y z : 0 <= x -> max y z * x = max (y * x) (z * x).
+Proof. by move=> *; rewrite mulrC maxr_pmulr // ![_ * x]mulrC. Qed.
+
+Lemma maxr_nmull x y z : x <= 0 -> max y z * x = min (y * x) (z * x).
+Proof. by move=> *; rewrite mulrC maxr_nmulr // ![_ * x]mulrC. Qed.
+
+Lemma maxrN x : max x (- x) = `|x|.
+Proof.
+case: ger0P=> hx; first by rewrite maxr_l // ge0_cp //.
+by rewrite maxr_r // le0_cp // ltrW.
+Qed.
+
+Lemma maxNr x : max (- x) x = `|x|.
+Proof. by rewrite maxrC maxrN. Qed.
+
+Lemma minrN x : min x (- x) = - `|x|.
+Proof. by rewrite -[minr _ _]opprK oppr_min opprK maxNr. Qed.
+
+Lemma minNr x : min (- x) x = - `|x|.
+Proof. by rewrite -[minr _ _]opprK oppr_min opprK maxrN. Qed.
+
+End MinMax.
+
+Section PolyBounds.
+
+Variable p : {poly R}.
+
+Lemma poly_itv_bound a b : {ub | forall x, a <= x <= b -> `|p.[x]| <= ub}.
+Proof.
+have [ub le_p_ub] := poly_disk_bound p (Num.max `|a| `|b|).
+exists ub => x /andP[le_a_x le_x_b]; rewrite le_p_ub // ler_maxr !ler_normr.
+by have [_|_] := ler0P x; rewrite ?ler_opp2 ?le_a_x ?le_x_b orbT.
+Qed.
+
+Lemma monic_Cauchy_bound : p \is monic -> {b | forall x, x >= b -> p.[x] > 0}.
+Proof.
+move/monicP=> mon_p; pose n := (size p - 2)%N.
+have [p_le1 | p_gt1] := leqP (size p) 1.
+ exists 0 => x _; rewrite (size1_polyC p_le1) hornerC.
+ by rewrite -[p`_0]lead_coefC -size1_polyC // mon_p ltr01.
+pose lb := \sum_(j < n.+1) `|p`_j|; exists (lb + 1) => x le_ub_x.
+have x_ge1: 1 <= x; last have x_gt0 := ltr_le_trans ltr01 x_ge1.
+ by rewrite -(ler_add2l lb) ler_paddl ?sumr_ge0 // => j _; apply: normr_ge0.
+rewrite horner_coef -(subnK p_gt1) -/n addnS big_ord_recr /= addn1.
+rewrite [in p`__]subnSK // subn1 -lead_coefE mon_p mul1r -ltr_subl_addl sub0r.
+apply: ler_lt_trans (_ : lb * x ^+ n < _); last first.
+ rewrite exprS ltr_pmul2r ?exprn_gt0 ?(ltr_le_trans ltr01) //.
+ by rewrite -(ltr_add2r 1) ltr_spaddr ?ltr01.
+rewrite -sumrN mulr_suml ler_sum // => j _; apply: ler_trans (ler_norm _) _.
+rewrite normrN normrM ler_wpmul2l ?normr_ge0 // normrX.
+by rewrite ger0_norm ?(ltrW x_gt0) // ler_weexpn2l ?leq_ord.
+Qed.
+
+End PolyBounds.
+
+End RealDomainOperations.
+
+Section RealField.
+
+Variables (F : realFieldType) (x y : F).
+
+Lemma lerif_mean_square : x * y <= (x ^+ 2 + y ^+ 2) / 2%:R ?= iff (x == y).
+Proof. by apply: real_lerif_mean_square; apply: num_real. Qed.
+
+Lemma lerif_AGM2 : x * y <= ((x + y) / 2%:R)^+ 2 ?= iff (x == y).
+Proof. by apply: real_lerif_AGM2; apply: num_real. Qed.
+
+End RealField.
+
+Section ArchimedeanFieldTheory.
+
+Variables (F : archiFieldType) (x : F).
+
+Lemma archi_boundP : 0 <= x -> x < (bound x)%:R.
+Proof. by move/ger0_norm=> {1}<-; rewrite /bound; case: (sigW _). Qed.
+
+Lemma upper_nthrootP i : (bound x <= i)%N -> x < 2%:R ^+ i.
+Proof.
+rewrite /bound; case: (sigW _) => /= b le_x_b le_b_i.
+apply: ler_lt_trans (ler_norm x) (ltr_trans le_x_b _ ).
+by rewrite -natrX ltr_nat (leq_ltn_trans le_b_i) // ltn_expl.
+Qed.
+
+End ArchimedeanFieldTheory.
+
+Section RealClosedFieldTheory.
+
+Variable R : rcfType.
+Implicit Types a x y : R.
+
+Lemma poly_ivt : real_closed_axiom R. Proof. exact: poly_ivt. Qed.
+
+(* Square Root theory *)
+
+Lemma sqrtr_ge0 a : 0 <= sqrt a.
+Proof. by rewrite /sqrt; case: (sig2W _). Qed.
+Hint Resolve sqrtr_ge0.
+
+Lemma sqr_sqrtr a : 0 <= a -> sqrt a ^+ 2 = a.
+Proof.
+by rewrite /sqrt => a_ge0; case: (sig2W _) => /= x _; rewrite a_ge0 => /eqP.
+Qed.
+
+Lemma ler0_sqrtr a : a <= 0 -> sqrt a = 0.
+Proof.
+rewrite /sqrtr; case: (sig2W _) => x /= _.
+by have [//|_ /eqP//|->] := ltrgt0P a; rewrite mulf_eq0 orbb => /eqP.
+Qed.
+
+Lemma ltr0_sqrtr a : a < 0 -> sqrt a = 0.
+Proof. by move=> /ltrW; apply: ler0_sqrtr. Qed.
+
+CoInductive sqrtr_spec a : R -> bool -> bool -> R -> Type :=
+| IsNoSqrtr of a < 0 : sqrtr_spec a a false true 0
+| IsSqrtr b of 0 <= b : sqrtr_spec a (b ^+ 2) true false b.
+
+Lemma sqrtrP a : sqrtr_spec a a (0 <= a) (a < 0) (sqrt a).
+Proof.
+have [a_ge0|a_lt0] := ger0P a.
+ by rewrite -{1 2}[a]sqr_sqrtr //; constructor.
+by rewrite ltr0_sqrtr //; constructor.
+Qed.
+
+Lemma sqrtr_sqr a : sqrt (a ^+ 2) = `|a|.
+Proof.
+have /eqP : sqrt (a ^+ 2) ^+ 2 = `|a| ^+ 2.
+ by rewrite -normrX ger0_norm ?sqr_sqrtr ?sqr_ge0.
+rewrite eqf_sqr => /predU1P[-> //|ha].
+have := sqrtr_ge0 (a ^+ 2); rewrite (eqP ha) oppr_ge0 normr_le0 => /eqP ->.
+by rewrite normr0 oppr0.
+Qed.
+
+Lemma sqrtrM a b : 0 <= a -> sqrt (a * b) = sqrt a * sqrt b.
+Proof.
+case: (sqrtrP a) => // {a} a a_ge0 _; case: (sqrtrP b) => [b_lt0 | {b} b b_ge0].
+ by rewrite mulr0 ler0_sqrtr // nmulr_lle0 ?mulr_ge0.
+by rewrite mulrACA sqrtr_sqr ger0_norm ?mulr_ge0.
+Qed.
+
+Lemma sqrtr0 : sqrt 0 = 0 :> R.
+Proof. by move: (sqrtr_sqr 0); rewrite exprS mul0r => ->; rewrite normr0. Qed.
+
+Lemma sqrtr1 : sqrt 1 = 1 :> R.
+Proof. by move: (sqrtr_sqr 1); rewrite expr1n => ->; rewrite normr1. Qed.
+
+Lemma sqrtr_eq0 a : (sqrt a == 0) = (a <= 0).
+Proof.
+case: sqrtrP => [/ltrW ->|b]; first by rewrite eqxx.
+case: ltrgt0P => [b_gt0|//|->]; last by rewrite exprS mul0r lerr.
+by rewrite ltr_geF ?pmulr_rgt0.
+Qed.
+
+Lemma sqrtr_gt0 a : (0 < sqrt a) = (0 < a).
+Proof. by rewrite lt0r sqrtr_ge0 sqrtr_eq0 -ltrNge andbT. Qed.
+
+Lemma eqr_sqrt a b : 0 <= a -> 0 <= b -> (sqrt a == sqrt b) = (a == b).
+Proof.
+move=> a_ge0 b_ge0; apply/eqP/eqP=> [HS|->] //.
+by move: (sqr_sqrtr a_ge0); rewrite HS (sqr_sqrtr b_ge0).
+Qed.
+
+Lemma ler_wsqrtr : {homo @sqrt R : a b / a <= b}.
+Proof.
+move=> a b /= le_ab; case: (boolP (0 <= a))=> [pa|]; last first.
+ by rewrite -ltrNge; move/ltrW; rewrite -sqrtr_eq0; move/eqP->.
+rewrite -(@ler_pexpn2r R 2) ?nnegrE ?sqrtr_ge0 //.
+by rewrite !sqr_sqrtr // (ler_trans pa).
+Qed.
+
+Lemma ler_psqrt : {in @pos R &, {mono sqrt : a b / a <= b}}.
+Proof.
+apply: homo_mono_in => x y x_gt0 y_gt0.
+rewrite !ltr_neqAle => /andP[neq_xy le_xy].
+by rewrite ler_wsqrtr // eqr_sqrt ?ltrW // neq_xy.
+Qed.
+
+Lemma ler_sqrt a b : 0 < b -> (sqrt a <= sqrt b) = (a <= b).
+Proof.
+move=> b_gt0; have [a_le0|a_gt0] := ler0P a; last by rewrite ler_psqrt.
+by rewrite ler0_sqrtr // sqrtr_ge0 (ler_trans a_le0) ?ltrW.
+Qed.
+
+Lemma ltr_sqrt a b : 0 < b -> (sqrt a < sqrt b) = (a < b).
+Proof.
+move=> b_gt0; have [a_le0|a_gt0] := ler0P a; last first.
+ by rewrite (lerW_mono_in ler_psqrt).
+by rewrite ler0_sqrtr // sqrtr_gt0 b_gt0 (ler_lt_trans a_le0).
+Qed.
+
+End RealClosedFieldTheory.
+
+End Theory.
+
+Module RealMixin.
+
+Section RealMixins.
+
+Variables (R : idomainType) (le : rel R) (lt : rel R) (norm : R -> R).
+Local Infix "<=" := le.
+Local Infix "<" := lt.
+Local Notation "`| x |" := (norm x) : ring_scope.
+
+Section LeMixin.
+
+Hypothesis le0_add : forall x y, 0 <= x -> 0 <= y -> 0 <= x + y.
+Hypothesis le0_mul : forall x y, 0 <= x -> 0 <= y -> 0 <= x * y.
+Hypothesis le0_anti : forall x, 0 <= x -> x <= 0 -> x = 0.
+Hypothesis sub_ge0 : forall x y, (0 <= y - x) = (x <= y).
+Hypothesis le0_total : forall x, (0 <= x) || (x <= 0).
+Hypothesis normN: forall x, `|- x| = `|x|.
+Hypothesis ge0_norm : forall x, 0 <= x -> `|x| = x.
+Hypothesis lt_def : forall x y, (x < y) = (y != x) && (x <= y).
+
+Let le0N x : (0 <= - x) = (x <= 0). Proof. by rewrite -sub0r sub_ge0. Qed.
+Let leN_total x : 0 <= x \/ 0 <= - x.
+Proof. by apply/orP; rewrite le0N le0_total. Qed.
+
+Let le00 : (0 <= 0). Proof. by have:= le0_total 0; rewrite orbb. Qed.
+Let le01 : (0 <= 1).
+Proof.
+by case/orP: (le0_total 1)=> // ?; rewrite -[1]mul1r -mulrNN le0_mul ?le0N.
+Qed.
+
+Fact lt0_add x y : 0 < x -> 0 < y -> 0 < x + y.
+Proof.
+rewrite !lt_def => /andP[x_neq0 l0x] /andP[y_neq0 l0y]; rewrite le0_add //.
+rewrite andbT addr_eq0; apply: contraNneq x_neq0 => hxy.
+by rewrite [x]le0_anti // hxy -le0N opprK.
+Qed.
+
+Fact eq0_norm x : `|x| = 0 -> x = 0.
+Proof.
+case: (leN_total x) => /ge0_norm => [-> // | Dnx nx0].
+by rewrite -[x]opprK -Dnx normN nx0 oppr0.
+Qed.
+
+Fact le_def x y : (x <= y) = (`|y - x| == y - x).
+Proof.
+wlog ->: x y / x = 0 by move/(_ 0 (y - x)); rewrite subr0 sub_ge0 => ->.
+rewrite {x}subr0; apply/idP/eqP=> [/ge0_norm// | Dy].
+by have [//| ny_ge0] := leN_total y; rewrite -Dy -normN ge0_norm.
+Qed.
+
+Fact normM : {morph norm : x y / x * y}.
+Proof.
+move=> x y /=; wlog x_ge0 : x / 0 <= x.
+ by move=> IHx; case: (leN_total x) => /IHx//; rewrite mulNr !normN.
+wlog y_ge0 : y / 0 <= y; last by rewrite ?ge0_norm ?le0_mul.
+by move=> IHy; case: (leN_total y) => /IHy//; rewrite mulrN !normN.
+Qed.
+
+Fact le_normD x y : `|x + y| <= `|x| + `|y|.
+Proof.
+wlog x_ge0 : x y / 0 <= x.
+ by move=> IH; case: (leN_total x) => /IH// /(_ (- y)); rewrite -opprD !normN.
+rewrite -sub_ge0 ge0_norm //; have [y_ge0 | ny_ge0] := leN_total y.
+ by rewrite !ge0_norm ?subrr ?le0_add.
+rewrite -normN ge0_norm //; have [hxy|hxy] := leN_total (x + y).
+ by rewrite ge0_norm // opprD addrCA -addrA addKr le0_add.
+by rewrite -normN ge0_norm // opprK addrCA addrNK le0_add.
+Qed.
+
+Lemma le_total x y : (x <= y) || (y <= x).
+Proof. by rewrite -sub_ge0 -opprB le0N orbC -sub_ge0 le0_total. Qed.
+
+Definition Le :=
+ Mixin le_normD lt0_add eq0_norm (in2W le_total) normM le_def lt_def.
+
+Lemma Real (R' : numDomainType) & phant R' :
+ R' = NumDomainType R Le -> real_axiom R'.
+Proof. by move->. Qed.
+
+End LeMixin.
+
+Section LtMixin.
+
+Hypothesis lt0_add : forall x y, 0 < x -> 0 < y -> 0 < x + y.
+Hypothesis lt0_mul : forall x y, 0 < x -> 0 < y -> 0 < x * y.
+Hypothesis lt0_ngt0 : forall x, 0 < x -> ~~ (x < 0).
+Hypothesis sub_gt0 : forall x y, (0 < y - x) = (x < y).
+Hypothesis lt0_total : forall x, x != 0 -> (0 < x) || (x < 0).
+Hypothesis normN : forall x, `|- x| = `|x|.
+Hypothesis ge0_norm : forall x, 0 <= x -> `|x| = x.
+Hypothesis le_def : forall x y, (x <= y) = (y == x) || (x < y).
+
+Fact le0_add x y : 0 <= x -> 0 <= y -> 0 <= x + y.
+Proof.
+rewrite !le_def => /predU1P[->|x_gt0]; first by rewrite add0r.
+by case/predU1P=> [->|y_gt0]; rewrite ?addr0 ?x_gt0 ?lt0_add // orbT.
+Qed.
+
+Fact le0_mul x y : 0 <= x -> 0 <= y -> 0 <= x * y.
+Proof.
+rewrite !le_def => /predU1P[->|x_gt0]; first by rewrite mul0r eqxx.
+by case/predU1P=> [->|y_gt0]; rewrite ?mulr0 ?eqxx // orbC lt0_mul.
+Qed.
+
+Fact le0_anti x : 0 <= x -> x <= 0 -> x = 0.
+Proof. by rewrite !le_def => /predU1P[] // /lt0_ngt0/negPf-> /predU1P[]. Qed.
+
+Fact sub_ge0 x y : (0 <= y - x) = (x <= y).
+Proof. by rewrite !le_def subr_eq0 sub_gt0. Qed.
+
+Fact lt_def x y : (x < y) = (y != x) && (x <= y).
+Proof.
+rewrite le_def; case: eqP => //= ->; rewrite -sub_gt0 subrr.
+by apply/idP=> lt00; case/negP: (lt0_ngt0 lt00).
+Qed.
+
+Fact le0_total x : (0 <= x) || (x <= 0).
+Proof. by rewrite !le_def [0 == _]eq_sym; have [|/lt0_total] := altP eqP. Qed.
+
+Definition Lt :=
+ Le le0_add le0_mul le0_anti sub_ge0 le0_total normN ge0_norm lt_def.
+
+End LtMixin.
+
+End RealMixins.
+
+End RealMixin.
+
+End Num.
+
+Export Num.NumDomain.Exports Num.NumField.Exports Num.ClosedField.Exports.
+Export Num.RealDomain.Exports Num.RealField.Exports.
+Export Num.ArchimedeanField.Exports Num.RealClosedField.Exports.
+Export Num.Syntax Num.PredInstances.
+
+Notation RealLeMixin := Num.RealMixin.Le.
+Notation RealLtMixin := Num.RealMixin.Lt.
+Notation RealLeAxiom R := (Num.RealMixin.Real (Phant R) (erefl _)).
diff --git a/mathcomp/algebra/vector.v b/mathcomp/algebra/vector.v
new file mode 100644
index 0000000..2cb59e9
--- /dev/null
+++ b/mathcomp/algebra/vector.v
@@ -0,0 +1,2040 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype bigop.
+Require Import finfun tuple ssralg matrix mxalgebra zmodp.
+
+(******************************************************************************)
+(* * Finite dimensional vector spaces *)
+(* vectType R == interface structure for finite dimensional (more *)
+(* precisely, detachable) vector spaces over R, which *)
+(* should be at least a ringType. *)
+(* Vector.axiom n M <-> type M is linearly isomorphic to 'rV_n. *)
+(* := {v2r : M -> 'rV_n| linear v2r & bijective v2r}. *)
+(* VectMixin isoM == packages a proof isoV of Vector.axiom n M as the *)
+(* vectType mixin for an n-dimensional R-space *)
+(* structure on a type M that is an lmodType R. *)
+(* VectType K M mT == packs the vectType mixin mT to into a vectType K *)
+(* instance for T; T should have an lmodType K *)
+(* canonical instance. *)
+(* [vectType R of T for vS] == a copy of the vS : vectType R structure where *)
+(* the sort is replaced by T; vS : lmodType R should *)
+(* be convertible to a canonical lmodType for T. *)
+(* [vectType R of V] == a clone of an existing vectType R structure on V. *)
+(* {vspace vT} == the type of (detachable) subspaces of vT; vT *)
+(* should have a vectType structure over a fieldType. *)
+(* subvs_of U == the subtype of elements of V in the subspace U. *)
+(* This is canonically a vectType. *)
+(* vsval u == linear injection of u : subvs_of U into V. *)
+(* vsproj U v == linear projection of v : V in subvs U. *)
+(* 'Hom(aT, rT) == the type of linear functions (homomorphisms) from *)
+(* aT to rT, where aT and rT ARE vectType structures. *)
+(* Elements of 'Hom(aT, rT) coerce to Coq functions. *)
+(* --> Caveat: aT and rT must denote actual vectType structures, not their *)
+(* projections on Type. *)
+(* linfun f == a vector linear function in 'Hom(aT, rT) that *)
+(* coincides with f : aT -> rT when f is linear. *)
+(* 'End(vT) == endomorphisms of vT (:= 'Hom(vT, vT)). *)
+(* --> The types subvs_of U, 'Hom(aT, rT), 'End(vT), K^o, 'M[K]_(m, n), *)
+(* vT * wT, {ffun I -> vT}, vT ^ n all have canonical vectType instances. *)
+(* *)
+(* Functions: *)
+(* <[v]>%VS == the vector space generated by v (a line if v != 0).*)
+(* 0%VS == the trivial vector subspace. *)
+(* fullv, {:vT} == the complete vector subspace (displays as fullv). *)
+(* (U + V)%VS == the join (sum) of two subspaces U and V. *)
+(* (U :&: V)%VS == intersection of vector subspaces U and V. *)
+(* (U^C)%VS == a complement of the vector subspace U. *)
+(* (U :\: V)%VS == a local complement to U :& V in the subspace U. *)
+(* \dim U == dimension of a vector space U. *)
+(* span X, <<X>>%VS == the subspace spanned by the vector sequence X. *)
+(* coord X i v == i'th coordinate of v on X, when v \in <<X>>%VS and *)
+(* where X : n.-tuple vT and i : 'I_n. Note that *)
+(* coord X i is a scalar function. *)
+(* vpick U == a nonzero element of U if U= 0%VS, or 0 if U = 0. *)
+(* vbasis U == a (\dim U).-tuple that is a basis of U. *)
+(* \1%VF == the identity linear function. *)
+(* (f \o g)%VF == the composite of two linear functions f and g. *)
+(* (f^-1)%VF == a linear function that is a right inverse to the *)
+(* linear function f on the codomain of f. *)
+(* (f @: U)%VS == the image of vs by the linear function f. *)
+(* (f @^-1: U)%VS == the pre-image of vs by the linear function f. *)
+(* lker f == the kernel of the linear function f. *)
+(* limg f == the image of the linear function f. *)
+(* fixedSpace f == the fixed space of a linear endomorphism f *)
+(* daddv_pi U V == projection onto U along V if U and V are disjoint; *)
+(* daddv_pi U V + daddv_pi V U is then a projection *)
+(* onto the direct sum (U + V)%VS. *)
+(* projv U == projection onto U (along U^C, := daddv_pi U U^C). *)
+(* addv_pi1 U V == projection onto the subspace U :\: V of U along V. *)
+(* addv_pi2 U V == projection onto V along U :\: V; note that *)
+(* addv_pi1 U V and addv_pi2 U V are (asymmetrical) *)
+(* complementary projections on (U + V)%VS. *)
+(* sumv_pi_for defV i == for defV : V = (V \sum_(j <- r | P j) Vs j)%VS, *)
+(* j ranging over an eqType, this is a projection on *)
+(* a subspace of Vs i, along a complement in V, such *)
+(* that \sum_(j <- r | P j) sumv_pi_for defV j is a *)
+(* projection onto V if filter P r is duplicate-free *)
+(* (e.g., when V := \sum_(j | P j) Vs j). *)
+(* sumv_pi V i == notation the above when defV == erefl V, and V is *)
+(* convertible to \sum_(j <- r | P j) Vs j)%VS. *)
+(* *)
+(* Predicates: *)
+(* v \in U == v belongs to U (:= (<[v]> <= U)%VS). *)
+(* (U <= V)%VS == U is a subspace of V. *)
+(* free B == B is a sequence of nonzero linearly independent *)
+(* vectors. *)
+(* basis_of U b == b is a basis of the subspace U. *)
+(* directv S == S is the expression for a direct sum of subspaces. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+Open Local Scope ring_scope.
+
+Reserved Notation "{ 'vspace' T }" (at level 0, format "{ 'vspace' T }").
+Reserved Notation "''Hom' ( T , rT )" (at level 8, format "''Hom' ( T , rT )").
+Reserved Notation "''End' ( T )" (at level 8, format "''End' ( T )").
+Reserved Notation "\dim A" (at level 10, A at level 8, format "\dim A").
+
+Delimit Scope vspace_scope with VS.
+
+Import GRing.Theory.
+
+(* Finite dimension vector space *)
+Module Vector.
+
+Section ClassDef.
+Variable R : ringType.
+
+Definition axiom_def n (V : lmodType R) of phant V :=
+ {v2r : V -> 'rV[R]_n | linear v2r & bijective v2r}.
+
+Inductive mixin_of (V : lmodType R) := Mixin dim & axiom_def dim (Phant V).
+
+Structure class_of V := Class {
+ base : GRing.Lmodule.class_of R V;
+ mixin : mixin_of (GRing.Lmodule.Pack _ base V)
+}.
+Local Coercion base : class_of >-> GRing.Lmodule.class_of.
+
+Structure type (phR : phant R) := Pack {sort; _ : class_of sort; _ : Type}.
+Local Coercion sort : type >-> Sortclass.
+Variables (phR : phant R) (T : Type) (cT : type phR).
+Definition class := let: Pack _ c _ := cT return class_of cT in c.
+Definition clone c of phant_id class c := @Pack phR T c T.
+Let xT := let: Pack T _ _ := cT in T.
+Notation xclass := (class : class_of xT).
+Definition dim := let: Mixin n _ := mixin class in n.
+
+Definition pack b0 (m0 : mixin_of (@GRing.Lmodule.Pack R _ T b0 T)) :=
+ fun bT b & phant_id (@GRing.Lmodule.class _ phR bT) b =>
+ fun m & phant_id m0 m => Pack phR (@Class T b m) T.
+
+Definition eqType := @Equality.Pack cT xclass xT.
+Definition choiceType := @Choice.Pack cT xclass xT.
+Definition zmodType := @GRing.Zmodule.Pack cT xclass xT.
+Definition lmodType := @GRing.Lmodule.Pack R phR cT xclass xT.
+
+End ClassDef.
+Notation axiom n V := (axiom_def n (Phant V)).
+
+Section OtherDefs.
+Local Coercion sort : type >-> Sortclass.
+Local Coercion dim : type >-> nat.
+Inductive space (K : fieldType) (vT : type (Phant K)) (phV : phant vT) :=
+ Space (mx : 'M[K]_vT) & <<mx>>%MS == mx.
+Inductive hom (R : ringType) (vT wT : type (Phant R)) :=
+ Hom of 'M[R]_(vT, wT).
+End OtherDefs.
+
+Module Import Exports.
+
+Coercion base : class_of >-> GRing.Lmodule.class_of.
+Coercion mixin : class_of >-> mixin_of.
+Coercion sort : type >-> Sortclass.
+Coercion eqType: type >-> Equality.type.
+Bind Scope ring_scope with sort.
+Canonical eqType.
+Coercion choiceType: type >-> Choice.type.
+Canonical choiceType.
+Coercion zmodType: type >-> GRing.Zmodule.type.
+Canonical zmodType.
+Coercion lmodType: type>-> GRing.Lmodule.type.
+Canonical lmodType.
+Notation vectType R := (@type _ (Phant R)).
+Notation VectType R V mV :=
+ (@pack _ (Phant R) V _ mV _ _ id _ id).
+Notation VectMixin := Mixin.
+Notation "[ 'vectType' R 'of' T 'for' cT ]" := (@clone _ (Phant R) T cT _ idfun)
+ (at level 0, format "[ 'vectType' R 'of' T 'for' cT ]") : form_scope.
+Notation "[ 'vectType' R 'of' T ]" := (@clone _ (Phant R) T _ _ idfun)
+ (at level 0, format "[ 'vectType' R 'of' T ]") : form_scope.
+
+Notation "{ 'vspace' vT }" := (space (Phant vT)) : type_scope.
+Notation "''Hom' ( aT , rT )" := (hom aT rT) : type_scope.
+Notation "''End' ( vT )" := (hom vT vT) : type_scope.
+
+Prenex Implicits Hom.
+
+Delimit Scope vspace_scope with VS.
+Bind Scope vspace_scope with space.
+Delimit Scope lfun_scope with VF.
+Bind Scope lfun_scope with hom.
+
+End Exports.
+
+(* The contents of this module exposes the matrix encodings, and should *)
+(* therefore not be used outside of the vector library implementation. *)
+Module InternalTheory.
+
+Section Iso.
+Variables (R : ringType) (vT rT : vectType R).
+Local Coercion dim : vectType >-> nat.
+
+Fact v2r_subproof : axiom vT vT. Proof. by case: vT => T [bT []]. Qed.
+Definition v2r := s2val v2r_subproof.
+
+Let v2r_bij : bijective v2r := s2valP' v2r_subproof.
+Fact r2v_subproof : {r2v | cancel r2v v2r}.
+Proof.
+have r2vP r: {v | v2r v = r}.
+ by apply: sig_eqW; have [v _ vK] := v2r_bij; exists (v r).
+by exists (fun r => sval (r2vP r)) => r; case: (r2vP r).
+Qed.
+Definition r2v := sval r2v_subproof.
+
+Lemma r2vK : cancel r2v v2r. Proof. exact: (svalP r2v_subproof). Qed.
+Lemma r2v_inj : injective r2v. Proof. exact: can_inj r2vK. Qed.
+Lemma v2rK : cancel v2r r2v. Proof. by have/bij_can_sym:= r2vK; apply. Qed.
+Lemma v2r_inj : injective v2r. Proof. exact: can_inj v2rK. Qed.
+
+Canonical v2r_linear := Linear (s2valP v2r_subproof : linear v2r).
+Canonical r2v_linear := Linear (can2_linear v2rK r2vK).
+End Iso.
+
+Section Vspace.
+Variables (K : fieldType) (vT : vectType K).
+Local Coercion dim : vectType >-> nat.
+
+Definition b2mx n (X : n.-tuple vT) := \matrix_i v2r (tnth X i).
+Lemma b2mxK n (X : n.-tuple vT) i : r2v (row i (b2mx X)) = X`_i.
+Proof. by rewrite rowK v2rK -tnth_nth. Qed.
+
+Definition vs2mx {phV} (U : @space K vT phV) := let: Space mx _ := U in mx.
+Lemma gen_vs2mx (U : {vspace vT}) : <<vs2mx U>>%MS = vs2mx U.
+Proof. by apply/eqP; rewrite /vs2mx; case: U. Qed.
+
+Fact mx2vs_subproof m (A : 'M[K]_(m, vT)) : <<(<<A>>)>>%MS == <<A>>%MS.
+Proof. by rewrite genmx_id. Qed.
+Definition mx2vs {m} A : {vspace vT} := Space _ (@mx2vs_subproof m A).
+
+Canonical space_subType := [subType for @vs2mx (Phant vT)].
+Lemma vs2mxK : cancel vs2mx mx2vs.
+Proof. by move=> v; apply: val_inj; rewrite /= gen_vs2mx. Qed.
+Lemma mx2vsK m (M : 'M_(m, vT)) : (vs2mx (mx2vs M) :=: M)%MS.
+Proof. exact: genmxE. Qed.
+End Vspace.
+
+Section Hom.
+Variables (R : ringType) (aT rT : vectType R).
+Definition f2mx (f : 'Hom(aT, rT)) := let: Hom A := f in A.
+Canonical hom_subType := [newType for f2mx].
+End Hom.
+
+Arguments Scope mx2vs [_ _ nat_scope matrix_set_scope].
+Prenex Implicits v2r r2v v2rK r2vK b2mx vs2mx vs2mxK f2mx.
+
+End InternalTheory.
+
+End Vector.
+Export Vector.Exports.
+Import Vector.InternalTheory.
+
+Section VspaceDefs.
+
+Variables (K : fieldType) (vT : vectType K).
+Implicit Types (u : vT) (X : seq vT) (U V : {vspace vT}).
+
+Definition space_eqMixin := Eval hnf in [eqMixin of {vspace vT} by <:].
+Canonical space_eqType := EqType {vspace vT} space_eqMixin.
+Definition space_choiceMixin := Eval hnf in [choiceMixin of {vspace vT} by <:].
+Canonical space_choiceType := ChoiceType {vspace vT} space_choiceMixin.
+
+Definition dimv U := \rank (vs2mx U).
+Definition subsetv U V := (vs2mx U <= vs2mx V)%MS.
+Definition vline u := mx2vs (v2r u).
+
+(* Vspace membership is defined as line inclusion. *)
+Definition pred_of_vspace phV (U : Vector.space phV) : pred_class :=
+ fun v => (vs2mx (vline v) <= vs2mx U)%MS.
+Canonical vspace_predType :=
+ @mkPredType _ (unkeyed {vspace vT}) (@pred_of_vspace _).
+
+Definition fullv : {vspace vT} := mx2vs 1%:M.
+Definition addv U V := mx2vs (vs2mx U + vs2mx V).
+Definition capv U V := mx2vs (vs2mx U :&: vs2mx V).
+Definition complv U := mx2vs (vs2mx U)^C.
+Definition diffv U V := mx2vs (vs2mx U :\: vs2mx V).
+Definition vpick U := r2v (nz_row (vs2mx U)).
+Fact span_key : unit. Proof. by []. Qed.
+Definition span_expanded_def X := mx2vs (b2mx (in_tuple X)).
+Definition span := locked_with span_key span_expanded_def.
+Canonical span_unlockable := [unlockable fun span].
+Definition vbasis_def U :=
+ [tuple r2v (row i (row_base (vs2mx U))) | i < dimv U].
+Definition vbasis := locked_with span_key vbasis_def.
+Canonical vbasis_unlockable := [unlockable fun vbasis].
+
+(* coord and directv are defined in the VectorTheory section. *)
+
+Definition free X := dimv (span X) == size X.
+Definition basis_of U X := (span X == U) && free X.
+
+End VspaceDefs.
+
+Coercion pred_of_vspace : Vector.space >-> pred_class.
+Notation "\dim U" := (dimv U) : nat_scope.
+Notation "U <= V" := (subsetv U V) : vspace_scope.
+Notation "U <= V <= W" := (subsetv U V && subsetv V W) : vspace_scope.
+Notation "<[ v ] >" := (vline v) : vspace_scope.
+Notation "<< X >>" := (span X) : vspace_scope.
+Notation "0" := (vline 0) : vspace_scope.
+Implicit Arguments fullv [[K] [vT]].
+Prenex Implicits subsetv addv capv complv diffv span free basis_of.
+
+Notation "U + V" := (addv U V) : vspace_scope.
+Notation "U :&: V" := (capv U V) : vspace_scope.
+Notation "U ^C" := (complv U) (at level 8, format "U ^C") : vspace_scope.
+Notation "U :\: V" := (diffv U V) : vspace_scope.
+Notation "{ : vT }" := (@fullv _ vT) (only parsing) : vspace_scope.
+
+Notation "\sum_ ( i <- r | P ) U" :=
+ (\big[addv/0%VS]_(i <- r | P%B) U%VS) : vspace_scope.
+Notation "\sum_ ( i <- r ) U" :=
+ (\big[addv/0%VS]_(i <- r) U%VS) : vspace_scope.
+Notation "\sum_ ( m <= i < n | P ) U" :=
+ (\big[addv/0%VS]_(m <= i < n | P%B) U%VS) : vspace_scope.
+Notation "\sum_ ( m <= i < n ) U" :=
+ (\big[addv/0%VS]_(m <= i < n) U%VS) : vspace_scope.
+Notation "\sum_ ( i | P ) U" :=
+ (\big[addv/0%VS]_(i | P%B) U%VS) : vspace_scope.
+Notation "\sum_ i U" :=
+ (\big[addv/0%VS]_i U%VS) : vspace_scope.
+Notation "\sum_ ( i : t | P ) U" :=
+ (\big[addv/0%VS]_(i : t | P%B) U%VS) (only parsing) : vspace_scope.
+Notation "\sum_ ( i : t ) U" :=
+ (\big[addv/0%VS]_(i : t) U%VS) (only parsing) : vspace_scope.
+Notation "\sum_ ( i < n | P ) U" :=
+ (\big[addv/0%VS]_(i < n | P%B) U%VS) : vspace_scope.
+Notation "\sum_ ( i < n ) U" :=
+ (\big[addv/0%VS]_(i < n) U%VS) : vspace_scope.
+Notation "\sum_ ( i 'in' A | P ) U" :=
+ (\big[addv/0%VS]_(i in A | P%B) U%VS) : vspace_scope.
+Notation "\sum_ ( i 'in' A ) U" :=
+ (\big[addv/0%VS]_(i in A) U%VS) : vspace_scope.
+
+Notation "\bigcap_ ( i <- r | P ) U" :=
+ (\big[capv/fullv]_(i <- r | P%B) U%VS) : vspace_scope.
+Notation "\bigcap_ ( i <- r ) U" :=
+ (\big[capv/fullv]_(i <- r) U%VS) : vspace_scope.
+Notation "\bigcap_ ( m <= i < n | P ) U" :=
+ (\big[capv/fullv]_(m <= i < n | P%B) U%VS) : vspace_scope.
+Notation "\bigcap_ ( m <= i < n ) U" :=
+ (\big[capv/fullv]_(m <= i < n) U%VS) : vspace_scope.
+Notation "\bigcap_ ( i | P ) U" :=
+ (\big[capv/fullv]_(i | P%B) U%VS) : vspace_scope.
+Notation "\bigcap_ i U" :=
+ (\big[capv/fullv]_i U%VS) : vspace_scope.
+Notation "\bigcap_ ( i : t | P ) U" :=
+ (\big[capv/fullv]_(i : t | P%B) U%VS) (only parsing) : vspace_scope.
+Notation "\bigcap_ ( i : t ) U" :=
+ (\big[capv/fullv]_(i : t) U%VS) (only parsing) : vspace_scope.
+Notation "\bigcap_ ( i < n | P ) U" :=
+ (\big[capv/fullv]_(i < n | P%B) U%VS) : vspace_scope.
+Notation "\bigcap_ ( i < n ) U" :=
+ (\big[capv/fullv]_(i < n) U%VS) : vspace_scope.
+Notation "\bigcap_ ( i 'in' A | P ) U" :=
+ (\big[capv/fullv]_(i in A | P%B) U%VS) : vspace_scope.
+Notation "\bigcap_ ( i 'in' A ) U" :=
+ (\big[capv/fullv]_(i in A) U%VS) : vspace_scope.
+
+Section VectorTheory.
+
+Variables (K : fieldType) (vT : vectType K).
+Implicit Types (a : K) (u v w : vT) (X Y : seq vT) (U V W : {vspace vT}).
+
+Local Notation subV := (@subsetv K vT) (only parsing).
+Local Notation addV := (@addv K vT) (only parsing).
+Local Notation capV := (@capv K vT) (only parsing).
+
+(* begin hide *)
+
+(* Internal theory facts *)
+Let vs2mxP U V : reflect (U = V) (vs2mx U == vs2mx V)%MS.
+Proof. by rewrite (sameP genmxP eqP) !gen_vs2mx; apply: eqP. Qed.
+
+Let memvK v U : (v \in U) = (v2r v <= vs2mx U)%MS.
+Proof. by rewrite -genmxE. Qed.
+
+Let mem_r2v rv U : (r2v rv \in U) = (rv <= vs2mx U)%MS.
+Proof. by rewrite memvK r2vK. Qed.
+
+Let vs2mx0 : @vs2mx K vT _ 0 = 0.
+Proof. by rewrite /= linear0 genmx0. Qed.
+
+Let vs2mxD U V : vs2mx (U + V) = (vs2mx U + vs2mx V)%MS.
+Proof. by rewrite /= genmx_adds !gen_vs2mx. Qed.
+
+Let vs2mx_sum := big_morph _ vs2mxD vs2mx0.
+
+Let vs2mxI U V : vs2mx (U :&: V) = (vs2mx U :&: vs2mx V)%MS.
+Proof. by rewrite /= genmx_cap !gen_vs2mx. Qed.
+
+Let vs2mxF : vs2mx {:vT} = 1%:M.
+Proof. by rewrite /= genmx1. Qed.
+
+Let row_b2mx n (X : n.-tuple vT) i : row i (b2mx X) = v2r X`_i.
+Proof. by rewrite -tnth_nth rowK. Qed.
+
+Let span_b2mx n (X : n.-tuple vT) : span X = mx2vs (b2mx X).
+Proof. by rewrite unlock tvalK; case: _ / (esym _). Qed.
+
+Let mul_b2mx n (X : n.-tuple vT) (rk : 'rV_n) :
+ \sum_i rk 0 i *: X`_i = r2v (rk *m b2mx X).
+Proof.
+rewrite mulmx_sum_row linear_sum; apply: eq_bigr => i _.
+by rewrite row_b2mx linearZ /= v2rK.
+Qed.
+
+Let lin_b2mx n (X : n.-tuple vT) k :
+ \sum_(i < n) k i *: X`_i = r2v (\row_i k i *m b2mx X).
+Proof. by rewrite -mul_b2mx; apply: eq_bigr => i _; rewrite mxE. Qed.
+
+Let free_b2mx n (X : n.-tuple vT) : free X = row_free (b2mx X).
+Proof. by rewrite /free /dimv span_b2mx genmxE size_tuple. Qed.
+(* end hide *)
+
+Fact vspace_key U : pred_key U. Proof. by []. Qed.
+Canonical vspace_keyed U := KeyedPred (vspace_key U).
+
+Lemma memvE v U : (v \in U) = (<[v]> <= U)%VS. Proof. by []. Qed.
+
+Lemma vlineP v1 v2 : reflect (exists k, v1 = k *: v2) (v1 \in <[v2]>)%VS.
+Proof.
+apply: (iffP idP) => [|[k ->]]; rewrite memvK genmxE ?linearZ ?scalemx_sub //.
+by case/sub_rVP=> k; rewrite -linearZ => /v2r_inj->; exists k.
+Qed.
+
+Fact memv_submod_closed U : submod_closed U.
+Proof.
+split=> [|a u v]; rewrite !memvK ?linear0 ?sub0mx // => Uu Uv.
+by rewrite linearP addmx_sub ?scalemx_sub.
+Qed.
+Canonical memv_opprPred U := OpprPred (memv_submod_closed U).
+Canonical memv_addrPred U := AddrPred (memv_submod_closed U).
+Canonical memv_zmodPred U := ZmodPred (memv_submod_closed U).
+Canonical memv_submodPred U := SubmodPred (memv_submod_closed U).
+
+Lemma mem0v U : 0 \in U. Proof. exact : rpred0. Qed.
+Lemma memvN U v : (- v \in U) = (v \in U). Proof. exact: rpredN. Qed.
+Lemma memvD U : {in U &, forall u v, u + v \in U}. Proof. exact : rpredD. Qed.
+Lemma memvB U : {in U &, forall u v, u - v \in U}. Proof. exact : rpredB. Qed.
+Lemma memvZ U k : {in U, forall v, k *: v \in U}. Proof. exact : rpredZ. Qed.
+
+Lemma memv_suml I r (P : pred I) vs U :
+ (forall i, P i -> vs i \in U) -> \sum_(i <- r | P i) vs i \in U.
+Proof. exact: rpred_sum. Qed.
+
+Lemma memv_line u : u \in <[u]>%VS.
+Proof. by apply/vlineP; exists 1; rewrite scale1r. Qed.
+
+Lemma subvP U V : reflect {subset U <= V} (U <= V)%VS.
+Proof.
+apply: (iffP rV_subP) => sU12 u.
+ by rewrite !memvE /subsetv !genmxE => /sU12.
+by have:= sU12 (r2v u); rewrite !memvE /subsetv !genmxE r2vK.
+Qed.
+
+Lemma subvv U : (U <= U)%VS. Proof. exact/subvP. Qed.
+Hint Resolve subvv.
+
+Lemma subv_trans : transitive subV.
+Proof. by move=> U V W /subvP sUV /subvP sVW; apply/subvP=> u /sUV/sVW. Qed.
+
+Lemma subv_anti : antisymmetric subV.
+Proof. by move=> U V; apply/vs2mxP. Qed.
+
+Lemma eqEsubv U V : (U == V) = (U <= V <= U)%VS.
+Proof. by apply/eqP/idP=> [-> | /subv_anti//]; rewrite subvv. Qed.
+
+Lemma vspaceP U V : U =i V <-> U = V.
+Proof.
+split=> [eqUV | -> //]; apply/subv_anti/andP.
+by split; apply/subvP=> v; rewrite eqUV.
+Qed.
+
+Lemma subvPn {U V} : reflect (exists2 u, u \in U & u \notin V) (~~ (U <= V)%VS).
+Proof.
+apply: (iffP idP) => [|[u Uu]]; last by apply: contra => /subvP->.
+case/row_subPn=> i; set vi := row i _ => V'vi.
+by exists (r2v vi); rewrite memvK r2vK ?row_sub.
+Qed.
+
+(* Empty space. *)
+Lemma sub0v U : (0 <= U)%VS.
+Proof. exact: mem0v. Qed.
+
+Lemma subv0 U : (U <= 0)%VS = (U == 0%VS).
+Proof. by rewrite eqEsubv sub0v andbT. Qed.
+
+Lemma memv0 v : v \in 0%VS = (v == 0).
+Proof. by apply/idP/eqP=> [/vlineP[k ->] | ->]; rewrite (scaler0, mem0v). Qed.
+
+(* Full space *)
+
+Lemma subvf U : (U <= fullv)%VS. Proof. by rewrite /subsetv vs2mxF submx1. Qed.
+Lemma memvf v : v \in fullv. Proof. exact: subvf. Qed.
+
+(* Picking a non-zero vector in a subspace. *)
+Lemma memv_pick U : vpick U \in U. Proof. by rewrite mem_r2v nz_row_sub. Qed.
+
+Lemma vpick0 U : (vpick U == 0) = (U == 0%VS).
+Proof. by rewrite -memv0 mem_r2v -subv0 /subV vs2mx0 !submx0 nz_row_eq0. Qed.
+
+(* Sum of subspaces. *)
+Lemma subv_add U V W : (U + V <= W)%VS = (U <= W)%VS && (V <= W)%VS.
+Proof. by rewrite /subV vs2mxD addsmx_sub. Qed.
+
+Lemma addvS U1 U2 V1 V2 : (U1 <= U2 -> V1 <= V2 -> U1 + V1 <= U2 + V2)%VS.
+Proof. by rewrite /subV !vs2mxD; apply: addsmxS. Qed.
+
+Lemma addvSl U V : (U <= U + V)%VS.
+Proof. by rewrite /subV vs2mxD addsmxSl. Qed.
+
+Lemma addvSr U V : (V <= U + V)%VS.
+Proof. by rewrite /subV vs2mxD addsmxSr. Qed.
+
+Lemma addvC : commutative addV.
+Proof. by move=> U V; apply/vs2mxP; rewrite !vs2mxD addsmxC submx_refl. Qed.
+
+Lemma addvA : associative addV.
+Proof. by move=> U V W; apply/vs2mxP; rewrite !vs2mxD addsmxA submx_refl. Qed.
+
+Lemma addv_idPl {U V}: reflect (U + V = U)%VS (V <= U)%VS.
+Proof. by rewrite /subV (sameP addsmx_idPl eqmxP) -vs2mxD; apply: vs2mxP. Qed.
+
+Lemma addv_idPr {U V} : reflect (U + V = V)%VS (U <= V)%VS.
+Proof. by rewrite addvC; apply: addv_idPl. Qed.
+
+Lemma addvv : idempotent addV.
+Proof. by move=> U; apply/addv_idPl. Qed.
+
+Lemma add0v : left_id 0%VS addV.
+Proof. by move=> U; apply/addv_idPr/sub0v. Qed.
+
+Lemma addv0 : right_id 0%VS addV.
+Proof. by move=> U; apply/addv_idPl/sub0v. Qed.
+
+Lemma sumfv : left_zero fullv addV.
+Proof. by move=> U; apply/addv_idPl/subvf. Qed.
+
+Lemma addvf : right_zero fullv addV.
+Proof. by move=> U; apply/addv_idPr/subvf. Qed.
+
+Canonical addv_monoid := Monoid.Law addvA add0v addv0.
+Canonical addv_comoid := Monoid.ComLaw addvC.
+
+Lemma memv_add u v U V : u \in U -> v \in V -> u + v \in (U + V)%VS.
+Proof. by rewrite !memvK genmxE linearD; apply: addmx_sub_adds. Qed.
+
+Lemma memv_addP {w U V} :
+ reflect (exists2 u, u \in U & exists2 v, v \in V & w = u + v)
+ (w \in U + V)%VS.
+Proof.
+apply: (iffP idP) => [|[u Uu [v Vv ->]]]; last exact: memv_add.
+rewrite memvK genmxE => /sub_addsmxP[r /(canRL v2rK)->].
+rewrite linearD /=; set u := r2v _; set v := r2v _.
+by exists u; last exists v; rewrite // mem_r2v submxMl.
+Qed.
+
+Section BigSum.
+Variable I : finType.
+Implicit Type P : pred I.
+
+Lemma sumv_sup i0 P U Vs :
+ P i0 -> (U <= Vs i0)%VS -> (U <= \sum_(i | P i) Vs i)%VS.
+Proof. by move=> Pi0 /subv_trans-> //; rewrite (bigD1 i0) ?addvSl. Qed.
+Implicit Arguments sumv_sup [P U Vs].
+
+Lemma subv_sumP {P Us V} :
+ reflect (forall i, P i -> Us i <= V)%VS (\sum_(i | P i) Us i <= V)%VS.
+Proof.
+apply: (iffP idP) => [sUV i Pi | sUV].
+ by apply: subv_trans sUV; apply: sumv_sup Pi _.
+by elim/big_rec: _ => [|i W Pi sWV]; rewrite ?sub0v // subv_add sUV.
+Qed.
+
+Lemma memv_sumr P vs (Us : I -> {vspace vT}) :
+ (forall i, P i -> vs i \in Us i) ->
+ \sum_(i | P i) vs i \in (\sum_(i | P i) Us i)%VS.
+Proof. by move=> Uv; apply/rpred_sum=> i Pi; apply/(sumv_sup i Pi)/Uv. Qed.
+
+Lemma memv_sumP {P} {Us : I -> {vspace vT}} {v} :
+ reflect (exists2 vs, forall i, P i -> vs i \in Us i
+ & v = \sum_(i | P i) vs i)
+ (v \in \sum_(i | P i) Us i)%VS.
+Proof.
+apply: (iffP idP) => [|[vs Uv ->]]; last exact: memv_sumr.
+rewrite memvK vs2mx_sum => /sub_sumsmxP[r /(canRL v2rK)->].
+pose f i := r2v (r i *m vs2mx (Us i)); rewrite linear_sum /=.
+by exists f => //= i _; rewrite mem_r2v submxMl.
+Qed.
+
+End BigSum.
+
+(* Intersection *)
+
+Lemma subv_cap U V W : (U <= V :&: W)%VS = (U <= V)%VS && (U <= W)%VS.
+Proof. by rewrite /subV vs2mxI sub_capmx. Qed.
+
+Lemma capvS U1 U2 V1 V2 : (U1 <= U2 -> V1 <= V2 -> U1 :&: V1 <= U2 :&: V2)%VS.
+Proof. by rewrite /subV !vs2mxI; apply: capmxS. Qed.
+
+Lemma capvSl U V : (U :&: V <= U)%VS.
+Proof. by rewrite /subV vs2mxI capmxSl. Qed.
+
+Lemma capvSr U V : (U :&: V <= V)%VS.
+Proof. by rewrite /subV vs2mxI capmxSr. Qed.
+
+Lemma capvC : commutative capV.
+Proof. by move=> U V; apply/vs2mxP; rewrite !vs2mxI capmxC submx_refl. Qed.
+
+Lemma capvA : associative capV.
+Proof. by move=> U V W; apply/vs2mxP; rewrite !vs2mxI capmxA submx_refl. Qed.
+
+Lemma capv_idPl {U V} : reflect (U :&: V = U)%VS (U <= V)%VS.
+Proof. by rewrite /subV(sameP capmx_idPl eqmxP) -vs2mxI; apply: vs2mxP. Qed.
+
+Lemma capv_idPr {U V} : reflect (U :&: V = V)%VS (V <= U)%VS.
+Proof. by rewrite capvC; apply: capv_idPl. Qed.
+
+Lemma capvv : idempotent capV.
+Proof. by move=> U; apply/capv_idPl. Qed.
+
+Lemma cap0v : left_zero 0%VS capV.
+Proof. by move=> U; apply/capv_idPl/sub0v. Qed.
+
+Lemma capv0 : right_zero 0%VS capV.
+Proof. by move=> U; apply/capv_idPr/sub0v. Qed.
+
+Lemma capfv : left_id fullv capV.
+Proof. by move=> U; apply/capv_idPr/subvf. Qed.
+
+Lemma capvf : right_id fullv capV.
+Proof. by move=> U; apply/capv_idPl/subvf. Qed.
+
+Canonical capv_monoid := Monoid.Law capvA capfv capvf.
+Canonical capv_comoid := Monoid.ComLaw capvC.
+
+Lemma memv_cap w U V : (w \in U :&: V)%VS = (w \in U) && (w \in V).
+Proof. by rewrite !memvE subv_cap. Qed.
+
+Lemma memv_capP {w U V} : reflect (w \in U /\ w \in V) (w \in U :&: V)%VS.
+Proof. by rewrite memv_cap; apply: andP. Qed.
+
+Lemma vspace_modl U V W : (U <= W -> U + (V :&: W) = (U + V) :&: W)%VS.
+Proof.
+by move=> sUV; apply/vs2mxP; rewrite !(vs2mxD, vs2mxI); exact/eqmxP/matrix_modl.
+Qed.
+
+Lemma vspace_modr U V W : (W <= U -> (U :&: V) + W = U :&: (V + W))%VS.
+Proof. by rewrite -!(addvC W) !(capvC U); apply: vspace_modl. Qed.
+
+Section BigCap.
+Variable I : finType.
+Implicit Type P : pred I.
+
+Lemma bigcapv_inf i0 P Us V :
+ P i0 -> (Us i0 <= V -> \bigcap_(i | P i) Us i <= V)%VS.
+Proof. by move=> Pi0; apply: subv_trans; rewrite (bigD1 i0) ?capvSl. Qed.
+
+Lemma subv_bigcapP {P U Vs} :
+ reflect (forall i, P i -> U <= Vs i)%VS (U <= \bigcap_(i | P i) Vs i)%VS.
+Proof.
+apply: (iffP idP) => [sUV i Pi | sUV].
+ by rewrite (subv_trans sUV) ?(bigcapv_inf Pi).
+by elim/big_rec: _ => [|i W Pi]; rewrite ?subvf // subv_cap sUV.
+Qed.
+
+End BigCap.
+
+(* Complement *)
+Lemma addv_complf U : (U + U^C)%VS = fullv.
+Proof.
+apply/vs2mxP; rewrite vs2mxD -gen_vs2mx -genmx_adds !genmxE submx1 sub1mx.
+exact: addsmx_compl_full.
+Qed.
+
+Lemma capv_compl U : (U :&: U^C = 0)%VS.
+Proof.
+apply/val_inj; rewrite [val]/= vs2mx0 vs2mxI -gen_vs2mx -genmx_cap.
+by rewrite capmx_compl genmx0.
+Qed.
+
+(* Difference *)
+Lemma diffvSl U V : (U :\: V <= U)%VS.
+Proof. by rewrite /subV genmxE diffmxSl. Qed.
+
+Lemma capv_diff U V : ((U :\: V) :&: V = 0)%VS.
+Proof.
+apply/val_inj; rewrite [val]/= vs2mx0 vs2mxI -(gen_vs2mx V) -genmx_cap.
+by rewrite capmx_diff genmx0.
+Qed.
+
+Lemma addv_diff_cap U V : (U :\: V + U :&: V)%VS = U.
+Proof.
+apply/vs2mxP; rewrite vs2mxD -genmx_adds !genmxE.
+exact/eqmxP/addsmx_diff_cap_eq.
+Qed.
+
+Lemma addv_diff U V : (U :\: V + V = U + V)%VS.
+Proof. by rewrite -{2}(addv_diff_cap U V) -addvA (addv_idPr (capvSr U V)). Qed.
+
+(* Subspace dimension. *)
+Lemma dimv0 : \dim (0%VS : {vspace vT}) = 0%N.
+Proof. by rewrite /dimv vs2mx0 mxrank0. Qed.
+
+Lemma dimv_eq0 U : (\dim U == 0%N) = (U == 0%VS).
+Proof. by rewrite /dimv /= mxrank_eq0 {2}/eq_op /= linear0 genmx0. Qed.
+
+Lemma dimvf : \dim {:vT} = Vector.dim vT.
+Proof. by rewrite /dimv vs2mxF mxrank1. Qed.
+
+Lemma dim_vline v : \dim <[v]> = (v != 0).
+Proof. by rewrite /dimv mxrank_gen rank_rV (can2_eq v2rK r2vK) linear0. Qed.
+
+Lemma dimvS U V : (U <= V)%VS -> \dim U <= \dim V.
+Proof. exact: mxrankS. Qed.
+
+Lemma dimv_leqif_sup U V : (U <= V)%VS -> \dim U <= \dim V ?= iff (V <= U)%VS.
+Proof. exact: mxrank_leqif_sup. Qed.
+
+Lemma dimv_leqif_eq U V : (U <= V)%VS -> \dim U <= \dim V ?= iff (U == V).
+Proof. by rewrite eqEsubv; apply: mxrank_leqif_eq. Qed.
+
+Lemma eqEdim U V : (U == V) = (U <= V)%VS && (\dim V <= \dim U).
+Proof. by apply/idP/andP=> [/eqP | [/dimv_leqif_eq/geq_leqif]] ->. Qed.
+
+Lemma dimv_compl U : \dim U^C = (\dim {:vT} - \dim U)%N.
+Proof. by rewrite dimvf /dimv mxrank_gen mxrank_compl. Qed.
+
+Lemma dimv_cap_compl U V : (\dim (U :&: V) + \dim (U :\: V))%N = \dim U.
+Proof. by rewrite /dimv !mxrank_gen mxrank_cap_compl. Qed.
+
+Lemma dimv_sum_cap U V : (\dim (U + V) + \dim (U :&: V) = \dim U + \dim V)%N.
+Proof. by rewrite /dimv !mxrank_gen mxrank_sum_cap. Qed.
+
+Lemma dimv_disjoint_sum U V :
+ (U :&: V = 0)%VS -> \dim (U + V) = (\dim U + \dim V)%N.
+Proof. by move=> dxUV; rewrite -dimv_sum_cap dxUV dimv0 addn0. Qed.
+
+Lemma dimv_add_leqif U V :
+ \dim (U + V) <= \dim U + \dim V ?= iff (U :&: V <= 0)%VS.
+Proof.
+by rewrite /dimv /subV !mxrank_gen vs2mx0 genmxE; apply: mxrank_adds_leqif.
+Qed.
+
+Lemma diffv_eq0 U V : (U :\: V == 0)%VS = (U <= V)%VS.
+Proof.
+rewrite -dimv_eq0 -(eqn_add2l (\dim (U :&: V))) addn0 dimv_cap_compl eq_sym.
+by rewrite (dimv_leqif_eq (capvSl _ _)) (sameP capv_idPl eqP).
+Qed.
+
+Lemma dimv_leq_sum I r (P : pred I) (Us : I -> {vspace vT}) :
+ \dim (\sum_(i <- r | P i) Us i) <= \sum_(i <- r | P i) \dim (Us i).
+Proof.
+elim/big_rec2: _ => [|i d vs _ le_vs_d]; first by rewrite dim_vline eqxx.
+by apply: (leq_trans (dimv_add_leqif _ _)); rewrite leq_add2l.
+Qed.
+
+Section SumExpr.
+
+(* The vector direct sum theory clones the interface types of the matrix *)
+(* direct sum theory (see mxalgebra for the technical details), but *)
+(* nevetheless reuses much of the matrix theory. *)
+Structure addv_expr := Sumv {
+ addv_val :> wrapped {vspace vT};
+ addv_dim : wrapped nat;
+ _ : mxsum_spec (vs2mx (unwrap addv_val)) (unwrap addv_dim)
+}.
+
+(* Piggyback on mxalgebra theory. *)
+Definition vs2mx_sum_expr_subproof (S : addv_expr) :
+ mxsum_spec (vs2mx (unwrap S)) (unwrap (addv_dim S)).
+Proof. by case: S. Qed.
+Canonical vs2mx_sum_expr S := ProperMxsumExpr (vs2mx_sum_expr_subproof S).
+
+Canonical trivial_addv U := @Sumv (Wrap U) (Wrap (\dim U)) (TrivialMxsum _).
+
+Structure proper_addv_expr := ProperSumvExpr {
+ proper_addv_val :> {vspace vT};
+ proper_addv_dim :> nat;
+ _ : mxsum_spec (vs2mx proper_addv_val) proper_addv_dim
+}.
+
+Definition proper_addvP (S : proper_addv_expr) :=
+ let: ProperSumvExpr _ _ termS := S return mxsum_spec (vs2mx S) S in termS.
+
+Canonical proper_addv (S : proper_addv_expr) :=
+ @Sumv (wrap (S : {vspace vT})) (wrap (S : nat)) (proper_addvP S).
+
+Section Binary.
+Variables S1 S2 : addv_expr.
+Fact binary_addv_subproof :
+ mxsum_spec (vs2mx (unwrap S1 + unwrap S2))
+ (unwrap (addv_dim S1) + unwrap (addv_dim S2)).
+Proof. by rewrite vs2mxD; apply: proper_mxsumP. Qed.
+Canonical binary_addv_expr := ProperSumvExpr binary_addv_subproof.
+End Binary.
+
+Section Nary.
+Variables (I : Type) (r : seq I) (P : pred I) (S_ : I -> addv_expr).
+Fact nary_addv_subproof :
+ mxsum_spec (vs2mx (\sum_(i <- r | P i) unwrap (S_ i)))
+ (\sum_(i <- r | P i) unwrap (addv_dim (S_ i))).
+Proof. by rewrite vs2mx_sum; apply: proper_mxsumP. Qed.
+Canonical nary_addv_expr := ProperSumvExpr nary_addv_subproof.
+End Nary.
+
+Definition directv_def S of phantom {vspace vT} (unwrap (addv_val S)) :=
+ \dim (unwrap S) == unwrap (addv_dim S).
+
+End SumExpr.
+
+Local Notation directv A := (directv_def (Phantom {vspace _} A%VS)).
+
+Lemma directvE (S : addv_expr) :
+ directv (unwrap S) = (\dim (unwrap S) == unwrap (addv_dim S)).
+Proof. by []. Qed.
+
+Lemma directvP {S : proper_addv_expr} : reflect (\dim S = S :> nat) (directv S).
+Proof. exact: eqnP. Qed.
+
+Lemma directv_trivial U : directv (unwrap (@trivial_addv U)).
+Proof. exact: eqxx. Qed.
+
+Lemma dimv_sum_leqif (S : addv_expr) :
+ \dim (unwrap S) <= unwrap (addv_dim S) ?= iff directv (unwrap S).
+Proof.
+rewrite directvE; case: S => [[U] [d] /= defUd]; split=> //=.
+rewrite /dimv; elim: {1}_ {U}_ d / defUd => // m1 m2 A1 A2 r1 r2 _ leA1 _ leA2.
+by apply: leq_trans (leq_add leA1 leA2); rewrite mxrank_adds_leqif.
+Qed.
+
+Lemma directvEgeq (S : addv_expr) :
+ directv (unwrap S) = (\dim (unwrap S) >= unwrap (addv_dim S)).
+Proof. by rewrite leq_eqVlt ltnNge eq_sym !dimv_sum_leqif orbF. Qed.
+
+Section BinaryDirect.
+
+Lemma directv_addE (S1 S2 : addv_expr) :
+ directv (unwrap S1 + unwrap S2)
+ = [&& directv (unwrap S1), directv (unwrap S2)
+ & unwrap S1 :&: unwrap S2 == 0]%VS.
+Proof.
+by rewrite /directv_def /dimv vs2mxD -mxdirectE mxdirect_addsE -vs2mxI -vs2mx0.
+Qed.
+
+Lemma directv_addP {U V} : reflect (U :&: V = 0)%VS (directv (U + V)).
+Proof. by rewrite directv_addE !directv_trivial; apply: eqP. Qed.
+
+Lemma directv_add_unique {U V} :
+ reflect (forall u1 u2 v1 v2, u1 \in U -> u2 \in U -> v1 \in V -> v2 \in V ->
+ (u1 + v1 == u2 + v2) = ((u1, v1) == (u2, v2)))
+ (directv (U + V)).
+Proof.
+apply: (iffP directv_addP) => [dxUV u1 u2 v1 v2 Uu1 Uu2 Vv1 Vv2 | dxUV].
+ apply/idP/idP=> [| /eqP[-> ->] //]; rewrite -subr_eq0 opprD addrACA addr_eq0.
+ move/eqP=> eq_uv; rewrite xpair_eqE -subr_eq0 eq_uv oppr_eq0 subr_eq0 andbb.
+ by rewrite -subr_eq0 -memv0 -dxUV memv_cap -memvN -eq_uv !memvB.
+apply/eqP; rewrite -subv0; apply/subvP=> v /memv_capP[U1v U2v].
+by rewrite memv0 -[v == 0]andbb {1}eq_sym -xpair_eqE -dxUV ?mem0v // addrC.
+Qed.
+
+End BinaryDirect.
+
+Section NaryDirect.
+
+Context {I : finType} {P : pred I}.
+
+Lemma directv_sumP {Us : I -> {vspace vT}} :
+ reflect (forall i, P i -> Us i :&: (\sum_(j | P j && (j != i)) Us j) = 0)%VS
+ (directv (\sum_(i | P i) Us i)).
+Proof.
+rewrite directvE /= /dimv vs2mx_sum -mxdirectE; apply: (equivP mxdirect_sumsP).
+by do [split=> dxU i /dxU; rewrite -vs2mx_sum -vs2mxI -vs2mx0] => [/val_inj|->].
+Qed.
+
+Lemma directv_sumE {Ss : I -> addv_expr} (xunwrap := unwrap) :
+ reflect [/\ forall i, P i -> directv (unwrap (Ss i))
+ & directv (\sum_(i | P i) xunwrap (Ss i))]
+ (directv (\sum_(i | P i) unwrap (Ss i))).
+Proof.
+by rewrite !directvE /= /dimv 2!{1}vs2mx_sum -!mxdirectE; apply: mxdirect_sumsE.
+Qed.
+
+Lemma directv_sum_independent {Us : I -> {vspace vT}} :
+ reflect (forall us,
+ (forall i, P i -> us i \in Us i) -> \sum_(i | P i) us i = 0 ->
+ (forall i, P i -> us i = 0))
+ (directv (\sum_(i | P i) Us i)).
+Proof.
+apply: (iffP directv_sumP) => [dxU us Uu u_0 i Pi | dxU i Pi].
+ apply/eqP; rewrite -memv0 -(dxU i Pi) memv_cap Uu //= -memvN -sub0r -{1}u_0.
+ by rewrite (bigD1 i) //= addrC addKr memv_sumr // => j /andP[/Uu].
+apply/eqP; rewrite -subv0; apply/subvP=> v.
+rewrite memv_cap memv0 => /andP[Uiv /memv_sumP[us Uu Dv]].
+have: \sum_(j | P j) [eta us with i |-> - v] j = 0.
+ rewrite (bigD1 i) //= eqxx {1}Dv addrC -sumrB big1 // => j /andP[_ i'j].
+ by rewrite (negPf i'j) subrr.
+move/dxU/(_ i Pi); rewrite /= eqxx -oppr_eq0 => -> // j Pj.
+by have [-> | i'j] := altP eqP; rewrite ?memvN // Uu ?Pj.
+Qed.
+
+Lemma directv_sum_unique {Us : I -> {vspace vT}} :
+ reflect (forall us vs,
+ (forall i, P i -> us i \in Us i) ->
+ (forall i, P i -> vs i \in Us i) ->
+ (\sum_(i | P i) us i == \sum_(i | P i) vs i)
+ = [forall (i | P i), us i == vs i])
+ (directv (\sum_(i | P i) Us i)).
+Proof.
+apply: (iffP directv_sum_independent) => [dxU us vs Uu Uv | dxU us Uu u_0 i Pi].
+ apply/idP/forall_inP=> [|eq_uv]; last by apply/eqP/eq_bigr => i /eq_uv/eqP.
+ rewrite -subr_eq0 -sumrB => /eqP/dxU eq_uv i Pi.
+ by rewrite -subr_eq0 eq_uv // => j Pj; apply: memvB; move: j Pj.
+apply/eqP; have:= esym (dxU us \0 Uu _); rewrite u_0 big1_eq eqxx.
+by move/(_ _)/forall_inP=> -> // j _; apply: mem0v.
+Qed.
+
+End NaryDirect.
+
+(* Linear span generated by a list of vectors *)
+Lemma memv_span X v : v \in X -> v \in <<X>>%VS.
+Proof.
+by case/seq_tnthP=> i {v}->; rewrite unlock memvK genmxE (eq_row_sub i) // rowK.
+Qed.
+
+Lemma memv_span1 v : v \in <<[:: v]>>%VS.
+Proof. by rewrite memv_span ?mem_head. Qed.
+
+Lemma dim_span X : \dim <<X>> <= size X.
+Proof. by rewrite unlock /dimv genmxE rank_leq_row. Qed.
+
+Lemma span_subvP {X U} : reflect {subset X <= U} (<<X>> <= U)%VS.
+Proof.
+rewrite /subV [@span _ _]unlock genmxE.
+apply: (iffP row_subP) => /= [sXU | sXU i].
+ by move=> _ /seq_tnthP[i ->]; have:= sXU i; rewrite rowK memvK.
+by rewrite rowK -memvK sXU ?mem_tnth.
+Qed.
+
+Lemma sub_span X Y : {subset X <= Y} -> (<<X>> <= <<Y>>)%VS.
+Proof. by move=> sXY; apply/span_subvP=> v /sXY/memv_span. Qed.
+
+Lemma eq_span X Y : X =i Y -> (<<X>> = <<Y>>)%VS.
+Proof.
+by move=> eqXY; apply: subv_anti; rewrite !sub_span // => u; rewrite eqXY.
+Qed.
+
+Lemma span_def X : span X = (\sum_(u <- X) <[u]>)%VS.
+Proof.
+apply/subv_anti/andP; split.
+ by apply/span_subvP=> v Xv; rewrite (big_rem v) // memvE addvSl.
+by rewrite big_tnth; apply/subv_sumP=> i _; rewrite -memvE memv_span ?mem_tnth.
+Qed.
+
+Lemma span_nil : (<<Nil vT>> = 0)%VS.
+Proof. by rewrite span_def big_nil. Qed.
+
+Lemma span_seq1 v : (<<[:: v]>> = <[v]>)%VS.
+Proof. by rewrite span_def big_seq1. Qed.
+
+Lemma span_cons v X : (<<v :: X>> = <[v]> + <<X>>)%VS.
+Proof. by rewrite !span_def big_cons. Qed.
+
+Lemma span_cat X Y : (<<X ++ Y>> = <<X>> + <<Y>>)%VS.
+Proof. by rewrite !span_def big_cat. Qed.
+
+(* Coordinates function; should perhaps be generalized to nat indices. *)
+
+Definition coord_expanded_def n (X : n.-tuple vT) i v :=
+ (v2r v *m pinvmx (b2mx X)) 0 i.
+Definition coord := locked_with span_key coord_expanded_def.
+Canonical coord_unlockable := [unlockable fun coord].
+
+Fact coord_is_scalar n (X : n.-tuple vT) i : scalar (coord X i).
+Proof. by move=> k u v; rewrite unlock linearP mulmxDl -scalemxAl !mxE. Qed.
+Canonical coord_addidive n Xn i := Additive (@coord_is_scalar n Xn i).
+Canonical coord_linear n Xn i := AddLinear (@coord_is_scalar n Xn i).
+
+Lemma coord_span n (X : n.-tuple vT) v :
+ v \in span X -> v = \sum_i coord X i v *: X`_i.
+Proof.
+rewrite memvK span_b2mx genmxE => Xv.
+by rewrite unlock_with mul_b2mx mulmxKpV ?v2rK.
+Qed.
+
+Lemma coord0 i v : coord [tuple 0] i v = 0.
+Proof.
+rewrite unlock /pinvmx rank_rV; case: negP => [[] | _].
+ by apply/eqP/rowP=> j; rewrite !mxE (tnth_nth 0) /= linear0 mxE.
+by rewrite pid_mx_0 !(mulmx0, mul0mx) mxE.
+Qed.
+
+(* Free generator sequences. *)
+
+Lemma nil_free : free (Nil vT).
+Proof. by rewrite /free span_nil dimv0. Qed.
+
+Lemma seq1_free v : free [:: v] = (v != 0).
+Proof. by rewrite /free span_seq1 dim_vline; case: (~~ _). Qed.
+
+Lemma perm_free X Y : perm_eq X Y -> free X = free Y.
+Proof.
+by move=> eqX; rewrite /free (perm_eq_size eqX) (eq_span (perm_eq_mem eqX)).
+Qed.
+
+Lemma free_directv X : free X = (0 \notin X) && directv (\sum_(v <- X) <[v]>).
+Proof.
+have leXi i (v := tnth (in_tuple X) i): true -> \dim <[v]> <= 1 ?= iff (v != 0).
+ by rewrite -seq1_free -span_seq1 => _; apply/leqif_eq/dim_span.
+have [_ /=] := leqif_trans (dimv_sum_leqif _) (leqif_sum leXi).
+rewrite sum1_card card_ord !directvE /= /free andbC span_def !(big_tnth _ _ X).
+by congr (_ = _ && _); rewrite -has_pred1 -all_predC -big_all big_tnth big_andE.
+Qed.
+
+Lemma free_not0 v X : free X -> v \in X -> v != 0.
+Proof. by rewrite free_directv andbC => /andP[_ /memPn]; apply. Qed.
+
+Lemma freeP n (X : n.-tuple vT) :
+ reflect (forall k, \sum_(i < n) k i *: X`_i = 0 -> (forall i, k i = 0))
+ (free X).
+Proof.
+rewrite free_b2mx; apply: (iffP idP) => [t_free k kt0 i | t_free].
+ suffices /rowP/(_ i): \row_i k i = 0 by rewrite !mxE.
+ by apply/(row_free_inj t_free)/r2v_inj; rewrite mul0mx -lin_b2mx kt0 linear0.
+rewrite -kermx_eq0; apply/rowV0P=> rk /sub_kermxP kt0.
+by apply/rowP=> i; rewrite mxE {}t_free // mul_b2mx kt0 linear0.
+Qed.
+
+Lemma coord_free n (X : n.-tuple vT) (i j : 'I_n) :
+ free X -> coord X j (X`_i) = (i == j)%:R.
+Proof.
+rewrite unlock free_b2mx => /row_freeP[Ct CtK]; rewrite -row_b2mx.
+by rewrite -row_mul -[pinvmx _]mulmx1 -CtK 2!mulmxA mulmxKpV // CtK !mxE.
+Qed.
+
+Lemma coord_sum_free n (X : n.-tuple vT) k j :
+ free X -> coord X j (\sum_(i < n) k i *: X`_i) = k j.
+Proof.
+move=> Xfree; rewrite linear_sum (bigD1 j) ?linearZ //= coord_free // eqxx.
+rewrite mulr1 big1 ?addr0 // => i /negPf j'i.
+by rewrite linearZ /= coord_free // j'i mulr0.
+Qed.
+
+Lemma cat_free X Y :
+ free (X ++ Y) = [&& free X, free Y & directv (<<X>> + <<Y>>)].
+Proof.
+rewrite !free_directv mem_cat directvE /= !big_cat -directvE directv_addE /=.
+rewrite negb_or -!andbA; do !bool_congr; rewrite -!span_def.
+by rewrite (sameP eqP directv_addP).
+Qed.
+
+Lemma catl_free Y X : free (X ++ Y) -> free X.
+Proof. by rewrite cat_free => /and3P[]. Qed.
+
+Lemma catr_free X Y : free (X ++ Y) -> free Y.
+Proof. by rewrite cat_free => /and3P[]. Qed.
+
+Lemma filter_free p X : free X -> free (filter p X).
+Proof.
+rewrite -(perm_free (etrans (perm_filterC p X _) (perm_eq_refl X))).
+exact: catl_free.
+Qed.
+
+Lemma free_cons v X : free (v :: X) = (v \notin <<X>>)%VS && free X.
+Proof.
+rewrite (cat_free [:: v]) seq1_free directvEgeq /= span_seq1 dim_vline.
+case: eqP => [-> | _] /=; first by rewrite mem0v.
+rewrite andbC ltnNge (geq_leqif (dimv_leqif_sup _)) ?addvSr //.
+by rewrite subv_add subvv andbT -memvE.
+Qed.
+
+Lemma freeE n (X : n.-tuple vT) :
+ free X = [forall i : 'I_n, X`_i \notin <<drop i.+1 X>>%VS].
+Proof.
+case: X => X /= /eqP <-{n}; rewrite -(big_andE xpredT) /=.
+elim: X => [|v X IH_X] /=; first by rewrite nil_free big_ord0.
+by rewrite free_cons IH_X big_ord_recl drop0.
+Qed.
+
+Lemma freeNE n (X : n.-tuple vT) :
+ ~~ free X = [exists i : 'I_n, X`_i \in <<drop i.+1 X>>%VS].
+Proof. by rewrite freeE -negb_exists negbK. Qed.
+
+Lemma free_uniq X : free X -> uniq X.
+Proof.
+elim: X => //= v b IH_X; rewrite free_cons => /andP[X'v /IH_X->].
+by rewrite (contra _ X'v) // => /memv_span.
+Qed.
+
+Lemma free_span X v (sumX := fun k => \sum_(x <- X) k x *: x) :
+ free X -> v \in <<X>>%VS ->
+ {k | v = sumX k & forall k1, v = sumX k1 -> {in X, k1 =1 k}}.
+Proof.
+rewrite -{2}[X]in_tupleE => freeX /coord_span def_v.
+pose k x := oapp (fun i => coord (in_tuple X) i v) 0 (insub (index x X)).
+exists k => [|k1 {def_v}def_v _ /(nthP 0)[i ltiX <-]].
+ rewrite /sumX (big_nth 0) big_mkord def_v; apply: eq_bigr => i _.
+ by rewrite /k index_uniq ?free_uniq // valK.
+rewrite /k /= index_uniq ?free_uniq // insubT //= def_v.
+by rewrite /sumX (big_nth 0) big_mkord coord_sum_free.
+Qed.
+
+Lemma linear_of_free (rT : lmodType K) X (fX : seq rT) :
+ {f : {linear vT -> rT} | free X -> size fX = size X -> map f X = fX}.
+Proof.
+pose f u := \sum_i coord (in_tuple X) i u *: fX`_i.
+have lin_f: linear f.
+ move=> k u v; rewrite scaler_sumr -big_split; apply: eq_bigr => i _.
+ by rewrite /= scalerA -scalerDl linearP.
+exists (Linear lin_f) => freeX eq_szX.
+apply/esym/(@eq_from_nth _ 0); rewrite ?size_map eq_szX // => i ltiX.
+rewrite (nth_map 0) //= /f (bigD1 (Ordinal ltiX)) //=.
+rewrite big1 => [|j /negbTE neqji]; rewrite (coord_free (Ordinal _)) //.
+ by rewrite eqxx scale1r addr0.
+by rewrite eq_sym neqji scale0r.
+Qed.
+
+(* Subspace bases *)
+
+Lemma span_basis U X : basis_of U X -> <<X>>%VS = U.
+Proof. by case/andP=> /eqP. Qed.
+
+Lemma basis_free U X : basis_of U X -> free X.
+Proof. by case/andP. Qed.
+
+Lemma coord_basis U n (X : n.-tuple vT) v :
+ basis_of U X -> v \in U -> v = \sum_i coord X i v *: X`_i.
+Proof. by move/span_basis <-; apply: coord_span. Qed.
+
+Lemma nil_basis : basis_of 0 (Nil vT).
+Proof. by rewrite /basis_of span_nil eqxx nil_free. Qed.
+
+Lemma seq1_basis v : v != 0 -> basis_of <[v]> [:: v].
+Proof. by move=> nz_v; rewrite /basis_of span_seq1 // eqxx seq1_free. Qed.
+
+Lemma basis_not0 x U X : basis_of U X -> x \in X -> x != 0.
+Proof. by move/basis_free/free_not0; apply. Qed.
+
+Lemma basis_mem x U X : basis_of U X -> x \in X -> x \in U.
+Proof. by move/span_basis=> <- /memv_span. Qed.
+
+Lemma cat_basis U V X Y :
+ directv (U + V) -> basis_of U X -> basis_of V Y -> basis_of (U + V) (X ++ Y).
+Proof.
+move=> dxUV /andP[/eqP defU freeX] /andP[/eqP defV freeY].
+by rewrite /basis_of span_cat cat_free defU defV // eqxx freeX freeY.
+Qed.
+
+Lemma size_basis U n (X : n.-tuple vT) : basis_of U X -> \dim U = n.
+Proof. by case/andP=> /eqP <- /eqnP->; apply: size_tuple. Qed.
+
+Lemma basisEdim X U : basis_of U X = (U <= <<X>>)%VS && (size X <= \dim U).
+Proof.
+apply/andP/idP=> [[defU /eqnP <-]| ]; first by rewrite -eqEdim eq_sym.
+case/andP=> sUX leXU; have leXX := dim_span X.
+rewrite /free eq_sym eqEdim sUX eqn_leq !(leq_trans leXX) //.
+by rewrite (leq_trans leXU) ?dimvS.
+Qed.
+
+Lemma basisEfree X U :
+ basis_of U X = [&& free X, (<<X>> <= U)%VS & \dim U <= size X].
+Proof.
+by rewrite andbC; apply: andb_id2r => freeX; rewrite eqEdim (eqnP freeX).
+Qed.
+
+Lemma perm_basis X Y U : perm_eq X Y -> basis_of U X = basis_of U Y.
+Proof.
+move=> eqXY; congr ((_ == _) && _); last exact: perm_free.
+by apply/eq_span; apply: perm_eq_mem.
+Qed.
+
+Lemma vbasisP U : basis_of U (vbasis U).
+Proof.
+rewrite /basis_of free_b2mx span_b2mx (sameP eqP (vs2mxP _ _)) !genmxE.
+have ->: b2mx (vbasis U) = row_base (vs2mx U).
+ by apply/row_matrixP=> i; rewrite unlock rowK tnth_mktuple r2vK.
+by rewrite row_base_free !eq_row_base submx_refl.
+Qed.
+
+Lemma vbasis_mem v U : v \in (vbasis U) -> v \in U.
+Proof. exact: (basis_mem (vbasisP _)). Qed.
+
+Lemma coord_vbasis v U :
+ v \in U -> v = \sum_(i < \dim U) coord (vbasis U) i v *: (vbasis U)`_i.
+Proof. exact: coord_basis (vbasisP U). Qed.
+
+Section BigSumBasis.
+
+Variables (I : finType) (P : pred I) (Xs : I -> seq vT).
+
+Lemma span_bigcat :
+ (<<\big[cat/[::]]_(i | P i) Xs i>> = \sum_(i | P i) <<Xs i>>)%VS.
+Proof. by rewrite (big_morph _ span_cat span_nil). Qed.
+
+Lemma bigcat_free :
+ directv (\sum_(i | P i) <<Xs i>>) ->
+ (forall i, P i -> free (Xs i)) -> free (\big[cat/[::]]_(i | P i) Xs i).
+Proof.
+rewrite /free directvE /= span_bigcat => /directvP-> /= freeXs.
+rewrite (big_morph _ (@size_cat _) (erefl _)) /=.
+by apply/eqP/eq_bigr=> i /freeXs/eqP.
+Qed.
+
+Lemma bigcat_basis Us (U := (\sum_(i | P i) Us i)%VS) :
+ directv U -> (forall i, P i -> basis_of (Us i) (Xs i)) ->
+ basis_of U (\big[cat/[::]]_(i | P i) Xs i).
+Proof.
+move=> dxU XsUs; rewrite /basis_of span_bigcat.
+have defUs i: P i -> span (Xs i) = Us i by case/XsUs/andP=> /eqP.
+rewrite (eq_bigr _ defUs) eqxx bigcat_free // => [|_ /XsUs/andP[]//].
+apply/directvP; rewrite /= (eq_bigr _ defUs) (directvP dxU) /=.
+by apply/eq_bigr=> i /defUs->.
+Qed.
+
+End BigSumBasis.
+
+End VectorTheory.
+
+Hint Resolve subvv.
+Implicit Arguments subvP [K vT U V].
+Implicit Arguments addv_idPl [K vT U V].
+Implicit Arguments addv_idPr [K vT U V].
+Implicit Arguments memv_addP [K vT U V w].
+Implicit Arguments sumv_sup [K vT I P U Vs].
+Implicit Arguments memv_sumP [K vT I P Us v].
+Implicit Arguments subv_sumP [K vT I P Us V].
+Implicit Arguments capv_idPl [K vT U V].
+Implicit Arguments capv_idPr [K vT U V].
+Implicit Arguments memv_capP [K vT U V w].
+Implicit Arguments bigcapv_inf [K vT I P Us V].
+Implicit Arguments subv_bigcapP [K vT I P U Vs].
+Implicit Arguments directvP [K vT S].
+Implicit Arguments directv_addP [K vT U V].
+Implicit Arguments directv_add_unique [K vT U V].
+Implicit Arguments directv_sumP [K vT I P Us].
+Implicit Arguments directv_sumE [K vT I P Ss].
+Implicit Arguments directv_sum_independent [K vT I P Us].
+Implicit Arguments directv_sum_unique [K vT I P Us].
+Implicit Arguments span_subvP [K vT X U].
+Implicit Arguments freeP [K vT n X].
+
+Prenex Implicits coord.
+Notation directv S := (directv_def (Phantom _ S%VS)).
+
+(* Linear functions over a vectType *)
+Section LfunDefs.
+
+Variable R : ringType.
+Implicit Types aT vT rT : vectType R.
+
+Fact lfun_key : unit. Proof. by []. Qed.
+Definition fun_of_lfun_def aT rT (f : 'Hom(aT, rT)) :=
+ r2v \o mulmxr (f2mx f) \o v2r.
+Definition fun_of_lfun := locked_with lfun_key fun_of_lfun_def.
+Canonical fun_of_lfun_unlockable := [unlockable fun fun_of_lfun].
+Definition linfun_def aT rT (f : aT -> rT) :=
+ Vector.Hom (lin1_mx (v2r \o f \o r2v)).
+Definition linfun := locked_with lfun_key linfun_def.
+Canonical linfun_unlockable := [unlockable fun linfun].
+
+Definition id_lfun vT := @linfun vT vT idfun.
+Definition comp_lfun aT vT rT (f : 'Hom(vT, rT)) (g : 'Hom(aT, vT)) :=
+ linfun (fun_of_lfun f \o fun_of_lfun g).
+
+End LfunDefs.
+
+Coercion fun_of_lfun : Vector.hom >-> Funclass.
+Notation "\1" := (@id_lfun _ _) : lfun_scope.
+Notation "f \o g" := (comp_lfun f g) : lfun_scope.
+
+Section LfunVspaceDefs.
+
+Variable K : fieldType.
+Implicit Types aT rT : vectType K.
+
+Definition inv_lfun aT rT (f : 'Hom(aT, rT)) := Vector.Hom (pinvmx (f2mx f)).
+Definition lker aT rT (f : 'Hom(aT, rT)) := mx2vs (kermx (f2mx f)).
+Fact lfun_img_key : unit. Proof. by []. Qed.
+Definition lfun_img_def aT rT f (U : {vspace aT}) : {vspace rT} :=
+ mx2vs (vs2mx U *m f2mx f).
+Definition lfun_img := locked_with lfun_img_key lfun_img_def.
+Canonical lfun_img_unlockable := [unlockable fun lfun_img].
+Definition lfun_preim aT rT (f : 'Hom(aT, rT)) W :=
+ (lfun_img (inv_lfun f) (W :&: lfun_img f fullv) + lker f)%VS.
+
+End LfunVspaceDefs.
+
+Prenex Implicits linfun lfun_img lker lfun_preim.
+Notation "f ^-1" := (inv_lfun f) : lfun_scope.
+Notation "f @: U" := (lfun_img f%VF%R U) (at level 24) : vspace_scope.
+Notation "f @^-1: W" := (lfun_preim f%VF%R W) (at level 24) : vspace_scope.
+Notation limg f := (lfun_img f fullv).
+
+Section LfunZmodType.
+
+Variables (R : ringType) (aT rT : vectType R).
+Implicit Types f g h : 'Hom(aT, rT).
+
+Canonical lfun_eqMixin := Eval hnf in [eqMixin of 'Hom(aT, rT) by <:].
+Canonical lfun_eqType := EqType 'Hom(aT, rT) lfun_eqMixin.
+Definition lfun_choiceMixin := [choiceMixin of 'Hom(aT, rT) by <:].
+Canonical lfun_choiceType := ChoiceType 'Hom(aT, rT) lfun_choiceMixin.
+
+Fact lfun_is_linear f : linear f.
+Proof. by rewrite unlock; apply: linearP. Qed.
+Canonical lfun_additive f := Additive (lfun_is_linear f).
+Canonical lfun_linear f := AddLinear (lfun_is_linear f).
+
+Lemma lfunE (ff : {linear aT -> rT}) : linfun ff =1 ff.
+Proof. by move=> v; rewrite 2!unlock /= mul_rV_lin1 /= !v2rK. Qed.
+
+Lemma fun_of_lfunK : cancel (@fun_of_lfun R aT rT) linfun.
+Proof.
+move=> f; apply/val_inj/row_matrixP=> i.
+by rewrite 2!unlock /= !rowE mul_rV_lin1 /= !r2vK.
+Qed.
+
+Lemma lfunP f g : f =1 g <-> f = g.
+Proof.
+split=> [eq_fg | -> //]; rewrite -[f]fun_of_lfunK -[g]fun_of_lfunK unlock.
+by apply/val_inj/row_matrixP=> i; rewrite !rowE !mul_rV_lin1 /= eq_fg.
+Qed.
+
+Definition zero_lfun : 'Hom(aT, rT) := linfun \0.
+Definition add_lfun f g := linfun (f \+ g).
+Definition opp_lfun f := linfun (-%R \o f).
+
+Fact lfun_addA : associative add_lfun.
+Proof. by move=> f g h; apply/lfunP=> v; rewrite !lfunE /= !lfunE addrA. Qed.
+
+Fact lfun_addC : commutative add_lfun.
+Proof. by move=> f g; apply/lfunP=> v; rewrite !lfunE /= addrC. Qed.
+
+Fact lfun_add0 : left_id zero_lfun add_lfun.
+Proof. by move=> f; apply/lfunP=> v; rewrite lfunE /= lfunE add0r. Qed.
+
+Lemma lfun_addN : left_inverse zero_lfun opp_lfun add_lfun.
+Proof. by move=> f; apply/lfunP=> v; rewrite !lfunE /= lfunE addNr. Qed.
+
+Definition lfun_zmodMixin := ZmodMixin lfun_addA lfun_addC lfun_add0 lfun_addN.
+Canonical lfun_zmodType := Eval hnf in ZmodType 'Hom(aT, rT) lfun_zmodMixin.
+
+Lemma zero_lfunE x : (0 : 'Hom(aT, rT)) x = 0. Proof. exact: lfunE. Qed.
+Lemma add_lfunE f g x : (f + g) x = f x + g x. Proof. exact: lfunE. Qed.
+Lemma opp_lfunE f x : (- f) x = - f x. Proof. exact: lfunE. Qed.
+Lemma sum_lfunE I (r : seq I) (P : pred I) (fs : I -> 'Hom(aT, rT)) x :
+ (\sum_(i <- r | P i) fs i) x = \sum_(i <- r | P i) fs i x.
+Proof. by elim/big_rec2: _ => [|i _ f _ <-]; rewrite lfunE. Qed.
+
+End LfunZmodType.
+
+Section LfunVectType.
+
+Variables (R : comRingType) (aT rT : vectType R).
+Implicit Types f : 'Hom(aT, rT).
+
+Definition scale_lfun k f := linfun (k \*: f).
+Local Infix "*:l" := scale_lfun (at level 40).
+
+Fact lfun_scaleA k1 k2 f : k1 *:l (k2 *:l f) = (k1 * k2) *:l f.
+Proof. by apply/lfunP=> v; rewrite !lfunE /= lfunE scalerA. Qed.
+
+Fact lfun_scale1 f : 1 *:l f = f.
+Proof. by apply/lfunP=> v; rewrite lfunE /= scale1r. Qed.
+
+Fact lfun_scaleDr k f1 f2 : k *:l (f1 + f2) = k *:l f1 + k *:l f2.
+Proof. by apply/lfunP=> v; rewrite !lfunE /= !lfunE scalerDr. Qed.
+
+Fact lfun_scaleDl f k1 k2 : (k1 + k2) *:l f = k1 *:l f + k2 *:l f.
+Proof. by apply/lfunP=> v; rewrite !lfunE /= !lfunE scalerDl. Qed.
+
+Definition lfun_lmodMixin :=
+ LmodMixin lfun_scaleA lfun_scale1 lfun_scaleDr lfun_scaleDl.
+Canonical lfun_lmodType := Eval hnf in LmodType R 'Hom(aT, rT) lfun_lmodMixin.
+
+Lemma scale_lfunE k f x : (k *: f) x = k *: f x. Proof. exact: lfunE. Qed.
+
+(* GG: exists (Vector.Hom \o vec_mx) fails in the proof below in 8.3, *)
+(* probably because of incomplete type unification. Will it work in 8.4? *)
+Fact lfun_vect_iso : Vector.axiom (Vector.dim aT * Vector.dim rT) 'Hom(aT, rT).
+Proof.
+exists (mxvec \o f2mx) => [a f g|].
+ rewrite /= -linearP /= -[A in _ = mxvec A]/(f2mx (Vector.Hom _)).
+ congr (mxvec (f2mx _)); apply/lfunP=> v; do 2!rewrite lfunE /=.
+ by rewrite unlock /= -linearP mulmxDr scalemxAr.
+apply: Bijective (Vector.Hom \o vec_mx) _ _ => [[A]|A] /=; last exact: vec_mxK.
+by rewrite mxvecK.
+Qed.
+
+Definition lfun_vectMixin := VectMixin lfun_vect_iso.
+Canonical lfun_vectType := VectType R 'Hom(aT, rT) lfun_vectMixin.
+
+End LfunVectType.
+
+Section CompLfun.
+
+Variables (R : ringType) (wT aT vT rT : vectType R).
+Implicit Types (f : 'Hom(vT, rT)) (g : 'Hom(aT, vT)) (h : 'Hom(wT, aT)).
+
+Lemma id_lfunE u: \1%VF u = u :> aT. Proof. exact: lfunE. Qed.
+Lemma comp_lfunE f g u : (f \o g)%VF u = f (g u). Proof. exact: lfunE. Qed.
+
+Lemma comp_lfunA f g h : (f \o (g \o h) = (f \o g) \o h)%VF.
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=. Qed.
+
+Lemma comp_lfun1l f : (\1 \o f)%VF = f.
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=. Qed.
+
+Lemma comp_lfun1r f : (f \o \1)%VF = f.
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=. Qed.
+
+Lemma comp_lfun0l g : (0 \o g)%VF = 0 :> 'Hom(aT, rT).
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=. Qed.
+
+Lemma comp_lfun0r f : (f \o 0)%VF = 0 :> 'Hom(aT, rT).
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=; rewrite linear0. Qed.
+
+Lemma comp_lfunDl f1 f2 g : ((f1 + f2) \o g = (f1 \o g) + (f2 \o g))%VF.
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=. Qed.
+
+Lemma comp_lfunDr f g1 g2 : (f \o (g1 + g2) = (f \o g1) + (f \o g2))%VF.
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=; rewrite linearD. Qed.
+
+Lemma comp_lfunNl f g : ((- f) \o g = - (f \o g))%VF.
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=. Qed.
+
+Lemma comp_lfunNr f g : (f \o (- g) = - (f \o g))%VF.
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=; rewrite linearN. Qed.
+
+End CompLfun.
+
+Definition lfun_simp :=
+ (comp_lfunE, scale_lfunE, opp_lfunE, add_lfunE, sum_lfunE, lfunE).
+
+Section ScaleCompLfun.
+
+Variables (R : comRingType) (aT vT rT : vectType R).
+Implicit Types (f : 'Hom(vT, rT)) (g : 'Hom(aT, vT)).
+
+Lemma comp_lfunZl k f g : (k *: (f \o g) = (k *: f) \o g)%VF.
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=. Qed.
+
+Lemma comp_lfunZr k f g : (k *: (f \o g) = f \o (k *: g))%VF.
+Proof. by apply/lfunP=> u; do !rewrite lfunE /=; rewrite linearZ. Qed.
+
+End ScaleCompLfun.
+
+Section LinearImage.
+
+Variables (K : fieldType) (aT rT : vectType K).
+Implicit Types (f g : 'Hom(aT, rT)) (U V : {vspace aT}) (W : {vspace rT}).
+
+Lemma limgS f U V : (U <= V)%VS -> (f @: U <= f @: V)%VS.
+Proof. by rewrite unlock /subsetv !genmxE; apply: submxMr. Qed.
+
+Lemma limg_line f v : (f @: <[v]> = <[f v]>)%VS.
+Proof.
+apply/eqP; rewrite 2!unlock eqEsubv /subsetv /= r2vK !genmxE.
+by rewrite !(eqmxMr _ (genmxE _)) submx_refl.
+Qed.
+
+Lemma limg0 f : (f @: 0 = 0)%VS. Proof. by rewrite limg_line linear0. Qed.
+
+Lemma memv_img f v U : v \in U -> f v \in (f @: U)%VS.
+Proof. by move=> Uv; rewrite memvE -limg_line limgS. Qed.
+
+Lemma memv_imgP f w U :
+ reflect (exists2 u, u \in U & w = f u) (w \in f @: U)%VS.
+Proof.
+apply: (iffP idP) => [|[u Uu ->]]; last exact: memv_img.
+rewrite 2!unlock memvE /subsetv !genmxE => /submxP[ku Drw].
+exists (r2v (ku *m vs2mx U)); last by rewrite /= r2vK -mulmxA -Drw v2rK.
+by rewrite memvE /subsetv !genmxE r2vK submxMl.
+Qed.
+
+Lemma lim0g U : (0 @: U = 0 :> {vspace rT})%VS.
+Proof.
+apply/eqP; rewrite -subv0; apply/subvP=> _ /memv_imgP[u _ ->].
+by rewrite lfunE rpred0.
+Qed.
+
+Lemma eq_in_limg V f g : {in V, f =1 g} -> (f @: V = g @: V)%VS.
+Proof.
+move=> eq_fg; apply/vspaceP=> y.
+by apply/memv_imgP/memv_imgP=> [][x Vx ->]; exists x; rewrite ?eq_fg.
+Qed.
+
+Lemma limg_add f : {morph lfun_img f : U V / U + V}%VS.
+Proof.
+move=> U V; apply/eqP; rewrite unlock eqEsubv /subsetv /= -genmx_adds.
+by rewrite !genmxE !(eqmxMr _ (genmxE _)) !addsmxMr submx_refl.
+Qed.
+
+Lemma limg_sum f I r (P : pred I) Us :
+ (f @: (\sum_(i <- r | P i) Us i) = \sum_(i <- r | P i) f @: Us i)%VS.
+Proof. exact: (big_morph _ (limg_add f) (limg0 f)). Qed.
+
+Lemma limg_cap f U V : (f @: (U :&: V) <= f @: U :&: f @: V)%VS.
+Proof. by rewrite subv_cap !limgS ?capvSl ?capvSr. Qed.
+
+Lemma limg_bigcap f I r (P : pred I) Us :
+ (f @: (\bigcap_(i <- r | P i) Us i) <= \bigcap_(i <- r | P i) f @: Us i)%VS.
+Proof.
+elim/big_rec2: _ => [|i V U _ sUV]; first exact: subvf.
+by rewrite (subv_trans (limg_cap f _ U)) ?capvS.
+Qed.
+
+Lemma limg_span f X : (f @: <<X>> = <<map f X>>)%VS.
+Proof.
+by rewrite !span_def big_map limg_sum; apply: eq_bigr => x _; rewrite limg_line.
+Qed.
+
+Lemma lfunPn f g : reflect (exists u, f u != g u) (f != g).
+Proof.
+apply: (iffP idP) => [f'g|[x]]; last by apply: contraNneq => /lfunP->.
+suffices /subvPn[_ /memv_imgP[u _ ->]]: ~~ (limg (f - g) <= 0)%VS.
+ by rewrite lfunE /= lfunE /= memv0 subr_eq0; exists u.
+apply: contra f'g => /subvP fg0; apply/eqP/lfunP=> u; apply/eqP.
+by rewrite -subr_eq0 -opp_lfunE -add_lfunE -memv0 fg0 ?memv_img ?memvf.
+Qed.
+
+Lemma inv_lfun_def f : (f \o f^-1 \o f)%VF = f.
+Proof.
+apply/lfunP=> u; do !rewrite lfunE /=; rewrite unlock /= !r2vK.
+by rewrite mulmxKpV ?submxMl.
+Qed.
+
+Lemma limg_lfunVK f : {in limg f, cancel f^-1%VF f}.
+Proof. by move=> _ /memv_imgP[u _ ->]; rewrite -!comp_lfunE inv_lfun_def. Qed.
+
+Lemma lkerE f U : (U <= lker f)%VS = (f @: U == 0)%VS.
+Proof.
+rewrite unlock -dimv_eq0 /dimv /subsetv !genmxE mxrank_eq0.
+by rewrite (sameP sub_kermxP eqP).
+Qed.
+
+Lemma memv_ker f v : (v \in lker f) = (f v == 0).
+Proof. by rewrite -memv0 !memvE subv0 lkerE limg_line. Qed.
+
+Lemma eqlfunP f g v : reflect (f v = g v) (v \in lker (f - g)).
+Proof. by rewrite memv_ker !lfun_simp subr_eq0; apply: eqP. Qed.
+
+Lemma eqlfun_inP V f g : reflect {in V, f =1 g} (V <= lker (f - g))%VS.
+Proof. by apply: (iffP subvP) => E x /E/eqlfunP. Qed.
+
+Lemma limg_ker_compl f U : (f @: (U :\: lker f) = f @: U)%VS.
+Proof.
+rewrite -{2}(addv_diff_cap U (lker f)) limg_add; apply/esym/addv_idPl.
+by rewrite (subv_trans _ (sub0v _)) // subv0 -lkerE capvSr.
+Qed.
+
+Lemma limg_ker_dim f U : (\dim (U :&: lker f) + \dim (f @: U) = \dim U)%N.
+Proof.
+rewrite unlock /dimv /= genmx_cap genmx_id -genmx_cap !genmxE.
+by rewrite addnC mxrank_mul_ker.
+Qed.
+
+Lemma limg_dim_eq f U : (U :&: lker f = 0)%VS -> \dim (f @: U) = \dim U.
+Proof. by rewrite -(limg_ker_dim f U) => ->; rewrite dimv0. Qed.
+
+Lemma limg_basis_of f U X :
+ (U :&: lker f = 0)%VS -> basis_of U X -> basis_of (f @: U) (map f X).
+Proof.
+move=> injUf /andP[/eqP defU /eqnP freeX].
+by rewrite /basis_of /free size_map -limg_span -freeX defU limg_dim_eq ?eqxx.
+Qed.
+
+Lemma lker0P f : reflect (injective f) (lker f == 0%VS).
+Proof.
+rewrite -subv0; apply: (iffP subvP) => [injf u v eq_fuv | injf u].
+ apply/eqP; rewrite -subr_eq0 -memv0 injf //.
+ by rewrite memv_ker linearB /= eq_fuv subrr.
+by rewrite memv_ker memv0 -(inj_eq injf) linear0.
+Qed.
+
+Lemma limg_ker0 f U V : lker f == 0%VS -> (f @: U <= f @: V)%VS = (U <= V)%VS.
+Proof.
+move/lker0P=> injf; apply/idP/idP=> [/subvP sfUV | ]; last exact: limgS.
+by apply/subvP=> u Uu; have /memv_imgP[v Vv /injf->] := sfUV _ (memv_img f Uu).
+Qed.
+
+Lemma eq_limg_ker0 f U V : lker f == 0%VS -> (f @: U == f @: V)%VS = (U == V).
+Proof. by move=> injf; rewrite !eqEsubv !limg_ker0. Qed.
+
+Lemma lker0_lfunK f : lker f == 0%VS -> cancel f f^-1%VF.
+Proof.
+by move/lker0P=> injf u; apply: injf; rewrite limg_lfunVK ?memv_img ?memvf.
+Qed.
+
+Lemma lker0_compVf f : lker f == 0%VS -> (f^-1 \o f = \1)%VF.
+Proof. by move/lker0_lfunK=> fK; apply/lfunP=> u; rewrite !lfunE /= fK. Qed.
+
+End LinearImage.
+
+Implicit Arguments memv_imgP [K aT rT f U w].
+Implicit Arguments lfunPn [K aT rT f g].
+Implicit Arguments lker0P [K aT rT f].
+Implicit Arguments eqlfunP [K aT rT f g v].
+Implicit Arguments eqlfun_inP [K aT rT f g V].
+
+Section FixedSpace.
+
+Variables (K : fieldType) (vT : vectType K).
+Implicit Types (f : 'End(vT)) (U : {vspace vT}).
+
+Definition fixedSpace f : {vspace vT} := lker (f - \1%VF).
+
+Lemma fixedSpaceP f a : reflect (f a = a) (a \in fixedSpace f).
+Proof.
+by rewrite memv_ker add_lfunE opp_lfunE id_lfunE subr_eq0; apply: eqP.
+Qed.
+
+Lemma fixedSpacesP f U : reflect {in U, f =1 id} (U <= fixedSpace f)%VS.
+Proof. by apply: (iffP subvP) => cUf x /cUf/fixedSpaceP. Qed.
+
+Lemma fixedSpace_limg f U : (U <= fixedSpace f -> f @: U = U)%VS.
+Proof.
+move/fixedSpacesP=> cUf; apply/vspaceP=> x.
+by apply/memv_imgP/idP=> [[{x} x Ux ->] | Ux]; last exists x; rewrite ?cUf.
+Qed.
+
+Lemma fixedSpace_id : fixedSpace \1 = {:vT}%VS.
+Proof.
+by apply/vspaceP=> x; rewrite memvf; apply/fixedSpaceP; rewrite lfunE.
+Qed.
+
+End FixedSpace.
+
+Implicit Arguments fixedSpaceP [K vT f a].
+Implicit Arguments fixedSpacesP [K vT f U].
+
+Section LinAut.
+
+Variables (K : fieldType) (vT : vectType K) (f : 'End(vT)).
+Hypothesis kerf0 : lker f == 0%VS.
+
+Lemma lker0_limgf : limg f = fullv.
+Proof.
+by apply/eqP; rewrite eqEdim subvf limg_dim_eq //= (eqP kerf0) capv0.
+Qed.
+
+Lemma lker0_lfunVK : cancel f^-1%VF f.
+Proof. by move=> u; rewrite limg_lfunVK // lker0_limgf memvf. Qed.
+
+Lemma lker0_compfV : (f \o f^-1 = \1)%VF.
+Proof. by apply/lfunP=> u; rewrite !lfunE /= lker0_lfunVK. Qed.
+
+Lemma lker0_compVKf aT g : (f \o (f^-1 \o g))%VF = g :> 'Hom(aT, vT).
+Proof. by rewrite comp_lfunA lker0_compfV comp_lfun1l. Qed.
+
+Lemma lker0_compKf aT g : (f^-1 \o (f \o g))%VF = g :> 'Hom(aT, vT).
+Proof. by rewrite comp_lfunA lker0_compVf ?comp_lfun1l. Qed.
+
+Lemma lker0_compfK rT h : ((h \o f) \o f^-1)%VF = h :> 'Hom(vT, rT).
+Proof. by rewrite -comp_lfunA lker0_compfV comp_lfun1r. Qed.
+
+Lemma lker0_compfVK rT h : ((h \o f^-1) \o f)%VF = h :> 'Hom(vT, rT).
+Proof. by rewrite -comp_lfunA lker0_compVf ?comp_lfun1r. Qed.
+
+End LinAut.
+
+Section LinearImageComp.
+
+Variables (K : fieldType) (aT vT rT : vectType K).
+Implicit Types (f : 'Hom(aT, vT)) (g : 'Hom(vT, rT)) (U : {vspace aT}).
+
+Lemma lim1g U : (\1 @: U)%VS = U.
+Proof.
+have /andP[/eqP <- _] := vbasisP U; rewrite limg_span map_id_in // => u _.
+by rewrite lfunE.
+Qed.
+
+Lemma limg_comp f g U : ((g \o f) @: U = g @: (f @: U))%VS.
+Proof.
+have /andP[/eqP <- _] := vbasisP U; rewrite !limg_span; congr (span _).
+by rewrite -map_comp; apply/eq_map => u; rewrite lfunE.
+Qed.
+
+End LinearImageComp.
+
+Section LinearPreimage.
+
+Variables (K : fieldType) (aT rT : vectType K).
+Implicit Types (f : 'Hom(aT, rT)) (U : {vspace aT}) (V W : {vspace rT}).
+
+Lemma lpreim_cap_limg f W : (f @^-1: (W :&: limg f))%VS = (f @^-1: W)%VS.
+Proof. by rewrite /lfun_preim -capvA capvv. Qed.
+
+Lemma lpreim0 f : (f @^-1: 0)%VS = lker f.
+Proof. by rewrite /lfun_preim cap0v limg0 add0v. Qed.
+
+Lemma lpreimS f V W : (V <= W)%VS-> (f @^-1: V <= f @^-1: W)%VS.
+Proof. by move=> sVW; rewrite addvS // limgS // capvS. Qed.
+
+Lemma lpreimK f W : (W <= limg f)%VS -> (f @: (f @^-1: W))%VS = W.
+Proof.
+move=> sWf; rewrite limg_add (capv_idPl sWf) // -limg_comp.
+have /eqP->: (f @: lker f == 0)%VS by rewrite -lkerE.
+have /andP[/eqP defW _] := vbasisP W; rewrite addv0 -defW limg_span.
+rewrite map_id_in // => x Xx; rewrite lfunE /= limg_lfunVK //.
+by apply: span_subvP Xx; rewrite defW.
+Qed.
+
+Lemma memv_preim f u W : (f u \in W) = (u \in f @^-1: W)%VS.
+Proof.
+apply/idP/idP=> [Wfu | /(memv_img f)]; last first.
+ by rewrite -lpreim_cap_limg lpreimK ?capvSr // => /memv_capP[].
+rewrite -[u](addNKr (f^-1%VF (f u))) memv_add ?memv_img //.
+ by rewrite memv_cap Wfu memv_img ?memvf.
+by rewrite memv_ker addrC linearB /= subr_eq0 limg_lfunVK ?memv_img ?memvf.
+Qed.
+
+End LinearPreimage.
+
+Section LfunAlgebra.
+(* This section is a bit of a place holder: the instances we build here can't *)
+(* be canonical because we are missing an interface for proper vectTypes, *)
+(* would sit between Vector and Falgebra. For now, we just supply structure *)
+(* definitions here and supply actual instances for F-algebras in a submodule *)
+(* of the algebra library (there is currently no actual use of the End(vT) *)
+(* algebra structure). Also note that the unit ring structure is missing. *)
+
+Variables (R : comRingType) (vT : vectType R).
+Hypothesis vT_proper : Vector.dim vT > 0.
+
+Fact lfun1_neq0 : \1%VF != 0 :> 'End(vT).
+Proof.
+apply/eqP=> /lfunP/(_ (r2v (const_mx 1))); rewrite !lfunE /= => /(canRL r2vK).
+by move=> /rowP/(_ (Ordinal vT_proper))/eqP; rewrite linear0 !mxE oner_eq0.
+Qed.
+
+Prenex Implicits comp_lfunA comp_lfun1l comp_lfun1r comp_lfunDl comp_lfunDr.
+
+Definition lfun_comp_ringMixin :=
+ RingMixin comp_lfunA comp_lfun1l comp_lfun1r comp_lfunDl comp_lfunDr
+ lfun1_neq0.
+Definition lfun_comp_ringType := RingType 'End(vT) lfun_comp_ringMixin.
+
+(* In the standard endomorphism ring product is categorical composition. *)
+Definition lfun_ringMixin : GRing.Ring.mixin_of (lfun_zmodType vT vT) :=
+ GRing.converse_ringMixin lfun_comp_ringType.
+Definition lfun_ringType := Eval hnf in RingType 'End(vT) lfun_ringMixin.
+Definition lfun_lalgType := Eval hnf in [lalgType R of 'End(vT)
+ for LalgType R lfun_ringType (fun k x y => comp_lfunZr k y x)].
+Definition lfun_algType := Eval hnf in [algType R of 'End(vT)
+ for AlgType R _ (fun k (x y : lfun_lalgType) => comp_lfunZl k y x)].
+
+End LfunAlgebra.
+
+Section Projection.
+
+Variables (K : fieldType) (vT : vectType K).
+Implicit Types U V : {vspace vT}.
+
+Definition daddv_pi U V := Vector.Hom (proj_mx (vs2mx U) (vs2mx V)).
+Definition projv U := daddv_pi U U^C.
+Definition addv_pi1 U V := daddv_pi (U :\: V) V.
+Definition addv_pi2 U V := daddv_pi V (U :\: V).
+
+Lemma memv_pi U V w : (daddv_pi U V) w \in U.
+Proof. by rewrite unlock memvE /subsetv genmxE /= r2vK proj_mx_sub. Qed.
+
+Lemma memv_proj U w : projv U w \in U. Proof. exact: memv_pi. Qed.
+
+Lemma memv_pi1 U V w : (addv_pi1 U V) w \in U.
+Proof. by rewrite (subvP (diffvSl U V)) ?memv_pi. Qed.
+
+Lemma memv_pi2 U V w : (addv_pi2 U V) w \in V. Proof. exact: memv_pi. Qed.
+
+Lemma daddv_pi_id U V u : (U :&: V = 0)%VS -> u \in U -> daddv_pi U V u = u.
+Proof.
+move/eqP; rewrite -dimv_eq0 memvE /subsetv /dimv !genmxE mxrank_eq0 => /eqP.
+by move=> dxUV Uu; rewrite unlock /= proj_mx_id ?v2rK.
+Qed.
+
+Lemma daddv_pi_proj U V w (pi := daddv_pi U V) :
+ (U :&: V = 0)%VS -> pi (pi w) = pi w.
+Proof. by move/daddv_pi_id=> -> //; apply: memv_pi. Qed.
+
+Lemma daddv_pi_add U V w :
+ (U :&: V = 0)%VS -> (w \in U + V)%VS -> daddv_pi U V w + daddv_pi V U w = w.
+Proof.
+move/eqP; rewrite -dimv_eq0 memvE /subsetv /dimv !genmxE mxrank_eq0 => /eqP.
+by move=> dxUW UVw; rewrite unlock /= -linearD /= add_proj_mx ?v2rK.
+Qed.
+
+Lemma projv_id U u : u \in U -> projv U u = u.
+Proof. exact: daddv_pi_id (capv_compl _). Qed.
+
+Lemma projv_proj U w : projv U (projv U w) = projv U w.
+Proof. exact: daddv_pi_proj (capv_compl _). Qed.
+
+Lemma memv_projC U w : w - projv U w \in (U^C)%VS.
+Proof.
+rewrite -{1}[w](daddv_pi_add (capv_compl U)) ?addv_complf ?memvf //.
+by rewrite addrC addKr memv_pi.
+Qed.
+
+Lemma limg_proj U : limg (projv U) = U.
+Proof.
+apply/vspaceP=> u; apply/memv_imgP/idP=> [[u1 _ ->] | ]; first exact: memv_proj.
+by exists (projv U u); rewrite ?projv_id ?memv_img ?memvf.
+Qed.
+
+Lemma lker_proj U : lker (projv U) = (U^C)%VS.
+Proof.
+apply/eqP; rewrite eqEdim andbC; apply/andP; split.
+ by rewrite dimv_compl -(limg_ker_dim (projv U) fullv) limg_proj addnK capfv.
+by apply/subvP=> v; rewrite memv_ker -{2}[v]subr0 => /eqP <-; apply: memv_projC.
+Qed.
+
+Lemma addv_pi1_proj U V w (pi1 := addv_pi1 U V) : pi1 (pi1 w) = pi1 w.
+Proof. by rewrite daddv_pi_proj // capv_diff. Qed.
+
+Lemma addv_pi2_id U V v : v \in V -> addv_pi2 U V v = v.
+Proof. by apply: daddv_pi_id; rewrite capvC capv_diff. Qed.
+
+Lemma addv_pi2_proj U V w (pi2 := addv_pi2 U V) : pi2 (pi2 w) = pi2 w.
+Proof. by rewrite addv_pi2_id ?memv_pi2. Qed.
+
+Lemma addv_pi1_pi2 U V w :
+ w \in (U + V)%VS -> addv_pi1 U V w + addv_pi2 U V w = w.
+Proof. by rewrite -addv_diff; apply: daddv_pi_add; apply: capv_diff. Qed.
+
+Section Sumv_Pi.
+
+Variables (I : eqType) (r0 : seq I) (P : pred I) (Vs : I -> {vspace vT}).
+
+Let sumv_pi_rec i :=
+ fix loop r := if r is j :: r1 then
+ let V1 := (\sum_(k <- r1) Vs k)%VS in
+ if j == i then addv_pi1 (Vs j) V1 else (loop r1 \o addv_pi2 (Vs j) V1)%VF
+ else 0.
+
+Notation sumV := (\sum_(i <- r0 | P i) Vs i)%VS.
+Definition sumv_pi_for V of V = sumV := fun i => sumv_pi_rec i (filter P r0).
+
+Variables (V : {vspace vT}) (defV : V = sumV).
+
+Lemma memv_sum_pi i v : sumv_pi_for defV i v \in Vs i.
+Proof.
+rewrite /sumv_pi_for.
+elim: (filter P r0) v => [|j r IHr] v /=; first by rewrite lfunE mem0v.
+by case: eqP => [->|_]; rewrite ?lfunE ?memv_pi1 /=.
+Qed.
+
+Lemma sumv_pi_uniq_sum v :
+ uniq (filter P r0) -> v \in V ->
+ \sum_(i <- r0 | P i) sumv_pi_for defV i v = v.
+Proof.
+rewrite /sumv_pi_for defV -!(big_filter r0 P).
+elim: (filter P r0) v => [|i r IHr] v /= => [_ | /andP[r'i /IHr{IHr}IHr]].
+ by rewrite !big_nil memv0 => /eqP.
+rewrite !big_cons eqxx => /addv_pi1_pi2; congr (_ + _ = v).
+rewrite -[_ v]IHr ?memv_pi2 //; apply: eq_big_seq => j /=.
+by case: eqP => [<- /idPn | _]; rewrite ?lfunE.
+Qed.
+
+End Sumv_Pi.
+
+End Projection.
+
+Prenex Implicits daddv_pi projv addv_pi1 addv_pi2.
+Notation sumv_pi V := (sumv_pi_for (erefl V)).
+
+Section SumvPi.
+
+Variable (K : fieldType) (vT : vectType K).
+
+Lemma sumv_pi_sum (I : finType) (P : pred I) Vs v (V : {vspace vT})
+ (defV : V = (\sum_(i | P i) Vs i)%VS) :
+ v \in V -> \sum_(i | P i) sumv_pi_for defV i v = v :> vT.
+Proof. by apply: sumv_pi_uniq_sum; apply: enum_uniq. Qed.
+
+Lemma sumv_pi_nat_sum m n (P : pred nat) Vs v (V : {vspace vT})
+ (defV : V = (\sum_(m <= i < n | P i) Vs i)%VS) :
+ v \in V -> \sum_(m <= i < n | P i) sumv_pi_for defV i v = v :> vT.
+Proof. by apply: sumv_pi_uniq_sum; apply/filter_uniq/iota_uniq. Qed.
+
+End SumvPi.
+
+Section SubVector.
+
+(* Turn a {vspace V} into a vectType *)
+Variable (K : fieldType) (vT : vectType K) (U : {vspace vT}).
+
+Inductive subvs_of : predArgType := Subvs u & u \in U.
+
+Definition vsval w := let: Subvs u _ := w in u.
+Canonical subvs_subType := Eval hnf in [subType for vsval].
+Definition subvs_eqMixin := Eval hnf in [eqMixin of subvs_of by <:].
+Canonical subvs_eqType := Eval hnf in EqType subvs_of subvs_eqMixin.
+Definition subvs_choiceMixin := [choiceMixin of subvs_of by <:].
+Canonical subvs_choiceType := ChoiceType subvs_of subvs_choiceMixin.
+Definition subvs_zmodMixin := [zmodMixin of subvs_of by <:].
+Canonical subvs_zmodType := ZmodType subvs_of subvs_zmodMixin.
+Definition subvs_lmodMixin := [lmodMixin of subvs_of by <:].
+Canonical subvs_lmodType := LmodType K subvs_of subvs_lmodMixin.
+
+Lemma subvsP w : vsval w \in U. Proof. exact: valP. Qed.
+Lemma subvs_inj : injective vsval. Proof. exact: val_inj. Qed.
+Lemma congr_subvs u v : u = v -> vsval u = vsval v. Proof. exact: congr1. Qed.
+
+Lemma vsval_is_linear : linear vsval. Proof. by []. Qed.
+Canonical vsval_additive := Additive vsval_is_linear.
+Canonical vsval_linear := AddLinear vsval_is_linear.
+
+Fact vsproj_key : unit. Proof. by []. Qed.
+Definition vsproj_def u := Subvs (memv_proj U u).
+Definition vsproj := locked_with vsproj_key vsproj_def.
+Canonical vsproj_unlockable := [unlockable fun vsproj].
+
+Lemma vsprojK : {in U, cancel vsproj vsval}.
+Proof. by rewrite unlock; apply: projv_id. Qed.
+Lemma vsvalK : cancel vsval vsproj.
+Proof. by move=> w; exact/val_inj/vsprojK/subvsP. Qed.
+
+Lemma vsproj_is_linear : linear vsproj.
+Proof. by move=> k w1 w2; apply: val_inj; rewrite unlock /= linearP. Qed.
+Canonical vsproj_additive := Additive vsproj_is_linear.
+Canonical vsproj_linear := AddLinear vsproj_is_linear.
+
+Fact subvs_vect_iso : Vector.axiom (\dim U) subvs_of.
+Proof.
+exists (fun w => \row_i coord (vbasis U) i (vsval w)).
+ by move=> k w1 w2; apply/rowP=> i; rewrite !mxE linearP.
+exists (fun rw : 'rV_(\dim U) => vsproj (\sum_i rw 0 i *: (vbasis U)`_i)).
+ move=> w /=; congr (vsproj _ = w): (vsvalK w).
+ by rewrite {1}(coord_vbasis (subvsP w)); apply: eq_bigr => i _; rewrite mxE.
+move=> rw; apply/rowP=> i; rewrite mxE vsprojK.
+ by rewrite coord_sum_free ?(basis_free (vbasisP U)).
+by apply: rpred_sum => j _; rewrite rpredZ ?vbasis_mem ?memt_nth.
+Qed.
+
+Definition subvs_vectMixin := VectMixin subvs_vect_iso.
+Canonical subvs_vectType := VectType K subvs_of subvs_vectMixin.
+
+End SubVector.
+Prenex Implicits vsval vsproj vsvalK.
+Implicit Arguments subvs_inj [[K] [vT] [U] x1 x2].
+Implicit Arguments vsprojK [[K] [vT] [U] x].
+
+Section MatrixVectType.
+
+Variables (R : ringType) (m n : nat).
+
+(* The apparently useless => /= in line 1 of the proof performs some evar *)
+(* expansions that the Ltac interpretation of exists is incapable of doing. *)
+Fact matrix_vect_iso : Vector.axiom (m * n) 'M[R]_(m, n).
+Proof.
+exists mxvec => /=; first exact: linearP.
+by exists vec_mx; [apply: mxvecK | apply: vec_mxK].
+Qed.
+Definition matrix_vectMixin := VectMixin matrix_vect_iso.
+Canonical matrix_vectType := VectType R 'M[R]_(m, n) matrix_vectMixin.
+
+End MatrixVectType.
+
+(* A ring is a one-dimension vector space *)
+Section RegularVectType.
+
+Variable R : ringType.
+
+Fact regular_vect_iso : Vector.axiom 1 R^o.
+Proof.
+exists (fun a => a%:M) => [a b c|]; first by rewrite rmorphD scale_scalar_mx.
+by exists (fun A : 'M_1 => A 0 0) => [a | A]; rewrite ?mxE // -mx11_scalar.
+Qed.
+Definition regular_vectMixin := VectMixin regular_vect_iso.
+Canonical regular_vectType := VectType R R^o regular_vectMixin.
+
+End RegularVectType.
+
+(* External direct product of two vectTypes. *)
+Section ProdVector.
+
+Variables (R : ringType) (vT1 vT2 : vectType R).
+
+Fact pair_vect_iso : Vector.axiom (Vector.dim vT1 + Vector.dim vT2) (vT1 * vT2).
+Proof.
+pose p2r (u : vT1 * vT2) := row_mx (v2r u.1) (v2r u.2).
+pose r2p w := (r2v (lsubmx w) : vT1, r2v (rsubmx w) : vT2).
+have r2pK : cancel r2p p2r by move=> w; rewrite /p2r !r2vK hsubmxK.
+have p2rK : cancel p2r r2p by case=> u v; rewrite /r2p row_mxKl row_mxKr !v2rK.
+have r2p_lin: linear r2p by move=> a u v; congr (_ , _); rewrite /= !linearP.
+by exists p2r; [apply: (@can2_linear _ _ _ (Linear r2p_lin)) | exists r2p].
+Qed.
+Definition pair_vectMixin := VectMixin pair_vect_iso.
+Canonical pair_vectType := VectType R (vT1 * vT2) pair_vectMixin.
+
+End ProdVector.
+
+(* Function from a finType into a ring form a vectype. *)
+Section FunVectType.
+
+Variable (I : finType) (R : ringType) (vT : vectType R).
+
+(* Type unification with exist is again a problem in this proof. *)
+Fact ffun_vect_iso : Vector.axiom (#|I| * Vector.dim vT) {ffun I -> vT}.
+Proof.
+pose fr (f : {ffun I -> vT}) := mxvec (\matrix_(i < #|I|) v2r (f (enum_val i))).
+exists fr => /= [k f g|].
+ rewrite /fr -linearP; congr (mxvec _); apply/matrixP=> i j.
+ by rewrite !mxE /= !ffunE linearP !mxE.
+exists (fun r => [ffun i => r2v (row (enum_rank i) (vec_mx r)) : vT]) => [g|r].
+ by apply/ffunP=> i; rewrite ffunE mxvecK rowK v2rK enum_rankK.
+by apply/(canLR vec_mxK)/matrixP=> i j; rewrite mxE ffunE r2vK enum_valK mxE.
+Qed.
+
+Definition ffun_vectMixin := VectMixin ffun_vect_iso.
+Canonical ffun_vectType := VectType R {ffun I -> vT} ffun_vectMixin.
+
+End FunVectType.
+
+Canonical exp_vectType (K : fieldType) (vT : vectType K) n :=
+ [vectType K of vT ^ n].
+
+(* Solving a tuple of linear equations. *)
+Section Solver.
+
+Variable (K : fieldType) (vT : vectType K).
+Variables (n : nat) (lhs : n.-tuple 'End(vT)) (rhs : n.-tuple vT).
+
+Let lhsf u := finfun ((tnth lhs)^~ u).
+Definition vsolve_eq U := finfun (tnth rhs) \in (linfun lhsf @: U)%VS.
+
+Lemma vsolve_eqP (U : {vspace vT}) :
+ reflect (exists2 u, u \in U & forall i, tnth lhs i u = tnth rhs i)
+ (vsolve_eq U).
+Proof.
+have lhsZ: linear lhsf by move=> a u v; apply/ffunP=> i; rewrite !ffunE linearP.
+apply: (iffP memv_imgP) => [] [u Uu sol_u]; exists u => //.
+ by move=> i; rewrite -[tnth rhs i]ffunE sol_u (lfunE (Linear lhsZ)) ffunE.
+by apply/ffunP=> i; rewrite (lfunE (Linear lhsZ)) !ffunE sol_u.
+Qed.
+
+End Solver.
+