diff options
| author | Enrico Tassi | 2015-03-09 11:07:53 +0100 |
|---|---|---|
| committer | Enrico Tassi | 2015-03-09 11:24:38 +0100 |
| commit | fc84c27eac260dffd8f2fb1cb56d599f1e3486d9 (patch) | |
| tree | c16205f1637c80833a4c4598993c29fa0fd8c373 /mathcomp/field | |
Initial commit
Diffstat (limited to 'mathcomp/field')
| -rw-r--r-- | mathcomp/field/algC.v | 1854 | ||||
| -rw-r--r-- | mathcomp/field/algebraics_fundamentals.v | 867 | ||||
| -rw-r--r-- | mathcomp/field/algnum.v | 835 | ||||
| -rw-r--r-- | mathcomp/field/all.v | 11 | ||||
| -rw-r--r-- | mathcomp/field/closed_field.v | 634 | ||||
| -rw-r--r-- | mathcomp/field/countalg.v | 1107 | ||||
| -rw-r--r-- | mathcomp/field/cyclotomic.v | 320 | ||||
| -rw-r--r-- | mathcomp/field/falgebra.v | 1199 | ||||
| -rw-r--r-- | mathcomp/field/fieldext.v | 1626 | ||||
| -rw-r--r-- | mathcomp/field/finfield.v | 585 | ||||
| -rw-r--r-- | mathcomp/field/galois.v | 1628 | ||||
| -rw-r--r-- | mathcomp/field/separable.v | 995 |
12 files changed, 11661 insertions, 0 deletions
diff --git a/mathcomp/field/algC.v b/mathcomp/field/algC.v new file mode 100644 index 0000000..e035bc9 --- /dev/null +++ b/mathcomp/field/algC.v @@ -0,0 +1,1854 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun ssrnat eqtype seq choice div fintype. +Require Import path bigop finset prime ssralg poly polydiv mxpoly. +Require Import generic_quotient countalg ssrnum ssrint rat intdiv. +Require Import algebraics_fundamentals. + +(******************************************************************************) +(* This file provides an axiomatic construction of the algebraic numbers. *) +(* The construction only assumes the existence of an algebraically closed *) +(* filed with an automorphism of order 2; this amounts to the purely *) +(* algebraic contents of the Fundamenta Theorem of Algebra. *) +(* algC == the closed, countable field of algebraic numbers. *) +(* algCeq, algCring, ..., algCnumField == structures for algC. *) +(* z^* == the complex conjugate of z (:= conjC z). *) +(* sqrtC z == a nonnegative square root of z, i.e., 0 <= sqrt x if 0 <= x. *) +(* n.-root z == more generally, for n > 0, an nth root of z, chosen with a *) +(* minimal non-negative argument for n > 1 (i.e., with a *) +(* maximal real part subject to a nonnegative imaginary part). *) +(* Note that n.-root (-1) is a primitive 2nth root of unity, *) +(* an thus not equal to -1 for n odd > 1 (this will be shown in *) +(* file cyclotomic.v). *) +(* The ssrnum interfaces are implemented for algC as follows: *) +(* x <= y <=> (y - x) is a nonnegative real *) +(* x < y <=> (y - x) is a (strictly) positive real *) +(* `|z| == the complex norm of z, i.e., sqrtC (z * z^* ). *) +(* Creal == the subset of real numbers (:= Num.real for algC). *) +(* In addition, we provide: *) +(* 'i == the imaginary number (:= sqrtC (-1)). *) +(* 'Re z == the real component of z. *) +(* 'Im z == the imaginary component of z. *) +(* Crat == the subset of rational numbers. *) +(* Cint == the subset of integers. *) +(* Cnat == the subset of natural integers. *) +(* getCrat z == some a : rat such that ratr a = z, provided z \in Crat. *) +(* floorC z == for z \in Creal, an m : int s.t. m%:~R <= z < (m + 1)%:~R. *) +(* truncC z == for z >= 0, an n : nat s.t. n%:R <= z < n.+1%:R, else 0%N. *) +(* minCpoly z == the minimal (monic) polynomial over Crat with root z. *) +(* algC_invaut nu == an inverse of nu : {rmorphism algC -> algC}. *) +(* (x %| y)%C <=> y is an integer (Cint) multiple of x; if x or y are *) +(* (x %| y)%Cx of type nat or int they are coerced to algC here. *) +(* The (x %| y)%Cx display form is a workaround for *) +(* design limitations of the Coq Notation facilities. *) +(* (x == y %[mod z])%C <=> x and y differ by an integer (Cint) multiple of z; *) +(* as above, arguments of type nat or int are cast to algC. *) +(* (x != y %[mod z])%C <=> x and y do not differ by an integer multiple of z. *) +(* Note that in file algnum we give an alternative definition of divisibility *) +(* based on algebraic integers, overloading the notation in the %A scope. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +(* The Num mixin for an algebraically closed field with an automorphism of *) +(* order 2, making it into a field of complex numbers. *) +Lemma ComplexNumMixin (L : closedFieldType) (conj : {rmorphism L -> L}) : + involutive conj -> ~ conj =1 id -> + {numL | forall x : NumDomainType L numL, `|x| ^+ 2 = x * conj x}. +Proof. +move=> conjK conj_nt. +have nz2: 2%:R != 0 :> L. + apply/eqP=> char2; apply: conj_nt => e; apply/eqP/idPn=> eJ. + have opp_id x: - x = x :> L. + by apply/esym/eqP; rewrite -addr_eq0 -mulr2n -mulr_natl char2 mul0r. + have{char2} char2: 2 \in [char L] by exact/eqP. + without loss{eJ} eJ: e / conj e = e + 1. + move/(_ (e / (e + conj e))); apply. + rewrite fmorph_div rmorphD conjK -{1}[conj e](addNKr e) mulrDl. + by rewrite opp_id (addrC e) divff // addr_eq0 opp_id. + pose a := e * conj e; have aJ: conj a = a by rewrite rmorphM conjK mulrC. + have [w Dw] := @solve_monicpoly _ 2 (nth 0 [:: e * a; - 1]) isT. + have{Dw} Dw: w ^+ 2 + w = e * a. + by rewrite Dw !big_ord_recl big_ord0 /= mulr1 mulN1r addr0 subrK. + pose b := w + conj w; have bJ: conj b = b by rewrite rmorphD conjK addrC. + have Db2: b ^+ 2 + b = a. + rewrite -Frobenius_autE // rmorphD addrACA Dw /= Frobenius_autE -rmorphX. + by rewrite -rmorphD Dw rmorphM aJ eJ -mulrDl -{1}[e]opp_id addKr mul1r. + have /eqP[] := oner_eq0 L; apply: (addrI b); rewrite addr0 -{2}bJ. + have: (b + e) * (b + conj e) == 0. + rewrite mulrDl 2!mulrDr -/a addrA addr_eq0 opp_id (mulrC e) -addrA. + by rewrite -mulrDr eJ addrAC -{2}[e]opp_id subrr add0r mulr1 Db2. + rewrite mulf_eq0 !addr_eq0 !opp_id => /pred2P[] -> //. + by rewrite {2}eJ rmorphD rmorph1. +have mul2I: injective (fun z : L => z *+ 2). + by move=> x y; rewrite /= -mulr_natl -(mulr_natl y) => /mulfI->. +pose sqrt x : L := sval (sig_eqW (@solve_monicpoly _ 2 (nth 0 [:: x]) isT)). +have sqrtK x: sqrt x ^+ 2 = x. + rewrite /sqrt; case: sig_eqW => /= y ->. + by rewrite !big_ord_recl big_ord0 /= mulr1 mul0r !addr0. +have sqrtE x y: y ^+ 2 = x -> {b : bool | y = (-1) ^+ b * sqrt x}. + move=> Dx; exists (y != sqrt x); apply/eqP; rewrite mulr_sign if_neg. + by case: ifPn => //; apply/implyP; rewrite implyNb -eqf_sqr Dx sqrtK. +pose i := sqrt (- 1). +have sqrMi x: (i * x) ^+ 2 = - x ^+ 2 by rewrite exprMn sqrtK mulN1r. +have iJ : conj i = - i. + have /sqrtE[b]: conj i ^+ 2 = - 1 by rewrite -rmorphX sqrtK rmorphN1. + rewrite mulr_sign -/i; case: b => // Ri. + case: conj_nt => z; wlog zJ: z / conj z = - z. + move/(_ (z - conj z)); rewrite !rmorphB conjK opprB => zJ. + by apply/mul2I/(canRL (subrK _)); rewrite -addrA zJ // addrC subrK. + have [-> | nz_z] := eqVneq z 0; first exact: rmorph0. + have [u Ru [v Rv Dz]]: + exists2 u, conj u = u & exists2 v, conj v = v & (u + z * v) ^+ 2 = z. + - pose y := sqrt z; exists ((y + conj y) / 2%:R). + by rewrite fmorph_div rmorphD conjK addrC rmorph_nat. + exists ((y - conj y) / (z *+ 2)). + rewrite fmorph_div rmorphMn zJ mulNrn invrN mulrN -mulNr rmorphB opprB. + by rewrite conjK. + rewrite -(mulr_natl z) invfM (mulrC z) !mulrA divfK // -mulrDl addrACA. + by rewrite subrr addr0 -mulr2n -mulr_natr mulfK ?Neq0 ?sqrtK. + suffices u0: u = 0 by rewrite -Dz u0 add0r rmorphX rmorphM Rv zJ mulNr sqrrN. + suffices [b Du]: exists b : bool, u = (-1) ^+ b * i * z * v. + apply: mul2I; rewrite mul0rn mulr2n -{2}Ru. + by rewrite Du !rmorphM rmorph_sign Rv Ri zJ !mulrN mulNr subrr. + have/eqP:= zJ; rewrite -addr_eq0 -{1 2}Dz rmorphX rmorphD rmorphM Ru Rv zJ. + rewrite mulNr sqrrB sqrrD addrACA (addrACA (u ^+ 2)) addNr addr0 -!mulr2n. + rewrite -mulrnDl -(mul0rn _ 2) (inj_eq mul2I) /= -[rhs in _ + rhs]opprK. + rewrite -sqrMi subr_eq0 eqf_sqr -mulNr !mulrA. + by case/pred2P=> ->; [exists false | exists true]; rewrite mulr_sign. +pose norm x := sqrt x * conj (sqrt x). +have normK x : norm x ^+ 2 = x * conj x by rewrite exprMn -rmorphX sqrtK. +have normE x y : y ^+ 2 = x -> norm x = y * conj y. + rewrite /norm => /sqrtE[b /(canLR (signrMK b)) <-]. + by rewrite !rmorphM rmorph_sign mulrACA -mulrA signrMK. +have norm_eq0 x : norm x = 0 -> x = 0. + by move/eqP; rewrite mulf_eq0 fmorph_eq0 -mulf_eq0 -expr2 sqrtK => /eqP. +have normM x y : norm (x * y) = norm x * norm y. + by rewrite mulrACA -rmorphM; apply: normE; rewrite exprMn !sqrtK. +have normN x : norm (- x) = norm x. + by rewrite -mulN1r normM {1}/norm iJ mulrN -expr2 sqrtK opprK mul1r. +pose le x y := norm (y - x) == y - x; pose lt x y := (y != x) && le x y. +have posE x: le 0 x = (norm x == x) by rewrite /le subr0. +have leB x y: le x y = le 0 (y - x) by rewrite posE. +have posP x : reflect (exists y, x = y * conj y) (le 0 x). + rewrite posE; apply: (iffP eqP) => [Dx | [y {x}->]]; first by exists (sqrt x). + by rewrite (normE _ _ (normK y)) rmorphM conjK (mulrC (conj _)) -expr2 normK. +have posJ x : le 0 x -> conj x = x. + by case/posP=> {x}u ->; rewrite rmorphM conjK mulrC. +have pos_linear x y : le 0 x -> le 0 y -> le x y || le y x. + move=> pos_x pos_y; rewrite leB -opprB orbC leB !posE normN -eqf_sqr. + by rewrite normK rmorphB !posJ ?subrr. +have sposDl x y : lt 0 x -> le 0 y -> lt 0 (x + y). + have sqrtJ z : le 0 z -> conj (sqrt z) = sqrt z. + rewrite posE -{2}[z]sqrtK -subr_eq0 -mulrBr mulf_eq0 subr_eq0. + by case/pred2P=> ->; rewrite ?rmorph0. + case/andP=> nz_x /sqrtJ uJ /sqrtJ vJ. + set u := sqrt x in uJ; set v := sqrt y in vJ; pose w := u + i * v. + have ->: x + y = w * conj w. + rewrite rmorphD rmorphM iJ uJ vJ mulNr mulrC -subr_sqr sqrMi opprK. + by rewrite !sqrtK. + apply/andP; split; last by apply/posP; exists w. + rewrite -normK expf_eq0 //=; apply: contraNneq nz_x => /norm_eq0 w0. + rewrite -[x]sqrtK expf_eq0 /= -/u -(inj_eq mul2I) !mulr2n -{2}(rmorph0 conj). + by rewrite -w0 rmorphD rmorphM iJ uJ vJ mulNr addrACA subrr addr0. +have sposD x y : lt 0 x -> lt 0 y -> lt 0 (x + y). + by move=> x_gt0 /andP[_]; apply: sposDl. +have normD x y : le (norm (x + y)) (norm x + norm y). + have sposM u v: lt 0 u -> le 0 (u * v) -> le 0 v. + by rewrite /lt !posE normM andbC => /andP[/eqP-> /mulfI/inj_eq->]. + have posD u v: le 0 u -> le 0 v -> le 0 (u + v). + have [-> | nz_u u_ge0 v_ge0] := eqVneq u 0; first by rewrite add0r. + by have /andP[]: lt 0 (u + v) by rewrite sposDl // /lt nz_u. + have le_sqr u v: conj u = u -> le 0 v -> le (u ^+ 2) (v ^+ 2) -> le u v. + move=> Ru v_ge0; have [-> // | nz_u] := eqVneq u 0. + have [u_gt0 | u_le0 _] := boolP (lt 0 u). + by rewrite leB (leB u) subr_sqr mulrC addrC; apply: sposM; apply: sposDl. + rewrite leB posD // posE normN -addr_eq0; apply/eqP. + rewrite /lt nz_u posE -subr_eq0 in u_le0; apply: (mulfI u_le0). + by rewrite mulr0 -subr_sqr normK Ru subrr. + have pos_norm z: le 0 (norm z) by apply/posP; exists (sqrt z). + rewrite le_sqr ?posJ ?posD // sqrrD !normK -normM rmorphD mulrDl !mulrDr. + rewrite addrA addrC !addrA -(addrC (y * conj y)) !addrA. + move: (y * _ + _) => u; rewrite -!addrA leB opprD addrACA {u}subrr add0r -leB. + rewrite {}le_sqr ?posD //. + by rewrite rmorphD !rmorphM !conjK addrC mulrC (mulrC y). + rewrite -mulr2n -mulr_natr exprMn normK -natrX mulr_natr sqrrD mulrACA. + rewrite -rmorphM (mulrC y x) addrAC leB mulrnA mulr2n opprD addrACA. + rewrite subrr addr0 {2}(mulrC x) rmorphM mulrACA -opprB addrAC -sqrrB -sqrMi. + apply/posP; exists (i * (x * conj y - y * conj x)); congr (_ * _). + rewrite !(rmorphM, rmorphB) iJ !conjK mulNr -mulrN opprB. + by rewrite (mulrC x) (mulrC y). +by exists (Num.Mixin normD sposD norm_eq0 pos_linear normM (rrefl _) (rrefl _)). +Qed. + +Module Algebraics. + +Module Type Specification. + +Parameter type : Type. + +Parameter eqMixin : Equality.class_of type. +Canonical eqType := EqType type eqMixin. + +Parameter choiceMixin : Choice.mixin_of type. +Canonical choiceType := ChoiceType type choiceMixin. + +Parameter countMixin : Countable.mixin_of type. +Canonical countType := CountType type countMixin. + +Parameter zmodMixin : GRing.Zmodule.mixin_of type. +Canonical zmodType := ZmodType type zmodMixin. +Canonical countZmodType := [countZmodType of type]. + +Parameter ringMixin : GRing.Ring.mixin_of zmodType. +Canonical ringType := RingType type ringMixin. +Canonical countRingType := [countRingType of type]. + +Parameter unitRingMixin : GRing.UnitRing.mixin_of ringType. +Canonical unitRingType := UnitRingType type unitRingMixin. + +Axiom mulC : @commutative ringType ringType *%R. +Canonical comRingType := ComRingType type mulC. +Canonical comUnitRingType := [comUnitRingType of type]. + +Axiom idomainAxiom : GRing.IntegralDomain.axiom ringType. +Canonical idomainType := IdomainType type idomainAxiom. + +Axiom fieldMixin : GRing.Field.mixin_of unitRingType. +Canonical fieldType := FieldType type fieldMixin. + +Parameter decFieldMixin : GRing.DecidableField.mixin_of unitRingType. +Canonical decFieldType := DecFieldType type decFieldMixin. + +Axiom closedFieldAxiom : GRing.ClosedField.axiom ringType. +Canonical closedFieldType := ClosedFieldType type closedFieldAxiom. + +Parameter numMixin : Num.mixin_of ringType. +Canonical numDomainType := NumDomainType type numMixin. +Canonical numFieldType := [numFieldType of type]. + +Parameter conj : {rmorphism type -> type}. +Axiom conjK : involutive conj. +Axiom normK : forall x, `|x| ^+ 2 = x * conj x. + +Axiom algebraic : integralRange (@ratr unitRingType). + +End Specification. + +Module Implementation : Specification. + +Definition L := tag Fundamental_Theorem_of_Algebraics. + +Definition conjL : {rmorphism L -> L} := + s2val (tagged Fundamental_Theorem_of_Algebraics). + +Fact conjL_K : involutive conjL. +Proof. exact: s2valP (tagged Fundamental_Theorem_of_Algebraics). Qed. + +Fact conjL_nt : ~ conjL =1 id. +Proof. exact: s2valP' (tagged Fundamental_Theorem_of_Algebraics). Qed. + +Definition LnumMixin := ComplexNumMixin conjL_K conjL_nt. +Definition Lnum := NumDomainType L (sval LnumMixin). + +Definition QtoL := [rmorphism of @ratr [numFieldType of Lnum]]. +Notation pQtoL := (map_poly QtoL). + +Definition rootQtoL p_j := + if p_j.1 == 0 then 0 else + (sval (closed_field_poly_normal (pQtoL p_j.1)))`_p_j.2. + +Definition eq_root p_j q_k := rootQtoL p_j == rootQtoL q_k. +Fact eq_root_is_equiv : equiv_class_of eq_root. +Proof. by rewrite /eq_root; split=> [ ? | ? ? | ? ? ? ] // /eqP->. Qed. +Canonical eq_root_equiv := EquivRelPack eq_root_is_equiv. +Definition type : Type := {eq_quot eq_root}%qT. + +Definition eqMixin : Equality.class_of type := EquivQuot.eqMixin _. +Canonical eqType := EqType type eqMixin. + +Definition choiceMixin : Choice.mixin_of type := EquivQuot.choiceMixin _. +Canonical choiceType := ChoiceType type choiceMixin. + +Definition countMixin : Countable.mixin_of type := CanCountMixin (@reprK _ _). +Canonical countType := CountType type countMixin. + +Definition CtoL (u : type) := rootQtoL (repr u). + +Fact CtoL_inj : injective CtoL. +Proof. by move=> u v /eqP eq_uv; rewrite -[u]reprK -[v]reprK; apply/eqmodP. Qed. + +Fact CtoL_P u : integralOver QtoL (CtoL u). +Proof. +rewrite /CtoL /rootQtoL; case: (repr u) => p j /=. +case: (closed_field_poly_normal _) => r Dp /=. +case: ifPn => [_ | nz_p]; first exact: integral0. +have [/(nth_default 0)-> | lt_j_r] := leqP (size r) j; first exact: integral0. +apply/integral_algebraic; exists p; rewrite // Dp -mul_polyC rootM orbC. +by rewrite root_prod_XsubC mem_nth. +Qed. + +Fact LtoC_subproof z : integralOver QtoL z -> {u | CtoL u = z}. +Proof. +case/sig2_eqW=> p mon_p pz0; rewrite /CtoL. +pose j := index z (sval (closed_field_poly_normal (pQtoL p))). +pose u := \pi_type%qT (p, j); exists u; have /eqmodP/eqP-> := reprK u. +rewrite /rootQtoL -if_neg monic_neq0 //; apply: nth_index => /=. +case: (closed_field_poly_normal _) => r /= Dp. +by rewrite Dp (monicP _) ?(monic_map QtoL) // scale1r root_prod_XsubC in pz0. +Qed. + +Definition LtoC z Az := sval (@LtoC_subproof z Az). +Fact LtoC_K z Az : CtoL (@LtoC z Az) = z. +Proof. exact: (svalP (LtoC_subproof Az)). Qed. + +Fact CtoL_K u : LtoC (CtoL_P u) = u. +Proof. by apply: CtoL_inj; rewrite LtoC_K. Qed. + +Definition zero := LtoC (integral0 _). +Definition add u v := LtoC (integral_add (CtoL_P u) (CtoL_P v)). +Definition opp u := LtoC (integral_opp (CtoL_P u)). + +Fact addA : associative add. +Proof. by move=> u v w; apply: CtoL_inj; rewrite !LtoC_K addrA. Qed. + +Fact addC : commutative add. +Proof. by move=> u v; apply: CtoL_inj; rewrite !LtoC_K addrC. Qed. + +Fact add0 : left_id zero add. +Proof. by move=> u; apply: CtoL_inj; rewrite !LtoC_K add0r. Qed. + +Fact addN : left_inverse zero opp add. +Proof. by move=> u; apply: CtoL_inj; rewrite !LtoC_K addNr. Qed. + +Definition zmodMixin := ZmodMixin addA addC add0 addN. +Canonical zmodType := ZmodType type zmodMixin. +Canonical countZmodType := [countZmodType of type]. + +Fact CtoL_is_additive : additive CtoL. +Proof. by move=> u v; rewrite !LtoC_K. Qed. +Canonical CtoL_additive := Additive CtoL_is_additive. + +Definition one := LtoC (integral1 _). +Definition mul u v := LtoC (integral_mul (CtoL_P u) (CtoL_P v)). +Definition inv u := LtoC (integral_inv (CtoL_P u)). + +Fact mulA : associative mul. +Proof. by move=> u v w; apply: CtoL_inj; rewrite !LtoC_K mulrA. Qed. + +Fact mulC : commutative mul. +Proof. by move=> u v; apply: CtoL_inj; rewrite !LtoC_K mulrC. Qed. + +Fact mul1 : left_id one mul. +Proof. by move=> u; apply: CtoL_inj; rewrite !LtoC_K mul1r. Qed. + +Fact mulD : left_distributive mul +%R. +Proof. by move=> u v w; apply: CtoL_inj; rewrite !LtoC_K mulrDl. Qed. + +Fact one_nz : one != 0 :> type. +Proof. by rewrite -(inj_eq CtoL_inj) !LtoC_K oner_eq0. Qed. + +Definition ringMixin := ComRingMixin mulA mulC mul1 mulD one_nz. +Canonical ringType := RingType type ringMixin. +Canonical comRingType := ComRingType type mulC. +Canonical countRingType := [countRingType of type]. + +Fact CtoL_is_multiplicative : multiplicative CtoL. +Proof. by split=> [u v|]; rewrite !LtoC_K. Qed. +Canonical CtoL_rmorphism := AddRMorphism CtoL_is_multiplicative. + +Fact mulVf : GRing.Field.axiom inv. +Proof. +move=> u; rewrite -(inj_eq CtoL_inj) rmorph0 => nz_u. +by apply: CtoL_inj; rewrite !LtoC_K mulVf. +Qed. +Fact inv0 : inv 0 = 0. Proof. by apply: CtoL_inj; rewrite !LtoC_K invr0. Qed. + +Definition unitRingMixin := FieldUnitMixin mulVf inv0. +Canonical unitRingType := UnitRingType type unitRingMixin. +Canonical comUnitRingType := [comUnitRingType of type]. + +Definition fieldMixin := @FieldMixin _ _ mulVf inv0. +Definition idomainAxiom := FieldIdomainMixin fieldMixin. +Canonical idomainType := IdomainType type idomainAxiom. +Canonical fieldType := FieldType type fieldMixin. + +Fact closedFieldAxiom : GRing.ClosedField.axiom ringType. +Proof. +move=> n a n_gt0; pose p := 'X^n - \poly_(i < n) CtoL (a i). +have Ap: {in p : seq L, integralRange QtoL}. + move=> _ /(nthP 0)[j _ <-]; rewrite coefB coefXn coef_poly. + apply: integral_sub; first exact: integral_nat. + by case: ifP => _; [apply: CtoL_P | apply: integral0]. +have sz_p: size p = n.+1. + by rewrite size_addl size_polyXn // size_opp ltnS size_poly. +have [z pz0]: exists z, root p z by apply/closed_rootP; rewrite sz_p eqSS -lt0n. +have Az: integralOver ratr z. + by apply: integral_root Ap; rewrite // -size_poly_gt0 sz_p. +exists (LtoC Az); apply/CtoL_inj; rewrite -[CtoL _]subr0 -(rootP pz0). +rewrite rmorphX /= LtoC_K hornerD hornerXn hornerN opprD addNKr opprK. +rewrite horner_poly rmorph_sum; apply: eq_bigr => k _. +by rewrite rmorphM rmorphX /= LtoC_K. +Qed. + +Definition decFieldMixin := closed_field.closed_fields_QEMixin closedFieldAxiom. +Canonical decFieldType := DecFieldType type decFieldMixin. +Canonical closedFieldType := ClosedFieldType type closedFieldAxiom. + +Fact conj_subproof u : integralOver QtoL (conjL (CtoL u)). +Proof. +have [p mon_p pu0] := CtoL_P u; exists p => //. +rewrite -(fmorph_root conjL) conjL_K map_poly_id // => _ /(nthP 0)[j _ <-]. +by rewrite coef_map fmorph_rat. +Qed. +Fact conj_is_rmorphism : rmorphism (fun u => LtoC (conj_subproof u)). +Proof. +do 2?split=> [u v|]; apply: CtoL_inj; last by rewrite !LtoC_K rmorph1. +- by rewrite LtoC_K 3!{1}rmorphB /= !LtoC_K. +by rewrite LtoC_K 3!{1}rmorphM /= !LtoC_K. +Qed. +Definition conj : {rmorphism type -> type} := RMorphism conj_is_rmorphism. +Lemma conjK : involutive conj. +Proof. by move=> u; apply: CtoL_inj; rewrite !LtoC_K conjL_K. Qed. + +Fact conj_nt : ~ conj =1 id. +Proof. +have [i i2]: exists i : type, i ^+ 2 = -1. + have [i] := @solve_monicpoly _ 2 (nth 0 [:: -1 : type]) isT. + by rewrite !big_ord_recl big_ord0 /= mul0r mulr1 !addr0; exists i. +move/(_ i)/(congr1 CtoL); rewrite LtoC_K => iL_J. +have/ltr_geF/idP[] := @ltr01 Lnum; rewrite -oppr_ge0 -(rmorphN1 CtoL_rmorphism). +rewrite -i2 rmorphX /= expr2 -{2}iL_J -(svalP LnumMixin). +by rewrite exprn_ge0 ?normr_ge0. +Qed. + +Definition numMixin := sval (ComplexNumMixin conjK conj_nt). +Canonical numDomainType := NumDomainType type numMixin. +Canonical numFieldType := [numFieldType of type]. + +Lemma normK u : `|u| ^+ 2 = u * conj u. +Proof. exact: svalP (ComplexNumMixin conjK conj_nt) u. Qed. + +Lemma algebraic : integralRange (@ratr unitRingType). +Proof. +move=> u; have [p mon_p pu0] := CtoL_P u; exists p => {mon_p}//. +rewrite -(fmorph_root CtoL_rmorphism) -map_poly_comp; congr (root _ _): pu0. +by apply/esym/eq_map_poly; apply: fmorph_eq_rat. +Qed. + +End Implementation. + +Definition divisor := Implementation.type. + +Module Internals. + +Import Implementation. + +Local Notation algC := type. +Local Notation "z ^*" := (conj z) (at level 2, format "z ^*") : ring_scope. +Local Notation QtoC := (ratr : rat -> algC). +Local Notation QtoCm := [rmorphism of QtoC]. +Local Notation pQtoC := (map_poly QtoC). +Local Notation ZtoQ := (intr : int -> rat). +Local Notation ZtoC := (intr : int -> algC). +Local Notation Creal := (Num.real : qualifier 0 algC). + +Fact algCi_subproof : {i : algC | i ^+ 2 = -1}. +Proof. exact: imaginary_exists. Qed. + +Let Re2 z := z + z^*. +Definition nnegIm z := 0 <= sval algCi_subproof * (z^* - z). +Definition argCle y z := nnegIm z ==> nnegIm y && (Re2 z <= Re2 y). + +CoInductive rootC_spec n (x : algC) : Type := + RootCspec (y : algC) of if (n > 0)%N then y ^+ n = x else y = 0 + & forall z, (n > 0)%N -> z ^+ n = x -> argCle y z. + +Fact rootC_subproof n x : rootC_spec n x. +Proof. +have realRe2 u : Re2 u \is Creal. + rewrite realEsqr expr2 {2}/Re2 -{2}[u]conjK addrC -rmorphD -normK. + by rewrite exprn_ge0 ?normr_ge0. +have argCtotal : total argCle. + move=> u v; rewrite /total /argCle. + by do 2!case: (nnegIm _) => //; rewrite ?orbT //= real_leVge. +have argCtrans : transitive argCle. + move=> u v w /implyP geZuv /implyP geZvw; apply/implyP. + by case/geZvw/andP=> /geZuv/andP[-> geRuv] /ler_trans->. +pose p := 'X^n - (x *+ (n > 0))%:P; have [r0 Dp] := closed_field_poly_normal p. +have sz_p: size p = n.+1. + rewrite size_addl ?size_polyXn // ltnS size_opp size_polyC mulrn_eq0. + by case: posnP => //; case: negP. +pose r := sort argCle r0; have r_arg: sorted argCle r by apply: sort_sorted. +have{Dp} Dp: p = \prod_(z <- r) ('X - z%:P). + rewrite Dp lead_coefE sz_p coefB coefXn coefC -mulrb -mulrnA mulnb lt0n andNb. + rewrite subr0 eqxx scale1r; apply: eq_big_perm. + by rewrite perm_eq_sym perm_sort. +have mem_rP z: (n > 0)%N -> reflect (z ^+ n = x) (z \in r). + move=> n_gt0; rewrite -root_prod_XsubC -Dp rootE !hornerE hornerXn n_gt0. + by rewrite subr_eq0; apply: eqP. +exists r`_0 => [|z n_gt0 /(mem_rP z n_gt0) r_z]. + have sz_r: size r = n by apply: succn_inj; rewrite -sz_p Dp size_prod_XsubC. + case: posnP => [n0 | n_gt0]; first by rewrite nth_default // sz_r n0. + by apply/mem_rP=> //; rewrite mem_nth ?sz_r. +case: {Dp mem_rP}r r_z r_arg => // y r1; rewrite inE => /predU1P[-> _|r1z]. + by apply/implyP=> ->; rewrite lerr. +by move/(order_path_min argCtrans)/allP->. +Qed. + +CoInductive getCrat_spec : Type := GetCrat_spec CtoQ of cancel QtoC CtoQ. + +Fact getCrat_subproof : getCrat_spec. +Proof. +have isQ := rat_algebraic_decidable algebraic. +exists (fun z => if isQ z is left Qz then sval (sig_eqW Qz) else 0) => a. +case: (isQ _) => [Qa | []]; last by exists a. +by case: (sig_eqW _) => b /= /fmorph_inj. +Qed. + +Fact floorC_subproof x : {m | x \is Creal -> ZtoC m <= x < ZtoC (m + 1)}. +Proof. +have [Rx | _] := boolP (x \is Creal); last by exists 0. +without loss x_ge0: x Rx / x >= 0. + have [x_ge0 | /ltrW x_le0] := real_ger0P Rx; first exact. + case/(_ (- x)) => [||m /(_ isT)]; rewrite ?rpredN ?oppr_ge0 //. + rewrite ler_oppr ltr_oppl -!rmorphN opprD /= ltr_neqAle ler_eqVlt. + case: eqP => [-> _ | _ /and3P[lt_x_m _ le_m_x]]. + by exists (- m) => _; rewrite lerr rmorphD ltr_addl ltr01. + by exists (- m - 1); rewrite le_m_x subrK. +have /ex_minnP[n lt_x_n1 min_n]: exists n, x < n.+1%:R. + have [n le_x_n] := rat_algebraic_archimedean algebraic x. + by exists n; rewrite -(ger0_norm x_ge0) (ltr_trans le_x_n) ?ltr_nat. +exists n%:Z => _; rewrite addrC -intS lt_x_n1 andbT. +case Dn: n => // [n1]; rewrite -Dn. +have [||//|] := @real_lerP _ n%:R x; rewrite ?rpred_nat //. +by rewrite Dn => /min_n; rewrite Dn ltnn. +Qed. + +Fact minCpoly_subproof (x : algC) : + {p | p \is monic & forall q, root (pQtoC q) x = (p %| q)%R}. +Proof. +have isQ := rat_algebraic_decidable algebraic. +have [p [mon_p px0 irr_p]] := minPoly_decidable_closure isQ (algebraic x). +exists p => // q; apply/idP/idP=> [qx0 | /dvdpP[r ->]]; last first. + by rewrite rmorphM rootM px0 orbT. +suffices /eqp_dvdl <-: gcdp p q %= p by apply: dvdp_gcdr. +rewrite irr_p ?dvdp_gcdl ?gtn_eqF // -(size_map_poly QtoCm) gcdp_map /=. +rewrite (@root_size_gt1 _ x) ?root_gcd ?px0 //. +by rewrite gcdp_eq0 negb_and map_poly_eq0 monic_neq0. +Qed. + +Definition algC_divisor (x : algC) := x : divisor. +Definition int_divisor m := m%:~R : divisor. +Definition nat_divisor n := n%:R : divisor. + +End Internals. + +Module Import Exports. + +Import Implementation Internals. + +Notation algC := type. +Notation conjC := conj. +Delimit Scope C_scope with C. +Delimit Scope C_core_scope with Cc. +Delimit Scope C_expanded_scope with Cx. +Open Scope C_core_scope. +Notation "x ^*" := (conjC x) (at level 2, format "x ^*") : C_core_scope. +Notation "x ^*" := x^* (only parsing) : C_scope. + +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical zmodType. +Canonical countZmodType. +Canonical ringType. +Canonical countRingType. +Canonical unitRingType. +Canonical comRingType. +Canonical comUnitRingType. +Canonical idomainType. +Canonical numDomainType. +Canonical fieldType. +Canonical numFieldType. +Canonical decFieldType. +Canonical closedFieldType. + +Notation algCeq := eqType. +Notation algCzmod := zmodType. +Notation algCring := ringType. +Notation algCuring := unitRingType. +Notation algCnum := numDomainType. +Notation algCfield := fieldType. +Notation algCnumField := numFieldType. + +Definition rootC n x := let: RootCspec y _ _ := rootC_subproof n x in y. +Notation "n .-root" := (rootC n) (at level 2, format "n .-root") : C_core_scope. +Notation "n .-root" := (rootC n) (only parsing) : C_scope. +Notation sqrtC := 2.-root. + +Definition algCi := sqrtC (-1). +Notation "'i" := algCi (at level 0) : C_core_scope. +Notation "'i" := 'i (only parsing) : C_scope. + +Definition algRe x := (x + x^*) / 2%:R. +Definition algIm x := 'i * (x^* - x) / 2%:R. +Notation "'Re z" := (algRe z) (at level 10, z at level 8) : C_core_scope. +Notation "'Im z" := (algIm z) (at level 10, z at level 8) : C_core_scope. +Notation "'Re z" := ('Re z) (only parsing) : C_scope. +Notation "'Im z" := ('Im z) (only parsing) : C_scope. + +Notation Creal := (@Num.Def.Rreal numDomainType). + +Definition getCrat := let: GetCrat_spec CtoQ _ := getCrat_subproof in CtoQ. +Definition Crat : pred_class := fun x : algC => ratr (getCrat x) == x. + +Definition floorC x := sval (floorC_subproof x). +Definition Cint : pred_class := fun x : algC => (floorC x)%:~R == x. + +Definition truncC x := if x >= 0 then `|floorC x|%N else 0%N. +Definition Cnat : pred_class := fun x : algC => (truncC x)%:R == x. + +Definition minCpoly x : {poly algC} := + let: exist2 p _ _ := minCpoly_subproof x in map_poly ratr p. + +Coercion nat_divisor : nat >-> divisor. +Coercion int_divisor : int >-> divisor. +Coercion algC_divisor : algC >-> divisor. + +Lemma nCdivE (p : nat) : p = p%:R :> divisor. Proof. by []. Qed. +Lemma zCdivE (p : int) : p = p%:~R :> divisor. Proof. by []. Qed. +Definition CdivE := (nCdivE, zCdivE). + +Definition dvdC (x : divisor) : pred_class := + fun y : algC => if x == 0 then y == 0 else y / x \in Cint. +Notation "x %| y" := (y \in dvdC x) : C_expanded_scope. +Notation "x %| y" := (@in_mem divisor y (mem (dvdC x))) : C_scope. + +Definition eqCmod (e x y : divisor) := (e %| x - y)%C. + +Notation "x == y %[mod e ]" := (eqCmod e x y) : C_scope. +Notation "x != y %[mod e ]" := (~~ (x == y %[mod e])%C) : C_scope. + +End Exports. + +End Algebraics. + +Export Algebraics.Exports. + +Section AlgebraicsTheory. + +Implicit Types (x y z : algC) (n : nat) (m : int) (b : bool). +Import Algebraics.Internals. + +Local Notation ZtoQ := (intr : int -> rat). +Local Notation ZtoC := (intr : int -> algC). +Local Notation QtoC := (ratr : rat -> algC). +Local Notation QtoCm := [rmorphism of QtoC]. +Local Notation CtoQ := getCrat. +Local Notation intrp := (map_poly intr). +Local Notation pZtoQ := (map_poly ZtoQ). +Local Notation pZtoC := (map_poly ZtoC). +Local Notation pQtoC := (map_poly ratr). +Local Hint Resolve (@intr_inj _ : injective ZtoC). + +(* Specialization of a few basic ssrnum order lemmas. *) + +Definition eqC_nat n p : (n%:R == p%:R :> algC) = (n == p) := eqr_nat _ n p. +Definition leC_nat n p : (n%:R <= p%:R :> algC) = (n <= p)%N := ler_nat _ n p. +Definition ltC_nat n p : (n%:R < p%:R :> algC) = (n < p)%N := ltr_nat _ n p. +Definition Cchar : [char algC] =i pred0 := @char_num _. + +(* This can be used in the converse direction to evaluate assertions over *) +(* manifest rationals, such as 3%:R^-1 + 7%:%^-1 < 2%:%^-1 :> algC. *) +(* Missing norm and integer exponent, due to gaps in ssrint and rat. *) +Definition CratrE := + let CnF := Algebraics.Implementation.numFieldType in + let QtoCm := ratr_rmorphism CnF in + ((rmorph0 QtoCm, rmorph1 QtoCm, rmorphMn QtoCm, rmorphN QtoCm, rmorphD QtoCm), + (rmorphM QtoCm, rmorphX QtoCm, fmorphV QtoCm), + (rmorphMz QtoCm, rmorphXz QtoCm, @ratr_norm CnF, @ratr_sg CnF), + =^~ (@ler_rat CnF, @ltr_rat CnF, (inj_eq (fmorph_inj QtoCm)))). + +Definition CintrE := + let CnF := Algebraics.Implementation.numFieldType in + let ZtoCm := intmul1_rmorphism CnF in + ((rmorph0 ZtoCm, rmorph1 ZtoCm, rmorphMn ZtoCm, rmorphN ZtoCm, rmorphD ZtoCm), + (rmorphM ZtoCm, rmorphX ZtoCm), + (rmorphMz ZtoCm, @intr_norm CnF, @intr_sg CnF), + =^~ (@ler_int CnF, @ltr_int CnF, (inj_eq (@intr_inj CnF)))). + +Let nz2 : 2%:R != 0 :> algC. Proof. by rewrite -!CintrE. Qed. + +(* Conjugation and norm. *) + +Definition conjCK : involutive conjC := Algebraics.Implementation.conjK. +Definition normCK x : `|x| ^+ 2 = x * x^* := Algebraics.Implementation.normK x. +Definition algC_algebraic x := Algebraics.Implementation.algebraic x. + +Lemma normCKC x : `|x| ^+ 2 = x^* * x. Proof. by rewrite normCK mulrC. Qed. + +Lemma mul_conjC_ge0 x : 0 <= x * x^*. +Proof. by rewrite -normCK exprn_ge0 ?normr_ge0. Qed. + +Lemma mul_conjC_gt0 x : (0 < x * x^*) = (x != 0). +Proof. +have [->|x_neq0] := altP eqP; first by rewrite rmorph0 mulr0. +by rewrite -normCK exprn_gt0 ?normr_gt0. +Qed. + +Lemma mul_conjC_eq0 x : (x * x^* == 0) = (x == 0). +Proof. by rewrite -normCK expf_eq0 normr_eq0. Qed. + +Lemma conjC_ge0 x : (0 <= x^*) = (0 <= x). +Proof. +wlog suffices: x / 0 <= x -> 0 <= x^*. + by move=> IH; apply/idP/idP=> /IH; rewrite ?conjCK. +rewrite le0r => /predU1P[-> | x_gt0]; first by rewrite rmorph0. +by rewrite -(pmulr_rge0 _ x_gt0) mul_conjC_ge0. +Qed. + +Lemma conjC_nat n : (n%:R)^* = n%:R. Proof. exact: rmorph_nat. Qed. +Lemma conjC0 : 0^* = 0. Proof. exact: rmorph0. Qed. +Lemma conjC1 : 1^* = 1. Proof. exact: rmorph1. Qed. +Lemma conjC_eq0 x : (x^* == 0) = (x == 0). Proof. exact: fmorph_eq0. Qed. + +Lemma invC_norm x : x^-1 = `|x| ^- 2 * x^*. +Proof. +have [-> | nx_x] := eqVneq x 0; first by rewrite conjC0 mulr0 invr0. +by rewrite normCK invfM divfK ?conjC_eq0. +Qed. + +(* Real number subset. *) + +Lemma Creal0 : 0 \is Creal. Proof. exact: rpred0. Qed. +Lemma Creal1 : 1 \is Creal. Proof. exact: rpred1. Qed. +Hint Resolve Creal0 Creal1. (* Trivial cannot resolve a general real0 hint. *) + +Lemma CrealE x : (x \is Creal) = (x^* == x). +Proof. +rewrite realEsqr ger0_def normrX normCK. +by have [-> | /mulfI/inj_eq-> //] := eqVneq x 0; rewrite rmorph0 !eqxx. +Qed. + +Lemma CrealP {x} : reflect (x^* = x) (x \is Creal). +Proof. by rewrite CrealE; apply: eqP. Qed. + +Lemma conj_Creal x : x \is Creal -> x^* = x. +Proof. by move/CrealP. Qed. + +Lemma conj_normC z : `|z|^* = `|z|. +Proof. by rewrite conj_Creal ?normr_real. Qed. + +Lemma geC0_conj x : 0 <= x -> x^* = x. +Proof. by move=> /ger0_real/CrealP. Qed. + +Lemma geC0_unit_exp x n : 0 <= x -> (x ^+ n.+1 == 1) = (x == 1). +Proof. by move=> x_ge0; rewrite pexpr_eq1. Qed. + +(* Elementary properties of roots. *) + +Ltac case_rootC := rewrite /rootC; case: (rootC_subproof _ _). + +Lemma root0C x : 0.-root x = 0. Proof. by case_rootC. Qed. + +Lemma rootCK n : (n > 0)%N -> cancel n.-root (fun x => x ^+ n). +Proof. by case: n => //= n _ x; case_rootC. Qed. + +Lemma root1C x : 1.-root x = x. Proof. exact: (@rootCK 1). Qed. + +Lemma rootC0 n : n.-root 0 = 0. +Proof. +have [-> | n_gt0] := posnP n; first by rewrite root0C. +by have /eqP := rootCK n_gt0 0; rewrite expf_eq0 n_gt0 /= => /eqP. +Qed. + +Lemma rootC_inj n : (n > 0)%N -> injective n.-root. +Proof. by move/rootCK/can_inj. Qed. + +Lemma eqr_rootC n : (n > 0)%N -> {mono n.-root : x y / x == y}. +Proof. by move/rootC_inj/inj_eq. Qed. + +Lemma rootC_eq0 n x : (n > 0)%N -> (n.-root x == 0) = (x == 0). +Proof. by move=> n_gt0; rewrite -{1}(rootC0 n) eqr_rootC. Qed. + +(* Rectangular coordinates. *) + +Lemma sqrCi : 'i ^+ 2 = -1. Proof. exact: rootCK. Qed. + +Lemma nonRealCi : 'i \isn't Creal. +Proof. by rewrite realEsqr sqrCi oppr_ge0 ltr_geF ?ltr01. Qed. + +Lemma neq0Ci : 'i != 0. +Proof. by apply: contraNneq nonRealCi => ->; apply: real0. Qed. + +Lemma normCi : `|'i| = 1. +Proof. +apply/eqP; rewrite -(@pexpr_eq1 _ _ 2) ?normr_ge0 //. +by rewrite -normrX sqrCi normrN1. +Qed. + +Lemma invCi : 'i^-1 = - 'i. +Proof. by rewrite -div1r -[1]opprK -sqrCi mulNr mulfK ?neq0Ci. Qed. + +Lemma conjCi : 'i^* = - 'i. +Proof. by rewrite -invCi invC_norm normCi expr1n invr1 mul1r. Qed. + +Lemma algCrect x : x = 'Re x + 'i * 'Im x. +Proof. +rewrite 2!mulrA -expr2 sqrCi mulN1r opprB -mulrDl addrACA subrr addr0. +by rewrite -mulr2n -mulr_natr mulfK. +Qed. + +Lemma Creal_Re x : 'Re x \is Creal. +Proof. by rewrite CrealE fmorph_div rmorph_nat rmorphD conjCK addrC. Qed. + +Lemma Creal_Im x : 'Im x \is Creal. +Proof. +rewrite CrealE fmorph_div rmorph_nat rmorphM rmorphB conjCK. +by rewrite conjCi -opprB mulrNN. +Qed. +Hint Resolve Creal_Re Creal_Im. + +Fact algRe_is_additive : additive algRe. +Proof. by move=> x y; rewrite /algRe rmorphB addrACA -opprD mulrBl. Qed. +Canonical algRe_additive := Additive algRe_is_additive. + +Fact algIm_is_additive : additive algIm. +Proof. +by move=> x y; rewrite /algIm rmorphB opprD addrACA -opprD mulrBr mulrBl. +Qed. +Canonical algIm_additive := Additive algIm_is_additive. + +Lemma Creal_ImP z : reflect ('Im z = 0) (z \is Creal). +Proof. +rewrite CrealE -subr_eq0 -(can_eq (mulKf neq0Ci)) mulr0. +by rewrite -(can_eq (divfK nz2)) mul0r; apply: eqP. +Qed. + +Lemma Creal_ReP z : reflect ('Re z = z) (z \in Creal). +Proof. +rewrite (sameP (Creal_ImP z) eqP) -(can_eq (mulKf neq0Ci)) mulr0. +by rewrite -(inj_eq (addrI ('Re z))) addr0 -algCrect eq_sym; apply: eqP. +Qed. + +Lemma algReMl : {in Creal, forall x, {morph algRe : z / x * z}}. +Proof. +by move=> x Rx z /=; rewrite /algRe rmorphM (conj_Creal Rx) -mulrDr -mulrA. +Qed. + +Lemma algReMr : {in Creal, forall x, {morph algRe : z / z * x}}. +Proof. by move=> x Rx z /=; rewrite mulrC algReMl // mulrC. Qed. + +Lemma algImMl : {in Creal, forall x, {morph algIm : z / x * z}}. +Proof. +by move=> x Rx z; rewrite /algIm rmorphM (conj_Creal Rx) -mulrBr mulrCA !mulrA. +Qed. + +Lemma algImMr : {in Creal, forall x, {morph algIm : z / z * x}}. +Proof. by move=> x Rx z /=; rewrite mulrC algImMl // mulrC. Qed. + +Lemma algRe_i : 'Re 'i = 0. Proof. by rewrite /algRe conjCi subrr mul0r. Qed. + +Lemma algIm_i : 'Im 'i = 1. +Proof. +rewrite /algIm conjCi -opprD mulrN -mulr2n mulrnAr ['i * _]sqrCi. +by rewrite mulNrn opprK divff. +Qed. + +Lemma algRe_conj z : 'Re z^* = 'Re z. +Proof. by rewrite /algRe addrC conjCK. Qed. + +Lemma algIm_conj z : 'Im z^* = - 'Im z. +Proof. by rewrite /algIm -mulNr -mulrN opprB conjCK. Qed. + +Lemma algRe_rect : {in Creal &, forall x y, 'Re (x + 'i * y) = x}. +Proof. +move=> x y Rx Ry; rewrite /= raddfD /= (Creal_ReP x Rx). +by rewrite algReMr // algRe_i mul0r addr0. +Qed. + +Lemma algIm_rect : {in Creal &, forall x y, 'Im (x + 'i * y) = y}. +Proof. +move=> x y Rx Ry; rewrite /= raddfD /= (Creal_ImP x Rx) add0r. +by rewrite algImMr // algIm_i mul1r. +Qed. + +Lemma conjC_rect : {in Creal &, forall x y, (x + 'i * y)^* = x - 'i * y}. +Proof. +by move=> x y Rx Ry; rewrite /= rmorphD rmorphM conjCi mulNr !conj_Creal. +Qed. + +Lemma addC_rect x1 y1 x2 y2 : + (x1 + 'i * y1) + (x2 + 'i * y2) = x1 + x2 + 'i * (y1 + y2). +Proof. by rewrite addrACA -mulrDr. Qed. + +Lemma oppC_rect x y : - (x + 'i * y) = - x + 'i * (- y). +Proof. by rewrite mulrN -opprD. Qed. + +Lemma subC_rect x1 y1 x2 y2 : + (x1 + 'i * y1) - (x2 + 'i * y2) = x1 - x2 + 'i * (y1 - y2). +Proof. by rewrite oppC_rect addC_rect. Qed. + +Lemma mulC_rect x1 y1 x2 y2 : + (x1 + 'i * y1) * (x2 + 'i * y2) + = x1 * x2 - y1 * y2 + 'i * (x1 * y2 + x2 * y1). +Proof. +rewrite mulrDl !mulrDr mulrCA -!addrA mulrAC -mulrA; congr (_ + _). +by rewrite mulrACA -expr2 sqrCi mulN1r addrA addrC. +Qed. + +Lemma normC2_rect : + {in Creal &, forall x y, `|x + 'i * y| ^+ 2 = x ^+ 2 + y ^+ 2}. +Proof. +move=> x y Rx Ry; rewrite /= normCK rmorphD rmorphM conjCi !conj_Creal //. +by rewrite mulrC mulNr -subr_sqr exprMn sqrCi mulN1r opprK. +Qed. + +Lemma normC2_Re_Im z : `|z| ^+ 2 = 'Re z ^+ 2 + 'Im z ^+ 2. +Proof. by rewrite -normC2_rect -?algCrect. Qed. + +Lemma invC_rect : + {in Creal &, forall x y, (x + 'i * y)^-1 = (x - 'i * y) / (x ^+ 2 + y ^+ 2)}. +Proof. +by move=> x y Rx Ry; rewrite /= invC_norm conjC_rect // mulrC normC2_rect. +Qed. + +Lemma lerif_normC_Re_Creal z : `|'Re z| <= `|z| ?= iff (z \is Creal). +Proof. +rewrite -(mono_in_lerif ler_sqr); try by rewrite qualifE normr_ge0. +rewrite normCK conj_Creal // normC2_Re_Im -expr2. +rewrite addrC -lerif_subLR subrr (sameP (Creal_ImP _) eqP) -sqrf_eq0 eq_sym. +by apply: lerif_eq; rewrite -realEsqr. +Qed. + +Lemma lerif_Re_Creal z : 'Re z <= `|z| ?= iff (0 <= z). +Proof. +have ubRe: 'Re z <= `|'Re z| ?= iff (0 <= 'Re z). + by rewrite ger0_def eq_sym; apply/lerif_eq/real_ler_norm. +congr (_ <= _ ?= iff _): (lerif_trans ubRe (lerif_normC_Re_Creal z)). +apply/andP/idP=> [[zRge0 /Creal_ReP <- //] | z_ge0]. +by have Rz := ger0_real z_ge0; rewrite (Creal_ReP _ _). +Qed. + +(* Equality from polar coordinates, for the upper plane. *) +Lemma eqC_semipolar x y : + `|x| = `|y| -> 'Re x = 'Re y -> 0 <= 'Im x * 'Im y -> x = y. +Proof. +move=> eq_norm eq_Re sign_Im. +rewrite [x]algCrect [y]algCrect eq_Re; congr (_ + 'i * _). +have /eqP := congr1 (fun z => z ^+ 2) eq_norm. +rewrite !normC2_Re_Im eq_Re (can_eq (addKr _)) eqf_sqr => /pred2P[] // eq_Im. +rewrite eq_Im mulNr -expr2 oppr_ge0 real_exprn_even_le0 //= in sign_Im. +by rewrite eq_Im (eqP sign_Im) oppr0. +Qed. + +(* Nth roots. *) + +Let argCleP y z : + reflect (0 <= 'Im z -> 0 <= 'Im y /\ 'Re z <= 'Re y) (argCle y z). +Proof. +suffices dIm x: nnegIm x = (0 <= 'Im x). + rewrite /argCle !dIm ler_pmul2r ?invr_gt0 ?ltr0n //. + by apply: (iffP implyP) => geZyz /geZyz/andP. +rewrite /('Im x) pmulr_lge0 ?invr_gt0 ?ltr0n //; congr (0 <= _ * _). +case Du: algCi_subproof => [u u2N1] /=. +have/eqP := u2N1; rewrite -sqrCi eqf_sqr => /pred2P[] //. +have:= conjCi; rewrite /'i; case_rootC => /= v v2n1 min_v conj_v Duv. +have{min_v} /idPn[] := min_v u isT u2N1; rewrite negb_imply /nnegIm Du /= Duv. +rewrite rmorphN conj_v opprK -opprD mulrNN mulNr -mulr2n mulrnAr -expr2 v2n1. +by rewrite mulNrn opprK ler0n oppr_ge0 (leC_nat 2 0). +Qed. + +Lemma rootC_Re_max n x y : + (n > 0)%N -> y ^+ n = x -> 0 <= 'Im y -> 'Re y <= 'Re (n.-root%C x). +Proof. +by move=> n_gt0 yn_x leI0y; case_rootC=> z /= _ /(_ y n_gt0 yn_x)/argCleP[]. +Qed. + +Let neg_unity_root n : (n > 1)%N -> exists2 w : algC, w ^+ n = 1 & 'Re w < 0. +Proof. +move=> n_gt1; have [|w /eqP pw_0] := closed_rootP (\poly_(i < n) (1 : algC)) _. + by rewrite size_poly_eq ?oner_eq0 // -(subnKC n_gt1). +rewrite horner_poly (eq_bigr _ (fun _ _ => mul1r _)) in pw_0. +have wn1: w ^+ n = 1 by apply/eqP; rewrite -subr_eq0 subrX1 pw_0 mulr0. +suffices /existsP[i ltRwi0]: [exists i : 'I_n, 'Re (w ^+ i) < 0]. + by exists (w ^+ i) => //; rewrite exprAC wn1 expr1n. +apply: contra_eqT (congr1 algRe pw_0); rewrite negb_exists => /forallP geRw0. +rewrite raddf_sum raddf0 /= (bigD1 (Ordinal (ltnW n_gt1))) //=. +rewrite (Creal_ReP _ _) ?rpred1 // gtr_eqF ?ltr_paddr ?ltr01 //=. +by apply: sumr_ge0 => i _; rewrite real_lerNgt. +Qed. + +Lemma Im_rootC_ge0 n x : (n > 1)%N -> 0 <= 'Im (n.-root x). +Proof. +set y := n.-root x => n_gt1; have n_gt0 := ltnW n_gt1. +apply: wlog_neg; rewrite -real_ltrNge // => ltIy0. +suffices [z zn_x leI0z]: exists2 z, z ^+ n = x & 'Im z >= 0. + by rewrite /y; case_rootC => /= y1 _ /(_ z n_gt0 zn_x)/argCleP[]. +have [w wn1 ltRw0] := neg_unity_root n_gt1. +wlog leRI0yw: w wn1 ltRw0 / 0 <= 'Re y * 'Im w. + move=> IHw; have: 'Re y * 'Im w \is Creal by rewrite rpredM. + case/real_ger0P=> [|/ltrW leRIyw0]; first exact: IHw. + apply: (IHw w^*); rewrite ?algRe_conj ?algIm_conj ?mulrN ?oppr_ge0 //. + by rewrite -rmorphX wn1 rmorph1. +exists (w * y); first by rewrite exprMn wn1 mul1r rootCK. +rewrite [w]algCrect [y]algCrect mulC_rect. +by rewrite algIm_rect ?rpredD ?rpredN 1?rpredM // addr_ge0 // ltrW ?nmulr_rgt0. +Qed. + +Lemma rootC_lt0 n x : (1 < n)%N -> (n.-root x < 0) = false. +Proof. +set y := n.-root x => n_gt1; have n_gt0 := ltnW n_gt1. +apply: negbTE; apply: wlog_neg => /negbNE lt0y; rewrite ler_gtF //. +have Rx: x \is Creal by rewrite -[x](rootCK n_gt0) rpredX // ltr0_real. +have Re_y: 'Re y = y by apply/Creal_ReP; rewrite ltr0_real. +have [z zn_x leR0z]: exists2 z, z ^+ n = x & 'Re z >= 0. + have [w wn1 ltRw0] := neg_unity_root n_gt1. + exists (w * y); first by rewrite exprMn wn1 mul1r rootCK. + by rewrite algReMr ?ltr0_real // ltrW // nmulr_lgt0. +without loss leI0z: z zn_x leR0z / 'Im z >= 0. + move=> IHz; have: 'Im z \is Creal by []. + case/real_ger0P=> [|/ltrW leIz0]; first exact: IHz. + apply: (IHz z^*); rewrite ?algRe_conj ?algIm_conj ?oppr_ge0 //. + by rewrite -rmorphX zn_x conj_Creal. +by apply: ler_trans leR0z _; rewrite -Re_y ?rootC_Re_max ?ltr0_real. +Qed. + +Lemma rootC_ge0 n x : (n > 0)%N -> (0 <= n.-root x) = (0 <= x). +Proof. +set y := n.-root x => n_gt0. +apply/idP/idP=> [/(exprn_ge0 n) | x_ge0]; first by rewrite rootCK. +rewrite -(ger_lerif (lerif_Re_Creal y)). +have Ray: `|y| \is Creal by apply: normr_real. +rewrite -(Creal_ReP _ Ray) rootC_Re_max ?(Creal_ImP _ Ray) //. +by rewrite -normrX rootCK // ger0_norm. +Qed. + +Lemma rootC_gt0 n x : (n > 0)%N -> (n.-root x > 0) = (x > 0). +Proof. by move=> n_gt0; rewrite !lt0r rootC_ge0 ?rootC_eq0. Qed. + +Lemma rootC_le0 n x : (1 < n)%N -> (n.-root x <= 0) = (x == 0). +Proof. +by move=> n_gt1; rewrite ler_eqVlt rootC_lt0 // orbF rootC_eq0 1?ltnW. +Qed. + +Lemma ler_rootCl n : (n > 0)%N -> {in Num.nneg, {mono n.-root : x y / x <= y}}. +Proof. +move=> n_gt0 x x_ge0 y; have [y_ge0 | not_y_ge0] := boolP (0 <= y). + by rewrite -(ler_pexpn2r n_gt0) ?qualifE ?rootC_ge0 ?rootCK. +rewrite (contraNF (@ler_trans _ _ 0 _ _)) ?rootC_ge0 //. +by rewrite (contraNF (ler_trans x_ge0)). +Qed. + +Lemma ler_rootC n : (n > 0)%N -> {in Num.nneg &, {mono n.-root : x y / x <= y}}. +Proof. by move=> n_gt0 x y x_ge0 _; apply: ler_rootCl. Qed. + +Lemma ltr_rootCl n : (n > 0)%N -> {in Num.nneg, {mono n.-root : x y / x < y}}. +Proof. by move=> n_gt0 x x_ge0 y; rewrite !ltr_def ler_rootCl ?eqr_rootC. Qed. + +Lemma ltr_rootC n : (n > 0)%N -> {in Num.nneg &, {mono n.-root : x y / x < y}}. +Proof. by move/ler_rootC/lerW_mono_in. Qed. + +Lemma exprCK n x : (0 < n)%N -> 0 <= x -> n.-root (x ^+ n) = x. +Proof. +move=> n_gt0 x_ge0; apply/eqP. +by rewrite -(eqr_expn2 n_gt0) ?rootC_ge0 ?exprn_ge0 ?rootCK. +Qed. + +Lemma norm_rootC n x : `|n.-root x| = n.-root `|x|. +Proof. +have [-> | n_gt0] := posnP n; first by rewrite !root0C normr0. +apply/eqP; rewrite -(eqr_expn2 n_gt0) ?rootC_ge0 ?normr_ge0 //. +by rewrite -normrX !rootCK. +Qed. + +Lemma rootCX n x k : (n > 0)%N -> 0 <= x -> n.-root (x ^+ k) = n.-root x ^+ k. +Proof. +move=> n_gt0 x_ge0; apply/eqP. +by rewrite -(eqr_expn2 n_gt0) ?(exprn_ge0, rootC_ge0) // 1?exprAC !rootCK. +Qed. + +Lemma rootC1 n : (n > 0)%N -> n.-root 1 = 1. +Proof. by move/(rootCX 0)/(_ ler01). Qed. + +Lemma rootCpX n x k : (k > 0)%N -> 0 <= x -> n.-root (x ^+ k) = n.-root x ^+ k. +Proof. +by case: n => [|n] k_gt0; [rewrite !root0C expr0n gtn_eqF | exact: rootCX]. +Qed. + +Lemma rootCV n x : (n > 0)%N -> 0 <= x -> n.-root x^-1 = (n.-root x)^-1. +Proof. +move=> n_gt0 x_ge0; apply/eqP. +by rewrite -(eqr_expn2 n_gt0) ?(invr_ge0, rootC_ge0) // !exprVn !rootCK. +Qed. + +Lemma rootC_eq1 n x : (n > 0)%N -> (n.-root x == 1) = (x == 1). +Proof. by move=> n_gt0; rewrite -{1}(rootC1 n_gt0) eqr_rootC. Qed. + +Lemma rootC_ge1 n x : (n > 0)%N -> (n.-root x >= 1) = (x >= 1). +Proof. +by move=> n_gt0; rewrite -{1}(rootC1 n_gt0) ler_rootCl // qualifE ler01. +Qed. + +Lemma rootC_gt1 n x : (n > 0)%N -> (n.-root x > 1) = (x > 1). +Proof. by move=> n_gt0; rewrite !ltr_def rootC_eq1 ?rootC_ge1. Qed. + +Lemma rootC_le1 n x : (n > 0)%N -> 0 <= x -> (n.-root x <= 1) = (x <= 1). +Proof. by move=> n_gt0 x_ge0; rewrite -{1}(rootC1 n_gt0) ler_rootCl. Qed. + +Lemma rootC_lt1 n x : (n > 0)%N -> 0 <= x -> (n.-root x < 1) = (x < 1). +Proof. by move=> n_gt0 x_ge0; rewrite !ltr_neqAle rootC_eq1 ?rootC_le1. Qed. + +Lemma rootCMl n x z : 0 <= x -> n.-root (x * z) = n.-root x * n.-root z. +Proof. +rewrite le0r => /predU1P[-> | x_gt0]; first by rewrite !(mul0r, rootC0). +have [| n_gt1 | ->] := ltngtP n 1; last by rewrite !root1C. + by case: n => //; rewrite !root0C mul0r. +have [x_ge0 n_gt0] := (ltrW x_gt0, ltnW n_gt1). +have nx_gt0: 0 < n.-root x by rewrite rootC_gt0. +have Rnx: n.-root x \is Creal by rewrite ger0_real ?ltrW. +apply: eqC_semipolar; last 1 first; try apply/eqP. +- by rewrite algImMl // !(Im_rootC_ge0, mulr_ge0, rootC_ge0). +- by rewrite -(eqr_expn2 n_gt0) ?normr_ge0 // -!normrX exprMn !rootCK. +rewrite eqr_le; apply/andP; split; last first. + rewrite rootC_Re_max ?exprMn ?rootCK ?algImMl //. + by rewrite mulr_ge0 ?Im_rootC_ge0 ?ltrW. +rewrite -[n.-root _](mulVKf (negbT (gtr_eqF nx_gt0))) !(algReMl Rnx) //. +rewrite ler_pmul2l // rootC_Re_max ?exprMn ?exprVn ?rootCK ?mulKf ?gtr_eqF //. +by rewrite algImMl ?rpredV // mulr_ge0 ?invr_ge0 ?Im_rootC_ge0 ?ltrW. +Qed. + +Lemma rootCMr n x z : 0 <= x -> n.-root (z * x) = n.-root z * n.-root x. +Proof. by move=> x_ge0; rewrite mulrC rootCMl // mulrC. Qed. + +(* More properties of n.-root will be established in cyclotomic.v. *) + +(* The proper form of the Arithmetic - Geometric Mean inequality. *) + +Lemma lerif_rootC_AGM (I : finType) (A : pred I) (n := #|A|) E : + {in A, forall i, 0 <= E i} -> + n.-root (\prod_(i in A) E i) <= (\sum_(i in A) E i) / n%:R + ?= iff [forall i in A, forall j in A, E i == E j]. +Proof. +move=> Ege0; have [n0 | n_gt0] := posnP n. + rewrite n0 root0C invr0 mulr0; apply/lerif_refl/forall_inP=> i. + by rewrite (card0_eq n0). +rewrite -(mono_in_lerif (ler_pexpn2r n_gt0)) ?rootCK //=; first 1 last. +- by rewrite qualifE rootC_ge0 // prodr_ge0. +- by rewrite rpred_div ?rpred_nat ?rpred_sum. +exact: lerif_AGM. +Qed. + +(* Square root. *) + +Lemma sqrtC0 : sqrtC 0 = 0. Proof. exact: rootC0. Qed. +Lemma sqrtC1 : sqrtC 1 = 1. Proof. exact: rootC1. Qed. +Lemma sqrtCK x : sqrtC x ^+ 2 = x. Proof. exact: rootCK. Qed. +Lemma sqrCK x : 0 <= x -> sqrtC (x ^+ 2) = x. Proof. exact: exprCK. Qed. + +Lemma sqrtC_ge0 x : (0 <= sqrtC x) = (0 <= x). Proof. exact: rootC_ge0. Qed. +Lemma sqrtC_eq0 x : (sqrtC x == 0) = (x == 0). Proof. exact: rootC_eq0. Qed. +Lemma sqrtC_gt0 x : (sqrtC x > 0) = (x > 0). Proof. exact: rootC_gt0. Qed. +Lemma sqrtC_lt0 x : (sqrtC x < 0) = false. Proof. exact: rootC_lt0. Qed. +Lemma sqrtC_le0 x : (sqrtC x <= 0) = (x == 0). Proof. exact: rootC_le0. Qed. + +Lemma ler_sqrtC : {in Num.nneg &, {mono sqrtC : x y / x <= y}}. +Proof. exact: ler_rootC. Qed. +Lemma ltr_sqrtC : {in Num.nneg &, {mono sqrtC : x y / x < y}}. +Proof. exact: ltr_rootC. Qed. +Lemma eqr_sqrtC : {mono sqrtC : x y / x == y}. +Proof. exact: eqr_rootC. Qed. +Lemma sqrtC_inj : injective sqrtC. +Proof. exact: rootC_inj. Qed. +Lemma sqrtCM : {in Num.nneg &, {morph sqrtC : x y / x * y}}. +Proof. by move=> x y _; apply: rootCMr. Qed. + +Lemma sqrCK_P x : reflect (sqrtC (x ^+ 2) = x) ((0 <= 'Im x) && ~~ (x < 0)). +Proof. +apply: (iffP andP) => [[leI0x not_gt0x] | <-]; last first. + by rewrite sqrtC_lt0 Im_rootC_ge0. +have /eqP := sqrtCK (x ^+ 2); rewrite eqf_sqr => /pred2P[] // defNx. +apply: sqrCK; rewrite -real_lerNgt // in not_gt0x; apply/Creal_ImP/ler_anti; +by rewrite leI0x -oppr_ge0 -raddfN -defNx Im_rootC_ge0. +Qed. + +Lemma normC_def x : `|x| = sqrtC (x * x^*). +Proof. by rewrite -normCK sqrCK ?normr_ge0. Qed. + +Lemma norm_conjC x : `|x^*| = `|x|. +Proof. by rewrite !normC_def conjCK mulrC. Qed. + +Lemma normC_rect : + {in Creal &, forall x y, `|x + 'i * y| = sqrtC (x ^+ 2 + y ^+ 2)}. +Proof. by move=> x y Rx Ry; rewrite /= normC_def -normCK normC2_rect. Qed. + +Lemma normC_Re_Im z : `|z| = sqrtC ('Re z ^+ 2 + 'Im z ^+ 2). +Proof. by rewrite normC_def -normCK normC2_Re_Im. Qed. + +(* Norm sum (in)equalities. *) + +Lemma normC_add_eq x y : + `|x + y| = `|x| + `|y| -> + {t : algC | `|t| == 1 & (x, y) = (`|x| * t, `|y| * t)}. +Proof. +move=> lin_xy; apply: sig2_eqW; pose u z := if z == 0 then 1 else z / `|z|. +have uE z: (`|u z| = 1) * (`|z| * u z = z). + rewrite /u; have [->|nz_z] := altP eqP; first by rewrite normr0 normr1 mul0r. + by rewrite normf_div normr_id mulrCA divff ?mulr1 ?normr_eq0. +have [->|nz_x] := eqVneq x 0; first by exists (u y); rewrite uE ?normr0 ?mul0r. +exists (u x); rewrite uE // /u (negPf nz_x); congr (_ , _). +have{lin_xy} def2xy: `|x| * `|y| *+ 2 = x * y ^* + y * x ^*. + apply/(addrI (x * x^*))/(addIr (y * y^*)); rewrite -2!{1}normCK -sqrrD. + by rewrite addrA -addrA -!mulrDr -mulrDl -rmorphD -normCK lin_xy. +have def_xy: x * y^* = y * x^*. + apply/eqP; rewrite -subr_eq0 -[_ == 0](@expf_eq0 _ _ 2). + rewrite (canRL (subrK _) (subr_sqrDB _ _)) opprK -def2xy exprMn_n exprMn. + by rewrite mulrN mulrAC mulrA -mulrA mulrACA -!normCK mulNrn addNr. +have{def_xy def2xy} def_yx: `|y * x| = y * x^*. + by apply: (mulIf nz2); rewrite !mulr_natr mulrC normrM def2xy def_xy. +rewrite -{1}(divfK nz_x y) invC_norm mulrCA -{}def_yx !normrM invfM. +by rewrite mulrCA divfK ?normr_eq0 // mulrAC mulrA. +Qed. + +Lemma normC_sum_eq (I : finType) (P : pred I) (F : I -> algC) : + `|\sum_(i | P i) F i| = \sum_(i | P i) `|F i| -> + {t : algC | `|t| == 1 & forall i, P i -> F i = `|F i| * t}. +Proof. +have [i /andP[Pi nzFi] | F0] := pickP [pred i | P i & F i != 0]; last first. + exists 1 => [|i Pi]; first by rewrite normr1. + by case/nandP: (F0 i) => [/negP[]// | /negbNE/eqP->]; rewrite normr0 mul0r. +rewrite !(bigD1 i Pi) /= => norm_sumF; pose Q j := P j && (j != i). +rewrite -normr_eq0 in nzFi; set c := F i / `|F i|; exists c => [|j Pj]. + by rewrite normrM normfV normr_id divff. +have [Qj | /nandP[/negP[]// | /negbNE/eqP->]] := boolP (Q j); last first. + by rewrite mulrC divfK. +have: `|F i + F j| = `|F i| + `|F j|. + do [rewrite !(bigD1 j Qj) /=; set z := \sum_(k | _) `|_|] in norm_sumF. + apply/eqP; rewrite eqr_le ler_norm_add -(ler_add2r z) -addrA -norm_sumF addrA. + by rewrite (ler_trans (ler_norm_add _ _)) // ler_add2l ler_norm_sum. +by case/normC_add_eq=> k _ [/(canLR (mulKf nzFi)) <-]; rewrite -(mulrC (F i)). +Qed. + +Lemma normC_sum_eq1 (I : finType) (P : pred I) (F : I -> algC) : + `|\sum_(i | P i) F i| = (\sum_(i | P i) `|F i|) -> + (forall i, P i -> `|F i| = 1) -> + {t : algC | `|t| == 1 & forall i, P i -> F i = t}. +Proof. +case/normC_sum_eq=> t t1 defF normF. +by exists t => // i Pi; rewrite defF // normF // mul1r. +Qed. + +Lemma normC_sum_upper (I : finType) (P : pred I) (F G : I -> algC) : + (forall i, P i -> `|F i| <= G i) -> + \sum_(i | P i) F i = \sum_(i | P i) G i -> + forall i, P i -> F i = G i. +Proof. +set sumF := \sum_(i | _) _; set sumG := \sum_(i | _) _ => leFG eq_sumFG. +have posG i: P i -> 0 <= G i by move/leFG; apply: ler_trans; exact: normr_ge0. +have norm_sumG: `|sumG| = sumG by rewrite ger0_norm ?sumr_ge0. +have norm_sumF: `|sumF| = \sum_(i | P i) `|F i|. + apply/eqP; rewrite eqr_le ler_norm_sum eq_sumFG norm_sumG -subr_ge0 -sumrB. + by rewrite sumr_ge0 // => i Pi; rewrite subr_ge0 ?leFG. +have [t _ defF] := normC_sum_eq norm_sumF. +have [/(psumr_eq0P posG) G0 i Pi | nz_sumG] := eqVneq sumG 0. + by apply/eqP; rewrite G0 // -normr_eq0 eqr_le normr_ge0 -(G0 i Pi) leFG. +have t1: t = 1. + apply: (mulfI nz_sumG); rewrite mulr1 -{1}norm_sumG -eq_sumFG norm_sumF. + by rewrite mulr_suml -(eq_bigr _ defF). +have /psumr_eq0P eqFG i: P i -> 0 <= G i - F i. + by move=> Pi; rewrite subr_ge0 defF // t1 mulr1 leFG. +move=> i /eqFG/(canRL (subrK _))->; rewrite ?add0r //. +by rewrite sumrB -/sumF eq_sumFG subrr. +Qed. + +Lemma normC_sub_eq x y : + `|x - y| = `|x| - `|y| -> {t | `|t| == 1 & (x, y) = (`|x| * t, `|y| * t)}. +Proof. +rewrite -{-1}(subrK y x) => /(canLR (subrK _))/esym-Dx; rewrite Dx. +by have [t ? [Dxy Dy]] := normC_add_eq Dx; exists t; rewrite // mulrDl -Dxy -Dy. +Qed. + +(* Integer subset. *) + +(* Not relying on the undocumented interval library, for now. *) +Lemma floorC_itv x : x \is Creal -> (floorC x)%:~R <= x < (floorC x + 1)%:~R. +Proof. by rewrite /floorC => Rx; case: (floorC_subproof x) => //= m; apply. Qed. + +Lemma floorC_def x m : m%:~R <= x < (m + 1)%:~R -> floorC x = m. +Proof. +case/andP=> lemx ltxm1; apply/eqP; rewrite eqr_le -!ltz_addr1. +have /floorC_itv/andP[lefx ltxf1]: x \is Creal. + by rewrite -[x](subrK m%:~R) rpredD ?realz ?ler_sub_real. +by rewrite -!(ltr_int [numFieldType of algC]) 2?(@ler_lt_trans _ x). +Qed. + +Lemma intCK : cancel intr floorC. +Proof. +by move=> m; apply: floorC_def; rewrite ler_int ltr_int ltz_addr1 lerr. +Qed. + +Lemma floorCK : {in Cint, cancel floorC intr}. Proof. by move=> z /eqP. Qed. + +Lemma floorC0 : floorC 0 = 0. Proof. exact: (intCK 0). Qed. +Lemma floorC1 : floorC 1 = 1. Proof. exact: (intCK 1). Qed. +Hint Resolve floorC0 floorC1. + +Lemma floorCpK (p : {poly algC}) : + p \is a polyOver Cint -> map_poly intr (map_poly floorC p) = p. +Proof. +move/(all_nthP 0)=> Zp; apply/polyP=> i. +rewrite coef_map coef_map_id0 //= -[p]coefK coef_poly. +by case: ifP => [/Zp/floorCK // | _]; rewrite floorC0. +Qed. + +Lemma floorCpP (p : {poly algC}) : + p \is a polyOver Cint -> {q | p = map_poly intr q}. +Proof. by exists (map_poly floorC p); rewrite floorCpK. Qed. + +Lemma Cint_int m : m%:~R \in Cint. +Proof. by rewrite unfold_in intCK. Qed. + +Lemma CintP x : reflect (exists m, x = m%:~R) (x \in Cint). +Proof. +by apply: (iffP idP) => [/eqP<-|[m ->]]; [exists (floorC x) | apply: Cint_int]. +Qed. + +Lemma floorCD : {in Cint & Creal, {morph floorC : x y / x + y}}. +Proof. +move=> _ y /CintP[m ->] Ry; apply: floorC_def. +by rewrite -addrA 2!rmorphD /= intCK ler_add2l ltr_add2l floorC_itv. +Qed. + +Lemma floorCN : {in Cint, {morph floorC : x / - x}}. +Proof. by move=> _ /CintP[m ->]; rewrite -rmorphN !intCK. Qed. + +Lemma floorCM : {in Cint &, {morph floorC : x y / x * y}}. +Proof. by move=> _ _ /CintP[m1 ->] /CintP[m2 ->]; rewrite -rmorphM !intCK. Qed. + +Lemma floorCX n : {in Cint, {morph floorC : x / x ^+ n}}. +Proof. by move=> _ /CintP[m ->]; rewrite -rmorphX !intCK. Qed. + +Lemma rpred_Cint S (ringS : subringPred S) (kS : keyed_pred ringS) x : + x \in Cint -> x \in kS. +Proof. by case/CintP=> m ->; apply: rpred_int. Qed. + +Lemma Cint0 : 0 \in Cint. Proof. exact: (Cint_int 0). Qed. +Lemma Cint1 : 1 \in Cint. Proof. exact: (Cint_int 1). Qed. +Hint Resolve Cint0 Cint1. + +Fact Cint_key : pred_key Cint. Proof. by []. Qed. +Fact Cint_subring : subring_closed Cint. +Proof. +by split=> // _ _ /CintP[m ->] /CintP[p ->]; + rewrite -(rmorphB, rmorphM) Cint_int. +Qed. +Canonical Cint_keyed := KeyedPred Cint_key. +Canonical Cint_opprPred := OpprPred Cint_subring. +Canonical Cint_addrPred := AddrPred Cint_subring. +Canonical Cint_mulrPred := MulrPred Cint_subring. +Canonical Cint_zmodPred := ZmodPred Cint_subring. +Canonical Cint_semiringPred := SemiringPred Cint_subring. +Canonical Cint_smulrPred := SmulrPred Cint_subring. +Canonical Cint_subringPred := SubringPred Cint_subring. + +Lemma Creal_Cint : {subset Cint <= Creal}. +Proof. by move=> _ /CintP[m ->]; apply: realz. Qed. + +Lemma conj_Cint x : x \in Cint -> x^* = x. +Proof. by move/Creal_Cint/conj_Creal. Qed. + +Lemma Cint_normK x : x \in Cint -> `|x| ^+ 2 = x ^+ 2. +Proof. by move/Creal_Cint/real_normK. Qed. + +Lemma CintEsign x : x \in Cint -> x = (-1) ^+ (x < 0)%C * `|x|. +Proof. by move/Creal_Cint/realEsign. Qed. + +(* Natural integer subset. *) + +Lemma truncC_itv x : 0 <= x -> (truncC x)%:R <= x < (truncC x).+1%:R. +Proof. +move=> x_ge0; have /andP[lemx ltxm1] := floorC_itv (ger0_real x_ge0). +rewrite /truncC x_ge0 -addn1 !pmulrn PoszD gez0_abs ?lemx //. +by rewrite -ltz_addr1 -(ltr_int [numFieldType of algC]) (ler_lt_trans x_ge0). +Qed. + +Lemma truncC_def x n : n%:R <= x < n.+1%:R -> truncC x = n. +Proof. +move=> ivt_n_x; have /andP[lenx _] := ivt_n_x. +by rewrite /truncC (ler_trans (ler0n _ n)) // (@floorC_def _ n) // addrC -intS. +Qed. + +Lemma natCK n : truncC n%:R = n. +Proof. by apply: truncC_def; rewrite lerr ltr_nat /=. Qed. + +Lemma CnatP x : reflect (exists n, x = n%:R) (x \in Cnat). +Proof. +by apply: (iffP eqP) => [<- | [n ->]]; [exists (truncC x) | rewrite natCK]. +Qed. + +Lemma truncCK : {in Cnat, cancel truncC (GRing.natmul 1)}. +Proof. by move=> x /eqP. Qed. + +Lemma truncC_gt0 x : (0 < truncC x)%N = (1 <= x). +Proof. +apply/idP/idP=> [m_gt0 | x_ge1]. + have /truncC_itv/andP[lemx _]: 0 <= x. + by move: m_gt0; rewrite /truncC; case: ifP. + by apply: ler_trans lemx; rewrite ler1n. +have /truncC_itv/andP[_ ltxm1]:= ler_trans ler01 x_ge1. +by rewrite -ltnS -ltC_nat (ler_lt_trans x_ge1). +Qed. + +Lemma truncC0Pn x : reflect (truncC x = 0%N) (~~ (1 <= x)). +Proof. by rewrite -truncC_gt0 -eqn0Ngt; apply: eqP. Qed. + +Lemma truncC0 : truncC 0 = 0%N. Proof. exact: (natCK 0). Qed. +Lemma truncC1 : truncC 1 = 1%N. Proof. exact: (natCK 1). Qed. + +Lemma truncCD : + {in Cnat & Num.nneg, {morph truncC : x y / x + y >-> (x + y)%N}}. +Proof. +move=> _ y /CnatP[n ->] y_ge0; apply: truncC_def. +by rewrite -addnS !natrD !natCK ler_add2l ltr_add2l truncC_itv. +Qed. + +Lemma truncCM : {in Cnat &, {morph truncC : x y / x * y >-> (x * y)%N}}. +Proof. by move=> _ _ /CnatP[n1 ->] /CnatP[n2 ->]; rewrite -natrM !natCK. Qed. + +Lemma truncCX n : {in Cnat, {morph truncC : x / x ^+ n >-> (x ^ n)%N}}. +Proof. by move=> _ /CnatP[n1 ->]; rewrite -natrX !natCK. Qed. + +Lemma rpred_Cnat S (ringS : semiringPred S) (kS : keyed_pred ringS) x : + x \in Cnat -> x \in kS. +Proof. by case/CnatP=> n ->; apply: rpred_nat. Qed. + +Lemma Cnat_nat n : n%:R \in Cnat. Proof. by apply/CnatP; exists n. Qed. +Lemma Cnat0 : 0 \in Cnat. Proof. exact: (Cnat_nat 0). Qed. +Lemma Cnat1 : 1 \in Cnat. Proof. exact: (Cnat_nat 1). Qed. +Hint Resolve Cnat_nat Cnat0 Cnat1. + +Fact Cnat_key : pred_key Cnat. Proof. by []. Qed. +Fact Cnat_semiring : semiring_closed Cnat. +Proof. +by do 2![split] => //= _ _ /CnatP[n ->] /CnatP[m ->]; rewrite -(natrD, natrM). +Qed. +Canonical Cnat_keyed := KeyedPred Cnat_key. +Canonical Cnat_addrPred := AddrPred Cnat_semiring. +Canonical Cnat_mulrPred := MulrPred Cnat_semiring. +Canonical Cnat_semiringPred := SemiringPred Cnat_semiring. + +Lemma Cnat_ge0 x : x \in Cnat -> 0 <= x. +Proof. by case/CnatP=> n ->; apply: ler0n. Qed. + +Lemma Cnat_gt0 x : x \in Cnat -> (0 < x) = (x != 0). +Proof. by case/CnatP=> n ->; rewrite pnatr_eq0 ltr0n lt0n. Qed. + +Lemma conj_Cnat x : x \in Cnat -> x^* = x. +Proof. by case/CnatP=> n ->; apply: rmorph_nat. Qed. + +Lemma norm_Cnat x : x \in Cnat -> `|x| = x. +Proof. by move/Cnat_ge0/ger0_norm. Qed. + +Lemma Creal_Cnat : {subset Cnat <= Creal}. +Proof. by move=> z /conj_Cnat/CrealP. Qed. + +Lemma Cnat_sum_eq1 (I : finType) (P : pred I) (F : I -> algC) : + (forall i, P i -> F i \in Cnat) -> \sum_(i | P i) F i = 1 -> + {i : I | [/\ P i, F i = 1 & forall j, j != i -> P j -> F j = 0]}. +Proof. +move=> natF sumF1; pose nF i := truncC (F i). +have{natF} defF i: P i -> F i = (nF i)%:R by move/natF/eqP. +have{sumF1} /eqP sumF1: (\sum_(i | P i) nF i == 1)%N. + by rewrite -eqC_nat natr_sum -(eq_bigr _ defF) sumF1. +have [i Pi nZfi]: {i : I | P i & nF i != 0%N}. + by apply/sig2W/exists_inP; rewrite -negb_forall_in -sum_nat_eq0 sumF1. +have F'ge0 := (leq0n _, etrans (eq_sym _ _) (sum_nat_eq0 (predD1 P i) nF)). +rewrite -lt0n in nZfi; have [_] := (leqif_add (leqif_eq nZfi) (F'ge0 _)). +rewrite /= big_andbC -bigD1 // sumF1 => /esym/andP/=[/eqP Fi1 /forall_inP Fi'0]. +exists i; split=> // [|j neq_ji Pj]; first by rewrite defF // -Fi1. +by rewrite defF // (eqP (Fi'0 j _)) // neq_ji. +Qed. + +Lemma Cnat_mul_eq1 x y : + x \in Cnat -> y \in Cnat -> (x * y == 1) = (x == 1) && (y == 1). +Proof. by do 2!move/truncCK <-; rewrite -natrM !pnatr_eq1 muln_eq1. Qed. + +Lemma Cnat_prod_eq1 (I : finType) (P : pred I) (F : I -> algC) : + (forall i, P i -> F i \in Cnat) -> \prod_(i | P i) F i = 1 -> + forall i, P i -> F i = 1. +Proof. +move=> natF prodF1; apply/eqfun_inP; rewrite -big_andE. +move: prodF1; elim/(big_load (fun x => x \in Cnat)): _. +elim/big_rec2: _ => // i all1x x /natF N_Fi [Nx x1all1]. +by split=> [|/eqP]; rewrite ?rpredM ?Cnat_mul_eq1 // => /andP[-> /eqP]. +Qed. + +(* Relating Cint and Cnat. *) + +Lemma Cint_Cnat : {subset Cnat <= Cint}. +Proof. by move=> _ /CnatP[n ->]; rewrite pmulrn Cint_int. Qed. + +Lemma CintE x : (x \in Cint) = (x \in Cnat) || (- x \in Cnat). +Proof. +apply/idP/idP=> [/CintP[[n | n] ->] | ]; first by rewrite Cnat_nat. + by rewrite NegzE opprK Cnat_nat orbT. +by case/pred2P=> [<- | /(canLR (@opprK _)) <-]; rewrite ?rpredN rpred_nat. +Qed. + +Lemma Cnat_norm_Cint x : x \in Cint -> `|x| \in Cnat. +Proof. +case/CintP=> [m ->]; rewrite [m]intEsign rmorphM rmorph_sign. +by rewrite normrM normr_sign mul1r normr_nat rpred_nat. +Qed. + +Lemma CnatEint x : (x \in Cnat) = (x \in Cint) && (0 <= x). +Proof. +apply/idP/andP=> [Nx | [Zx x_ge0]]; first by rewrite Cint_Cnat ?Cnat_ge0. +by rewrite -(ger0_norm x_ge0) Cnat_norm_Cint. +Qed. + +Lemma CintEge0 x : 0 <= x -> (x \in Cint) = (x \in Cnat). +Proof. by rewrite CnatEint andbC => ->. Qed. + +Lemma Cnat_exp_even x n : ~~ odd n -> x \in Cint -> x ^+ n \in Cnat. +Proof. +rewrite -dvdn2 => /dvdnP[m ->] Zx; rewrite mulnC exprM -Cint_normK ?rpredX //. +exact: Cnat_norm_Cint. +Qed. + +Lemma norm_Cint_ge1 x : x \in Cint -> x != 0 -> 1 <= `|x|. +Proof. +rewrite -normr_eq0 => /Cnat_norm_Cint/CnatP[n ->]. +by rewrite pnatr_eq0 ler1n lt0n. +Qed. + +Lemma sqr_Cint_ge1 x : x \in Cint -> x != 0 -> 1 <= x ^+ 2. +Proof. +by move=> Zx nz_x; rewrite -Cint_normK // expr_ge1 ?normr_ge0 ?norm_Cint_ge1. +Qed. + +Lemma Cint_ler_sqr x : x \in Cint -> x <= x ^+ 2. +Proof. +move=> Zx; have [-> | nz_x] := eqVneq x 0; first by rewrite expr0n. +apply: ler_trans (_ : `|x| <= _); first by rewrite real_ler_norm ?Creal_Cint. +by rewrite -Cint_normK // ler_eexpr // norm_Cint_ge1. +Qed. + +(* Integer divisibility. *) + +Lemma dvdCP x y : reflect (exists2 z, z \in Cint & y = z * x) (x %| y)%C. +Proof. +rewrite unfold_in; have [-> | nz_x] := altP eqP. + by apply: (iffP eqP) => [-> | [z _ ->]]; first exists 0; rewrite ?mulr0. +apply: (iffP idP) => [Zyx | [z Zz ->]]; last by rewrite mulfK. +by exists (y / x); rewrite ?divfK. +Qed. + +Lemma dvdCP_nat x y : 0 <= x -> 0 <= y -> (x %| y)%C -> {n | y = n%:R * x}. +Proof. +move=> x_ge0 y_ge0 x_dv_y; apply: sig_eqW. +case/dvdCP: x_dv_y => z Zz -> in y_ge0 *; move: x_ge0 y_ge0 Zz. +rewrite ler_eqVlt => /predU1P[<- | ]; first by exists 22; rewrite !mulr0. +by move=> /pmulr_lge0-> /CintEge0-> /CnatP[n ->]; exists n. +Qed. + +Lemma dvdC0 x : (x %| 0)%C. +Proof. by apply/dvdCP; exists 0; rewrite ?mul0r. Qed. + +Lemma dvd0C x : (0 %| x)%C = (x == 0). +Proof. by rewrite unfold_in eqxx. Qed. + +Lemma dvdC_mull x y z : y \in Cint -> (x %| z)%C -> (x %| y * z)%C. +Proof. +move=> Zy /dvdCP[m Zm ->]; apply/dvdCP. +by exists (y * m); rewrite ?mulrA ?rpredM. +Qed. + +Lemma dvdC_mulr x y z : y \in Cint -> (x %| z)%C -> (x %| z * y)%C. +Proof. by rewrite mulrC; apply: dvdC_mull. Qed. + +Lemma dvdC_mul2r x y z : y != 0 -> (x * y %| z * y)%C = (x %| z)%C. +Proof. +move=> nz_y; rewrite !unfold_in !(mulIr_eq0 _ (mulIf nz_y)). +by rewrite mulrAC invfM mulrA divfK. +Qed. + +Lemma dvdC_mul2l x y z : y != 0 -> (y * x %| y * z)%C = (x %| z)%C. +Proof. by rewrite !(mulrC y); apply: dvdC_mul2r. Qed. + +Lemma dvdC_trans x y z : (x %| y)%C -> (y %| z)%C -> (x %| z)%C. +Proof. by move=> x_dv_y /dvdCP[m Zm ->]; apply: dvdC_mull. Qed. + +Lemma dvdC_refl x : (x %| x)%C. +Proof. by apply/dvdCP; exists 1; rewrite ?mul1r. Qed. +Hint Resolve dvdC_refl. + +Fact dvdC_key x : pred_key (dvdC x). Proof. by []. Qed. +Lemma dvdC_zmod x : zmod_closed (dvdC x). +Proof. +split=> [| _ _ /dvdCP[y Zy ->] /dvdCP[z Zz ->]]; first exact: dvdC0. +by rewrite -mulrBl dvdC_mull ?rpredB. +Qed. +Canonical dvdC_keyed x := KeyedPred (dvdC_key x). +Canonical dvdC_opprPred x := OpprPred (dvdC_zmod x). +Canonical dvdC_addrPred x := AddrPred (dvdC_zmod x). +Canonical dvdC_zmodPred x := ZmodPred (dvdC_zmod x). + +Lemma dvdC_nat (p n : nat) : (p %| n)%C = (p %| n)%N. +Proof. +rewrite unfold_in CintEge0 ?divr_ge0 ?invr_ge0 ?ler0n // !pnatr_eq0. +have [-> | nz_p] := altP eqP; first by rewrite dvd0n. +apply/CnatP/dvdnP=> [[q def_q] | [q ->]]; exists q. + by apply/eqP; rewrite -eqC_nat natrM -def_q divfK ?pnatr_eq0. +by rewrite [num in num / _]natrM mulfK ?pnatr_eq0. +Qed. + +Lemma dvdC_int (p : nat) x : x \in Cint -> (p %| x)%C = (p %| `|floorC x|)%N. +Proof. +move=> Zx; rewrite -{1}(floorCK Zx) {1}[floorC x]intEsign. +by rewrite rmorphMsign rpredMsign dvdC_nat. +Qed. + +(* Elementary modular arithmetic. *) + +Lemma eqCmod_refl e x : (x == x %[mod e])%C. +Proof. by rewrite /eqCmod subrr rpred0. Qed. + +Lemma eqCmodm0 e : (e == 0 %[mod e])%C. Proof. by rewrite /eqCmod subr0. Qed. +Hint Resolve eqCmod_refl eqCmodm0. + +Lemma eqCmod0 e x : (x == 0 %[mod e])%C = (e %| x)%C. +Proof. by rewrite /eqCmod subr0. Qed. + +Lemma eqCmod_sym e x y : ((x == y %[mod e]) = (y == x %[mod e]))%C. +Proof. by rewrite /eqCmod -opprB rpredN. Qed. + +Lemma eqCmod_trans e y x z : + (x == y %[mod e] -> y == z %[mod e] -> x == z %[mod e])%C. +Proof. by move=> Exy Eyz; rewrite /eqCmod -[x](subrK y) -addrA rpredD. Qed. + +Lemma eqCmod_transl e x y z : + (x == y %[mod e])%C -> (x == z %[mod e])%C = (y == z %[mod e])%C. +Proof. by move/(sym_left_transitive (eqCmod_sym e) (@eqCmod_trans e)). Qed. + +Lemma eqCmod_transr e x y z : + (x == y %[mod e])%C -> (z == x %[mod e])%C = (z == y %[mod e])%C. +Proof. by move/(sym_right_transitive (eqCmod_sym e) (@eqCmod_trans e)). Qed. + +Lemma eqCmodN e x y : (- x == y %[mod e])%C = (x == - y %[mod e])%C. +Proof. by rewrite eqCmod_sym /eqCmod !opprK addrC. Qed. + +Lemma eqCmodDr e x y z : (y + x == z + x %[mod e])%C = (y == z %[mod e])%C. +Proof. by rewrite /eqCmod addrAC opprD !addrA subrK. Qed. + +Lemma eqCmodDl e x y z : (x + y == x + z %[mod e])%C = (y == z %[mod e])%C. +Proof. by rewrite !(addrC x) eqCmodDr. Qed. + +Lemma eqCmodD e x1 x2 y1 y2 : + (x1 == x2 %[mod e] -> y1 == y2 %[mod e] -> x1 + y1 == x2 + y2 %[mod e])%C. +Proof. rewrite -(eqCmodDl e x2 y1) -(eqCmodDr e y1); exact: eqCmod_trans. Qed. + +Lemma eqCmod_nat (e m n : nat) : (m == n %[mod e])%C = (m == n %[mod e]). +Proof. +without loss lenm: m n / (n <= m)%N. + by move=> IH; case/orP: (leq_total m n) => /IH //; rewrite eqCmod_sym eq_sym. +by rewrite /eqCmod -natrB // dvdC_nat eqn_mod_dvd. +Qed. + +Lemma eqCmod0_nat (e m : nat) : (m == 0 %[mod e])%C = (e %| m)%N. +Proof. by rewrite eqCmod0 dvdC_nat. Qed. + +Lemma eqCmodMr e : + {in Cint, forall z x y, x == y %[mod e] -> x * z == y * z %[mod e]}%C. +Proof. by move=> z Zz x y; rewrite /eqCmod -mulrBl => /dvdC_mulr->. Qed. + +Lemma eqCmodMl e : + {in Cint, forall z x y, x == y %[mod e] -> z * x == z * y %[mod e]}%C. +Proof. by move=> z Zz x y Exy; rewrite !(mulrC z) eqCmodMr. Qed. + +Lemma eqCmodMl0 e : {in Cint, forall x, x * e == 0 %[mod e]}%C. +Proof. by move=> x Zx; rewrite -(mulr0 x) eqCmodMl. Qed. + +Lemma eqCmodMr0 e : {in Cint, forall x, e * x == 0 %[mod e]}%C. +Proof. by move=> x Zx; rewrite /= mulrC eqCmodMl0. Qed. + +Lemma eqCmod_addl_mul e : {in Cint, forall x y, x * e + y == y %[mod e]}%C. +Proof. by move=> x Zx y; rewrite -{2}[y]add0r eqCmodDr eqCmodMl0. Qed. + +Lemma eqCmodM e : {in Cint & Cint, forall x1 y2 x2 y1, + x1 == x2 %[mod e] -> y1 == y2 %[mod e] -> x1 * y1 == x2 * y2 %[mod e]}%C. +Proof. +move=> x1 y2 Zx1 Zy2 x2 y1 eq_x /(eqCmodMl Zx1)/eqCmod_trans-> //. +exact: eqCmodMr. +Qed. + +(* Rational number subset. *) + +Lemma ratCK : cancel QtoC CtoQ. +Proof. by rewrite /getCrat; case: getCrat_subproof. Qed. + +Lemma getCratK : {in Crat, cancel CtoQ QtoC}. +Proof. by move=> x /eqP. Qed. + +Lemma Crat_rat (a : rat) : QtoC a \in Crat. +Proof. by rewrite unfold_in ratCK. Qed. + +Lemma CratP x : reflect (exists a, x = QtoC a) (x \in Crat). +Proof. +by apply: (iffP eqP) => [<- | [a ->]]; [exists (CtoQ x) | rewrite ratCK]. +Qed. + +Lemma Crat0 : 0 \in Crat. Proof. by apply/CratP; exists 0; rewrite rmorph0. Qed. +Lemma Crat1 : 1 \in Crat. Proof. by apply/CratP; exists 1; rewrite rmorph1. Qed. +Hint Resolve Crat0 Crat1. + +Fact Crat_key : pred_key Crat. Proof. by []. Qed. +Fact Crat_divring_closed : divring_closed Crat. +Proof. +split=> // _ _ /CratP[x ->] /CratP[y ->]. + by rewrite -rmorphB Crat_rat. +by rewrite -fmorph_div Crat_rat. +Qed. +Canonical Crat_keyed := KeyedPred Crat_key. +Canonical Crat_opprPred := OpprPred Crat_divring_closed. +Canonical Crat_addrPred := AddrPred Crat_divring_closed. +Canonical Crat_mulrPred := MulrPred Crat_divring_closed. +Canonical Crat_zmodPred := ZmodPred Crat_divring_closed. +Canonical Crat_semiringPred := SemiringPred Crat_divring_closed. +Canonical Crat_smulrPred := SmulrPred Crat_divring_closed. +Canonical Crat_divrPred := DivrPred Crat_divring_closed. +Canonical Crat_subringPred := SubringPred Crat_divring_closed. +Canonical Crat_sdivrPred := SdivrPred Crat_divring_closed. +Canonical Crat_divringPred := DivringPred Crat_divring_closed. + +Lemma rpred_Crat S (ringS : divringPred S) (kS : keyed_pred ringS) : + {subset Crat <= kS}. +Proof. by move=> _ /CratP[a ->]; apply: rpred_rat. Qed. + +Lemma conj_Crat z : z \in Crat -> z^* = z. +Proof. by move/getCratK <-; rewrite fmorph_div !rmorph_int. Qed. + +Lemma Creal_Crat : {subset Crat <= Creal}. +Proof. by move=> x /conj_Crat/CrealP. Qed. + +Lemma Cint_rat a : (QtoC a \in Cint) = (a \in Qint). +Proof. +apply/idP/idP=> [Za | /numqK <-]; last by rewrite rmorph_int Cint_int. +apply/QintP; exists (floorC (QtoC a)); apply: (can_inj ratCK). +by rewrite rmorph_int floorCK. +Qed. + +Lemma minCpolyP x : + {p | minCpoly x = pQtoC p /\ p \is monic + & forall q, root (pQtoC q) x = (p %| q)%R}. +Proof. by rewrite /minCpoly; case: (minCpoly_subproof x) => p; exists p. Qed. + +Lemma minCpoly_monic x : minCpoly x \is monic. +Proof. by have [p [-> mon_p] _] := minCpolyP x; rewrite map_monic. Qed. + +Lemma minCpoly_eq0 x : (minCpoly x == 0) = false. +Proof. exact/negbTE/monic_neq0/minCpoly_monic. Qed. + +Lemma root_minCpoly x : root (minCpoly x) x. +Proof. by have [p [-> _] ->] := minCpolyP x. Qed. + +Lemma size_minCpoly x : (1 < size (minCpoly x))%N. +Proof. by apply: root_size_gt1 (root_minCpoly x); rewrite ?minCpoly_eq0. Qed. + +(* Basic properties of automorphisms. *) +Section AutC. + +Implicit Type nu : {rmorphism algC -> algC}. + +Lemma aut_Cnat nu : {in Cnat, nu =1 id}. +Proof. by move=> _ /CnatP[n ->]; apply: rmorph_nat. Qed. + +Lemma aut_Cint nu : {in Cint, nu =1 id}. +Proof. by move=> _ /CintP[m ->]; apply: rmorph_int. Qed. + +Lemma aut_Crat nu : {in Crat, nu =1 id}. +Proof. by move=> _ /CratP[a ->]; apply: fmorph_rat. Qed. + +Lemma Cnat_aut nu x : (nu x \in Cnat) = (x \in Cnat). +Proof. +by do [apply/idP/idP=> Nx; have:= aut_Cnat nu Nx] => [/fmorph_inj <- | ->]. +Qed. + +Lemma Cint_aut nu x : (nu x \in Cint) = (x \in Cint). +Proof. by rewrite !CintE -rmorphN !Cnat_aut. Qed. + +Lemma Crat_aut nu x : (nu x \in Crat) = (x \in Crat). +Proof. +apply/idP/idP=> /CratP[a] => [|->]; last by rewrite fmorph_rat Crat_rat. +by rewrite -(fmorph_rat nu) => /fmorph_inj->; apply: Crat_rat. +Qed. + +Lemma algC_invaut_subproof nu x : {y | nu y = x}. +Proof. +have [r Dp] := closed_field_poly_normal (minCpoly x). +suffices /mapP/sig2_eqW[y _ ->]: x \in map nu r by exists y. +rewrite -root_prod_XsubC; congr (root _ x): (root_minCpoly x). +have [q [Dq _] _] := minCpolyP x; rewrite Dq -(eq_map_poly (fmorph_rat nu)). +rewrite (map_poly_comp nu) -{q}Dq Dp (monicP (minCpoly_monic x)) scale1r. +rewrite rmorph_prod big_map; apply: eq_bigr => z _. +by rewrite rmorphB /= map_polyX map_polyC. +Qed. +Definition algC_invaut nu x := sval (algC_invaut_subproof nu x). + +Lemma algC_invautK nu : cancel (algC_invaut nu) nu. +Proof. by move=> x; rewrite /algC_invaut; case: algC_invaut_subproof. Qed. + +Lemma algC_autK nu : cancel nu (algC_invaut nu). +Proof. exact: inj_can_sym (algC_invautK nu) (fmorph_inj nu). Qed. + +Fact algC_invaut_is_rmorphism nu : rmorphism (algC_invaut nu). +Proof. exact: can2_rmorphism (algC_autK nu) (algC_invautK nu). Qed. +Canonical algC_invaut_additive nu := Additive (algC_invaut_is_rmorphism nu). +Canonical algC_invaut_rmorphism nu := RMorphism (algC_invaut_is_rmorphism nu). + +Lemma minCpoly_aut nu x : minCpoly (nu x) = minCpoly x. +Proof. +wlog suffices dvd_nu: nu x / (minCpoly x %| minCpoly (nu x))%R. + apply/eqP; rewrite -eqp_monic ?minCpoly_monic //; apply/andP; split=> //. + by rewrite -{2}(algC_autK nu x) dvd_nu. +have [[q [Dq _] min_q] [q1 [Dq1 _] _]] := (minCpolyP x, minCpolyP (nu x)). +rewrite Dq Dq1 dvdp_map -min_q -(fmorph_root nu) -map_poly_comp. +by rewrite (eq_map_poly (fmorph_rat nu)) -Dq1 root_minCpoly. +Qed. + +End AutC. + +Section AutLmodC. + +Variables (U V : lmodType algC) (f : {additive U -> V}). + +Lemma raddfZ_Cnat a u : a \in Cnat -> f (a *: u) = a *: f u. +Proof. by case/CnatP=> n ->; exact: raddfZnat. Qed. + +Lemma raddfZ_Cint a u : a \in Cint -> f (a *: u) = a *: f u. +Proof. by case/CintP=> m ->; rewrite !scaler_int raddfMz. Qed. + +End AutLmodC. + +Section PredCmod. + +Variable V : lmodType algC. + +Lemma rpredZ_Cnat S (addS : @addrPred V S) (kS : keyed_pred addS) : + {in Cnat & kS, forall z u, z *: u \in kS}. +Proof. by move=> _ u /CnatP[n ->]; apply: rpredZnat. Qed. + +Lemma rpredZ_Cint S (subS : @zmodPred V S) (kS : keyed_pred subS) : + {in Cint & kS, forall z u, z *: u \in kS}. +Proof. by move=> _ u /CintP[m ->]; apply: rpredZint. Qed. + +End PredCmod. + +End AlgebraicsTheory. +Hint Resolve Creal0 Creal1 Cnat_nat Cnat0 Cnat1 Cint0 Cint1 floorC0 Crat0 Crat1. +Hint Resolve dvdC0 dvdC_refl eqCmod_refl eqCmodm0. diff --git a/mathcomp/field/algebraics_fundamentals.v b/mathcomp/field/algebraics_fundamentals.v new file mode 100644 index 0000000..9dc672d --- /dev/null +++ b/mathcomp/field/algebraics_fundamentals.v @@ -0,0 +1,867 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun ssrnat eqtype seq choice div fintype. +Require Import path tuple bigop finset prime ssralg poly polydiv mxpoly. +Require Import countalg ssrnum ssrint rat intdiv. +Require Import fingroup finalg zmodp cyclic pgroup sylow. +Require Import vector falgebra fieldext separable galois. + +(******************************************************************************) +(* The main result in this file is the existence theorem that underpins the *) +(* construction of the algebraic numbers in file algC.v. This theorem simply *) +(* asserts the existence of an algebraically closed field with an *) +(* automorphism of order 2, and dubbed the Fundamental_Theorem_of_Algebraics *) +(* because it is essentially the Fundamental Theorem of Algebra for algebraic *) +(* numbers (the more familiar version for complex numbers can be derived by *) +(* continuity). *) +(* Although our proof does indeed construct exactly the algebraics, we *) +(* choose not to expose this in the statement of our Theorem. In algC.v we *) +(* construct the norm and partial order of the "complex field" introduced by *) +(* the Theorem; as these imply is has characteristic 0, we then get the *) +(* algebraics as a subfield. To avoid some duplication a few basic properties *) +(* of the algebraics, such as the existence of minimal polynomials, that are *) +(* required by the proof of the Theorem, are also proved here. *) +(* The main theorem of countalg.v supplies us directly with an algebraic *) +(* closure of the rationals (as the rationals are a countable field), so all *) +(* we really need to construct is a conjugation automorphism that exchanges *) +(* the two roots (i and -i) of X^2 + 1, and fixes a (real) subfield of *) +(* index 2. This does not require actually constructing this field: the *) +(* kHomExtend construction from galois.v supplies us with an automorphism *) +(* conj_n of the number field Q[z_n] = Q[x_n, i] for any x_n such that Q[x_n] *) +(* does not contain i (e.g., such that Q[x_n] is real). As conj_n will extend *) +(* conj_m when Q[x_n] contains x_m, it therefore suffices to construct a *) +(* sequence x_n such that *) +(* (1) For each n, Q[x_n] is a REAL field containing Q[x_m] for all m <= n. *) +(* (2) Each z in C belongs to Q[z_n] = Q[x_n, i] for large enough n. *) +(* This, of course, amounts to proving the Fundamental Theorem of Algebra. *) +(* Indeed, we use a constructive variant of Artin's algebraic proof of that *) +(* Theorem to replace (2) by *) +(* (3) Each monic polynomial over Q[x_m] whose constant term is -c^2 for some *) +(* c in Q[x_m] has a root in Q[x_n] for large enough n. *) +(* We then ensure (3) by setting Q[x_n+1] = Q[x_n, y] where y is the root of *) +(* of such a polynomial p found by dichotomy in some interval [0, b] with b *) +(* suitably large (such that p[b] >= 0), and p is obtained by decoding n into *) +(* a triple (m, p, c) that satisfies the conditions of (3) (taking x_n+1=x_n *) +(* if this is not the case), thereby ensuring that all such triples are *) +(* ultimately considered. *) +(* In more detail, the 600-line proof consists in six (uneven) parts: *) +(* (A) - Construction of number fields (~ 100 lines): in order to make use of *) +(* the theory developped in falgebra, fieldext, separable and galois we *) +(* construct a separate fielExtType Q z for the number field Q[z], with *) +(* z in C, the closure of rat supplied by countable_algebraic_closure. *) +(* The morphism (ofQ z) maps Q z to C, and the Primitive Element Theorem *) +(* lets us define a predicate sQ z characterizing the image of (ofQ z), *) +(* as well as a partial inverse (inQ z) to (ofQ z). *) +(* (B) - Construction of the real extension Q[x, y] (~ 230 lines): here y has *) +(* to be a root of a polynomial p over Q[x] satisfying the conditions of *) +(* (3), and Q[x] should be real and archimedean, which we represent by *) +(* a morphism from Q x to some archimedean field R, as the ssrnum and *) +(* fieldext structures are not compatible. The construction starts by *) +(* weakening the condition p[0] = -c^2 to p[0] <= 0 (in R), then reducing *) +(* to the case where p is the minimal polynomial over Q[x] of some y (in *) +(* some Q[w] that contains x and all roots of p). Then we only need to *) +(* construct a realFieldType structure for Q[t] = Q[x,y] (we don't even *) +(* need to show it is consistent with that of R). This amounts to fixing *) +(* the sign of all z != 0 in Q[t], consistently with arithmetic in Q[t]. *) +(* Now any such z is equal to q[y] for some q in Q[x][X] coprime with p. *) +(* Then up + vq = 1 for Bezout coefficients u and v. As p is monic, there *) +(* is some b0 >= 0 in R such that p changes sign in ab0 = [0; b0]. As R *) +(* is archimedean, some iteration of the binary search for a root of p in *) +(* ab0 will yield an interval ab_n such that |up[d]| < 1/2 for d in ab_n. *) +(* Then |q[d]| > 1/2M > 0 for any upper bound M on |v[X]| in ab0, so q *) +(* cannot change sign in ab_n (as then root-finding in ab_n would yield a *) +(* d with |Mq[d]| < 1/2), so we can fix the sign of z to that of q in *) +(* ab_n. *) +(* (C) - Construction of the x_n and z_n (~50 lines): x_ n is obtained by *) +(* iterating (B), starting with x_0 = 0, and then (A) and the PET yield *) +(* z_ n. We establish (1) and (3), and that the minimal polynomial of the *) +(* preimage i_ n of i over the preimage R_ n of Q[x_n] is X^2 + 1. *) +(* (D) - Establish (2), i.e., prove the FTA (~180 lines). We must depart from *) +(* Artin's proof because deciding membership in the union of the Q[x_n] *) +(* requires the FTA, i.e., we cannot (yet) construct a maximal real *) +(* subfield of C. We work around this issue by first reducing to the case *) +(* where Q[z] is Galois over Q and contains i, then using induction over *) +(* the degree of z over Q[z_ n] (i.e., the degree of a monic polynomial *) +(* over Q[z_n] that has z as a root). We can assume that z is not in *) +(* Q[z_n]; then it suffices to find some y in Q[z_n, z] \ Q[z_n] that is *) +(* also in Q[z_m] for some m > n, as then we can apply induction with the *) +(* minimal polynomial of z over Q[z_n, y]. In any Galois extension Q[t] *) +(* of Q that contains both z and z_n, Q[x_n, z] = Q[z_n, z] is Galois *) +(* over both Q[x_n] and Q[z_n]. If Gal(Q[x_n,z] / Q[x_n]) isn't a 2-group *) +(* take one of its Sylow 2-groups P; the minimal polynomial p of any *) +(* generator of the fixed field F of P over Q[x_n] has odd degree, hence *) +(* by (3) - p[X]p[-X] and thus p has a root y in some Q[x_m], hence in *) +(* Q[z_m]. As F is normal, y is in F, with minimal polynomial p, and y *) +(* is not in Q[z_n] = Q[x_n, i] since p has odd degree. Otherwise, *) +(* Gal(Q[z_n,z] / Q[z_n]) is a proper 2-group, and has a maximal subgroup *) +(* P of index 2. The fixed field F of P has a generator w over Q[z_n] *) +(* with w^2 in Q[z_n] \ Q[x_n], i.e. w^2 = u + 2iv with v != 0. From (3) *) +(* X^4 - uX^2 - v^2 has a root x in some Q[x_m]; then x != 0 as v != 0, *) +(* hence w^2 = y^2 for y = x + iv/x in Q[z_m], and y generates F. *) +(* (E) - Construct conj and conclude (~40 lines): conj z is defined as *) +(* conj_ n z with the n provided by (2); since each conj_ m is a morphism *) +(* of order 2 and conj z = conj_ m z for any m >= n, it follows that conj *) +(* is also a morphism of order 2. *) +(* Note that (C), (D) and (E) only depend on Q[x_n] not containing i; the *) +(* order structure is not used (hence we need not prove that the ordering of *) +(* Q[x_m] is consistent with that of Q[x_n] for m >= n). *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Local Notation "p ^ f" := (map_poly f p) : ring_scope. +Local Notation "p ^@" := (p ^ in_alg _) (at level 2, format "p ^@"): ring_scope. +Local Notation "<< E ; u >>" := <<E; u>>%VS. +Local Notation Qmorphism C := {rmorphism rat -> C}. + +Lemma rat_algebraic_archimedean (C : numFieldType) (QtoC : Qmorphism C) : + integralRange QtoC -> Num.archimedean_axiom C. +Proof. +move=> algC x. +without loss x_ge0: x / 0 <= x by rewrite -normr_id; apply; apply: normr_ge0. +have [-> | nz_x] := eqVneq x 0; first by exists 1%N; rewrite normr0. +have [p mon_p px0] := algC x; exists (\sum_(j < size p) `|numq p`_j|)%N. +rewrite ger0_norm // real_ltrNge ?rpred_nat ?ger0_real //. +apply: contraL px0 => lb_x; rewrite rootE gtr_eqF // horner_coef size_map_poly. +have x_gt0 k: 0 < x ^+ k by rewrite exprn_gt0 // ltr_def nz_x. +move: lb_x; rewrite polySpred ?monic_neq0 // !big_ord_recr coef_map /=. +rewrite -lead_coefE (monicP mon_p) natrD rmorph1 mul1r => lb_x. +case: _.-1 (lb_x) => [|n]; first by rewrite !big_ord0 !add0r ltr01. +rewrite -ltr_subl_addl add0r -(ler_pmul2r (x_gt0 n)) -exprS. +apply: ltr_le_trans; rewrite mulrDl mul1r ltr_spaddr // -sumrN. +rewrite natr_sum mulr_suml ler_sum // => j _. +rewrite coef_map /= fmorph_eq_rat (ler_trans (real_ler_norm _)) //. + by rewrite rpredN rpredM ?rpred_rat ?rpredX // ger0_real. +rewrite normrN normrM ler_pmul //=. + rewrite normf_div -!intr_norm -!abszE ler_pimulr ?ler0n //. + by rewrite invf_le1 ?ler1n ?ltr0n ?absz_gt0 ?denq_eq0. +rewrite normrX ger0_norm ?(ltrW x_gt0) // ler_weexpn2l ?leq_ord //. +by rewrite (ler_trans _ lb_x) // -natrD addn1 ler1n. +Qed. + +Definition decidable_embedding sT T (f : sT -> T) := + forall y, decidable (exists x, y = f x). + +Lemma rat_algebraic_decidable (C : fieldType) (QtoC : Qmorphism C) : + integralRange QtoC -> decidable_embedding QtoC. +Proof. +have QtoCinj: injective QtoC by apply: fmorph_inj. +pose ZtoQ : int -> rat := intr; pose ZtoC : int -> C := intr. +have ZtoQinj: injective ZtoQ by apply: intr_inj. +have defZtoC: ZtoC =1 QtoC \o ZtoQ by move=> m; rewrite /= rmorph_int. +move=> algC x; have /sig2_eqW[q mon_q qx0] := algC x; pose d := (size q).-1. +have [n ub_n]: {n | forall y, root q y -> `|y| < n}. + have [n1 ub_n1] := monic_Cauchy_bound mon_q. + have /monic_Cauchy_bound[n2 ub_n2]: (-1) ^+ d *: (q \Po - 'X) \is monic. + rewrite monicE lead_coefZ lead_coef_comp ?size_opp ?size_polyX // -/d. + by rewrite lead_coef_opp lead_coefX (monicP mon_q) (mulrC 1) signrMK. + exists (Num.max n1 n2) => y; rewrite ltrNge ler_normr !ler_maxl rootE. + apply: contraL => /orP[]/andP[] => [/ub_n1/gtr_eqF->// | _ /ub_n2/gtr_eqF]. + by rewrite hornerZ horner_comp !hornerE opprK mulf_eq0 signr_eq0 => /= ->. +have [p [a nz_a Dq]] := rat_poly_scale q; pose N := Num.bound `|n * a%:~R|. +pose xa : seq rat := [seq (m%:R - N%:R) / a%:~R | m <- iota 0 N.*2]. +have [/sig2_eqW[y _ ->] | xa'x] := @mapP _ _ QtoC xa x; first by left; exists y. +right=> [[y Dx]]; case: xa'x; exists y => //. +have{x Dx qx0} qy0: root q y by rewrite Dx fmorph_root in qx0. +have /dvdzP[b Da]: (denq y %| a)%Z. + have /Gauss_dvdzl <-: coprimez (denq y) (numq y ^+ d). + by rewrite coprimez_sym coprimez_expl //; apply: coprime_num_den. + pose p1 : {poly int} := a *: 'X^d - p. + have Dp1: p1 ^ intr = a%:~R *: ('X^d - q). + by rewrite rmorphB linearZ /= map_polyXn scalerBr Dq scalerKV ?intr_eq0. + apply/dvdzP; exists (\sum_(i < d) p1`_i * numq y ^+ i * denq y ^+ (d - i.+1)). + apply: ZtoQinj; rewrite /ZtoQ rmorphM mulr_suml rmorph_sum /=. + transitivity ((p1 ^ intr).[y] * (denq y ^+ d)%:~R). + rewrite Dp1 !hornerE hornerXn (rootP qy0) subr0. + by rewrite !rmorphX /= numqE exprMn mulrA. + have sz_p1: (size (p1 ^ ZtoQ)%R <= d)%N. + rewrite Dp1 size_scale ?intr_eq0 //; apply/leq_sizeP=> i. + rewrite leq_eqVlt eq_sym -polySpred ?monic_neq0 // coefB coefXn. + case: eqP => [-> _ | _ /(nth_default 0)->//]. + by rewrite -lead_coefE (monicP mon_q). + rewrite (horner_coef_wide _ sz_p1) mulr_suml; apply: eq_bigr => i _. + rewrite -!mulrA -exprSr coef_map !rmorphM !rmorphX /= numqE exprMn -mulrA. + by rewrite -exprD -addSnnS subnKC. +pose m := `|(numq y * b + N)%R|%N. +have Dm: m%:R = `|y * a%:~R + N%:R|. + by rewrite pmulrn abszE intr_norm Da rmorphD !rmorphM /= numqE mulrAC mulrA. +have ltr_Qnat n1 n2 : (n1%:R < n2%:R :> rat = _) := ltr_nat _ n1 n2. +have ub_y: `|y * a%:~R| < N%:R. + apply: ler_lt_trans (archi_boundP (normr_ge0 _)); rewrite !normrM. + by rewrite ler_pmul ?normr_ge0 // (ler_trans _ (ler_norm n)) ?ltrW ?ub_n. +apply/mapP; exists m. + rewrite mem_iota /= add0n -addnn -ltr_Qnat Dm natrD. + by rewrite (ler_lt_trans (ler_norm_add _ _)) // normr_nat ltr_add2r. +rewrite Dm ger0_norm ?addrK ?mulfK ?intr_eq0 // -ler_subl_addl sub0r. +by rewrite (ler_trans (ler_norm _)) ?normrN ?ltrW. +Qed. + +Lemma minPoly_decidable_closure + (F : fieldType) (L : closedFieldType) (FtoL : {rmorphism F -> L}) x : + decidable_embedding FtoL -> integralOver FtoL x -> + {p | [/\ p \is monic, root (p ^ FtoL) x & irreducible_poly p]}. +Proof. +move=> isF /sig2W[p /monicP mon_p px0]. +have [r Dp] := closed_field_poly_normal (p ^ FtoL); pose n := size r. +rewrite lead_coef_map {}mon_p rmorph1 scale1r in Dp. +pose Fpx q := (q \is a polyOver isF) && root q x. +have FpxF q: Fpx (q ^ FtoL) = root (q ^ FtoL) x. + by rewrite /Fpx polyOver_poly // => j _; apply/sumboolP; exists q`_j. +pose p_ (I : {set 'I_n}) := \prod_(i <- enum I) ('X - (r`_i)%:P). +have{px0 Dp} /ex_minset[I /minsetP[/andP[FpI pIx0] minI]]: exists I, Fpx (p_ I). + exists setT; suffices ->: p_ setT = p ^ FtoL by rewrite FpxF. + by rewrite Dp (big_nth 0) big_mkord /p_ (eq_enum (in_set _)) big_filter. +have{p} [p DpI]: {p | p_ I = p ^ FtoL}. + exists (p_ I ^ (fun y => if isF y is left Fy then sval (sig_eqW Fy) else 0)). + rewrite -map_poly_comp map_poly_id // => y /(allP FpI) /=. + by rewrite unfold_in; case: (isF y) => // Fy _; case: (sig_eqW _). +have mon_pI: p_ I \is monic by apply: monic_prod_XsubC. +have mon_p: p \is monic by rewrite -(map_monic FtoL) -DpI. +exists p; rewrite -DpI; split=> //; split=> [|q nCq q_dv_p]. + by rewrite -(size_map_poly FtoL) -DpI (root_size_gt1 _ pIx0) ?monic_neq0. +rewrite -dvdp_size_eqp //; apply/eqP. +without loss mon_q: q nCq q_dv_p / q \is monic. + move=> IHq; pose a := lead_coef q; pose q1 := a^-1 *: q. + have nz_a: a != 0 by rewrite lead_coef_eq0 (dvdpN0 q_dv_p) ?monic_neq0. + have /IHq IHq1: q1 \is monic by rewrite monicE lead_coefZ mulVf. + by rewrite -IHq1 ?size_scale ?dvdp_scalel ?invr_eq0. +without loss{nCq} qx0: q mon_q q_dv_p / root (q ^ FtoL) x. + have /dvdpP[q1 Dp] := q_dv_p; rewrite DpI Dp rmorphM rootM -implyNb in pIx0. + have mon_q1: q1 \is monic by rewrite Dp monicMr in mon_p. + move=> IH; apply: (IH) (implyP pIx0 _) => //; apply: contra nCq => /IH IHq1. + rewrite -(subnn (size q1)) {1}IHq1 ?Dp ?dvdp_mulr // polySpred ?monic_neq0 //. + by rewrite eqSS size_monicM ?monic_neq0 // -!subn1 subnAC addKn. +have /dvdp_prod_XsubC[m Dq]: q ^ FtoL %| p_ I by rewrite DpI dvdp_map. +pose B := [set j in mask m (enum I)]; have{Dq} Dq: q ^ FtoL = p_ B. + apply/eqP; rewrite -eqp_monic ?monic_map ?monic_prod_XsubC //. + congr (_ %= _): Dq; apply: eq_big_perm => //. + by rewrite uniq_perm_eq ?mask_uniq ?enum_uniq // => j; rewrite mem_enum inE. +rewrite -!(size_map_poly FtoL) Dq -DpI (minI B) // -?Dq ?FpxF //. +by apply/subsetP=> j; rewrite inE => /mem_mask; rewrite mem_enum. +Qed. + +Lemma alg_integral (F : fieldType) (L : fieldExtType F) : + integralRange (in_alg L). +Proof. +move=> x; have [/polyOver1P[p Dp]] := (minPolyOver 1 x, monic_minPoly 1 x). +by rewrite Dp map_monic; exists p; rewrite // -Dp root_minPoly. +Qed. +Prenex Implicits alg_integral. + +Lemma imaginary_exists (C : closedFieldType) : {i : C | i ^+ 2 = -1}. +Proof. +have /sig_eqW[i Di2] := @solve_monicpoly C 2 (nth 0 [:: -1]) isT. +by exists i; rewrite Di2 big_ord_recl big_ord1 mul0r mulr1 !addr0. +Qed. + +Import DefaultKeying GRing.DefaultPred. +Implicit Arguments map_poly_inj [[F] [R] x1 x2]. + +Theorem Fundamental_Theorem_of_Algebraics : + {L : closedFieldType & + {conj : {rmorphism L -> L} | involutive conj & ~ conj =1 id}}. +Proof. +have maxn3 n1 n2 n3: {m | [/\ n1 <= m, n2 <= m & n3 <= m]%N}. + by exists (maxn n1 (maxn n2 n3)); apply/and3P; rewrite -!geq_max. +have [C [/= QtoC algC]] := countable_algebraic_closure [countFieldType of rat]. +exists C; have [i Di2] := imaginary_exists C. +pose Qfield := fieldExtType rat; pose Cmorph (L : Qfield) := {rmorphism L -> C}. +have charQ (L : Qfield): [char L] =i pred0 := ftrans (char_lalg L) (char_num _). +have sepQ (L : Qfield) (K E : {subfield L}): separable K E. + by apply/separableP=> u _; apply: charf0_separable. +pose genQfield z L := {LtoC : Cmorph L & {u | LtoC u = z & <<1; u>> = fullv}}. +have /all_tag[Q /all_tag[ofQ genQz]] z: {Qz : Qfield & genQfield z Qz}. + have [|p [/monic_neq0 nzp pz0 irr_p]] := minPoly_decidable_closure _ (algC z). + exact: rat_algebraic_decidable. + pose Qz := SubFieldExtType pz0 irr_p. + pose QzC := subfx_inj_rmorphism QtoC z p. + exists Qz, QzC, (subfx_root QtoC z p); first exact: subfx_inj_root. + apply/vspaceP=> u; rewrite memvf; apply/Fadjoin1_polyP. + by have [q] := subfxEroot pz0 nzp u; exists q. +have pQof z p: p^@ ^ ofQ z = p ^ QtoC. + by rewrite -map_poly_comp; apply: eq_map_poly => x; rewrite !fmorph_eq_rat. +have pQof2 z p u: ofQ z p^@.[u] = (p ^ QtoC).[ofQ z u]. + by rewrite -horner_map pQof. +have PET_Qz z (E : {subfield Q z}): {u | <<1; u>> = E}. + exists (separable_generator 1 E). + by rewrite -eq_adjoin_separable_generator ?sub1v. +pose gen z x := exists q, x = (q ^ QtoC).[z]. +have PET2 x y: {z | gen z x & gen z y}. + pose Gxy := (x, y) = let: (p, q, z) := _ in ((p ^ QtoC).[z], (q ^ QtoC).[z]). + suffices [[[p q] z] []]: {w | Gxy w} by exists z; [exists p | exists q]. + apply/sig_eqW; have /integral_algebraic[px nz_px pxx0] := algC x. + have /integral_algebraic[py nz_py pyy0] := algC y. + have [n [[p Dx] [q Dy]]] := char0_PET nz_px pxx0 nz_py pyy0 (char_num _). + by exists (p, q, y *+ n - x); congr (_, _). +have gen_inQ z x: gen z x -> {u | ofQ z u = x}. + have [u Dz _] := genQz z => /sig_eqW[q ->]. + by exists q^@.[u]; rewrite pQof2 Dz. +have gen_ofP z u v: reflect (gen (ofQ z u) (ofQ z v)) (v \in <<1; u>>). + apply: (iffP Fadjoin1_polyP) => [[q ->]|]; first by rewrite pQof2; exists q. + by case=> q; rewrite -pQof2 => /fmorph_inj->; exists q. +have /all_tag[sQ genP] z: {s : pred C & forall x, reflect (gen z x) (x \in s)}. + apply: all_tag (fun x => reflect (gen z x)) _ => x. + have [w /gen_inQ[u <-] /gen_inQ[v <-]] := PET2 z x. + by exists (v \in <<1; u>>)%VS; apply: gen_ofP. +have sQtrans: transitive (fun x z => x \in sQ z). + move=> x y z /genP[p ->] /genP[q ->]; apply/genP; exists (p \Po q). + by rewrite map_comp_poly horner_comp. +have sQid z: z \in sQ z by apply/genP; exists 'X; rewrite map_polyX hornerX. +have{gen_ofP} sQof2 z u v: (ofQ z u \in sQ (ofQ z v)) = (u \in <<1; v>>%VS). + exact/genP/(gen_ofP z). +have sQof z v: ofQ z v \in sQ z. + by have [u Dz defQz] := genQz z; rewrite -[in sQ z]Dz sQof2 defQz memvf. +have{gen_inQ} sQ_inQ z x z_x := gen_inQ z x (genP z x z_x). +have /all_sig[inQ inQ_K] z: {inQ | {in sQ z, cancel inQ (ofQ z)}}. + by apply: all_sig_cond (fun x u => ofQ z u = x) 0 _ => x /sQ_inQ. +have ofQ_K z: cancel (ofQ z) (inQ z). + by move=> x; have /inQ_K/fmorph_inj := sQof z x. +have sQring z: divring_closed (sQ z). + have sQ_1: 1 \in sQ z by rewrite -(rmorph1 (ofQ z)) sQof. + by split=> // x y /inQ_K<- /inQ_K<- /=; rewrite -(rmorphB, fmorph_div) sQof. +have sQopp z : oppr_closed (sQ z) := sQring z. +have sQadd z : addr_closed (sQ z) := sQring z. +have sQmul z : mulr_closed (sQ z) := sQring z. +have sQinv z : invr_closed (sQ z) := sQring z. +pose morph_ofQ x z Qxz := forall u, ofQ z (Qxz u) = ofQ x u. +have QtoQ z x: x \in sQ z -> {Qxz : 'AHom(Q x, Q z) | morph_ofQ x z Qxz}. + move=> z_x; pose Qxz u := inQ z (ofQ x u). + have QxzE u: ofQ z (Qxz u) = ofQ x u by apply/inQ_K/(sQtrans x). + suffices /rat_lrmorphism QxzM: rmorphism Qxz. + by exists (linfun_ahom (LRMorphism QxzM)) => u; rewrite lfunE QxzE. + split=> [u v|]; first by apply: (canLR (ofQ_K z)); rewrite !rmorphB !QxzE. + by split=> [u v|]; apply: (canLR (ofQ_K z)); rewrite ?rmorph1 ?rmorphM ?QxzE. +pose sQs z s := all (mem (sQ z)) s. +have inQsK z s: sQs z s -> map (ofQ z) (map (inQ z) s) = s. + by rewrite -map_comp => /allP/(_ _ _)/inQ_K; apply: map_id_in. +have inQpK z p: p \is a polyOver (sQ z) -> (p ^ inQ z) ^ ofQ z = p. + by move=> /allP/(_ _ _)/inQ_K/=/map_poly_id; rewrite -map_poly_comp. +have{gen PET2 genP} PET s: {z | sQs z s & <<1 & map (inQ z) s>>%VS = fullv}. + have [y /inQsK Ds]: {y | sQs y s}. + elim: s => [|x s /= [y IHs]]; first by exists 0. + have [z /genP z_x /genP z_y] := PET2 x y. + by exists z; rewrite /= {x}z_x; apply: sub_all IHs => x /sQtrans/= ->. + have [w defQs] := PET_Qz _ <<1 & map (inQ y) s>>%AS; pose z := ofQ y w. + have z_s: sQs z s. + rewrite -Ds /sQs all_map; apply/allP=> u s_u /=. + by rewrite sQof2 defQs seqv_sub_adjoin. + have [[u Dz defQz] [Qzy QzyE]] := (genQz z, QtoQ y z (sQof y w)). + exists z => //; apply/eqP; rewrite eqEsubv subvf /= -defQz. + rewrite -(limg_ker0 _ _ (AHom_lker0 Qzy)) aimg_adjoin_seq aimg_adjoin aimg1. + rewrite -[map _ _](mapK (ofQ_K y)) -(map_comp (ofQ y)) (eq_map QzyE) inQsK //. + by rewrite -defQs -(canLR (ofQ_K y) Dz) -QzyE ofQ_K. +pose rp s := \prod_(z <- s) ('X - z%:P). +have map_rp (f : {rmorphism _}) s: rp _ s ^ f = rp _ (map f s). + rewrite rmorph_prod /rp big_map; apply: eq_bigr => x _. + by rewrite rmorphB /= map_polyX map_polyC. +pose is_Gal z := SplittingField.axiom (Q z). +have galQ x: {z | x \in sQ z & is_Gal z}. + have /sig2W[p mon_p pz0] := algC x. + have [s Dp] := closed_field_poly_normal (p ^ QtoC). + rewrite (monicP _) ?monic_map // scale1r in Dp; have [z z_s defQz] := PET s. + exists z; first by apply/(allP z_s); rewrite -root_prod_XsubC -Dp. + exists p^@; first exact: alg_polyOver. + exists (map (inQ z) s); last by apply/vspaceP=> u; rewrite defQz memvf. + by rewrite -(eqp_map (ofQ z)) pQof Dp map_rp inQsK ?eqpxx. +pose is_realC x := {R : archiFieldType & {rmorphism Q x -> R}}. +pose realC := {x : C & is_realC x}. +pose has_Rroot (xR : realC) p c (Rx := sQ (tag xR)) := + [&& p \is a polyOver Rx, p \is monic, c \in Rx & p.[0] == - c ^+ 2]. +pose root_in (xR : realC) p := exists2 w, w \in sQ (tag xR) & root p w. +pose extendsR (xR yR : realC) := tag xR \in sQ (tag yR). +have add_Rroot xR p c: {yR | extendsR xR yR & has_Rroot xR p c -> root_in yR p}. + rewrite {}/extendsR; case: (has_Rroot xR p c) / and4P; last by exists xR. + case: xR => x [R QxR] /= [/inQpK <-]; move: (p ^ _) => {p}p mon_p /inQ_K<- Dc. + have{c Dc} p0_le0: (p ^ QxR).[0] <= 0. + rewrite horner_coef0 coef_map -[p`_0]ofQ_K -coef_map -horner_coef0 (eqP Dc). + by rewrite -rmorphX -rmorphN ofQ_K /= rmorphN rmorphX oppr_le0 sqr_ge0. + have [s Dp] := closed_field_poly_normal (p ^ ofQ x). + have{Dp} /all_and2[s_p p_s] y: root (p ^ ofQ x) y <-> (y \in s). + by rewrite Dp (monicP mon_p) scale1r root_prod_XsubC. + rewrite map_monic in mon_p; have [z /andP[z_x /allP/=z_s] _] := PET (x :: s). + have{z_x} [[Qxz QxzE] Dx] := (QtoQ z x z_x, inQ_K z x z_x). + pose Qx := <<1; inQ z x>>%AS; pose QxzM := [rmorphism of Qxz]. + have pQwx q1: q1 \is a polyOver Qx -> {q | q1 = q ^ Qxz}. + move/polyOverP=> Qx_q1; exists ((q1 ^ ofQ z) ^ inQ x). + apply: (map_poly_inj (ofQ z)); rewrite -map_poly_comp (eq_map_poly QxzE). + by rewrite inQpK ?polyOver_poly // => j _; rewrite -Dx sQof2 Qx_q1. + have /all_sig[t_ Dt] u: {t | <<1; t>> = <<Qx; u>>} by apply: PET_Qz. + suffices{p_s}[u Ry px0]: {u : Q z & is_realC (ofQ z (t_ u)) & ofQ z u \in s}. + exists (Tagged is_realC Ry) => [|_] /=. + by rewrite -Dx sQof2 Dt subvP_adjoin ?memv_adjoin. + by exists (ofQ z u); rewrite ?p_s // sQof2 Dt memv_adjoin. + without loss{z_s s_p} [u Dp s_y]: p mon_p p0_le0 / + {u | minPoly Qx u = p ^ Qxz & ofQ z u \in s}. + - move=> IHp; move: {2}_.+1 (ltnSn (size p)) => d. + elim: d => // d IHd in p mon_p s_p p0_le0 *; rewrite ltnS => le_p_d. + have /closed_rootP/sig_eqW[y py0]: size (p ^ ofQ x) != 1%N. + rewrite size_map_poly size_poly_eq1 eqp_monic ?rpred1 //. + by apply: contraTneq p0_le0 => ->; rewrite rmorph1 hornerC ltr_geF ?ltr01. + have /s_p s_y := py0; have /z_s/sQ_inQ[u Dy] := s_y. + have /pQwx[q Dq] := minPolyOver Qx u. + have mon_q: q \is monic by have:= monic_minPoly Qx u; rewrite Dq map_monic. + have /dvdpP/sig_eqW[r Dp]: q %| p. + rewrite -(dvdp_map QxzM) -Dq minPoly_dvdp //. + by apply: polyOver_poly => j _; rewrite -sQof2 QxzE Dx. + by rewrite -(fmorph_root (ofQ z)) Dy -map_poly_comp (eq_map_poly QxzE). + have mon_r: r \is monic by rewrite Dp monicMr in mon_p. + have [q0_le0 | q0_gt0] := lerP ((q ^ QxR).[0]) 0. + by apply: (IHp q) => //; exists u; rewrite ?Dy. + have r0_le0: (r ^ QxR).[0] <= 0. + by rewrite -(ler_pmul2r q0_gt0) mul0r -hornerM -rmorphM -Dp. + apply: (IHd r mon_r) => // [w rw0|]. + by rewrite s_p // Dp rmorphM rootM rw0. + apply: leq_trans le_p_d; rewrite Dp size_Mmonic ?monic_neq0 // addnC. + by rewrite -(size_map_poly QxzM q) -Dq size_minPoly !ltnS leq_addl. + exists u => {s s_y}//; set y := ofQ z (t_ u); set p1 := minPoly Qx u in Dp. + have /QtoQ[Qyz QyzE]: y \in sQ z := sQof z (t_ u). + pose q1_ v := Fadjoin_poly Qx u (Qyz v). + have{QyzE} QyzE v: Qyz v = (q1_ v).[u]. + by rewrite Fadjoin_poly_eq // -Dt -sQof2 QyzE sQof. + have /all_sig2[q_ coqp Dq] v: {q | v != 0 -> coprimep p q & q ^ Qxz = q1_ v}. + have /pQwx[q Dq]: q1_ v \is a polyOver Qx by apply: Fadjoin_polyOver. + exists q => // nz_v; rewrite -(coprimep_map QxzM) -Dp -Dq -gcdp_eqp1. + have /minPoly_irr/orP[] // := dvdp_gcdl p1 (q1_ v). + by rewrite gcdp_polyOver ?minPolyOver ?Fadjoin_polyOver. + rewrite -/p1 {1}/eqp dvdp_gcd => /and3P[_ _ /dvdp_leq/=/implyP]. + rewrite size_minPoly ltnNge size_poly (contraNneq _ nz_v) // => q1v0. + by rewrite -(fmorph_eq0 [rmorphism of Qyz]) /= QyzE q1v0 horner0. + pose h2 : R := 2%:R^-1; have nz2: 2%:R != 0 :> R by rewrite pnatr_eq0. + pose itv ab := [pred c : R | ab.1 <= c <= ab.2]. + pose wid ab : R := ab.2 - ab.1; pose mid ab := (ab.1 + ab.2) * h2. + pose sub_itv ab cd := cd.1 <= ab.1 :> R /\ ab.2 <= cd.2 :> R. + pose xup q ab := [/\ q.[ab.1] <= 0, q.[ab.2] >= 0 & ab.1 <= ab.2 :> R]. + pose narrow q ab (c := mid ab) := if q.[c] >= 0 then (ab.1, c) else (c, ab.2). + pose find k q := iter k (narrow q). + have findP k q ab (cd := find k q ab): + xup q ab -> [/\ xup q cd, sub_itv cd ab & wid cd = wid ab / (2 ^ k)%:R]. + - rewrite {}/cd; case: ab => a b xq_ab. + elim: k => /= [|k]; first by rewrite divr1. + case: (find k q _) => c d [[/= qc_le0 qd_ge0 le_cd] [/= le_ac le_db] Dcd]. + have [/= le_ce le_ed] := midf_le le_cd; set e := _ / _ in le_ce le_ed. + rewrite expnSr natrM invfM mulrA -{}Dcd /narrow /= -[mid _]/e. + have [qe_ge0 // | /ltrW qe_le0] := lerP 0 q.[e]. + do ?split=> //=; [exact: (ler_trans le_ed) | apply: canRL (mulfK nz2) _]. + by rewrite mulrBl divfK // mulr_natr opprD addrACA subrr add0r. + do ?split=> //=; [exact: (ler_trans le_ac) | apply: canRL (mulfK nz2) _]. + by rewrite mulrBl divfK // mulr_natr opprD addrACA subrr addr0. + have find_root r q ab: + xup q ab -> {n | forall x, x \in itv (find n q ab) ->`|(r * q).[x]| < h2}. + - move=> xab; have ub_ab := poly_itv_bound _ ab.1 ab.2. + have [Mu MuP] := ub_ab r; have /all_sig[Mq MqP] j := ub_ab q^`N(j). + pose d := wid ab; pose dq := \poly_(i < (size q).-1) Mq i.+1. + have d_ge0: 0 <= d by rewrite subr_ge0; case: xab. + have [Mdq MdqP] := poly_disk_bound dq d. + pose n := Num.bound (Mu * Mdq * d); exists n => c /= /andP[]. + have{xab} [[]] := findP n _ _ xab; case: (find n q ab) => a1 b1 /=. + rewrite -/d => qa1_le0 qb1_ge0 le_ab1 [/= le_aa1 le_b1b] Dab1 le_a1c le_cb1. + have /MuP lbMu: c \in itv ab. + by rewrite !inE (ler_trans le_aa1) ?(ler_trans le_cb1). + have Mu_ge0: 0 <= Mu by rewrite (ler_trans _ lbMu) ?normr_ge0. + have Mdq_ge0: 0 <= Mdq. + by rewrite (ler_trans _ (MdqP 0 _)) ?normr_ge0 ?normr0. + suffices lb1 a2 b2 (ab1 := (a1, b1)) (ab2 := (a2, b2)) : + xup q ab2 /\ sub_itv ab2 ab1 -> q.[b2] - q.[a2] <= Mdq * wid ab1. + + apply: ler_lt_trans (_ : Mu * Mdq * wid (a1, b1) < h2); last first. + rewrite {}Dab1 mulrA ltr_pdivr_mulr ?ltr0n ?expn_gt0 //. + rewrite (ltr_le_trans (archi_boundP _)) ?mulr_ge0 ?ltr_nat // -/n. + rewrite ler_pdivl_mull ?ltr0n // -natrM ler_nat. + by case: n => // n; rewrite expnS leq_pmul2l // ltn_expl. + rewrite -mulrA hornerM normrM ler_pmul ?normr_ge0 //. + have [/ltrW qc_le0 | qc_ge0] := ltrP q.[c] 0. + by apply: ler_trans (lb1 c b1 _); rewrite ?ler0_norm ?ler_paddl. + by apply: ler_trans (lb1 a1 c _); rewrite ?ger0_norm ?ler_paddr ?oppr_ge0. + case{c le_a1c le_cb1 lbMu}=> [[/=qa2_le0 qb2_ge0 le_ab2] [/=le_a12 le_b21]]. + pose h := b2 - a2; have h_ge0: 0 <= h by rewrite subr_ge0. + have [-> | nz_q] := eqVneq q 0. + by rewrite !horner0 subrr mulr_ge0 ?subr_ge0. + rewrite -(subrK a2 b2) (addrC h) (nderiv_taylor q (mulrC a2 h)). + rewrite (polySpred nz_q) big_ord_recl /= mulr1 nderivn0 addrC addKr. + have [le_aa2 le_b2b] := (ler_trans le_aa1 le_a12, ler_trans le_b21 le_b1b). + have /MqP MqPx1: a2 \in itv ab by rewrite inE le_aa2 (ler_trans le_ab2). + apply: ler_trans (ler_trans (ler_norm _) (ler_norm_sum _ _ _)) _. + apply: ler_trans (_ : `|dq.[h] * h| <= _); last first. + by rewrite normrM ler_pmul ?normr_ge0 ?MdqP // ?ger0_norm ?ler_sub ?h_ge0. + rewrite horner_poly ger0_norm ?mulr_ge0 ?sumr_ge0 // => [|j _]; last first. + by rewrite mulr_ge0 ?exprn_ge0 // (ler_trans _ (MqPx1 _)) ?normr_ge0. + rewrite mulr_suml ler_sum // => j _; rewrite normrM -mulrA -exprSr. + by rewrite ler_pmul ?normr_ge0 // normrX ger0_norm. + have [ab0 xab0]: {ab | xup (p ^ QxR) ab}. + have /monic_Cauchy_bound[b pb_gt0]: p ^ QxR \is monic by apply: monic_map. + by exists (0, `|b|); rewrite /xup normr_ge0 p0_le0 ltrW ?pb_gt0 ?ler_norm. + pose ab_ n := find n (p ^ QxR) ab0; pose Iab_ n := itv (ab_ n). + pose lim v a := (q_ v ^ QxR).[a]; pose nlim v n := lim v (ab_ n).2. + have lim0 a: lim 0 a = 0. + rewrite /lim; suffices /eqP ->: q_ 0 == 0 by rewrite rmorph0 horner0. + by rewrite -(map_poly_eq0 QxzM) Dq /q1_ !raddf0. + have limN v a: lim (- v) a = - lim v a. + rewrite /lim; suffices ->: q_ (- v) = - q_ v by rewrite rmorphN hornerN. + by apply: (map_poly_inj QxzM); rewrite Dq /q1_ !raddfN /= Dq. + pose lim_nz n v := exists2 e, e > 0 & {in Iab_ n, forall a, e < `|lim v a| }. + have /(all_sig_cond 0%N)[n_ nzP] v: v != 0 -> {n | lim_nz n v}. + move=> nz_v; do [move/(_ v nz_v); rewrite -(coprimep_map QxR)] in coqp. + have /sig_eqW[r r_pq_1] := Bezout_eq1_coprimepP _ _ coqp. + have /(find_root r.1)[n ub_rp] := xab0; exists n. + have [M Mgt0 ubM]: {M | 0 < M & {in Iab_ n, forall a, `|r.2.[a]| <= M}}. + have [M ubM] := poly_itv_bound r.2 (ab_ n).1 (ab_ n).2. + exists (Num.max 1 M) => [|s /ubM vM]; first by rewrite ltr_maxr ltr01. + by rewrite ler_maxr orbC vM. + exists (h2 / M) => [|a xn_a]; first by rewrite divr_gt0 ?invr_gt0 ?ltr0n. + rewrite ltr_pdivr_mulr // -(ltr_add2l h2) -mulr2n -mulr_natl divff //. + rewrite -normr1 -(hornerC 1 a) -[1%:P]r_pq_1 hornerD. + rewrite ?(ler_lt_trans (ler_norm_add _ _)) ?ltr_le_add ?ub_rp //. + by rewrite mulrC hornerM normrM ler_wpmul2l ?ubM. + have ab_le m n: (m <= n)%N -> (ab_ n).2 \in Iab_ m. + move/subnKC=> <-; move: {n}(n - m)%N => n; rewrite /ab_. + have /(findP m)[/(findP n)[[_ _]]] := xab0. + rewrite /find -iter_add -!/(find _ _) -!/(ab_ _) addnC !inE. + by move: (ab_ _) => /= ab_mn le_ab_mn [/ler_trans->]. + pose lt v w := 0 < nlim (w - v) (n_ (w - v)). + have posN v: lt 0 (- v) = lt v 0 by rewrite /lt subr0 add0r. + have posB v w: lt 0 (w - v) = lt v w by rewrite /lt subr0. + have posE n v: (n_ v <= n)%N -> lt 0 v = (0 < nlim v n). + rewrite /lt subr0 /nlim => /ab_le; set a := _.2; set b := _.2 => Iv_a. + have [-> | /nzP[e e_gt0]] := eqVneq v 0; first by rewrite !lim0 ltrr. + move: (n_ v) => m in Iv_a b * => v_gte. + without loss lt0v: v v_gte / 0 < lim v b. + move=> IHv; apply/idP/idP => [v_gt0 | /ltrW]; first by rewrite -IHv. + rewrite ltr_def -normr_gt0 ?(ltr_trans _ (v_gte _ _)) ?ab_le //=. + rewrite !lerNgt -!oppr_gt0 -!limN; apply: contra => v_lt0. + by rewrite -IHv // => c /v_gte; rewrite limN normrN. + rewrite lt0v (ltr_trans e_gt0) ?(ltr_le_trans (v_gte a Iv_a)) //. + rewrite ger0_norm // lerNgt; apply/negP=> /ltrW lev0. + have [le_a le_ab] : _ /\ a <= b := andP Iv_a. + have xab: xup (q_ v ^ QxR) (a, b) by move/ltrW in lt0v. + have /(find_root (h2 / e)%:P)[n1] := xab; have /(findP n1)[[_ _]] := xab. + case: (find _ _ _) => c d /= le_cd [/= le_ac le_db] _ /(_ c)/implyP. + rewrite inE lerr le_cd hornerM hornerC normrM ler_gtF //. + rewrite ger0_norm ?divr_ge0 ?invr_ge0 ?ler0n ?(ltrW e_gt0) // mulrAC. + rewrite ler_pdivl_mulr // ler_wpmul2l ?invr_ge0 ?ler0n // ltrW // v_gte //=. + by rewrite inE -/b (ler_trans le_a) //= (ler_trans le_cd). + pose lim_pos m v := exists2 e, e > 0 & forall n, (m <= n)%N -> e < nlim v n. + have posP v: reflect (exists m, lim_pos m v) (lt 0 v). + apply: (iffP idP) => [v_gt0|[m [e e_gt0 v_gte]]]; last first. + by rewrite (posE _ _ (leq_maxl _ m)) (ltr_trans e_gt0) ?v_gte ?leq_maxr. + have [|e e_gt0 v_gte] := nzP v. + by apply: contraTneq v_gt0 => ->; rewrite /lt subr0 /nlim lim0 ltrr. + exists (n_ v), e => // n le_vn; rewrite (posE n) // in v_gt0. + by rewrite -(ger0_norm (ltrW v_gt0)) v_gte ?ab_le. + have posNneg v: lt 0 v -> ~~ lt v 0. + case/posP=> m [d d_gt0 v_gtd]; rewrite -posN. + apply: contraL d_gt0 => /posP[n [e e_gt0 nv_gte]]. + rewrite ltr_gtF // (ltr_trans (v_gtd _ (leq_maxl m n))) // -oppr_gt0. + by rewrite /nlim -limN (ltr_trans e_gt0) ?nv_gte ?leq_maxr. + have posVneg v: v != 0 -> lt 0 v || lt v 0. + case/nzP=> e e_gt0 v_gte; rewrite -posN; set w := - v. + have [m [le_vm le_wm _]] := maxn3 (n_ v) (n_ w) 0%N; rewrite !(posE m) //. + by rewrite /nlim limN -ltr_normr (ltr_trans e_gt0) ?v_gte ?ab_le. + have posD v w: lt 0 v -> lt 0 w -> lt 0 (v + w). + move=> /posP[m [d d_gt0 v_gtd]] /posP[n [e e_gt0 w_gte]]. + apply/posP; exists (maxn m n), (d + e) => [|k]; first exact: addr_gt0. + rewrite geq_max => /andP[le_mk le_nk]; rewrite /nlim /lim. + have ->: q_ (v + w) = q_ v + q_ w. + by apply: (map_poly_inj QxzM); rewrite rmorphD /= !{1}Dq /q1_ !raddfD. + by rewrite rmorphD hornerD ltr_add ?v_gtd ?w_gte. + have posM v w: lt 0 v -> lt 0 w -> lt 0 (v * w). + move=> /posP[m [d d_gt0 v_gtd]] /posP[n [e e_gt0 w_gte]]. + have /dvdpP[r /(canRL (subrK _))Dqvw]: p %| q_ (v * w) - q_ v * q_ w. + rewrite -(dvdp_map QxzM) rmorphB rmorphM /= !Dq -Dp minPoly_dvdp //. + by rewrite rpredB 1?rpredM ?Fadjoin_polyOver. + by rewrite rootE !hornerE -!QyzE rmorphM subrr. + have /(find_root ((d * e)^-1 *: r ^ QxR))[N ub_rp] := xab0. + pose f := d * e * h2; apply/posP; exists (maxn N (maxn m n)), f => [|k]. + by rewrite !mulr_gt0 ?invr_gt0 ?ltr0n. + rewrite !geq_max => /and3P[/ab_le/ub_rp{ub_rp}ub_rp le_mk le_nk]. + rewrite -(ltr_add2r f) -mulr2n -mulr_natr divfK // /nlim /lim Dqvw. + rewrite rmorphD hornerD /= -addrA -ltr_subl_addl ler_lt_add //. + by rewrite rmorphM hornerM ler_pmul ?ltrW ?v_gtd ?w_gte. + rewrite -ltr_pdivr_mull ?mulr_gt0 // (ler_lt_trans _ ub_rp) //. + by rewrite -scalerAl hornerZ -rmorphM mulrN -normrN ler_norm. + pose le v w := (w == v) || lt v w. + pose abs v := if le 0 v then v else - v. + have absN v: abs (- v) = abs v. + rewrite /abs /le oppr_eq0 opprK posN. + have [-> | /posVneg/orP[v_gt0 | v_lt0]] := altP eqP; first by rewrite oppr0. + by rewrite v_gt0 /= -if_neg posNneg. + by rewrite v_lt0 /= -if_neg -(opprK v) posN posNneg ?posN. + have absE v: le 0 v -> abs v = v by rewrite /abs => ->. + pose QyNum := RealLtMixin posD posM posNneg posB posVneg absN absE (rrefl _). + pose QyNumField := [numFieldType of NumDomainType (Q y) QyNum]. + pose Ry := [realFieldType of RealDomainType _ (RealLeAxiom QyNumField)]. + have archiRy := @rat_algebraic_archimedean Ry _ alg_integral. + by exists (ArchiFieldType Ry archiRy); apply: [rmorphism of idfun]. +have some_realC: realC. + suffices /all_sig[f QfK] x: {a | in_alg (Q 0) a = x}. + exists 0, [archiFieldType of rat], f. + exact: can2_rmorphism (inj_can_sym QfK (fmorph_inj _)) QfK. + have /Fadjoin1_polyP/sig_eqW[q]: x \in <<1; 0>>%VS by rewrite -sQof2 rmorph0. + by exists q.[0]; rewrite -horner_map rmorph0. +pose fix xR n : realC := + if n isn't n'.+1 then some_realC else + if unpickle (nth 0%N (CodeSeq.decode n') 1) isn't Some (p, c) then xR n' else + tag (add_Rroot (xR n') p c). +pose x_ n := tag (xR n). +have sRle m n: (m <= n)%N -> {subset sQ (x_ m) <= sQ (x_ n)}. + move/subnK <-; elim: {n}(n - m)%N => // n IHn x /IHn{IHn}Rx. + rewrite addSn /x_ /=; case: (unpickle _) => [[p c]|] //=. + by case: (add_Rroot _ _ _) => yR /= /(sQtrans _ x)->. +have xRroot n p c: has_Rroot (xR n) p c -> {m | n <= m & root_in (xR m) p}%N. + case/and4P=> Rp mon_p Rc Dc; pose m := CodeSeq.code [:: n; pickle (p, c)]. + have le_n_m: (n <= m)%N by apply/ltnW/(allP (CodeSeq.ltn_code _))/mem_head. + exists m.+1; rewrite ?leqW /x_ //= CodeSeq.codeK pickleK. + case: (add_Rroot _ _ _) => yR /= _; apply; apply/and4P. + by split=> //; first apply: polyOverS Rp; apply: (sRle n). +have /all_sig[z_ /all_and3[Ri_R Ri_i defRi]] n (x := x_ n): + {z | [/\ x \in sQ z, i \in sQ z & <<<<1; inQ z x>>; inQ z i>> = fullv]}. +- have [z /and3P[z_x z_i _] Dzi] := PET [:: x; i]. + by exists z; rewrite -adjoin_seq1 -adjoin_cons. +pose i_ n := inQ (z_ n) i; pose R_ n := <<1; inQ (z_ n) (x_ n)>>%AS. +have memRi n: <<R_ n; i_ n>> =i predT by move=> u; rewrite defRi memvf. +have sCle m n: (m <= n)%N -> {subset sQ (z_ m) <= sQ (z_ n)}. + move/sRle=> Rmn _ /sQ_inQ[u <-]. + have /Fadjoin_polyP[p /polyOverP Rp ->] := memRi m u. + rewrite -horner_map inQ_K ?rpred_horner //=; apply/polyOver_poly=> j _. + by apply: sQtrans (Ri_R n); rewrite Rmn // -(inQ_K _ _ (Ri_R m)) sQof2. +have R'i n: i \notin sQ (x_ n). + rewrite /x_; case: (xR n) => x [Rn QxR] /=. + apply: contraL (@ltr01 Rn) => /sQ_inQ[v Di]. + suffices /eqP <-: - QxR v ^+ 2 == 1 by rewrite oppr_gt0 -lerNgt sqr_ge0. + rewrite -rmorphX -rmorphN fmorph_eq1 -(fmorph_eq1 (ofQ x)) rmorphN eqr_oppLR. + by rewrite rmorphX Di Di2. +have szX2_1: size ('X^2 + 1) = 3. + by move=> R; rewrite size_addl ?size_polyXn ?size_poly1. +have minp_i n (p_i := minPoly (R_ n) (i_ n)): p_i = 'X^2 + 1. + have p_dv_X2_1: p_i %| 'X^2 + 1. + rewrite minPoly_dvdp ?rpredD ?rpredX ?rpred1 ?polyOverX //. + rewrite -(fmorph_root (ofQ _)) inQ_K // rmorphD rmorph1 /= map_polyXn. + by rewrite rootE hornerD hornerXn hornerC Di2 addNr. + apply/eqP; rewrite -eqp_monic ?monic_minPoly //; last first. + by rewrite monicE lead_coefE szX2_1 coefD coefXn coefC addr0. + rewrite -dvdp_size_eqp // eqn_leq dvdp_leq -?size_poly_eq0 ?szX2_1 //= ltnNge. + by rewrite size_minPoly ltnS leq_eqVlt orbF adjoin_deg_eq1 -sQof2 !inQ_K. +have /all_sig[n_ FTA] z: {n | z \in sQ (z_ n)}. + without loss [z_i gal_z]: z / i \in sQ z /\ is_Gal z. + have [y /and3P[/sQtrans y_z /sQtrans y_i _] _] := PET [:: z; i]. + have [t /sQtrans t_y gal_t] := galQ y. + by case/(_ t)=> [|n]; last exists n; rewrite ?y_z ?y_i ?t_y. + apply/sig_eqW; have n := 0%N. + have [p]: exists p, [&& p \is monic, root p z & p \is a polyOver (sQ (z_ n))]. + have [p mon_p pz0] := algC z; exists (p ^ QtoC). + by rewrite map_monic mon_p pz0 -(pQof (z_ n)); apply/polyOver_poly. + elim: {p}_.+1 {-2}p n (ltnSn (size p)) => // d IHd p n lepd pz0. + have [t [t_C t_z gal_t]]: exists t, [/\ z_ n \in sQ t, z \in sQ t & is_Gal t]. + have [y /and3P[y_C y_z _]] := PET [:: z_ n; z]. + by have [t /(sQtrans y)t_y] := galQ y; exists t; rewrite !t_y. + pose Qt := SplittingFieldType rat (Q t) gal_t; have /QtoQ[CnQt CnQtE] := t_C. + pose Rn : {subfield Qt} := (CnQt @: R_ n)%AS; pose i_t : Qt := CnQt (i_ n). + pose Cn : {subfield Qt} := <<Rn; i_t>>%AS. + have defCn: Cn = limg CnQt :> {vspace Q t} by rewrite /= -aimg_adjoin defRi. + have memRn u: (u \in Rn) = (ofQ t u \in sQ (x_ n)). + by rewrite /= aimg_adjoin aimg1 -sQof2 CnQtE inQ_K. + have memCn u: (u \in Cn) = (ofQ t u \in sQ (z_ n)). + have [v Dv genCn] := genQz (z_ n). + by rewrite -Dv -CnQtE sQof2 defCn -genCn aimg_adjoin aimg1. + have Dit: ofQ t i_t = i by rewrite CnQtE inQ_K. + have Dit2: i_t ^+ 2 = -1. + by apply: (fmorph_inj (ofQ t)); rewrite rmorphX rmorphN1 Dit. + have dimCn: \dim_Rn Cn = 2. + rewrite -adjoin_degreeE adjoin_degree_aimg. + by apply: succn_inj; rewrite -size_minPoly minp_i. + have /sQ_inQ[u_z Dz] := t_z; pose Rz := <<Cn; u_z>>%AS. + have{p lepd pz0} le_Rz_d: (\dim_Cn Rz < d)%N. + rewrite -ltnS -adjoin_degreeE -size_minPoly (leq_trans _ lepd) // !ltnS. + have{pz0} [mon_p pz0 Cp] := and3P pz0. + have{Cp} Dp: ((p ^ inQ (z_ n)) ^ CnQt) ^ ofQ t = p. + by rewrite -map_poly_comp (eq_map_poly CnQtE) inQpK. + rewrite -Dp size_map_poly dvdp_leq ?monic_neq0 -?(map_monic (ofQ _)) ?Dp //. + rewrite defCn minPoly_dvdp //; try by rewrite -(fmorph_root (ofQ t)) Dz Dp. + by apply/polyOver_poly=> j _; rewrite memv_img ?memvf. + have [sRCn sCnRz]: (Rn <= Cn)%VS /\ (Cn <= Rz)%VS by rewrite !subv_adjoin. + have sRnRz := subv_trans sRCn sCnRz. + have{gal_z} galRz: galois Rn Rz. + apply/and3P; split=> //; apply/splitting_normalField=> //. + pose u : SplittingFieldType rat (Q z) gal_z := inQ z z. + have /QtoQ[Qzt QztE] := t_z; exists (minPoly 1 u ^ Qzt). + have /polyOver1P[q ->] := minPolyOver 1 u; apply/polyOver_poly=> j _. + by rewrite coef_map linearZZ rmorph1 rpredZ ?rpred1. + have [s /eqP Ds] := splitting_field_normal 1 u. + rewrite Ds; exists (map Qzt s); first by rewrite map_rp eqpxx. + apply/eqP; rewrite eqEsubv; apply/andP; split. + apply/Fadjoin_seqP; split=> // _ /mapP[w s_w ->]. + by rewrite (subvP (adjoinSl u_z (sub1v _))) // -sQof2 Dz QztE. + rewrite /= adjoinC (Fadjoin_idP _) -/Rz; last first. + by rewrite (subvP (adjoinSl _ (sub1v _))) // -sQof2 Dz Dit. + rewrite /= -adjoin_seq1 adjoin_seqSr //; apply/allP=> /=; rewrite andbT. + rewrite -(mem_map (fmorph_inj (ofQ _))) -map_comp (eq_map QztE); apply/mapP. + by exists u; rewrite ?inQ_K // -root_prod_XsubC -Ds root_minPoly. + have galCz: galois Cn Rz by rewrite (galoisS _ galRz) ?sRCn. + have [Cz | C'z]:= boolP (u_z \in Cn); first by exists n; rewrite -Dz -memCn. + pose G := 'Gal(Rz / Cn)%G; have{C'z} ntG: G :!=: 1%g. + rewrite trivg_card1 -galois_dim 1?(galoisS _ galCz) ?subvv //=. + by rewrite -adjoin_degreeE adjoin_deg_eq1. + pose extRz m := exists2 w, ofQ t w \in sQ (z_ m) & w \in [predD Rz & Cn]. + suffices [m le_n_m [w Cw /andP[C'w Rz_w]]]: exists2 m, (n <= m)%N & extRz m. + pose p := minPoly <<Cn; w>> u_z; apply: (IHd (p ^ ofQ t) m). + apply: leq_trans le_Rz_d; rewrite size_map_poly size_minPoly ltnS. + rewrite adjoin_degreeE adjoinC (addv_idPl Rz_w) agenv_id. + rewrite ltn_divLR ?adim_gt0 // mulnC. + rewrite muln_divCA ?field_dimS ?subv_adjoin // ltn_Pmulr ?adim_gt0 //. + by rewrite -adjoin_degreeE ltnNge leq_eqVlt orbF adjoin_deg_eq1. + rewrite map_monic monic_minPoly -Dz fmorph_root root_minPoly /=. + have /polyOverP Cw_p: p \is a polyOver <<Cn; w>>%VS by apply: minPolyOver. + apply/polyOver_poly=> j _; have /Fadjoin_polyP[q Cq {j}->] := Cw_p j. + rewrite -horner_map rpred_horner //; apply/polyOver_poly=> j _. + by rewrite (sCle n) // -memCn (polyOverP Cq). + have [evenG | oddG] := boolP (2.-group G); last first. + have [P /and3P[sPG evenP oddPG]] := Sylow_exists 2 'Gal(Rz / Rn). + have [w defQw] := PET_Qz t [aspace of fixedField P]. + pose pw := minPoly Rn w; pose p := (- pw * (pw \Po - 'X)) ^ ofQ t. + have sz_pw: (size pw).-1 = #|'Gal(Rz / Rn) : P|. + rewrite size_minPoly adjoin_degreeE -dim_fixed_galois //= -defQw. + congr (\dim_Rn _); apply/esym/eqP; rewrite eqEsubv adjoinSl ?sub1v //=. + by apply/FadjoinP; rewrite memv_adjoin /= defQw -galois_connection. + have mon_p: p \is monic. + have mon_pw: pw \is monic := monic_minPoly _ _. + rewrite map_monic mulNr -mulrN monicMl // monicE. + rewrite !(lead_coef_opp, lead_coef_comp) ?size_opp ?size_polyX //. + by rewrite lead_coefX sz_pw -signr_odd odd_2'nat oddPG mulrN1 opprK. + have Dp0: p.[0] = - ofQ t pw.[0] ^+ 2. + rewrite -(rmorph0 (ofQ t)) horner_map hornerM rmorphM. + by rewrite horner_comp !hornerN hornerX oppr0 rmorphN mulNr. + have Rpw: pw \is a polyOver Rn by apply: minPolyOver. + have Rp: p \is a polyOver (sQ (x_ n)). + apply/polyOver_poly=> j _; rewrite -memRn; apply: polyOverP j => /=. + by rewrite rpredM 1?polyOver_comp ?rpredN ?polyOverX. + have Rp0: ofQ t pw.[0] \in sQ (x_ n) by rewrite -memRn rpred_horner ?rpred0. + have [|{mon_p Rp Rp0 Dp0}m lenm p_Rm_0] := xRroot n p (ofQ t pw.[0]). + by rewrite /has_Rroot mon_p Rp Rp0 -Dp0 /=. + have{p_Rm_0} [y Ry pw_y]: {y | y \in sQ (x_ m) & root (pw ^ ofQ t) y}. + apply/sig2W; have [y Ry] := p_Rm_0. + rewrite [p]rmorphM /= map_comp_poly !rmorphN /= map_polyX. + rewrite rootM rootN root_comp hornerN hornerX. + by case/orP; [exists y | exists (- y)]; rewrite ?rpredN. + have [u Rz_u Dy]: exists2 u, u \in Rz & y = ofQ t u. + have Rz_w: w \in Rz by rewrite -sub_adjoin1v defQw capvSl. + have [sg [Gsg _ Dpw]] := galois_factors sRnRz galRz w Rz_w. + set s := map _ sg in Dpw. + have /mapP[u /mapP[g Gg Du] ->]: y \in map (ofQ t) s. + by rewrite -root_prod_XsubC -/(rp C _) -map_rp -[rp _ _]Dpw. + by exists u; rewrite // Du memv_gal. + have{pw_y} pw_u: root pw u by rewrite -(fmorph_root (ofQ t)) -Dy. + exists m => //; exists u; first by rewrite -Dy; apply: sQtrans Ry _. + rewrite inE /= Rz_u andbT; apply: contra oddG => Cu. + suffices: 2.-group 'Gal(Rz / Rn). + apply: pnat_dvd; rewrite -!galois_dim // ?(galoisS _ galQr) ?sRCz //. + rewrite dvdn_divLR ?field_dimS ?adim_gt0 //. + by rewrite mulnC muln_divCA ?field_dimS ?dvdn_mulr. + congr (2.-group _): evenP; apply/eqP. + rewrite eqEsubset sPG -indexg_eq1 (pnat_1 _ oddPG) // -sz_pw. + have (pu := minPoly Rn u): (pu %= pw) || (pu %= 1). + by rewrite minPoly_irr ?minPoly_dvdp ?minPolyOver. + rewrite /= -size_poly_eq1 {1}size_minPoly orbF => /eqp_size <-. + rewrite size_minPoly /= adjoin_degreeE (@pnat_dvd _ 2) // -dimCn. + rewrite dvdn_divLR ?divnK ?adim_gt0 ?field_dimS ?subv_adjoin //. + exact/FadjoinP. + have [w Rz_w deg_w]: exists2 w, w \in Rz & adjoin_degree Cn w = 2. + have [P sPG iPG]: exists2 P : {group gal_of Rz}, P \subset G & #|G : P| = 2. + have [_ _ [k oG]] := pgroup_pdiv evenG ntG. + have [P [sPG _ oP]] := normal_pgroup evenG (normal_refl G) (leq_pred _). + by exists P => //; rewrite -divgS // oP oG pfactorK // -expnB ?subSnn. + have [w defQw] := PET_Qz _ [aspace of fixedField P]. + exists w; first by rewrite -sub_adjoin1v defQw capvSl. + rewrite adjoin_degreeE -iPG -dim_fixed_galois // -defQw; congr (\dim_Cn _). + apply/esym/eqP; rewrite eqEsubv adjoinSl ?sub1v //=; apply/FadjoinP. + by rewrite memv_adjoin /= defQw -galois_connection. + have nz2: 2%:R != 0 :> Qt by move/charf0P: (charQ (Q t)) => ->. + without loss{deg_w} [C'w Cw2]: w Rz_w / w \notin Cn /\ w ^+ 2 \in Cn. + pose p := minPoly Cn w; pose v := p`_1 / 2%:R. + have /polyOverP Cp: p \is a polyOver Cn := minPolyOver Cn w. + have Cv: v \in Cn by rewrite rpred_div ?rpred_nat ?Cp. + move/(_ (v + w)); apply; first by rewrite rpredD // subvP_adjoin. + split; first by rewrite rpredDl // -adjoin_deg_eq1 deg_w. + rewrite addrC -[_ ^+ 2]subr0 -(rootP (root_minPoly Cn w)) -/p. + rewrite sqrrD [_ - _]addrAC rpredD ?rpredX // -mulr_natr -mulrA divfK //. + rewrite [w ^+ 2 + _]addrC mulrC -rpredN opprB horner_coef. + have /monicP := monic_minPoly Cn w; rewrite lead_coefE size_minPoly deg_w. + by rewrite 2!big_ord_recl big_ord1 => ->; rewrite mulr1 mul1r addrK Cp. + without loss R'w2: w Rz_w C'w Cw2 / w ^+ 2 \notin Rn. + move=> IHw; have [Rw2 | /IHw] := boolP (w ^+ 2 \in Rn); last exact. + have R'it: i_t \notin Rn by rewrite memRn Dit. + pose v := 1 + i_t; have R'v: v \notin Rn by rewrite rpredDl ?rpred1. + have Cv: v \in Cn by rewrite rpredD ?rpred1 ?memv_adjoin. + have nz_v: v != 0 by rewrite (memPnC R'v) ?rpred0. + apply: (IHw (v * w)); last 1 [|] || by rewrite fpredMl // subvP_adjoin. + by rewrite exprMn rpredM // rpredX. + rewrite exprMn fpredMr //=; last by rewrite expf_eq0 (memPnC C'w) ?rpred0. + by rewrite sqrrD Dit2 expr1n addrC addKr -mulrnAl fpredMl ?rpred_nat. + pose rect_w2 u v := [/\ u \in Rn, v \in Rn & u + i_t * (v * 2%:R) = w ^+ 2]. + have{Cw2} [u [v [Ru Rv Dw2]]]: {u : Qt & {v | rect_w2 u v}}. + rewrite /rect_w2 -(Fadjoin_poly_eq Cw2); set p := Fadjoin_poly Rn i_t _. + have /polyOverP Rp: p \is a polyOver Rn by apply: Fadjoin_polyOver. + exists p`_0, (p`_1 / 2%:R); split; rewrite ?rpred_div ?rpred_nat //. + rewrite divfK // (horner_coef_wide _ (size_Fadjoin_poly _ _ _)) -/p. + by rewrite adjoin_degreeE dimCn big_ord_recl big_ord1 mulr1 mulrC. + pose p := Poly [:: - (ofQ t v ^+ 2); 0; - ofQ t u; 0; 1]. + have [|m lenm [x Rx px0]] := xRroot n p (ofQ t v). + rewrite /has_Rroot 2!unfold_in lead_coefE horner_coef0 -memRn Rv. + rewrite (@PolyK _ 1) ?oner_eq0 //= !eqxx !rpred0 ?rpred1 ?rpredN //=. + by rewrite !andbT rpredX -memRn. + suffices [y Cy Dy2]: {y | y \in sQ (z_ m) & ofQ t w ^+ 2 == y ^+ 2}. + exists m => //; exists w; last by rewrite inE C'w. + by move: Dy2; rewrite eqf_sqr => /pred2P[]->; rewrite ?rpredN. + exists (x + i * (ofQ t v / x)). + rewrite rpredD 1?rpredM ?rpred_div //= (sQtrans (x_ m)) //. + by rewrite (sRle n) // -memRn. + rewrite rootE /horner (@PolyK _ 1) ?oner_eq0 //= ?addr0 ?mul0r in px0. + rewrite add0r mul1r -mulrA -expr2 subr_eq0 in px0. + have nz_x2: x ^+ 2 != 0. + apply: contraNneq R'w2 => y2_0; rewrite -Dw2 mulrCA. + suffices /eqP->: v == 0 by rewrite mul0r addr0. + by rewrite y2_0 mulr0 eq_sym sqrf_eq0 fmorph_eq0 in px0. + apply/eqP/esym/(mulIf nz_x2); rewrite -exprMn -rmorphX -Dw2 rmorphD rmorphM. + rewrite Dit mulrDl -expr2 mulrA divfK; last by rewrite expf_eq0 in nz_x2. + rewrite mulr_natr addrC sqrrD exprMn Di2 mulN1r -(eqP px0) -mulNr opprB. + by rewrite -mulrnAl -mulrnAr -rmorphMn -!mulrDl addrAC subrK. +have inFTA n z: (n_ z <= n)%N -> z = ofQ (z_ n) (inQ (z_ n) z). + by move/sCle=> le_zn; rewrite inQ_K ?le_zn. +pose is_cj n cj := {in R_ n, cj =1 id} /\ cj (i_ n) = - i_ n. +have /all_sig[cj_ /all_and2[cj_R cj_i]] n: {cj : 'AEnd(Q (z_ n)) | is_cj n cj}. + have cj_P: root (minPoly (R_ n) (i_ n) ^ \1%VF) (- i_ n). + rewrite minp_i -(fmorph_root (ofQ _)) !rmorphD !rmorph1 /= !map_polyXn. + by rewrite rmorphN inQ_K // rootE hornerD hornerXn hornerC sqrrN Di2 addNr. + have cj_M: ahom_in fullv (kHomExtend (R_ n) \1 (i_ n) (- i_ n)). + by rewrite -defRi -k1HomE kHomExtendP ?sub1v ?kHom1. + exists (AHom cj_M); split=> [y /kHomExtend_id->|]; first by rewrite ?id_lfunE. + by rewrite (kHomExtend_val (kHom1 1 _)). +pose conj_ n z := ofQ _ (cj_ n (inQ _ z)); pose conj z := conj_ (n_ z) z. +have conjK n m z: (n_ z <= n)%N -> (n <= m)%N -> conj_ m (conj_ n z) = z. + move/sCle=> le_z_n le_n_m; have /le_z_n/sQ_inQ[u <-] := FTA z. + have /QtoQ[Qmn QmnE]: z_ n \in sQ (z_ m) by rewrite (sCle n). + rewrite /conj_ ofQ_K -!QmnE !ofQ_K -!comp_lfunE; congr (ofQ _ _). + move: u (memRi n u); apply/eqlfun_inP/FadjoinP; split=> /=. + apply/eqlfun_inP=> y Ry; rewrite !comp_lfunE !cj_R //. + by move: Ry; rewrite -!sQof2 QmnE !inQ_K //; apply: sRle. + apply/eqlfunP; rewrite !comp_lfunE cj_i !linearN /=. + suffices ->: Qmn (i_ n) = i_ m by rewrite cj_i ?opprK. + by apply: (fmorph_inj (ofQ _)); rewrite QmnE !inQ_K. +have conjE n z: (n_ z <= n)%N -> conj z = conj_ n z. + move/leq_trans=> le_zn; set x := conj z; set y := conj_ n z. + have [m [le_xm le_ym le_nm]] := maxn3 (n_ x) (n_ y) n. + by have /conjK/=/can_in_inj := leqnn m; apply; rewrite ?conjK // le_zn. +suffices conjM: rmorphism conj. + exists (RMorphism conjM) => [z | /(_ i)/eqP/idPn[]] /=. + by have [n [/conjE-> /(conjK (n_ z))->]] := maxn3 (n_ (conj z)) (n_ z) 0%N. + rewrite /conj/conj_ cj_i rmorphN inQ_K // eq_sym -addr_eq0 -mulr2n -mulr_natl. + rewrite mulf_neq0 ?(memPnC (R'i 0%N)) ?rpred0 //. + by have /charf0P-> := ftrans (fmorph_char QtoC) (char_num _). +do 2?split=> [x y|]; last pose n1 := n_ 1. +- have [m [le_xm le_ym le_xym]] := maxn3 (n_ x) (n_ y) (n_ (x - y)). + by rewrite !(conjE m) // (inFTA m x) // (inFTA m y) -?rmorphB /conj_ ?ofQ_K. +- have [m [le_xm le_ym le_xym]] := maxn3 (n_ x) (n_ y) (n_ (x * y)). + by rewrite !(conjE m) // (inFTA m x) // (inFTA m y) -?rmorphM /conj_ ?ofQ_K. +by rewrite /conj -/n1 -(rmorph1 (ofQ (z_ n1))) /conj_ ofQ_K !rmorph1. +Qed. diff --git a/mathcomp/field/algnum.v b/mathcomp/field/algnum.v new file mode 100644 index 0000000..cff3197 --- /dev/null +++ b/mathcomp/field/algnum.v @@ -0,0 +1,835 @@ +(* (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 finalg zmodp poly. +Require Import ssrnum ssrint rat polydiv intdiv algC matrix mxalgebra mxpoly. +Require Import vector falgebra fieldext separable galois cyclotomic. + +(******************************************************************************) +(* This file provides a few basic results and constructions in algebraic *) +(* number theory, that are used in the character theory library. Most of *) +(* these could be generalized to a more abstract setting. Note that the type *) +(* of abstract number fields is simply extFieldType rat. We define here: *) +(* x \in Crat_span X <=> x is a Q-linear combination of elements of *) +(* X : seq algC. *) +(* x \in Cint_span X <=> x is a Z-linear combination of elements of *) +(* X : seq algC. *) +(* x \in Aint <=> x : algC is an algebraic integer, i.e., the (monic) *) +(* polynomial of x over Q has integer coeficients. *) +(* (e %| a)%A <=> e divides a with respect to algebraic integers, *) +(* (e %| a)%Ax i.e., a is in the algebraic integer ideal generated *) +(* by e. This is is notation for a \in dvdA e, where *) +(* dvdv is the (collective) predicate for the Aint *) +(* ideal generated by e. As in the (e %| a)%C notation *) +(* e and a can be coerced to algC from nat or int. *) +(* The (e %| a)%Ax display form is a workaround for *) +(* design limitations of the Coq Notation facilities. *) +(* (a == b %[mod e])%A, (a != b %[mod e])%A <=> *) +(* a is equal (resp. not equal) to b mod e, i.e., a and *) +(* b belong to the same e * Aint class. We do not *) +(* force a, b and e to be algebraic integers. *) +(* #[x]%C == the multiplicative order of x, i.e., the n such that *) +(* x is an nth primitive root of unity, or 0 if x is not *) +(* a root of unity. *) +(* In addition several lemmas prove the (constructive) existence of number *) +(* fields and of automorphisms of algC. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Local Notation ZtoQ := (intr : int -> rat). +Local Notation ZtoC := (intr : int -> algC). +Local Notation QtoC := (ratr : rat -> algC). + +Local Notation intrp := (map_poly intr). +Local Notation pZtoQ := (map_poly ZtoQ). +Local Notation pZtoC := (map_poly ZtoC). +Local Notation pQtoC := (map_poly ratr). + +Local Hint Resolve (@intr_inj _ : injective ZtoC). +Local Notation QtoCm := [rmorphism of QtoC]. + +(* Number fields and rational spans. *) +Lemma algC_PET (s : seq algC) : + {z | exists a : nat ^ size s, z = \sum_(i < size s) s`_i *+ a i + & exists ps, s = [seq (pQtoC p).[z] | p <- ps]}. +Proof. +elim: s => [|x s [z /sig_eqW[a Dz] /sig_eqW[ps Ds]]]. + by exists 0; [exists [ffun _ => 2]; rewrite big_ord0 | exists nil]. +have r_exists (y : algC): {r | r != 0 & root (pQtoC r) y}. + have [r [_ mon_r] dv_r] := minCpolyP y. + by exists r; rewrite ?monic_neq0 ?dv_r. +suffices /sig_eqW[[n [|px [|pz []]]]// [Dpx Dpz]]: + exists np, let zn := x *+ np.1 + z in + [:: x; z] = [seq (pQtoC p).[zn] | p <- np.2]. +- exists (x *+ n + z). + exists [ffun i => oapp a n (unlift ord0 i)]. + rewrite /= big_ord_recl ffunE unlift_none Dz; congr (_ + _). + by apply: eq_bigr => i _; rewrite ffunE liftK. + exists (px :: [seq p \Po pz | p <- ps]); rewrite /= -Dpx; congr (_ :: _). + rewrite -map_comp Ds; apply: eq_map => p /=. + by rewrite map_comp_poly horner_comp -Dpz. +have [rx nz_rx rx0] := r_exists x. +have [rz nz_rz rz0] := r_exists (- z). +have char0_Q: [char rat] =i pred0 by exact: char_num. +have [n [[pz Dpz] [px Dpx]]] := char0_PET nz_rz rz0 nz_rx rx0 char0_Q. +by exists (n, [:: px; - pz]); rewrite /= !raddfN hornerN -[z]opprK Dpz Dpx. +Qed. + +Canonical subfx_unitAlgType (F L : fieldType) iota (z : L) p := + Eval hnf in [unitAlgType F of subFExtend iota z p]. + +Lemma num_field_exists (s : seq algC) : + {Qs : fieldExtType rat & {QsC : {rmorphism Qs -> algC} + & {s1 : seq Qs | map QsC s1 = s & <<1 & s1>>%VS = fullv}}}. +Proof. +have [z /sig_eqW[a Dz] /sig_eqW[ps Ds]] := algC_PET s. +suffices [Qs [QsC [z1 z1C z1gen]]]: + {Qs : fieldExtType rat & {QsC : {rmorphism Qs -> algC} & + {z1 : Qs | QsC z1 = z & forall xx, exists p, fieldExt_horner z1 p = xx}}}. +- set inQs := fieldExt_horner z1 in z1gen *; pose s1 := map inQs ps. + have inQsK p: QsC (inQs p) = (pQtoC p).[z]. + rewrite /= -horner_map z1C -map_poly_comp; congr _.[z]. + apply: eq_map_poly => b /=; apply: canRL (mulfK _) _. + by rewrite intr_eq0 denq_eq0. + rewrite /= mulrzr -rmorphMz scalerMzl -{1}[b]divq_num_den -mulrzr. + by rewrite divfK ?intr_eq0 ?denq_eq0 // scaler_int rmorph_int. + exists Qs, QsC, s1; first by rewrite -map_comp Ds (eq_map inQsK). + have sz_ps: size ps = size s by rewrite Ds size_map. + apply/vspaceP=> x; rewrite memvf; have [p {x}<-] := z1gen x. + elim/poly_ind: p => [|p b ApQs]; first by rewrite /inQs rmorph0 mem0v. + rewrite /inQs rmorphD rmorphM /= fieldExt_hornerX fieldExt_hornerC -/inQs /=. + suffices ->: z1 = \sum_(i < size s) s1`_i *+ a i. + rewrite memvD ?memvZ ?mem1v ?memvM ?memv_suml // => i _. + by rewrite rpredMn ?seqv_sub_adjoin ?mem_nth // size_map sz_ps. + apply: (fmorph_inj QsC); rewrite z1C Dz rmorph_sum; apply: eq_bigr => i _. + by rewrite rmorphMn {1}Ds !(nth_map 0) ?sz_ps //= inQsK. +have [r [Dr /monic_neq0 nz_r] dv_r] := minCpolyP z. +have rz0: root (pQtoC r) z by rewrite dv_r. +have irr_r: irreducible_poly r. + by apply/(subfx_irreducibleP rz0 nz_r)=> q qz0 nzq; rewrite dvdp_leq // -dv_r. +exists (SubFieldExtType rz0 irr_r), (subfx_inj_rmorphism QtoCm z r). +exists (subfx_root _ z r) => [|x]; first exact: subfx_inj_root. +by have{x} [p ->] := subfxEroot rz0 nz_r x; exists p. +Qed. + +Definition in_Crat_span s x := + exists a : rat ^ size s, x = \sum_i QtoC (a i) * s`_i. + +Fact Crat_span_subproof s x : decidable (in_Crat_span s x). +Proof. +have [Qxs [QxsC [[|x1 s1] // [<- <-] {x s} _]]] := num_field_exists (x :: s). +have QxsC_Z a zz: QxsC (a *: zz) = QtoC a * QxsC zz. + rewrite mulrAC; apply: (canRL (mulfK _)); first by rewrite intr_eq0 denq_eq0. + by rewrite mulrzr mulrzl -!rmorphMz scalerMzl -mulrzr -numqE scaler_int. +apply: decP (x1 \in <<in_tuple s1>>%VS) _; rewrite /in_Crat_span size_map. +apply: (iffP idP) => [/coord_span-> | [a Dx]]. + move: (coord _) => a; exists [ffun i => a i x1]; rewrite rmorph_sum. + by apply: eq_bigr => i _; rewrite ffunE (nth_map 0). +have{Dx} ->: x1 = \sum_i a i *: s1`_i. + apply: (fmorph_inj QxsC); rewrite Dx rmorph_sum. + by apply: eq_bigr => i _; rewrite QxsC_Z (nth_map 0). +by apply: memv_suml => i _; rewrite memvZ ?memv_span ?mem_nth. +Qed. + +Definition Crat_span s : pred algC := Crat_span_subproof s. +Lemma Crat_spanP s x : reflect (in_Crat_span s x) (x \in Crat_span s). +Proof. exact: sumboolP. Qed. +Fact Crat_span_key s : pred_key (Crat_span s). Proof. by []. Qed. +Canonical Crat_span_keyed s := KeyedPred (Crat_span_key s). + +Lemma mem_Crat_span s : {subset s <= Crat_span s}. +Proof. +move=> _ /(nthP 0)[ix ltxs <-]; pose i0 := Ordinal ltxs. +apply/Crat_spanP; exists [ffun i => (i == i0)%:R]. +rewrite (bigD1 i0) //= ffunE eqxx // rmorph1 mul1r. +by rewrite big1 ?addr0 // => i; rewrite ffunE rmorph_nat mulr_natl => /negbTE->. +Qed. + +Fact Crat_span_zmod_closed s : zmod_closed (Crat_span s). +Proof. +split=> [|_ _ /Crat_spanP[x ->] /Crat_spanP[y ->]]. + apply/Crat_spanP; exists 0. + by apply/esym/big1=> i _; rewrite ffunE rmorph0 mul0r. +apply/Crat_spanP; exists (x - y); rewrite -sumrB; apply: eq_bigr => i _. +by rewrite -mulrBl -rmorphB !ffunE. +Qed. +Canonical Crat_span_opprPred s := OpprPred (Crat_span_zmod_closed s). +Canonical Crat_span_addrPred s := AddrPred (Crat_span_zmod_closed s). +Canonical Crat_span_zmodPred s := ZmodPred (Crat_span_zmod_closed s). + +Section MoreAlgCaut. + +Implicit Type rR : unitRingType. + +Lemma alg_num_field (Qz : fieldExtType rat) a : a%:A = ratr a :> Qz. +Proof. by rewrite -in_algE fmorph_eq_rat. Qed. + +Lemma rmorphZ_num (Qz : fieldExtType rat) rR (f : {rmorphism Qz -> rR}) a x : + f (a *: x) = ratr a * f x. +Proof. by rewrite -mulr_algl rmorphM alg_num_field fmorph_rat. Qed. + +Lemma fmorph_numZ (Qz1 Qz2 : fieldExtType rat) (f : {rmorphism Qz1 -> Qz2}) : + scalable f. +Proof. by move=> a x; rewrite rmorphZ_num -alg_num_field mulr_algl. Qed. +Definition NumLRmorphism Qz1 Qz2 f := AddLRMorphism (@fmorph_numZ Qz1 Qz2 f). + +End MoreAlgCaut. + +Section NumFieldProj. + +Variables (Qn : fieldExtType rat) (QnC : {rmorphism Qn -> algC}). + +Lemma Crat_spanZ b a : {in Crat_span b, forall x, ratr a * x \in Crat_span b}. +Proof. +move=> _ /Crat_spanP[a1 ->]; apply/Crat_spanP; exists [ffun i => a * a1 i]. +by rewrite mulr_sumr; apply: eq_bigr => i _; rewrite ffunE mulrA -rmorphM. +Qed. + +Lemma Crat_spanM b : {in Crat & Crat_span b, forall a x, a * x \in Crat_span b}. +Proof. by move=> _ x /CratP[a ->]; exact: Crat_spanZ. Qed. + +(* In principle CtoQn could be taken to be additive and Q-linear, but this *) +(* would require a limit construction. *) +Lemma num_field_proj : {CtoQn | CtoQn 0 = 0 & cancel QnC CtoQn}. +Proof. +pose b := vbasis {:Qn}. +have Qn_bC (u : {x | x \in Crat_span (map QnC b)}): {y | QnC y = sval u}. + case: u => _ /= /Crat_spanP/sig_eqW[a ->]. + exists (\sum_i a i *: b`_i); rewrite rmorph_sum; apply: eq_bigr => i _. + by rewrite rmorphZ_num (nth_map 0) // -(size_map QnC). +pose CtoQn x := oapp (fun u => sval (Qn_bC u)) 0 (insub x). +suffices QnCK: cancel QnC CtoQn by exists CtoQn; rewrite // -(rmorph0 QnC). +move=> x; rewrite /CtoQn insubT => /= [|Qn_x]; last first. + by case: (Qn_bC _) => x1 /= /fmorph_inj. +rewrite (coord_vbasis (memvf x)) rmorph_sum rpred_sum // => i _. +rewrite rmorphZ_num Crat_spanZ ?mem_Crat_span // -/b. +by rewrite -tnth_nth -tnth_map mem_tnth. +Qed. + +Lemma restrict_aut_to_num_field (nu : {rmorphism algC -> algC}) : + (forall x, exists y, nu (QnC x) = QnC y) -> + {nu0 : {lrmorphism Qn -> Qn} | {morph QnC : x / nu0 x >-> nu x}}. +Proof. +move=> Qn_nu; pose nu0 x := sval (sig_eqW (Qn_nu x)). +have QnC_nu0: {morph QnC : x / nu0 x >-> nu x}. + by rewrite /nu0 => x; case: (sig_eqW _). +suffices nu0M: rmorphism nu0 by exists (NumLRmorphism (RMorphism nu0M)). +do 2?split=> [x y|]; apply: (fmorph_inj QnC); rewrite ?QnC_nu0 ?rmorph1 //. + by rewrite ?(rmorphB, QnC_nu0). +by rewrite ?(rmorphM, QnC_nu0). +Qed. + +Lemma map_Qnum_poly (nu : {rmorphism algC -> algC}) p : + p \in polyOver 1%VS -> map_poly (nu \o QnC) p = (map_poly QnC p). +Proof. +move=> Qp; apply/polyP=> i; rewrite /= !coef_map /=. +have /vlineP[a ->]: p`_i \in 1%VS by exact: polyOverP. +by rewrite alg_num_field !fmorph_rat. +Qed. + +End NumFieldProj. + +Lemma restrict_aut_to_normal_num_field (Qn : splittingFieldType rat) + (QnC : {rmorphism Qn -> algC})(nu : {rmorphism algC -> algC}) : + {nu0 : {lrmorphism Qn -> Qn} | {morph QnC : x / nu0 x >-> nu x}}. +Proof. +apply: restrict_aut_to_num_field => x. +case: (splitting_field_normal 1%AS x) => rs /eqP Hrs. +have: root (map_poly (nu \o QnC) (minPoly 1%AS x)) (nu (QnC x)). + by rewrite fmorph_root root_minPoly. +rewrite map_Qnum_poly ?minPolyOver // Hrs. +rewrite [map_poly _ _](_:_ = \prod_(y <- map QnC rs) ('X - y%:P)); last first. + rewrite big_map rmorph_prod; apply eq_bigr => i _. + by rewrite rmorphB /= map_polyX map_polyC. +rewrite root_prod_XsubC. +by case/mapP => y _ ?; exists y. +Qed. + +(* Integral spans. *) + +Lemma dec_Cint_span (V : vectType algC) m (s : m.-tuple V) 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 IzT := {: 'I_m * 'I_(\dim <<s>>)}; pose Iz := 'I_#|IzT|. +pose b := vbasis <<s>>. +pose z_s := [seq coord b ij.2 (tnth s ij.1) | ij : IzT]. +pose rank2 j i: Iz := enum_rank (i, j); pose val21 (p : Iz) := (enum_val p).1. +pose inQzs w := [forall j, Crat_span z_s (coord b j w)]. +have enum_pairK j: {in predT, cancel (rank2 j) val21}. + by move=> i; rewrite /val21 enum_rankK. +have Qz_Zs a: inQzs (\sum_(i < m) s`_i *~ a i). + apply/forallP=> j; apply/Crat_spanP; rewrite /in_Crat_span size_map -cardE. + exists [ffun ij => (a (val21 ij))%:Q *+ ((enum_val ij).2 == j)]. + rewrite linear_sum {1}(reindex_onto _ _ (enum_pairK j)). + rewrite big_mkcond; apply: eq_bigr => ij _ /=; rewrite nth_image (tnth_nth 0). + rewrite (can2_eq (@enum_rankK _) (@enum_valK _)) ffunE -scaler_int /val21. + case Dij: (enum_val ij) => [i j1]; rewrite xpair_eqE eqxx /= eq_sym -mulrb. + by rewrite linearZ rmorphMn rmorph_int mulrnAl; case: eqP => // ->. +case Qz_v: (inQzs v); last by right=> [[a Dv]]; rewrite Dv Qz_Zs in Qz_v. +have [Qz [QzC [z1s Dz_s _]]] := num_field_exists z_s. +have sz_z1s: size z1s = #|IzT| by rewrite -(size_map QzC) Dz_s size_map cardE. +have xv j: {x | coord b j v = QzC x}. + apply: sig_eqW; have /Crat_spanP[x ->] := forallP Qz_v j. + exists (\sum_ij x ij *: z1s`_ij); rewrite rmorph_sum. + apply: eq_bigr => ij _; rewrite mulrAC. + apply: canLR (mulfK _) _; first by rewrite intr_eq0 denq_neq0. + rewrite mulrzr -rmorphMz scalerMzl -(mulrzr (x _)) -numqE scaler_int. + by rewrite rmorphMz mulrzl -(nth_map _ 0) ?Dz_s // -(size_map QzC) Dz_s. +pose sz := [tuple [ffun j => z1s`_(rank2 j i)] | i < m]. +have [Zsv | Zs'v] := dec_Qint_span sz [ffun j => sval (xv j)]. + left; have{Zsv} [a Dv] := Zsv; exists a. + transitivity (\sum_j \sum_(i < m) QzC ((sz`_i *~ a i) j) *: b`_j). + rewrite {1}(coord_vbasis s_v) -/b; apply: eq_bigr => j _. + rewrite -scaler_suml; congr (_ *: _). + have{Dv} /ffunP/(_ j) := Dv; rewrite sum_ffunE !ffunE -rmorph_sum => <-. + by case: (xv j). + rewrite exchange_big; apply: eq_bigr => i _. + rewrite (coord_vbasis (s_s i)) -/b mulrz_suml; apply: eq_bigr => j _. + rewrite scalerMzl ffunMzE rmorphMz; congr ((_ *~ _) *: _). + rewrite nth_mktuple ffunE -(nth_map _ 0) ?sz_z1s // Dz_s. + by rewrite nth_image enum_rankK /= (tnth_nth 0). +right=> [[a Dv]]; case: Zs'v; exists a. +apply/ffunP=> j; rewrite sum_ffunE !ffunE; apply: (fmorph_inj QzC). +case: (xv j) => /= _ <-; rewrite Dv linear_sum rmorph_sum. +apply: eq_bigr => i _; rewrite nth_mktuple raddfMz !ffunMzE rmorphMz ffunE. +by rewrite -(nth_map _ 0 QzC) ?sz_z1s // Dz_s nth_image enum_rankK -tnth_nth. +Qed. + +Definition Cint_span (s : seq algC) : pred algC := + fun x => dec_Cint_span (in_tuple [seq \row_(i < 1) y | y <- s]) (\row_i x). +Fact Cint_span_key s : pred_key (Cint_span s). Proof. by []. Qed. +Canonical Cint_span_keyed s := KeyedPred (Cint_span_key s). + +Lemma Cint_spanP n (s : n.-tuple algC) x : + reflect (inIntSpan s x) (x \in Cint_span s). +Proof. +rewrite unfold_in; case: (dec_Cint_span _ _) => [Zs_x | Zs'x] /=. + left; have{Zs_x} [] := Zs_x; rewrite /= size_map size_tuple => a /rowP/(_ 0). + rewrite !mxE => ->; exists a; rewrite summxE; apply: eq_bigr => i _. + by rewrite -scaler_int (nth_map 0) ?size_tuple // !mxE mulrzl. +right=> [[a Dx]]; have{Zs'x} [] := Zs'x. +rewrite /inIntSpan /= size_map size_tuple; exists a. +apply/rowP=> i0; rewrite !mxE summxE Dx; apply: eq_bigr => i _. +by rewrite -scaler_int mxE mulrzl (nth_map 0) ?size_tuple // !mxE. +Qed. + +Lemma mem_Cint_span s : {subset s <= Cint_span s}. +Proof. +move=> _ /(nthP 0)[ix ltxs <-]; apply/(Cint_spanP (in_tuple s)). +exists [ffun i => i == Ordinal ltxs : int]. +rewrite (bigD1 (Ordinal ltxs)) //= ffunE eqxx. +by rewrite big1 ?addr0 // => i; rewrite ffunE => /negbTE->. +Qed. + +Lemma Cint_span_zmod_closed s : zmod_closed (Cint_span s). +Proof. +have sP := Cint_spanP (in_tuple s); split=> [|_ _ /sP[x ->] /sP[y ->]]. + by apply/sP; exists 0; rewrite big1 // => i; rewrite ffunE. +apply/sP; exists (x - y); rewrite -sumrB; apply: eq_bigr => i _. +by rewrite !ffunE raddfB. +Qed. +Canonical Cint_span_opprPred s := OpprPred (Cint_span_zmod_closed s). +Canonical Cint_span_addrPred s := AddrPred (Cint_span_zmod_closed s). +Canonical Cint_span_zmodPred s := ZmodPred (Cint_span_zmod_closed s). + +(* Automorphism extensions. *) +Lemma extend_algC_subfield_aut (Qs : fieldExtType rat) + (QsC : {rmorphism Qs -> algC}) (phi : {rmorphism Qs -> Qs}) : + {nu : {rmorphism algC -> algC} | {morph QsC : x / phi x >-> nu x}}. +Proof. +pose numF_inj (Qr : fieldExtType rat) := {rmorphism Qr -> algC}. +pose subAut := {Qr : _ & numF_inj Qr * {lrmorphism Qr -> Qr}}%type. +pose SubAut := existS _ _ (_, _) : subAut. +pose Sdom (mu : subAut) := projS1 mu. +pose Sinj (mu : subAut) : {rmorphism Sdom mu -> algC} := (projS2 mu).1. +pose Saut (mu : subAut) : {rmorphism Sdom mu -> Sdom mu} := (projS2 mu).2. +have SinjZ Qr (QrC : numF_inj Qr) a x: QrC (a *: x) = QtoC a * QrC x. + rewrite mulrAC; apply: canRL (mulfK _) _. + by rewrite intr_eq0 denq_neq0. + by rewrite mulrzr mulrzl -!rmorphMz scalerMzl -scaler_int -mulrzr -numqE. +have Sinj_poly Qr (QrC : numF_inj Qr) p: + map_poly QrC (map_poly (in_alg Qr) p) = pQtoC p. +- rewrite -map_poly_comp; apply: eq_map_poly => a. + by rewrite /= SinjZ rmorph1 mulr1. +have ext1 mu0 x: {mu1 | exists y, x = Sinj mu1 y + & exists2 in01 : {lrmorphism _}, Sinj mu0 =1 Sinj mu1 \o in01 + & {morph in01: y / Saut mu0 y >-> Saut mu1 y}}. +- pose b0 := vbasis {:Sdom mu0}. + have [z _ /sig_eqW[[|px ps] // [Dx Ds]]] := algC_PET (x :: map (Sinj mu0) b0). + have [p [_ mon_p] /(_ p) pz0] := minCpolyP z; rewrite dvdpp in pz0. + have [r Dr] := closed_field_poly_normal (pQtoC p : {poly algC}). + rewrite lead_coef_map {mon_p}(monicP mon_p) rmorph1 scale1r in Dr. + have{pz0} rz: z \in r by rewrite -root_prod_XsubC -Dr. + have [Qr [QrC [rr Drr genQr]]] := num_field_exists r. + have{rz} [zz Dz]: {zz | QrC zz = z}. + by move: rz; rewrite -Drr => /mapP/sig2_eqW[zz]; exists zz. + have{ps Ds} [in01 Din01]: {in01 : {lrmorphism _} | Sinj mu0 =1 QrC \o in01}. + have in01P y: {yy | Sinj mu0 y = QrC yy}. + exists (\sum_i coord b0 i y *: (map_poly (in_alg Qr) ps`_i).[zz]). + rewrite {1}(coord_vbasis (memvf y)) !rmorph_sum; apply: eq_bigr => i _. + rewrite !SinjZ; congr (_ * _); rewrite -(nth_map _ 0) ?size_tuple // Ds. + rewrite -horner_map Dz Sinj_poly (nth_map 0) //. + by have:= congr1 size Ds; rewrite !size_map size_tuple => <-. + pose in01 y := sval (in01P y). + have Din01 y: Sinj mu0 y = QrC (in01 y) by rewrite /in01; case: (in01P y). + suffices in01M: lrmorphism in01 by exists (LRMorphism in01M). + pose rwM := (=^~ Din01, SinjZ, rmorph1, rmorphB, rmorphM). + by do 3?split; try move=> ? ?; apply: (fmorph_inj QrC); rewrite !rwM. + have {z zz Dz px Dx} Dx: exists xx, x = QrC xx. + exists (map_poly (in_alg Qr) px).[zz]. + by rewrite -horner_map Dz Sinj_poly Dx. + pose lin01 := linfun in01; pose K := (lin01 @: fullv)%VS. + have memK y: reflect (exists yy, y = in01 yy) (y \in K). + apply: (iffP memv_imgP) => [[yy _ ->] | [yy ->]]; + by exists yy; rewrite ?lfunE ?memvf. + have algK: is_aspace K. + rewrite /is_aspace has_algid1; last first. + by apply/memK; exists 1; rewrite rmorph1. + apply/prodvP=> _ _ /memK[y1 ->] /memK[y2 ->]. + by apply/memK; exists (y1 * y2); rewrite rmorphM. + have ker_in01: lker lin01 == 0%VS. + by apply/lker0P=> y1 y2; rewrite !lfunE; apply: fmorph_inj. + pose f := (lin01 \o linfun (Saut mu0) \o lin01^-1)%VF. + have Df y: f (in01 y) = in01 (Saut mu0 y). + transitivity (f (lin01 y)); first by rewrite !lfunE. + by do 4!rewrite lfunE /=; rewrite lker0_lfunK. + have hom_f: kHom 1 (ASpace algK) f. + apply/kHomP; split=> [_ _ /memK[y1 ->] /memK[y2 ->] |_ /vlineP[a ->]]. + by rewrite -rmorphM !Df !rmorphM. + by rewrite -(rmorph1 in01) -linearZ /= Df {1}linearZ /= rmorph1. + pose pr := map_poly (in_alg Qr) p. + have Qpr: pr \is a polyOver 1%VS. + by apply/polyOverP=> i; rewrite coef_map memvZ ?memv_line. + have splitQr: splittingFieldFor K pr fullv. + apply: splittingFieldForS (sub1v (Sub K algK)) (subvf _) _; exists rr => //. + congr (_ %= _): (eqpxx pr); apply: (@map_poly_inj _ _ QrC). + rewrite Sinj_poly Dr -Drr big_map rmorph_prod; apply: eq_bigr => zz _. + by rewrite rmorphB /= map_polyX map_polyC. + have [f1 aut_f1 Df1]:= kHom_extends (sub1v (ASpace algK)) hom_f Qpr splitQr. + pose nu := LRMorphism (kHom_lrmorphism aut_f1). + exists (SubAut Qr QrC nu) => //; exists in01 => //= y. + by rewrite -Df -Df1 //; apply/memK; exists y. +have phiZ: scalable phi. + move=> a y; do 2!rewrite -mulr_algl -in_algE. + by rewrite -[a]divq_num_den !(fmorph_div, rmorphM, rmorph_int). +pose fix ext n := + if n is i.+1 then oapp (fun x => s2val (ext1 (ext i) x)) (ext i) (unpickle i) + else SubAut Qs QsC (AddLRMorphism phiZ). +have mem_ext x n: (pickle x < n)%N -> {xx | Sinj (ext n) xx = x}. + move=> ltxn; apply: sig_eqW; elim: n ltxn => // n IHn. + rewrite ltnS leq_eqVlt => /predU1P[<- | /IHn[xx <-]] /=. + by rewrite pickleK /=; case: (ext1 _ x) => mu [xx]; exists xx. + case: (unpickle n) => /= [y|]; last by exists xx. + case: (ext1 _ y) => mu /= _ [in_mu inj_in_mu _]. + by exists (in_mu xx); rewrite inj_in_mu. +pose nu x := Sinj _ (Saut _ (sval (mem_ext x _ (ltnSn _)))). +have nu_inj n y: nu (Sinj (ext n) y) = Sinj (ext n) (Saut (ext n) y). + rewrite /nu; case: (mem_ext _ _ _); move: _.+1 => n1 y1 Dy /=. + without loss /subnK Dn1: n n1 y y1 Dy / (n <= n1)%N. + by move=> IH; case/orP: (leq_total n n1) => /IH => [/(_ y) | /(_ y1)]->. + elim: {n}(_ - n)%N {-1}n => [|k IHk] n in Dn1 y Dy *. + by move: y1 Dy; rewrite -Dn1 => y1 /fmorph_inj ->. + rewrite addSnnS in Dn1; move/IHk: Dn1 => /=. + case: (unpickle _) => [z|] /=; last exact. + case: (ext1 _ _) => mu /= _ [in_mu Dinj Daut]. + by rewrite Dy => /(_ _ (Dinj _))->; rewrite -Daut Dinj. +suffices nuM: rmorphism nu. + by exists (RMorphism nuM) => x; rewrite /= (nu_inj 0%N). +pose le_nu (x : algC) n := (pickle x < n)%N. +have max3 x1 x2 x3: exists n, [/\ le_nu x1 n, le_nu x2 n & le_nu x3 n]. + exists (maxn (pickle x1) (maxn (pickle x2) (pickle x3))).+1. + by apply/and3P; rewrite /le_nu !ltnS -!geq_max. +do 2?split; try move=> x1 x2. +- have [n] := max3 (x1 - x2) x1 x2. + case=> /mem_ext[y Dx] /mem_ext[y1 Dx1] /mem_ext[y2 Dx2]. + rewrite -Dx nu_inj; rewrite -Dx1 -Dx2 -rmorphB in Dx. + by rewrite (fmorph_inj _ Dx) !rmorphB -!nu_inj Dx1 Dx2. +- have [n] := max3 (x1 * x2) x1 x2. + case=> /mem_ext[y Dx] /mem_ext[y1 Dx1] /mem_ext[y2 Dx2]. + rewrite -Dx nu_inj; rewrite -Dx1 -Dx2 -rmorphM in Dx. + by rewrite (fmorph_inj _ Dx) !rmorphM -!nu_inj Dx1 Dx2. +by rewrite -(rmorph1 QsC) (nu_inj 0%N) !rmorph1. +Qed. + +(* Extended automorphisms of Q_n. *) +Lemma Qn_aut_exists k n : + coprime k n -> + {u : {rmorphism algC -> algC} | forall z, z ^+ n = 1 -> u z = z ^+ k}. +Proof. +have [-> /eqnP | n_gt0 co_k_n] := posnP n. + by rewrite gcdn0 => ->; exists [rmorphism of idfun]. +have [z prim_z] := C_prim_root_exists n_gt0. +have [Qn [QnC [[|zn []] // [Dz]]] genQn] := num_field_exists [:: z]. +pose phi := kHomExtend 1 \1 zn (zn ^+ k). +have homQn1: kHom 1 1 (\1%VF : 'End(Qn)) by rewrite kHom1. +have pzn_zk0: root (map_poly \1%VF (minPoly 1 zn)) (zn ^+ k). + rewrite -(fmorph_root QnC) rmorphX Dz -map_poly_comp. + rewrite (@eq_map_poly _ _ _ QnC) => [|a]; last by rewrite /= id_lfunE. + set p1 := map_poly _ _. + have [q1 Dp1]: exists q1, p1 = pQtoC q1. + have aP i: (minPoly 1 zn)`_i \in 1%VS. + by apply/polyOverP; exact: minPolyOver. + have{aP} a_ i := sig_eqW (vlineP _ _ (aP i)). + exists (\poly_(i < size (minPoly 1 zn)) sval (a_ i)). + apply/polyP=> i; rewrite coef_poly coef_map coef_poly /=. + case: ifP => _; rewrite ?rmorph0 //; case: (a_ i) => a /= ->. + apply: canRL (mulfK _) _; first by rewrite intr_eq0 denq_eq0. + by rewrite mulrzr -rmorphMz scalerMzl -mulrzr -numqE scaler_int rmorph_int. + have: root p1 z by rewrite -Dz fmorph_root root_minPoly. + rewrite Dp1; have [q2 [Dq2 _] ->] := minCpolyP z. + case/dvdpP=> r1 ->; rewrite rmorphM rootM /= -Dq2; apply/orP; right. + rewrite (minCpoly_cyclotomic prim_z) /cyclotomic. + rewrite (bigD1 (Ordinal (ltn_pmod k n_gt0))) ?coprime_modl //=. + by rewrite rootM root_XsubC prim_expr_mod ?eqxx. +have phiM: lrmorphism phi. + by apply/kHom_lrmorphism; rewrite -genQn span_seq1 /= kHomExtendP. +have [nu Dnu] := extend_algC_subfield_aut QnC (RMorphism phiM). +exists nu => _ /(prim_rootP prim_z)[i ->]. +rewrite rmorphX exprAC -Dz -Dnu /= -{1}[zn]hornerX /phi. +rewrite (kHomExtend_poly homQn1) ?polyOverX //. +rewrite map_polyE map_id_in => [|?]; last by rewrite id_lfunE. +by rewrite polyseqK hornerX rmorphX. +Qed. + +(* Algebraic integers. *) + +Definition Aint : pred_class := + fun x : algC => minCpoly x \is a polyOver Cint. +Fact Aint_key : pred_key Aint. Proof. by []. Qed. +Canonical Aint_keyed := KeyedPred Aint_key. + +Lemma root_monic_Aint p x : + root p x -> p \is monic -> p \is a polyOver Cint -> x \in Aint. +Proof. +have pZtoQtoC pz: pQtoC (pZtoQ pz) = pZtoC pz. + by rewrite -map_poly_comp; apply: eq_map_poly => b; rewrite /= rmorph_int. +move=> px0 mon_p /floorCpP[pz Dp]; rewrite unfold_in. +move: px0; rewrite Dp -pZtoQtoC; have [q [-> mon_q] ->] := minCpolyP x. +case/dvdpP_rat_int=> qz [a nz_a Dq] [r]. +move/(congr1 (fun q1 => lead_coef (a *: pZtoQ q1))). +rewrite rmorphM scalerAl -Dq lead_coefZ lead_coefM /=. +have /monicP->: pZtoQ pz \is monic by rewrite -(map_monic QtoCm) pZtoQtoC -Dp. +rewrite (monicP mon_q) mul1r mulr1 lead_coef_map_inj //; last exact: intr_inj. +rewrite Dq => ->; apply/polyOverP=> i; rewrite !(coefZ, coef_map). +by rewrite -rmorphM /= rmorph_int Cint_int. +Qed. + +Lemma Cint_rat_Aint z : z \in Crat -> z \in Aint -> z \in Cint. +Proof. +case/CratP=> a ->{z} /polyOverP/(_ 0%N). +have [p [Dp mon_p] dv_p] := minCpolyP (ratr a); rewrite Dp coef_map. +suffices /eqP->: p == 'X - a%:P by rewrite polyseqXsubC /= rmorphN rpredN. +rewrite -eqp_monic ?monicXsubC // irredp_XsubC //. + by rewrite -(size_map_poly QtoCm) -Dp neq_ltn size_minCpoly orbT. +by rewrite -dv_p fmorph_root root_XsubC. +Qed. + +Lemma Aint_Cint : {subset Cint <= Aint}. +Proof. +move=> x; rewrite -polyOverXsubC. +by apply: root_monic_Aint; rewrite ?monicXsubC ?root_XsubC. +Qed. + +Lemma Aint_int x : x%:~R \in Aint. +Proof. by rewrite Aint_Cint ?Cint_int. Qed. + +Lemma Aint0 : 0 \in Aint. Proof. exact: (Aint_int 0). Qed. +Lemma Aint1 : 1 \in Aint. Proof. exact: (Aint_int 1). Qed. +Hint Resolve Aint0 Aint1. + +Lemma Aint_unity_root n x : (n > 0)%N -> n.-unity_root x -> x \in Aint. +Proof. +move=> n_gt0 xn1; apply: root_monic_Aint xn1 (monic_Xn_sub_1 _ n_gt0) _. +by apply/polyOverP=> i; rewrite coefB coefC -mulrb coefXn /= rpredB ?rpred_nat. +Qed. + +Lemma Aint_prim_root n z : n.-primitive_root z -> z \in Aint. +Proof. +move=> pr_z; apply/(Aint_unity_root (prim_order_gt0 pr_z))/unity_rootP. +exact: prim_expr_order. +Qed. + +Lemma Aint_Cnat : {subset Cnat <= Aint}. +Proof. by move=> z /Cint_Cnat/Aint_Cint. Qed. + +(* This is Isaacs, Lemma (3.3) *) +Lemma Aint_subring_exists (X : seq algC) : + {subset X <= Aint} -> + {S : pred algC & + (*a*) subring_closed S + /\ (*b*) {subset X <= S} + & (*c*) {Y : {n : nat & n.-tuple algC} & + {subset tagged Y <= S} + & forall x, reflect (inIntSpan (tagged Y) x) (x \in S)}}. +Proof. +move=> AZ_X; pose m := (size X).+1. +pose n (i : 'I_m) := (size (minCpoly X`_i)).-2; pose N := (\max_i n i).+1. +pose IY := family (fun i => [pred e : 'I_N | e <= n i]%N). +have IY_0: 0 \in IY by apply/familyP=> // i; rewrite ffunE. +pose inIY := enum_rank_in IY_0. +pose Y := [seq \prod_(i < m) X`_i ^+ (f : 'I_N ^ m) i | f in IY]. +have S_P := Cint_spanP [tuple of Y]; set S := Cint_span _ in S_P. +have sYS: {subset Y <= S} by exact: mem_Cint_span. +have S_1: 1 \in S. + by apply/sYS/imageP; exists 0 => //; rewrite big1 // => i; rewrite ffunE. +have SmulX (i : 'I_m): {in S, forall x, x * X`_i \in S}. + move=> _ /S_P[x ->]; rewrite mulr_suml rpred_sum // => j _. + rewrite mulrzAl rpredMz {x}// nth_image mulrC (bigD1 i) //= mulrA -exprS. + move: {j}(enum_val j) (familyP (enum_valP j)) => f fP. + have:= fP i; rewrite inE /= leq_eqVlt => /predU1P[-> | fi_ltn]; last first. + apply/sYS/imageP; have fiK: (inord (f i).+1 : 'I_N) = (f i).+1 :> nat. + by rewrite inordK // ltnS (bigmax_sup i). + exists (finfun [eta f with i |-> inord (f i).+1]). + apply/familyP=> i1; rewrite inE ffunE /= fun_if fiK. + by case: eqP => [-> // | _]; exact: fP. + rewrite (bigD1 i isT) ffunE /= eqxx fiK; congr (_ * _). + by apply: eq_bigr => i1; rewrite ffunE /= => /negPf->. + have [/monicP ] := (minCpoly_monic X`_i, root_minCpoly X`_i). + rewrite /root horner_coef lead_coefE -(subnKC (size_minCpoly _)) subn2. + rewrite big_ord_recr /= addrC addr_eq0 => ->; rewrite mul1r => /eqP->. + have /floorCpP[p Dp]: X`_i \in Aint. + by have [/(nth_default 0)-> | /(mem_nth 0)/AZ_X] := leqP (size X) i. + rewrite -/(n i) Dp mulNr rpredN // mulr_suml rpred_sum // => [[e le_e]] /= _. + rewrite coef_map -mulrA mulrzl rpredMz ?sYS //; apply/imageP. + have eK: (inord e : 'I_N) = e :> nat by rewrite inordK // ltnS (bigmax_sup i). + exists (finfun [eta f with i |-> inord e]). + apply/familyP=> i1; rewrite inE ffunE /= fun_if eK. + by case: eqP => [-> // | _]; exact: fP. + rewrite (bigD1 i isT) ffunE /= eqxx eK; congr (_ * _). + by apply: eq_bigr => i1; rewrite ffunE /= => /negPf->. +exists S; last by exists (Tagged (fun n => n.-tuple _) [tuple of Y]). +split=> [|x Xx]; last first. + by rewrite -[x]mul1r -(nth_index 0 Xx) (SmulX (Ordinal _)) // ltnS index_size. +split=> // x y Sx Sy; first by rewrite rpredB. +case/S_P: Sy => {y}[y ->]; rewrite mulr_sumr rpred_sum //= => j. +rewrite mulrzAr rpredMz {y}// nth_image; move: {j}(enum_val j) => f. +elim/big_rec: _ => [|i y _ IHy] in x Sx *; first by rewrite mulr1. +rewrite mulrA {y}IHy //. +elim: {f}(f i : nat) => [|e IHe] in x Sx *; first by rewrite mulr1. +by rewrite exprS mulrA IHe // SmulX. +Qed. + +Section AlgIntSubring. + +Import DefaultKeying GRing.DefaultPred perm. + +(* This is Isaacs, Theorem (3.4). *) +Theorem fin_Csubring_Aint S n (Y : n.-tuple algC) : + mulr_closed S -> (forall x, reflect (inIntSpan Y x) (x \in S)) -> + {subset S <= Aint}. +Proof. +have ZP_C c: (ZtoC c)%:P \is a polyOver Cint by rewrite raddfMz rpred_int. +move=> mulS S_P x Sx; pose v := \row_(i < n) Y`_i. +have [v0 | nz_v] := eqVneq v 0. + case/S_P: Sx => {x}x ->; rewrite big1 ?isAlgInt0 // => i _. + by have /rowP/(_ i) := v0; rewrite !mxE => ->; rewrite mul0rz. +have sYS (i : 'I_n): x * Y`_i \in S. + by rewrite rpredM //; apply/S_P/Cint_spanP/mem_Cint_span/memt_nth. +pose A := \matrix_(i, j < n) sval (sig_eqW (S_P _ (sYS j))) i. +pose p := char_poly (map_mx ZtoC A). +have: p \is a polyOver Cint. + rewrite rpred_sum // => s _; rewrite rpredMsign rpred_prod // => j _. + by rewrite !mxE /= rpredB ?rpredMn ?polyOverX. +apply: root_monic_Aint (char_poly_monic _). +rewrite -eigenvalue_root_char; apply/eigenvalueP; exists v => //. +apply/rowP=> j; case dAj: (sig_eqW (S_P _ (sYS j))) => [a DxY]. +by rewrite !mxE DxY; apply: eq_bigr => i _; rewrite !mxE dAj /= mulrzr. +Qed. + +(* This is Isaacs, Corollary (3.5). *) +Corollary Aint_subring : subring_closed Aint. +Proof. +suff rAZ: {in Aint &, forall x y, (x - y \in Aint) * (x * y \in Aint)}. + by split=> // x y AZx AZy; rewrite rAZ. +move=> x y AZx AZy. +have [|S [ringS] ] := @Aint_subring_exists [:: x; y]; first exact/allP/and3P. +move=> /allP/and3P[Sx Sy _] [Y _ genYS]. +have AZ_S := fin_Csubring_Aint ringS genYS. +by have [_ S_B S_M] := ringS; rewrite !AZ_S ?S_B ?S_M. +Qed. +Canonical Aint_opprPred := OpprPred Aint_subring. +Canonical Aint_addrPred := AddrPred Aint_subring. +Canonical Aint_mulrPred := MulrPred Aint_subring. +Canonical Aint_zmodPred := ZmodPred Aint_subring. +Canonical Aint_semiringPred := SemiringPred Aint_subring. +Canonical Aint_smulrPred := SmulrPred Aint_subring. +Canonical Aint_subringPred := SubringPred Aint_subring. + +End AlgIntSubring. + +Lemma Aint_aut (nu : {rmorphism algC -> algC}) x : + (nu x \in Aint) = (x \in Aint). +Proof. by rewrite !unfold_in minCpoly_aut. Qed. + +Definition dvdA (e : Algebraics.divisor) : pred_class := + fun z : algC => if e == 0 then z == 0 else z / e \in Aint. +Fact dvdA_key e : pred_key (dvdA e). Proof. by []. Qed. +Canonical dvdA_keyed e := KeyedPred (dvdA_key e). +Delimit Scope algC_scope with A. +Delimit Scope algC_expanded_scope with Ax. +Notation "e %| x" := (x \in dvdA e) : algC_expanded_scope. +Notation "e %| x" := (@in_mem Algebraics.divisor x (mem (dvdA e))) : algC_scope. + +Fact dvdA_zmod_closed e : zmod_closed (dvdA e). +Proof. +split=> [|x y]; first by rewrite unfold_in mul0r eqxx rpred0 ?if_same. +rewrite ![(e %| _)%A]unfold_in. +case: ifP => [_ x0 /eqP-> | _]; first by rewrite subr0. +by rewrite mulrBl; apply: rpredB. +Qed. +Canonical dvdA_opprPred e := OpprPred (dvdA_zmod_closed e). +Canonical dvdA_addrPred e := AddrPred (dvdA_zmod_closed e). +Canonical dvdA_zmodPred e := ZmodPred (dvdA_zmod_closed e). + +Definition eqAmod (e x y : Algebraics.divisor) := (e %| x - y)%A. +Notation "x == y %[mod e ]" := (eqAmod e x y) : algC_scope. +Notation "x != y %[mod e ]" := (~~ (eqAmod e x y)) : algC_scope. + +Lemma eqAmod_refl e x : (x == x %[mod e])%A. +Proof. by rewrite /eqAmod subrr rpred0. Qed. +Hint Resolve eqAmod_refl. + +Lemma eqAmod_sym e x y : ((x == y %[mod e]) = (y == x %[mod e]))%A. +Proof. by rewrite /eqAmod -opprB rpredN. Qed. + +Lemma eqAmod_trans e y x z : + (x == y %[mod e] -> y == z %[mod e] -> x == z %[mod e])%A. +Proof. by move=> Exy Eyz; rewrite /eqAmod -[x](subrK y) -addrA rpredD. Qed. + +Lemma eqAmod_transl e x y z : + (x == y %[mod e])%A -> (x == z %[mod e])%A = (y == z %[mod e])%A. +Proof. by move/(sym_left_transitive (eqAmod_sym e) (@eqAmod_trans e)). Qed. + +Lemma eqAmod_transr e x y z : + (x == y %[mod e])%A -> (z == x %[mod e])%A = (z == y %[mod e])%A. +Proof. by move/(sym_right_transitive (eqAmod_sym e) (@eqAmod_trans e)). Qed. + +Lemma eqAmod0 e x : (x == 0 %[mod e])%A = (e %| x)%A. +Proof. by rewrite /eqAmod subr0. Qed. + +Lemma eqAmodN e x y : (- x == y %[mod e])%A = (x == - y %[mod e])%A. +Proof. by rewrite eqAmod_sym /eqAmod !opprK addrC. Qed. + +Lemma eqAmodDr e x y z : (y + x == z + x %[mod e])%A = (y == z %[mod e])%A. +Proof. by rewrite /eqAmod addrAC opprD !addrA subrK. Qed. + +Lemma eqAmodDl e x y z : (x + y == x + z %[mod e])%A = (y == z %[mod e])%A. +Proof. by rewrite !(addrC x) eqAmodDr. Qed. + +Lemma eqAmodD e x1 x2 y1 y2 : + (x1 == x2 %[mod e] -> y1 == y2 %[mod e] -> x1 + y1 == x2 + y2 %[mod e])%A. +Proof. rewrite -(eqAmodDl e x2 y1) -(eqAmodDr e y1); exact: eqAmod_trans. Qed. + +Lemma eqAmodm0 e : (e == 0 %[mod e])%A. +Proof. by rewrite /eqAmod subr0 unfold_in; case: ifPn => // /divff->. Qed. +Hint Resolve eqAmodm0. + +Lemma eqAmodMr e : + {in Aint, forall z x y, x == y %[mod e] -> x * z == y * z %[mod e]}%A. +Proof. +move=> z Zz x y. +rewrite /eqAmod -mulrBl ![(e %| _)%A]unfold_in mulf_eq0 mulrAC. +by case: ifP => [_ -> // | _ Exy]; apply: rpredM. +Qed. + +Lemma eqAmodMl e : + {in Aint, forall z x y, x == y %[mod e] -> z * x == z * y %[mod e]}%A. +Proof. by move=> z Zz x y Exy; rewrite !(mulrC z) eqAmodMr. Qed. + +Lemma eqAmodMl0 e : {in Aint, forall x, x * e == 0 %[mod e]}%A. +Proof. by move=> x Zx; rewrite -(mulr0 x) eqAmodMl. Qed. + +Lemma eqAmodMr0 e : {in Aint, forall x, e * x == 0 %[mod e]}%A. +Proof. by move=> x Zx; rewrite /= mulrC eqAmodMl0. Qed. + +Lemma eqAmod_addl_mul e : {in Aint, forall x y, x * e + y == y %[mod e]}%A. +Proof. by move=> x Zx y; rewrite -{2}[y]add0r eqAmodDr eqAmodMl0. Qed. + +Lemma eqAmodM e : {in Aint &, forall x1 y2 x2 y1, + x1 == x2 %[mod e] -> y1 == y2 %[mod e] -> x1 * y1 == x2 * y2 %[mod e]}%A. +Proof. +move=> x1 y2 Zx1 Zy2 x2 y1 eq_x /(eqAmodMl Zx1)/eqAmod_trans-> //. +exact: eqAmodMr. +Qed. + +Lemma eqAmod_rat : + {in Crat & &, forall e m n, (m == n %[mod e])%A = (m == n %[mod e])%C}. +Proof. +move=> e m n Qe Qm Qn; rewrite /eqCmod unfold_in /eqAmod unfold_in. +case: ifPn => // nz_e; apply/idP/idP=> [/Cint_rat_Aint | /Aint_Cint] -> //. +by rewrite rpred_div ?rpredB. +Qed. + +Lemma eqAmod0_rat : {in Crat &, forall e n, (n == 0 %[mod e])%A = (e %| n)%C}. +Proof. by move=> e n Qe Qn; rewrite /= eqAmod_rat /eqCmod ?subr0 ?Crat0. Qed. + +Lemma eqAmod_nat (e m n : nat) : (m == n %[mod e])%A = (m == n %[mod e])%N. +Proof. by rewrite eqAmod_rat ?rpred_nat // eqCmod_nat. Qed. + +Lemma eqAmod0_nat (e m : nat) : (m == 0 %[mod e])%A = (e %| m)%N. +Proof. by rewrite eqAmod0_rat ?rpred_nat // dvdC_nat. Qed. + +(* Multiplicative order. *) + +Definition orderC x := + let p := minCpoly x in + oapp val 0%N [pick n : 'I_(2 * size p ^ 2) | p == intrp 'Phi_n]. + +Notation "#[ x ]" := (orderC x) : C_scope. + +Lemma exp_orderC x : x ^+ #[x]%C = 1. +Proof. +rewrite /orderC; case: pickP => //= [] [n _] /= /eqP Dp. +have n_gt0: (0 < n)%N. + rewrite lt0n; apply: contraTneq (size_minCpoly x) => n0. + by rewrite Dp n0 Cyclotomic0 rmorph1 size_poly1. +have [z prim_z] := C_prim_root_exists n_gt0. +rewrite prim_expr_order // -(root_cyclotomic prim_z). +by rewrite -Cintr_Cyclotomic // -Dp root_minCpoly. +Qed. + +Lemma dvdn_orderC x n : (#[x]%C %| n)%N = (x ^+ n == 1). +Proof. +apply/idP/eqP=> [|x_n_1]; first by apply: expr_dvd; apply: exp_orderC. +have [-> | n_gt0] := posnP n; first by rewrite dvdn0. +have [m prim_x m_dv_n] := prim_order_exists n_gt0 x_n_1. +have{n_gt0} m_gt0 := dvdn_gt0 n_gt0 m_dv_n; congr (_ %| n)%N: m_dv_n. +pose p := minCpoly x; have Dp: p = cyclotomic x m := minCpoly_cyclotomic prim_x. +rewrite /orderC; case: pickP => /= [k /eqP Dp_k | no_k]; last first. + suffices lt_m_2p: (m < 2 * size p ^ 2)%N. + have /eqP[] := no_k (Ordinal lt_m_2p). + by rewrite /= -/p Dp -Cintr_Cyclotomic. + rewrite Dp size_cyclotomic (sqrnD 1) addnAC mulnDr -add1n leq_add //. + suffices: (m <= \prod_(q <- primes m | q == 2) q * totient m ^ 2)%N. + have [m_even | m_odd] := boolP (2 \in primes m). + by rewrite -big_filter filter_pred1_uniq ?primes_uniq // big_seq1. + by rewrite big_hasC ?has_pred1 // => /leq_trans-> //; apply: leq_addl. + rewrite big_mkcond totientE // -mulnn -!big_split /=. + rewrite {1}[m]prod_prime_decomp // prime_decompE big_map /= !big_seq. + elim/big_ind2: _ => // [n1 m1 n2 m2 | q]; first exact: leq_mul. + rewrite mem_primes => /and3P[q_pr _ q_dv_m]. + rewrite lognE q_pr m_gt0 q_dv_m /=; move: (logn q _) => k. + rewrite !mulnA expnS leq_mul //. + case: (ltngtP q) => // [|q_gt2 | ->]; first by rewrite ltnNge prime_gt1. + rewrite mul1n mulnAC mulnn -{1}[q]muln1 leq_mul ?expn_gt0 ?prime_gt0 //. + by rewrite -(subnKC q_gt2) (ltn_exp2l 1). + by rewrite !muln1 -expnS (ltn_exp2l 0). +have k_prim_x: k.-primitive_root x. + have k_gt0: (0 < k)%N. + rewrite lt0n; apply: contraTneq (size_minCpoly x) => k0. + by rewrite Dp_k k0 Cyclotomic0 rmorph1 size_poly1. + have [z prim_z] := C_prim_root_exists k_gt0. + rewrite -(root_cyclotomic prim_z) -Cintr_Cyclotomic //. + by rewrite -Dp_k root_minCpoly. +apply/eqP; rewrite eqn_dvd !(@prim_order_dvd _ _ x) //. +by rewrite !prim_expr_order ?eqxx. +Qed. diff --git a/mathcomp/field/all.v b/mathcomp/field/all.v new file mode 100644 index 0000000..a57ac19 --- /dev/null +++ b/mathcomp/field/all.v @@ -0,0 +1,11 @@ +Require Export algC. +Require Export algebraics_fundamentals. +Require Export algnum. +Require Export closed_field. +Require Export countalg. +Require Export cyclotomic. +Require Export falgebra. +Require Export fieldext. +Require Export finfield. +Require Export galois. +Require Export separable. diff --git a/mathcomp/field/closed_field.v b/mathcomp/field/closed_field.v new file mode 100644 index 0000000..edd27c5 --- /dev/null +++ b/mathcomp/field/closed_field.v @@ -0,0 +1,634 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. +Require Import bigop ssralg poly polydiv. + +(******************************************************************************) +(* A proof that algebraically closed field enjoy quantifier elimination, *) +(* as described in *) +(* ``A formal quantifier elimination for algebraically closed fields'', *) +(* proceedings of Calculemus 2010, by Cyril Cohen and Assia Mahboubi *) +(* *) +(* This file constructs an instance of quantifier elimination mixin, *) +(* (see the ssralg library) from the theory of polynomials with coefficients *) +(* is an algebraically closed field (see the polydiv library). *) +(* *) +(* This file hence deals with the transformation of formulae part, which we *) +(* address by implementing one CPS style formula transformer per effective *) +(* operation involved in the proof of quantifier elimination. See the paper *) +(* for more details. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GRing. +Local Open Scope ring_scope. + +Import Pdiv.Ring. +Import PreClosedField. + +Section ClosedFieldQE. + +Variable F : Field.type. + +Variable axiom : ClosedField.axiom F. + +Notation fF := (formula F). +Notation qf f := (qf_form f && rformula f). + +Definition polyF := seq (term F). + +Fixpoint eval_poly (e : seq F) pf := + if pf is c::q then (eval_poly e q)*'X + (eval e c)%:P else 0. + +Definition rpoly (p : polyF) := all (@rterm F) p. + +Fixpoint sizeT (k : nat -> fF) (p : polyF) := + if p is c::q then + sizeT (fun n => + if n is m.+1 then k m.+2 else + GRing.If (c == 0) (k 0%N) (k 1%N)) q + else k O%N. + + +Lemma sizeTP (k : nat -> formula F) (pf : polyF) (e : seq F) : + qf_eval e (sizeT k pf) = qf_eval e (k (size (eval_poly e pf))). +Proof. +elim: pf e k; first by move=> *; rewrite size_poly0. +move=> c qf Pqf e k; rewrite Pqf. +rewrite size_MXaddC -(size_poly_eq0 (eval_poly _ _)). +by case: (size (eval_poly e qf))=> //=; case: eqP; rewrite // orbF. +Qed. + +Lemma sizeT_qf (k : nat -> formula F) (p : polyF) : + (forall n, qf (k n)) -> rpoly p -> qf (sizeT k p). +Proof. +elim: p k => /= [|c q ihp] k kP rp; first exact: kP. +case/andP: rp=> rc rq. +apply: ihp; rewrite ?rq //; case=> [|n]; last exact: kP. +have [/andP[qf0 rf0] /andP[qf1 rf1]] := (kP 0, kP 1)%N. +by rewrite If_form_qf ?If_form_rf //= andbT. +Qed. + +Definition isnull (k : bool -> fF) (p : polyF) := + sizeT (fun n => k (n == 0%N)) p. + +Lemma isnullP (k : bool -> formula F) (p : polyF) (e : seq F) : + qf_eval e (isnull k p) = qf_eval e (k (eval_poly e p == 0)). +Proof. by rewrite sizeTP size_poly_eq0. Qed. + +Lemma isnull_qf (k : bool -> formula F) (p : polyF) : + (forall b, qf (k b)) -> rpoly p -> qf (isnull k p). +Proof. by move=> *; apply: sizeT_qf. Qed. + +Definition lt_sizeT (k : bool -> fF) (p q : polyF) : fF := + sizeT (fun n => sizeT (fun m => k (n<m)) q) p. + +Definition lift (p : {poly F}) := let: q := p in map Const q. + +Lemma eval_lift (e : seq F) (p : {poly F}) : eval_poly e (lift p) = p. +Proof. +elim/poly_ind: p => [|p c]; first by rewrite /lift polyseq0. +rewrite -cons_poly_def /lift polyseq_cons /nilp. +case pn0: (_ == _) => /=; last by move->; rewrite -cons_poly_def. +move=> _; rewrite polyseqC. +case c0: (_==_)=> /=. + move: pn0; rewrite (eqP c0) size_poly_eq0; move/eqP->. + by apply:val_inj=> /=; rewrite polyseq_cons // polyseq0. +by rewrite mul0r add0r; apply:val_inj=> /=; rewrite polyseq_cons // /nilp pn0. +Qed. + +Fixpoint lead_coefT (k : term F -> fF) p := + if p is c::q then + lead_coefT (fun l => GRing.If (l == 0) (k c) (k l)) q + else k (Const 0). + +Lemma lead_coefTP (k : term F -> formula F) : + (forall x e, qf_eval e (k x) = qf_eval e (k (Const (eval e x)))) -> + forall (p : polyF) (e : seq F), + qf_eval e (lead_coefT k p) = qf_eval e (k (Const (lead_coef (eval_poly e p)))). +Proof. +move=> Pk p e; elim: p k Pk => /= [*|a p' Pp' k Pk]; first by rewrite lead_coef0. +rewrite Pp'; last by move=> *; rewrite //= -Pk. +rewrite GRing.eval_If /= lead_coef_eq0. +case p'0: (_ == _); first by rewrite (eqP p'0) mul0r add0r lead_coefC -Pk. +rewrite lead_coefDl ?lead_coefMX // polyseqC size_mul ?p'0 //; last first. + by rewrite -size_poly_eq0 size_polyX. +rewrite size_polyX addnC /=; case: (_ == _)=> //=. +by rewrite ltnS lt0n size_poly_eq0 p'0. +Qed. + +Lemma lead_coefT_qf (k : term F -> formula F) (p : polyF) : + (forall c, rterm c -> qf (k c)) -> rpoly p -> qf (lead_coefT k p). +Proof. +elim: p k => /= [|c q ihp] k kP rp; first exact: kP. +move: rp; case/andP=> rc rq; apply: ihp; rewrite ?rq // => l rl. +have [/andP[qfc rfc] /andP[qfl rfl]] := (kP c rc, kP l rl). +by rewrite If_form_qf ?If_form_rf //= andbT. +Qed. + +Fixpoint amulXnT (a : term F) (n : nat) : polyF := + if n is n'.+1 then (Const 0) :: (amulXnT a n') else [::a]. + +Lemma eval_amulXnT (a : term F) (n : nat) (e : seq F) : + eval_poly e (amulXnT a n) = (eval e a)%:P * 'X^n. +Proof. +elim: n=> [|n] /=; first by rewrite expr0 mulr1 mul0r add0r. +by move->; rewrite addr0 -mulrA -exprSr. +Qed. + +Lemma ramulXnT: forall a n, rterm a -> rpoly (amulXnT a n). +Proof. by move=> a n; elim: n a=> [a /= -> //|n ihn a ra]; apply: ihn. Qed. + +Fixpoint sumpT (p q : polyF) := + if p is a::p' then + if q is b::q' then (Add a b)::(sumpT p' q') + else p + else q. + +Lemma eval_sumpT (p q : polyF) (e : seq F) : + eval_poly e (sumpT p q) = (eval_poly e p) + (eval_poly e q). +Proof. +elim: p q => [|a p Hp] q /=; first by rewrite add0r. +case: q => [|b q] /=; first by rewrite addr0. +rewrite Hp mulrDl -!addrA; congr (_+_); rewrite polyC_add addrC -addrA. +by congr (_+_); rewrite addrC. +Qed. + +Lemma rsumpT (p q : polyF) : rpoly p -> rpoly q -> rpoly (sumpT p q). +Proof. +elim: p q=> [|a p ihp] q rp rq //; move: rp; case/andP=> ra rp. +case: q rq => [|b q]; rewrite /= ?ra ?rp //=. +by case/andP=> -> rq //=; apply: ihp. +Qed. + +Fixpoint mulpT (p q : polyF) := + if p is a :: p' then sumpT (map (Mul a) q) (Const 0::(mulpT p' q)) else [::]. + +Lemma eval_mulpT (p q : polyF) (e : seq F) : + eval_poly e (mulpT p q) = (eval_poly e p) * (eval_poly e q). +Proof. +elim: p q=> [|a p Hp] q /=; first by rewrite mul0r. +rewrite eval_sumpT /= Hp addr0 mulrDl addrC mulrAC; congr (_+_). +elim: q=> [|b q Hq] /=; first by rewrite mulr0. +by rewrite Hq polyC_mul mulrDr mulrA. +Qed. + +Lemma rpoly_map_mul (t : term F) (p : polyF) (rt : rterm t) : + rpoly (map (Mul t) p) = rpoly p. +Proof. +by rewrite /rpoly all_map /= (@eq_all _ _ (@rterm _)) // => x; rewrite /= rt. +Qed. + +Lemma rmulpT (p q : polyF) : rpoly p -> rpoly q -> rpoly (mulpT p q). +Proof. +elim: p q=> [|a p ihp] q rp rq //=; move: rp; case/andP=> ra rp /=. +apply: rsumpT; last exact: ihp. +by rewrite rpoly_map_mul. +Qed. + +Definition opppT := map (Mul (@Const F (-1))). + +Lemma eval_opppT (p : polyF) (e : seq F) : + eval_poly e (opppT p) = - eval_poly e p. +Proof. +by elim: p; rewrite /= ?oppr0 // => ? ? ->; rewrite !mulNr opprD polyC_opp mul1r. +Qed. + +Definition natmulpT n := map (Mul (@NatConst F n)). + +Lemma eval_natmulpT (p : polyF) (n : nat) (e : seq F) : + eval_poly e (natmulpT n p) = (eval_poly e p) *+ n. +Proof. +elim: p; rewrite //= ?mul0rn // => c p ->. +rewrite mulrnDl mulr_natl polyC_muln; congr (_+_). +by rewrite -mulr_natl mulrAC -mulrA mulr_natl mulrC. +Qed. + +Fixpoint redivp_rec_loopT (q : polyF) sq cq (k : nat * polyF * polyF -> fF) + (c : nat) (qq r : polyF) (n : nat) {struct n}:= + sizeT (fun sr => + if sr < sq then k (c, qq, r) else + lead_coefT (fun lr => + let m := amulXnT lr (sr - sq) in + let qq1 := sumpT (mulpT qq [::cq]) m in + let r1 := sumpT (mulpT r ([::cq])) (opppT (mulpT m q)) in + if n is n1.+1 then redivp_rec_loopT q sq cq k c.+1 qq1 r1 n1 + else k (c.+1, qq1, r1) + ) r + ) r. + +Fixpoint redivp_rec_loop (q : {poly F}) sq cq + (k : nat) (qq r : {poly F})(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 redivp_rec_loop q sq cq k.+1 qq1 r1 n1 else + (k.+1, qq1, r1). + +Lemma redivp_rec_loopTP (k : nat * polyF * polyF -> formula F) : + (forall c qq r e, qf_eval e (k (c,qq,r)) + = qf_eval e (k (c, lift (eval_poly e qq), lift (eval_poly e r)))) + -> forall q sq cq c qq r n e + (d := redivp_rec_loop (eval_poly e q) sq (eval e cq) + c (eval_poly e qq) (eval_poly e r) n), + qf_eval e (redivp_rec_loopT q sq cq k c qq r n) + = qf_eval e (k (d.1.1, lift d.1.2, lift d.2)). +Proof. +move=> Pk q sq cq c qq r n e /=. +elim: n c qq r k Pk e => [|n Pn] c qq r k Pk e; rewrite sizeTP. + case ltrq : (_ < _); first by rewrite /= ltrq /= -Pk. + rewrite lead_coefTP => [|a p]; rewrite Pk. + rewrite ?(eval_mulpT,eval_amulXnT,eval_sumpT,eval_opppT) //=. + by rewrite ltrq //= mul_polyC ?(mul0r,add0r). + by symmetry; rewrite Pk ?(eval_mulpT,eval_amulXnT,eval_sumpT, eval_opppT). +case ltrq : (_<_); first by rewrite /= ltrq Pk. +rewrite lead_coefTP. + rewrite Pn ?(eval_mulpT,eval_amulXnT,eval_sumpT,eval_opppT) //=. + by rewrite ltrq //= mul_polyC ?(mul0r,add0r). +rewrite -/redivp_rec_loopT => x e'. +rewrite Pn; last by move=>*; rewrite Pk. +symmetry; rewrite Pn; last by move=>*; rewrite Pk. +rewrite Pk ?(eval_lift,eval_mulpT,eval_amulXnT,eval_sumpT,eval_opppT). +by rewrite mul_polyC ?(mul0r,add0r). +Qed. + +Lemma redivp_rec_loopT_qf (q : polyF) (sq : nat) (cq : term F) + (k : nat * polyF * polyF -> formula F) (c : nat) (qq r : polyF) (n : nat) : + (forall r, [&& rpoly r.1.2 & rpoly r.2] -> qf (k r)) -> + rpoly q -> rterm cq -> rpoly qq -> rpoly r -> + qf (redivp_rec_loopT q sq cq k c qq r n). +Proof. +elim: n q sq cq k c qq r => [|n ihn] q sq cq k c qq r kP rq rcq rqq rr. + apply: sizeT_qf=> // n; case: (_ < _); first by apply: kP; rewrite // rqq rr. + apply: lead_coefT_qf=> // l rl; apply: kP. + by rewrite /= ?(rsumpT,rmulpT,ramulXnT,rpoly_map_mul) //= rcq. +apply: sizeT_qf=> // m; case: (_ < _); first by apply: kP => //=; rewrite rqq rr. +apply: lead_coefT_qf=> // l rl; apply: ihn; rewrite //= ?rcq //. + by rewrite ?(rsumpT,rmulpT,ramulXnT,rpoly_map_mul) //= rcq. +by rewrite ?(rsumpT,rmulpT,ramulXnT,rpoly_map_mul) //= rcq. +Qed. + +Definition redivpT (p : polyF) (k : nat * polyF * polyF -> fF) + (q : polyF) : fF := + isnull (fun b => + if b then k (0%N, [::Const 0], p) else + sizeT (fun sq => + sizeT (fun sp => + lead_coefT (fun lq => + redivp_rec_loopT q sq lq k 0 [::Const 0] p sp + ) q + ) p + ) q + ) q. + +Lemma redivp_rec_loopP (q : {poly F}) (c : nat) (qq r : {poly F}) (n : nat) : + redivp_rec q c qq r n = redivp_rec_loop q (size q) (lead_coef q) c qq r n. +Proof. by elim: n c qq r => [| n Pn] c qq r //=; rewrite Pn. Qed. + +Lemma redivpTP (k : nat * polyF * polyF -> formula F) : + (forall c qq r e, + qf_eval e (k (c,qq,r)) = + qf_eval e (k (c, lift (eval_poly e qq), lift (eval_poly e r)))) -> + forall p q e (d := redivp (eval_poly e p) (eval_poly e q)), + qf_eval e (redivpT p k q) = qf_eval e (k (d.1.1, lift d.1.2, lift d.2)). +Proof. +move=> Pk p q e /=; rewrite isnullP unlock. +case q0 : (_ == _); first by rewrite Pk /= mul0r add0r polyC0. +rewrite !sizeTP lead_coefTP /=; last by move=> *; rewrite !redivp_rec_loopTP. +rewrite redivp_rec_loopTP /=; last by move=> *; rewrite Pk. +by rewrite mul0r add0r polyC0 redivp_rec_loopP. +Qed. + +Lemma redivpT_qf (p : polyF) (k : nat * polyF * polyF -> formula F) (q : polyF) : + (forall r, [&& rpoly r.1.2 & rpoly r.2] -> qf (k r)) -> + rpoly p -> rpoly q -> qf (redivpT p k q). +Proof. +move=> kP rp rq; rewrite /redivpT; apply: isnull_qf=> // [[]]; first exact: kP. +apply: sizeT_qf => // sq; apply: sizeT_qf=> // sp. +apply: lead_coefT_qf=> // lq rlq; exact: redivp_rec_loopT_qf. +Qed. + +Definition rmodpT (p : polyF) (k : polyF -> fF) (q : polyF) : fF := + redivpT p (fun d => k d.2) q. +Definition rdivpT (p : polyF) (k:polyF -> fF) (q : polyF) : fF := + redivpT p (fun d => k d.1.2) q. +Definition rscalpT (p : polyF) (k: nat -> fF) (q : polyF) : fF := + redivpT p (fun d => k d.1.1) q. +Definition rdvdpT (p : polyF) (k:bool -> fF) (q : polyF) : fF := + rmodpT p (isnull k) q. + +Fixpoint rgcdp_loop n (pp qq : {poly F}) {struct n} := + if rmodp pp qq == 0 then qq + else if n is n1.+1 then rgcdp_loop n1 qq (rmodp pp qq) + else rmodp pp qq. + +Fixpoint rgcdp_loopT (pp : polyF) (k : polyF -> formula F) n (qq : polyF) := + rmodpT pp (isnull + (fun b => if b then (k qq) + else (if n is n1.+1 + then rmodpT pp (rgcdp_loopT qq k n1) qq + else rmodpT pp k qq) + ) + ) qq. + +Lemma rgcdp_loopP (k : polyF -> formula F) : + (forall p e, qf_eval e (k p) = qf_eval e (k (lift (eval_poly e p)))) -> + forall n p q e, + qf_eval e (rgcdp_loopT p k n q) = + qf_eval e (k (lift (rgcdp_loop n (eval_poly e p) (eval_poly e q)))). +Proof. +move=> Pk n p q e. +elim: n p q e => /= [| m Pm] p q e. + rewrite redivpTP; last by move=>*; rewrite !isnullP eval_lift. + rewrite isnullP eval_lift; case: (_ == 0); first by rewrite Pk. + by rewrite redivpTP; last by move=>*; rewrite Pk. +rewrite redivpTP; last by move=>*; rewrite !isnullP eval_lift. +rewrite isnullP eval_lift; case: (_ == 0); first by rewrite Pk. +by rewrite redivpTP; move=>*; rewrite ?Pm !eval_lift. +Qed. + +Lemma rgcdp_loopT_qf (p : polyF) (k : polyF -> formula F) (q : polyF) (n : nat) : + (forall r, rpoly r -> qf (k r)) -> + rpoly p -> rpoly q -> qf (rgcdp_loopT p k n q). +elim: n p k q => [|n ihn] p k q kP rp rq. + apply: redivpT_qf=> // r; case/andP=> _ rr. + apply: isnull_qf=> // [[]]; first exact: kP. + by apply: redivpT_qf=> // r'; case/andP=> _ rr'; apply: kP. +apply: redivpT_qf=> // r; case/andP=> _ rr. +apply: isnull_qf=> // [[]]; first exact: kP. +by apply: redivpT_qf=> // r'; case/andP=> _ rr'; apply: ihn. +Qed. + +Definition rgcdpT (p : polyF) k (q : polyF) : fF := + let aux p1 k q1 := isnull + (fun b => if b + then (k q1) + else (sizeT (fun n => (rgcdp_loopT p1 k n q1)) p1)) p1 + in (lt_sizeT (fun b => if b then (aux q k p) else (aux p k q)) p q). + +Lemma rgcdpTP (k : seq (term F) -> formula F) : + (forall p e, qf_eval e (k p) = qf_eval e (k (lift (eval_poly e p)))) -> + forall p q e, qf_eval e (rgcdpT p k q) = + qf_eval e (k (lift (rgcdp (eval_poly e p) (eval_poly e q)))). +Proof. +move=> Pk p q e; rewrite /rgcdpT !sizeTP; case lqp: (_ < _). + rewrite isnullP; case q0: (_ == _); first by rewrite Pk (eqP q0) rgcdp0. + rewrite sizeTP rgcdp_loopP => [|e' p']; last by rewrite Pk. + by rewrite /rgcdp lqp q0. +rewrite isnullP; case p0: (_ == _); first by rewrite Pk (eqP p0) rgcd0p. +rewrite sizeTP rgcdp_loopP => [|e' p']; last by rewrite Pk. +by rewrite /rgcdp lqp p0. +Qed. + +Lemma rgcdpT_qf (p : polyF) (k : polyF -> formula F) (q : polyF) : + (forall r, rpoly r -> qf (k r)) -> rpoly p -> rpoly q -> qf (rgcdpT p k q). +Proof. +move=> kP rp rq; apply: sizeT_qf=> // n; apply: sizeT_qf=> // m. +by case:(_ < _); + apply: isnull_qf=> //; case; do ?apply: kP=> //; + apply: sizeT_qf=> // n'; apply: rgcdp_loopT_qf. +Qed. + +Fixpoint rgcdpTs k (ps : seq polyF) : fF := + if ps is p::pr then rgcdpTs (rgcdpT p k) pr else k [::Const 0]. + +Lemma rgcdpTsP (k : polyF -> formula F) : + (forall p e, qf_eval e (k p) = qf_eval e (k (lift (eval_poly e p)))) -> + forall ps e, + qf_eval e (rgcdpTs k ps) = + qf_eval e (k (lift (\big[@rgcdp _/0%:P]_(i <- ps)(eval_poly e i)))). +Proof. +move=> Pk ps e. +elim: ps k Pk; first by move=> p Pk; rewrite /= big_nil Pk /= mul0r add0r. +move=> p ps Pps /= k Pk /=; rewrite big_cons Pps => [|p' e']. + by rewrite rgcdpTP // eval_lift. +by rewrite !rgcdpTP // Pk !eval_lift . +Qed. + +Definition rseq_poly ps := all rpoly ps. + +Lemma rgcdpTs_qf (k : polyF -> formula F) (ps : seq polyF) : + (forall r, rpoly r -> qf (k r)) -> rseq_poly ps -> qf (rgcdpTs k ps). +Proof. +elim: ps k=> [|c p ihp] k kP rps=> /=; first exact: kP. +by move: rps; case/andP=> rc rp; apply: ihp=> // r rr; apply: rgcdpT_qf. +Qed. + +Fixpoint rgdcop_recT (q : polyF) k (p : polyF) n := + if n is m.+1 then + rgcdpT p (sizeT (fun sd => + if sd == 1%N then k p + else rgcdpT p (rdivpT p (fun r => rgdcop_recT q k r m)) q + )) q + else isnull (fun b => k [::Const b%:R]) q. + + +Lemma rgdcop_recTP (k : polyF -> formula F) : + (forall p e, qf_eval e (k p) = qf_eval e (k (lift (eval_poly e p)))) + -> forall p q n e, qf_eval e (rgdcop_recT p k q n) + = qf_eval e (k (lift (rgdcop_rec (eval_poly e p) (eval_poly e q) n))). +Proof. +move=> Pk p q n e. +elim: n k Pk p q e => [|n Pn] k Pk p q e /=. + rewrite isnullP /=. + by case: (_ == _); rewrite Pk /= mul0r add0r ?(polyC0, polyC1). +rewrite rgcdpTP ?sizeTP ?eval_lift //. + rewrite /rcoprimep; case se : (_==_); rewrite Pk //. + do ?[rewrite (rgcdpTP,Pn,eval_lift,redivpTP) | move=> * //=]. +by do ?[rewrite (sizeTP,eval_lift) | move=> * //=]. +Qed. + +Lemma rgdcop_recT_qf (p : polyF) (k : polyF -> formula F) (q : polyF) (n : nat) : + (forall r, rpoly r -> qf (k r)) -> + rpoly p -> rpoly q -> qf (rgdcop_recT p k q n). +Proof. +elim: n p k q => [|n ihn] p k q kP rp rq /=. +apply: isnull_qf=> //; first by case; rewrite kP. +apply: rgcdpT_qf=> // g rg; apply: sizeT_qf=> // n'. +case: (_ == _); first exact: kP. +apply: rgcdpT_qf=> // g' rg'; apply: redivpT_qf=> // r; case/andP=> rr _. +exact: ihn. +Qed. + +Definition rgdcopT q k p := sizeT (rgdcop_recT q k p) p. + +Lemma rgdcopTP (k : polyF -> formula F) : + (forall p e, qf_eval e (k p) = qf_eval e (k (lift (eval_poly e p)))) -> + forall p q e, qf_eval e (rgdcopT p k q) = + qf_eval e (k (lift (rgdcop (eval_poly e p) (eval_poly e q)))). +Proof. by move=> *; rewrite sizeTP rgdcop_recTP 1?Pk. Qed. + +Lemma rgdcopT_qf (p : polyF) (k : polyF -> formula F) (q : polyF) : + (forall r, rpoly r -> qf (k r)) -> rpoly p -> rpoly q -> qf (rgdcopT p k q). +Proof. +by move=> kP rp rq; apply: sizeT_qf => // n; apply: rgdcop_recT_qf. +Qed. + + +Definition ex_elim_seq (ps : seq polyF) (q : polyF) := + (rgcdpTs (rgdcopT q (sizeT (fun n => Bool (n != 1%N)))) ps). + +Lemma ex_elim_seqP (ps : seq polyF) (q : polyF) (e : seq F) : + let gp := (\big[@rgcdp _/0%:P]_(p <- ps)(eval_poly e p)) in + qf_eval e (ex_elim_seq ps q) = (size (rgdcop (eval_poly e q) gp) != 1%N). +Proof. +by do ![rewrite (rgcdpTsP,rgdcopTP,sizeTP,eval_lift) //= | move=> * //=]. +Qed. + +Lemma ex_elim_seq_qf (ps : seq polyF) (q : polyF) : + rseq_poly ps -> rpoly q -> qf (ex_elim_seq ps q). +Proof. +move=> rps rq; apply: rgcdpTs_qf=> // g rg; apply: rgdcopT_qf=> // d rd. +exact : sizeT_qf. +Qed. + +Fixpoint abstrX (i : nat) (t : term F) := + match t with + | (Var n) => if n == i then [::Const 0; Const 1] else [::t] + | (Opp x) => opppT (abstrX i x) + | (Add x y) => sumpT (abstrX i x) (abstrX i y) + | (Mul x y) => mulpT (abstrX i x) (abstrX i y) + | (NatMul x n) => natmulpT n (abstrX i x) + | (Exp x n) => let ax := (abstrX i x) in + iter n (mulpT ax) [::Const 1] + | _ => [::t] + end. + +Lemma abstrXP (i : nat) (t : term F) (e : seq F) (x : F) : + rterm t -> (eval_poly e (abstrX i t)).[x] = eval (set_nth 0 e i x) t. +Proof. +elim: t => [n | r | n | t tP s sP | t tP | t tP n | t tP s sP | t tP | t tP n] h. +- move=> /=; case ni: (_ == _); + rewrite //= ?(mul0r,add0r,addr0,polyC1,mul1r,hornerX,hornerC); + by rewrite // nth_set_nth /= ni. +- by rewrite /= mul0r add0r hornerC. +- by rewrite /= mul0r add0r hornerC. +- by case/andP: h => *; rewrite /= eval_sumpT hornerD tP ?sP. +- by rewrite /= eval_opppT hornerN tP. +- by rewrite /= eval_natmulpT hornerMn tP. +- by case/andP: h => *; rewrite /= eval_mulpT hornerM tP ?sP. +- by []. +- elim: n h => [|n ihn] rt; first by rewrite /= expr0 mul0r add0r hornerC. + by rewrite /= eval_mulpT exprSr hornerM ihn // mulrC tP. +Qed. + +Lemma rabstrX (i : nat) (t : term F) : rterm t -> rpoly (abstrX i t). +Proof. +elim: t; do ?[ by move=> * //=; do ?case: (_ == _)]. +- move=> t irt s irs /=; case/andP=> rt rs. + by apply: rsumpT; rewrite ?irt ?irs //. +- by move=> t irt /= rt; rewrite rpoly_map_mul ?irt //. +- by move=> t irt /= n rt; rewrite rpoly_map_mul ?irt //. +- move=> t irt s irs /=; case/andP=> rt rs. + by apply: rmulpT; rewrite ?irt ?irs //. +- move=> t irt /= n rt; move: (irt rt)=> {rt} rt; elim: n => [|n ihn] //=. + exact: rmulpT. +Qed. + +Implicit Types tx ty : term F. + +Lemma abstrX_mulM (i : nat) : {morph abstrX i : x y / Mul x y >-> mulpT x y}. +Proof. done. Qed. + +Lemma abstrX1 (i : nat) : abstrX i (Const 1) = [::Const 1]. +Proof. done. Qed. + +Lemma eval_poly_mulM e : {morph eval_poly e : x y / mulpT x y >-> mul x y}. +Proof. by move=> x y; rewrite eval_mulpT. Qed. + +Lemma eval_poly1 e : eval_poly e [::Const 1] = 1. +Proof. by rewrite /= mul0r add0r. Qed. + +Notation abstrX_bigmul := (big_morph _ (abstrX_mulM _) (abstrX1 _)). +Notation eval_bigmul := (big_morph _ (eval_poly_mulM _) (eval_poly1 _)). +Notation bigmap_id := (big_map _ (fun _ => true) id). + +Lemma rseq_poly_map (x : nat) (ts : seq (term F)) : + all (@rterm _) ts -> rseq_poly (map (abstrX x) ts). +Proof. +by elim: ts => //= t ts iht; case/andP=> rt rts; rewrite rabstrX // iht. +Qed. + +Definition ex_elim (x : nat) (pqs : seq (term F) * seq (term F)) := + ex_elim_seq (map (abstrX x) pqs.1) + (abstrX x (\big[Mul/Const 1]_(q <- pqs.2) q)). + +Lemma ex_elim_qf (x : nat) (pqs : seq (term F) * seq (term F)) : + dnf_rterm pqs -> qf (ex_elim x pqs). +case: pqs => ps qs; case/andP=> /= rps rqs. +apply: ex_elim_seq_qf; first exact: rseq_poly_map. +apply: rabstrX=> /=. +elim: qs rqs=> [|t ts iht] //=; first by rewrite big_nil. +by case/andP=> rt rts; rewrite big_cons /= rt /= iht. +Qed. + +Lemma holds_conj : forall e i x ps, all (@rterm _) ps -> + (holds (set_nth 0 e i x) (foldr (fun t : term F => And (t == 0)) True ps) + <-> all ((@root _)^~ x) (map (eval_poly e \o abstrX i) ps)). +Proof. +move=> e i x; elim=> [|p ps ihps] //=. +case/andP=> rp rps; rewrite rootE abstrXP //. +constructor; first by case=> -> hps; rewrite eqxx /=; apply/ihps. +by case/andP; move/eqP=> -> psr; split=> //; apply/ihps. +Qed. + +Lemma holds_conjn (e : seq F) (i : nat) (x : F) (ps : seq (term F)) : + all (@rterm _) ps -> + (holds (set_nth 0 e i x) (foldr (fun t : term F => And (t != 0)) True ps) <-> + all (fun p => ~~root p x) (map (eval_poly e \o abstrX i) ps)). +Proof. +elim: ps => [|p ps ihps] //=. +case/andP=> rp rps; rewrite rootE abstrXP //. +constructor; first by case=> /eqP-> hps /=; apply/ihps. +by case/andP=> pr psr; split; first apply/eqP=> //; apply/ihps. +Qed. + +Lemma holds_ex_elim : GRing.valid_QE_proj ex_elim. +Proof. +move=> i [ps qs] /= e; case/andP=> /= rps rqs. +rewrite ex_elim_seqP big_map. +have -> : \big[@rgcdp _/0%:P]_(j <- ps) eval_poly e (abstrX i j) = + \big[@rgcdp _/0%:P]_(j <- (map (eval_poly e) (map (abstrX i) (ps)))) j. + by rewrite !big_map. +rewrite -!map_comp. + have aux I (l : seq I) (P : I -> {poly F}) : + \big[(@gcdp F)/0]_(j <- l) P j %= \big[(@rgcdp F)/0]_(j <- l) P j. + elim: l => [| u l ihl] /=; first by rewrite !big_nil eqpxx. + rewrite !big_cons; move: ihl; move/(eqp_gcdr (P u)) => h. + apply: eqp_trans h _; rewrite eqp_sym; exact: eqp_rgcd_gcd. +case g0: (\big[(@rgcdp F)/0%:P]_(j <- map (eval_poly e \o abstrX i) ps) j == 0). + rewrite (eqP g0) rgdcop0. + case m0 : (_ == 0)=> //=; rewrite ?(size_poly1,size_poly0) //=. + rewrite abstrX_bigmul eval_bigmul -bigmap_id in m0. + constructor=> [[x] // []] //. + case=> _; move/holds_conjn=> hc; move/hc:rqs. + by rewrite -root_bigmul //= (eqP m0) root0. + constructor; move/negP:m0; move/negP=>m0. + case: (closed_nonrootP axiom _ m0) => x {m0}. + rewrite abstrX_bigmul eval_bigmul -bigmap_id root_bigmul=> m0. + exists x; do 2?constructor=> //; last by apply/holds_conjn. + apply/holds_conj; rewrite //= -root_biggcd. + by rewrite (eqp_root (aux _ _ _ )) (eqP g0) root0. +apply:(iffP (closed_rootP axiom _)); case=> x Px; exists x; move:Px => //=. + rewrite (eqp_root (eqp_rgdco_gdco _ _)) root_gdco ?g0 //. + rewrite -(eqp_root (aux _ _ _ )) root_biggcd abstrX_bigmul eval_bigmul. + rewrite -bigmap_id root_bigmul; case/andP=> psr qsr. + do 2?constructor; first by apply/holds_conj. + by apply/holds_conjn. +rewrite (eqp_root (eqp_rgdco_gdco _ _)) root_gdco ?g0 // -(eqp_root (aux _ _ _)). +rewrite root_biggcd abstrX_bigmul eval_bigmul -bigmap_id. +rewrite root_bigmul=> [[] // [hps hqs]]; apply/andP. +constructor; first by apply/holds_conj. +by apply/holds_conjn. +Qed. + +Lemma wf_ex_elim : GRing.wf_QE_proj ex_elim. +Proof. by move=> i bc /= rbc; apply: ex_elim_qf. Qed. + +Definition closed_fields_QEMixin := + QEdecFieldMixin wf_ex_elim holds_ex_elim. + +End ClosedFieldQE. diff --git a/mathcomp/field/countalg.v b/mathcomp/field/countalg.v new file mode 100644 index 0000000..68dd16f --- /dev/null +++ b/mathcomp/field/countalg.v @@ -0,0 +1,1107 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import bigop ssralg finalg zmodp matrix mxalgebra. +Require Import poly polydiv mxpoly generic_quotient ring_quotient closed_field. +Require Import ssrint rat. + +(*****************************************************************************) +(* This file clones part of ssralg hierachy for countable types; it does not *) +(* cover the left module / algebra interfaces, providing only *) +(* countZmodType == countable zmodType interface. *) +(* countRingType == countable ringType interface. *) +(* countComRingType == countable comRingType interface. *) +(* countUnitRingType == countable unitRingType interface. *) +(* countComUnitRingType == countable comUnitRingType interface. *) +(* countIdomainType == countable idomainType interface. *) +(* countFieldType == countable fieldType interface. *) +(* countDecFieldType == countable decFieldType interface. *) +(* countClosedFieldType == countable closedFieldType interface. *) +(* The interface cloning syntax is extended to these structures *) +(* [countZmodType of M] == countZmodType structure for an M that has both *) +(* zmodType and countType structures. *) +(* ... etc *) +(* This file provides constructions for both simple extension and algebraic *) +(* closure of countable fields. *) +(*****************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope ring_scope. +Import GRing.Theory CodeSeq. + +Module CountRing. + +Local Notation mixin_of T := (Countable.mixin_of T). + +Section Generic. + +(* Implicits *) +Variables (type base_type : Type) (class_of base_of : Type -> Type). +Variable base_sort : base_type -> Type. + +(* Explicits *) +Variable Pack : forall T, class_of T -> Type -> type. +Variable Class : forall T, base_of T -> mixin_of T -> 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 c m & phant_id (Countable.class fT) (Countable.Class c m) => + Pack (@Class T b m) T. + +End Generic. + +Implicit Arguments gen_pack [type base_type class_of base_of base_sort]. +Local Notation cnt_ c := (@Countable.Class _ c c). +Local Notation do_pack pack T := (pack T _ _ id _ _ _ id). +Import GRing.Theory. + +Module Zmodule. + +Section ClassDef. + +Record class_of M := + Class { base : GRing.Zmodule.class_of M; mixin : mixin_of M }. +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 (cnt_ xclass) xT. +Definition zmodType := @GRing.Zmodule.Pack cT xclass xT. + +Definition join_countType := @Countable.Pack zmodType (cnt_ xclass) xT. + +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 zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Canonical join_countType. +Notation countZmodType := type. +Notation "[ 'countZmodType' 'of' T ]" := (do_pack pack T) + (at level 0, format "[ 'countZmodType' 'of' T ]") : form_scope. +End Exports. + +End Zmodule. +Import Zmodule.Exports. + +Module Ring. + +Section ClassDef. + +Record class_of R := Class { base : GRing.Ring.class_of R; mixin : mixin_of R }. +Definition base2 R (c : class_of R) := Zmodule.Class (base c) (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 (cnt_ xclass) xT. +Definition zmodType := @GRing.Zmodule.Pack cT xclass cT. +Definition countZmodType := @Zmodule.Pack cT xclass xT. +Definition ringType := @GRing.Ring.Pack cT xclass xT. +Definition join_countType := @Countable.Pack ringType (cnt_ xclass) xT. +Definition join_countZmodType := @Zmodule.Pack ringType xclass xT. + +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 zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Coercion countZmodType : type >-> Zmodule.type. +Canonical countZmodType. +Coercion ringType : type >-> GRing.Ring.type. +Canonical ringType. +Canonical join_countType. +Canonical join_countZmodType. +Notation countRingType := type. +Notation "[ 'countRingType' 'of' T ]" := (do_pack pack T) + (at level 0, format "[ 'countRingType' 'of' T ]") : form_scope. +End Exports. + +End Ring. +Import Ring.Exports. + +Module ComRing. + +Section ClassDef. + +Record class_of R := + Class { base : GRing.ComRing.class_of R; mixin : mixin_of R }. +Definition base2 R (c : class_of R) := Ring.Class (base c) (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 (cnt_ xclass) xT. +Definition zmodType := @GRing.Zmodule.Pack cT xclass xT. +Definition countZmodType := @Zmodule.Pack cT xclass xT. +Definition ringType := @GRing.Ring.Pack cT xclass xT. +Definition countRingType := @Ring.Pack cT xclass xT. +Definition comRingType := @GRing.ComRing.Pack cT xclass xT. +Definition join_countType := @Countable.Pack comRingType (cnt_ xclass) xT. +Definition join_countZmodType := @Zmodule.Pack comRingType xclass xT. +Definition join_countRingType := @Ring.Pack comRingType xclass xT. + +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 zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Coercion countZmodType : type >-> Zmodule.type. +Canonical countZmodType. +Coercion ringType : type >-> GRing.Ring.type. +Canonical ringType. +Coercion countRingType : type >-> Ring.type. +Canonical countRingType. +Coercion comRingType : type >-> GRing.ComRing.type. +Canonical comRingType. +Canonical join_countType. +Canonical join_countZmodType. +Canonical join_countRingType. +Notation countComRingType := CountRing.ComRing.type. +Notation "[ 'countComRingType' 'of' T ]" := (do_pack pack T) + (at level 0, format "[ 'countComRingType' 'of' T ]") : form_scope. +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 }. +Definition base2 R (c : class_of R) := Ring.Class (base c) (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 (cnt_ xclass) xT. +Definition zmodType := @GRing.Zmodule.Pack cT xclass xT. +Definition countZmodType := @Zmodule.Pack cT xclass xT. +Definition ringType := @GRing.Ring.Pack cT xclass xT. +Definition countRingType := @Ring.Pack cT xclass xT. +Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT. + +Definition join_countType := @Countable.Pack unitRingType (cnt_ xclass) xT. +Definition join_countZmodType := @Zmodule.Pack unitRingType xclass xT. +Definition join_countRingType := @Ring.Pack unitRingType xclass xT. + +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 zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Coercion countZmodType : type >-> Zmodule.type. +Canonical countZmodType. +Coercion ringType : type >-> GRing.Ring.type. +Canonical ringType. +Coercion countRingType : type >-> Ring.type. +Canonical countRingType. +Coercion unitRingType : type >-> GRing.UnitRing.type. +Canonical unitRingType. +Canonical join_countType. +Canonical join_countZmodType. +Canonical join_countRingType. +Notation countUnitRingType := CountRing.UnitRing.type. +Notation "[ 'countUnitRingType' 'of' T ]" := (do_pack pack T) + (at level 0, format "[ 'countUnitRingType' 'of' T ]") : form_scope. +End Exports. + +End UnitRing. +Import UnitRing.Exports. + +Module ComUnitRing. + +Section ClassDef. + +Record class_of R := + Class { base : GRing.ComUnitRing.class_of R; mixin : mixin_of R }. +Definition base2 R (c : class_of R) := ComRing.Class (base c) (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 (cnt_ xclass) xT. +Definition zmodType := @GRing.Zmodule.Pack cT xclass xT. +Definition countZmodType := @Zmodule.Pack cT xclass xT. +Definition ringType := @GRing.Ring.Pack cT xclass xT. +Definition countRingType := @Ring.Pack cT xclass xT. +Definition comRingType := @GRing.ComRing.Pack cT xclass xT. +Definition countComRingType := @ComRing.Pack cT xclass xT. +Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT. +Definition countUnitRingType := @UnitRing.Pack cT xclass xT. +Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT. + +Definition join_countType := @Countable.Pack comUnitRingType (cnt_ xclass) xT. +Definition join_countZmodType := @Zmodule.Pack comUnitRingType xclass xT. +Definition join_countRingType := @Ring.Pack comUnitRingType xclass xT. +Definition join_countComRingType := @ComRing.Pack comUnitRingType xclass xT. +Definition join_countUnitRingType := @UnitRing.Pack comUnitRingType xclass xT. +Definition ujoin_countComRingType := @ComRing.Pack unitRingType xclass xT. +Definition cjoin_countUnitRingType := @UnitRing.Pack comRingType xclass xT. +Definition ccjoin_countUnitRingType := + @UnitRing.Pack countComRingType xclass xT. + +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 zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Coercion countZmodType : type >-> Zmodule.type. +Canonical countZmodType. +Coercion ringType : type >-> GRing.Ring.type. +Canonical ringType. +Coercion countRingType : type >-> Ring.type. +Canonical countRingType. +Coercion comRingType : type >-> GRing.ComRing.type. +Canonical comRingType. +Coercion countComRingType : type >-> ComRing.type. +Canonical countComRingType. +Coercion unitRingType : type >-> GRing.UnitRing.type. +Canonical unitRingType. +Coercion countUnitRingType : type >-> UnitRing.type. +Canonical countUnitRingType. +Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. +Canonical comUnitRingType. +Canonical join_countType. +Canonical join_countZmodType. +Canonical join_countRingType. +Canonical join_countComRingType. +Canonical join_countUnitRingType. +Canonical ujoin_countComRingType. +Canonical cjoin_countUnitRingType. +Canonical ccjoin_countUnitRingType. +Notation countComUnitRingType := CountRing.ComUnitRing.type. +Notation "[ 'countComUnitRingType' 'of' T ]" := (do_pack pack T) + (at level 0, format "[ 'countComUnitRingType' 'of' T ]") : form_scope. +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 }. +Definition base2 R (c : class_of R) := ComUnitRing.Class (base c) (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 (cnt_ xclass) xT. +Definition zmodType := @GRing.Zmodule.Pack cT xclass xT. +Definition countZmodType := @Zmodule.Pack cT xclass xT. +Definition ringType := @GRing.Ring.Pack cT xclass xT. +Definition countRingType := @Ring.Pack cT xclass xT. +Definition comRingType := @GRing.ComRing.Pack cT xclass xT. +Definition countComRingType := @ComRing.Pack cT xclass xT. +Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT. +Definition countUnitRingType := @UnitRing.Pack cT xclass xT. +Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT. +Definition countComUnitRingType := @ComUnitRing.Pack cT xclass xT. +Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT. + +Definition join_countType := @Countable.Pack idomainType (cnt_ xclass) xT. +Definition join_countZmodType := @Zmodule.Pack idomainType xclass xT. +Definition join_countRingType := @Ring.Pack idomainType xclass xT. +Definition join_countUnitRingType := @UnitRing.Pack idomainType xclass xT. +Definition join_countComRingType := @ComRing.Pack idomainType xclass xT. +Definition join_countComUnitRingType := @ComUnitRing.Pack idomainType xclass xT. + +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 zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Coercion countZmodType : type >-> Zmodule.type. +Canonical countZmodType. +Coercion ringType : type >-> GRing.Ring.type. +Canonical ringType. +Coercion countRingType : type >-> Ring.type. +Canonical countRingType. +Coercion comRingType : type >-> GRing.ComRing.type. +Canonical comRingType. +Coercion countComRingType : type >-> ComRing.type. +Canonical countComRingType. +Coercion unitRingType : type >-> GRing.UnitRing.type. +Canonical unitRingType. +Coercion countUnitRingType : type >-> UnitRing.type. +Canonical countUnitRingType. +Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. +Canonical comUnitRingType. +Coercion countComUnitRingType : type >-> ComUnitRing.type. +Canonical countComUnitRingType. +Coercion idomainType : type >-> GRing.IntegralDomain.type. +Canonical idomainType. +Canonical join_countType. +Canonical join_countZmodType. +Canonical join_countRingType. +Canonical join_countComRingType. +Canonical join_countUnitRingType. +Canonical join_countComUnitRingType. +Notation countIdomainType := CountRing.IntegralDomain.type. +Notation "[ 'countIdomainType' 'of' T ]" := (do_pack pack T) + (at level 0, format "[ 'countIdomainType' 'of' T ]") : form_scope. +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 }. +Definition base2 R (c : class_of R) := IntegralDomain.Class (base c) (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 (cnt_ xclass) xT. +Definition zmodType := @GRing.Zmodule.Pack cT xclass xT. +Definition countZmodType := @Zmodule.Pack cT xclass xT. +Definition ringType := @GRing.Ring.Pack cT xclass xT. +Definition countRingType := @Ring.Pack cT xclass xT. +Definition comRingType := @GRing.ComRing.Pack cT xclass xT. +Definition countComRingType := @ComRing.Pack cT xclass xT. +Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT. +Definition countUnitRingType := @UnitRing.Pack cT xclass xT. +Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT. +Definition countComUnitRingType := @ComUnitRing.Pack cT xclass xT. +Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT. +Definition countIdomainType := @IntegralDomain.Pack cT xclass xT. +Definition fieldType := @GRing.Field.Pack cT xclass xT. + +Definition join_countType := @Countable.Pack fieldType (cnt_ xclass) xT. +Definition join_countZmodType := @Zmodule.Pack fieldType xclass xT. +Definition join_countRingType := @Ring.Pack fieldType xclass xT. +Definition join_countUnitRingType := @UnitRing.Pack fieldType xclass xT. +Definition join_countComRingType := @ComRing.Pack fieldType xclass xT. +Definition join_countComUnitRingType := @ComUnitRing.Pack fieldType xclass xT. +Definition join_countIdomainType := @IntegralDomain.Pack fieldType xclass xT. + +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 zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Coercion countZmodType : type >-> Zmodule.type. +Canonical countZmodType. +Coercion ringType : type >-> GRing.Ring.type. +Canonical ringType. +Coercion countRingType : type >-> Ring.type. +Canonical countRingType. +Coercion comRingType : type >-> GRing.ComRing.type. +Canonical comRingType. +Coercion countComRingType : type >-> ComRing.type. +Canonical countComRingType. +Coercion unitRingType : type >-> GRing.UnitRing.type. +Canonical unitRingType. +Coercion countUnitRingType : type >-> UnitRing.type. +Canonical countUnitRingType. +Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. +Canonical comUnitRingType. +Coercion countComUnitRingType : type >-> ComUnitRing.type. +Canonical countComUnitRingType. +Coercion idomainType : type >-> GRing.IntegralDomain.type. +Canonical idomainType. +Coercion countIdomainType : type >-> IntegralDomain.type. +Canonical countIdomainType. +Coercion fieldType : type >-> GRing.Field.type. +Canonical fieldType. +Canonical join_countType. +Canonical join_countZmodType. +Canonical join_countRingType. +Canonical join_countComRingType. +Canonical join_countUnitRingType. +Canonical join_countComUnitRingType. +Canonical join_countIdomainType. +Notation countFieldType := CountRing.Field.type. +Notation "[ 'countFieldType' 'of' T ]" := (do_pack pack T) + (at level 0, format "[ 'countFieldType' 'of' T ]") : form_scope. +End Exports. + +End Field. +Import Field.Exports. + +Module DecidableField. + +Section ClassDef. + +Record class_of R := + Class { base : GRing.DecidableField.class_of R; mixin : mixin_of R }. +Definition base2 R (c : class_of R) := Field.Class (base c) (mixin c). +Local Coercion base : class_of >-> GRing.DecidableField.class_of. +Local Coercion base2 : class_of >-> Field.class_of. + +Structure type := Pack {sort; _ : class_of sort; _ : Type}. +Local Coercion sort : type >-> Sortclass. +Definition pack := gen_pack Pack Class GRing.DecidableField.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 (cnt_ xclass) xT. +Definition zmodType := @GRing.Zmodule.Pack cT xclass xT. +Definition countZmodType := @Zmodule.Pack cT xclass xT. +Definition ringType := @GRing.Ring.Pack cT xclass xT. +Definition countRingType := @Ring.Pack cT xclass xT. +Definition comRingType := @GRing.ComRing.Pack cT xclass xT. +Definition countComRingType := @ComRing.Pack cT xclass xT. +Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT. +Definition countUnitRingType := @UnitRing.Pack cT xclass xT. +Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT. +Definition countComUnitRingType := @ComUnitRing.Pack cT xclass xT. +Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT. +Definition countIdomainType := @IntegralDomain.Pack cT xclass xT. +Definition fieldType := @GRing.Field.Pack cT xclass xT. +Definition countFieldType := @Field.Pack cT xclass xT. +Definition decFieldType := @GRing.DecidableField.Pack cT xclass xT. + +Definition join_countType := @Countable.Pack decFieldType (cnt_ xclass) xT. +Definition join_countZmodType := @Zmodule.Pack decFieldType xclass xT. +Definition join_countRingType := @Ring.Pack decFieldType xclass xT. +Definition join_countUnitRingType := @UnitRing.Pack decFieldType xclass xT. +Definition join_countComRingType := @ComRing.Pack decFieldType xclass xT. +Definition join_countComUnitRingType := + @ComUnitRing.Pack decFieldType xclass xT. +Definition join_countIdomainType := @IntegralDomain.Pack decFieldType xclass xT. +Definition join_countFieldType := @Field.Pack decFieldType xclass xT. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> GRing.DecidableField.class_of. +Coercion base2 : class_of >-> Field.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 zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Coercion countZmodType : type >-> Zmodule.type. +Canonical countZmodType. +Coercion ringType : type >-> GRing.Ring.type. +Canonical ringType. +Coercion countRingType : type >-> Ring.type. +Canonical countRingType. +Coercion comRingType : type >-> GRing.ComRing.type. +Canonical comRingType. +Coercion countComRingType : type >-> ComRing.type. +Canonical countComRingType. +Coercion unitRingType : type >-> GRing.UnitRing.type. +Canonical unitRingType. +Coercion countUnitRingType : type >-> UnitRing.type. +Canonical countUnitRingType. +Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. +Canonical comUnitRingType. +Coercion countComUnitRingType : type >-> ComUnitRing.type. +Canonical countComUnitRingType. +Coercion idomainType : type >-> GRing.IntegralDomain.type. +Canonical idomainType. +Coercion countIdomainType : type >-> IntegralDomain.type. +Canonical countIdomainType. +Coercion fieldType : type >-> GRing.Field.type. +Canonical fieldType. +Coercion countFieldType : type >-> Field.type. +Canonical countFieldType. +Coercion decFieldType : type >-> GRing.DecidableField.type. +Canonical decFieldType. +Canonical join_countType. +Canonical join_countZmodType. +Canonical join_countRingType. +Canonical join_countComRingType. +Canonical join_countUnitRingType. +Canonical join_countComUnitRingType. +Canonical join_countIdomainType. +Canonical join_countFieldType. +Notation countDecFieldType := CountRing.DecidableField.type. +Notation "[ 'countDecFieldType' 'of' T ]" := (do_pack pack T) + (at level 0, format "[ 'countDecFieldType' 'of' T ]") : form_scope. +End Exports. + +End DecidableField. +Import DecidableField.Exports. + +Module ClosedField. + +Section ClassDef. + +Record class_of R := + Class { base : GRing.ClosedField.class_of R; mixin : mixin_of R }. +Definition base2 R (c : class_of R) := DecidableField.Class (base c) (mixin c). +Local Coercion base : class_of >-> GRing.ClosedField.class_of. +Local Coercion base2 : class_of >-> DecidableField.class_of. + +Structure type := Pack {sort; _ : class_of sort; _ : Type}. +Local Coercion sort : type >-> Sortclass. +Definition pack := gen_pack Pack Class GRing.ClosedField.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 (cnt_ xclass) xT. +Definition zmodType := @GRing.Zmodule.Pack cT xclass xT. +Definition countZmodType := @Zmodule.Pack cT xclass xT. +Definition ringType := @GRing.Ring.Pack cT xclass xT. +Definition countRingType := @Ring.Pack cT xclass xT. +Definition comRingType := @GRing.ComRing.Pack cT xclass xT. +Definition countComRingType := @ComRing.Pack cT xclass xT. +Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT. +Definition countUnitRingType := @UnitRing.Pack cT xclass xT. +Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT. +Definition countComUnitRingType := @ComUnitRing.Pack cT xclass xT. +Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT. +Definition countIdomainType := @IntegralDomain.Pack cT xclass xT. +Definition fieldType := @GRing.Field.Pack cT xclass xT. +Definition countFieldType := @Field.Pack cT xclass xT. +Definition decFieldType := @GRing.DecidableField.Pack cT xclass xT. +Definition countDecFieldType := @DecidableField.Pack cT xclass xT. +Definition closedFieldType := @GRing.ClosedField.Pack cT xclass xT. + +Definition join_countType := @Countable.Pack closedFieldType (cnt_ xclass) xT. +Definition join_countZmodType := @Zmodule.Pack closedFieldType xclass xT. +Definition join_countRingType := @Ring.Pack closedFieldType xclass xT. +Definition join_countUnitRingType := @UnitRing.Pack closedFieldType xclass xT. +Definition join_countComRingType := @ComRing.Pack closedFieldType xclass xT. +Definition join_countComUnitRingType := + @ComUnitRing.Pack closedFieldType xclass xT. +Definition join_countIdomainType := + @IntegralDomain.Pack closedFieldType xclass xT. +Definition join_countFieldType := @Field.Pack closedFieldType xclass xT. +Definition join_countDecFieldType := + @DecidableField.Pack closedFieldType xclass xT. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> GRing.ClosedField.class_of. +Coercion base2 : 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 countType : type >-> Countable.type. +Canonical countType. +Coercion zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Coercion countZmodType : type >-> Zmodule.type. +Canonical countZmodType. +Coercion ringType : type >-> GRing.Ring.type. +Canonical ringType. +Coercion countRingType : type >-> Ring.type. +Canonical countRingType. +Coercion comRingType : type >-> GRing.ComRing.type. +Canonical comRingType. +Coercion countComRingType : type >-> ComRing.type. +Canonical countComRingType. +Coercion unitRingType : type >-> GRing.UnitRing.type. +Canonical unitRingType. +Coercion countUnitRingType : type >-> UnitRing.type. +Canonical countUnitRingType. +Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. +Canonical comUnitRingType. +Coercion countComUnitRingType : type >-> ComUnitRing.type. +Canonical countComUnitRingType. +Coercion idomainType : type >-> GRing.IntegralDomain.type. +Canonical idomainType. +Coercion fieldType : type >-> GRing.Field.type. +Canonical fieldType. +Coercion countFieldType : type >-> Field.type. +Canonical countFieldType. +Coercion decFieldType : type >-> GRing.DecidableField.type. +Canonical decFieldType. +Coercion countDecFieldType : type >-> DecidableField.type. +Canonical countDecFieldType. +Coercion closedFieldType : type >-> GRing.ClosedField.type. +Canonical closedFieldType. +Canonical join_countType. +Canonical join_countZmodType. +Canonical join_countRingType. +Canonical join_countComRingType. +Canonical join_countUnitRingType. +Canonical join_countComUnitRingType. +Canonical join_countIdomainType. +Canonical join_countFieldType. +Canonical join_countDecFieldType. +Notation countClosedFieldType := CountRing.ClosedField.type. +Notation "[ 'countClosedFieldType' 'of' T ]" := (do_pack pack T) + (at level 0, format "[ 'countClosedFieldType' 'of' T ]") : form_scope. +End Exports. + +End ClosedField. +Import ClosedField.Exports. + +End CountRing. + +Import CountRing. +Export Zmodule.Exports Ring.Exports ComRing.Exports UnitRing.Exports. +Export ComUnitRing.Exports IntegralDomain.Exports. +Export Field.Exports DecidableField.Exports ClosedField.Exports. + +Require Import poly polydiv generic_quotient ring_quotient. +Require Import mxpoly polyXY. +Import GRing.Theory. +Require Import closed_field. + +Canonical Zp_countZmodType m := [countZmodType of 'I_m.+1]. +Canonical Zp_countRingType m := [countRingType of 'I_m.+2]. +Canonical Zp_countComRingType m := [countComRingType of 'I_m.+2]. +Canonical Zp_countUnitRingType m := [countUnitRingType of 'I_m.+2]. +Canonical Zp_countComUnitRingType m := [countComUnitRingType of 'I_m.+2]. +Canonical Fp_countIdomainType p := [countIdomainType of 'F_p]. +Canonical Fp_countFieldType p := [countFieldType of 'F_p]. +Canonical Fp_countDecFieldType p := [countDecFieldType of 'F_p]. + +Canonical matrix_countZmodType (M : countZmodType) m n := + [countZmodType of 'M[M]_(m, n)]. +Canonical matrix_countRingType (R : countRingType) n := + [countRingType of 'M[R]_n.+1]. +Canonical matrix_countUnitRingType (R : countComUnitRingType) n := + [countUnitRingType of 'M[R]_n.+1]. + +Definition poly_countMixin (R : countRingType) := + [countMixin of polynomial R by <:]. +Canonical polynomial_countType R := CountType _ (poly_countMixin R). +Canonical poly_countType (R : countRingType) := [countType of {poly R}]. +Canonical polynomial_countZmodType (R : countRingType) := + [countZmodType of polynomial R]. +Canonical poly_countZmodType (R : countRingType) := [countZmodType of {poly R}]. +Canonical polynomial_countRingType (R : countRingType) := + [countRingType of polynomial R]. +Canonical poly_countRingType (R : countRingType) := [countRingType of {poly R}]. +Canonical polynomial_countComRingType (R : countComRingType) := + [countComRingType of polynomial R]. +Canonical poly_countComRingType (R : countComRingType) := + [countComRingType of {poly R}]. +Canonical polynomial_countUnitRingType (R : countIdomainType) := + [countUnitRingType of polynomial R]. +Canonical poly_countUnitRingType (R : countIdomainType) := + [countUnitRingType of {poly R}]. +Canonical polynomial_countComUnitRingType (R : countIdomainType) := + [countComUnitRingType of polynomial R]. +Canonical poly_countComUnitRingType (R : countIdomainType) := + [countComUnitRingType of {poly R}]. +Canonical polynomial_countIdomainType (R : countIdomainType) := + [countIdomainType of polynomial R]. +Canonical poly_countIdomainType (R : countIdomainType) := + [countIdomainType of {poly R}]. + +Canonical int_countZmodType := [countZmodType of int]. +Canonical int_countRingType := [countRingType of int]. +Canonical int_countComRingType := [countComRingType of int]. +Canonical int_countUnitRingType := [countUnitRingType of int]. +Canonical int_countComUnitRingType := [countComUnitRingType of int]. +Canonical int_countIdomainType := [countIdomainType of int]. + +Canonical rat_countZmodType := [countZmodType of rat]. +Canonical rat_countRingType := [countRingType of rat]. +Canonical rat_countComRingType := [countComRingType of rat]. +Canonical rat_countUnitRingType := [countUnitRingType of rat]. +Canonical rat_countComUnitRingType := [countComUnitRingType of rat]. +Canonical rat_countIdomainType := [countIdomainType of rat]. +Canonical rat_countFieldType := [countFieldType of rat]. + +Lemma countable_field_extension (F : countFieldType) (p : {poly F}) : + size p > 1 -> + {E : countFieldType & {FtoE : {rmorphism F -> E} & + {w : E | root (map_poly FtoE p) w + & forall u : E, exists q, u = (map_poly FtoE q).[w]}}}. +Proof. +pose fix d i := + if i is i1.+1 then + let d1 := oapp (gcdp (d i1)) 0 (unpickle i1) in + if size d1 > 1 then d1 else d i1 + else p. +move=> p_gt1; have sz_d i: size (d i) > 1 by elim: i => //= i IHi; case: ifP. +have dv_d i j: i <= j -> d j %| d i. + move/subnK <-; elim: {j}(j - i)%N => //= j IHj; case: ifP => //=. + case: (unpickle _) => /= [q _|]; last by rewrite size_poly0. + exact: dvdp_trans (dvdp_gcdl _ _) IHj. +pose I : pred {poly F} := [pred q | d (pickle q).+1 %| q]. +have I'co q i: q \notin I -> i > pickle q -> coprimep q (d i). + rewrite inE => I'q /dv_d/coprimep_dvdl-> //; apply: contraR I'q. + rewrite coprimep_sym /coprimep /= pickleK /= neq_ltn. + case: ifP => [_ _| ->]; first exact: dvdp_gcdr. + rewrite orbF ltnS leqn0 size_poly_eq0 gcdp_eq0 -size_poly_eq0. + by rewrite -leqn0 leqNgt ltnW //. +have memI q: reflect (exists i, d i %| q) (q \in I). + apply: (iffP idP) => [|[i dv_di_q]]; first by exists (pickle q).+1. + have [le_i_q | /I'co i_co_q] := leqP i (pickle q). + rewrite inE /= pickleK /=; case: ifP => _; first exact: dvdp_gcdr. + exact: dvdp_trans (dv_d _ _ le_i_q) dv_di_q. + apply: contraR i_co_q _. + by rewrite /coprimep (eqp_size (dvdp_gcd_idr dv_di_q)) neq_ltn sz_d orbT. +have I_ideal : idealr_closed I. + split=> [||a q1 q2 Iq1 Iq2]; first exact: dvdp0. + by apply/memI=> [[i /idPn[]]]; rewrite dvdp1 neq_ltn sz_d orbT. + apply/memI; exists (maxn (pickle q1).+1 (pickle q2).+1); apply: dvdp_add. + by apply: dvdp_mull; apply: dvdp_trans Iq1; apply/dv_d/leq_maxl. + by apply: dvdp_trans Iq2; apply/dv_d/leq_maxr. +pose Iaddkey := GRing.Pred.Add (DefaultPredKey I) I_ideal. +pose Iidkey := MkIdeal (GRing.Pred.Zmod Iaddkey I_ideal) I_ideal. +pose E := ComRingType _ (@Quotient.mulqC _ _ _ (KeyedPred Iidkey)). +pose PtoE : {rmorphism {poly F} -> E} := [rmorphism of \pi_E%qT : {poly F} -> E]. +have PtoEd i: PtoE (d i) = 0. + by apply/eqP; rewrite piE Quotient.equivE subr0; apply/memI; exists i. +pose Einv (z : E) (q := repr z) (dq := d (pickle q).+1) := + let q_unitP := Bezout_eq1_coprimepP q dq in + if q_unitP is ReflectT ex_uv then PtoE (sval (sig_eqW ex_uv)).1 else 0. +have Einv0: Einv 0 = 0. + rewrite /Einv; case: Bezout_eq1_coprimepP => // ex_uv. + case/negP: (oner_neq0 E); rewrite piE -[_ 1]/(PtoE 1); have [uv <-] := ex_uv. + by rewrite rmorphD !rmorphM PtoEd /= reprK !mulr0 addr0. +have EmulV: GRing.Field.axiom Einv. + rewrite /Einv=> z nz_z; case: Bezout_eq1_coprimepP => [ex_uv |]; last first. + move/Bezout_eq1_coprimepP; rewrite I'co //. + by rewrite piE -{1}[z]reprK -Quotient.idealrBE subr0 in nz_z. + apply/eqP; case: sig_eqW => {ex_uv} [uv uv1]; set i := _.+1 in uv1 *. + rewrite piE /= -[z]reprK -(rmorphM PtoE) -Quotient.idealrBE. + by rewrite -uv1 opprD addNKr -mulNr; apply/memI; exists i; exact: dvdp_mull. +pose EringU := [comUnitRingType of UnitRingType _ (FieldUnitMixin EmulV Einv0)]. +have Eunitf := @FieldMixin _ _ EmulV Einv0. +pose Efield := FieldType (IdomainType EringU (FieldIdomainMixin Eunitf)) Eunitf. +pose Ecount := CountType Efield (CanCountMixin (@reprK _ _)). +pose FtoE := [rmorphism of PtoE \o polyC]; pose w : E := PtoE 'X. +have defPtoE q: (map_poly FtoE q).[w] = PtoE q. + by rewrite map_poly_comp horner_map [_.['X]]comp_polyXr. +exists [countFieldType of Ecount], FtoE, w => [|u]. + by rewrite /root defPtoE (PtoEd 0%N). +by exists (repr u); rewrite defPtoE /= reprK. +Qed. + +Lemma countable_algebraic_closure (F : countFieldType) : + {K : countClosedFieldType & {FtoK : {rmorphism F -> K} | integralRange FtoK}}. +Proof. +pose minXp (R : ringType) (p : {poly R}) := if size p > 1 then p else 'X. +have minXp_gt1 R p: size (minXp R p) > 1. + by rewrite /minXp; case: ifP => // _; rewrite size_polyX. +have minXpE (R : ringType) (p : {poly R}) : size p > 1 -> minXp R p = p. + by rewrite /minXp => ->. +have ext1 p := countable_field_extension (minXp_gt1 _ p). +pose ext1fT E p := tag (ext1 E p). +pose ext1to E p : {rmorphism _ -> ext1fT E p} := tag (tagged (ext1 E p)). +pose ext1w E p : ext1fT E p := s2val (tagged (tagged (ext1 E p))). +have ext1root E p: root (map_poly (ext1to E p) (minXp E p)) (ext1w E p). + by rewrite /ext1w; case: (tagged (tagged (ext1 E p))). +have ext1gen E p u: {q | u = (map_poly (ext1to E p) q).[ext1w E p]}. + by apply: sig_eqW; rewrite /ext1w; case: (tagged (tagged (ext1 E p))) u. +pose pExtEnum (E : countFieldType) := nat -> {poly E}. +pose Ext := {E : countFieldType & pExtEnum E}; pose MkExt : Ext := Tagged _ _. +pose EtoInc (E : Ext) i := ext1to (tag E) (tagged E i). +pose incEp E i j := + let v := map_poly (EtoInc E i) (tagged E j) in + if decode j is [:: i1; k] then + if i1 == i then odflt v (unpickle k) else v + else v. +pose fix E_ i := if i is i1.+1 then MkExt _ (incEp (E_ i1) i1) else MkExt F \0. +pose E i := tag (E_ i); pose Krep := {i : nat & E i}. +pose fix toEadd i k : {rmorphism E i -> E (k + i)%N} := + if k is k1.+1 then [rmorphism of EtoInc _ (k1 + i)%N \o toEadd _ _] + else [rmorphism of idfun]. +pose toE i j (le_ij : i <= j) := + ecast j {rmorphism E i -> E j} (subnK le_ij) (toEadd i (j - i)%N). +have toEeq i le_ii: toE i i le_ii =1 id. + by rewrite /toE; move: (subnK _); rewrite subnn => ?; rewrite eq_axiomK. +have toEleS i j leij leiSj z: toE i j.+1 leiSj z = EtoInc _ _ (toE i j leij z). + rewrite /toE; move: (j - i)%N {leij leiSj}(subnK _) (subnK _) => k. + by case: j /; rewrite (addnK i k.+1) => eq_kk; rewrite [eq_kk]eq_axiomK. +have toEirr := congr1 ((toE _ _)^~ _) (bool_irrelevance _ _). +have toEtrans j i k leij lejk leik z: + toE i k leik z = toE j k lejk (toE i j leij z). +- elim: k leik lejk => [|k IHk] leiSk lejSk. + by case: j => // in leij lejSk *; rewrite toEeq. + have:= lejSk; rewrite {1}leq_eqVlt ltnS => /predU1P[Dk | lejk]. + by rewrite -Dk in leiSk lejSk *; rewrite toEeq. + by have leik := leq_trans leij lejk; rewrite !toEleS -IHk. +have [leMl leMr] := (leq_maxl, leq_maxr); pose le_max := (leq_max, leqnn, orbT). +pose pairK (x y : Krep) (m := maxn _ _) := + (toE _ m (leMl _ _) (tagged x), toE _ m (leMr _ _) (tagged y)). +pose eqKrep x y := prod_curry (@eq_op _) (pairK x y). +have eqKrefl : reflexive eqKrep by move=> z; apply/eqP; apply: toEirr. +have eqKsym : symmetric eqKrep. + move=> z1 z2; rewrite {1}/eqKrep /= eq_sym; move: (leMl _ _) (leMr _ _). + by rewrite maxnC => lez1m lez2m; congr (_ == _); apply: toEirr. +have eqKtrans : transitive eqKrep. + rewrite /eqKrep /= => z2 z1 z3 /eqP eq_z12 /eqP eq_z23. + rewrite -(inj_eq (fmorph_inj (toE _ _ (leMr (tag z2) _)))). + rewrite -!toEtrans ?le_max // maxnCA maxnA => lez3m lez1m. + rewrite {lez1m}(toEtrans (maxn (tag z1) (tag z2))) // {}eq_z12. + do [rewrite -toEtrans ?le_max // -maxnA => lez2m] in lez3m *. + by rewrite (toEtrans (maxn (tag z2) (tag z3))) // eq_z23 -toEtrans. +pose K := {eq_quot (EquivRel _ eqKrefl eqKsym eqKtrans)}%qT. +have cntK : Countable.mixin_of K := CanCountMixin (@reprK _ _). +pose EtoKrep i (x : E i) : K := \pi%qT (Tagged E x). +have [EtoK piEtoK]: {EtoK | forall i, EtoKrep i =1 EtoK i} by exists EtoKrep. +pose FtoK := EtoK 0%N; rewrite {}/EtoKrep in piEtoK. +have eqEtoK i j x y: + toE i _ (leMl i j) x = toE j _ (leMr i j) y -> EtoK i x = EtoK j y. +- by move/eqP=> eq_xy; rewrite -!piEtoK; apply/eqmodP. +have toEtoK j i leij x : EtoK j (toE i j leij x) = EtoK i x. + by apply: eqEtoK; rewrite -toEtrans. +have EtoK_0 i: EtoK i 0 = FtoK 0 by apply: eqEtoK; rewrite !rmorph0. +have EtoK_1 i: EtoK i 1 = FtoK 1 by apply: eqEtoK; rewrite !rmorph1. +have EtoKeq0 i x: (EtoK i x == FtoK 0) = (x == 0). + by rewrite /FtoK -!piEtoK eqmodE /= /eqKrep /= rmorph0 fmorph_eq0. +have toErepr m i leim x lerm: + toE _ m lerm (tagged (repr (EtoK i x))) = toE i m leim x. +- have: (Tagged E x == repr (EtoK i x) %[mod K])%qT by rewrite reprK piEtoK. + rewrite eqmodE /= /eqKrep; case: (repr _) => j y /= in lerm * => /eqP /=. + have leijm: maxn i j <= m by rewrite geq_max leim. + by move/(congr1 (toE _ _ leijm)); rewrite -!toEtrans. +pose Kadd (x y : K) := EtoK _ (prod_curry +%R (pairK (repr x) (repr y))). +pose Kopp (x : K) := EtoK _ (- tagged (repr x)). +pose Kmul (x y : K) := EtoK _ (prod_curry *%R (pairK (repr x) (repr y))). +pose Kinv (x : K) := EtoK _ (tagged (repr x))^-1. +have EtoK_D i: {morph EtoK i : x y / x + y >-> Kadd x y}. + move=> x y; apply: eqEtoK; set j := maxn (tag _) _; rewrite !rmorphD. + by rewrite -!toEtrans ?le_max // => lexm leym; rewrite !toErepr. +have EtoK_N i: {morph EtoK i : x / - x >-> Kopp x}. + by move=> x; apply: eqEtoK; set j := tag _; rewrite !rmorphN toErepr. +have EtoK_M i: {morph EtoK i : x y / x * y >-> Kmul x y}. + move=> x y; apply: eqEtoK; set j := maxn (tag _) _; rewrite !rmorphM. + by rewrite -!toEtrans ?le_max // => lexm leym; rewrite !toErepr. +have EtoK_V i: {morph EtoK i : x / x^-1 >-> Kinv x}. + by move=> x; apply: eqEtoK; set j := tag _; rewrite !fmorphV toErepr. +case: {toErepr}I in (Kadd) (Kopp) (Kmul) (Kinv) EtoK_D EtoK_N EtoK_M EtoK_V. +pose inEi i z := {x : E i | z = EtoK i x}; have KtoE z: {i : nat & inEi i z}. + by elim/quotW: z => [[i x] /=]; exists i, x; rewrite piEtoK. +have inEle i j z: i <= j -> inEi i z -> inEi j z. + by move=> leij [x ->]; exists (toE i j leij x); rewrite toEtoK. +have KtoE2 z1 z2: {i : nat & inEi i z1 & inEi i z2}. + have [[i1 Ez1] [i2 Ez2]] := (KtoE z1, KtoE z2). + by exists (maxn i1 i2); [apply: inEle Ez1 | apply: inEle Ez2]. +have KtoE3 z1 z2 z3: {i : nat & inEi i z1 & inEi i z2 * inEi i z3}%type. + have [[i1 Ez1] [i2 Ez2 Ez3]] := (KtoE z1, KtoE2 z2 z3). + by exists (maxn i1 i2); [apply: inEle Ez1 | split; apply: inEle (leMr _ _) _]. +have KaddC: commutative Kadd. + by move=> u v; have [i [x ->] [y ->]] := KtoE2 u v; rewrite -!EtoK_D addrC. +have KaddA: associative Kadd. + move=> u v w; have [i [x ->] [[y ->] [z ->]]] := KtoE3 u v w. + by rewrite -!EtoK_D addrA. +have Kadd0: left_id (FtoK 0) Kadd. + by move=> u; have [i [x ->]] := KtoE u; rewrite -(EtoK_0 i) -EtoK_D add0r. +have KaddN: left_inverse (FtoK 0) Kopp Kadd. + by move=> u; have [i [x ->]] := KtoE u; rewrite -EtoK_N -EtoK_D addNr EtoK_0. +pose Kzmod := ZmodType K (ZmodMixin KaddA KaddC Kadd0 KaddN). +have KmulC: commutative Kmul. + by move=> u v; have [i [x ->] [y ->]] := KtoE2 u v; rewrite -!EtoK_M mulrC. +have KmulA: @associative Kzmod Kmul. + move=> u v w; have [i [x ->] [[y ->] [z ->]]] := KtoE3 u v w. + by rewrite -!EtoK_M mulrA. +have Kmul1: left_id (FtoK 1) Kmul. + by move=> u; have [i [x ->]] := KtoE u; rewrite -(EtoK_1 i) -EtoK_M mul1r. +have KmulD: left_distributive Kmul Kadd. + move=> u v w; have [i [x ->] [[y ->] [z ->]]] := KtoE3 u v w. + by rewrite -!(EtoK_M, EtoK_D) mulrDl. +have Kone_nz: FtoK 1 != FtoK 0 by rewrite EtoKeq0 oner_neq0. +pose KringMixin := ComRingMixin KmulA KmulC Kmul1 KmulD Kone_nz. +pose Kring := ComRingType (RingType Kzmod KringMixin) KmulC. +have KmulV: @GRing.Field.axiom Kring Kinv. + move=> u; have [i [x ->]] := KtoE u; rewrite EtoKeq0 => nz_x. + by rewrite -EtoK_V -[_ * _]EtoK_M mulVf ?EtoK_1. +have Kinv0: Kinv (FtoK 0) = FtoK 0 by rewrite -EtoK_V invr0. +pose Kuring := [comUnitRingType of UnitRingType _ (FieldUnitMixin KmulV Kinv0)]. +pose KfieldMixin := @FieldMixin _ _ KmulV Kinv0. +pose Kidomain := IdomainType Kuring (FieldIdomainMixin KfieldMixin). +pose Kfield := FieldType Kidomain KfieldMixin. +have EtoKrmorphism i: rmorphism (EtoK i : E i -> Kfield). + by do 2?split=> [x y|]; rewrite ?EtoK_D ?EtoK_N ?EtoK_M ?EtoK_1. +pose EtoKM := RMorphism (EtoKrmorphism _); have EtoK_E: EtoK _ = EtoKM _ by []. +have toEtoKp := @eq_map_poly _ Kring _ _(toEtoK _ _ _). +have Kclosed: GRing.ClosedField.axiom Kfield. + move=> n pK n_gt0; pose m0 := \max_(i < n) tag (KtoE (pK i)); pose m := m0.+1. + have /fin_all_exists[pE DpE] (i : 'I_n): exists y, EtoK m y = pK i. + pose u := KtoE (pK i); have leum0: tag u <= m0 by rewrite (bigmax_sup i). + by have [y ->] := tagged u; exists (toE _ _ (leqW leum0) y); rewrite toEtoK. + pose p := 'X^n - rVpoly (\row_i pE i); pose j := code [:: m0; pickle p]. + pose pj := tagged (E_ j) j; pose w : E j.+1 := ext1w (E j) pj. + have lemj: m <= j by rewrite (allP (ltn_code _)) ?mem_head. + exists (EtoKM j.+1 w); apply/eqP; rewrite -subr_eq0; apply/eqP. + transitivity (EtoKM j.+1 (map_poly (toE m j.+1 (leqW lemj)) p).[w]). + rewrite -horner_map -map_poly_comp toEtoKp EtoK_E; move/EtoKM: w => w. + rewrite rmorphB [_ 'X^n]map_polyXn !hornerE hornerXn; congr (_ - _ : Kring). + rewrite (@horner_coef_wide _ n) ?size_map_poly ?size_poly //. + by apply: eq_bigr => i _; rewrite coef_map coef_rVpoly valK mxE /= DpE. + suffices Dpj: map_poly (toE m j lemj) p = pj. + apply/eqP; rewrite EtoKeq0 (eq_map_poly (toEleS _ _ _ _)) map_poly_comp Dpj. + rewrite -rootE -[pj]minXpE ?ext1root // -Dpj size_map_poly. + by rewrite size_addl ?size_polyXn ltnS ?size_opp ?size_poly. + rewrite {w}/pj; elim: {-9}j lemj => // k IHk lemSk. + move: lemSk (lemSk); rewrite {1}leq_eqVlt ltnS => /predU1P[<- | lemk] lemSk. + rewrite {k IHk lemSk}(eq_map_poly (toEeq m _)) map_poly_id //= /incEp. + by rewrite codeK eqxx pickleK. + rewrite (eq_map_poly (toEleS _ _ _ _)) map_poly_comp {}IHk //= /incEp codeK. + by rewrite -if_neg neq_ltn lemk. +suffices{Kclosed} algF_K: {FtoK : {rmorphism F -> Kfield} | integralRange FtoK}. + pose Kdec := DecFieldType Kfield (closed_fields_QEMixin Kclosed). + pose KclosedField := ClosedFieldType Kdec Kclosed. + by exists [countClosedFieldType of CountType KclosedField cntK]. +exists (EtoKM 0%N) => /= z; have [i [{z}z ->]] := KtoE z. +suffices{z} /(_ z)[p mon_p]: integralRange (toE 0%N i isT). + by rewrite -(fmorph_root (EtoKM i)) -map_poly_comp toEtoKp; exists p. +rewrite /toE /E; clear - minXp_gt1 ext1root ext1gen. +move: (i - 0)%N (subnK _) => n; case: i /. +elim: n => [|n IHn] /= z; first exact: integral_id. +have{z} [q ->] := ext1gen _ _ z; set pn := tagged (E_ _) _. +apply: integral_horner. + by apply/integral_poly=> i; rewrite coef_map; apply: integral_rmorph. +apply: integral_root (ext1root _ _) _. + by rewrite map_poly_eq0 -size_poly_gt0 ltnW. +by apply/integral_poly=> i; rewrite coef_map; apply: integral_rmorph. +Qed. diff --git a/mathcomp/field/cyclotomic.v b/mathcomp/field/cyclotomic.v new file mode 100644 index 0000000..edd33c2 --- /dev/null +++ b/mathcomp/field/cyclotomic.v @@ -0,0 +1,320 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import fintype tuple finfun bigop prime ssralg poly finset. +Require Import fingroup finalg zmodp cyclic. +Require Import ssrnum ssrint polydiv rat intdiv. +Require Import mxpoly vector falgebra fieldext separable galois algC. + +(******************************************************************************) +(* This file provides few basic properties of cyclotomic polynomials. *) +(* We define: *) +(* cyclotomic z n == the factorization of the nth cyclotomic polynomial in *) +(* a ring R in which z is an nth primitive root of unity. *) +(* 'Phi_n == the nth cyclotomic polynomial in int. *) +(* This library is quite limited, and should be extended in the future. In *) +(* particular the irreducibity of 'Phi_n is only stated indirectly, as the *) +(* fact that its embedding in the algebraics (algC) is the minimal polynomial *) +(* of an nth primitive root of unity. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GRing.Theory Num.Theory. +Local Open Scope ring_scope. + +Section CyclotomicPoly. + +Section Ring. + +Variable R : ringType. + +Definition cyclotomic (z : R) n := + \prod_(k < n | coprime k n) ('X - (z ^+ k)%:P). + +Lemma cyclotomic_monic z n : cyclotomic z n \is monic. +Proof. exact: monic_prod_XsubC. Qed. + +Lemma size_cyclotomic z n : size (cyclotomic z n) = (totient n).+1. +Proof. +rewrite /cyclotomic -big_filter filter_index_enum size_prod_XsubC; congr _.+1. +rewrite -cardE -sum1_card totient_count_coprime -big_mkcond big_mkord. +by apply: eq_bigl => k; rewrite coprime_sym. +Qed. + +End Ring. + +Lemma separable_Xn_sub_1 (R : idomainType) n : + n%:R != 0 :> R -> @separable_poly R ('X^n - 1). +Proof. +case: n => [/eqP// | n nz_n]; rewrite /separable_poly linearB /=. +rewrite derivC subr0 derivXn -scaler_nat coprimep_scaler //= exprS -scaleN1r. +rewrite coprimep_sym coprimep_addl_mul coprimep_scaler ?coprimep1 //. +by rewrite (signr_eq0 _ 1). +Qed. + +Section Field. + +Variables (F : fieldType) (n : nat) (z : F). +Hypothesis prim_z : n.-primitive_root z. +Let n_gt0 := prim_order_gt0 prim_z. + +Lemma root_cyclotomic x : root (cyclotomic z n) x = n.-primitive_root x. +Proof. +rewrite /cyclotomic -big_filter filter_index_enum. +rewrite -(big_map _ xpredT (fun y => 'X - y%:P)) root_prod_XsubC. +apply/imageP/idP=> [[k co_k_n ->] | prim_x]. + by rewrite prim_root_exp_coprime. +have [k Dx] := prim_rootP prim_z (prim_expr_order prim_x). +exists (Ordinal (ltn_pmod k n_gt0)) => /=. + by rewrite unfold_in /= coprime_modl -(prim_root_exp_coprime k prim_z) -Dx. +by rewrite prim_expr_mod. +Qed. + +Lemma prod_cyclotomic : + 'X^n - 1 = \prod_(d <- divisors n) cyclotomic (z ^+ (n %/ d)) d. +Proof. +have in_d d: (d %| n)%N -> val (@inord n d) = d by move/dvdn_leq/inordK=> /= ->. +have dv_n k: (n %/ gcdn k n %| n)%N. + by rewrite -{3}(divnK (dvdn_gcdr k n)) dvdn_mulr. +have [uDn _ inDn] := divisors_correct n_gt0. +have defDn: divisors n = map val (map (@inord n) (divisors n)). + by rewrite -map_comp map_id_in // => d; rewrite inDn => /in_d. +rewrite defDn big_map big_uniq /=; last first. + by rewrite -(map_inj_uniq val_inj) -defDn. +pose h (k : 'I_n) : 'I_n.+1 := inord (n %/ gcdn k n). +rewrite -(factor_Xn_sub_1 prim_z) big_mkord. +rewrite (partition_big h (dvdn^~ n)) /= => [|k _]; last by rewrite in_d ?dv_n. +apply: eq_big => d; first by rewrite -(mem_map val_inj) -defDn inDn. +set q := (n %/ d)%N => d_dv_n. +have [q_gt0 d_gt0]: (0 < q /\ 0 < d)%N by apply/andP; rewrite -muln_gt0 divnK. +have fP (k : 'I_d): (q * k < n)%N by rewrite divn_mulAC ?ltn_divLR ?ltn_pmul2l. +rewrite (reindex (fun k => Ordinal (fP k))); last first. + have f'P (k : 'I_n): (k %/ q < d)%N by rewrite ltn_divLR // mulnC divnK. + exists (fun k => Ordinal (f'P k)) => [k _ | k /eqnP/=]. + by apply: val_inj; rewrite /= mulKn. + rewrite in_d // => Dd; apply: val_inj; rewrite /= mulnC divnK // /q -Dd. + by rewrite divnA ?mulKn ?dvdn_gcdl ?dvdn_gcdr. +apply: eq_big => k; rewrite ?exprM // -val_eqE in_d //=. +rewrite -eqn_mul ?dvdn_gcdr ?gcdn_gt0 ?n_gt0 ?orbT //. +rewrite -[n in gcdn _ n](divnK d_dv_n) -muln_gcdr mulnCA mulnA divnK //. +by rewrite mulnC eqn_mul // divnn n_gt0 eq_sym. +Qed. + +End Field. + +End CyclotomicPoly. + +Local Notation ZtoQ := (intr : int -> rat). +Local Notation ZtoC := (intr : int -> algC). +Local Notation QtoC := (ratr : rat -> algC). + +Local Notation intrp := (map_poly intr). +Local Notation pZtoQ := (map_poly ZtoQ). +Local Notation pZtoC := (map_poly ZtoC). +Local Notation pQtoC := (map_poly ratr). + +Local Hint Resolve (@intr_inj [numDomainType of algC]). +Local Notation QtoC_M := (ratr_rmorphism [numFieldType of algC]). + +Lemma C_prim_root_exists n : (n > 0)%N -> {z : algC | n.-primitive_root z}. +Proof. +pose p : {poly algC} := 'X^n - 1; have [r Dp] := closed_field_poly_normal p. +move=> n_gt0; apply/sigW; rewrite (monicP _) ?monic_Xn_sub_1 // scale1r in Dp. +have rn1: all n.-unity_root r by apply/allP=> z; rewrite -root_prod_XsubC -Dp. +have sz_r: (n < (size r).+1)%N. + by rewrite -(size_prod_XsubC r id) -Dp size_Xn_sub_1. +have [|z] := hasP (has_prim_root n_gt0 rn1 _ sz_r); last by exists z. +by rewrite -separable_prod_XsubC -Dp separable_Xn_sub_1 // pnatr_eq0 -lt0n. +Qed. + +(* (Integral) Cyclotomic polynomials. *) + +Definition Cyclotomic n : {poly int} := + let: exist z _ := C_prim_root_exists (ltn0Sn n.-1) in + map_poly floorC (cyclotomic z n). + +Notation "''Phi_' n" := (Cyclotomic n) + (at level 8, n at level 2, format "''Phi_' n"). + +Lemma Cyclotomic_monic n : 'Phi_n \is monic. +Proof. +rewrite /'Phi_n; case: (C_prim_root_exists _) => z /= _. +rewrite monicE lead_coefE coef_map_id0 ?(int_algC_K 0) ?getCint0 //. +by rewrite size_poly_eq -lead_coefE (monicP (cyclotomic_monic _ _)) (intCK 1). +Qed. + +Lemma Cintr_Cyclotomic n z : + n.-primitive_root z -> pZtoC 'Phi_n = cyclotomic z n. +Proof. +elim: {n}_.+1 {-2}n z (ltnSn n) => // m IHm n z0 le_mn prim_z0. +rewrite /'Phi_n; case: (C_prim_root_exists _) => z /=. +have n_gt0 := prim_order_gt0 prim_z0; rewrite prednK // => prim_z. +have [uDn _ inDn] := divisors_correct n_gt0. +pose q := \prod_(d <- rem n (divisors n)) 'Phi_d. +have mon_q: q \is monic by apply: monic_prod => d _; exact: Cyclotomic_monic. +have defXn1: cyclotomic z n * pZtoC q = 'X^n - 1. + rewrite (prod_cyclotomic prim_z) (big_rem n) ?inDn //=. + rewrite divnn n_gt0 rmorph_prod /=; congr (_ * _). + apply: eq_big_seq => d; rewrite mem_rem_uniq ?inE //= inDn => /andP[n'd ddvn]. + rewrite -IHm ?dvdn_prim_root // -ltnS (leq_ltn_trans _ le_mn) //. + by rewrite ltn_neqAle n'd dvdn_leq. +have mapXn1 (R1 R2 : ringType) (f : {rmorphism R1 -> R2}): + map_poly f ('X^n - 1) = 'X^n - 1. +- by rewrite rmorphB /= rmorph1 map_polyXn. +have nz_q: pZtoC q != 0. + by rewrite -size_poly_eq0 size_map_inj_poly // size_poly_eq0 monic_neq0. +have [r def_zn]: exists r, cyclotomic z n = pZtoC r. + have defZtoC: ZtoC =1 QtoC \o ZtoQ by move=> a; rewrite /= rmorph_int. + have /dvdpP[r0 Dr0]: map_poly ZtoQ q %| 'X^n - 1. + rewrite -(dvdp_map QtoC_M) mapXn1 -map_poly_comp. + by rewrite -(eq_map_poly defZtoC) -defXn1 dvdp_mull. + have [r [a nz_a Dr]] := rat_poly_scale r0. + exists (zprimitive r); apply: (mulIf nz_q); rewrite defXn1. + rewrite -rmorphM -(zprimitive_monic mon_q) -zprimitiveM /=. + have ->: r * q = a *: ('X^n - 1). + apply: (map_inj_poly (intr_inj : injective ZtoQ)) => //. + rewrite map_polyZ mapXn1 Dr0 Dr -scalerAl scalerKV ?intr_eq0 //. + by rewrite rmorphM. + by rewrite zprimitiveZ // zprimitive_monic ?monic_Xn_sub_1 ?mapXn1. +rewrite floorCpK; last by apply/polyOverP=> i; rewrite def_zn coef_map Cint_int. +pose f e (k : 'I_n) := Ordinal (ltn_pmod (k * e) n_gt0). +have [e Dz0] := prim_rootP prim_z (prim_expr_order prim_z0). +have co_e_n: coprime e n by rewrite -(prim_root_exp_coprime e prim_z) -Dz0. +have injf: injective (f e). + apply: can_inj (f (egcdn e n).1) _ => k; apply: val_inj => /=. + rewrite modnMml -mulnA -modnMmr -{1}(mul1n e). + by rewrite (chinese_modr co_e_n 0) modnMmr muln1 modn_small. +rewrite [_ n](reindex_inj injf); apply: eq_big => k /=. + by rewrite coprime_modl coprime_mull co_e_n andbT. +by rewrite prim_expr_mod // mulnC exprM -Dz0. +Qed. + +Lemma prod_Cyclotomic n : + (n > 0)%N -> \prod_(d <- divisors n) 'Phi_d = 'X^n - 1. +Proof. +move=> n_gt0; have [z prim_z] := C_prim_root_exists n_gt0. +apply: (map_inj_poly (intr_inj : injective ZtoC)) => //. +rewrite rmorphB rmorph1 rmorph_prod /= map_polyXn (prod_cyclotomic prim_z). +apply: eq_big_seq => d; rewrite -dvdn_divisors // => d_dv_n. +by rewrite -Cintr_Cyclotomic ?dvdn_prim_root. +Qed. + +Lemma Cyclotomic0 : 'Phi_0 = 1. +Proof. +rewrite /'Phi_0; case: (C_prim_root_exists _) => z /= _. +by rewrite -[1]polyseqK /cyclotomic big_ord0 map_polyE !polyseq1 /= (intCK 1). +Qed. + +Lemma size_Cyclotomic n : size 'Phi_n = (totient n).+1. +Proof. +have [-> | n_gt0] := posnP n; first by rewrite Cyclotomic0 polyseq1. +have [z prim_z] := C_prim_root_exists n_gt0. +rewrite -(size_map_inj_poly (can_inj intCK)) //. +rewrite (Cintr_Cyclotomic prim_z) -[_ n]big_filter filter_index_enum. +rewrite size_prod_XsubC -cardE totient_count_coprime big_mkord -big_mkcond /=. +by rewrite (eq_card (fun _ => coprime_sym _ _)) sum1_card. +Qed. + +Lemma minCpoly_cyclotomic n z : + n.-primitive_root z -> minCpoly z = cyclotomic z n. +Proof. +move=> prim_z; have n_gt0 := prim_order_gt0 prim_z. +have Dpz := Cintr_Cyclotomic prim_z; set pz := cyclotomic z n in Dpz *. +have mon_pz: pz \is monic by exact: cyclotomic_monic. +have pz0: root pz z by rewrite root_cyclotomic. +have [pf [Dpf mon_pf] dv_pf] := minCpolyP z. +have /dvdpP_rat_int[f [af nz_af Df] [g /esym Dfg]]: pf %| pZtoQ 'Phi_n. + rewrite -dv_pf; congr (root _ z): pz0; rewrite -Dpz -map_poly_comp. + by apply: eq_map_poly => b; rewrite /= rmorph_int. +without loss{nz_af} [mon_f mon_g]: af f g Df Dfg / f \is monic /\ g \is monic. + move=> IH; pose cf := lead_coef f; pose cg := lead_coef g. + have cfg1: cf * cg = 1. + by rewrite -lead_coefM Dfg (monicP (Cyclotomic_monic n)). + apply: (IH (af *~ cf) (f *~ cg) (g *~ cf)). + - by rewrite rmorphMz -scalerMzr scalerMzl -mulrzA cfg1. + - by rewrite mulrzAl mulrzAr -mulrzA cfg1. + by rewrite !(intz, =^~ scaler_int) !monicE !lead_coefZ mulrC cfg1. +have{af Df} Df: pQtoC pf = pZtoC f. + have:= congr1 lead_coef Df. + rewrite lead_coefZ lead_coef_map_inj //; last exact: intr_inj. + rewrite !(monicP _) // mulr1 Df => <-; rewrite scale1r -map_poly_comp. + by apply: eq_map_poly => b; rewrite /= rmorph_int. +have [/size1_polyC Dg | g_gt1] := leqP (size g) 1. + rewrite monicE Dg lead_coefC in mon_g. + by rewrite -Dpz -Dfg Dg (eqP mon_g) mulr1 Dpf. +have [zk gzk0]: exists zk, root (pZtoC g) zk. + have [rg] := closed_field_poly_normal (pZtoC g). + rewrite lead_coef_map_inj // (monicP mon_g) scale1r => Dg. + rewrite -(size_map_inj_poly (can_inj intCK)) // Dg in g_gt1. + rewrite size_prod_XsubC in g_gt1. + by exists rg`_0; rewrite Dg root_prod_XsubC mem_nth. +have [k cokn Dzk]: exists2 k, coprime k n & zk = z ^+ k. + have: root pz zk by rewrite -Dpz -Dfg rmorphM rootM gzk0 orbT. + rewrite -[pz]big_filter -(big_map _ xpredT (fun a => 'X - a%:P)). + by rewrite root_prod_XsubC => /imageP[k]; exists k. +have co_fg (R : idomainType): n%:R != 0 :> R -> @coprimep R (intrp f) (intrp g). + move=> nz_n; have: separable_poly (intrp ('X^n - 1) : {poly R}). + by rewrite rmorphB rmorph1 /= map_polyXn separable_Xn_sub_1. + rewrite -prod_Cyclotomic // (big_rem n) -?dvdn_divisors //= -Dfg. + by rewrite !rmorphM /= !separable_mul => /and3P[] /and3P[]. +suffices fzk0: root (pZtoC f) zk. + have [] // := negP (coprimep_root (co_fg _ _) fzk0). + by rewrite pnatr_eq0 -lt0n. +move: gzk0 cokn; rewrite {zk}Dzk; elim: {k}_.+1 {-2}k (ltnSn k) => // m IHm k. +rewrite ltnS => lekm gzk0 cokn. +have [|k_gt1] := leqP k 1; last have [p p_pr /dvdnP[k1 Dk]] := pdivP k_gt1. + rewrite -[leq k 1](mem_iota 0 2) !inE => /pred2P[k0 | ->]; last first. + by rewrite -Df dv_pf. + have /eqP := size_Cyclotomic n; rewrite -Dfg size_Mmonic ?monic_neq0 //. + rewrite k0 /coprime gcd0n in cokn; rewrite (eqP cokn). + rewrite -(size_map_inj_poly (can_inj intCK)) // -Df -Dpf. + by rewrite -(subnKC g_gt1) -(subnKC (size_minCpoly z)) !addnS. +move: cokn; rewrite Dk coprime_mull => /andP[cok1n]. +rewrite prime_coprime // (dvdn_charf (char_Fp p_pr)) => /co_fg {co_fg}. +have charFpX: p \in [char {poly 'F_p}]. + by rewrite (rmorph_char (polyC_rmorphism _)) ?char_Fp. +rewrite -(coprimep_pexpr _ _ (prime_gt0 p_pr)) -(Frobenius_autE charFpX). +rewrite -[g]comp_polyXr map_comp_poly -horner_map /= Frobenius_autE -rmorphX. +rewrite -!map_poly_comp (@eq_map_poly _ _ _ (polyC \o *~%R 1)); last first. + by move=> a; rewrite /= !rmorph_int. +rewrite map_poly_comp -[_.[_]]map_comp_poly /= => co_fg. +suffices: coprimep (pZtoC f) (pZtoC (g \Po 'X^p)). + move/coprimep_root=> /=/(_ (z ^+ k1))/implyP. + rewrite map_comp_poly map_polyXn horner_comp hornerXn. + rewrite -exprM -Dk [_ == 0]gzk0 implybF => /negP[]. + have: root pz (z ^+ k1). + by rewrite root_cyclotomic // prim_root_exp_coprime. + rewrite -Dpz -Dfg rmorphM rootM => /orP[] //= /IHm-> //. + rewrite (leq_trans _ lekm) // -[k1]muln1 Dk ltn_pmul2l ?prime_gt1 //. + by have:= ltnW k_gt1; rewrite Dk muln_gt0 => /andP[]. +suffices: coprimep f (g \Po 'X^p). + case/Bezout_coprimepP=> [[u v]]; rewrite -size_poly_eq1. + rewrite -(size_map_inj_poly (can_inj intCK)) // rmorphD !rmorphM /=. + rewrite size_poly_eq1 => {co_fg}co_fg; apply/Bezout_coprimepP. + by exists (pZtoC u, pZtoC v). +apply: contraLR co_fg => /coprimepPn[|d]; first exact: monic_neq0. +rewrite andbC -size_poly_eq1 dvdp_gcd => /and3P[sz_d]. +pose d1 := zprimitive d. +have d_dv_mon h: d %| h -> h \is monic -> exists h1, h = d1 * h1. + case/Pdiv.Idomain.dvdpP=> [[c h1] /= nz_c Dh] mon_h; exists (zprimitive h1). + by rewrite -zprimitiveM mulrC -Dh zprimitiveZ ?zprimitive_monic. +case/d_dv_mon=> // f1 Df1 /d_dv_mon[|f2 ->]. + rewrite monicE lead_coefE size_comp_poly size_polyXn /=. + rewrite comp_polyE coef_sum polySpred ?monic_neq0 //= mulnC. + rewrite big_ord_recr /= -lead_coefE (monicP mon_g) scale1r. + rewrite -exprM coefXn eqxx big1 ?add0r // => i _. + rewrite coefZ -exprM coefXn eqn_pmul2l ?prime_gt0 //. + by rewrite eqn_leq leqNgt ltn_ord mulr0. +have monFp h: h \is monic -> size (map_poly intr h) = size h. + by move=> mon_h; rewrite size_poly_eq // -lead_coefE (monicP mon_h) oner_eq0. +apply/coprimepPn; last exists (map_poly intr d1). + by rewrite -size_poly_eq0 monFp // size_poly_eq0 monic_neq0. +rewrite Df1 !rmorphM dvdp_gcd !dvdp_mulr //= -size_poly_eq1. +rewrite monFp ?size_zprimitive //. +rewrite monicE [_ d1]intEsg sgz_lead_primitive -zprimitive_eq0 -/d1. +rewrite -lead_coef_eq0 -absz_eq0. +have/esym/eqP := congr1 (absz \o lead_coef) Df1. +by rewrite /= (monicP mon_f) lead_coefM abszM muln_eq1 => /andP[/eqP-> _]. +Qed. diff --git a/mathcomp/field/falgebra.v b/mathcomp/field/falgebra.v new file mode 100644 index 0000000..d7dfd85 --- /dev/null +++ b/mathcomp/field/falgebra.v @@ -0,0 +1,1199 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq path choice fintype. +Require Import div tuple finfun bigop ssralg finalg zmodp matrix vector poly. + +(******************************************************************************) +(* Finite dimensional free algebras, usually known as F-algebras. *) +(* FalgType K == the interface type for F-algebras over K; it simply *) +(* joins the unitAlgType K and vectType K interfaces. *) +(* [FalgType K of aT] == an FalgType K structure for a type aT that has both *) +(* unitAlgType K and vectType K canonical structures. *) +(* [FalgType K of aT for vT] == an FalgType K structure for a type aT with a *) +(* unitAlgType K canonical structure, given a structure *) +(* vT : vectType K whose lmodType K projection matches *) +(* the canonical lmodType for aT. *) +(* FalgUnitRingType T == a default unitRingType structure for a type T with *) +(* both algType and vectType structures. *) +(* Any aT with an FalgType structure inherits all the Vector, Ring and *) +(* Algebra operations, and supports the following additional operations: *) +(* \dim_A M == (\dim M %/ dim A)%N -- free module dimension. *) +(* amull u == the linear function v |-> u * v, for u, v : aT. *) +(* amulr u == the linear function v |-> v * u, for u, v : aT. *) +(* 1, f * g, f ^+ n == the identity function, the composite g \o f, the nth *) +(* iterate of f, for 1, f, g in 'End(aT). This is just *) +(* the usual F-algebra structure on 'End(aT). It is NOT *) +(* canonical by default, but can be activated by the *) +(* line Import FalgLfun. Beware also that (f^-1)%VF is *) +(* the linear function inverse, not the ring inverse of *) +(* f (though they do coincide when f is injective). *) +(* 1%VS == the line generated by 1 : aT. *) +(* (U * V)%VS == the smallest subspace of aT that contains all *) +(* products u * v for u in U, v in V. *) +(* (U ^+ n)%VS == (U * U * ... * U), n-times. U ^+ 0 = 1%VS *) +(* 'C[u]%VS == the centraliser subspace of the vector u. *) +(* 'C_U[v]%VS := (U :&: 'C[v])%VS. *) +(* 'C(V)%VS == the centraliser subspace of the subspace V. *) +(* 'C_U(V)%VS := (U :&: 'C(V))%VS. *) +(* 'Z(V)%VS == the center subspace of the subspace V. *) +(* agenv U == the smallest subalgebra containing U ^+ n for all n. *) +(* <<U; v>>%VS == agenv (U + <[v]>) (adjoin v to U). *) +(* <<U & vs>>%VS == agenv (U + <<vs>>) (adjoin vs to U). *) +(* {aspace aT} == a subType of {vspace aT} consisting of sub-algebras *) +(* of aT (see below); for A : {aspace aT}, subvs_of A *) +(* has a canonical FalgType K structure. *) +(* is_aspace U <=> the characteristic predicate of {aspace aT} stating *) +(* that U is closed under product and contains an *) +(* identity element, := has_algid U && (U * U <= U)%VS. *) +(* algid A == the identity element of A : {aspace aT}, which need *) +(* not be equal to 1 (indeed, in a Wedderburn *) +(* decomposition it is not even a unit in aT). *) +(* is_algid U e <-> e : aT is an identity element for the subspace U: *) +(* e in U, e != 0 & e * u = u * e = u for all u in U. *) +(* has_algid U <=> there is an e such that is_algid U e. *) +(* [aspace of U] == a clone of an existing {aspace aT} structure on *) +(* U : {vspace aT} (more instances of {aspace aT} will *) +(* be defined in extFieldType). *) +(* [aspace of U for A] == a clone of A : {aspace aT} for U : {vspace aT}. *) +(* 1%AS == the canonical sub-algebra 1%VS. *) +(* {:aT}%AS == the canonical full algebra. *) +(* <<U>>%AS == the canonical algebra for agenv U; note that this is *) +(* unrelated to <<vs>>%VS, the subspace spanned by vs. *) +(* <<U; v>>%AS == the canonical algebra for <<U; v>>%VS. *) +(* <<U & vs>>%AS == the canonical algebra for <<U & vs>>%VS. *) +(* ahom_in U f <=> f : 'Hom(aT, rT) is a multiplicative homomorphism *) +(* inside U, and in addition f 1 = 1 (even if U doesn't *) +(* contain 1). Note that f @: U need not be a *) +(* subalgebra when U is, as f could annilate U. *) +(* 'AHom(aT, rT) == the type of algebra homomorphisms from aT to rT, *) +(* where aT and rT ARE FalgType structures. Elements of *) +(* 'AHom(aT, rT) coerce to 'End(aT, rT) and aT -> rT. *) +(* --> Caveat: aT and rT must denote actual FalgType structures, not their *) +(* projections on Type. *) +(* 'AEnd(aT) == algebra endomorphisms of aT (:= 'AHom(aT, aT)). *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. +Open Local Scope ring_scope. + +Reserved Notation "{ 'aspace' T }" (at level 0, format "{ 'aspace' T }"). +Reserved Notation "<< U & vs >>" (at level 0, format "<< U & vs >>"). +Reserved Notation "<< U ; x >>" (at level 0, format "<< U ; x >>"). +Reserved Notation "''AHom' ( T , rT )" + (at level 8, format "''AHom' ( T , rT )"). +Reserved Notation "''AEnd' ( T )" (at level 8, format "''AEnd' ( T )"). + +Notation "\dim_ E V" := (divn (\dim V) (\dim E)) + (at level 10, E at level 2, V at level 8, format "\dim_ E V") : nat_scope. + +Import GRing.Theory. + +(* Finite dimensional algebra *) +Module Falgebra. + +(* Supply a default unitRing mixin for the default unitAlgType base type. *) +Section DefaultBase. + +Variables (K : fieldType) (A : algType K). + +Lemma BaseMixin : Vector.mixin_of A -> GRing.UnitRing.mixin_of A. +Proof. +move=> vAm; pose vA := VectType K A vAm. +pose am u := linfun (u \o* idfun : vA -> vA). +have amE u v : am u v = v * u by rewrite lfunE. +pose uam := [pred u | lker (am u) == 0%VS]. +pose vam := [fun u => if u \in uam then (am u)^-1%VF 1 else u]. +have vamKl: {in uam, left_inverse 1 vam *%R}. + by move=> u Uu; rewrite /= Uu -amE lker0_lfunVK. +exists uam vam => // [u Uu | u v [_ uv1] | u /negbTE/= -> //]. + by apply/(lker0P Uu); rewrite !amE -mulrA vamKl // mul1r mulr1. +by apply/lker0P=> w1 w2 /(congr1 (am v)); rewrite !amE -!mulrA uv1 !mulr1. +Qed. + +Definition BaseType T := + fun c vAm & phant_id c (GRing.UnitRing.Class (BaseMixin vAm)) => + fun (vT : vectType K) & phant vT + & phant_id (Vector.mixin (Vector.class vT)) vAm => + @GRing.UnitRing.Pack T c T. + +End DefaultBase. + +Section ClassDef. +Variable R : ringType. +Implicit Type phR : phant R. + +Record class_of A := Class { + base1 : GRing.UnitAlgebra.class_of R A; + mixin : Vector.mixin_of (GRing.Lmodule.Pack _ base1 A) +}. +Local Coercion base1 : class_of >-> GRing.UnitAlgebra.class_of. +Definition base2 A c := @Vector.Class _ _ (@base1 A c) (mixin c). +Local Coercion base2 : class_of >-> Vector.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. +Let xT := let: Pack T _ _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack := + fun bT b & phant_id (@GRing.UnitAlgebra.class R phR bT) + (b : GRing.UnitAlgebra.class_of R T) => + fun mT m & phant_id (@Vector.class R phR mT) (@Vector.Class R T b 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 := @GRing.Zmodule.Pack cT xclass xT. +Definition lmodType := @GRing.Lmodule.Pack R phR cT xclass xT. +Definition ringType := @GRing.Ring.Pack cT xclass xT. +Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT. +Definition lalgType := @GRing.Lalgebra.Pack R phR cT xclass xT. +Definition algType := @GRing.Algebra.Pack R phR cT xclass xT. +Definition unitAlgType := @GRing.UnitAlgebra.Pack R phR cT xclass xT. +Definition vectType := @Vector.Pack R phR cT xclass cT. +Definition vect_ringType := @GRing.Ring.Pack vectType xclass xT. +Definition vect_unitRingType := @GRing.UnitRing.Pack vectType xclass xT. +Definition vect_lalgType := @GRing.Lalgebra.Pack R phR vectType xclass xT. +Definition vect_algType := @GRing.Algebra.Pack R phR vectType xclass xT. +Definition vect_unitAlgType := @GRing.UnitAlgebra.Pack R phR vectType xclass xT. + +End ClassDef. + +Module Exports. + +Coercion base1 : class_of >-> GRing.UnitAlgebra.class_of. +Coercion base2 : class_of >-> Vector.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 lmodType : type>-> GRing.Lmodule.type. +Canonical lmodType. +Coercion ringType : type >-> GRing.Ring.type. +Canonical ringType. +Coercion unitRingType : type >-> GRing.UnitRing.type. +Canonical unitRingType. +Coercion lalgType : type >-> GRing.Lalgebra.type. +Canonical lalgType. +Coercion algType : type >-> GRing.Algebra.type. +Canonical algType. +Coercion unitAlgType : type >-> GRing.UnitAlgebra.type. +Canonical unitAlgType. +Coercion vectType : type >-> Vector.type. +Canonical vectType. +Canonical vect_ringType. +Canonical vect_unitRingType. +Canonical vect_lalgType. +Canonical vect_algType. +Canonical vect_unitAlgType. +Notation FalgType R := (type (Phant R)). +Notation "[ 'FalgType' R 'of' A ]" := (@pack _ (Phant R) A _ _ id _ _ id) + (at level 0, format "[ 'FalgType' R 'of' A ]") : form_scope. +Notation "[ 'FalgType' R 'of' A 'for' vT ]" := + (@pack _ (Phant R) A _ _ id vT _ idfun) + (at level 0, format "[ 'FalgType' R 'of' A 'for' vT ]") : form_scope. +Notation FalgUnitRingType T := (@BaseType _ _ T _ _ id _ (Phant T) id). +End Exports. + +End Falgebra. +Export Falgebra.Exports. + +Notation "1" := (vline 1) : vspace_scope. + +Canonical matrix_FalgType (K : fieldType) n := [FalgType K of 'M[K]_n.+1]. + +Section Proper. + +Variables (R : ringType) (aT : FalgType R). +Import Vector.InternalTheory. + +Lemma FalgType_proper : Vector.dim aT > 0. +Proof. +rewrite lt0n; apply: contraNneq (oner_neq0 aT) => aT0. +by apply/eqP/v2r_inj; do 2!move: (v2r _); rewrite aT0 => u v; rewrite !thinmx0. +Qed. + +End Proper. + +Module FalgLfun. + +Section FalgLfun. + +Variable (R : comRingType) (aT : FalgType R). +Implicit Types f g : 'End(aT). + +Canonical Falg_fun_ringType := lfun_ringType (FalgType_proper aT). +Canonical Falg_fun_lalgType := lfun_lalgType (FalgType_proper aT). +Canonical Falg_fun_algType := lfun_algType (FalgType_proper aT). + +Lemma lfun_mulE f g u : (f * g) u = g (f u). Proof. exact: lfunE. Qed. +Lemma lfun_compE f g : (g \o f)%VF = f * g. Proof. by []. Qed. + +End FalgLfun. + +Section InvLfun. + +Variable (K : fieldType) (aT : FalgType K). +Implicit Types f g : 'End(aT). + +Definition lfun_invr f := if lker f == 0%VS then f^-1%VF else f. + +Lemma lfun_mulVr f : lker f == 0%VS -> f^-1%VF * f = 1. +Proof. exact: lker0_compfV. Qed. + +Lemma lfun_mulrV f : lker f == 0%VS -> f * f^-1%VF = 1. +Proof. exact: lker0_compVf. Qed. + +Fact lfun_mulRVr f : lker f == 0%VS -> lfun_invr f * f = 1. +Proof. by move=> Uf; rewrite /lfun_invr Uf lfun_mulVr. Qed. + +Fact lfun_mulrRV f : lker f == 0%VS -> f * lfun_invr f = 1. +Proof. by move=> Uf; rewrite /lfun_invr Uf lfun_mulrV. Qed. + +Fact lfun_unitrP f g : g * f = 1 /\ f * g = 1 -> lker f == 0%VS. +Proof. +case=> _ fK; apply/lker0P; apply: can_inj (g) _ => u. +by rewrite -lfun_mulE fK lfunE. +Qed. + +Lemma lfun_invr_out f : lker f != 0%VS -> lfun_invr f = f. +Proof. by rewrite /lfun_invr => /negPf->. Qed. + +Definition lfun_unitRingMixin := + UnitRingMixin lfun_mulRVr lfun_mulrRV lfun_unitrP lfun_invr_out. +Canonical lfun_unitRingType := UnitRingType 'End(aT) lfun_unitRingMixin. +Canonical lfun_unitAlgType := [unitAlgType K of 'End(aT)]. +Canonical Falg_fun_FalgType := [FalgType K of 'End(aT)]. + +Lemma lfun_invE f : lker f == 0%VS -> f^-1%VF = f^-1. +Proof. by rewrite /f^-1 /= /lfun_invr => ->. Qed. + +End InvLfun. + +End FalgLfun. + +Section FalgebraTheory. + +Variables (K : fieldType) (aT : FalgType K). +Implicit Types (u v : aT) (U V W : {vspace aT}). + +Import FalgLfun. + +Definition amull u : 'End(aT) := linfun (u \*o @idfun aT). +Definition amulr u : 'End(aT) := linfun (u \o* @idfun aT). + +Lemma amull_inj : injective amull. +Proof. by move=> u v /lfunP/(_ 1); rewrite !lfunE /= !mulr1. Qed. + +Lemma amulr_inj : injective amulr. +Proof. by move=> u v /lfunP/(_ 1); rewrite !lfunE /= !mul1r. Qed. + +Fact amull_is_linear : linear amull. +Proof. +move=> a u v; apply/lfunP => w. +by rewrite !lfunE /= scale_lfunE !lfunE /= mulrDl scalerAl. +Qed. +Canonical amull_additive := Eval hnf in Additive amull_is_linear. +Canonical amull_linear := Eval hnf in AddLinear amull_is_linear. + +(* amull is a converse ring morphism *) +Lemma amull1 : amull 1 = \1%VF. +Proof. by apply/lfunP => z; rewrite id_lfunE lfunE /= mul1r. Qed. + +Lemma amullM u v : (amull (u * v) = amull v * amull u)%VF. +Proof. by apply/lfunP => w; rewrite comp_lfunE !lfunE /= mulrA. Qed. + +Lemma amulr_is_lrmorphism : lrmorphism amulr. +Proof. +split=> [|a u]; last by apply/lfunP=> w; rewrite scale_lfunE !lfunE /= scalerAr. +split=> [u v|]; first by apply/lfunP => w; do 3!rewrite !lfunE /= ?mulrBr. +split=> [u v|]; last by apply/lfunP=> w; rewrite id_lfunE !lfunE /= mulr1. +by apply/lfunP=> w; rewrite comp_lfunE !lfunE /= mulrA. +Qed. +Canonical amulr_additive := Eval hnf in Additive amulr_is_lrmorphism. +Canonical amulr_linear := Eval hnf in AddLinear amulr_is_lrmorphism. +Canonical amulr_rmorphism := Eval hnf in AddRMorphism amulr_is_lrmorphism. +Canonical amulr_lrmorphism := Eval hnf in LRMorphism amulr_is_lrmorphism. + +Lemma lker0_amull u : u \is a GRing.unit -> lker (amull u) == 0%VS. +Proof. by move=> Uu; apply/lker0P=> v w; rewrite !lfunE; apply: mulrI. Qed. + +Lemma lker0_amulr u : u \is a GRing.unit -> lker (amulr u) == 0%VS. +Proof. by move=> Uu; apply/lker0P=> v w; rewrite !lfunE; apply: mulIr. Qed. + +Lemma lfun1_poly (p : {poly aT}) : map_poly \1%VF p = p. +Proof. by apply: map_poly_id => u _; apply: id_lfunE. Qed. + +Fact prodv_key : unit. Proof. by []. Qed. +Definition prodv := + locked_with prodv_key (fun U V => <<allpairs *%R (vbasis U) (vbasis V)>>%VS). +Canonical prodv_unlockable := [unlockable fun prodv]. +Local Notation "A * B" := (prodv A B) : vspace_scope. + +Lemma memv_mul U V : {in U & V, forall u v, u * v \in (U * V)%VS}. +Proof. +move=> u v /coord_vbasis-> /coord_vbasis->. +rewrite mulr_suml; apply: memv_suml => i _. +rewrite mulr_sumr; apply: memv_suml => j _. +rewrite -scalerAl -scalerAr !memvZ // [prodv]unlock memv_span //. +by apply/allpairsP; exists ((vbasis U)`_i, (vbasis V)`_j); rewrite !memt_nth. +Qed. + +Lemma prodvP {U V W} : + reflect {in U & V, forall u v, u * v \in W} (U * V <= W)%VS. +Proof. +apply: (iffP idP) => [sUVW u v Uu Vv | sUVW]. + by rewrite (subvP sUVW) ?memv_mul. +rewrite [prodv]unlock; apply/span_subvP=> _ /allpairsP[[u v] /= [Uu Vv ->]]. +by rewrite sUVW ?vbasis_mem. +Qed. + +Lemma prodv_line u v : (<[u]> * <[v]> = <[u * v]>)%VS. +Proof. +apply: subv_anti; rewrite -memvE memv_mul ?memv_line // andbT. +apply/prodvP=> _ _ /vlineP[a ->] /vlineP[b ->]. +by rewrite -scalerAr -scalerAl !memvZ ?memv_line. +Qed. + +Lemma dimv1: \dim (1%VS : {vspace aT}) = 1%N. +Proof. by rewrite dim_vline oner_neq0. Qed. + +Lemma dim_prodv U V : \dim (U * V) <= \dim U * \dim V. +Proof. by rewrite unlock (leq_trans (dim_span _)) ?size_tuple. Qed. + +Lemma vspace1_neq0 : (1 != 0 :> {vspace aT})%VS. +Proof. by rewrite -dimv_eq0 dimv1. Qed. + +Lemma vbasis1 : exists2 k, k != 0 & vbasis 1 = [:: k%:A] :> seq aT. +Proof. +move: (vbasis 1) (@vbasisP K aT 1); rewrite dim_vline oner_neq0. +case/tupleP=> x X0; rewrite {X0}tuple0 => defX; have Xx := mem_head x nil. +have /vlineP[k def_x] := basis_mem defX Xx; exists k; last by rewrite def_x. +by have:= basis_not0 defX Xx; rewrite def_x scaler_eq0 oner_eq0 orbF. +Qed. + +Lemma prod0v : left_zero 0%VS prodv. +Proof. +move=> U; apply/eqP; rewrite -dimv_eq0 -leqn0 (leq_trans (dim_prodv 0 U)) //. +by rewrite dimv0. +Qed. + +Lemma prodv0 : right_zero 0%VS prodv. +Proof. +move=> U; apply/eqP; rewrite -dimv_eq0 -leqn0 (leq_trans (dim_prodv U 0)) //. +by rewrite dimv0 muln0. +Qed. + +Canonical prodv_muloid := Monoid.MulLaw prod0v prodv0. + +Lemma prod1v : left_id 1%VS prodv. +Proof. +move=> U; apply/subv_anti/andP; split. + by apply/prodvP=> _ u /vlineP[a ->] Uu; rewrite mulr_algl memvZ. +by apply/subvP=> u Uu; rewrite -[u]mul1r memv_mul ?memv_line. +Qed. + +Lemma prodv1 : right_id 1%VS prodv. +Proof. +move=> U; apply/subv_anti/andP; split. + by apply/prodvP=> u _ Uu /vlineP[a ->]; rewrite mulr_algr memvZ. +by apply/subvP=> u Uu; rewrite -[u]mulr1 memv_mul ?memv_line. +Qed. + +Lemma prodvS U1 U2 V1 V2 : (U1 <= U2 -> V1 <= V2 -> U1 * V1 <= U2 * V2)%VS. +Proof. +move/subvP=> sU12 /subvP sV12; apply/prodvP=> u v Uu Vv. +by rewrite memv_mul ?sU12 ?sV12. +Qed. + +Lemma prodvSl U1 U2 V : (U1 <= U2 -> U1 * V <= U2 * V)%VS. +Proof. by move/prodvS->. Qed. + +Lemma prodvSr U V1 V2 : (V1 <= V2 -> U * V1 <= U * V2)%VS. +Proof. exact: prodvS. Qed. + +Lemma prodvDl : left_distributive prodv addv. +Proof. +move=> U1 U2 V; apply/esym/subv_anti/andP; split. + by rewrite subv_add 2?prodvS ?addvSl ?addvSr. +apply/prodvP=> _ v /memv_addP[u1 Uu1 [u2 Uu2 ->]] Vv. +by rewrite mulrDl memv_add ?memv_mul. +Qed. + +Lemma prodvDr : right_distributive prodv addv. +Proof. +move=> U V1 V2; apply/esym/subv_anti/andP; split. + by rewrite subv_add 2?prodvS ?addvSl ?addvSr. +apply/prodvP=> u _ Uu /memv_addP[v1 Vv1 [v2 Vv2 ->]]. +by rewrite mulrDr memv_add ?memv_mul. +Qed. + +Canonical addv_addoid := Monoid.AddLaw prodvDl prodvDr. + +Lemma prodvA : associative prodv. +Proof. +move=> U V W; rewrite -(span_basis (vbasisP U)) span_def !big_distrl /=. +apply: eq_bigr => u _; rewrite -(span_basis (vbasisP W)) span_def !big_distrr. +apply: eq_bigr => w _; rewrite -(span_basis (vbasisP V)) span_def /=. +rewrite !(big_distrl, big_distrr) /=; apply: eq_bigr => v _. +by rewrite !prodv_line mulrA. +Qed. + +Canonical prodv_monoid := Monoid.Law prodvA prod1v prodv1. + +Definition expv U n := iterop n.+1.-1 prodv U 1%VS. +Local Notation "A ^+ n" := (expv A n) : vspace_scope. + +Lemma expv0 U : (U ^+ 0 = 1)%VS. Proof. by []. Qed. +Lemma expv1 U : (U ^+ 1 = U)%VS. Proof. by []. Qed. +Lemma expv2 U : (U ^+ 2 = U * U)%VS. Proof. by []. Qed. + +Lemma expvSl U n : (U ^+ n.+1 = U * U ^+ n)%VS. +Proof. by case: n => //; rewrite prodv1. Qed. + +Lemma expv0n n : (0 ^+ n = if n is _.+1 then 0 else 1)%VS. +Proof. by case: n => // n; rewrite expvSl prod0v. Qed. + +Lemma expv1n n : (1 ^+ n = 1)%VS. +Proof. by elim: n => // n IHn; rewrite expvSl IHn prodv1. Qed. + +Lemma expvD U m n : (U ^+ (m + n) = U ^+ m * U ^+ n)%VS. +Proof. by elim: m => [|m IHm]; rewrite ?prod1v // !expvSl IHm prodvA. Qed. + +Lemma expvSr U n : (U ^+ n.+1 = U ^+ n * U)%VS. +Proof. by rewrite -addn1 expvD. Qed. + +Lemma expvM U m n : (U ^+ (m * n) = U ^+ m ^+ n)%VS. +Proof. by elim: n => [|n IHn]; rewrite ?muln0 // mulnS expvD IHn expvSl. Qed. + +Lemma expvS U V n : (U <= V -> U ^+ n <= V ^+ n)%VS. +Proof. +move=> sUV; elim: n => [|n IHn]; first by rewrite !expv0 subvv. +by rewrite !expvSl prodvS. +Qed. + +Lemma expv_line u n : (<[u]> ^+ n = <[u ^+ n]>)%VS. +Proof. +elim: n => [|n IH]; first by rewrite expr0 expv0. +by rewrite exprS expvSl IH prodv_line. +Qed. + +(* Centralisers and centers. *) + +Definition centraliser1_vspace u := lker (amulr u - amull u). +Local Notation "'C [ u ]" := (centraliser1_vspace u) : vspace_scope. +Definition centraliser_vspace V := (\bigcap_i 'C[tnth (vbasis V) i])%VS. +Local Notation "'C ( V )" := (centraliser_vspace V) : vspace_scope. +Definition center_vspace V := (V :&: 'C(V))%VS. +Local Notation "'Z ( V )" := (center_vspace V) : vspace_scope. + +Lemma cent1vP u v : reflect (u * v = v * u) (u \in 'C[v]%VS). +Proof. by rewrite (sameP eqlfunP eqP) !lfunE /=; apply: eqP. Qed. + +Lemma cent1v1 u : 1 \in 'C[u]%VS. Proof. by apply/cent1vP; rewrite commr1. Qed. +Lemma cent1v_id u : u \in 'C[u]%VS. Proof. exact/cent1vP. Qed. +Lemma cent1vX u n : u ^+ n \in 'C[u]%VS. Proof. exact/cent1vP/esym/commrX. Qed. +Lemma cent1vC u v : (u \in 'C[v])%VS = (v \in 'C[u])%VS. +Proof. exact/cent1vP/cent1vP. Qed. + +Lemma centvP u V : reflect {in V, forall v, u * v = v * u} (u \in 'C(V))%VS. +Proof. +apply: (iffP subv_bigcapP) => [cVu y /coord_vbasis-> | cVu i _]. + apply/esym/cent1vP/rpred_sum=> i _; apply: rpredZ. + by rewrite -tnth_nth cent1vC memvE cVu. +exact/cent1vP/cVu/vbasis_mem/mem_tnth. +Qed. +Lemma centvsP U V : reflect {in U & V, commutative *%R} (U <= 'C(V))%VS. +Proof. by apply: (iffP subvP) => [cUV u v | cUV u] /cUV-/centvP; apply. Qed. + +Lemma subv_cent1 U v : (U <= 'C[v])%VS = (v \in 'C(U)%VS). +Proof. +by apply/subvP/centvP=> cUv u Uu; apply/cent1vP; rewrite 1?cent1vC cUv. +Qed. + +Lemma centv1 V : 1 \in 'C(V)%VS. +Proof. by apply/centvP=> v _; rewrite commr1. Qed. +Lemma centvX V u n : u \in 'C(V)%VS -> u ^+ n \in 'C(V)%VS. +Proof. by move/centvP=> cVu; apply/centvP=> v /cVu/esym/commrX->. Qed. +Lemma centvC U V : (U <= 'C(V))%VS = (V <= 'C(U))%VS. +Proof. by apply/centvsP/centvsP=> cUV u v UVu /cUV->. Qed. + +Lemma centerv_sub V : ('Z(V) <= V)%VS. Proof. exact: capvSl. Qed. +Lemma cent_centerv V : (V <= 'C('Z(V)))%VS. +Proof. by rewrite centvC capvSr. Qed. + +(* Building the predicate that checks is a vspace has a unit *) +Definition is_algid e U := + [/\ e \in U, e != 0 & {in U, forall u, e * u = u /\ u * e = u}]. + +Fact algid_decidable U : decidable (exists e, is_algid e U). +Proof. +have [-> | nzU] := eqVneq U 0%VS. + by right=> [[e []]]; rewrite memv0 => ->. +pose X := vbasis U; pose feq f1 f2 := [tuple of map f1 X ++ map f2 X]. +have feqL f i: tnth (feq _ f _) (lshift _ i) = f X`_i. + set v := f _; rewrite (tnth_nth v) /= nth_cat size_map size_tuple. + by rewrite ltn_ord (nth_map 0) ?size_tuple. +have feqR f i: tnth (feq _ _ f) (rshift _ i) = f X`_i. + set v := f _; rewrite (tnth_nth v) /= nth_cat size_map size_tuple. + by rewrite ltnNge leq_addr addKn /= (nth_map 0) ?size_tuple. +apply: decP (vsolve_eq (feq _ amulr amull) (feq _ id id) U) _. +apply: (iffP (vsolve_eqP _ _ _)) => [[e Ue id_e] | [e [Ue _ id_e]]]. + suffices idUe: {in U, forall u, e * u = u /\ u * e = u}. + exists e; split=> //; apply: contraNneq nzU => e0; rewrite -subv0. + by apply/subvP=> u /idUe[<- _]; rewrite e0 mul0r mem0v. + move=> u /coord_vbasis->; rewrite mulr_sumr mulr_suml. + split; apply/eq_bigr=> i _; rewrite -(scalerAr, scalerAl); congr (_ *: _). + by have:= id_e (lshift _ i); rewrite !feqL lfunE. + by have:= id_e (rshift _ i); rewrite !feqR lfunE. +have{id_e} /all_and2[ideX idXe]:= id_e _ (vbasis_mem (mem_tnth _ X)). +exists e => // k; rewrite -[k]splitK. +by case: (split k) => i; rewrite !(feqL, feqR) lfunE /= -tnth_nth. +Qed. + +Definition has_algid : pred {vspace aT} := algid_decidable. + +Lemma has_algidP {U} : reflect (exists e, is_algid e U) (has_algid U). +Proof. exact: sumboolP. Qed. + +Lemma has_algid1 U : 1 \in U -> has_algid U. +Proof. +move=> U1; apply/has_algidP; exists 1; split; rewrite ?oner_eq0 // => u _. +by rewrite mulr1 mul1r. +Qed. + +Definition is_aspace U := has_algid U && (U * U <= U)%VS. +Structure aspace := ASpace {asval :> {vspace aT}; _ : is_aspace asval}. +Definition aspace_of of phant aT := aspace. +Local Notation "{ 'aspace' T }" := (aspace_of (Phant T)) : type_scope. + +Canonical aspace_subType := Eval hnf in [subType for asval]. +Definition aspace_eqMixin := [eqMixin of aspace by <:]. +Canonical aspace_eqType := Eval hnf in EqType aspace aspace_eqMixin. +Definition aspace_choiceMixin := [choiceMixin of aspace by <:]. +Canonical aspace_choiceType := Eval hnf in ChoiceType aspace aspace_choiceMixin. + +Canonical aspace_of_subType := Eval hnf in [subType of {aspace aT}]. +Canonical aspace_of_eqType := Eval hnf in [eqType of {aspace aT}]. +Canonical aspace_of_choiceType := Eval hnf in [choiceType of {aspace aT}]. + +Definition clone_aspace U (A : {aspace aT}) := + fun algU & phant_id algU (valP A) => @ASpace U algU : {aspace aT}. + +Fact aspace1_subproof : is_aspace 1. +Proof. by rewrite /is_aspace prod1v -memvE has_algid1 memv_line. Qed. +Canonical aspace1 : {aspace aT} := ASpace aspace1_subproof. + +Lemma aspacef_subproof : is_aspace fullv. +Proof. by rewrite /is_aspace subvf has_algid1 ?memvf. Qed. +Canonical aspacef : {aspace aT} := ASpace aspacef_subproof. + +Lemma polyOver1P p : + reflect (exists q, p = map_poly (in_alg aT) q) (p \is a polyOver 1%VS). +Proof. +apply: (iffP idP) => [/allP/=Qp | [q ->]]; last first. + by apply/polyOverP=> j; rewrite coef_map rpredZ ?memv_line. +exists (map_poly (coord [tuple 1] 0) p). +rewrite -map_poly_comp map_poly_id // => _ /Qp/vlineP[a ->] /=. +by rewrite linearZ /= (coord_free 0) ?mulr1 // seq1_free ?oner_eq0. +Qed. + +End FalgebraTheory. + +Delimit Scope aspace_scope with AS. +Bind Scope aspace_scope with aspace. +Bind Scope aspace_scope with aspace_of. +Arguments Scope asval [_ _ aspace_scope]. +Arguments Scope clone_aspace [_ _ vspace_scope aspace_scope _ _]. + +Notation "{ 'aspace' T }" := (aspace_of (Phant T)) : type_scope. +Notation "A * B" := (prodv A B) : vspace_scope. +Notation "A ^+ n" := (expv A n) : vspace_scope. +Notation "'C [ u ]" := (centraliser1_vspace u) : vspace_scope. +Notation "'C_ U [ v ]" := (capv U 'C[v]) : vspace_scope. +Notation "'C_ ( U ) [ v ]" := (capv U 'C[v]) (only parsing) : vspace_scope. +Notation "'C ( V )" := (centraliser_vspace V) : vspace_scope. +Notation "'C_ U ( V )" := (capv U 'C(V)) : vspace_scope. +Notation "'C_ ( U ) ( V )" := (capv U 'C(V)) (only parsing) : vspace_scope. +Notation "'Z ( V )" := (center_vspace V) : vspace_scope. + +Notation "1" := (aspace1 _) : aspace_scope. +Notation "{ : aT }" := (aspacef aT) : aspace_scope. +Notation "[ 'aspace' 'of' U ]" := (@clone_aspace _ _ U _ _ id) + (at level 0, format "[ 'aspace' 'of' U ]") : form_scope. +Notation "[ 'aspace' 'of' U 'for' A ]" := (@clone_aspace _ _ U A _ idfun) + (at level 0, format "[ 'aspace' 'of' U 'for' A ]") : form_scope. + +Implicit Arguments prodvP [K aT U V W]. +Implicit Arguments cent1vP [K aT u v]. +Implicit Arguments centvP [K aT u V]. +Implicit Arguments centvsP [K aT U V]. +Implicit Arguments has_algidP [K aT U]. +Implicit Arguments polyOver1P [K aT p]. + +Section AspaceTheory. + +Variables (K : fieldType) (aT : FalgType K). +Implicit Types (u v e : aT) (U V : {vspace aT}) (A B : {aspace aT}). +Import FalgLfun. + +Lemma algid_subproof U : + {e | e \in U + & has_algid U ==> (U <= lker (amull e - 1) :&: lker (amulr e - 1))%VS}. +Proof. +apply: sig2W; case: has_algidP => [[e]|]; last by exists 0; rewrite ?mem0v. +case=> Ae _ idAe; exists e => //; apply/subvP=> u /idAe[eu_u ue_u]. +by rewrite memv_cap !memv_ker !lfun_simp /= eu_u ue_u subrr eqxx. +Qed. + +Definition algid U := s2val (algid_subproof U). + +Lemma memv_algid U : algid U \in U. +Proof. by rewrite /algid; case: algid_subproof. Qed. + +Lemma algidl A : {in A, left_id (algid A) *%R}. +Proof. +rewrite /algid; case: algid_subproof => e _ /=; have /andP[-> _] := valP A. +move/subvP=> idAe u /idAe/memv_capP[]. +by rewrite memv_ker !lfun_simp /= subr_eq0 => /eqP. +Qed. + +Lemma algidr A : {in A, right_id (algid A) *%R}. +Proof. +rewrite /algid; case: algid_subproof => e _ /=; have /andP[-> _] := valP A. +move/subvP=> idAe u /idAe/memv_capP[_]. +by rewrite memv_ker !lfun_simp /= subr_eq0 => /eqP. +Qed. + +Lemma unitr_algid1 A u : u \in A -> u \is a GRing.unit -> algid A = 1. +Proof. by move=> Eu /mulrI; apply; rewrite mulr1 algidr. Qed. + +Lemma algid_eq1 A : (algid A == 1) = (1 \in A). +Proof. by apply/eqP/idP=> [<- | /algidr <-]; rewrite ?memv_algid ?mul1r. Qed. + +Lemma algid_neq0 A : algid A != 0. +Proof. +have /andP[/has_algidP[u [Au nz_u _]] _] := valP A. +by apply: contraNneq nz_u => e0; rewrite -(algidr Au) e0 mulr0. +Qed. + +Lemma dim_algid A : \dim <[algid A]> = 1%N. +Proof. by rewrite dim_vline algid_neq0. Qed. + +Lemma adim_gt0 A : (0 < \dim A)%N. +Proof. by rewrite -(dim_algid A) dimvS // -memvE ?memv_algid. Qed. + +Lemma not_asubv0 A : ~~ (A <= 0)%VS. +Proof. by rewrite subv0 -dimv_eq0 -lt0n adim_gt0. Qed. + +Lemma adim1P {A} : reflect (A = <[algid A]>%VS :> {vspace aT}) (\dim A == 1%N). +Proof. +rewrite eqn_leq adim_gt0 -(memv_algid A) andbC -(dim_algid A) -eqEdim eq_sym. +exact: eqP. +Qed. + +Lemma asubv A : (A * A <= A)%VS. +Proof. by have /andP[] := valP A. Qed. + +Lemma memvM A : {in A &, forall u v, u * v \in A}. +Proof. exact/prodvP/asubv. Qed. + +Lemma prodv_id A : (A * A)%VS = A. +Proof. +apply/eqP; rewrite eqEsubv asubv; apply/subvP=> u Au. +by rewrite -(algidl Au) memv_mul // memv_algid. +Qed. + +Lemma prodv_sub U V A : (U <= A -> V <= A -> U * V <= A)%VS. +Proof. by move=> sUA sVA; rewrite -prodv_id prodvS. Qed. + +Lemma expv_id A n : (A ^+ n.+1)%VS = A. +Proof. by elim: n => // n IHn; rewrite !expvSl prodvA prodv_id -expvSl. Qed. + +Lemma limg_amulr U v : (amulr v @: U = U * <[v]>)%VS. +Proof. +rewrite -(span_basis (vbasisP U)) limg_span !span_def big_distrl /= big_map. +by apply: eq_bigr => u; rewrite prodv_line lfunE. +Qed. + +Lemma memv_cosetP {U v w} : + reflect (exists2 u, u\in U & w = u * v) (w \in U * <[v]>)%VS. +Proof. +rewrite -limg_amulr. +by apply: (iffP memv_imgP) => [] [u] Uu ->; exists u; rewrite ?lfunE. +Qed. + +Lemma dim_cosetv_unit V u : u \is a GRing.unit -> \dim (V * <[u]>) = \dim V. +Proof. +by move/lker0_amulr/eqP=> Uu; rewrite -limg_amulr limg_dim_eq // Uu capv0. +Qed. + +Lemma memvV A u : (u^-1 \in A) = (u \in A). +Proof. +suffices{u} invA: invr_closed A by apply/idP/idP=> /invA; rewrite ?invrK. +move=> u Au; have [Uu | /invr_out-> //] := boolP (u \is a GRing.unit). +rewrite memvE -(limg_ker0 _ _ (lker0_amulr Uu)) limg_line lfunE /= mulVr //. +suff ->: (amulr u @: A)%VS = A by rewrite -memvE -algid_eq1 (unitr_algid1 Au). +by apply/eqP; rewrite limg_amulr -dimv_leqif_eq ?prodv_sub ?dim_cosetv_unit. +Qed. + +Fact aspace_cap_subproof A B : algid A \in B -> is_aspace (A :&: B). +Proof. +move=> BeA; apply/andP. +split; [apply/has_algidP | by rewrite subv_cap !prodv_sub ?capvSl ?capvSr]. +exists (algid A); rewrite /is_algid algid_neq0 memv_cap memv_algid. +by split=> // u /memv_capP[Au _]; rewrite ?algidl ?algidr. +Qed. +Definition aspace_cap A B BeA := ASpace (@aspace_cap_subproof A B BeA). + +Fact centraliser1_is_aspace u : is_aspace 'C[u]. +Proof. +rewrite /is_aspace has_algid1 ?cent1v1 //=. +apply/prodvP=> v w /cent1vP-cuv /cent1vP-cuw. +by apply/cent1vP; rewrite -mulrA cuw !mulrA cuv. +Qed. +Canonical centraliser1_aspace u := ASpace (centraliser1_is_aspace u). + +Fact centraliser_is_aspace V : is_aspace 'C(V). +Proof. +rewrite /is_aspace has_algid1 ?centv1 //=. +apply/prodvP=> u w /centvP-cVu /centvP-cVw. +by apply/centvP=> v Vv; rewrite /= -mulrA cVw // !mulrA cVu. +Qed. +Canonical centraliser_aspace V := ASpace (centraliser_is_aspace V). + +Lemma centv_algid A : algid A \in 'C(A)%VS. +Proof. by apply/centvP=> u Au; rewrite algidl ?algidr. Qed. +Canonical center_aspace A := [aspace of 'Z(A) for aspace_cap (centv_algid A)]. + +Lemma algid_center A : algid 'Z(A) = algid A. +Proof. +rewrite -(algidl (subvP (centerv_sub A) _ (memv_algid _))) algidr //=. +by rewrite memv_cap memv_algid centv_algid. +Qed. + +Lemma Falgebra_FieldMixin : + GRing.IntegralDomain.axiom aT -> GRing.Field.mixin_of aT. +Proof. +move=> domT u nz_u; apply/unitrP. +have kerMu: lker (amulr u) == 0%VS. + rewrite eqEsubv sub0v andbT; apply/subvP=> v; rewrite memv_ker lfunE /=. + by move/eqP/domT; rewrite (negPf nz_u) orbF memv0. +have /memv_imgP[v _ vu1]: 1 \in limg (amulr u); last rewrite lfunE /= in vu1. + suffices /eqP->: limg (amulr u) == fullv by rewrite memvf. + by rewrite -dimv_leqif_eq ?subvf ?limg_dim_eq // (eqP kerMu) capv0. +exists v; split=> //; apply: (lker0P kerMu). +by rewrite !lfunE /= -mulrA -vu1 mulr1 mul1r. +Qed. + +Section SkewField. + +Hypothesis fieldT : GRing.Field.mixin_of aT. + +Lemma skew_field_algid1 A : algid A = 1. +Proof. by rewrite (unitr_algid1 (memv_algid A)) ?fieldT ?algid_neq0. Qed. + +Lemma skew_field_module_semisimple A M : + let sumA X := (\sum_(x <- X) A * <[x]>)%VS in + (A * M <= M)%VS -> {X | [/\ sumA X = M, directv (sumA X) & 0 \notin X]}. +Proof. +move=> sumA sAM_M; pose X := Nil aT; pose k := (\dim (A * M) - \dim (sumA X))%N. +have: (\dim (A * M) - \dim (sumA X) < k.+1)%N by []. +have: [/\ (sumA X <= A * M)%VS, directv (sumA X) & 0 \notin X]. + by rewrite /sumA directvE /= !big_nil sub0v dimv0. +elim: {X k}k.+1 (X) => // k IHk X [sAX_AM dxAX nzX]; rewrite ltnS => leAXk. +have [sM_AX | /subvPn/sig2W[y My notAXy]] := boolP (M <= sumA X)%VS. + by exists X; split=> //; apply/eqP; rewrite eqEsubv (subv_trans sAX_AM). +have nz_y: y != 0 by rewrite (memPnC notAXy) ?mem0v. +pose AY := sumA (y :: X). +have sAY_AM: (AY <= A * M)%VS by rewrite [AY]big_cons subv_add ?prodvSr. +have dxAY: directv AY. + rewrite directvE /= !big_cons [_ == _]directv_addE dxAX directvE eqxx /=. + rewrite -/(sumA X) eqEsubv sub0v andbT -limg_amulr. + apply/subvP=> _ /memv_capP[/memv_imgP[a Aa ->]]; rewrite lfunE /= => AXay. + rewrite memv0 (mulIr_eq0 a (mulIr _)) ?fieldT //. + apply: contraR notAXy => /fieldT-Ua; rewrite -[y](mulKr Ua) /sumA. + by rewrite -big_distrr -(prodv_id A) /= -prodvA big_distrr memv_mul ?memvV. +apply: (IHk (y :: X)); first by rewrite !inE eq_sym negb_or nz_y. +rewrite -subSn ?dimvS // (directvP dxAY) /= big_cons -(directvP dxAX) /=. +rewrite subnDA (leq_trans _ leAXk) ?leq_sub2r // leq_subLR -add1n leq_add2r. +by rewrite dim_cosetv_unit ?fieldT ?adim_gt0. +Qed. + +Lemma skew_field_module_dimS A M : (A * M <= M)%VS -> \dim A %| \dim M. +Proof. +case/skew_field_module_semisimple=> X [<- /directvP-> nzX] /=. +rewrite big_seq prime.dvdn_sum // => x /(memPn nzX)nz_x. +by rewrite dim_cosetv_unit ?fieldT. +Qed. + +Lemma skew_field_dimS A B : (A <= B)%VS -> \dim A %| \dim B. +Proof. by move=> sAB; rewrite skew_field_module_dimS ?prodv_sub. Qed. + +End SkewField. + +End AspaceTheory. + +(* Note that local centraliser might not be proper sub-algebras. *) +Notation "'C [ u ]" := (centraliser1_aspace u) : aspace_scope. +Notation "'C ( V )" := (centraliser_aspace V) : aspace_scope. +Notation "'Z ( A )" := (center_aspace A) : aspace_scope. + +Implicit Arguments adim1P [K aT A]. +Implicit Arguments memv_cosetP [K aT U v w]. + +Section Closure. + +Variables (K : fieldType) (aT : FalgType K). +Implicit Types (u v : aT) (U V W : {vspace aT}). + +(* Subspaces of an F-algebra form a Kleene algebra *) +Definition agenv U := (\sum_(i < \dim {:aT}) U ^+ i)%VS. +Local Notation "<< U & vs >>" := (agenv (U + <<vs>>)) : vspace_scope. +Local Notation "<< U ; x >>" := (agenv (U + <[x]>)) : vspace_scope. + +Lemma agenvEl U : agenv U = (1 + U * agenv U)%VS. +Proof. +pose f V := (1 + U * V)%VS; rewrite -/(f _); pose n := \dim {:aT}. +have ->: agenv U = iter n f 0%VS. + rewrite /agenv -/n; elim: n => [|n IHn]; first by rewrite big_ord0. + rewrite big_ord_recl /= -{}IHn; congr (1 + _)%VS; rewrite big_distrr /=. + by apply: eq_bigr => i; rewrite expvSl. +have fS i j: i <= j -> (iter i f 0 <= iter j f 0)%VS. + by elim: i j => [|i IHi] [|j] leij; rewrite ?sub0v //= addvS ?prodvSr ?IHi. +suffices /(@trajectP _ f _ n.+1)[i le_i_n Dfi]: looping f 0%VS n.+1. + by apply/eqP; rewrite eqEsubv -iterS fS // Dfi fS. +apply: contraLR (dimvS (subvf (iter n.+1 f 0%VS))); rewrite -/n -ltnNge. +rewrite -looping_uniq; elim: n.+1 => // i IHi; rewrite trajectSr rcons_uniq. +rewrite {1}trajectSr mem_rcons inE negb_or eq_sym eqEdim fS ?leqW // -ltnNge. +by rewrite -andbA => /and3P[lt_fi _ /IHi/leq_ltn_trans->]. +Qed. + +Lemma agenvEr U : agenv U = (1 + agenv U * U)%VS. +Proof. +rewrite [lhs in lhs = _]agenvEl big_distrr big_distrl /=; congr (_ + _)%VS. +by apply: eq_bigr => i _ /=; rewrite -expvSr -expvSl. +Qed. + +Lemma agenv_modl U V : (U * V <= V -> agenv U * V <= V)%VS. +Proof. +rewrite big_distrl /= => idlU_V; apply/subv_sumP=> [[i _] /= _]. +elim: i => [|i]; first by rewrite expv0 prod1v. +by apply: subv_trans; rewrite expvSr -prodvA prodvSr. +Qed. + +Lemma agenv_modr U V : (V * U <= V -> V * agenv U <= V)%VS. +Proof. +rewrite big_distrr /= => idrU_V; apply/subv_sumP=> [[i _] /= _]. +elim: i => [|i]; first by rewrite expv0 prodv1. +by apply: subv_trans; rewrite expvSl prodvA prodvSl. +Qed. + +Fact agenv_is_aspace U : is_aspace (agenv U). +Proof. +rewrite /is_aspace has_algid1; last by rewrite memvE agenvEl addvSl. +by rewrite agenv_modl // [V in (_ <= V)%VS]agenvEl addvSr. +Qed. +Canonical agenv_aspace U : {aspace aT} := ASpace (agenv_is_aspace U). + +Lemma agenvE U : agenv U = agenv_aspace U. Proof. by []. Qed. + +(* Kleene algebra properties *) + +Lemma agenvM U : (agenv U * agenv U)%VS = agenv U. Proof. exact: prodv_id. Qed. +Lemma agenvX n U : (agenv U ^+ n.+1)%VS = agenv U. Proof. exact: expv_id. Qed. + +Lemma sub1_agenv U : (1 <= agenv U)%VS. Proof. by rewrite agenvEl addvSl. Qed. + +Lemma sub_agenv U : (U <= agenv U)%VS. +Proof. by rewrite 2!agenvEl addvC prodvDr prodv1 -addvA addvSl. Qed. + +Lemma subX_agenv U n : (U ^+ n <= agenv U)%VS. +Proof. +by case: n => [|n]; rewrite ?sub1_agenv // -(agenvX n) expvS // sub_agenv. +Qed. + +Lemma agenv_sub_modl U V : (1 <= V -> U * V <= V -> agenv U <= V)%VS. +Proof. +move=> s1V /agenv_modl; apply: subv_trans. +by rewrite -[Us in (Us <= _)%VS]prodv1 prodvSr. +Qed. + +Lemma agenv_sub_modr U V : (1 <= V -> V * U <= V -> agenv U <= V)%VS. +Proof. +move=> s1V /agenv_modr; apply: subv_trans. +by rewrite -[Us in (Us <= _)%VS]prod1v prodvSl. +Qed. + +Lemma agenv_id U : agenv (agenv U) = agenv U. +Proof. +apply/eqP; rewrite eqEsubv sub_agenv andbT. +by rewrite agenv_sub_modl ?sub1_agenv ?agenvM. +Qed. + +Lemma agenvS U V : (U <= V -> agenv U <= agenv V)%VS. +Proof. +move=> sUV; rewrite agenv_sub_modl ?sub1_agenv //. +by rewrite -[Vs in (_ <= Vs)%VS]agenvM prodvSl ?(subv_trans sUV) ?sub_agenv. +Qed. + +Lemma agenv_add_id U V : agenv (agenv U + V) = agenv (U + V). +Proof. +apply/eqP; rewrite eqEsubv andbC agenvS ?addvS ?sub_agenv //=. +rewrite agenv_sub_modl ?sub1_agenv //. +rewrite -[rhs in (_ <= rhs)%VS]agenvM prodvSl // subv_add agenvS ?addvSl //=. +exact: subv_trans (addvSr U V) (sub_agenv _). +Qed. + +Lemma subv_adjoin U x : (U <= <<U; x>>)%VS. +Proof. by rewrite (subv_trans (sub_agenv _)) ?agenvS ?addvSl. Qed. + +Lemma subv_adjoin_seq U xs : (U <= <<U & xs>>)%VS. +Proof. by rewrite (subv_trans (sub_agenv _)) // ?agenvS ?addvSl. Qed. + +Lemma memv_adjoin U x : x \in <<U; x>>%VS. +Proof. by rewrite memvE (subv_trans (sub_agenv _)) ?agenvS ?addvSr. Qed. + +Lemma seqv_sub_adjoin U xs : {subset xs <= <<U & xs>>%VS}. +Proof. +by apply/span_subvP; rewrite (subv_trans (sub_agenv _)) ?agenvS ?addvSr. +Qed. + +Lemma subvP_adjoin U x y : y \in U -> y \in <<U; x>>%VS. +Proof. exact/subvP/subv_adjoin. Qed. + +Lemma adjoin_nil V : <<V & [::]>>%VS = agenv V. +Proof. by rewrite span_nil addv0. Qed. + +Lemma adjoin_cons V x rs : <<V & x :: rs>>%VS = << <<V; x>> & rs>>%VS. +Proof. by rewrite span_cons addvA agenv_add_id. Qed. + +Lemma adjoin_rcons V rs x : <<V & rcons rs x>>%VS = << <<V & rs>>%VS; x>>%VS. +Proof. by rewrite -cats1 span_cat addvA span_seq1 agenv_add_id. Qed. + +Lemma adjoin_seq1 V x : <<V & [:: x]>>%VS = <<V; x>>%VS. +Proof. by rewrite adjoin_cons adjoin_nil agenv_id. Qed. + +Lemma adjoinC V x y : << <<V; x>>; y>>%VS = << <<V; y>>; x>>%VS. +Proof. by rewrite !agenv_add_id -!addvA (addvC <[x]>%VS). Qed. + +Lemma adjoinSl U V x : (U <= V -> <<U; x>> <= <<V; x>>)%VS. +Proof. by move=> sUV; rewrite agenvS ?addvS. Qed. + +Lemma adjoin_seqSl U V rs : (U <= V -> <<U & rs>> <= <<V & rs>>)%VS. +Proof. by move=> sUV; rewrite agenvS ?addvS. Qed. + +Lemma adjoin_seqSr U rs1 rs2 : + {subset rs1 <= rs2} -> (<<U & rs1>> <= <<U & rs2>>)%VS. +Proof. by move/sub_span=> s_rs12; rewrite agenvS ?addvS. Qed. + +End Closure. + +Notation "<< U >>" := (agenv_aspace U) : aspace_scope. +Notation "<< U & vs >>" := (agenv (U + <<vs>>)) : vspace_scope. +Notation "<< U ; x >>" := (agenv (U + <[x]>)) : vspace_scope. +Notation "<< U & vs >>" := << U + <<vs>> >>%AS : aspace_scope. +Notation "<< U ; x >>" := << U + <[x]> >>%AS : aspace_scope. + +Section SubFalgType. + +(* The FalgType structure of subvs_of A for A : {aspace aT}. *) +(* We can't use the rpred-based mixin, because A need not contain 1. *) +Variable (K : fieldType) (aT : FalgType K) (A : {aspace aT}). + +Definition subvs_one := Subvs (memv_algid A). +Definition subvs_mul (u v : subvs_of A) := + Subvs (subv_trans (memv_mul (subvsP u) (subvsP v)) (asubv _)). + +Fact subvs_mulA : associative subvs_mul. +Proof. by move=> x y z; apply/val_inj/mulrA. Qed. +Fact subvs_mu1l : left_id subvs_one subvs_mul. +Proof. by move=> x; apply/val_inj/algidl/(valP x). Qed. +Fact subvs_mul1 : right_id subvs_one subvs_mul. +Proof. by move=> x; apply/val_inj/algidr/(valP x). Qed. +Fact subvs_mulDl : left_distributive subvs_mul +%R. +Proof. move=> x y z; apply/val_inj/mulrDl. Qed. +Fact subvs_mulDr : right_distributive subvs_mul +%R. +Proof. move=> x y z; apply/val_inj/mulrDr. Qed. + +Definition subvs_ringMixin := + RingMixin subvs_mulA subvs_mu1l subvs_mul1 subvs_mulDl subvs_mulDr + (algid_neq0 _). +Canonical subvs_ringType := Eval hnf in RingType (subvs_of A) subvs_ringMixin. + +Lemma subvs_scaleAl k (x y : subvs_of A) : k *: (x * y) = (k *: x) * y. +Proof. exact/val_inj/scalerAl. Qed. +Canonical subvs_lalgType := Eval hnf in LalgType K (subvs_of A) subvs_scaleAl. + +Lemma subvs_scaleAr k (x y : subvs_of A) : k *: (x * y) = x * (k *: y). +Proof. exact/val_inj/scalerAr. Qed. +Canonical subvs_algType := Eval hnf in AlgType K (subvs_of A) subvs_scaleAr. + +Canonical subvs_unitRingType := Eval hnf in FalgUnitRingType (subvs_of A). +Canonical subvs_unitAlgType := Eval hnf in [unitAlgType K of subvs_of A]. +Canonical subvs_FalgType := Eval hnf in [FalgType K of subvs_of A]. + +Implicit Type w : subvs_of A. + +Lemma vsval_unitr w : vsval w \is a GRing.unit -> w \is a GRing.unit. +Proof. +case: w => /= u Au Uu; have Au1: u^-1 \in A by rewrite memvV. +apply/unitrP; exists (Subvs Au1). +by split; apply: val_inj; rewrite /= ?mulrV ?mulVr ?(unitr_algid1 Au). +Qed. + +Lemma vsval_invr w : vsval w \is a GRing.unit -> val w^-1 = (val w)^-1. +Proof. +move=> Uu; have def_w: w / w * w = w by rewrite divrK ?vsval_unitr. +by apply: (mulrI Uu); rewrite -[in u in u / _]def_w ?mulrK. +Qed. + +End SubFalgType. + +Section AHom. + +Variable K : fieldType. + +Section Class_Def. + +Variables aT rT : FalgType K. + +Definition ahom_in (U : {vspace aT}) (f : 'Hom(aT, rT)) := + let fM_at x y := f (x * y) == f x * f y in + all (fun x => all (fM_at x) (vbasis U)) (vbasis U) && (f 1 == 1). + +Lemma ahom_inP {f : 'Hom(aT, rT)} {U : {vspace aT}} : + reflect ({in U &, {morph f : x y / x * y >-> x * y}} * (f 1 = 1)) + (ahom_in U f). +Proof. +apply: (iffP andP) => [[/allP fM /eqP f1] | [fM f1]]; last first. + rewrite f1; split=> //; apply/allP=> x Ax; apply/allP=> y Ay. + by rewrite fM // vbasis_mem. +split=> // x y /coord_vbasis -> /coord_vbasis ->. +rewrite !mulr_suml ![f _]linear_sum mulr_suml; apply: eq_bigr => i _ /=. +rewrite !mulr_sumr linear_sum; apply: eq_bigr => j _ /=. +rewrite !linearZ -!scalerAr -!scalerAl 2!linearZ /=; congr (_ *: (_ *: _)). +by apply/eqP/(allP (fM _ _)); apply: memt_nth. +Qed. + +Lemma ahomP {f : 'Hom(aT, rT)} : reflect (lrmorphism f) (ahom_in {:aT} f). +Proof. +apply: (iffP ahom_inP) => [[fM f1] | fRM_P]; last first. + pose fRM := LRMorphism fRM_P. + by split; [apply: in2W (rmorphM fRM) | apply: (rmorph1 fRM)]. +split; last exact: linearZZ; split; first exact: linearB. +by split=> // x y; rewrite fM ?memvf. +Qed. + +Structure ahom := AHom {ahval :> 'Hom(aT, rT); _ : ahom_in {:aT} ahval}. + +Canonical ahom_subType := Eval hnf in [subType for ahval]. +Definition ahom_eqMixin := [eqMixin of ahom by <:]. +Canonical ahom_eqType := Eval hnf in EqType ahom ahom_eqMixin. + +Definition ahom_choiceMixin := [choiceMixin of ahom by <:]. +Canonical ahom_choiceType := Eval hnf in ChoiceType ahom ahom_choiceMixin. + +Fact linfun_is_ahom (f : {lrmorphism aT -> rT}) : ahom_in {:aT} (linfun f). +Proof. by apply/ahom_inP; split=> [x y|]; rewrite !lfunE ?rmorphM ?rmorph1. Qed. +Canonical linfun_ahom f := AHom (linfun_is_ahom f). + +End Class_Def. + +Implicit Arguments ahom_in [aT rT]. +Implicit Arguments ahom_inP [aT rT f U]. +Implicit Arguments ahomP [aT rT f]. + +Section LRMorphism. + +Variables aT rT sT : FalgType K. + +Fact ahom_is_lrmorphism (f : ahom aT rT) : lrmorphism f. +Proof. by apply/ahomP; case: f. Qed. +Canonical ahom_rmorphism f := Eval hnf in AddRMorphism (ahom_is_lrmorphism f). +Canonical ahom_lrmorphism f := Eval hnf in AddLRMorphism (ahom_is_lrmorphism f). + +Lemma ahomWin (f : ahom aT rT) U : ahom_in U f. +Proof. +by apply/ahom_inP; split; [apply: in2W (rmorphM _) | apply: rmorph1]. +Qed. + +Lemma id_is_ahom (V : {vspace aT}) : ahom_in V \1. +Proof. by apply/ahom_inP; split=> [x y|] /=; rewrite !id_lfunE. Qed. +Canonical id_ahom := AHom (id_is_ahom (aspacef aT)). + +Lemma comp_is_ahom (V : {vspace aT}) (f : 'Hom(rT, sT)) (g : 'Hom(aT, rT)) : + ahom_in {:rT} f -> ahom_in V g -> ahom_in V (f \o g). +Proof. +move=> /ahom_inP fM /ahom_inP gM; apply/ahom_inP. +by split=> [x y Vx Vy|] /=; rewrite !comp_lfunE gM // fM ?memvf. +Qed. +Canonical comp_ahom (f : ahom rT sT) (g : ahom aT rT) := + AHom (comp_is_ahom (valP f) (valP g)). + +Lemma aimgM (f : ahom aT rT) U V : (f @: (U * V) = f @: U * f @: V)%VS. +Proof. +apply/eqP; rewrite eqEsubv; apply/andP; split; last first. + apply/prodvP=> _ _ /memv_imgP[u Hu ->] /memv_imgP[v Hv ->]. + by rewrite -rmorphM memv_img // memv_mul. +apply/subvP=> _ /memv_imgP[w UVw ->]; rewrite memv_preim (subvP _ w UVw) //. +by apply/prodvP=> u v Uu Vv; rewrite -memv_preim rmorphM memv_mul // memv_img. +Qed. + +Lemma aimg1 (f : ahom aT rT) : (f @: 1 = 1)%VS. +Proof. by rewrite limg_line rmorph1. Qed. + +Lemma aimgX (f : ahom aT rT) U n : (f @: (U ^+ n) = f @: U ^+ n)%VS. +Proof. +elim: n => [|n IH]; first by rewrite !expv0 aimg1. +by rewrite !expvSl aimgM IH. +Qed. + +Lemma aimg_agen (f : ahom aT rT) U : (f @: agenv U)%VS = agenv (f @: U). +Proof. +apply/eqP; rewrite eqEsubv; apply/andP; split. + by rewrite limg_sum; apply/subv_sumP => i _; rewrite aimgX subX_agenv. +apply: agenv_sub_modl; first by rewrite -(aimg1 f) limgS // sub1_agenv. +by rewrite -aimgM limgS // [rhs in (_ <= rhs)%VS]agenvEl addvSr. +Qed. + +Lemma aimg_adjoin (f : ahom aT rT) U x : (f @: <<U; x>> = <<f @: U; f x>>)%VS. +Proof. by rewrite aimg_agen limg_add limg_line. Qed. + +Lemma aimg_adjoin_seq (f : ahom aT rT) U xs : + (f @: <<U & xs>> = <<f @: U & map f xs>>)%VS. +Proof. by rewrite aimg_agen limg_add limg_span. Qed. + +Fact ker_sub_ahom_is_aspace (f g : ahom aT rT) : + is_aspace (lker (ahval f - ahval g)). +Proof. +rewrite /is_aspace has_algid1; last by apply/eqlfunP; rewrite !rmorph1. +apply/prodvP=> a b /eqlfunP Dfa /eqlfunP Dfb. +by apply/eqlfunP; rewrite !rmorphM /= Dfa Dfb. +Qed. +Canonical ker_sub_ahom_aspace f g := ASpace (ker_sub_ahom_is_aspace f g). + +End LRMorphism. + +Canonical fixedSpace_aspace aT (f : ahom aT aT) := [aspace of fixedSpace f]. + +End AHom. + +Implicit Arguments ahom_in [K aT rT]. + +Notation "''AHom' ( aT , rT )" := (ahom aT rT) : type_scope. +Notation "''AEnd' ( aT )" := (ahom aT aT) : type_scope. + +Delimit Scope lrfun_scope with AF. +Bind Scope lrfun_scope with ahom. + +Notation "\1" := (@id_ahom _ _) : lrfun_scope. +Notation "f \o g" := (comp_ahom f g) : lrfun_scope. diff --git a/mathcomp/field/fieldext.v b/mathcomp/field/fieldext.v new file mode 100644 index 0000000..d9b181a --- /dev/null +++ b/mathcomp/field/fieldext.v @@ -0,0 +1,1626 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div choice fintype. +Require Import tuple finfun bigop ssralg finalg zmodp matrix vector falgebra. +Require Import poly polydiv mxpoly generic_quotient. + +(******************************************************************************) +(* * Finite dimensional field extentions *) +(* fieldExtType F == the interface type for finite field extensions of F *) +(* it simply combines the fieldType and FalgType F *) +(* interfaces. *) +(* [fieldExtType F of L] == a fieldExt F structure for a type L that has both *) +(* fieldType and FalgType F canonical structures. *) +(* [fieldExtType F of L for K] == a fieldExtType F structure for a type L *) +(* that has an FalgType F canonical structure, given *) +(* a K : fieldType whose unitRingType projection *) +(* coincides with the canonical unitRingType for F. *) +(* {subfield L} == the type of subfields of L that are also extensions *) +(* of F; since we are in a finite dimensional setting *) +(* these are exactly the F-subalgebras of L, and *) +(* indeed {subfield L} is just display notation for *) +(* {aspace L} when L is an extFieldType. *) +(* --> All aspace operations apply to {subfield L}, but there are several *) +(* additional lemmas and canonical instances specific to {subfield L} *) +(* spaces, e.g., subvs_of E is an extFieldType F when E : {subfield L}. *) +(* --> Also note that not all constructive subfields have type {subfield E} *) +(* in the same way that not all constructive subspaces have type *) +(* {vspace E}. These types only include the so called "detachable" *) +(* subspaces (and subalgebras). *) +(* *) +(* (E :&: F)%AS, (E * F)%AS == the intersection and product (meet and join) *) +(* of E and F as subfields. *) +(* subFExtend iota z p == Given a field morphism iota : F -> L, this is a *) +(* type for the field F^iota(z) obtained by *) +(* adjoining z to the image of F in L under iota. *) +(* The construction requires a non-zero polynomial *) +(* p in F such that z is a root of p^iota; it *) +(* returns the field F^iota if this is not so. *) +(* However, p need not be irredicible. *) +(* subfx_inj x == The injection of F^iota(z) into L. *) +(* inj_subfx iota z p x == The injection of F into F^iota(z). *) +(* subfx_eval iota z p q == Given q : {poly F} returns q.[z] as a value of *) +(* type F^iota(z). *) +(* subfx_root iota z p == The generator of F^iota(z) over F. *) +(* SubFieldExtType pz0 irr_p == A fieldExtType F structure for F^iota(z) *) +(* (more precisely, subFExtend iota z p), given *) +(* proofs pz0: root (map_poly iota p) z and *) +(* irr_p : irreducible_poly p. The corresponding *) +(* vectType substructure (SubfxVectType pz0 irr_p) *) +(* has dimension (size p).-1 over F. *) +(* minPoly K x == the monic minimal polynomial of x over the *) +(* subfield K. *) +(* adjoin_degree K x == the degree of the minimial polynomial or the *) +(* dimension of K(x)/K. *) +(* Fadjoin_poly K x y == a polynomial p over K such that y = p.[x]. *) +(* *) +(* fieldOver F == L, but with an extFieldType (subvs_of F) *) +(* structure, for F : {subfield L} *) +(* vspaceOver F V == the smallest subspace of fieldOver F containing *) +(* V; this coincides with V if V is an F-module. *) +(* baseFieldType L == L, but with an extFieldType F0 structure, when L *) +(* has a canonical extFieldType F structure and F *) +(* in turn has an extFieldType F0 structure. *) +(* baseVspace V == the subspace of baseFieldType L that coincides *) +(* with V : {vspace L}. *) +(* --> Some caution muse be exercised when using fieldOver and basFieldType, *) +(* because these are convertible to L while carrying different Lmodule *) +(* structures. This means that the safeguards engineered in the ssralg *) +(* library that normally curb the Coq kernel's inclination to diverge are *) +(* no longer effectcive, so additional precautions should be taken when *) +(* matching or rewriting terms of the form a *: u, because Coq may take *) +(* forever to realize it's dealing with a *: in the wrong structure. The *) +(* baseField_scaleE and fieldOver_scaleE lemmas should be used to expand *) +(* or fold such "trans-structure" operations explicitly beforehand. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope ring_scope. +Import GRing.Theory. + +Module FieldExt. + +Import GRing. + +Section FieldExt. + +Variable R : ringType. + +Record class_of T := Class { + base : Falgebra.class_of R T; + comm_ext : commutative (Ring.mul base); + idomain_ext : IntegralDomain.axiom (Ring.Pack base T); + field_ext : Field.mixin_of (UnitRing.Pack base T) +}. + +Local Coercion base : class_of >-> Falgebra.class_of. + +Section Bases. +Variables (T : Type) (c : class_of T). +Definition base1 := ComRing.Class (@comm_ext T c). +Definition base2 := @ComUnitRing.Class T base1 c. +Definition base3 := @IntegralDomain.Class T base2 (@idomain_ext T c). +Definition base4 := @Field.Class T base3 (@field_ext T c). +End Bases. +Local Coercion base1 : class_of >-> ComRing.class_of. +Local Coercion base2 : class_of >-> ComUnitRing.class_of. +Local Coercion base3 : class_of >-> IntegralDomain.class_of. +Local Coercion base4 : class_of >-> Field.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. +Let xT := let: Pack T _ _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack := + fun (bT : Falgebra.type phR) b + & phant_id (Falgebra.class bT : Falgebra.class_of R bT) + (b : Falgebra.class_of R T) => + fun mT Cm IDm Fm & phant_id (Field.class mT) (@Field.Class T + (@IntegralDomain.Class T (@ComUnitRing.Class T (@ComRing.Class T b + Cm) b) IDm) Fm) => Pack phR (@Class T b Cm IDm Fm) T. + +Definition pack_eta K := + let cK := Field.class K in let Cm := ComRing.mixin cK in + let IDm := IntegralDomain.mixin cK in let Fm := Field.mixin cK in + fun (bT : Falgebra.type phR) b & phant_id (Falgebra.class bT) b => + fun cT_ & phant_id (@Class T b) cT_ => @Pack phR T (cT_ Cm IDm Fm) 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 comRingType := @ComRing.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 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 unitAlgType := @UnitAlgebra.Pack R phR cT xclass xT. +Definition vectType := @Vector.Pack R phR cT xclass xT. +Definition FalgType := @Falgebra.Pack R phR cT xclass xT. + +Definition Falg_comRingType := @ComRing.Pack FalgType xclass xT. +Definition Falg_comUnitRingType := @ComUnitRing.Pack FalgType xclass xT. +Definition Falg_idomainType := @IntegralDomain.Pack FalgType xclass xT. +Definition Falg_fieldType := @Field.Pack FalgType xclass xT. + +Definition vect_comRingType := @ComRing.Pack vectType xclass xT. +Definition vect_comUnitRingType := @ComUnitRing.Pack vectType xclass xT. +Definition vect_idomainType := @IntegralDomain.Pack vectType xclass xT. +Definition vect_fieldType := @Field.Pack vectType xclass xT. + +Definition unitAlg_comRingType := @ComRing.Pack unitAlgType xclass xT. +Definition unitAlg_comUnitRingType := @ComUnitRing.Pack unitAlgType xclass xT. +Definition unitAlg_idomainType := @IntegralDomain.Pack unitAlgType xclass xT. +Definition unitAlg_fieldType := @Field.Pack unitAlgType xclass xT. + +Definition alg_comRingType := @ComRing.Pack algType xclass xT. +Definition alg_comUnitRingType := @ComUnitRing.Pack algType xclass xT. +Definition alg_idomainType := @IntegralDomain.Pack algType xclass xT. +Definition alg_fieldType := @Field.Pack algType xclass xT. + +Definition lalg_comRingType := @ComRing.Pack lalgType xclass xT. +Definition lalg_comUnitRingType := @ComUnitRing.Pack lalgType xclass xT. +Definition lalg_idomainType := @IntegralDomain.Pack lalgType xclass xT. +Definition lalg_fieldType := @Field.Pack lalgType xclass xT. + +Definition lmod_comRingType := @ComRing.Pack lmodType xclass xT. +Definition lmod_comUnitRingType := @ComUnitRing.Pack lmodType xclass xT. +Definition lmod_idomainType := @IntegralDomain.Pack lmodType xclass xT. +Definition lmod_fieldType := @Field.Pack lmodType xclass xT. + +End FieldExt. + +Module Exports. + +Coercion sort : type >-> Sortclass. +Bind Scope ring_scope with sort. +Coercion base : class_of >-> Falgebra.class_of. +Coercion base4 : class_of >-> Field.class_of. +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 comRingType : type >-> ComRing.type. +Canonical comRingType. +Coercion comUnitRingType : type >-> ComUnitRing.type. +Canonical comUnitRingType. +Coercion idomainType : type >-> IntegralDomain.type. +Canonical idomainType. +Coercion fieldType : type >-> Field.type. +Canonical fieldType. +Coercion lmodType : type >-> Lmodule.type. +Canonical lmodType. +Coercion lalgType : type >-> Lalgebra.type. +Canonical lalgType. +Coercion algType : type >-> Algebra.type. +Canonical algType. +Coercion unitAlgType : type >-> UnitAlgebra.type. +Canonical unitAlgType. +Coercion vectType : type >-> Vector.type. +Canonical vectType. +Coercion FalgType : type >-> Falgebra.type. +Canonical FalgType. + +Canonical Falg_comRingType. +Canonical Falg_comUnitRingType. +Canonical Falg_idomainType. +Canonical Falg_fieldType. +Canonical vect_comRingType. +Canonical vect_comUnitRingType. +Canonical vect_idomainType. +Canonical vect_fieldType. +Canonical unitAlg_comRingType. +Canonical unitAlg_comUnitRingType. +Canonical unitAlg_idomainType. +Canonical unitAlg_fieldType. +Canonical alg_comRingType. +Canonical alg_comUnitRingType. +Canonical alg_idomainType. +Canonical alg_fieldType. +Canonical lalg_comRingType. +Canonical lalg_comUnitRingType. +Canonical lalg_idomainType. +Canonical lalg_fieldType. +Canonical lmod_comRingType. +Canonical lmod_comUnitRingType. +Canonical lmod_idomainType. +Canonical lmod_fieldType. +Notation fieldExtType R := (type (Phant R)). + +Notation "[ 'fieldExtType' F 'of' L ]" := + (@pack _ (Phant F) L _ _ id _ _ _ _ id) + (at level 0, format "[ 'fieldExtType' F 'of' L ]") : form_scope. +(*Notation "[ 'fieldExtType' F 'of' L 'for' K ]" := + (@FieldExt.pack _ (Phant F) L _ _ id K _ _ _ idfun) + (at level 0, format "[ 'fieldExtType' F 'of' L 'for' K ]") : form_scope. +*) +Notation "[ 'fieldExtType' F 'of' L 'for' K ]" := + (@pack_eta _ (Phant F) L K _ _ id _ id) + (at level 0, format "[ 'fieldExtType' F 'of' L 'for' K ]") : form_scope. + +Notation "{ 'subfield' L }" := (@aspace_of _ (FalgType _) (Phant L)) + (at level 0, format "{ 'subfield' L }") : type_scope. + +End Exports. +End FieldExt. +Export FieldExt.Exports. + +Section FieldExtTheory. + +Variables (F0 : fieldType) (L : fieldExtType F0). +Implicit Types (U V M : {vspace L}) (E F K : {subfield L}). + +Lemma dim_cosetv U x : x != 0 -> \dim (U * <[x]>) = \dim U. +Proof. +move=> nz_x; rewrite -limg_amulr limg_dim_eq //. +apply/eqP; rewrite -subv0; apply/subvP=> y. +by rewrite memv_cap memv0 memv_ker lfunE mulf_eq0 (negPf nz_x) orbF => /andP[]. +Qed. + +Lemma prodvC : commutative (@prodv F0 L). +Proof. +move=> U V; without loss suffices subC: U V / (U * V <= V * U)%VS. + by apply/eqP; rewrite eqEsubv !{1}subC. +by apply/prodvP=> x y Ux Vy; rewrite mulrC memv_mul. +Qed. +Canonical prodv_comoid := Monoid.ComLaw prodvC. + +Lemma prodvCA : left_commutative (@prodv F0 L). +Proof. exact: Monoid.mulmCA. Qed. + +Lemma prodvAC : right_commutative (@prodv F0 L). +Proof. exact: Monoid.mulmAC. Qed. + +Lemma algid1 K : algid K = 1. Proof. exact/skew_field_algid1/fieldP. Qed. + +Lemma mem1v K : 1 \in K. Proof. by rewrite -algid_eq1 algid1. Qed. +Lemma sub1v K : (1 <= K)%VS. Proof. exact: mem1v. Qed. + +Lemma subfield_closed K : agenv K = K. +Proof. +by apply/eqP; rewrite eqEsubv sub_agenv agenv_sub_modr ?sub1v ?asubv. +Qed. + +Lemma AHom_lker0 (rT : FalgType F0) (f : 'AHom(L, rT)) : lker f == 0%VS. +Proof. by apply/lker0P; apply: fmorph_inj. Qed. + +Lemma AEnd_lker0 (f : 'AEnd(L)) : lker f == 0%VS. Proof. exact: AHom_lker0. Qed. + +Fact aimg_is_aspace (rT : FalgType F0) (f : 'AHom(L, rT)) (E : {subfield L}) : + is_aspace (f @: E). +Proof. +rewrite /is_aspace -aimgM limgS ?prodv_id // has_algid1 //. +by apply/memv_imgP; exists 1; rewrite ?mem1v ?rmorph1. +Qed. +Canonical aimg_aspace rT f E := ASpace (@aimg_is_aspace rT f E). + +Lemma Fadjoin_idP {K x} : reflect (<<K; x>>%VS = K) (x \in K). +Proof. +apply: (iffP idP) => [/addv_idPl-> | <-]; first exact: subfield_closed. +exact: memv_adjoin. +Qed. + +Lemma Fadjoin0 K : <<K; 0>>%VS = K. +Proof. by rewrite addv0 subfield_closed. Qed. + +Lemma Fadjoin_nil K : <<K & [::]>>%VS = K. +Proof. by rewrite adjoin_nil subfield_closed. Qed. + +Lemma FadjoinP {K x E} : + reflect (K <= E /\ x \in E)%VS (<<K; x>>%AS <= E)%VS. +Proof. +apply: (iffP idP) => [sKxE | /andP]. + by rewrite (subvP sKxE) ?memv_adjoin // (subv_trans _ sKxE) ?subv_adjoin. +by rewrite -subv_add => /agenvS; rewrite subfield_closed. +Qed. + +Lemma Fadjoin_seqP {K} {rs : seq L} {E} : + reflect (K <= E /\ {subset rs <= E})%VS (<<K & rs>> <= E)%VS. +Proof. +apply: (iffP idP) => [sKrsE | [sKE /span_subvP/(conj sKE)/andP]]. + split=> [|x rs_x]; first exact: subv_trans (subv_adjoin_seq _ _) sKrsE. + by rewrite (subvP sKrsE) ?seqv_sub_adjoin. +by rewrite -subv_add => /agenvS; rewrite subfield_closed. +Qed. + +Lemma alg_polyOver E p : map_poly (in_alg L) p \is a polyOver E. +Proof. by apply/(polyOverS (subvP (sub1v _)))/polyOver1P; exists p. Qed. + +Lemma sub_adjoin1v x E : (<<1; x>> <= E)%VS = (x \in E)%VS. +Proof. by rewrite (sameP FadjoinP andP) sub1v. Qed. + +Fact vsval_multiplicative K : multiplicative (vsval : subvs_of K -> L). +Proof. by split => //=; apply: algid1. Qed. +Canonical vsval_rmorphism K := AddRMorphism (vsval_multiplicative K). +Canonical vsval_lrmorphism K := [lrmorphism of (vsval : subvs_of K -> L)]. + +Lemma vsval_invf K (w : subvs_of K) : val w^-1 = (vsval w)^-1. +Proof. +have [-> | Uv] := eqVneq w 0; first by rewrite !invr0. +by apply: vsval_invr; rewrite unitfE. +Qed. + +Fact aspace_divr_closed K : divr_closed K. +Proof. by split=> [|u v Ku Kv]; rewrite ?mem1v ?memvM ?memvV. Qed. +Canonical aspace_mulrPred K := MulrPred (aspace_divr_closed K). +Canonical aspace_divrPred K := DivrPred (aspace_divr_closed K). +Canonical aspace_smulrPred K := SmulrPred (aspace_divr_closed K). +Canonical aspace_sdivrPred K := SdivrPred (aspace_divr_closed K). +Canonical aspace_semiringPred K := SemiringPred (aspace_divr_closed K). +Canonical aspace_subringPred K := SubringPred (aspace_divr_closed K). +Canonical aspace_subalgPred K := SubalgPred (memv_submod_closed K). +Canonical aspace_divringPred K := DivringPred (aspace_divr_closed K). +Canonical aspace_divalgPred K := DivalgPred (memv_submod_closed K). + +Definition subvs_mulC K := [comRingMixin of subvs_of K by <:]. +Canonical subvs_comRingType K := + Eval hnf in ComRingType (subvs_of K) (@subvs_mulC K). +Canonical subvs_comUnitRingType K := + Eval hnf in [comUnitRingType of subvs_of K]. +Definition subvs_mul_eq0 K := [idomainMixin of subvs_of K by <:]. +Canonical subvs_idomainType K := + Eval hnf in IdomainType (subvs_of K) (@subvs_mul_eq0 K). +Lemma subvs_fieldMixin K : GRing.Field.mixin_of (@subvs_idomainType K). +Proof. +by move=> w nz_w; rewrite unitrE -val_eqE /= vsval_invf algid1 divff. +Qed. +Canonical subvs_fieldType K := + Eval hnf in FieldType (subvs_of K) (@subvs_fieldMixin K). +Canonical subvs_fieldExtType K := Eval hnf in [fieldExtType F0 of subvs_of K]. + +Lemma polyOver_subvs {K} {p : {poly L}} : + reflect (exists q : {poly subvs_of K}, p = map_poly vsval q) + (p \is a polyOver K). +Proof. +apply: (iffP polyOverP) => [Hp | [q ->] i]; last by rewrite coef_map // subvsP. +exists (\poly_(i < size p) (Subvs (Hp i))); rewrite -{1}[p]coefK. +by apply/polyP=> i; rewrite coef_map !coef_poly; case: ifP. +Qed. + +Lemma divp_polyOver K : {in polyOver K &, forall p q, p %/ q \is a polyOver K}. +Proof. +move=> _ _ /polyOver_subvs[p ->] /polyOver_subvs[q ->]. +by apply/polyOver_subvs; exists (p %/ q); rewrite map_divp. +Qed. + +Lemma modp_polyOver K : {in polyOver K &, forall p q, p %% q \is a polyOver K}. +Proof. +move=> _ _ /polyOver_subvs[p ->] /polyOver_subvs[q ->]. +by apply/polyOver_subvs; exists (p %% q); rewrite map_modp. +Qed. + +Lemma gcdp_polyOver K : + {in polyOver K &, forall p q, gcdp p q \is a polyOver K}. +Proof. +move=> _ _ /polyOver_subvs[p ->] /polyOver_subvs[q ->]. +by apply/polyOver_subvs; exists (gcdp p q); rewrite gcdp_map. +Qed. + +Fact prodv_is_aspace E F : is_aspace (E * F). +Proof. +rewrite /is_aspace prodvCA -!prodvA prodvA !prodv_id has_algid1 //=. +by rewrite -[1]mulr1 memv_mul ?mem1v. +Qed. +Canonical prodv_aspace E F : {subfield L} := ASpace (prodv_is_aspace E F). + +Fact field_mem_algid E F : algid E \in F. Proof. by rewrite algid1 mem1v. Qed. +Canonical capv_aspace E F : {subfield L} := aspace_cap (field_mem_algid E F). + +Lemma polyOverSv U V : (U <= V)%VS -> {subset polyOver U <= polyOver V}. +Proof. by move/subvP=> sUV; apply: polyOverS. Qed. + +Lemma field_subvMl F U : (U <= F * U)%VS. +Proof. by rewrite -{1}[U]prod1v prodvSl ?sub1v. Qed. + +Lemma field_subvMr U F : (U <= U * F)%VS. +Proof. by rewrite prodvC field_subvMl. Qed. + +Lemma field_module_eq F M : (F * M <= M)%VS -> (F * M)%VS = M. +Proof. by move=> modM; apply/eqP; rewrite eqEsubv modM field_subvMl. Qed. + +Lemma sup_field_module F E : (F * E <= E)%VS = (F <= E)%VS. +Proof. +apply/idP/idP; first exact: subv_trans (field_subvMr F E). +by move/(prodvSl E)/subv_trans->; rewrite ?asubv. +Qed. + +Lemma field_module_dimS F M : (F * M <= M)%VS -> (\dim F %| \dim M)%N. +Proof. exact/skew_field_module_dimS/fieldP. Qed. + +Lemma field_dimS F E : (F <= E)%VS -> (\dim F %| \dim E)%N. +Proof. exact/skew_field_dimS/fieldP. Qed. + +Lemma dim_field_module F M : (F * M <= M)%VS -> \dim M = (\dim_F M * \dim F)%N. +Proof. by move/field_module_dimS/divnK. Qed. + +Lemma dim_sup_field F E : (F <= E)%VS -> \dim E = (\dim_F E * \dim F)%N. +Proof. by move/field_dimS/divnK. Qed. + +Lemma field_module_semisimple F M (m := \dim_F M) : + (F * M <= M)%VS -> + {X : m.-tuple L | {subset X <= M} /\ 0 \notin X + & let FX := (\sum_(i < m) F * <[X`_i]>)%VS in FX = M /\ directv FX}. +Proof. +move=> modM; have dimM: (m * \dim F)%N = \dim M by rewrite -dim_field_module. +have [X [defM dxFX nzX]] := skew_field_module_semisimple (@fieldP L) modM. +have szX: size X == m. + rewrite -(eqn_pmul2r (adim_gt0 F)) dimM -defM (directvP dxFX) /=. + rewrite -sum1_size big_distrl; apply/eqP/eq_big_seq => x Xx /=. + by rewrite mul1n dim_cosetv ?(memPn nzX). +rewrite directvE /= !(big_nth 0) (eqP szX) !big_mkord -directvE /= in defM dxFX. +exists (Tuple szX) => //; split=> // _ /tnthP[i ->]; rewrite (tnth_nth 0) /=. +by rewrite -defM memvE (sumv_sup i) ?field_subvMl. +Qed. + +Section FadjoinPolyDefinitions. + +Variables (U : {vspace L}) (x : L). + +Definition adjoin_degree := (\dim_U <<U; x>>).-1.+1. +Local Notation n := adjoin_degree. + +Definition Fadjoin_sum := (\sum_(i < n) U * <[x ^+ i]>)%VS. + +Definition Fadjoin_poly v : {poly L} := + \poly_(i < n) (sumv_pi Fadjoin_sum (inord i) v / x ^+ i). + +Definition minPoly : {poly L} := 'X^n - Fadjoin_poly (x ^+ n). + +Lemma size_Fadjoin_poly v : size (Fadjoin_poly v) <= n. +Proof. exact: size_poly. Qed. + +Lemma Fadjoin_polyOver v : Fadjoin_poly v \is a polyOver U. +Proof. +apply/(all_nthP 0) => i _; rewrite coef_poly /=. +case: ifP => lti; last exact: mem0v. +have /memv_cosetP[y Uy ->] := memv_sum_pi (erefl Fadjoin_sum) (inord i) v. +rewrite inordK //; have [-> | /mulfK-> //] := eqVneq (x ^+ i) 0. +by rewrite mulr0 mul0r mem0v. +Qed. + +Fact Fadjoin_poly_is_linear : linear_for (in_alg L \; *:%R) Fadjoin_poly. +Proof. +move=> a u v; apply/polyP=> i; rewrite coefD coefZ !coef_poly. +case: ifP => lti; last by rewrite mulr0 addr0. +by rewrite linearP mulrA -mulrDl mulr_algl. +Qed. +Canonical Fadjoin_poly_additive := Additive Fadjoin_poly_is_linear. +Canonical Fadjoin_poly_linear := AddLinear Fadjoin_poly_is_linear. + +Lemma size_minPoly : size minPoly = n.+1. +Proof. by rewrite size_addl ?size_polyXn // size_opp ltnS size_poly. Qed. + +Lemma monic_minPoly : minPoly \is monic. +Proof. +rewrite monicE /lead_coef size_minPoly coefB coefXn eqxx. +by rewrite nth_default ?subr0 ?size_poly. +Qed. + +End FadjoinPolyDefinitions. + +Section FadjoinPoly. + +Variables (K : {subfield L}) (x : L). +Local Notation n := (adjoin_degree (asval K) x). +Local Notation sumKx := (Fadjoin_sum (asval K) x). + +Lemma adjoin_degreeE : n = \dim_K <<K; x>>. +Proof. by rewrite [n]prednK // divn_gt0 ?adim_gt0 // dimvS ?subv_adjoin. Qed. + +Lemma dim_Fadjoin : \dim <<K; x>> = (n * \dim K)%N. +Proof. by rewrite adjoin_degreeE -dim_sup_field ?subv_adjoin. Qed. + +Lemma adjoin0_deg : adjoin_degree K 0 = 1%N. +Proof. by rewrite /adjoin_degree addv0 subfield_closed divnn adim_gt0. Qed. + +Lemma adjoin_deg_eq1 : (n == 1%N) = (x \in K). +Proof. +rewrite (sameP Fadjoin_idP eqP) adjoin_degreeE; have sK_Kx := subv_adjoin K x. +apply/eqP/idP=> [dimKx1 | /eqP->]; last by rewrite divnn adim_gt0. +by rewrite eq_sym eqEdim sK_Kx /= (dim_sup_field sK_Kx) dimKx1 mul1n. +Qed. + +Lemma Fadjoin_sum_direct : directv sumKx. +Proof. +rewrite directvE /=; case Dn: {-2}n (leqnn n) => // [m] {Dn}. +elim: m => [|m IHm] ltm1n; rewrite ?big_ord1 // !(big_ord_recr m.+1) /=. +do [move/(_ (ltnW ltm1n))/eqP; set S := (\sum_i _)%VS] in IHm *. +rewrite -IHm dimv_add_leqif; apply/subvP=> z; rewrite memv_cap => /andP[Sz]. +case/memv_cosetP=> y Ky Dz; rewrite memv0 Dz mulf_eq0 expf_eq0 /=. +apply: contraLR ltm1n => /norP[nz_y nz_x]. +rewrite -leqNgt -(leq_pmul2r (adim_gt0 K)) -dim_Fadjoin. +have{IHm} ->: (m.+1 * \dim K)%N = \dim S. + rewrite -[m.+1]card_ord -sum_nat_const IHm. + by apply: eq_bigr => i; rewrite dim_cosetv ?expf_neq0. +apply/dimvS/agenv_sub_modl; first by rewrite (sumv_sup 0) //= prodv1 sub1v. +rewrite prodvDl subv_add -[S]big_distrr prodvA prodv_id subvv !big_distrr /=. +apply/subv_sumP=> i _; rewrite -expv_line prodvCA -expvSl expv_line. +have [ltim | lemi] := ltnP i m; first by rewrite (sumv_sup (Sub i.+1 _)). +have{lemi} /eqP->: i == m :> nat by rewrite eqn_leq leq_ord. +rewrite -big_distrr -2!{2}(prodv_id K) /= -!prodvA big_distrr -/S prodvSr //=. +by rewrite -(canLR (mulKf nz_y) Dz) -memvE memv_mul ?rpredV. +Qed. + +Let nz_x_i (i : 'I_n) : x ^+ i != 0. +Proof. +by rewrite expf_eq0; case: eqP i => [->|_] [[]] //; rewrite adjoin0_deg. +Qed. + +Lemma Fadjoin_eq_sum : <<K; x>>%VS = sumKx. +Proof. +apply/esym/eqP; rewrite eqEdim eq_leq ?andbT. + apply/subv_sumP=> i _; rewrite -agenvM prodvS ?subv_adjoin //. + by rewrite -expv_line (subv_trans (subX_agenv _ _)) ?agenvS ?addvSr. +rewrite dim_Fadjoin -[n]card_ord -sum_nat_const (directvP Fadjoin_sum_direct). +by apply: eq_bigr => i _; rewrite /= dim_cosetv. +Qed. + +Lemma Fadjoin_poly_eq v : v \in <<K; x>>%VS -> (Fadjoin_poly K x v).[x] = v. +Proof. +move/(sumv_pi_sum Fadjoin_eq_sum)=> {2}<-; rewrite horner_poly. +by apply: eq_bigr => i _; rewrite inord_val mulfVK. +Qed. + +Lemma mempx_Fadjoin p : p \is a polyOver K -> p.[x] \in <<K; x>>%VS. +Proof. +move=> Kp; rewrite rpred_horner ?memv_adjoin ?(polyOverS _ Kp) //. +exact: subvP_adjoin. +Qed. + +Lemma Fadjoin_polyP {v} : + reflect (exists2 p, p \in polyOver K & v = p.[x]) (v \in <<K; x>>%VS). +Proof. +apply: (iffP idP) => [Kx_v | [p Kp ->]]; last exact: mempx_Fadjoin. +by exists (Fadjoin_poly K x v); rewrite ?Fadjoin_polyOver ?Fadjoin_poly_eq. +Qed. + +Lemma Fadjoin_poly_unique p v : + p \is a polyOver K -> size p <= n -> p.[x] = v -> Fadjoin_poly K x v = p. +Proof. +have polyKx q i: q \is a polyOver K -> q`_i * x ^+ i \in (K * <[x ^+ i]>)%VS. + by move/polyOverP=> Kq; rewrite memv_mul ?Kq ?memv_line. +move=> Kp szp Dv; have /Fadjoin_poly_eq/eqP := mempx_Fadjoin Kp. +rewrite {1}Dv {Dv} !(@horner_coef_wide _ n) ?size_poly //. +move/polyKx in Kp; have /polyKx K_pv := Fadjoin_polyOver K x v. +rewrite (directv_sum_unique Fadjoin_sum_direct) // => /eqfunP eq_pq. +apply/polyP=> i; have [leni|?] := leqP n i; last exact: mulIf (eq_pq (Sub i _)). +by rewrite !nth_default ?(leq_trans _ leni) ?size_poly. +Qed. + +Lemma Fadjoin_polyC v : v \in K -> Fadjoin_poly K x v = v%:P. +Proof. +move=> Kv; apply: Fadjoin_poly_unique; rewrite ?polyOverC ?hornerC //. +by rewrite size_polyC (leq_trans (leq_b1 _)). +Qed. + +Lemma Fadjoin_polyX : x \notin K -> Fadjoin_poly K x x = 'X. +Proof. +move=> K'x; apply: Fadjoin_poly_unique; rewrite ?polyOverX ?hornerX //. +by rewrite size_polyX ltn_neqAle andbT eq_sym adjoin_deg_eq1. +Qed. + +Lemma minPolyOver : minPoly K x \is a polyOver K. +Proof. by rewrite /minPoly rpredB ?rpredX ?polyOverX ?Fadjoin_polyOver. Qed. + +Lemma minPolyxx : (minPoly K x).[x] = 0. +Proof. +by rewrite !hornerE hornerXn Fadjoin_poly_eq ?subrr ?rpredX ?memv_adjoin. +Qed. + +Lemma root_minPoly : root (minPoly K x) x. Proof. exact/rootP/minPolyxx. Qed. + +Lemma Fadjoin_poly_mod p : + p \is a polyOver K -> Fadjoin_poly K x p.[x] = p %% minPoly K x. +Proof. +move=> Kp; rewrite {1}(divp_eq p (minPoly K x)) 2!hornerE minPolyxx mulr0 add0r. +apply: Fadjoin_poly_unique => //; first by rewrite modp_polyOver // minPolyOver. +by rewrite -ltnS -size_minPoly ltn_modp // monic_neq0 ?monic_minPoly. +Qed. + +Lemma minPoly_XsubC : reflect (minPoly K x = 'X - x%:P) (x \in K). +Proof. +set p := minPoly K x; apply: (iffP idP) => [Kx | Dp]; last first. + suffices ->: x = - p`_0 by rewrite rpredN (polyOverP minPolyOver). + by rewrite Dp coefB coefX coefC add0r opprK. +rewrite (@all_roots_prod_XsubC _ p [:: x]) /= ?root_minPoly //. + by rewrite big_seq1 (monicP (monic_minPoly K x)) scale1r. +by apply/eqP; rewrite size_minPoly eqSS adjoin_deg_eq1. +Qed. + +Lemma root_small_adjoin_poly p : + p \is a polyOver K -> size p <= n -> root p x = (p == 0). +Proof. +move=> Kp szp; apply/rootP/eqP=> [px0 | ->]; last by rewrite horner0. +rewrite -(Fadjoin_poly_unique Kp szp px0). +by apply: Fadjoin_poly_unique; rewrite ?polyOver0 ?size_poly0 ?horner0. +Qed. + +Lemma minPoly_irr p : + p \is a polyOver K -> p %| minPoly K x -> (p %= minPoly K x) || (p %= 1). +Proof. +rewrite dvdp_eq; set q := _ %/ _ => Kp def_pq. +have Kq: q \is a polyOver K by rewrite divp_polyOver // minPolyOver. +move: q Kq def_pq root_minPoly (size_minPoly K x) => q Kq /eqP->. +rewrite rootM => pqx0 szpq. +have [nzq nzp]: q != 0 /\ p != 0. + by apply/norP; rewrite -mulf_eq0 -size_poly_eq0 szpq. +without loss{pqx0} qx0: q p Kp Kq nzp nzq szpq / root q x. + move=> IH; case/orP: pqx0 => /IH{IH}IH; first exact: IH. + have{IH} /orP[]: (q %= p * q) || (q %= 1) by apply: IH => //; rewrite mulrC. + by rewrite orbC -{1}[q]mul1r eqp_mul2r // eqp_sym => ->. + by rewrite -{1}[p]mul1r eqp_sym eqp_mul2r // => ->. +apply/orP; right; rewrite -size_poly_eq1 eqn_leq lt0n size_poly_eq0 nzp andbT. +rewrite -(leq_add2r (size q)) -leq_subLR subn1 -size_mul // mulrC szpq. +by rewrite ltnNge; apply: contra nzq => /(root_small_adjoin_poly Kq) <-. +Qed. + +Lemma minPoly_dvdp p : p \is a polyOver K -> root p x -> (minPoly K x) %| p. +Proof. +move=> Kp rootp. +have gcdK : gcdp (minPoly K x) p \is a polyOver K. + by rewrite gcdp_polyOver ?minPolyOver. +have /orP[gcd_eqK|gcd_eq1] := minPoly_irr gcdK (dvdp_gcdl (minPoly K x) p). + by rewrite -(eqp_dvdl _ gcd_eqK) dvdp_gcdr. +case/negP: (root1 x). +by rewrite -(eqp_root gcd_eq1) root_gcd rootp root_minPoly. +Qed. + +End FadjoinPoly. + +Lemma minPolyS K E a : (K <= E)%VS -> minPoly E a %| minPoly K a. +Proof. +move=> sKE; apply: minPoly_dvdp; last exact: root_minPoly. +by apply: (polyOverSv sKE); rewrite minPolyOver. +Qed. + +Implicit Arguments Fadjoin_polyP [K x v]. +Lemma Fadjoin1_polyP x v : + reflect (exists p, v = (map_poly (in_alg L) p).[x]) (v \in <<1; x>>%VS). +Proof. +apply: (iffP Fadjoin_polyP) => [[_ /polyOver1P]|] [p ->]; first by exists p. +by exists (map_poly (in_alg L) p) => //; apply: alg_polyOver. +Qed. + +Section Horner. + +Variables z : L. + +Definition fieldExt_horner := horner_morph (fun x => mulrC z (in_alg L x)). +Canonical fieldExtHorner_additive := [additive of fieldExt_horner]. +Canonical fieldExtHorner_rmorphism := [rmorphism of fieldExt_horner]. +Lemma fieldExt_hornerC b : fieldExt_horner b%:P = b%:A. +Proof. exact: horner_morphC. Qed. +Lemma fieldExt_hornerX : fieldExt_horner 'X = z. +Proof. exact: horner_morphX. Qed. +Fact fieldExt_hornerZ : scalable fieldExt_horner. +Proof. +move=> a p; rewrite -mul_polyC rmorphM /= fieldExt_hornerC. +by rewrite -scalerAl mul1r. +Qed. +Canonical fieldExt_horner_linear := AddLinear fieldExt_hornerZ. +Canonical fieldExt_horner_lrmorhism := [lrmorphism of fieldExt_horner]. + +End Horner. + +End FieldExtTheory. + +Notation "E :&: F" := (capv_aspace E F) : aspace_scope. +Notation "'C_ E [ x ]" := (capv_aspace E 'C[x]) : aspace_scope. +Notation "'C_ ( E ) [ x ]" := (capv_aspace E 'C[x]) + (only parsing) : aspace_scope. +Notation "'C_ E ( V )" := (capv_aspace E 'C(V)) : aspace_scope. +Notation "'C_ ( E ) ( V )" := (capv_aspace E 'C(V)) + (only parsing) : aspace_scope. +Notation "E * F" := (prodv_aspace E F) : aspace_scope. +Notation "f @: E" := (aimg_aspace f E) : aspace_scope. + +Implicit Arguments Fadjoin_idP [F0 L K x]. +Implicit Arguments FadjoinP [F0 L K x E]. +Implicit Arguments Fadjoin_seqP [F0 L K rs E]. +Implicit Arguments polyOver_subvs [F0 L K p]. +Implicit Arguments Fadjoin_polyP [F0 L K x v]. +Implicit Arguments Fadjoin1_polyP [F0 L x v]. +Implicit Arguments minPoly_XsubC [F0 L K x]. + +Section MapMinPoly. + +Variables (F0 : fieldType) (L rL : fieldExtType F0) (f : 'AHom(L, rL)). +Variables (K : {subfield L}) (x : L). + +Lemma adjoin_degree_aimg : adjoin_degree (f @: K) (f x) = adjoin_degree K x. +Proof. +rewrite !adjoin_degreeE -aimg_adjoin. +by rewrite !limg_dim_eq ?(eqP (AHom_lker0 f)) ?capv0. +Qed. + +Lemma map_minPoly : map_poly f (minPoly K x) = minPoly (f @: K) (f x). +Proof. +set fp := minPoly (f @: K) (f x); pose fM := [rmorphism of f]. +have [p Kp Dp]: exists2 p, p \is a polyOver K & map_poly f p = fp. + have Kfp: fp \is a polyOver (f @: K)%VS by apply: minPolyOver. + exists (map_poly f^-1%VF fp). + apply/polyOver_poly=> j _; have /memv_imgP[y Ky ->] := polyOverP Kfp j. + by rewrite lker0_lfunK ?AHom_lker0. + rewrite -map_poly_comp map_poly_id // => _ /(allP Kfp)/memv_imgP[y _ ->]. + by rewrite /= limg_lfunVK ?memv_img ?memvf. +apply/eqP; rewrite -eqp_monic ?monic_map ?monic_minPoly // -Dp eqp_map. +have: ~~ (p %= 1) by rewrite -size_poly_eq1 -(size_map_poly fM) Dp size_minPoly. +apply: implyP; rewrite implyNb orbC eqp_sym minPoly_irr //. +rewrite -(dvdp_map fM) Dp minPoly_dvdp ?fmorph_root ?root_minPoly //. +by apply/polyOver_poly=> j _; apply/memv_img/polyOverP/minPolyOver. +Qed. + +End MapMinPoly. + +(* Changing up the reference field of a fieldExtType. *) +Section FieldOver. + +Variables (F0 : fieldType) (L : fieldExtType F0) (F : {subfield L}). + +Definition fieldOver of {vspace L} : Type := L. +Local Notation K_F := (subvs_of F). +Local Notation L_F := (fieldOver F). + +Canonical fieldOver_eqType := [eqType of L_F]. +Canonical fieldOver_choiceType := [choiceType of L_F]. +Canonical fieldOver_zmodType := [zmodType of L_F]. +Canonical fieldOver_ringType := [ringType of L_F]. +Canonical fieldOver_unitRingType := [unitRingType of L_F]. +Canonical fieldOver_comRingType := [comRingType of L_F]. +Canonical fieldOver_comUnitRingType := [comUnitRingType of L_F]. +Canonical fieldOver_idomainType := [idomainType of L_F]. +Canonical fieldOver_fieldType := [fieldType of L_F]. + +Definition fieldOver_scale (a : K_F) (u : L_F) : L_F := vsval a * u. +Local Infix "*F:" := fieldOver_scale (at level 40). + +Fact fieldOver_scaleA a b u : a *F: (b *F: u) = (a * b) *F: u. +Proof. exact: mulrA. Qed. + +Fact fieldOver_scale1 u : 1 *F: u = u. +Proof. by rewrite /(1 *F: u) /= algid1 mul1r. Qed. + +Fact fieldOver_scaleDr a u v : a *F: (u + v) = a *F: u + a *F: v. +Proof. exact: mulrDr. Qed. + +Fact fieldOver_scaleDl v a b : (a + b) *F: v = a *F: v + b *F: v. +Proof. exact: mulrDl. Qed. + +Definition fieldOver_lmodMixin := + LmodMixin fieldOver_scaleA fieldOver_scale1 + fieldOver_scaleDr fieldOver_scaleDl. + +Canonical fieldOver_lmodType := LmodType K_F L_F fieldOver_lmodMixin. + +Lemma fieldOver_scaleE a (u : L) : a *: (u : L_F) = vsval a * u. +Proof. by []. Qed. + +Fact fieldOver_scaleAl a u v : a *F: (u * v) = (a *F: u) * v. +Proof. exact: mulrA. Qed. + +Canonical fieldOver_lalgType := LalgType K_F L_F fieldOver_scaleAl. + +Fact fieldOver_scaleAr a u v : a *F: (u * v) = u * (a *F: v). +Proof. exact: mulrCA. Qed. + +Canonical fieldOver_algType := AlgType K_F L_F fieldOver_scaleAr. +Canonical fieldOver_unitAlgType := [unitAlgType K_F of L_F]. + +Fact fieldOver_vectMixin : Vector.mixin_of fieldOver_lmodType. +Proof. +have [bL [_ nz_bL] [defL dxSbL]] := field_module_semisimple (subvf (F * _)). +do [set n := \dim_F {:L} in bL nz_bL *; set SbL := (\sum_i _)%VS] in defL dxSbL. +have in_bL i (a : K_F) : val a * (bL`_i : L_F) \in (F * <[bL`_i]>)%VS. + by rewrite memv_mul ?(valP a) ?memv_line. +have nz_bLi (i : 'I_n): bL`_i != 0 by rewrite (memPn nz_bL) ?memt_nth. +pose r2v (v : 'rV[K_F]_n) : L_F := \sum_i v 0 i *: (bL`_i : L_F). +have r2v_lin: linear r2v. + move=> a u v; rewrite /r2v scaler_sumr -big_split /=; apply: eq_bigr => i _. + by rewrite scalerA -scalerDl !mxE. +have v2rP x: {r : 'rV[K_F]_n | x = r2v r}. + apply: sig_eqW; have /memv_sumP[y Fy ->]: x \in SbL by rewrite defL memvf. + have /fin_all_exists[r Dr] i: exists r, y i = r *: (bL`_i : L_F). + by have /memv_cosetP[a Fa ->] := Fy i isT; exists (Subvs Fa). + by exists (\row_i r i); apply: eq_bigr => i _; rewrite mxE. +pose v2r x := sval (v2rP x). +have v2rK: cancel v2r (Linear r2v_lin) by rewrite /v2r => x; case: (v2rP x). +suffices r2vK: cancel r2v v2r. + by exists n, v2r; [exact: can2_linear v2rK | exists r2v]. +move=> r; apply/rowP=> i; apply/val_inj/(mulIf (nz_bLi i))/eqP; move: i isT. +by apply/forall_inP; move/directv_sum_unique: dxSbL => <- //; exact/eqP/v2rK. +Qed. + +Canonical fieldOver_vectType := VectType K_F L_F fieldOver_vectMixin. +Canonical fieldOver_FalgType := [FalgType K_F of L_F]. +Canonical fieldOver_fieldExtType := [fieldExtType K_F of L_F]. + +Implicit Types (V : {vspace L}) (E : {subfield L}). + +Lemma trivial_fieldOver : (1%VS : {vspace L_F}) =i F. +Proof. +move=> x; apply/vlineP/idP=> [[{x}x ->] | Fx]. + by rewrite fieldOver_scaleE mulr1 (valP x). +by exists (vsproj F x); rewrite fieldOver_scaleE mulr1 vsprojK. +Qed. + +Definition vspaceOver V := <<vbasis V : seq L_F>>%VS. + +Lemma mem_vspaceOver V : vspaceOver V =i (F * V)%VS. +Proof. +move=> y; apply/idP/idP; last rewrite unlock; move=> /coord_span->. + rewrite (@memv_suml F0 L) // => i _. + by rewrite memv_mul ?subvsP // vbasis_mem ?memt_nth. +rewrite memv_suml // => ij _; rewrite -tnth_nth; set x := tnth _ ij. +have/allpairsP[[u z] /= [Fu Vz {x}->]]: x \in _ := mem_tnth ij _. +by rewrite scalerAl (memvZ (Subvs _)) ?memvZ ?memv_span //= vbasis_mem. +Qed. + +Lemma mem_aspaceOver E : (F <= E)%VS -> vspaceOver E =i E. +Proof. +by move=> sFE y; rewrite mem_vspaceOver field_module_eq ?sup_field_module. +Qed. + +Fact aspaceOver_suproof E : is_aspace (vspaceOver E). +Proof. +rewrite /is_aspace has_algid1; last by rewrite mem_vspaceOver (@mem1v _ L). +by apply/prodvP=> u v; rewrite !mem_vspaceOver; exact: memvM. +Qed. +Canonical aspaceOver E := ASpace (aspaceOver_suproof E). + +Lemma dim_vspaceOver M : (F * M <= M)%VS -> \dim (vspaceOver M) = \dim_F M. +Proof. +move=> modM; have [] := field_module_semisimple modM. +set n := \dim_F M => b [Mb nz_b] [defM dx_b]. +suff: basis_of (vspaceOver M) b by apply: size_basis. +apply/andP; split. + rewrite eqEsubv; apply/andP; split; apply/span_subvP=> u. + by rewrite mem_vspaceOver field_module_eq // => /Mb. + move/(@vbasis_mem _ _ _ M); rewrite -defM => /memv_sumP[{u}u Fu ->]. + apply: memv_suml => i _; have /memv_cosetP[a Fa ->] := Fu i isT. + by apply: (memvZ (Subvs Fa)); rewrite memv_span ?memt_nth. +apply/freeP=> a /(directv_sum_independent dx_b) a_0 i. +have{a_0}: a i *: (b`_i : L_F) == 0. + by rewrite a_0 {i}// => i _; rewrite memv_mul ?memv_line ?subvsP. +by rewrite scaler_eq0=> /predU1P[] // /idPn[]; rewrite (memPn nz_b) ?memt_nth. +Qed. + +Lemma dim_aspaceOver E : (F <= E)%VS -> \dim (vspaceOver E) = \dim_F E. +Proof. by rewrite -sup_field_module; exact: dim_vspaceOver. Qed. + +Lemma vspaceOverP V_F : + {V | [/\ V_F = vspaceOver V, (F * V <= V)%VS & V_F =i V]}. +Proof. +pose V := (F * <<vbasis V_F : seq L>>)%VS. +have idV: (F * V)%VS = V by rewrite prodvA prodv_id. +suffices defVF: V_F = vspaceOver V. + by exists V; split=> [||u]; rewrite ?defVF ?mem_vspaceOver ?idV. +apply/vspaceP=> v; rewrite mem_vspaceOver idV. +do [apply/idP/idP; last rewrite /V unlock] => [/coord_vbasis|/coord_span] ->. + by apply: memv_suml => i _; rewrite memv_mul ?subvsP ?memv_span ?memt_nth. +apply: memv_suml => i _; rewrite -tnth_nth; set xu := tnth _ i. +have /allpairsP[[x u] /=]: xu \in _ := mem_tnth i _. +case=> /vbasis_mem Fx /vbasis_mem Vu ->. +rewrite scalerAl (coord_span Vu) mulr_sumr memv_suml // => j_. +by rewrite -scalerCA (memvZ (Subvs _)) ?memvZ // vbasis_mem ?memt_nth. +Qed. + +Lemma aspaceOverP (E_F : {subfield L_F}) : + {E | [/\ E_F = aspaceOver E, (F <= E)%VS & E_F =i E]}. +Proof. +have [V [defEF modV memV]] := vspaceOverP E_F. +have algE: has_algid V && (V * V <= V)%VS. + rewrite has_algid1; last by rewrite -memV mem1v. + by apply/prodvP=> u v; rewrite -!memV; exact: memvM. +by exists (ASpace algE); rewrite -sup_field_module; split; first exact: val_inj. +Qed. + +End FieldOver. + +(* Changing the reference field to a smaller field. *) +Section BaseField. + +Variables (F0 : fieldType) (F : fieldExtType F0) (L : fieldExtType F). + +Definition baseField_type of phant L : Type := L. +Notation L0 := (baseField_type (Phant (FieldExt.sort L))). + +Canonical baseField_eqType := [eqType of L0]. +Canonical baseField_choiceType := [choiceType of L0]. +Canonical baseField_zmodType := [zmodType of L0]. +Canonical baseField_ringType := [ringType of L0]. +Canonical baseField_unitRingType := [unitRingType of L0]. +Canonical baseField_comRingType := [comRingType of L0]. +Canonical baseField_comUnitRingType := [comUnitRingType of L0]. +Canonical baseField_idomainType := [idomainType of L0]. +Canonical baseField_fieldType := [fieldType of L0]. + +Definition baseField_scale (a : F0) (u : L0) : L0 := in_alg F a *: u. +Local Infix "*F0:" := baseField_scale (at level 40). + +Fact baseField_scaleA a b u : a *F0: (b *F0: u) = (a * b) *F0: u. +Proof. by rewrite [_ *F0: _]scalerA -rmorphM. Qed. + +Fact baseField_scale1 u : 1 *F0: u = u. +Proof. by rewrite /(1 *F0: u) rmorph1 scale1r. Qed. + +Fact baseField_scaleDr a u v : a *F0: (u + v) = a *F0: u + a *F0: v. +Proof. exact: scalerDr. Qed. + +Fact baseField_scaleDl v a b : (a + b) *F0: v = a *F0: v + b *F0: v. +Proof. by rewrite -scalerDl -rmorphD. Qed. + +Definition baseField_lmodMixin := + LmodMixin baseField_scaleA baseField_scale1 + baseField_scaleDr baseField_scaleDl. + +Canonical baseField_lmodType := LmodType F0 L0 baseField_lmodMixin. + +Lemma baseField_scaleE a (u : L) : a *: (u : L0) = a%:A *: u. +Proof. by []. Qed. + +Fact baseField_scaleAl a (u v : L0) : a *F0: (u * v) = (a *F0: u) * v. +Proof. exact: scalerAl. Qed. + +Canonical baseField_lalgType := LalgType F0 L0 baseField_scaleAl. + +Fact baseField_scaleAr a u v : a *F0: (u * v) = u * (a *F0: v). +Proof. exact: scalerAr. Qed. + +Canonical baseField_algType := AlgType F0 L0 baseField_scaleAr. +Canonical baseField_unitAlgType := [unitAlgType F0 of L0]. + +Let n := \dim {:F}. +Let bF : n.-tuple F := vbasis {:F}. +Let coordF (x : F) := (coord_vbasis (memvf x)). + +Fact baseField_vectMixin : Vector.mixin_of baseField_lmodType. +Proof. +pose bL := vbasis {:L}; set m := \dim {:L} in bL. +pose v2r (x : L0) := mxvec (\matrix_(i, j) coord bF j (coord bL i x)). +have v2r_lin: linear v2r. + move=> a x y; rewrite -linearP; congr (mxvec _); apply/matrixP=> i j. + by rewrite !mxE linearP mulr_algl linearP. +pose r2v r := \sum_(i < m) (\sum_(j < n) vec_mx r i j *: bF`_j) *: bL`_i. +have v2rK: cancel v2r r2v. + move=> x; transitivity (\sum_(i < m) coord bL i x *: bL`_i); last first. + by rewrite -coord_vbasis ?memvf. + (* GG: rewrite {2}(coord_vbasis (memvf x)) -/m would take 8s; *) + (* The -/m takes 8s, and without it then apply: eq_bigr takes 12s. *) + (* The time drops to 2s with a -[GRing.Field.ringType F]/(F : fieldType) *) + apply: eq_bigr => i _; rewrite mxvecK; congr (_ *: _ : L). + by rewrite (coordF (coord bL i x)); apply: eq_bigr => j _; rewrite mxE. +exists (m * n)%N, v2r => //; exists r2v => // r. +apply: (canLR vec_mxK); apply/matrixP=> i j; rewrite mxE. +by rewrite !coord_sum_free ?(basis_free (vbasisP _)). +Qed. + +Canonical baseField_vectType := VectType F0 L0 baseField_vectMixin. +Canonical baseField_FalgType := [FalgType F0 of L0]. +Canonical baseField_extFieldType := [fieldExtType F0 of L0]. + +Let F0ZEZ a x v : a *: ((x *: v : L) : L0) = (a *: x) *: v. +Proof. by rewrite [a *: _]scalerA -scalerAl mul1r. Qed. + +Let baseVspace_basis V : seq L0 := + [seq tnth bF ij.2 *: tnth (vbasis V) ij.1 | ij : 'I_(\dim V) * 'I_n]. +Definition baseVspace V := <<baseVspace_basis V>>%VS. + +Lemma mem_baseVspace V : baseVspace V =i V. +Proof. +move=> y; apply/idP/idP=> [/coord_span->|/coord_vbasis->]; last first. + apply: memv_suml => i _; rewrite (coordF (coord _ i (y : L))) scaler_suml -/n. + apply: memv_suml => j _; rewrite -/bF -F0ZEZ memvZ ?memv_span // -!tnth_nth. + by apply/imageP; exists (i, j). + (* GG: the F0ZEZ lemma avoids serious performance issues here. *) +apply: memv_suml => k _; rewrite nth_image; case: (enum_val k) => i j /=. +by rewrite F0ZEZ memvZ ?vbasis_mem ?mem_tnth. +Qed. + +Lemma dim_baseVspace V : \dim (baseVspace V) = (\dim V * n)%N. +Proof. +pose bV0 := baseVspace_basis V; set m := \dim V in bV0 *. +suffices /size_basis->: basis_of (baseVspace V) bV0. + by rewrite card_prod !card_ord. +rewrite /basis_of eqxx. +apply/freeP=> s sb0 k; rewrite -(enum_valK k); case/enum_val: k => i j. +have free_baseP := freeP (basis_free (vbasisP _)). +move: j; apply: (free_baseP _ _ fullv); move: i; apply: (free_baseP _ _ V). +transitivity (\sum_i \sum_j s (enum_rank (i, j)) *: bV0`_(enum_rank (i, j))). + apply: eq_bigr => i _; rewrite scaler_suml; apply: eq_bigr => j _. + by rewrite -F0ZEZ nth_image enum_rankK -!tnth_nth. +rewrite pair_bigA (reindex _ (onW_bij _ (enum_val_bij _))); apply: etrans sb0. +by apply: eq_bigr => k _; rewrite -{5 6}[k](enum_valK k); case/enum_val: k. +Qed. + +Fact baseAspace_suproof (E : {subfield L}) : is_aspace (baseVspace E). +Proof. +rewrite /is_aspace has_algid1; last by rewrite mem_baseVspace (mem1v E). +by apply/prodvP=> u v; rewrite !mem_baseVspace; exact: memvM. +Qed. +Canonical baseAspace E := ASpace (baseAspace_suproof E). + +Fact refBaseField_key : unit. Proof. by []. Qed. +Definition refBaseField := locked_with refBaseField_key (baseAspace 1). +Canonical refBaseField_unlockable := [unlockable of refBaseField]. +Notation F1 := refBaseField. + +Lemma dim_refBaseField : \dim F1 = n. +Proof. by rewrite [F1]unlock dim_baseVspace dimv1 mul1n. Qed. + +Lemma baseVspace_module V (V0 := baseVspace V) : (F1 * V0 <= V0)%VS. +Proof. +apply/prodvP=> u v; rewrite [F1]unlock !mem_baseVspace => /vlineP[x ->] Vv. +by rewrite -(@scalerAl F L) mul1r; exact: memvZ. +Qed. + +Lemma sub_baseField (E : {subfield L}) : (F1 <= baseVspace E)%VS. +Proof. by rewrite -sup_field_module baseVspace_module. Qed. + +Lemma vspaceOver_refBase V : vspaceOver F1 (baseVspace V) =i V. +Proof. +move=> v; rewrite mem_vspaceOver field_module_eq ?baseVspace_module //. +by rewrite mem_baseVspace. +Qed. + +Lemma module_baseVspace M0 : + (F1 * M0 <= M0)%VS -> {V | M0 = baseVspace V & M0 =i V}. +Proof. +move=> modM0; pose V := <<vbasis (vspaceOver F1 M0) : seq L>>%VS. +suffices memM0: M0 =i V. + by exists V => //; apply/vspaceP=> v; rewrite mem_baseVspace memM0. +move=> v; rewrite -{1}(field_module_eq modM0) -(mem_vspaceOver M0) {}/V. +move: (vspaceOver F1 M0) => M. +apply/idP/idP=> [/coord_vbasis|/coord_span]->; apply/memv_suml=> i _. + rewrite /(_ *: _) /= /fieldOver_scale; case: (coord _ i _) => /= x. + rewrite {1}[F1]unlock mem_baseVspace => /vlineP[{x}x ->]. + by rewrite -(@scalerAl F L) mul1r memvZ ?memv_span ?memt_nth. +move: (coord _ i _) => x; rewrite -[_`_i]mul1r scalerAl -tnth_nth. +have F1x: x%:A \in F1. + by rewrite [F1]unlock mem_baseVspace (@memvZ F L) // mem1v. +by congr (_ \in M): (memvZ (Subvs F1x) (vbasis_mem (mem_tnth i _))). +Qed. + +Lemma module_baseAspace (E0 : {subfield L0}) : + (F1 <= E0)%VS -> {E | E0 = baseAspace E & E0 =i E}. +Proof. +rewrite -sup_field_module => /module_baseVspace[E defE0 memE0]. +suffices algE: is_aspace E by exists (ASpace algE); first exact: val_inj. +rewrite /is_aspace has_algid1 -?memE0 ?mem1v //. +by apply/prodvP=> u v; rewrite -!memE0; apply: memvM. +Qed. + +End BaseField. + +Notation baseFieldType L := (baseField_type (Phant L)). + +(* Base of fieldOver, finally. *) +Section MoreFieldOver. + +Variables (F0 : fieldType) (L : fieldExtType F0) (F : {subfield L}). + +Lemma base_vspaceOver V : baseVspace (vspaceOver F V) =i (F * V)%VS. +Proof. by move=> v; rewrite mem_baseVspace mem_vspaceOver. Qed. + +Lemma base_moduleOver V : (F * V <= V)%VS -> baseVspace (vspaceOver F V) =i V. +Proof. by move=> /field_module_eq defV v; rewrite base_vspaceOver defV. Qed. + +Lemma base_aspaceOver (E : {subfield L}) : + (F <= E)%VS -> baseVspace (vspaceOver F E) =i E. +Proof. by rewrite -sup_field_module; apply: base_moduleOver. Qed. + +End MoreFieldOver. + +Section SubFieldExtension. + +Local Open Scope quotient_scope. + +Variables (F L : fieldType) (iota : {rmorphism F -> L}). +Variables (z : L) (p : {poly F}). + +Local Notation "p ^iota" := (map_poly (GRing.RMorphism.apply iota) p) + (at level 2, format "p ^iota") : ring_scope. + +Let wf_p := (p != 0) && root p^iota z. +Let p0 : {poly F} := if wf_p then (lead_coef p)^-1 *: p else 'X. +Let z0 := if wf_p then z else 0. +Let n := (size p0).-1. + +Let p0_mon : p0 \is monic. +Proof. +rewrite /p0; case: ifP => [/andP[nz_p _] | _]; last exact: monicX. +by rewrite monicE lead_coefZ mulVf ?lead_coef_eq0. +Qed. + +Let nz_p0 : p0 != 0. Proof. by rewrite monic_neq0 // p0_mon. Qed. + +Let p0z0 : root p0^iota z0. +Proof. +rewrite /p0 /z0; case: ifP => [/andP[_ pz0]|]; last by rewrite map_polyX rootX. +by rewrite map_polyZ rootE hornerZ (rootP pz0) mulr0. +Qed. + +Let n_gt0: 0 < n. +Proof. +rewrite /n -subn1 subn_gt0 -(size_map_poly iota). +by rewrite (root_size_gt1 _ p0z0) ?map_poly_eq0. +Qed. + +Let z0Ciota : commr_rmorph iota z0. Proof. by move=> x; apply: mulrC. Qed. +Local Notation iotaPz := (horner_morph z0Ciota). +Let iotaFz (x : 'rV[F]_n) := iotaPz (rVpoly x). + +Definition equiv_subfext x y := (iotaFz x == iotaFz y). + +Fact equiv_subfext_is_equiv : equiv_class_of equiv_subfext. +Proof. by rewrite /equiv_subfext; split=> x // y w /eqP->. Qed. + +Canonical equiv_subfext_equiv := EquivRelPack equiv_subfext_is_equiv. +Canonical equiv_subfext_encModRel := defaultEncModRel equiv_subfext. + +Definition subFExtend := {eq_quot equiv_subfext}. +Canonical subFExtend_eqType := [eqType of subFExtend]. +Canonical subFExtend_choiceType := [choiceType of subFExtend]. +Canonical subFExtend_quotType := [quotType of subFExtend]. +Canonical subFExtend_eqQuotType := [eqQuotType equiv_subfext of subFExtend]. + +Definition subfx_inj := lift_fun1 subFExtend iotaFz. + +Fact pi_subfx_inj : {mono \pi : x / iotaFz x >-> subfx_inj x}. +Proof. +unlock subfx_inj => x; apply/eqP; rewrite -/(equiv_subfext _ x). +by rewrite -eqmodE reprK. +Qed. +Canonical pi_subfx_inj_morph := PiMono1 pi_subfx_inj. + +Let iotaPz_repr x : iotaPz (rVpoly (repr (\pi_(subFExtend) x))) = iotaFz x. +Proof. by rewrite -/(iotaFz _) -!pi_subfx_inj reprK. Qed. + +Definition subfext0 := lift_cst subFExtend 0. +Canonical subfext0_morph := PiConst subfext0. + +Definition subfext_add := lift_op2 subFExtend +%R. +Fact pi_subfext_add : {morph \pi : x y / x + y >-> subfext_add x y}. +Proof. +unlock subfext_add => x y /=; apply/eqmodP/eqP. +by rewrite /iotaFz !linearD /= !iotaPz_repr. +Qed. +Canonical pi_subfx_add_morph := PiMorph2 pi_subfext_add. + +Definition subfext_opp := lift_op1 subFExtend -%R. +Fact pi_subfext_opp : {morph \pi : x / - x >-> subfext_opp x}. +Proof. +unlock subfext_opp => y /=; apply/eqmodP/eqP. +by rewrite /iotaFz !linearN /= !iotaPz_repr. +Qed. +Canonical pi_subfext_opp_morph := PiMorph1 pi_subfext_opp. + +Fact addfxA : associative subfext_add. +Proof. by move=> x y t; rewrite -[x]reprK -[y]reprK -[t]reprK !piE addrA. Qed. + +Fact addfxC : commutative subfext_add. +Proof. by move=> x y; rewrite -[x]reprK -[y]reprK !piE addrC. Qed. + +Fact add0fx : left_id subfext0 subfext_add. +Proof. by move=> x; rewrite -[x]reprK !piE add0r. Qed. + +Fact addfxN : left_inverse subfext0 subfext_opp subfext_add. +Proof. by move=> x; rewrite -[x]reprK !piE addNr. Qed. + +Definition subfext_zmodMixin := ZmodMixin addfxA addfxC add0fx addfxN. +Canonical subfext_zmodType := + Eval hnf in ZmodType subFExtend subfext_zmodMixin. + +Let poly_rV_modp_K q : rVpoly (poly_rV (q %% p0) : 'rV[F]_n) = q %% p0. +Proof. by apply: poly_rV_K; rewrite -ltnS -polySpred // ltn_modp. Qed. + +Let iotaPz_modp q : iotaPz (q %% p0) = iotaPz q. +Proof. +rewrite {2}(divp_eq q p0) rmorphD rmorphM /=. +by rewrite [iotaPz p0](rootP p0z0) mulr0 add0r. +Qed. + +Definition subfx_mul_rep (x y : 'rV[F]_n) : 'rV[F]_n := + poly_rV ((rVpoly x) * (rVpoly y) %% p0). + +Definition subfext_mul := lift_op2 subFExtend subfx_mul_rep. +Fact pi_subfext_mul : + {morph \pi : x y / subfx_mul_rep x y >-> subfext_mul x y}. +Proof. +unlock subfext_mul => x y /=; apply/eqmodP/eqP. +by rewrite /iotaFz !poly_rV_modp_K !iotaPz_modp !rmorphM /= !iotaPz_repr. +Qed. +Canonical pi_subfext_mul_morph := PiMorph2 pi_subfext_mul. + +Definition subfext1 := lift_cst subFExtend (poly_rV 1). +Canonical subfext1_morph := PiConst subfext1. + +Fact mulfxA : associative (subfext_mul). +Proof. +elim/quotW=> x; elim/quotW=> y; elim/quotW=> w; rewrite !piE /subfx_mul_rep. +by rewrite !poly_rV_modp_K [_ %% p0 * _]mulrC !modp_mul // mulrA mulrC. +Qed. + +Fact mulfxC : commutative subfext_mul. +Proof. +by elim/quotW=> x; elim/quotW=> y; rewrite !piE /subfx_mul_rep /= mulrC. +Qed. + +Fact mul1fx : left_id subfext1 subfext_mul. +Proof. +elim/quotW=> x; rewrite !piE /subfx_mul_rep poly_rV_K ?size_poly1 // mul1r. +by rewrite modp_small ?rVpolyK // (polySpred nz_p0) ltnS size_poly. +Qed. + +Fact mulfx_addl : left_distributive subfext_mul subfext_add. +Proof. +elim/quotW=> x; elim/quotW=> y; elim/quotW=> w; rewrite !piE /subfx_mul_rep. +by rewrite linearD /= mulrDl modp_add linearD. +Qed. + +Fact nonzero1fx : subfext1 != subfext0. +Proof. +rewrite !piE /equiv_subfext /iotaFz !linear0. +by rewrite poly_rV_K ?rmorph1 ?oner_eq0 // size_poly1. +Qed. + +Definition subfext_comRingMixin := + ComRingMixin mulfxA mulfxC mul1fx mulfx_addl nonzero1fx. +Canonical subfext_Ring := Eval hnf in RingType subFExtend subfext_comRingMixin. +Canonical subfext_comRing := Eval hnf in ComRingType subFExtend mulfxC. + +Definition subfx_poly_inv (q : {poly F}) : {poly F} := + if iotaPz q == 0 then 0 else + let r := gdcop q p0 in let: (u, v) := egcdp q r in + ((u * q + v * r)`_0)^-1 *: u. + +Let subfx_poly_invE q : iotaPz (subfx_poly_inv q) = (iotaPz q)^-1. +Proof. +rewrite /subfx_poly_inv. +have [-> | nzq] := altP eqP; first by rewrite rmorph0 invr0. +rewrite [nth]lock -[_^-1]mul1r; apply: canRL (mulfK nzq) _; rewrite -rmorphM /=. +have rz0: iotaPz (gdcop q p0) = 0. + by apply/rootP; rewrite gdcop_map root_gdco ?map_poly_eq0 // p0z0 nzq. +do [case: gdcopP => r _; rewrite (negPf nz_p0) orbF => co_r_q _] in rz0 *. +case: (egcdp q r) (egcdpE q r) => u v /=/eqp_size/esym/eqP. +rewrite coprimep_size_gcd 1?coprimep_sym // => /size_poly1P[a nz_a Da]. +rewrite Da -scalerAl (canRL (addrK _) Da) -lock coefC linearZ linearB /=. +by rewrite rmorphM /= rz0 mulr0 subr0 horner_morphC -rmorphM mulVf ?rmorph1. +Qed. + +Definition subfx_inv_rep (x : 'rV[F]_n) : 'rV[F]_n := + poly_rV (subfx_poly_inv (rVpoly x) %% p0). + +Definition subfext_inv := lift_op1 subFExtend subfx_inv_rep. +Fact pi_subfext_inv : {morph \pi : x / subfx_inv_rep x >-> subfext_inv x}. +Proof. +unlock subfext_inv => x /=; apply/eqmodP/eqP; rewrite /iotaFz. +by rewrite 2!{1}poly_rV_modp_K 2!{1}iotaPz_modp !subfx_poly_invE iotaPz_repr. +Qed. +Canonical pi_subfext_inv_morph := PiMorph1 pi_subfext_inv. + +Fact subfx_fieldAxiom : + GRing.Field.axiom (subfext_inv : subFExtend -> subFExtend). +Proof. +elim/quotW=> x; apply: contraNeq; rewrite !piE /equiv_subfext /iotaFz !linear0. +apply: contraR => nz_x; rewrite poly_rV_K ?size_poly1 // !poly_rV_modp_K. +by rewrite iotaPz_modp rmorph1 rmorphM /= iotaPz_modp subfx_poly_invE mulVf. +Qed. + +Fact subfx_inv0 : subfext_inv (0 : subFExtend) = (0 : subFExtend). +Proof. +apply/eqP; rewrite !piE /equiv_subfext /iotaFz /subfx_inv_rep !linear0. +by rewrite /subfx_poly_inv rmorph0 eqxx mod0p !linear0. +Qed. + +Definition subfext_unitRingMixin := FieldUnitMixin subfx_fieldAxiom subfx_inv0. +Canonical subfext_unitRingType := + Eval hnf in UnitRingType subFExtend subfext_unitRingMixin. +Canonical subfext_comUnitRing := Eval hnf in [comUnitRingType of subFExtend]. +Definition subfext_fieldMixin := @FieldMixin _ _ subfx_fieldAxiom subfx_inv0. +Definition subfext_idomainMixin := FieldIdomainMixin subfext_fieldMixin. +Canonical subfext_idomainType := + Eval hnf in IdomainType subFExtend subfext_idomainMixin. +Canonical subfext_fieldType := + Eval hnf in FieldType subFExtend subfext_fieldMixin. + +Fact subfx_inj_is_rmorphism : rmorphism subfx_inj. +Proof. +do 2?split; last by rewrite piE /iotaFz poly_rV_K ?rmorph1 ?size_poly1. + by elim/quotW=> x; elim/quotW=> y; rewrite !piE /iotaFz linearB rmorphB. +elim/quotW=> x; elim/quotW=> y; rewrite !piE /subfx_mul_rep /iotaFz. +by rewrite poly_rV_modp_K iotaPz_modp rmorphM. +Qed. +Canonical subfx_inj_additive := Additive subfx_inj_is_rmorphism. +Canonical subfx_inj_rmorphism := RMorphism subfx_inj_is_rmorphism. + +Definition subfx_eval := lift_embed subFExtend (fun q => poly_rV (q %% p0)). +Canonical subfx_eval_morph := PiEmbed subfx_eval. + +Definition subfx_root := subfx_eval 'X. + +Lemma subfx_eval_is_rmorphism : rmorphism subfx_eval. +Proof. +do 2?split=> [x y|] /=; apply/eqP; rewrite piE. +- by rewrite -linearB modp_add modNp. +- by rewrite /subfx_mul_rep !poly_rV_modp_K !(modp_mul, mulrC _ y). +by rewrite modp_small // size_poly1 -subn_gt0 subn1. +Qed. +Canonical subfx_eval_additive := Additive subfx_eval_is_rmorphism. +Canonical subfx_eval_rmorphism := AddRMorphism subfx_eval_is_rmorphism. + +Definition inj_subfx := (subfx_eval \o polyC). +Canonical inj_subfx_addidive := [additive of inj_subfx]. +Canonical inj_subfx_rmorphism := [rmorphism of inj_subfx]. + +Lemma subfxE x: exists p, x = subfx_eval p. +Proof. +elim/quotW: x => x; exists (rVpoly x); apply/eqP; rewrite piE /equiv_subfext. +by rewrite /iotaFz poly_rV_modp_K iotaPz_modp. +Qed. + +Definition subfx_scale a x := inj_subfx a * x. +Fact subfx_scalerA a b x : + subfx_scale a (subfx_scale b x) = subfx_scale (a * b) x. +Proof. by rewrite /subfx_scale rmorphM mulrA. Qed. +Fact subfx_scaler1r : left_id 1 subfx_scale. +Proof. by move=> x; rewrite /subfx_scale rmorph1 mul1r. Qed. +Fact subfx_scalerDr : right_distributive subfx_scale +%R. +Proof. by move=> a; exact: mulrDr. Qed. +Fact subfx_scalerDl x : {morph subfx_scale^~ x : a b / a + b}. +Proof. by move=> a b; rewrite /subfx_scale rmorphD mulrDl. Qed. +Definition subfx_lmodMixin := + LmodMixin subfx_scalerA subfx_scaler1r subfx_scalerDr subfx_scalerDl. +Canonical subfx_lmodType := LmodType F subFExtend subfx_lmodMixin. + +Fact subfx_scaleAl : GRing.Lalgebra.axiom ( *%R : subFExtend -> _). +Proof. by move=> a; apply: mulrA. Qed. +Canonical subfx_lalgType := LalgType F subFExtend subfx_scaleAl. + +Fact subfx_scaleAr : GRing.Algebra.axiom subfx_lalgType. +Proof. by move=> a; apply: mulrCA. Qed. +Canonical subfx_algType := AlgType F subFExtend subfx_scaleAr. +Canonical subfext_unitAlgType := [unitAlgType F of subFExtend]. + +Fact subfx_evalZ : scalable subfx_eval. +Proof. by move=> a q; rewrite -mul_polyC rmorphM. Qed. +Canonical subfx_eval_linear := AddLinear subfx_evalZ. +Canonical subfx_eval_lrmorphism := [lrmorphism of subfx_eval]. + +Hypothesis (pz0 : root p^iota z). + +Section NonZero. + +Hypothesis nz_p : p != 0. + +Lemma subfx_inj_eval q : subfx_inj (subfx_eval q) = q^iota.[z]. +Proof. +by rewrite piE /iotaFz poly_rV_modp_K iotaPz_modp /iotaPz /z0 /wf_p nz_p pz0. +Qed. + +Lemma subfx_inj_root : subfx_inj subfx_root = z. +Proof. by rewrite subfx_inj_eval // map_polyX hornerX. Qed. + +Lemma subfx_injZ b x : subfx_inj (b *: x) = iota b * subfx_inj x. +Proof. by rewrite rmorphM /= subfx_inj_eval // map_polyC hornerC. Qed. + +Lemma subfx_inj_base b : subfx_inj b%:A = iota b. +Proof. by rewrite subfx_injZ rmorph1 mulr1. Qed. + +Lemma subfxEroot x : {q | x = (map_poly (in_alg subFExtend) q).[subfx_root]}. +Proof. +have /sig_eqW[q ->] := subfxE x; exists q. +apply: (fmorph_inj subfx_inj_rmorphism). +rewrite -horner_map /= subfx_inj_root subfx_inj_eval //. +by rewrite -map_poly_comp (eq_map_poly subfx_inj_base). +Qed. + +Lemma subfx_irreducibleP : + (forall q, root q^iota z -> q != 0 -> size p <= size q) <-> irreducible_poly p. +Proof. +split=> [min_p | irr_p q qz0 nz_q]. + split=> [|q nonC_q q_dv_p]. + by rewrite -(size_map_poly iota) (root_size_gt1 _ pz0) ?map_poly_eq0. + have /dvdpP[r Dp] := q_dv_p; rewrite -dvdp_size_eqp // eqn_leq dvdp_leq //=. + have [nz_r nz_q]: r != 0 /\ q != 0 by apply/norP; rewrite -mulf_eq0 -Dp. + have: root r^iota z || root q^iota z by rewrite -rootM -rmorphM -Dp. + case/orP=> /min_p; [case/(_ _)/idPn=> // | exact]. + rewrite polySpred // -leqNgt Dp size_mul //= polySpred // -subn2 ltn_subRL. + by rewrite addSnnS addnC ltn_add2l ltn_neqAle eq_sym nonC_q size_poly_gt0. +pose r := gcdp p q; have nz_r: r != 0 by rewrite gcdp_eq0 (negPf nz_p). +suffices /eqp_size <-: r %= p by rewrite dvdp_leq ?dvdp_gcdr. +rewrite (irr_p _) ?dvdp_gcdl // -(size_map_poly iota) gtn_eqF //. +by rewrite (@root_size_gt1 _ z) ?map_poly_eq0 // gcdp_map root_gcd pz0. +Qed. + +End NonZero. + +Section Irreducible. + +Hypothesis irr_p : irreducible_poly p. +Let nz_p : p != 0. Proof. exact: irredp_neq0. Qed. + +(* The Vector axiom requires irreducibility. *) +Lemma min_subfx_vectAxiom : Vector.axiom (size p).-1 subfx_lmodType. +Proof. +move/subfx_irreducibleP: irr_p => /=/(_ nz_p) min_p; set d := (size p).-1. +have Dd: d.+1 = size p by rewrite polySpred. +pose Fz2v x : 'rV_d := poly_rV (sval (sig_eqW (subfxE x)) %% p). +pose vFz : 'rV_d -> subFExtend := subfx_eval \o @rVpoly F d. +have FLinj: injective subfx_inj by apply: fmorph_inj. +have Fz2vK: cancel Fz2v vFz. + move=> x; rewrite /vFz /Fz2v; case: (sig_eqW _) => /= q ->. + apply: FLinj; rewrite !subfx_inj_eval // {2}(divp_eq q p) rmorphD rmorphM /=. + by rewrite !hornerE (eqP pz0) mulr0 add0r poly_rV_K // -ltnS Dd ltn_modpN0. +suffices vFzK: cancel vFz Fz2v. + by exists Fz2v; [apply: can2_linear Fz2vK | exists vFz]. +apply: inj_can_sym Fz2vK _ => v1 v2 /(congr1 subfx_inj)/eqP. +rewrite -subr_eq0 -!raddfB /= subfx_inj_eval // => /min_p/implyP. +rewrite leqNgt implybNN -Dd ltnS size_poly linearB subr_eq0 /=. +by move/eqP/(can_inj (@rVpolyK _ _)). +Qed. + +Definition SubfxVectMixin := VectMixin min_subfx_vectAxiom. +Definition SubfxVectType := VectType F subFExtend SubfxVectMixin. +Definition SubfxFalgType := Eval simpl in [FalgType F of SubfxVectType]. +Definition SubFieldExtType := Eval simpl in [fieldExtType F of SubfxFalgType]. + +End Irreducible. + +End SubFieldExtension. + +Prenex Implicits subfx_inj. + +Lemma irredp_FAdjoin (F : fieldType) (p : {poly F}) : + irreducible_poly p -> + {L : fieldExtType F & \dim {:L} = (size p).-1 & + {z | root (map_poly (in_alg L) p) z & <<1; z>>%VS = fullv}}. +Proof. +case=> p_gt1 irr_p; set n := (size p).-1; pose vL := [vectType F of 'rV_n]. +have Dn: n.+1 = size p := ltn_predK p_gt1. +have nz_p: p != 0 by rewrite -size_poly_eq0 -Dn. +suffices [L dimL [toPF [toL toPF_K toL_K]]]: + {L : fieldExtType F & \dim {:L} = (size p).-1 + & {toPF : {linear L -> {poly F}} & {toL : {lrmorphism {poly F} -> L} | + cancel toPF toL & forall q, toPF (toL q) = q %% p}}}. +- exists L => //; pose z := toL 'X; set iota := in_alg _. + suffices q_z q: toPF (map_poly iota q).[z] = q %% p. + exists z; first by rewrite /root -(can_eq toPF_K) q_z modpp linear0. + apply/vspaceP=> x; rewrite memvf; apply/Fadjoin_polyP. + exists (map_poly iota (toPF x)). + by apply/polyOverP=> i; rewrite coef_map memvZ ?mem1v. + by apply: (can_inj toPF_K); rewrite q_z -toL_K toPF_K. + elim/poly_ind: q => [|a q IHq]. + by rewrite map_poly0 horner0 linear0 mod0p. + rewrite rmorphD rmorphM /= map_polyX map_polyC hornerMXaddC linearD /=. + rewrite linearZ /= -(rmorph1 toL) toL_K -modp_scalel alg_polyC modp_add. + congr (_ + _); rewrite -toL_K rmorphM /= -/z; congr (toPF (_ * z)). + by apply: (can_inj toPF_K); rewrite toL_K. +pose toL q : vL := poly_rV (q %% p); pose toPF (x : vL) := rVpoly x. +have toL_K q : toPF (toL q) = q %% p. + by rewrite /toPF poly_rV_K // -ltnS Dn ?ltn_modp -?Dn. +have toPF_K: cancel toPF toL. + by move=> x; rewrite /toL modp_small ?rVpolyK // -Dn ltnS size_poly. +have toPinj := can_inj toPF_K. +pose mul x y := toL (toPF x * toPF y); pose L1 := toL 1. +have L1K: toPF L1 = 1 by rewrite toL_K modp_small ?size_poly1. +have mulC: commutative mul by rewrite /mul => x y; rewrite mulrC. +have mulA: associative mul. + by move=> x y z; apply: toPinj; rewrite -!(mulC z) !toL_K !modp_mul mulrCA. +have mul1: left_id L1 mul. + by move=> x; apply: toPinj; rewrite mulC !toL_K modp_mul mulr1 -toL_K toPF_K. +have mulD: left_distributive mul +%R. + move=> x y z; apply: toPinj; rewrite /toPF raddfD /= -!/(toPF _). + by rewrite !toL_K /toPF raddfD mulrDl modp_add. +have nzL1: L1 != 0 by rewrite -(inj_eq toPinj) L1K /toPF raddf0 oner_eq0. +pose mulM := ComRingMixin mulA mulC mul1 mulD nzL1. +pose rL := ComRingType (RingType vL mulM) mulC. +have mulZl: GRing.Lalgebra.axiom mul. + move=> a x y; apply: toPinj; rewrite toL_K /toPF !linearZ /= -!/(toPF _). + by rewrite toL_K -scalerAl modp_scalel. +have mulZr: GRing.Algebra.axiom (LalgType F rL mulZl). + by move=> a x y; rewrite !(mulrC x) scalerAl. +pose aL := AlgType F _ mulZr; pose urL := FalgUnitRingType aL. +pose uaL := [unitAlgType F of AlgType F urL mulZr]. +pose faL := [FalgType F of uaL]. +have unitE: GRing.Field.mixin_of urL. + move=> x nz_x; apply/unitrP; set q := toPF x. + have nz_q: q != 0 by rewrite -(inj_eq toPinj) /toPF raddf0 in nz_x. + have /Bezout_eq1_coprimepP[u upq1]: coprimep p q. + apply: contraLR (leq_gcdpr p nz_q) => /irr_p/implyP. + rewrite dvdp_gcdl -ltnNge /= => /eqp_size->. + by rewrite (polySpred nz_p) ltnS size_poly. + suffices: x * toL u.2 = 1 by exists (toL u.2); rewrite mulrC. + apply: toPinj; rewrite !toL_K -upq1 modp_mul modp_add mulrC. + by rewrite modp_mull add0r. +pose ucrL := [comUnitRingType of ComRingType urL mulC]. +have mul0 := GRing.Field.IdomainMixin unitE. +pose fL := FieldType (IdomainType ucrL mul0) unitE. +exists [fieldExtType F of faL for fL]; first by rewrite dimvf; apply: mul1n. +exists [linear of toPF as @rVpoly _ _]. +suffices toLM: lrmorphism (toL : {poly F} -> aL) by exists (LRMorphism toLM). +have toLlin: linear toL. + by move=> a q1 q2; rewrite -linearP -modp_scalel -modp_add. +do ?split; try exact: toLlin; move=> q r /=. +by apply: toPinj; rewrite !toL_K modp_mul -!(mulrC r) modp_mul. +Qed. + +(*Coq 8.3 processes this shorter proof correctly, but then crashes on Qed. +Lemma Xirredp_FAdjoin' (F : fieldType) (p : {poly F}) : + irreducible_poly p -> + {L : fieldExtType F & Vector.dim L = (size p).-1 & + {z | root (map_poly (in_alg L) p) z & <<1; z>>%VS = fullv}}. +Proof. +case=> p_gt1 irr_p; set n := (size p).-1; pose vL := [vectType F of 'rV_n]. +have Dn: n.+1 = size p := ltn_predK p_gt1. +have nz_p: p != 0 by rewrite -size_poly_eq0 -Dn. +pose toL q : vL := poly_rV (q %% p). +have toL_K q : rVpoly (toL q) = q %% p. + by rewrite poly_rV_K // -ltnS Dn ?ltn_modp -?Dn. +pose mul (x y : vL) : vL := toL (rVpoly x * rVpoly y). +pose L1 : vL := poly_rV 1. +have L1K: rVpoly L1 = 1 by rewrite poly_rV_K // size_poly1 -ltnS Dn. +have mulC: commutative mul by rewrite /mul => x y; rewrite mulrC. +have mulA: associative mul. + by move=> x y z; rewrite -!(mulC z) /mul !toL_K /toL !modp_mul mulrCA. +have mul1: left_id L1 mul. + move=> x; rewrite /mul L1K mul1r /toL modp_small ?rVpolyK // -Dn ltnS. + by rewrite size_poly. +have mulD: left_distributive mul +%R. + move=> x y z; apply: canLR (@rVpolyK _ _) _. + by rewrite !raddfD mulrDl /= !toL_K /toL modp_add. +have nzL1: L1 != 0 by rewrite -(can_eq (@rVpolyK _ _)) L1K raddf0 oner_eq0. +pose mulM := ComRingMixin mulA mulC mul1 mulD nzL1. +pose rL := ComRingType (RingType vL mulM) mulC. +have mulZl: GRing.Lalgebra.axiom mul. + move=> a x y; apply: canRL (@rVpolyK _ _) _; rewrite !linearZ /= toL_K. + by rewrite -scalerAl modp_scalel. +have mulZr: @GRing.Algebra.axiom _ (LalgType F rL mulZl). + by move=> a x y; rewrite !(mulrC x) scalerAl. +pose aL := AlgType F _ mulZr; pose urL := FalgUnitRingType aL. +pose uaL := [unitAlgType F of AlgType F urL mulZr]. +pose faL := [FalgType F of uaL]. +have unitE: GRing.Field.mixin_of urL. + move=> x nz_x; apply/unitrP; set q := rVpoly x. + have nz_q: q != 0 by rewrite -(can_eq (@rVpolyK _ _)) raddf0 in nz_x. + have /Bezout_eq1_coprimepP[u upq1]: coprimep p q. + have /contraR := irr_p _ _ (dvdp_gcdl p q); apply. + have: size (gcdp p q) <= size q by exact: leq_gcdpr. + rewrite leqNgt;apply:contra;move/eqp_size ->. + by rewrite (polySpred nz_p) ltnS size_poly. + suffices: x * toL u.2 = 1 by exists (toL u.2); rewrite mulrC. + congr (poly_rV _); rewrite toL_K modp_mul mulrC (canRL (addKr _) upq1). + by rewrite -mulNr modp_addl_mul_small ?size_poly1. +pose ucrL := [comUnitRingType of ComRingType urL mulC]. +pose fL := FieldType (IdomainType ucrL (GRing.Field.IdomainMixin unitE)) unitE. +exists [fieldExtType F of faL for fL]; first exact: mul1n. +pose z : vL := toL 'X; set iota := in_alg _. +have q_z q: rVpoly (map_poly iota q).[z] = q %% p. + elim/poly_ind: q => [|a q IHq]. + by rewrite map_poly0 horner0 linear0 mod0p. + rewrite rmorphD rmorphM /= map_polyX map_polyC hornerMXaddC linearD /=. + rewrite linearZ /= L1K alg_polyC modp_add; congr (_ + _); last first. + by rewrite modp_small // size_polyC; case: (~~ _) => //; apply: ltnW. + by rewrite !toL_K IHq mulrC modp_mul mulrC modp_mul. +exists z; first by rewrite /root -(can_eq (@rVpolyK _ _)) q_z modpp linear0. +apply/vspaceP=> x; rewrite memvf; apply/Fadjoin_polyP. +exists (map_poly iota (rVpoly x)). + by apply/polyOverP=> i; rewrite coef_map memvZ ?mem1v. +apply: (can_inj (@rVpolyK _ _)). +by rewrite q_z modp_small // -Dn ltnS size_poly. +Qed. +*) diff --git a/mathcomp/field/finfield.v b/mathcomp/field/finfield.v new file mode 100644 index 0000000..14a02ef --- /dev/null +++ b/mathcomp/field/finfield.v @@ -0,0 +1,585 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype div. +Require Import tuple bigop prime finset fingroup ssralg poly polydiv. +Require Import morphism action finalg zmodp cyclic center pgroup abelian. +Require Import matrix mxabelem vector falgebra fieldext separable galois. +Require ssrnum ssrint algC cyclotomic. + +(******************************************************************************) +(* Additional constructions and results on finite fields. *) +(* *) +(* FinFieldExtType L == A FinFieldType structure on the carrier of L, *) +(* where L IS a fieldExtType F structure for an *) +(* F that has a finFieldType structure. This *) +(* does not take any existing finType structure *) +(* on L; this should not be made canonical. *) +(* FinSplittingFieldType F L == A SplittingFieldType F structure on the *) +(* carrier of L, where L IS a fieldExtType F for *) +(* an F with a finFieldType structure; this *) +(* should not be made canonical. *) +(* Import FinVector :: Declares canonical default finType, finRing, *) +(* etc structures (including FinFieldExtType *) +(* above) for abstract vectType, FalgType and *) +(* fieldExtType over a finFieldType. This should *) +(* be used with caution (e.g., local to a proof) *) +(* as the finType so obtained may clash with the *) +(* canonical one for standard types like matrix. *) +(* PrimeCharType charRp == The carrier of a ringType R such that *) +(* charRp : p \in [char R] holds. This type has *) +(* canonical ringType, ..., fieldType structures *) +(* compatible with those of R, as well as *) +(* canonical lmodType 'F_p, ..., algType 'F_p *) +(* structures, plus an FalgType structure if R *) +(* is a finUnitRingType and a splittingFieldType *) +(* struture if R is a finFieldType. *) +(* FinDomainFieldType domR == A finFieldType structure on a finUnitRingType *) +(* R, given domR : GRing.IntegralDomain.axiom R. *) +(* This is intended to be used inside proofs, *) +(* where one cannot declare Canonical instances. *) +(* Otherwise one should construct explicitly the *) +(* intermediate structures using the ssralg and *) +(* finalg constructors, and finDomain_mulrC domR *) +(* finDomain_fieldP domR to prove commutativity *) +(* and field axioms (the former is Wedderburn's *) +(* little theorem). *) +(* FinDomainSplittingFieldType domR charRp == A splittingFieldType structure *) +(* that repackages the two constructions above. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory FinRing.Theory. +Local Open Scope ring_scope. + +Section FinRing. + +Variable R : finRingType. + +(* GG: Coq v8.3 fails to unify FinGroup.arg_sort _ with FinRing.sort R here *) +(* because it expands the latter rather than FinGroup.arg_sort, which would *) +(* expose the FinGroup.sort projection, thereby enabling canonical structure *) +(* expansion. We should check whether the improved heuristics in Coq 8.4 have *) +(* resolved this issue. *) +Lemma finRing_nontrivial : [set: R] != 1%g. +Proof. by apply/(@trivgPn R); exists 1; rewrite ?inE ?oner_neq0. Qed. + +(* GG: same issue here. *) +Lemma finRing_gt1 : 1 < #|R|. +Proof. by rewrite -cardsT (@cardG_gt1 R) finRing_nontrivial. Qed. + +End FinRing. + +Section FinField. + +Variable F : finFieldType. + +Lemma card_finField_unit : #|[set: {unit F}]| = #|F|.-1. +Proof. +by rewrite -(cardC1 0) cardsT card_sub; apply: eq_card => x; rewrite unitfE. +Qed. + +Fact finField_unit_subproof x : x != 0 :> F -> x \is a GRing.unit. +Proof. by rewrite unitfE. Qed. + +Definition finField_unit x nz_x := + FinRing.unit F (@finField_unit_subproof x nz_x). + +Lemma expf_card x : x ^+ #|F| = x :> F. +Proof. +apply/eqP; rewrite -{2}[x]mulr1 -[#|F|]prednK; last by rewrite (cardD1 x). +rewrite exprS -subr_eq0 -mulrBr mulf_eq0 -implyNb -unitfE subr_eq0. +apply/implyP=> Ux; rewrite -(val_unitX _ (FinRing.unit F Ux)) -val_unit1. +by rewrite val_eqE -order_dvdn -card_finField_unit order_dvdG ?inE. +Qed. + +Lemma finField_genPoly : 'X^#|F| - 'X = \prod_x ('X - x%:P) :> {poly F}. +Proof. +set n := #|F|; set Xn := 'X^n; set NX := - 'X; set pF := Xn + NX. +have lnNXn: size NX <= n by rewrite size_opp size_polyX finRing_gt1. +have UeF: uniq_roots (enum F) by rewrite uniq_rootsE enum_uniq. +rewrite [pF](all_roots_prod_XsubC _ _ UeF) ?size_addl ?size_polyXn -?cardE //. + by rewrite enumT lead_coefDl ?size_polyXn // (monicP (monicXn _ _)) scale1r. +by apply/allP=> x _; rewrite rootE !hornerE hornerXn expf_card subrr. +Qed. + +Lemma finCharP : {p | prime p & p \in [char F]}. +Proof. +pose e := exponent [set: F]; have e_gt0: e > 0 := exponent_gt0 _. +have: e%:R == 0 :> F by rewrite -zmodXgE expg_exponent // inE. +by case/natf0_char/sigW=> // p charFp; exists p; rewrite ?(charf_prime charFp). +Qed. + +Lemma finField_is_abelem : is_abelem [set: F]. +Proof. +have [p pr_p charFp] := finCharP. +apply/is_abelemP; exists p; rewrite ?abelemE ?zmod_abelian //=. +by apply/exponentP=> x _; rewrite zmodXgE mulrn_char. +Qed. + +Lemma card_finCharP p n : #|F| = (p ^ n)%N -> prime p -> p \in [char F]. +Proof. +move=> oF pr_p; rewrite inE pr_p -order_dvdn. +rewrite (abelem_order_p finField_is_abelem) ?inE ?oner_neq0 //=. +have n_gt0: n > 0 by rewrite -(ltn_exp2l _ _ (prime_gt1 pr_p)) -oF finRing_gt1. +by rewrite cardsT oF -(prednK n_gt0) pdiv_pfactor. +Qed. + +End FinField. + +Section CardVspace. + +Variables (F : finFieldType) (T : finType). + +Section Vector. + +Variable cvT : Vector.class_of F T. +Let vT := Vector.Pack (Phant F) cvT T. + +Lemma card_vspace (V : {vspace vT}) : #|V| = (#|F| ^ \dim V)%N. +Proof. +set n := \dim V; pose V2rV v := \row_i coord (vbasis V) i v. +pose rV2V (rv : 'rV_n) := \sum_i rv 0 i *: (vbasis V)`_i. +have rV2V_K: cancel rV2V V2rV. + have freeV: free (vbasis V) := basis_free (vbasisP V). + by move=> rv; apply/rowP=> i; rewrite mxE coord_sum_free. +rewrite -[n]mul1n -card_matrix -(card_imset _ (can_inj rV2V_K)). +apply: eq_card => v; apply/idP/imsetP=> [/coord_vbasis-> | [rv _ ->]]. + by exists (V2rV v) => //; apply: eq_bigr => i _; rewrite mxE. +by apply: (@rpred_sum vT) => i _; rewrite rpredZ ?vbasis_mem ?memt_nth. +Qed. + +Lemma card_vspacef : #|{: vT}%VS| = #|T|. +Proof. by apply: eq_card => v; rewrite (@memvf _ vT). Qed. + +End Vector. + +Variable caT : Falgebra.class_of F T. +Let aT := Falgebra.Pack (Phant F) caT T. + +Lemma card_vspace1 : #|(1%VS : {vspace aT})| = #|F|. +Proof. by rewrite card_vspace (dimv1 aT). Qed. + +End CardVspace. + +Lemma VectFinMixin (R : finRingType) (vT : vectType R) : Finite.mixin_of vT. +Proof. +have v2rK := @Vector.InternalTheory.v2rK R vT. +exact: CanFinMixin (v2rK : @cancel _ (CountType vT (CanCountMixin v2rK)) _ _). +Qed. + +(* These instacnces are not exported by default because they conflict with *) +(* existing finType instances such as matrix_finType or primeChar_finType. *) +Module FinVector. +Section Interfaces. + +Variable F : finFieldType. +Implicit Types (vT : vectType F) (aT : FalgType F) (fT : fieldExtType F). + +Canonical vect_finType vT := FinType vT (VectFinMixin vT). +Canonical Falg_finType aT := FinType aT (VectFinMixin aT). +Canonical fieldExt_finType fT := FinType fT (VectFinMixin fT). + +Canonical Falg_finRingType aT := [finRingType of aT]. +Canonical fieldExt_finRingType fT := [finRingType of fT]. +Canonical fieldExt_finFieldType fT := [finFieldType of fT]. + +Lemma finField_splittingField_axiom fT : SplittingField.axiom fT. +Proof. +exists ('X^#|fT| - 'X); first by rewrite rpredB 1?rpredX ?polyOverX. +exists (enum fT); first by rewrite enumT finField_genPoly eqpxx. +by apply/vspaceP=> x; rewrite memvf seqv_sub_adjoin ?mem_enum. +Qed. + +End Interfaces. +End FinVector. + +Notation FinFieldExtType := FinVector.fieldExt_finFieldType. +Notation FinSplittingFieldAxiom := (FinVector.finField_splittingField_axiom _). +Notation FinSplittingFieldType F L := + (SplittingFieldType F L FinSplittingFieldAxiom). + +Section PrimeChar. + +Variable p : nat. + +Section PrimeCharRing. + +Variable R0 : ringType. + +Definition PrimeCharType of p \in [char R0] : predArgType := R0. + +Hypothesis charRp : p \in [char R0]. +Local Notation R := (PrimeCharType charRp). +Implicit Types (a b : 'F_p) (x y : R). + +Canonical primeChar_eqType := [eqType of R]. +Canonical primeChar_choiceType := [choiceType of R]. +Canonical primeChar_zmodType := [zmodType of R]. +Canonical primeChar_ringType := [ringType of R]. + +Definition primeChar_scale a x := a%:R * x. +Local Infix "*p:" := primeChar_scale (at level 40). + +Let natrFp n : (inZp n : 'F_p)%:R = n%:R :> R. +Proof. +rewrite {2}(divn_eq n p) natrD mulrnA (mulrn_char charRp) add0r. +by rewrite /= (Fp_cast (charf_prime charRp)). +Qed. + +Lemma primeChar_scaleA a b x : a *p: (b *p: x) = (a * b) *p: x. +Proof. by rewrite /primeChar_scale mulrA -natrM natrFp. Qed. + +Lemma primeChar_scale1 : left_id 1 primeChar_scale. +Proof. by move=> x; rewrite /primeChar_scale mul1r. Qed. + +Lemma primeChar_scaleDr : right_distributive primeChar_scale +%R. +Proof. by move=> a x y /=; rewrite /primeChar_scale mulrDr. Qed. + +Lemma primeChar_scaleDl x : {morph primeChar_scale^~ x: a b / a + b}. +Proof. by move=> a b; rewrite /primeChar_scale natrFp natrD mulrDl. Qed. + +Definition primeChar_lmodMixin := + LmodMixin primeChar_scaleA primeChar_scale1 + primeChar_scaleDr primeChar_scaleDl. +Canonical primeChar_lmodType := LmodType 'F_p R primeChar_lmodMixin. + +Lemma primeChar_scaleAl : GRing.Lalgebra.axiom ( *%R : R -> R -> R). +Proof. by move=> a x y; apply: mulrA. Qed. +Canonical primeChar_LalgType := LalgType 'F_p R primeChar_scaleAl. + +Lemma primeChar_scaleAr : GRing.Algebra.axiom primeChar_LalgType. +Proof. by move=> a x y; rewrite ![a *: _]mulr_natl mulrnAr. Qed. +Canonical primeChar_algType := AlgType 'F_p R primeChar_scaleAr. + +End PrimeCharRing. + +Local Notation type := @PrimeCharType. + +Canonical primeChar_unitRingType (R : unitRingType) charRp := + [unitRingType of type R charRp]. +Canonical primeChar_unitAlgType (R : unitRingType) charRp := + [unitAlgType 'F_p of type R charRp]. +Canonical primeChar_comRingType (R : comRingType) charRp := + [comRingType of type R charRp]. +Canonical primeChar_comUnitRingType (R : comUnitRingType) charRp := + [comUnitRingType of type R charRp]. +Canonical primeChar_idomainType (R : idomainType) charRp := + [idomainType of type R charRp]. +Canonical primeChar_fieldType (F : fieldType) charFp := + [fieldType of type F charFp]. + +Section FinRing. + +Variables (R0 : finRingType) (charRp : p \in [char R0]). +Local Notation R := (type _ charRp). + +Canonical primeChar_finType := [finType of R]. +Canonical primeChar_finZmodType := [finZmodType of R]. +Canonical primeChar_baseGroupType := [baseFinGroupType of R for +%R]. +Canonical primeChar_groupType := [finGroupType of R for +%R]. +Canonical primeChar_finRingType := [finRingType of R]. + +Let pr_p : prime p. Proof. exact: charf_prime charRp. Qed. + +Lemma primeChar_abelem : p.-abelem [set: R]. +Proof. +rewrite abelemE ?zmod_abelian //=. +by apply/exponentP=> x _; rewrite zmodXgE mulrn_char. +Qed. + +Lemma primeChar_pgroup : p.-group [set: R]. +Proof. by case/and3P: primeChar_abelem. Qed. + +Lemma order_primeChar x : x != 0 :> R -> #[x]%g = p. +Proof. by apply: (abelem_order_p primeChar_abelem); rewrite inE. Qed. + +Let n := logn p #|R|. + +Lemma card_primeChar : #|R| = (p ^ n)%N. +Proof. by rewrite /n -cardsT {1}(card_pgroup primeChar_pgroup). Qed. + +Lemma primeChar_vectAxiom : Vector.axiom n (primeChar_lmodType charRp). +Proof. +have /isog_isom/=[f /isomP[injf im_f]]: [set: R] \isog [set: 'rV['F_p]_n]. + have [abelR ntR] := (primeChar_abelem, finRing_nontrivial R0). + by rewrite /n -cardsT -(dim_abelemE abelR) ?isog_abelem_rV. +exists f; last by exists (invm injf) => x; rewrite ?invmE ?invmK ?im_f ?inE. +move=> a x y; rewrite [a *: _]mulr_natl morphM ?morphX ?inE // zmodXgE. +by congr (_ + _); rewrite -scaler_nat natr_Zp. +Qed. + +Definition primeChar_vectMixin := Vector.Mixin primeChar_vectAxiom. +Canonical primeChar_vectType := VectType 'F_p R primeChar_vectMixin. + +Lemma primeChar_dimf : \dim {:primeChar_vectType} = n. +Proof. by rewrite dimvf. Qed. + +End FinRing. + +Canonical primeChar_finUnitRingType (R : finUnitRingType) charRp := + [finUnitRingType of type R charRp]. +Canonical primeChar_finUnitAlgType (R : finUnitRingType) charRp := + [finUnitAlgType 'F_p of type R charRp]. +Canonical primeChar_FalgType (R : finUnitRingType) charRp := + [FalgType 'F_p of type R charRp]. +Canonical primeChar_finComRingType (R : finComRingType) charRp := + [finComRingType of type R charRp]. +Canonical primeChar_finComUnitRingType (R : finComUnitRingType) charRp := + [finComUnitRingType of type R charRp]. +Canonical primeChar_finIdomainType (R : finIdomainType) charRp := + [finIdomainType of type R charRp]. + +Section FinField. + +Variables (F0 : finFieldType) (charFp : p \in [char F0]). +Local Notation F := (type _ charFp). + +Canonical primeChar_finFieldType := [finFieldType of F]. +(* We need to use the eta-long version of the constructor here as projections *) +(* of the Canonical fieldType of F cannot be computed syntactically. *) +Canonical primeChar_fieldExtType := [fieldExtType 'F_p of F for F0]. +Canonical primeChar_splittingFieldType := FinSplittingFieldType 'F_p F. + +End FinField. + +End PrimeChar. + +Section FinSplittingField. + +Variable F : finFieldType. + +(* By card_vspace order K = #|K| for any finType structure on L; however we *) +(* do not want to impose the FinVector instance here. *) +Let order (L : vectType F) (K : {vspace L}) := (#|F| ^ \dim K)%N. + +Section FinGalois. + +Variable L : splittingFieldType F. +Implicit Types (a b : F) (x y : L) (K E : {subfield L}). + +Let galL K : galois K {:L}. +Proof. +without loss {K} ->: K / K = 1%AS. + by move=> IH; apply: galoisS (IH _ (erefl _)); rewrite sub1v subvf. +apply/splitting_galoisField; pose finL := FinFieldExtType L. +exists ('X^#|finL| - 'X); split; first by rewrite rpredB 1?rpredX ?polyOverX. + rewrite (finField_genPoly finL) -big_filter. + by rewrite separable_prod_XsubC ?(enum_uniq finL). +exists (enum finL); first by rewrite enumT (finField_genPoly finL) eqpxx. +by apply/vspaceP=> x; rewrite memvf seqv_sub_adjoin ?(mem_enum finL). +Qed. + +Fact galLgen K : + {alpha | generator 'Gal({:L} / K) alpha & forall x, alpha x = x ^+ order K}. +Proof. +without loss{K} ->: K / K = 1%AS; last rewrite /order dimv1 expn1. + rewrite /generator => /(_ _ (erefl _))[alpha /eqP defGalL]. + rewrite /order dimv1 expn1 => Dalpha. + exists (alpha ^+ \dim K)%g => [|x]; last first. + elim: (\dim K) => [|n IHn]; first by rewrite gal_id. + by rewrite expgSr galM ?memvf // IHn Dalpha expnSr exprM. + rewrite (eq_subG_cyclic (cycle_cyclic alpha)) ?cycleX //=; last first. + by rewrite -defGalL galS ?sub1v. + rewrite eq_sym -orderE orderXdiv orderE -defGalL -{1}(galois_dim (galL 1)). + by rewrite dimv1 divn1 galois_dim. + by rewrite dimv1 divn1 field_dimS // subvf. +pose f x := x ^+ #|F|. +have idfP x: reflect (f x = x) (x \in 1%VS). + rewrite /f; apply: (iffP (vlineP _ _)) => [[a ->] | xFx]. + by rewrite -in_algE -rmorphX expf_card. + pose q := map_poly (in_alg L) ('X^#|F| - 'X). + have: root q x. + rewrite /q rmorphB /= map_polyXn map_polyX. + by rewrite rootE !(hornerE, hornerXn) xFx subrr. + have{q} ->: q = \prod_(z <- [seq b%:A | b : F]) ('X - z%:P). + rewrite /q finField_genPoly rmorph_prod big_map enumT. + by apply: eq_bigr => b _; rewrite rmorphB /= map_polyX map_polyC. + by rewrite root_prod_XsubC => /mapP[a]; exists a. +have fM: rmorphism f. + rewrite /f; do 2?split=> [x y|]; rewrite ?exprMn ?expr1n //. + have [p _ charFp] := finCharP F; rewrite (card_primeChar charFp). + elim: (logn _ _) => // n IHn; rewrite expnSr !exprM {}IHn. + by rewrite -(char_lalg L) in charFp; rewrite -Frobenius_autE rmorphB. +have fZ: linear f. + move=> a x y; rewrite -mulr_algl [f _](rmorphD (RMorphism fM)) rmorphM /=. + by rewrite (idfP _ _) ?mulr_algl ?memvZ // memv_line. +have /kAut_to_gal[alpha galLalpha Dalpha]: kAut 1 {:L} (linfun (Linear fZ)). + rewrite kAutfE; apply/kHomP; split=> [x y _ _ | x /idfP]; rewrite !lfunE //=. + exact: (rmorphM (RMorphism fM)). +exists alpha => [|a]; last by rewrite -Dalpha ?memvf ?lfunE. +suffices <-: fixedField [set alpha] = 1%AS by rewrite gal_generated /generator. +apply/vspaceP => x; apply/fixedFieldP/idfP; rewrite ?memvf // => id_x. + by rewrite -{2}(id_x _ (set11 _)) -Dalpha ?lfunE ?memvf. +by move=> _ /set1P->; rewrite -Dalpha ?memvf ?lfunE. +Qed. + +Lemma finField_galois K E : (K <= E)%VS -> galois K E. +Proof. +move=> sKE; have /galois_fixedField <- := galL E. +rewrite normal_fixedField_galois // -sub_abelian_normal ?galS //. +apply: abelianS (galS _ (sub1v _)) _. +by have [alpha /('Gal(_ / _) =P _)-> _] := galLgen 1; apply: cycle_abelian. +Qed. + +Lemma finField_galois_generator K E : + (K <= E)%VS -> + {alpha | generator 'Gal(E / K) alpha + & {in E, forall x, alpha x = x ^+ order K}}. +Proof. +move=> sKE; have [alpha defGalLK Dalpha] := galLgen K. +have inKL_E: (K <= E <= {:L})%VS by rewrite sKE subvf. +have nKE: normalField K E by have/and3P[] := finField_galois sKE. +have galLKalpha: alpha \in 'Gal({:L} / K). + by rewrite (('Gal(_ / _) =P _) defGalLK) cycle_id. +exists (normalField_cast _ alpha) => [|x Ex]; last first. + by rewrite (normalField_cast_eq inKL_E). +rewrite /generator -(morphim_cycle (normalField_cast_morphism inKL_E nKE)) //. +by rewrite -((_ =P <[alpha]>) defGalLK) normalField_img. +Qed. + +End FinGalois. + +Lemma Fermat's_little_theorem (L : fieldExtType F) (K : {subfield L}) a : + (a \in K) = (a ^+ order K == a). +Proof. +move: K a; wlog [{L}L -> K a]: L / exists galL : splittingFieldType F, L = galL. + by pose galL := (FinSplittingFieldType F L) => /(_ galL); apply; exists galL. +have /galois_fixedField fixLK := finField_galois (subvf K). +have [alpha defGalLK Dalpha] := finField_galois_generator (subvf K). +rewrite -Dalpha ?memvf // -{1}fixLK (('Gal(_ / _) =P _) defGalLK). +rewrite /cycle -gal_generated (galois_fixedField _) ?fixedField_galois //. +by apply/fixedFieldP/eqP=> [|-> | alpha_x _ /set1P->]; rewrite ?memvf ?set11. +Qed. + +End FinSplittingField. + +Section FinDomain. + +Import ssrnum ssrint algC cyclotomic Num.Theory. +Local Infix "%|" := dvdn. (* Hide polynomial divisibility. *) + +Variable R : finUnitRingType. + +Hypothesis domR : GRing.IntegralDomain.axiom R. +Implicit Types x y : R. + +Let lregR x : x != 0 -> GRing.lreg x. +Proof. by move=> xnz; apply: mulrI0_lreg => y /domR/orP[/idPn | /eqP]. Qed. + +Lemma finDomain_field : GRing.Field.mixin_of R. +Proof. +move=> x /lregR-regx; apply/unitrP; exists (invF regx 1). +by split; first apply: (regx); rewrite ?mulrA f_invF // mulr1 mul1r. +Qed. + +(* This is Witt's proof of Wedderburn's little theorem. *) +Theorem finDomain_mulrC : @commutative R R *%R. +Proof. +have fieldR := finDomain_field. +have [p p_pr charRp]: exists2 p, prime p & p \in [char R]. + have [e /prod_prime_decomp->]: {e | (e > 0)%N & e%:R == 0 :> R}. + by exists #|[set: R]%G|; rewrite // -order_dvdn order_dvdG ?inE. + rewrite big_seq; elim/big_rec: _ => [|[p m] /= n]; first by rewrite oner_eq0. + case/mem_prime_decomp=> p_pr _ _ IHn. + elim: m => [|m IHm]; rewrite ?mul1n {IHn}// expnS -mulnA natrM. + by case/eqP/domR/orP=> //; exists p; last exact/andP. +pose Rp := PrimeCharType charRp; pose L : {vspace Rp} := fullv. +pose G := [set: {unit R}]; pose ofG : {unit R} -> Rp := val. +pose projG (E : {vspace Rp}) := [preim ofG of E]. +have inG t nzt: Sub t (finDomain_field nzt) \in G by rewrite inE. +have card_projG E: #|projG E| = (p ^ \dim E - 1)%N. + transitivity #|E|.-1; last by rewrite subn1 card_vspace card_Fp. + rewrite (cardD1 0) mem0v (card_preim val_inj) /=. + apply: eq_card => x; congr (_ && _); rewrite [LHS]codom_val. + by apply/idP/idP=> [/(memPn _ _)-> | /fieldR]; rewrite ?unitr0. +pose C u := 'C[ofG u]%AS; pose Q := 'C(L)%AS; pose q := (p ^ \dim Q)%N. +have defC u: 'C[u] =i projG (C u). + by move=> v; rewrite cent1E !inE (sameP cent1vP eqP). +have defQ: 'Z(G) =i projG Q. + move=> u; rewrite !inE. + apply/centP/centvP=> cGu v _; last exact/val_inj/cGu/memvf. + by have [-> | /inG/cGu[]] := eqVneq v 0; first by rewrite commr0. +have q_gt1: (1 < q)%N by rewrite (ltn_exp2l 0) ?prime_gt1 ?adim_gt0. +pose n := \dim_Q L; have oG: #|G| = (q ^ n - 1)%N. + rewrite -expnM mulnC divnK ?skew_field_dimS ?subvf // -card_projG. + by apply: eq_card => u; rewrite !inE memvf. +have oZ: #|'Z(G)| = (q - 1)%N by rewrite -card_projG; apply: eq_card. +suffices n_le1: (n <= 1)%N. + move=> u v; apply/centvsP: (memvf (u : Rp)) (memvf (v : Rp)) => {u v}. + rewrite -(geq_leqif (dimv_leqif_sup (subvf Q))) -/L. + by rewrite leq_divLR ?mul1n ?skew_field_dimS ?subvf in n_le1. +without loss n_gt1: / (1 < n)%N by rewrite ltnNge; apply: wlog_neg. +have [q_gt0 n_gt0] := (ltnW q_gt1, ltnW n_gt1). +have [z z_prim] := C_prim_root_exists n_gt0. +have zn1: z ^+ n = 1 by apply: prim_expr_order. +have /eqP-n1z: `|z| == 1. + by rewrite -(pexpr_eq1 n_gt0) ?normr_ge0 // -normrX zn1 normr1. +suffices /eqP/normC_sub_eq[t n1t [Dq Dz]]: `|q%:R - z| == `|q%:R| - `|z|. + suffices z1: z == 1 by rewrite leq_eqVlt -dvdn1 (prim_order_dvd z_prim) z1. + by rewrite Dz n1z mul1r -(eqr_pmuln2r q_gt0) Dq normr_nat mulr_natl. +pose aq d : algC := (cyclotomic (z ^+ (n %/ d)) d).[q%:R]. +suffices: `|aq n| <= (q - 1)%:R. + rewrite eqr_le ler_sub_dist andbT n1z normr_nat natrB //; apply: ler_trans. + rewrite {}/aq horner_prod divnn n_gt0 expr1 normr_prod. + rewrite (bigD1 (Ordinal n_gt1)) ?coprime1n //= !hornerE ler_pemulr //. + elim/big_ind: _ => // [|d _]; first exact: mulr_ege1. + rewrite !hornerE; apply: ler_trans (ler_sub_dist _ _). + by rewrite normr_nat normrX n1z expr1n ler_subr_addl (leC_nat 2). +have Zaq d: d %| n -> aq d \in Cint. + move/(dvdn_prim_root z_prim)=> zd_prim. + rewrite rpred_horner ?rpred_nat //= -Cintr_Cyclotomic //. + by apply/polyOverP=> i; rewrite coef_map ?rpred_int. +suffices: (aq n %| (q - 1)%:R)%C. + rewrite {1}[aq n]CintEsign ?Zaq // -(rpredMsign _ (aq n < 0)%R). + rewrite dvdC_mul2l ?signr_eq0 //. + have /CnatP[m ->]: `|aq n| \in Cnat by rewrite Cnat_norm_Cint ?Zaq. + by rewrite leC_nat dvdC_nat; apply: dvdn_leq; rewrite subn_gt0. +have prod_aq m: m %| n -> \prod_(d < n.+1 | d %| m) aq d = (q ^ m - 1)%:R. + move=> m_dv_n; transitivity ('X^m - 1).[q%:R : algC]; last first. + by rewrite !hornerE hornerXn -natrX natrB ?expn_gt0 ?prime_gt0. + rewrite (prod_cyclotomic (dvdn_prim_root z_prim m_dv_n)). + have def_divm: perm_eq (divisors m) [seq d <- index_iota 0 n.+1 | d %| m]. + rewrite uniq_perm_eq ?divisors_uniq ?filter_uniq ?iota_uniq // => d. + rewrite -dvdn_divisors ?(dvdn_gt0 n_gt0) // mem_filter mem_iota ltnS /=. + by apply/esym/andb_idr=> d_dv_m; rewrite dvdn_leq ?(dvdn_trans d_dv_m). + rewrite (eq_big_perm _ def_divm) big_filter big_mkord horner_prod. + by apply: eq_bigr => d d_dv_m; rewrite -exprM muln_divA ?divnK. +have /rpredBl<-: (aq n %| #|G|%:R)%C. + rewrite oG -prod_aq // (bigD1 ord_max) //= dvdC_mulr //. + by apply: rpred_prod => d /andP[/Zaq]. +rewrite center_class_formula addrC oZ natrD addKr natr_sum /=. +apply: rpred_sum => _ /imsetP[u /setDP[_ Z'u] ->]; rewrite -/G /=. +have sQC: (Q <= C u)%VS by apply/subvP=> v /centvP-cLv; apply/cent1vP/cLv/memvf. +have{sQC} /dvdnP[m Dm]: \dim Q %| \dim (C u) by apply: skew_field_dimS. +have m_dv_n: m %| n by rewrite dvdn_divRL // -?Dm ?skew_field_dimS ?subvf. +have m_gt0: (0 < m)%N := dvdn_gt0 n_gt0 m_dv_n. +have{Dm} oCu: #|'C[u]| = (q ^ m - 1)%N. + by rewrite -expnM mulnC -Dm (eq_card (defC u)) card_projG. +have ->: #|u ^: G|%:R = \prod_(d < n.+1 | d %| n) (aq d / aq d ^+ (d %| m)). + rewrite -index_cent1 natf_indexg ?subsetT //= setTI prodf_div prod_aq // -oG. + congr (_ / _); rewrite big_mkcond oCu -prod_aq //= big_mkcond /=. + by apply: eq_bigr => d _; case: ifP => [/dvdn_trans->| _]; rewrite ?if_same. +rewrite (bigD1 ord_max) //= [n %| m](contraNF _ Z'u) => [|n_dv_m]; last first. + rewrite -sub_cent1 subEproper eq_sym eqEcard subsetT oG oCu leq_sub2r //. + by rewrite leq_exp2l // dvdn_leq. +rewrite divr1 dvdC_mulr //; apply/rpred_prod => d /andP[/Zaq-Zaqd _]. +have [-> | nz_aqd] := eqVneq (aq d) 0; first by rewrite mul0r. +by rewrite -[aq d]expr1 -exprB ?leq_b1 ?unitfE ?rpredX. +Qed. + +Definition FinDomainFieldType : finFieldType := + let fin_unit_class := FinRing.UnitRing.class R in + let com_class := GRing.ComRing.Class finDomain_mulrC in + let com_unit_class := @GRing.ComUnitRing.Class R com_class fin_unit_class in + let dom_class := @GRing.IntegralDomain.Class R com_unit_class domR in + let field_class := @GRing.Field.Class R dom_class finDomain_field in + let finfield_class := @FinRing.Field.Class R field_class fin_unit_class in + FinRing.Field.Pack finfield_class R. + +Definition FinDomainSplittingFieldType p (charRp : p \in [char R]) := + let RoverFp := @primeChar_splittingFieldType p FinDomainFieldType charRp in + [splittingFieldType 'F_p of R for RoverFp]. + +End FinDomain. diff --git a/mathcomp/field/galois.v b/mathcomp/field/galois.v new file mode 100644 index 0000000..ccdd8e4 --- /dev/null +++ b/mathcomp/field/galois.v @@ -0,0 +1,1628 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div choice fintype. +Require Import tuple finfun bigop ssralg poly polydiv. +Require Import finset fingroup morphism quotient perm action zmodp cyclic. +Require Import matrix mxalgebra vector falgebra fieldext separable. + +(******************************************************************************) +(* This file develops some basic Galois field theory, defining: *) +(* splittingFieldFor K p E <-> E is the smallest field over K that splits p *) +(* into linear factors. *) +(* kHom K E f <=> f : 'End(L) is a ring morphism on E and fixes K. *) +(* kAut K E f <=> f : 'End(L) is a kHom K E and f @: E == E. *) +(* kHomExtend E f x y == a kHom K <<E; x>> that extends f and maps x to y, *) +(* when f \is a kHom K E and root (minPoly E x) y. *) +(* *) +(* splittingFieldFor K p E <-> E is splitting field for p over K: p splits in *) +(* E and its roots generate E from K. *) +(* splittingFieldType F == the interface type of splitting field extensions *) +(* of F, that is, extensions generated by all the *) +(* algebraic roots of some polynomial, or, *) +(* equivalently, normal field extensions of F. *) +(* SplittingField.axiom F L == the axiom stating that L is a splitting field. *) +(* SplittingFieldType F L FsplitL == packs a proof FsplitL of the splitting *) +(* field axiom for L into a splitingFieldType F, *) +(* provided L has a fieldExtType F structure. *) +(* [splittingFieldType F of L] == a clone of the canonical splittingFieldType *) +(* structure for L. *) +(*[splittingFieldType F of L for M] == an L-clone of the canonical *) +(* splittingFieldType structure on M. *) +(* *) +(* gal_of E == the group_type of automorphisms of E over the *) +(* base field F. *) +(* 'Gal(E / K) == the group of automorphisms of E that fix K. *) +(* fixedField s == the field fixed by the set of automorphisms s. *) +(* fixedField set0 = E when set0 : {set: gal_of E} *) +(* normalField K E <=> E is invariant for every 'Gal(L / K) for every L. *) +(* galois K E <=> E is a normal and separable field extension of K. *) +(* galTrace K E a == \sum_(f in 'Gal(E / K)) (f a). *) +(* galNorm K E a == \prod_(f in 'Gal(E / K)) (f a). *) +(* *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Reserved Notation "''Gal' ( A / B )" + (at level 8, A at level 35, format "''Gal' ( A / B )"). + +Import GroupScope GRing.Theory. +Local Open Scope ring_scope. + +Section SplittingFieldFor. + +Variables (F : fieldType) (L : fieldExtType F). + +Definition splittingFieldFor (U : {vspace L}) (p : {poly L}) (V : {vspace L}) := + exists2 rs, p %= \prod_(z <- rs) ('X - z%:P) & <<U & rs>>%VS = V. + +Lemma splittingFieldForS (K M E : {subfield L}) p : + (K <= M)%VS -> (M <= E)%VS -> + splittingFieldFor K p E -> splittingFieldFor M p E. +Proof. +move=> sKM sKE [rs Dp genL]; exists rs => //; apply/eqP. +rewrite eqEsubv -[in X in _ && (X <= _)%VS]genL adjoin_seqSl // andbT. +by apply/Fadjoin_seqP; split; rewrite // -genL; apply: seqv_sub_adjoin. +Qed. + +End SplittingFieldFor. + +Section kHom. + +Variables (F : fieldType) (L : fieldExtType F). +Implicit Types (U V : {vspace L}) (K E : {subfield L}) (f g : 'End(L)). + +Definition kHom U V f := ahom_in V f && (U <= fixedSpace f)%VS. + +Lemma kHomP {K V f} : + reflect [/\ {in V &, forall x y, f (x * y) = f x * f y} + & {in K, forall x, f x = x}] + (kHom K V f). +Proof. +apply: (iffP andP) => [[/ahom_inP[fM _] /subvP idKf] | [fM idKf]]. + by split=> // x /idKf/fixedSpaceP. +split; last by apply/subvP=> x /idKf/fixedSpaceP. +by apply/ahom_inP; split=> //; rewrite idKf ?mem1v. +Qed. + +Lemma kAHomP {U V} {f : 'AEnd(L)} : + reflect {in U, forall x, f x = x} (kHom U V f). +Proof. by rewrite /kHom ahomWin; apply: fixedSpacesP. Qed. + +Lemma kHom1 U V : kHom U V \1. +Proof. by apply/kAHomP => u _; rewrite lfunE. Qed. + +Lemma k1HomE V f : kHom 1 V f = ahom_in V f. +Proof. by apply: andb_idr => /ahom_inP[_ f1]; apply/fixedSpaceP. Qed. + +Lemma kHom_lrmorphism (f : 'End(L)) : reflect (lrmorphism f) (kHom 1 {:L} f). +Proof. by rewrite k1HomE; apply: ahomP. Qed. + +Lemma k1AHom V (f : 'AEnd(L)) : kHom 1 V f. +Proof. by rewrite k1HomE ahomWin. Qed. + +Lemma kHom_poly_id K E f p : + kHom K E f -> p \is a polyOver K -> map_poly f p = p. +Proof. +by case/kHomP=> _ idKf /polyOverP Kp; apply/polyP=> i; rewrite coef_map /= idKf. +Qed. + +Lemma kHomSl U1 U2 V f : (U1 <= U2)%VS -> kHom U2 V f -> kHom U1 V f. +Proof. by rewrite /kHom => sU12 /andP[-> /(subv_trans sU12)]. Qed. + +Lemma kHomSr K V1 V2 f : (V1 <= V2)%VS -> kHom K V2 f -> kHom K V1 f. +Proof. by move/subvP=> sV12 /kHomP[/(sub_in2 sV12)fM idKf]; apply/kHomP. Qed. + +Lemma kHomS K1 K2 V1 V2 f : + (K1 <= K2)%VS -> (V1 <= V2)%VS -> kHom K2 V2 f -> kHom K1 V1 f. +Proof. by move=> sK12 sV12 /(kHomSl sK12)/(kHomSr sV12). Qed. + +Lemma kHom_eq K E f g : + (K <= E)%VS -> {in E, f =1 g} -> kHom K E f = kHom K E g. +Proof. +move/subvP=> sKE eq_fg; wlog suffices: f g eq_fg / kHom K E f -> kHom K E g. + by move=> IH; apply/idP/idP; apply: IH => x /eq_fg. +case/kHomP=> fM idKf; apply/kHomP. +by split=> [x y Ex Ey | x Kx]; rewrite -!eq_fg ?fM ?rpredM // ?idKf ?sKE. +Qed. + +Lemma kHom_inv K E f : kHom K E f -> {in E, {morph f : x / x^-1}}. +Proof. +case/kHomP=> fM idKf x Ex. +case (eqVneq x 0) => [-> | nz_x]; first by rewrite linear0 invr0 linear0. +have fxV: f x * f x^-1 = 1 by rewrite -fM ?rpredV ?divff // idKf ?mem1v. +have Ufx: f x \is a GRing.unit by apply/unitrPr; exists (f x^-1). +by apply: (mulrI Ufx); rewrite divrr. +Qed. + +Lemma kHom_dim K E f : kHom K E f -> \dim (f @: E) = \dim E. +Proof. +move=> homKf; have [fM idKf] := kHomP homKf. +apply/limg_dim_eq/eqP; rewrite -subv0; apply/subvP=> v. +rewrite memv_cap memv0 memv_ker => /andP[Ev]; apply: contraLR => nz_v. +by rewrite -unitfE unitrE -(kHom_inv homKf) // -fM ?rpredV ?divff ?idKf ?mem1v. +Qed. + +Lemma kHom_is_rmorphism K E f : + kHom K E f -> rmorphism (f \o vsval : subvs_of E -> L). +Proof. +case/kHomP=> fM idKf; split=> [a b|]; first exact: raddfB. +by split=> [a b|] /=; [rewrite /= fM ?subvsP | rewrite algid1 idKf // mem1v]. +Qed. +Definition kHom_rmorphism K E f homKEf := + RMorphism (@kHom_is_rmorphism K E f homKEf). + +Lemma kHom_horner K E f p x : + kHom K E f -> p \is a polyOver E -> x \in E -> f p.[x] = (map_poly f p).[f x]. +Proof. +move=> homKf /polyOver_subvs[{p}p -> Ex]; pose fRM := kHom_rmorphism homKf. +by rewrite (horner_map _ _ (Subvs Ex)) -[f _](horner_map fRM) map_poly_comp. +Qed. + +Lemma kHom_root K E f p x : + kHom K E f -> p \is a polyOver E -> x \in E -> root p x -> + root (map_poly f p) (f x). +Proof. +by move/kHom_horner=> homKf Ep Ex /rootP px0; rewrite /root -homKf ?px0 ?raddf0. +Qed. + +Lemma kHom_root_id K E f p x : + (K <= E)%VS -> kHom K E f -> p \is a polyOver K -> x \in E -> root p x -> + root p (f x). +Proof. +move=> sKE homKf Kp Ex /(kHom_root homKf (polyOverSv sKE Kp) Ex). +by rewrite (kHom_poly_id homKf). +Qed. + +Section kHomExtend. + +Variables (K E : {subfield L}) (f : 'End(L)) (x y : L). + +Fact kHomExtend_subproof : + linear (fun z => (map_poly f (Fadjoin_poly E x z)).[y]). +Proof. +move=> k a b; rewrite linearP /= raddfD hornerE; congr (_ + _). +rewrite -[rhs in _ = rhs]mulr_algl -hornerZ /=; congr _.[_]. +by apply/polyP => i; rewrite !(coefZ, coef_map) /= !mulr_algl linearZ. +Qed. +Definition kHomExtend := linfun (Linear kHomExtend_subproof). + +Lemma kHomExtendE z : kHomExtend z = (map_poly f (Fadjoin_poly E x z)).[y]. +Proof. by rewrite lfunE. Qed. + +Hypotheses (sKE : (K <= E)%VS) (homKf : kHom K E f). +Local Notation Px := (minPoly E x). +Hypothesis fPx_y_0 : root (map_poly f Px) y. + +Lemma kHomExtend_id z : z \in E -> kHomExtend z = f z. +Proof. by move=> Ez; rewrite kHomExtendE Fadjoin_polyC ?map_polyC ?hornerC. Qed. + +Lemma kHomExtend_val : kHomExtend x = y. +Proof. +have fX: map_poly f 'X = 'X by rewrite (kHom_poly_id homKf) ?polyOverX. +have [Ex | E'x] := boolP (x \in E); last first. + by rewrite kHomExtendE Fadjoin_polyX // fX hornerX. +have:= fPx_y_0; rewrite (minPoly_XsubC Ex) raddfB /= map_polyC fX root_XsubC /=. +by rewrite (kHomExtend_id Ex) => /eqP->. +Qed. + +Lemma kHomExtend_poly p : + p \in polyOver E -> kHomExtend p.[x] = (map_poly f p).[y]. +Proof. +move=> Ep; rewrite kHomExtendE (Fadjoin_poly_mod x) //. +rewrite (divp_eq (map_poly f p) (map_poly f Px)). +rewrite !hornerE (rootP fPx_y_0) mulr0 add0r. +have [p1 ->] := polyOver_subvs Ep. +have [Px1 ->] := polyOver_subvs (minPolyOver E x). +by rewrite -map_modp -!map_poly_comp (map_modp (kHom_rmorphism homKf)). +Qed. + +Lemma kHomExtendP : kHom K <<E; x>> kHomExtend. +Proof. +have [fM idKf] := kHomP homKf. +apply/kHomP; split=> [|z Kz]; last by rewrite kHomExtend_id ?(subvP sKE) ?idKf. +move=> _ _ /Fadjoin_polyP[p Ep ->] /Fadjoin_polyP[q Eq ->]. +rewrite -hornerM !kHomExtend_poly ?rpredM // -hornerM; congr _.[_]. +apply/polyP=> i; rewrite coef_map !coefM /= linear_sum /=. +by apply: eq_bigr => j _; rewrite !coef_map /= fM ?(polyOverP _). +Qed. + +End kHomExtend. + +Definition kAut U V f := kHom U V f && (f @: V == V)%VS. + +Lemma kAutE K E f : kAut K E f = kHom K E f && (f @: E <= E)%VS. +Proof. +apply/andP/andP=> [[-> /eqP->] // | [homKf EfE]]. +by rewrite eqEdim EfE /= (kHom_dim homKf). +Qed. + +Lemma kAutS U1 U2 V f : (U1 <= U2)%VS -> kAut U2 V f -> kAut U1 V f. +Proof. by move=> sU12 /andP[/(kHomSl sU12)homU1f EfE]; apply/andP. Qed. + +Lemma kHom_kAut_sub K E f : kAut K E f -> kHom K E f. Proof. by case/andP. Qed. + +Lemma kAut_eq K E (f g : 'End(L)) : + (K <= E)%VS -> {in E, f =1 g} -> kAut K E f = kAut K E g. +Proof. +by move=> sKE eq_fg; rewrite !kAutE (kHom_eq sKE eq_fg) (eq_in_limg eq_fg). +Qed. + +Lemma kAutfE K f : kAut K {:L} f = kHom K {:L} f. +Proof. by rewrite kAutE subvf andbT. Qed. + +Lemma kAut1E E (f : 'AEnd(L)) : kAut 1 E f = (f @: E <= E)%VS. +Proof. by rewrite kAutE k1AHom. Qed. + +Lemma kAutf_lker0 K f : kHom K {:L} f -> lker f == 0%VS. +Proof. +move/(kHomSl (sub1v _))/kHom_lrmorphism=> fM. +by apply/lker0P; apply: (fmorph_inj (RMorphism fM)). +Qed. + +Lemma inv_kHomf K f : kHom K {:L} f -> kHom K {:L} f^-1. +Proof. +move=> homKf; have [[fM idKf] kerf0] := (kHomP homKf, kAutf_lker0 homKf). +have f1K: cancel f^-1%VF f by apply: lker0_lfunVK. +apply/kHomP; split=> [x y _ _ | x Kx]; apply: (lker0P kerf0). + by rewrite fM ?memvf ?{1}f1K. +by rewrite f1K idKf. +Qed. + +Lemma inv_is_ahom (f : 'AEnd(L)) : ahom_in {:L} f^-1. +Proof. +have /ahomP/kHom_lrmorphism hom1f := valP f. +exact/ahomP/kHom_lrmorphism/inv_kHomf. +Qed. + +Canonical inv_ahom (f : 'AEnd(L)) : 'AEnd(L) := AHom (inv_is_ahom f). +Notation "f ^-1" := (inv_ahom f) : lrfun_scope. + +Lemma comp_kHom_img K E f g : + kHom K (g @: E) f -> kHom K E g -> kHom K E (f \o g). +Proof. +move=> /kHomP[fM idKf] /kHomP[gM idKg]; apply/kHomP; split=> [x y Ex Ey | x Kx]. + by rewrite !lfunE /= gM // fM ?memv_img. +by rewrite lfunE /= idKg ?idKf. +Qed. + +Lemma comp_kHom K E f g : kHom K {:L} f -> kHom K E g -> kHom K E (f \o g). +Proof. by move/(kHomSr (subvf (g @: E))); apply: comp_kHom_img. Qed. + +Lemma kHom_extends K E f p U : + (K <= E)%VS -> kHom K E f -> + p \is a polyOver K -> splittingFieldFor E p U -> + {g | kHom K U g & {in E, f =1 g}}. +Proof. +move=> sKE homEf Kp /sig2_eqW[rs Dp <-{U}]. +set r := rs; have rs_r: all (mem rs) r by apply/allP. +elim: r rs_r => [_|z r IHr /=/andP[rs_z rs_r]] /= in E f sKE homEf *. + by exists f; rewrite ?Fadjoin_nil. +set Ez := <<E; z>>%AS; pose fpEz := map_poly f (minPoly E z). +suffices{IHr} /sigW[y fpEz_y]: exists y, root fpEz y. + have homEz_fz: kHom K Ez (kHomExtend E f z y) by apply: kHomExtendP. + have sKEz: (K <= Ez)%VS := subv_trans sKE (subv_adjoin E z). + have [g homGg Dg] := IHr rs_r _ _ sKEz homEz_fz. + exists g => [|x Ex]; first by rewrite adjoin_cons. + by rewrite -Dg ?subvP_adjoin // kHomExtend_id. +have [m DfpEz]: {m | fpEz %= \prod_(w <- mask m rs) ('X - w%:P)}. + apply: dvdp_prod_XsubC; rewrite -(eqp_dvdr _ Dp) -(kHom_poly_id homEf Kp). + have /polyOver_subvs[q Dq] := polyOverSv sKE Kp. + have /polyOver_subvs[qz Dqz] := minPolyOver E z. + rewrite /fpEz Dq Dqz -2?{1}map_poly_comp (dvdp_map (kHom_rmorphism homEf)). + rewrite -(dvdp_map [rmorphism of @vsval _ _ E]) -Dqz -Dq. + by rewrite minPoly_dvdp ?(polyOverSv sKE) // (eqp_root Dp) root_prod_XsubC. +exists (mask m rs)`_0; rewrite (eqp_root DfpEz) root_prod_XsubC mem_nth //. +rewrite -ltnS -(size_prod_XsubC _ id) -(eqp_size DfpEz). +rewrite size_poly_eq -?lead_coefE ?size_minPoly // (monicP (monic_minPoly E z)). +by have [_ idKf] := kHomP homEf; rewrite idKf ?mem1v ?oner_eq0. +Qed. + +End kHom. + +Notation "f ^-1" := (inv_ahom f) : lrfun_scope. + +Implicit Arguments kHomP [F L K V f]. +Implicit Arguments kAHomP [F L U V f]. +Implicit Arguments kHom_lrmorphism [F L f]. + +Module SplittingField. + +Import GRing. + +Section ClassDef. + +Variable F : fieldType. + +Definition axiom (L : fieldExtType F) := + exists2 p : {poly L}, p \is a polyOver 1%VS & splittingFieldFor 1 p {:L}. + +Record class_of (L : Type) : Type := + Class {base : FieldExt.class_of F L; _ : axiom (FieldExt.Pack _ base L)}. +Local Coercion base : class_of >-> FieldExt.class_of. + +Structure type (phF : phant F) := Pack {sort; _ : class_of sort; _ : Type}. +Local Coercion sort : type >-> Sortclass. +Variable (phF : phant F) (T : Type) (cT : type phF). +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 phF T c T. + +Definition pack b0 (ax0 : axiom (@FieldExt.Pack F (Phant F) T b0 T)) := + fun bT b & phant_id (@FieldExt.class F phF bT) b => + fun ax & phant_id ax0 ax => Pack (Phant F) (@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 unitRingType := @UnitRing.Pack cT xclass xT. +Definition comRingType := @ComRing.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 lmodType := @Lmodule.Pack F phF cT xclass xT. +Definition lalgType := @Lalgebra.Pack F phF cT xclass xT. +Definition algType := @Algebra.Pack F phF cT xclass xT. +Definition unitAlgType := @UnitAlgebra.Pack F phF cT xclass xT. +Definition vectType := @Vector.Pack F phF cT xclass xT. +Definition FalgType := @Falgebra.Pack F phF cT xclass xT. +Definition fieldExtType := @FieldExt.Pack F phF cT xclass xT. + +End ClassDef. + +Module Exports. + +Coercion sort : type >-> Sortclass. +Bind Scope ring_scope with sort. +Coercion base : class_of >-> FieldExt.class_of. +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 comRingType : type >-> ComRing.type. +Canonical comRingType. +Coercion comUnitRingType : type >-> ComUnitRing.type. +Canonical comUnitRingType. +Coercion idomainType : type >-> IntegralDomain.type. +Canonical idomainType. +Coercion fieldType : type >-> Field.type. +Canonical fieldType. +Coercion lmodType : type >-> Lmodule.type. +Canonical lmodType. +Coercion lalgType : type >-> Lalgebra.type. +Canonical lalgType. +Coercion algType : type >-> Algebra.type. +Canonical algType. +Coercion unitAlgType : type >-> UnitAlgebra.type. +Canonical unitAlgType. +Coercion vectType : type >-> Vector.type. +Canonical vectType. +Coercion FalgType : type >-> Falgebra.type. +Canonical FalgType. +Coercion fieldExtType : type >-> FieldExt.type. +Canonical fieldExtType. + +Notation splittingFieldType F := (type (Phant F)). +Notation SplittingFieldType F L ax := (@pack _ (Phant F) L _ ax _ _ id _ id). +Notation "[ 'splittingFieldType' F 'of' L 'for' K ]" := + (@clone _ (Phant F) L K _ idfun) + (at level 0, format "[ 'splittingFieldType' F 'of' L 'for' K ]") + : form_scope. +Notation "[ 'splittingFieldType' F 'of' L ]" := + (@clone _ (Phant F) L _ _ id) + (at level 0, format "[ 'splittingFieldType' F 'of' L ]") : form_scope. + +End Exports. +End SplittingField. +Export SplittingField.Exports. + +Lemma normal_field_splitting (F : fieldType) (L : fieldExtType F) : + (forall (K : {subfield L}) x, + exists r, minPoly K x == \prod_(y <- r) ('X - y%:P)) -> + SplittingField.axiom L. +Proof. +move=> normalL; pose r i := sval (sigW (normalL 1%AS (tnth (vbasis {:L}) i))). +have sz_r i: size (r i) <= \dim {:L}. + rewrite -ltnS -(size_prod_XsubC _ id) /r; case: sigW => _ /= /eqP <-. + rewrite size_minPoly ltnS; move: (tnth _ _) => x. + by rewrite adjoin_degreeE dimv1 divn1 dimvS // subvf. +pose mkf (z : L) := 'X - z%:P. +exists (\prod_i \prod_(j < \dim {:L} | j < size (r i)) mkf (r i)`_j). + apply: rpred_prod => i _; rewrite big_ord_narrow /= /r; case: sigW => rs /=. + by rewrite (big_nth 0) big_mkord => /eqP <- {rs}; apply: minPolyOver. +rewrite pair_big_dep /= -big_filter filter_index_enum -(big_map _ xpredT mkf). +set rF := map _ _; exists rF; first exact: eqpxx. +apply/eqP; rewrite eqEsubv subvf -(span_basis (vbasisP {:L})). +apply/span_subvP=> _ /tnthP[i ->]; set x := tnth _ i. +have /tnthP[j ->]: x \in in_tuple (r i). + by rewrite -root_prod_XsubC /r; case: sigW => _ /=/eqP<-; apply: root_minPoly. +apply/seqv_sub_adjoin/imageP; rewrite (tnth_nth 0) /in_mem/=. +by exists (i, widen_ord (sz_r i) j) => /=. +Qed. + +Section SplittingFieldTheory. + +Variables (F : fieldType) (L : splittingFieldType F). + +Implicit Types (U V W : {vspace L}). +Implicit Types (K M E : {subfield L}). + +Lemma splittingFieldP : SplittingField.axiom L. +Proof. by case: L => ? []. Qed. + +Lemma splittingPoly : + {p : {poly L} | p \is a polyOver 1%VS & splittingFieldFor 1 p {:L}}. +Proof. +pose factF p s := (p \is a polyOver 1%VS) && (p %= \prod_(z <- s) ('X - z%:P)). +suffices [[p rs] /andP[]]: {ps | factF F L ps.1 ps.2 & <<1 & ps.2>> = {:L}}%VS. + by exists p; last exists rs. +apply: sig2_eqW; have [p F0p [rs splitLp genLrs]] := splittingFieldP. +by exists (p, rs); rewrite // /factF F0p splitLp. +Qed. + +Fact fieldOver_splitting E : SplittingField.axiom (fieldOver_fieldExtType E). +Proof. +have [p Fp [r Dp defL]] := splittingFieldP; exists p. + apply/polyOverP=> j; rewrite trivial_fieldOver. + by rewrite (subvP (sub1v E)) ?(polyOverP Fp). +exists r => //; apply/vspaceP=> x; rewrite memvf. +have [L0 [_ _ defL0]] := @aspaceOverP _ _ E <<1 & r : seq (fieldOver E)>>. +rewrite defL0; have: x \in <<1 & r>>%VS by rewrite defL (@memvf _ L). +apply: subvP; apply/Fadjoin_seqP; rewrite -memvE -defL0 mem1v. +by split=> // y r_y; rewrite -defL0 seqv_sub_adjoin. +Qed. +Canonical fieldOver_splittingFieldType E := + SplittingFieldType (subvs_of E) (fieldOver E) (fieldOver_splitting E). + +Lemma enum_AEnd : {kAutL : seq 'AEnd(L) | forall f, f \in kAutL}. +Proof. +pose isAutL (s : seq 'AEnd(L)) (f : 'AEnd(L)) := kHom 1 {:L} f = (f \in s). +suffices [kAutL in_kAutL] : {kAutL : seq 'AEnd(L) | forall f, isAutL kAutL f}. + by exists kAutL => f; rewrite -in_kAutL k1AHom. +have [p Kp /sig2_eqW[rs Dp defL]] := splittingPoly. +do [rewrite {}/isAutL -(erefl (asval 1)); set r := rs; set E := 1%AS] in defL *. +have [sKE rs_r]: (1 <= E)%VS /\ all (mem rs) r by split; last apply/allP. +elim: r rs_r => [_|z r IHr /=/andP[rs_z rs_r]] /= in (E) sKE defL *. + rewrite Fadjoin_nil in defL; exists [tuple \1%AF] => f; rewrite defL inE. + apply/idP/eqP=> [/kAHomP f1 | ->]; last exact: kHom1. + by apply/val_inj/lfunP=> x; rewrite id_lfunE f1 ?memvf. +do [set Ez := <<E; z>>%VS; rewrite adjoin_cons] in defL. +have sEEz: (E <= Ez)%VS := subv_adjoin E z; have sKEz := subv_trans sKE sEEz. +have{IHr} [homEz DhomEz] := IHr rs_r _ sKEz defL. +have Ep: p \in polyOver E := polyOverSv sKE Kp. +have{rs_z} pz0: root p z by rewrite (eqp_root Dp) root_prod_XsubC. +pose pEz := minPoly E z; pose n := \dim_E Ez. +have{pz0} [rz DpEz]: {rz : n.-tuple L | pEz %= \prod_(w <- rz) ('X - w%:P)}. + have /dvdp_prod_XsubC[m DpEz]: pEz %| \prod_(w <- rs) ('X - w%:P). + by rewrite -(eqp_dvdr _ Dp) minPoly_dvdp ?(polyOverSv sKE). + suffices sz_rz: size (mask m rs) == n by exists (Tuple sz_rz). + rewrite -[n]adjoin_degreeE -eqSS -size_minPoly. + by rewrite (eqp_size DpEz) size_prod_XsubC. +have fEz i (y := tnth rz i): {f : 'AEnd(L) | kHom E {:L} f & f z = y}. + have homEfz: kHom E Ez (kHomExtend E \1 z y). + rewrite kHomExtendP ?kHom1 // lfun1_poly. + by rewrite (eqp_root DpEz) -/rz root_prod_XsubC mem_tnth. + have splitFp: splittingFieldFor Ez p {:L}. + exists rs => //; apply/eqP; rewrite eqEsubv subvf -defL adjoin_seqSr //. + exact/allP. + have [f homLf Df] := kHom_extends sEEz homEfz Ep splitFp. + have [ahomf _] := andP homLf; exists (AHom ahomf) => //. + rewrite -Df ?memv_adjoin ?(kHomExtend_val (kHom1 E E)) // lfun1_poly. + by rewrite (eqp_root DpEz) root_prod_XsubC mem_tnth. +exists [seq (s2val (fEz i) \o f)%AF| i <- enum 'I_n, f <- homEz] => f. +apply/idP/allpairsP => [homLf | [[i g] [_ Hg ->]] /=]; last first. + by case: (fEz i) => fi /= /comp_kHom->; rewrite ?(kHomSl sEEz) ?DhomEz. +have /tnthP[i Dfz]: f z \in rz. + rewrite memtE /= -root_prod_XsubC -(eqp_root DpEz). + by rewrite (kHom_root_id _ homLf) ?memvf ?subvf ?minPolyOver ?root_minPoly. +case Dfi: (fEz i) => [fi homLfi fi_z]; have kerfi0 := kAutf_lker0 homLfi. +set fj := (fi ^-1 \o f)%AF; suffices Hfj : fj \in homEz. + exists (i, fj) => //=; rewrite mem_enum inE Hfj; split => //. + by apply/val_inj; rewrite {}Dfi /= (lker0_compVKf kerfi0). +rewrite -DhomEz; apply/kAHomP => _ /Fadjoin_polyP[q Eq ->]. +have homLfj: kHom E {:L} fj := comp_kHom (inv_kHomf homLfi) homLf. +have /kHom_lrmorphism fjM := kHomSl (sub1v _) homLfj. +rewrite -[fj _](horner_map (RMorphism fjM)) (kHom_poly_id homLfj) //=. +by rewrite lfunE /= Dfz -fi_z lker0_lfunK. +Qed. + +Lemma splitting_field_normal K x : + exists r, minPoly K x == \prod_(y <- r) ('X - y%:P). +Proof. +pose q1 := minPoly 1 x; pose fx_root q (f : 'AEnd(L)) := root q (f x). +have [[p F0p splitLp] [autL DautL]] := (splittingFieldP, enum_AEnd). +suffices{K} autL_px q: q != 0 -> q %| q1 -> size q > 1 -> has (fx_root q) autL. + set q := minPoly K x; have: q \is monic := monic_minPoly K x. + have: q %| q1 by rewrite minPolyS // sub1v. + elim: {q}_.+1 {-2}q (ltnSn (size q)) => // d IHd q leqd q_dv_q1 mon_q. + have nz_q: q != 0 := monic_neq0 mon_q. + have [|q_gt1|q_1] := ltngtP (size q) 1; last first; last by rewrite polySpred. + by exists nil; rewrite big_nil -eqp_monic ?monic1 // -size_poly_eq1 q_1. + have /hasP[f autLf /factor_theorem[q2 Dq]] := autL_px q nz_q q_dv_q1 q_gt1. + have mon_q2: q2 \is monic by rewrite -(monicMr _ (monicXsubC (f x))) -Dq. + rewrite Dq size_monicM -?size_poly_eq0 ?size_XsubC ?addn2 //= ltnS in leqd. + have q2_dv_q1: q2 %| q1 by rewrite (dvdp_trans _ q_dv_q1) // Dq dvdp_mulr. + rewrite Dq; have [r /eqP->] := IHd q2 leqd q2_dv_q1 mon_q2. + by exists (f x :: r); rewrite big_cons mulrC. +elim: {q}_.+1 {-2}q (ltnSn (size q)) => // d IHd q leqd nz_q q_dv_q1 q_gt1. +without loss{d leqd IHd nz_q q_gt1} irr_q: q q_dv_q1 / irreducible_poly q. + move=> IHq; apply: wlog_neg => not_autLx_q; apply: IHq => //. + split=> // q2 q2_neq1 q2_dv_q; rewrite -dvdp_size_eqp // eqn_leq dvdp_leq //=. + rewrite leqNgt; apply: contra not_autLx_q => ltq2q. + have nz_q2: q2 != 0 by apply: contraTneq q2_dv_q => ->; rewrite dvd0p. + have{q2_neq1} q2_gt1: size q2 > 1 by rewrite neq_ltn polySpred in q2_neq1 *. + have{leqd ltq2q} ltq2d: size q2 < d by apply: leq_trans ltq2q _. + apply: sub_has (IHd _ ltq2d nz_q2 (dvdp_trans q2_dv_q q_dv_q1) q2_gt1) => f. + by rewrite /fx_root !root_factor_theorem => /dvdp_trans->. +have{irr_q} [Lz [inLz [z qz0]]]: {Lz : fieldExtType F & + {inLz : 'AHom(L, Lz) & {z : Lz | root (map_poly inLz q) z}}}. +- have [Lz0 _ [z qz0 defLz]] := irredp_FAdjoin irr_q. + pose Lz := baseField_extFieldType Lz0. + pose inLz : {rmorphism L -> Lz} := [rmorphism of in_alg Lz0]. + have inLzL_linear: linear (locked inLz). + move=> a u v; rewrite -(@mulr_algl F Lz) baseField_scaleE. + by rewrite -{1}mulr_algl rmorphD rmorphM -lock. + have ihLzZ: ahom_in {:L} (linfun (Linear inLzL_linear)). + by apply/ahom_inP; split=> [u v|]; rewrite !lfunE (rmorphM, rmorph1). + exists Lz, (AHom ihLzZ), z; congr (root _ z): qz0. + by apply: eq_map_poly => y; rewrite lfunE /= -lock. +pose imL := [aspace of limg inLz]; pose pz := map_poly inLz p. +have in_imL u: inLz u \in imL by rewrite memv_img ?memvf. +have F0pz: pz \is a polyOver 1%VS. + apply/polyOverP=> i; rewrite -(aimg1 inLz) coef_map /= memv_img //. + exact: (polyOverP F0p). +have{splitLp} splitLpz: splittingFieldFor 1 pz imL. + have [r def_p defL] := splitLp; exists (map inLz r) => [|{def_p}]. + move: def_p; rewrite -(eqp_map [rmorphism of inLz]) rmorph_prod. + rewrite big_map; congr (_ %= _); apply: eq_big => // y _. + by rewrite rmorphB /= map_polyX map_polyC. + apply/eqP; rewrite eqEsubv /= -{2}defL {defL}; apply/andP; split. + by apply/Fadjoin_seqP; rewrite sub1v; split=> // _ /mapP[y r_y ->]. + elim/last_ind: r => [|r y IHr] /=; first by rewrite !Fadjoin_nil aimg1. + rewrite map_rcons !adjoin_rcons /=. + apply/subvP=> _ /memv_imgP[_ /Fadjoin_polyP[p1 r_p1 ->] ->]. + rewrite -horner_map /= mempx_Fadjoin //=; apply/polyOverP=> i. + by rewrite coef_map (subvP IHr) //= memv_img ?(polyOverP r_p1). +have [f homLf fxz]: exists2 f : 'End(Lz), kHom 1 imL f & f (inLz x) = z. + pose q1z := minPoly 1 (inLz x). + have Dq1z: map_poly inLz q1 %| q1z. + have F0q1z i: exists a, q1z`_i = a%:A by apply/vlineP/polyOverP/minPolyOver. + have [q2 Dq2]: exists q2, q1z = map_poly inLz q2. + exists (\poly_(i < size q1z) (sval (sig_eqW (F0q1z i)))%:A). + rewrite -{1}[q1z]coefK; apply/polyP=> i; rewrite coef_map !{1}coef_poly. + by case: sig_eqW => a; case: ifP; rewrite /= ?rmorph0 ?linearZ ?rmorph1. + rewrite Dq2 dvdp_map minPoly_dvdp //. + apply/polyOverP=> i; have[a] := F0q1z i. + rewrite -(rmorph1 [rmorphism of inLz]) -linearZ. + by rewrite Dq2 coef_map => /fmorph_inj->; rewrite rpredZ ?mem1v. + by rewrite -(fmorph_root [rmorphism of inLz]) -Dq2 root_minPoly. + have q1z_z: root q1z z. + rewrite !root_factor_theorem in qz0 *. + by apply: dvdp_trans qz0 (dvdp_trans _ Dq1z); rewrite dvdp_map. + have map1q1z_z: root (map_poly \1%VF q1z) z. + by rewrite map_poly_id => // ? _; rewrite lfunE. + pose f0 := kHomExtend 1 \1 (inLz x) z. + have{map1q1z_z} hom_f0 : kHom 1 <<1; inLz x>> f0. + by apply: kHomExtendP map1q1z_z => //; apply: kHom1. + have{splitLpz} splitLpz: splittingFieldFor <<1; inLz x>> pz imL. + have [r def_pz defLz] := splitLpz; exists r => //. + apply/eqP; rewrite eqEsubv -{2}defLz adjoin_seqSl ?sub1v // andbT. + apply/Fadjoin_seqP; split; last first. + by rewrite /= -[limg _]defLz; apply: seqv_sub_adjoin. + by apply/FadjoinP/andP; rewrite sub1v memv_img ?memvf. + have [f homLzf Df] := kHom_extends (sub1v _) hom_f0 F0pz splitLpz. + have [-> | x'z] := eqVneq (inLz x) z. + by exists \1%VF; rewrite ?lfunE ?kHom1. + exists f => //; rewrite -Df ?memv_adjoin ?(kHomExtend_val (kHom1 1 1)) //. + by rewrite lfun1_poly. +pose f1 := (inLz^-1 \o f \o inLz)%VF; have /kHomP[fM fFid] := homLf. +have Df1 u: inLz (f1 u) = f (inLz u). + rewrite !comp_lfunE limg_lfunVK //= -[limg _]/(asval imL). + have [r def_pz defLz] := splitLpz. + have []: all (mem r) r /\ inLz u \in imL by split; first apply/allP. + rewrite -{1}defLz; elim/last_ind: {-1}r {u}(inLz u) => [|r1 y IHr1] u. + by rewrite Fadjoin_nil => _ Fu; rewrite fFid // (subvP (sub1v _)). + rewrite all_rcons adjoin_rcons => /andP[rr1 ry] /Fadjoin_polyP[pu r1pu ->]. + rewrite (kHom_horner homLf) -defLz; last exact: seqv_sub_adjoin; last first. + by apply: polyOverS r1pu; apply/subvP/adjoin_seqSr/allP. + apply: rpred_horner. + by apply/polyOverP=> i; rewrite coef_map /= defLz IHr1 ?(polyOverP r1pu). + rewrite seqv_sub_adjoin // -root_prod_XsubC -(eqp_root def_pz). + rewrite (kHom_root_id _ homLf) ?sub1v //. + by rewrite -defLz seqv_sub_adjoin. + by rewrite (eqp_root def_pz) root_prod_XsubC. +suffices f1_is_ahom : ahom_in {:L} f1. + apply/hasP; exists (AHom f1_is_ahom); first exact: DautL. + by rewrite /fx_root -(fmorph_root [rmorphism of inLz]) /= Df1 fxz. +apply/ahom_inP; split=> [a b _ _|]; apply: (fmorph_inj [rmorphism of inLz]). + by rewrite rmorphM /= !Df1 rmorphM fM ?in_imL. +by rewrite /= Df1 /= fFid ?rmorph1 ?mem1v. +Qed. + +Lemma kHom_to_AEnd K E f : kHom K E f -> {g : 'AEnd(L) | {in E, f =1 val g}}. +Proof. +move=> homKf; have{homKf} [homFf sFE] := (kHomSl (sub1v K) homKf, sub1v E). +have [p Fp /(splittingFieldForS sFE (subvf E))splitLp] := splittingPoly. +have [g0 homLg0 eq_fg] := kHom_extends sFE homFf Fp splitLp. +by apply: exist (Sub g0 _) _ => //; apply/ahomP/kHom_lrmorphism. +Qed. + +End SplittingFieldTheory. + +(* Hide the finGroup structure on 'AEnd(L) in a module so that we can control *) +(* when it is exported. Most people will want to use the finGroup structure *) +(* on 'Gal(E / K) and will not need this module. *) +Module Import AEnd_FinGroup. +Section AEnd_FinGroup. + +Variables (F : fieldType) (L : splittingFieldType F). +Implicit Types (U V W : {vspace L}) (K M E : {subfield L}). + +Definition inAEnd f := SeqSub (svalP (enum_AEnd L) f). +Fact inAEndK : cancel inAEnd val. Proof. by []. Qed. + +Definition AEnd_countMixin := Eval hnf in CanCountMixin inAEndK. +Canonical AEnd_countType := Eval hnf in CountType 'AEnd(L) AEnd_countMixin. +Canonical AEnd_subCountType := Eval hnf in [subCountType of 'AEnd(L)]. +Definition AEnd_finMixin := Eval hnf in CanFinMixin inAEndK. +Canonical AEnd_finType := Eval hnf in FinType 'AEnd(L) AEnd_finMixin. +Canonical AEnd_subFinType := Eval hnf in [subFinType of 'AEnd(L)]. + +(* the group operation is the categorical composition operation *) +Definition comp_AEnd (f g : 'AEnd(L)) : 'AEnd(L) := (g \o f)%AF. + +Fact comp_AEndA : associative comp_AEnd. +Proof. by move=> f g h; apply: val_inj; symmetry; apply: comp_lfunA. Qed. + +Fact comp_AEnd1l : left_id \1%AF comp_AEnd. +Proof. by move=> f; apply/val_inj/comp_lfun1r. Qed. + +Fact comp_AEndK : left_inverse \1%AF (@inv_ahom _ L) comp_AEnd. +Proof. by move=> f; apply/val_inj; rewrite /= lker0_compfV ?AEnd_lker0. Qed. + +Definition AEnd_baseFinGroupMixin := + FinGroup.Mixin comp_AEndA comp_AEnd1l comp_AEndK. +Canonical AEnd_baseFinGroupType := + BaseFinGroupType 'AEnd(L) AEnd_baseFinGroupMixin. +Canonical AEnd_finGroupType := FinGroupType comp_AEndK. + +Definition kAEnd U V := [set f : 'AEnd(L) | kAut U V f]. +Definition kAEndf U := kAEnd U {:L}. + +Lemma kAEnd_group_set K E : group_set (kAEnd K E). +Proof. +apply/group_setP; split=> [|f g]; first by rewrite inE /kAut kHom1 lim1g eqxx. +rewrite !inE !kAutE => /andP[homKf EfE] /andP[/(kHomSr EfE)homKg EgE]. +by rewrite (comp_kHom_img homKg homKf) limg_comp (subv_trans _ EgE) ?limgS. +Qed. +Canonical kAEnd_group K E := group (kAEnd_group_set K E). +Canonical kAEndf_group K := [group of kAEndf K]. + +Lemma kAEnd_norm K E : kAEnd K E \subset 'N(kAEndf E)%g. +Proof. +apply/subsetP=> x; rewrite -groupV 2!in_set => /andP[_ /eqP ExE]. +apply/subsetP=> _ /imsetP[y homEy ->]; rewrite !in_set !kAutfE in homEy *. +apply/kAHomP=> u Eu; have idEy := kAHomP homEy; rewrite -ExE in idEy. +by rewrite !lfunE /= lfunE /= idEy ?memv_img // lker0_lfunVK ?AEnd_lker0. +Qed. + +Lemma mem_kAut_coset K E (g : 'AEnd(L)) : + kAut K E g -> g \in coset (kAEndf E) g. +Proof. +move=> autEg; rewrite val_coset ?rcoset_refl //. +by rewrite (subsetP (kAEnd_norm K E)) // inE. +Qed. + +Lemma aut_mem_eqP E (x y : coset_of (kAEndf E)) f g : + f \in x -> g \in y -> reflect {in E, f =1 g} (x == y). +Proof. +move=> x_f y_g; rewrite -(coset_mem x_f) -(coset_mem y_g). +have [Nf Ng] := (subsetP (coset_norm x) f x_f, subsetP (coset_norm y) g y_g). +rewrite (sameP eqP (rcoset_kercosetP Nf Ng)) mem_rcoset inE kAutfE. +apply: (iffP kAHomP) => idEfg u Eu. + by rewrite -(mulgKV g f) lfunE /= idEfg. +by rewrite lfunE /= idEfg // lker0_lfunK ?AEnd_lker0. +Qed. + +End AEnd_FinGroup. +End AEnd_FinGroup. + +Section GaloisTheory. + +Variables (F : fieldType) (L : splittingFieldType F). + +Implicit Types (U V W : {vspace L}). +Implicit Types (K M E : {subfield L}). + +(* We take Galois automorphisms for a subfield E to be automorphisms of the *) +(* full field {:L} that operate in E taken modulo those that fix E pointwise. *) +(* The type of Galois automorphisms of E is then the subtype of elements of *) +(* the quotient kAEnd 1 E / kAEndf E, which we encapsulate in a specific *) +(* wrapper to ensure stability of the gal_repr coercion insertion. *) +Section gal_of_Definition. + +Variable V : {vspace L}. + +(* The <<_>>, which becomes redundant when V is a {subfield L}, ensures that *) +(* the argument of [subg _] is syntactically a group. *) +Inductive gal_of := Gal of [subg kAEnd_group 1 <<V>> / kAEndf (agenv V)]. +Definition gal (f : 'AEnd(L)) := Gal (subg _ (coset _ f)). +Definition gal_sgval x := let: Gal u := x in u. + +Fact gal_sgvalK : cancel gal_sgval Gal. Proof. by case. Qed. +Let gal_sgval_inj := can_inj gal_sgvalK. + +Definition gal_eqMixin := CanEqMixin gal_sgvalK. +Canonical gal_eqType := Eval hnf in EqType gal_of gal_eqMixin. +Definition gal_choiceMixin := CanChoiceMixin gal_sgvalK. +Canonical gal_choiceType := Eval hnf in ChoiceType gal_of gal_choiceMixin. +Definition gal_countMixin := CanCountMixin gal_sgvalK. +Canonical gal_countType := Eval hnf in CountType gal_of gal_countMixin. +Definition gal_finMixin := CanFinMixin gal_sgvalK. +Canonical gal_finType := Eval hnf in FinType gal_of gal_finMixin. + +Definition gal_one := Gal 1%g. +Definition gal_inv x := Gal (gal_sgval x)^-1. +Definition gal_mul x y := Gal (gal_sgval x * gal_sgval y). +Fact gal_oneP : left_id gal_one gal_mul. +Proof. by move=> x; apply/gal_sgval_inj/mul1g. Qed. +Fact gal_invP : left_inverse gal_one gal_inv gal_mul. +Proof. by move=> x; apply/gal_sgval_inj/mulVg. Qed. +Fact gal_mulP : associative gal_mul. +Proof. by move=> x y z; apply/gal_sgval_inj/mulgA. Qed. + +Definition gal_finGroupMixin := + FinGroup.Mixin gal_mulP gal_oneP gal_invP. +Canonical gal_finBaseGroupType := + Eval hnf in BaseFinGroupType gal_of gal_finGroupMixin. +Canonical gal_finGroupType := Eval hnf in FinGroupType gal_invP. + +Coercion gal_repr u : 'AEnd(L) := repr (sgval (gal_sgval u)). + +Fact gal_is_morphism : {in kAEnd 1 (agenv V) &, {morph gal : x y / x * y}%g}. +Proof. +move=> f g /= autEa autEb; congr (Gal _). +by rewrite !morphM ?mem_morphim // (subsetP (kAEnd_norm 1 _)). +Qed. +Canonical gal_morphism := Morphism gal_is_morphism. + +Lemma gal_reprK : cancel gal_repr gal. +Proof. by case=> x; rewrite /gal coset_reprK sgvalK. Qed. + +Lemma gal_repr_inj : injective gal_repr. +Proof. exact: can_inj gal_reprK. Qed. + +Lemma gal_AEnd x : gal_repr x \in kAEnd 1 (agenv V). +Proof. +rewrite /gal_repr; case/gal_sgval: x => _ /=/morphimP[g Ng autEg ->]. +rewrite val_coset //=; case: repr_rcosetP => f; rewrite groupMr // !inE kAut1E. +by rewrite kAutE -andbA => /and3P[_ /fixedSpace_limg-> _]. +Qed. + +End gal_of_Definition. + +Prenex Implicits gal_repr. + +Lemma gal_eqP E {x y : gal_of E} : reflect {in E, x =1 y} (x == y). +Proof. +by rewrite -{1}(subfield_closed E); apply: aut_mem_eqP; apply: mem_repr_coset. +Qed. + +Lemma galK E (f : 'AEnd(L)) : (f @: E <= E)%VS -> {in E, gal E f =1 f}. +Proof. +rewrite -kAut1E -{1 2}(subfield_closed E) => autEf. +apply: (aut_mem_eqP (mem_repr_coset _) _ (eqxx _)). +by rewrite subgK /= ?(mem_kAut_coset autEf) // ?mem_quotient ?inE. +Qed. + +Lemma eq_galP E (f g : 'AEnd(L)) : + (f @: E <= E)%VS -> (g @: E <= E)%VS -> + reflect {in E, f =1 g} (gal E f == gal E g). +Proof. +move=> EfE EgE. +by apply: (iffP gal_eqP) => Dfg a Ea; have:= Dfg a Ea; rewrite !{1}galK. +Qed. + +Lemma limg_gal E (x : gal_of E) : (x @: E)%VS = E. +Proof. by have:= gal_AEnd x; rewrite inE subfield_closed => /andP[_ /eqP]. Qed. + +Lemma memv_gal E (x : gal_of E) a : a \in E -> x a \in E. +Proof. by move/(memv_img x); rewrite limg_gal. Qed. + +Lemma gal_id E a : (1 : gal_of E)%g a = a. +Proof. by rewrite /gal_repr repr_coset1 id_lfunE. Qed. + +Lemma galM E (x y : gal_of E) a : a \in E -> (x * y)%g a = y (x a). +Proof. +rewrite /= -comp_lfunE; apply/eq_galP; rewrite ?limg_comp ?limg_gal //. +by rewrite morphM /= ?gal_reprK ?gal_AEnd. +Qed. + +Lemma galV E (x : gal_of E) : {in E, (x^-1)%g =1 x^-1%VF}. +Proof. +move=> a Ea; apply: canRL (lker0_lfunK (AEnd_lker0 _)) _. +by rewrite -galM // mulVg gal_id. +Qed. + +(* Standard mathematical notation for 'Gal(E / K) puts the larger field first.*) +Definition galoisG V U := gal V @* <<kAEnd (U :&: V) V>>. +Local Notation "''Gal' ( V / U )" := (galoisG V U) : group_scope. +Canonical galoisG_group E U := Eval hnf in [group of (galoisG E U)]. +Local Notation "''Gal' ( V / U )" := (galoisG_group V U) : Group_scope. + +Section Automorphism. + +Lemma gal_cap U V : 'Gal(V / U) = 'Gal(V / U :&: V). +Proof. by rewrite /galoisG -capvA capvv. Qed. + +Lemma gal_kAut K E x : (K <= E)%VS -> (x \in 'Gal(E / K)) = kAut K E x. +Proof. +move=> sKE; apply/morphimP/idP=> /= [[g EgE KautEg ->{x}] | KautEx]. + rewrite genGid !inE kAut1E /= subfield_closed (capv_idPl sKE) in KautEg EgE. + by apply: etrans KautEg; apply/(kAut_eq sKE); apply: galK. +exists (x : 'AEnd(L)); rewrite ?gal_reprK ?gal_AEnd //. +by rewrite (capv_idPl sKE) mem_gen ?inE. +Qed. + +Lemma gal_kHom K E x : (K <= E)%VS -> (x \in 'Gal(E / K)) = kHom K E x. +Proof. by move/gal_kAut->; rewrite /kAut limg_gal eqxx andbT. Qed. + +Lemma kAut_to_gal K E f : + kAut K E f -> {x : gal_of E | x \in 'Gal(E / K) & {in E, f =1 x}}. +Proof. +case/andP=> homKf EfE; have [g Df] := kHom_to_AEnd homKf. +have{homKf EfE} autEg: kAut (K :&: E) E g. + rewrite /kAut -(kHom_eq (capvSr _ _) Df) (kHomSl (capvSl _ _) homKf) /=. + by rewrite -(eq_in_limg Df). +have FautEg := kAutS (sub1v _) autEg. +exists (gal E g) => [|a Ea]; last by rewrite {f}Df // galK // -kAut1E. +by rewrite mem_morphim /= ?subfield_closed ?genGid ?inE. +Qed. + +Lemma fixed_gal K E x a : + (K <= E)%VS -> x \in 'Gal(E / K) -> a \in K -> x a = a. +Proof. by move/gal_kHom=> -> /kAHomP idKx /idKx. Qed. + +Lemma fixedPoly_gal K E x p : + (K <= E)%VS -> x \in 'Gal(E / K) -> p \is a polyOver K -> map_poly x p = p. +Proof. +move=> sKE galEKx /polyOverP Kp; apply/polyP => i. +by rewrite coef_map /= (fixed_gal sKE). +Qed. + +Lemma root_minPoly_gal K E x a : + (K <= E)%VS -> x \in 'Gal(E / K) -> a \in E -> root (minPoly K a) (x a). +Proof. +move=> sKE galEKx Ea; have homKx: kHom K E x by rewrite -gal_kHom. +have K_Pa := minPolyOver K a; rewrite -[minPoly K a](fixedPoly_gal _ galEKx) //. +by rewrite (kHom_root homKx) ?root_minPoly // (polyOverS (subvP sKE)). +Qed. + +End Automorphism. + +Lemma gal_adjoin_eq K a x y : + x \in 'Gal(<<K; a>> / K) -> y \in 'Gal(<<K; a>> / K) -> + (x == y) = (x a == y a). +Proof. +move=> galKa_x galKa_y; apply/idP/eqP=> [/eqP-> // | eq_xy_a]. +apply/gal_eqP => _ /Fadjoin_polyP[p Kp ->]. +by rewrite -!horner_map !(fixedPoly_gal (subv_adjoin K a)) //= eq_xy_a. +Qed. + +Lemma galS K M E : (K <= M)%VS -> 'Gal(E / M) \subset 'Gal(E / K). +Proof. +rewrite gal_cap (gal_cap K E) => sKM; apply/subsetP=> x. +by rewrite !gal_kAut ?capvSr //; apply: kAutS; apply: capvS. +Qed. + +Lemma gal_conjg K E x : 'Gal(E / K) :^ x = 'Gal(E / x @: K). +Proof. +without loss sKE: K / (K <= E)%VS. + move=> IH_K; rewrite gal_cap {}IH_K ?capvSr //. + transitivity 'Gal(E / x @: K :&: x @: E); last by rewrite limg_gal -gal_cap. + congr 'Gal(E / _); apply/eqP; rewrite eqEsubv limg_cap; apply/subvP=> a. + rewrite memv_cap => /andP[/memv_imgP[b Kb ->] /memv_imgP[c Ec] eq_bc]. + by rewrite memv_img // memv_cap Kb (lker0P (AEnd_lker0 _) _ _ eq_bc). +wlog suffices IHx: x K sKE / 'Gal(E / K) :^ x \subset 'Gal(E / x @: K). + apply/eqP; rewrite eqEsubset IHx // -sub_conjgV (subset_trans (IHx _ _ _)) //. + by apply/subvP=> _ /memv_imgP[a Ka ->]; rewrite memv_gal ?(subvP sKE). + rewrite -limg_comp (etrans (eq_in_limg _) (lim1g _)) // => a /(subvP sKE)Ka. + by rewrite !lfunE /= -galM // mulgV gal_id. +apply/subsetP=> _ /imsetP[y galEy ->]; rewrite gal_cap gal_kHom ?capvSr //=. +apply/kAHomP=> _ /memv_capP[/memv_imgP[a Ka ->] _]; have Ea := subvP sKE a Ka. +by rewrite -galM // -conjgC galM // (fixed_gal sKE galEy). +Qed. + +Definition fixedField V (A : {set gal_of V}) := + (V :&: \bigcap_(x in A) fixedSpace x)%VS. + +Lemma fixedFieldP E {A : {set gal_of E}} a : + a \in E -> reflect (forall x, x \in A -> x a = a) (a \in fixedField A). +Proof. +by rewrite memv_cap => ->; apply: (iffP subv_bigcapP) => cAa x /cAa/fixedSpaceP. +Qed. + +Lemma mem_fixedFieldP E (A : {set gal_of E}) a : + a \in fixedField A -> a \in E /\ (forall x, x \in A -> x a = a). +Proof. +by move=> fixAa; have [Ea _] := memv_capP fixAa; have:= fixedFieldP Ea fixAa. +Qed. + +Fact fixedField_is_aspace E (A : {set gal_of E}) : is_aspace (fixedField A). +Proof. +rewrite /fixedField; elim/big_rec: _ {1}E => [|x K _ IH_K] M. + exact: (valP (M :&: _)%AS). +by rewrite capvA IH_K. +Qed. +Canonical fixedField_aspace E A : {subfield L} := + ASpace (@fixedField_is_aspace E A). + +Lemma fixedField_bound E (A : {set gal_of E}) : (fixedField A <= E)%VS. +Proof. exact: capvSl. Qed. + +Lemma fixedFieldS E (A B : {set gal_of E}) : + A \subset B -> (fixedField B <= fixedField A)%VS. +Proof. +move/subsetP=> sAB; apply/subvP => a /mem_fixedFieldP[Ea cBa]. +by apply/fixedFieldP; last apply: sub_in1 cBa. +Qed. + +Lemma galois_connection_subv K E : + (K <= E)%VS -> (K <= fixedField ('Gal(E / K)))%VS. +Proof. +move=> sKE; apply/subvP => a Ka; have Ea := subvP sKE a Ka. +by apply/fixedFieldP=> // x galEx; apply: (fixed_gal sKE). +Qed. + +Lemma galois_connection_subset E (A : {set gal_of E}): + A \subset 'Gal(E / fixedField A). +Proof. +apply/subsetP => x Ax; rewrite gal_kAut ?capvSl // kAutE limg_gal subvv andbT. +by apply/kAHomP=> a /mem_fixedFieldP[_ ->]. +Qed. + +Lemma galois_connection K E (A : {set gal_of E}): + (K <= E)%VS -> (A \subset 'Gal(E / K)) = (K <= fixedField A)%VS. +Proof. +move=> sKE; apply/idP/idP => [/fixedFieldS | /(galS E)]. + by apply: subv_trans; apply galois_connection_subv. +by apply: subset_trans; apply: galois_connection_subset. +Qed. + +Definition galTrace U V a := \sum_(x in 'Gal(V / U)) (x a). + +Definition galNorm U V a := \prod_(x in 'Gal(V / U)) (x a). + +Section TraceAndNormMorphism. + +Variables U V : {vspace L}. + +Fact galTrace_is_additive : additive (galTrace U V). +Proof. +by move=> a b /=; rewrite -sumrB; apply: eq_bigr => x _; rewrite rmorphB. +Qed. +Canonical galTrace_additive := Additive galTrace_is_additive. + +Lemma galNorm1 : galNorm U V 1 = 1. +Proof. by apply: big1 => x _; rewrite rmorph1. Qed. + +Lemma galNormM : {morph galNorm U V : a b / a * b}. +Proof. +by move=> a b /=; rewrite -big_split; apply: eq_bigr => x _; rewrite rmorphM. +Qed. + +Lemma galNormV : {morph galNorm U V : a / a^-1}. +Proof. +by move=> a /=; rewrite -prodfV; apply: eq_bigr => x _; rewrite fmorphV. +Qed. + +Lemma galNormX n : {morph galNorm U V : a / a ^+ n}. +Proof. +move=> a; elim: n => [|n IHn]; first by apply: galNorm1. +by rewrite !exprS galNormM IHn. +Qed. + +Lemma galNorm_prod (I : Type) (r : seq I) (P : pred I) (B : I -> L) : + galNorm U V (\prod_(i <- r | P i) B i) + = \prod_(i <- r | P i) galNorm U V (B i). +Proof. exact: (big_morph _ galNormM galNorm1). Qed. + +Lemma galNorm0 : galNorm U V 0 = 0. +Proof. by rewrite /galNorm (bigD1 1%g) ?group1 // rmorph0 /= mul0r. Qed. + +Lemma galNorm_eq0 a : (galNorm U V a == 0) = (a == 0). +Proof. +apply/idP/eqP=> [/prodf_eq0[x _] | ->]; last by rewrite galNorm0. +by rewrite fmorph_eq0 => /eqP. +Qed. + +End TraceAndNormMorphism. + +Section TraceAndNormField. + +Variables K E : {subfield L}. + +Lemma galTrace_fixedField a : + a \in E -> galTrace K E a \in fixedField 'Gal(E / K). +Proof. +move=> Ea; apply/fixedFieldP=> [|x galEx]. + by apply: rpred_sum => x _; apply: memv_gal. +rewrite {2}/galTrace (reindex_acts 'R _ galEx) ?astabsR //=. +by rewrite rmorph_sum; apply: eq_bigr => y _; rewrite galM ?lfunE. +Qed. + +Lemma galTrace_gal a x : + a \in E -> x \in 'Gal(E / K) -> galTrace K E (x a) = galTrace K E a. +Proof. +move=> Ea galEx; rewrite {2}/galTrace (reindex_inj (mulgI x)). +by apply: eq_big => [b | b _]; rewrite ?groupMl // galM ?lfunE. +Qed. + +Lemma galNorm_fixedField a : + a \in E -> galNorm K E a \in fixedField 'Gal(E / K). +Proof. +move=> Ea; apply/fixedFieldP=> [|x galEx]. + by apply: rpred_prod => x _; apply: memv_gal. +rewrite {2}/galNorm (reindex_acts 'R _ galEx) ?astabsR //=. +by rewrite rmorph_prod; apply: eq_bigr => y _; rewrite galM ?lfunE. +Qed. + +Lemma galNorm_gal a x : + a \in E -> x \in 'Gal(E / K) -> galNorm K E (x a) = galNorm K E a. +Proof. +move=> Ea galEx; rewrite {2}/galNorm (reindex_inj (mulgI x)). +by apply: eq_big => [b | b _]; rewrite ?groupMl // galM ?lfunE. +Qed. + +End TraceAndNormField. + +Definition normalField U V := [forall x in kAEndf U, x @: V == V]%VS. + +Lemma normalField_kAut K M E f : + (K <= M <= E)%VS -> normalField K M -> kAut K E f -> kAut K M f. +Proof. +case/andP=> sKM sME nKM /kAut_to_gal[x galEx /(sub_in1 (subvP sME))Df]. +have sKE := subv_trans sKM sME; rewrite gal_kHom // in galEx. +rewrite (kAut_eq sKM Df) /kAut (kHomSr sME) //= (forall_inP nKM) // inE. +by rewrite kAutfE; apply/kAHomP; apply: (kAHomP galEx). +Qed. + +Lemma normalFieldP K E : + reflect {in E, forall a, exists2 r, + all (mem E) r & minPoly K a = \prod_(b <- r) ('X - b%:P)} + (normalField K E). +Proof. +apply: (iffP eqfun_inP) => [nKE a Ea | nKE x]; last first. + rewrite inE kAutfE => homKx; suffices: kAut K E x by case/andP=> _ /eqP. + rewrite kAutE (kHomSr (subvf E)) //=; apply/subvP=> _ /memv_imgP[a Ea ->]. + have [r /allP/=srE splitEa] := nKE a Ea. + rewrite srE // -root_prod_XsubC -splitEa. + by rewrite -(kHom_poly_id homKx (minPolyOver K a)) fmorph_root root_minPoly. +have [r /eqP splitKa] := splitting_field_normal K a. +exists r => //; apply/allP => b; rewrite -root_prod_XsubC -splitKa => pKa_b_0. +pose y := kHomExtend K \1 a b; have [hom1K lf1p] := (kHom1 K K, lfun1_poly). +have homKy: kHom K <<K; a>> y by apply/kHomExtendP; rewrite ?lf1p. +have [[g Dy] [_ idKy]] := (kHom_to_AEnd homKy, kHomP homKy). +have <-: g a = b by rewrite -Dy ?memv_adjoin // (kHomExtend_val hom1K) ?lf1p. +suffices /nKE <-: g \in kAEndf K by apply: memv_img. +by rewrite inE kAutfE; apply/kAHomP=> c Kc; rewrite -Dy ?subvP_adjoin ?idKy. +Qed. + +Lemma normalFieldf K : normalField K {:L}. +Proof. +apply/normalFieldP=> a _; have [r /eqP->] := splitting_field_normal K a. +by exists r => //; apply/allP=> b; rewrite /= memvf. +Qed. + +Lemma normalFieldS K M E : (K <= M)%VS -> normalField K E -> normalField M E. +Proof. +move=> sKM /normalFieldP nKE; apply/normalFieldP=> a Ea. +have [r /allP Er splitKa] := nKE a Ea. +have /dvdp_prod_XsubC[m splitMa]: minPoly M a %| \prod_(b <- r) ('X - b%:P). + by rewrite -splitKa minPolyS. +exists (mask m r); first by apply/allP=> b /mem_mask/Er. +by apply/eqP; rewrite -eqp_monic ?monic_prod_XsubC ?monic_minPoly. +Qed. + +Lemma splitting_normalField E K : + (K <= E)%VS -> + reflect (exists2 p, p \is a polyOver K & splittingFieldFor K p E) + (normalField K E). +Proof. +move=> sKE; apply: (iffP idP) => [nKE| [p Kp [rs Dp defE]]]; last first. + apply/forall_inP=> g; rewrite inE kAutE => /andP[homKg _]. + rewrite -dimv_leqif_eq ?limg_dim_eq ?(eqP (AEnd_lker0 g)) ?capv0 //. + rewrite -defE aimg_adjoin_seq; have [_ /fixedSpace_limg->] := andP homKg. + apply/adjoin_seqSr=> _ /mapP[a rs_a ->]. + rewrite -!root_prod_XsubC -!(eqp_root Dp) in rs_a *. + by apply: kHom_root_id homKg Kp _ rs_a; rewrite ?subvf ?memvf. +pose splitK a r := minPoly K a = \prod_(b <- r) ('X - b%:P). +have{nKE} rK_ a: {r | a \in E -> all (mem E) r /\ splitK a r}. + case Ea: (a \in E); last by exists [::]. + by have /sig2_eqW[r] := normalFieldP _ _ nKE a Ea; exists r. +have sXE := basis_mem (vbasisP E); set X : seq L := vbasis E in sXE. +exists (\prod_(a <- X) minPoly K a). + by apply: rpred_prod => a _; apply: minPolyOver. +exists (flatten [seq (sval (rK_ a)) | a <- X]). + move/allP: sXE; elim: X => [|a X IHX] ; first by rewrite !big_nil eqpxx. + rewrite big_cons /= big_cat /= => /andP[Ea sXE]. + by case: (rK_ a) => /= r [] // _ <-; apply/eqp_mull/IHX. +apply/eqP; rewrite eqEsubv; apply/andP; split. + apply/Fadjoin_seqP; split=> // b /flatten_mapP[a /sXE Ea]. + by apply/allP; case: rK_ => r /= []. +rewrite -{1}(span_basis (vbasisP E)); apply/span_subvP=> a Xa. +apply/seqv_sub_adjoin/flatten_mapP; exists a => //; rewrite -root_prod_XsubC. +by case: rK_ => /= r [| _ <-]; rewrite ?sXE ?root_minPoly. +Qed. + +Lemma kHom_to_gal K M E f : + (K <= M <= E)%VS -> normalField K E -> kHom K M f -> + {x | x \in 'Gal(E / K) & {in M, f =1 x}}. +Proof. +case/andP=> /subvP sKM /subvP sME nKE KhomMf. +have [[g Df] [_ idKf]] := (kHom_to_AEnd KhomMf, kHomP KhomMf). +suffices /kAut_to_gal[x galEx Dg]: kAut K E g. + by exists x => //= a Ma; rewrite Df // Dg ?sME. +have homKg: kHom K {:L} g by apply/kAHomP=> a Ka; rewrite -Df ?sKM ?idKf. +by rewrite /kAut (kHomSr (subvf _)) // (forall_inP nKE) // inE kAutfE. +Qed. + +Lemma normalField_root_minPoly K E a b : + (K <= E)%VS -> normalField K E -> a \in E -> root (minPoly K a) b -> + exists2 x, x \in 'Gal(E / K) & x a = b. +Proof. +move=> sKE nKE Ea pKa_b_0; pose f := kHomExtend K \1 a b. +have homKa_f: kHom K <<K; a>> f. + by apply: kHomExtendP; rewrite ?kHom1 ?lfun1_poly. +have sK_Ka_E: (K <= <<K; a>> <= E)%VS. + by rewrite subv_adjoin; apply/FadjoinP; rewrite sKE Ea. +have [x galEx Df] := kHom_to_gal sK_Ka_E nKE homKa_f; exists x => //. +by rewrite -Df ?memv_adjoin // (kHomExtend_val (kHom1 K K)) ?lfun1_poly. +Qed. + +Implicit Arguments normalFieldP [K E]. + +Lemma normalField_factors K E : + (K <= E)%VS -> + reflect {in E, forall a, exists2 r : seq (gal_of E), + r \subset 'Gal(E / K) + & minPoly K a = \prod_(x <- r) ('X - (x a)%:P)} + (normalField K E). +Proof. +move=> sKE; apply: (iffP idP) => [nKE a Ea | nKE]; last first. + apply/normalFieldP=> a Ea; have [r _ ->] := nKE a Ea. + exists [seq x a | x : gal_of E <- r]; last by rewrite big_map. + by rewrite all_map; apply/allP=> b _; apply: memv_gal. +have [r Er splitKa] := normalFieldP nKE a Ea. +pose f b := [pick x in 'Gal(E / K) | x a == b]. +exists (pmap f r). + apply/subsetP=> x; rewrite mem_pmap /f => /mapP[b _]. + by case: (pickP _) => // c /andP[galEc _] [->]. +rewrite splitKa; have{splitKa}: all (root (minPoly K a)) r. + by apply/allP => b; rewrite splitKa root_prod_XsubC. +elim: r Er => /= [|b r IHr]; first by rewrite !big_nil. +case/andP=> Eb Er /andP[pKa_b_0 /(IHr Er){IHr Er}IHr]. +have [x galE /eqP xa_b] := normalField_root_minPoly sKE nKE Ea pKa_b_0. +rewrite /(f b); case: (pickP _) => [y /andP[_ /eqP<-]|/(_ x)/andP[]//]. +by rewrite !big_cons IHr. +Qed. + +Definition galois U V := [&& (U <= V)%VS, separable U V & normalField U V]. + +Lemma galoisS K M E : (K <= M <= E)%VS -> galois K E -> galois M E. +Proof. +case/andP=> sKM sME /and3P[_ sepUV nUV]. +by rewrite /galois sME (separableSl sKM) ?(normalFieldS sKM). +Qed. + +Lemma galois_dim K E : galois K E -> \dim_K E = #|'Gal(E / K)|. +Proof. +case/and3P=> sKE /eq_adjoin_separable_generator-> // nKE. +set a := separable_generator K E in nKE *. +have [r /allP/=Er splitKa] := normalFieldP nKE a (memv_adjoin K a). +rewrite (dim_sup_field (subv_adjoin K a)) mulnK ?adim_gt0 //. +apply/eqP; rewrite -eqSS -adjoin_degreeE -size_minPoly splitKa size_prod_XsubC. +set n := size r; rewrite eqSS -[n]card_ord. +have x_ (i : 'I_n): {x | x \in 'Gal(<<K; a>> / K) & x a = r`_i}. + apply/sig2_eqW/normalField_root_minPoly; rewrite ?subv_adjoin ?memv_adjoin //. + by rewrite splitKa root_prod_XsubC mem_nth. +have /card_image <-: injective (fun i => s2val (x_ i)). + move=> i j /eqP; case: (x_ i) (x_ j) => y /= galEy Dya [z /= galEx Dza]. + rewrite gal_adjoin_eq // Dya Dza nth_uniq // => [/(i =P j)//|]. + by rewrite -separable_prod_XsubC -splitKa; apply: separable_generatorP. +apply/eqP/eq_card=> x; apply/codomP/idP=> [[i ->] | galEx]; first by case: x_. +have /(nthP 0) [i ltin Dxa]: x a \in r. + rewrite -root_prod_XsubC -splitKa. + by rewrite root_minPoly_gal ?memv_adjoin ?subv_adjoin. +exists (Ordinal ltin); apply/esym/eqP. +by case: x_ => y /= galEy /eqP; rewrite Dxa gal_adjoin_eq. +Qed. + +Lemma galois_factors K E : + (K <= E)%VS -> + reflect {in E, forall a, exists r, let r_a := [seq x a | x : gal_of E <- r] in + [/\ r \subset 'Gal(E / K), uniq r_a + & minPoly K a = \prod_(b <- r_a) ('X - b%:P)]} + (galois K E). +Proof. +move=> sKE; apply: (iffP and3P) => [[_ sepKE nKE] a Ea | galKE]. + have [r galEr splitEa] := normalField_factors sKE nKE a Ea. + exists r; rewrite /= -separable_prod_XsubC !big_map -splitEa. + by split=> //; apply: separableP Ea. +split=> //. + apply/separableP => a /galKE[r [_ Ur_a splitKa]]. + by rewrite /separable_element splitKa separable_prod_XsubC. +apply/(normalField_factors sKE)=> a /galKE[r [galEr _ ->]]. +by rewrite big_map; exists r. +Qed. + +Lemma splitting_galoisField K E : + reflect (exists p, [/\ p \is a polyOver K, separable_poly p + & splittingFieldFor K p E]) + (galois K E). +Proof. +apply: (iffP and3P) => [[sKE sepKE nKE]|[p [Kp sep_p [r Dp defE]]]]. + rewrite (eq_adjoin_separable_generator sepKE) // in nKE *. + set a := separable_generator K E in nKE *; exists (minPoly K a). + split; first 1 [exact: minPolyOver | exact/separable_generatorP]. + have [r /= /allP Er splitKa] := normalFieldP nKE a (memv_adjoin _ _). + exists r; first by rewrite splitKa eqpxx. + apply/eqP; rewrite eqEsubv; apply/andP; split. + by apply/Fadjoin_seqP; split => //; apply: subv_adjoin. + apply/FadjoinP; split; first exact: subv_adjoin_seq. + by rewrite seqv_sub_adjoin // -root_prod_XsubC -splitKa root_minPoly. +have sKE: (K <= E)%VS by rewrite -defE subv_adjoin_seq. +split=> //; last by apply/splitting_normalField=> //; exists p; last exists r. +rewrite -defE; apply/separable_Fadjoin_seq/allP=> a r_a. +by apply/separable_elementP; exists p; rewrite (eqp_root Dp) root_prod_XsubC. +Qed. + +Lemma galois_fixedField K E : + reflect (fixedField 'Gal(E / K) = K) (galois K E). +Proof. +apply (iffP idP) => [/and3P[sKE /separableP sepKE nKE] | fixedKE]. + apply/eqP; rewrite eqEsubv galois_connection_subv ?andbT //. + apply/subvP=> a /mem_fixedFieldP[Ea fixEa]; rewrite -adjoin_deg_eq1. + have [r /allP Er splitKa] := normalFieldP nKE a Ea. + rewrite -eqSS -size_minPoly splitKa size_prod_XsubC eqSS -/(size [:: a]). + have Ur: uniq r by rewrite -separable_prod_XsubC -splitKa; apply: sepKE. + rewrite -uniq_size_uniq {Ur}// => b; rewrite inE -root_prod_XsubC -splitKa. + apply/eqP/idP=> [-> | pKa_b_0]; first exact: root_minPoly. + by have [x /fixEa-> ->] := normalField_root_minPoly sKE nKE Ea pKa_b_0. +have sKE: (K <= E)%VS by rewrite -fixedKE capvSl. +apply/galois_factors=> // a Ea. +pose r_pKa := [seq x a | x : gal_of E in 'Gal(E / K)]. +have /fin_all_exists2[x_ galEx_ Dx_a] (b : seq_sub r_pKa) := imageP (valP b). +exists (codom x_); rewrite -map_comp; set r := map _ _. +have r_xa x: x \in 'Gal(E / K) -> x a \in r. + move=> galEx; have r_pKa_xa: x a \in r_pKa by apply/imageP; exists x. + by rewrite [x a](Dx_a (SeqSub r_pKa_xa)); apply: codom_f. +have Ur: uniq r by apply/injectiveP=> b c /=; rewrite -!Dx_a => /val_inj. +split=> //; first by apply/subsetP=> _ /codomP[b ->]. +apply/eqP; rewrite -eqp_monic ?monic_minPoly ?monic_prod_XsubC //. +apply/andP; split; last first. + rewrite uniq_roots_dvdp ?uniq_rootsE // all_map. + by apply/allP=> b _ /=; rewrite root_minPoly_gal. +apply: minPoly_dvdp; last by rewrite root_prod_XsubC -(gal_id E a) r_xa ?group1. +rewrite -fixedKE; apply/polyOverP => i; apply/fixedFieldP=> [|x galEx]. + rewrite (polyOverP _) // big_map rpred_prod // => b _. + by rewrite polyOverXsubC memv_gal. +rewrite -coef_map rmorph_prod; congr (_ : {poly _})`_i. +symmetry; rewrite (eq_big_perm (map x r)) /= ?(big_map x). + by apply: eq_bigr => b _; rewrite rmorphB /= map_polyX map_polyC. +have Uxr: uniq (map x r) by rewrite map_inj_uniq //; apply: fmorph_inj. +have /leq_size_perm: {subset map x r <= r}. + by rewrite -map_comp => _ /codomP[b ->] /=; rewrite -galM // r_xa ?groupM. +by rewrite (size_map x) perm_eq_sym; case=> // /uniq_perm_eq->. +Qed. + +Lemma mem_galTrace K E a : galois K E -> a \in E -> galTrace K E a \in K. +Proof. by move/galois_fixedField => {2}<- /galTrace_fixedField. Qed. + +Lemma mem_galNorm K E a : galois K E -> a \in E -> galNorm K E a \in K. +Proof. by move/galois_fixedField=> {2}<- /galNorm_fixedField. Qed. + +Lemma gal_independent_contra E (P : pred (gal_of E)) (c_ : gal_of E -> L) x : + P x -> c_ x != 0 -> + exists2 a, a \in E & \sum_(y | P y) c_ y * y a != 0. +Proof. +elim: {P}_.+1 c_ x {-2}P (ltnSn #|P|) => // n IHn c_ x P lePn Px nz_cx. +rewrite ltnS (cardD1x Px) in lePn; move/IHn: lePn => {n IHn}/=IH_P. +have [/eqfun_inP c_Px'_0 | ] := boolP [forall (y | P y && (y != x)), c_ y == 0]. + exists 1; rewrite ?mem1v // (bigD1 x Px) /= rmorph1 mulr1. + by rewrite big1 ?addr0 // => y /c_Px'_0->; rewrite mul0r. +rewrite negb_forall_in => /exists_inP[y Px'y nz_cy]. +have [Py /gal_eqP/eqlfun_inP/subvPn[a Ea]] := andP Px'y. +rewrite memv_ker !lfun_simp => nz_yxa; pose d_ y := c_ y * (y a - x a). +have /IH_P[//|b Eb nz_sumb]: d_ y != 0 by rewrite mulf_neq0. +have [sumb_0|] := eqVneq (\sum_(z | P z) c_ z * z b) 0; last by exists b. +exists (a * b); first exact: rpredM. +rewrite -subr_eq0 -[z in _ - z](mulr0 (x a)) -[in z in _ - z]sumb_0. +rewrite mulr_sumr -sumrB (bigD1 x Px) rmorphM /= mulrCA subrr add0r. +congr (_ != 0): nz_sumb; apply: eq_bigr => z _. +by rewrite mulrCA rmorphM -mulrBr -mulrBl mulrA. +Qed. + +Lemma gal_independent E (P : pred (gal_of E)) (c_ : gal_of E -> L) : + (forall a, a \in E -> \sum_(x | P x) c_ x * x a = 0) -> + (forall x, P x -> c_ x = 0). +Proof. +move=> sum_cP_0 x Px; apply/eqP/idPn=> /(gal_independent_contra Px)[a Ea]. +by rewrite sum_cP_0 ?eqxx. +Qed. + +Lemma Hilbert's_theorem_90 K E x a : + generator 'Gal(E / K) x -> a \in E -> + reflect (exists2 b, b \in E /\ b != 0 & a = b / x b) (galNorm K E a == 1). +Proof. +move/(_ =P <[x]>)=> DgalE Ea. +have galEx: x \in 'Gal(E / K) by rewrite DgalE cycle_id. +apply: (iffP eqP) => [normEa1 | [b [Eb nzb] ->]]; last first. + by rewrite galNormM galNormV galNorm_gal // mulfV // galNorm_eq0. +have [x1 | ntx] := eqVneq x 1%g. + exists 1; first by rewrite mem1v oner_neq0. + by rewrite -{1}normEa1 /galNorm DgalE x1 cycle1 big_set1 !gal_id divr1. +pose c_ y := \prod_(i < invm (injm_Zpm x) y) (x ^+ i)%g a. +have nz_c1: c_ 1%g != 0 by rewrite /c_ morph1 big_ord0 oner_neq0. +have [d] := @gal_independent_contra _ (mem 'Gal(E / K)) _ _ (group1 _) nz_c1. +set b := \sum_(y in _) _ => Ed nz_b; exists b. + split=> //; apply: rpred_sum => y galEy. + by apply: rpredM; first apply: rpred_prod => i _; apply: memv_gal. +apply: canRL (mulfK _) _; first by rewrite fmorph_eq0. +rewrite rmorph_sum mulr_sumr [b](reindex_acts 'R _ galEx) ?astabsR //=. +apply: eq_bigr => y galEy; rewrite galM // rmorphM mulrA; congr (_ * _). +have /morphimP[/= i _ _ ->] /=: y \in Zpm @* Zp #[x] by rewrite im_Zpm -DgalE. +have <-: Zpm (i + 1) = (Zpm i * x)%g by rewrite morphM ?mem_Zp ?order_gt1. +rewrite /c_ !invmE ?mem_Zp ?order_gt1 //= addn1; set n := _.+2. +transitivity (\prod_(j < i.+1) (x ^+ j)%g a). + rewrite big_ord_recl gal_id rmorph_prod; congr (_ * _). + by apply: eq_bigr => j _; rewrite expgSr galM ?lfunE. +have [/modn_small->//||->] := ltngtP i.+1 n; first by rewrite ltnNge ltn_ord. +rewrite modnn big_ord0; apply: etrans normEa1; rewrite /galNorm DgalE -im_Zpm. +rewrite morphimEdom big_imset /=; last exact/injmP/injm_Zpm. +by apply: eq_bigl => j /=; rewrite mem_Zp ?order_gt1. +Qed. + +Section Matrix. + +Variable (E : {subfield L}) (A : {set gal_of E}). + +Let K := fixedField A. + +Lemma gal_matrix : + {w : #|A|.-tuple L | {subset w <= E} /\ 0 \notin w & + [/\ \matrix_(i, j < #|A|) enum_val i (tnth w j) \in unitmx, + directv (\sum_i K * <[tnth w i]>) & + group_set A -> (\sum_i K * <[tnth w i]>)%VS = E] }. +Proof. +pose nzE (w : #|A|.-tuple L) := {subset w <= E} /\ 0 \notin w. +pose M w := \matrix_(i, j < #|A|) nth 1%g (enum A) i (tnth w j). +have [w [Ew nzw] uM]: {w : #|A|.-tuple L | nzE w & M w \in unitmx}. + rewrite {}/nzE {}/M cardE; have: uniq (enum A) := enum_uniq _. + elim: (enum A) => [|x s IHs] Uxs. + by exists [tuple]; rewrite // flatmx0 -(flatmx0 1%:M) unitmx1. + have [s'x Us]: x \notin s /\ uniq s by apply/andP. + have{IHs} [w [Ew nzw] uM] := IHs Us; set M := \matrix_(i, j) _ in uM. + pose a := \row_i x (tnth w i) *m invmx M. + pose c_ y := oapp (a 0) (-1) (insub (index y s)). + have cx_n1 : c_ x = -1 by rewrite /c_ insubN ?index_mem. + have nz_cx : c_ x != 0 by rewrite cx_n1 oppr_eq0 oner_neq0. + have Px: [pred y in x :: s] x := mem_head x s. + have{Px nz_cx} /sig2W[w0 Ew0 nzS] := gal_independent_contra Px nz_cx. + exists [tuple of cons w0 w]. + split; first by apply/allP; rewrite /= Ew0; apply/allP. + rewrite inE negb_or (contraNneq _ nzS) // => <-. + by rewrite big1 // => y _; rewrite rmorph0 mulr0. + rewrite unitmxE -[\det _]mul1r; set M1 := \matrix_(i, j < 1 + size s) _. + have <-: \det (block_mx 1 (- a) 0 1%:M) = 1 by rewrite det_ublock !det1 mulr1. + rewrite -det_mulmx -[M1]submxK mulmx_block !mul0mx !mul1mx !add0r !mulNmx. + have ->: drsubmx M1 = M by apply/matrixP => i j; rewrite !mxE !(tnth_nth 0). + have ->: ursubmx M1 - a *m M = 0. + by apply/rowP=> i; rewrite mulmxKV // !mxE !(tnth_nth 0) subrr. + rewrite det_lblock unitrM andbC -unitmxE uM unitfE -oppr_eq0. + congr (_ != 0): nzS; rewrite [_ - _]mx11_scalar det_scalar !mxE opprB /=. + rewrite -big_uniq // big_cons /= cx_n1 mulN1r addrC; congr (_ + _). + rewrite (big_nth 1%g) big_mkord; apply: eq_bigr => j _. + by rewrite /c_ index_uniq // valK; congr (_ * _); rewrite !mxE. +exists w => [//|]; split=> [||gA]. +- by congr (_ \in unitmx): uM; apply/matrixP=> i j; rewrite !mxE -enum_val_nth. +- apply/directv_sum_independent=> kw_ Kw_kw sum_kw_0 j _. + have /fin_all_exists2[k_ Kk_ Dk_] i := memv_cosetP (Kw_kw i isT). + pose kv := \col_i k_ i. + transitivity (kv j 0 * tnth w j); first by rewrite !mxE. + suffices{j}/(canRL (mulKmx uM))->: M w *m kv = 0 by rewrite mulmx0 mxE mul0r. + apply/colP=> i; rewrite !mxE; pose Ai := nth 1%g (enum A) i. + transitivity (Ai (\sum_j kw_ j)); last by rewrite sum_kw_0 rmorph0. + rewrite rmorph_sum; apply: eq_bigr => j _; rewrite !mxE /= -/Ai. + rewrite Dk_ mulrC rmorphM /=; congr (_ * _). + by have /mem_fixedFieldP[_ -> //] := Kk_ j; rewrite -mem_enum mem_nth -?cardE. +pose G := group gA; have G_1 := group1 G; pose iG := enum_rank_in G_1. +apply/eqP; rewrite eqEsubv; apply/andP; split. + apply/subv_sumP=> i _; apply: subv_trans (asubv _). + by rewrite prodvS ?capvSl // -memvE Ew ?mem_tnth. +apply/subvP=> w0 Ew0; apply/memv_sumP. +pose wv := \col_(i < #|A|) enum_val i w0; pose v := invmx (M w) *m wv. +exists (fun i => tnth w i * v i 0) => [i _|]; last first. + transitivity (wv (iG 1%g) 0); first by rewrite mxE enum_rankK_in ?gal_id. + rewrite -[wv](mulKVmx uM) -/v; rewrite mxE; apply: eq_bigr => i _. + by congr (_ * _); rewrite !mxE -enum_val_nth enum_rankK_in ?gal_id. +rewrite mulrC memv_mul ?memv_line //; apply/fixedFieldP=> [|x Gx]. + rewrite mxE rpred_sum // => j _; rewrite !mxE rpredM //; last exact: memv_gal. + have E_M k l: M w k l \in E by rewrite mxE memv_gal // Ew ?mem_tnth. + have Edet n (N : 'M_n) (E_N : forall i j, N i j \in E): \det N \in E. + by apply: rpred_sum => sigma _; rewrite rpredMsign rpred_prod. + rewrite /invmx uM 2!mxE mulrC rpred_div ?Edet //. + by rewrite rpredMsign Edet // => k l; rewrite 2!mxE. +suffices{i} {2}<-: map_mx x v = v by rewrite [map_mx x v i 0]mxE. +have uMx: map_mx x (M w) \in unitmx by rewrite map_unitmx. +rewrite map_mxM map_invmx /=; apply: canLR {uMx}(mulKmx uMx) _. +apply/colP=> i; rewrite !mxE; pose ix := iG (enum_val i * x)%g. +have Dix b: b \in E -> enum_val ix b = x (enum_val i b). + by move=> Eb; rewrite enum_rankK_in ?groupM ?enum_valP // galM ?lfunE. +transitivity ((M w *m v) ix 0); first by rewrite mulKVmx // mxE Dix. +rewrite mxE; apply: eq_bigr => j _; congr (_ * _). +by rewrite !mxE -!enum_val_nth Dix // ?Ew ?mem_tnth. +Qed. + +End Matrix. + +Lemma dim_fixedField E (G : {group gal_of E}) : #|G| = \dim_(fixedField G) E. +Proof. +have [w [_ nzw] [_ Edirect /(_ (groupP G))defE]] := gal_matrix G. +set n := #|G|; set m := \dim (fixedField G); rewrite -defE (directvP Edirect). +rewrite -[n]card_ord -(@mulnK #|'I_n| m) ?adim_gt0 //= -sum_nat_const. +congr (_ %/ _)%N; apply: eq_bigr => i _. +by rewrite dim_cosetv ?(memPn nzw) ?mem_tnth. +Qed. + +Lemma dim_fixed_galois K E (G : {group gal_of E}) : + galois K E -> G \subset 'Gal(E / K) -> + \dim_K (fixedField G) = #|'Gal(E / K) : G|. +Proof. +move=> galE sGgal; have [sFE _ _] := and3P galE; apply/eqP. +rewrite -divgS // eqn_div ?cardSg // dim_fixedField -galois_dim //. +by rewrite mulnC muln_divA ?divnK ?field_dimS ?capvSl -?galois_connection. +Qed. + +Lemma gal_fixedField E (G : {group gal_of E}): 'Gal(E / fixedField G) = G. +Proof. +apply/esym/eqP; rewrite eqEcard galois_connection_subset /= (dim_fixedField G). +rewrite galois_dim //; apply/galois_fixedField/eqP. +rewrite eqEsubv galois_connection_subv ?capvSl //. +by rewrite fixedFieldS ?galois_connection_subset. +Qed. + +Lemma gal_generated E (A : {set gal_of E}) : 'Gal(E / fixedField A) = <<A>>. +Proof. +apply/eqP; rewrite eqEsubset gen_subG galois_connection_subset. +by rewrite -[<<A>>]gal_fixedField galS // fixedFieldS // subset_gen. +Qed. + +Lemma fixedField_galois E (A : {set gal_of E}): galois (fixedField A) E. +Proof. +have: galois (fixedField <<A>>) E. + by apply/galois_fixedField; rewrite gal_fixedField. +by apply: galoisS; rewrite capvSl fixedFieldS // subset_gen. +Qed. + +Section FundamentalTheoremOfGaloisTheory. + +Variables E K : {subfield L}. +Hypothesis galKE : galois K E. + +Section IntermediateField. + +Variable M : {subfield L}. +Hypothesis (sKME : (K <= M <= E)%VS) (nKM : normalField K M). + +Lemma normalField_galois : galois K M. +Proof. +have [[sKM sME] [_ sepKE nKE]] := (andP sKME, and3P galKE). +by rewrite /galois sKM (separableSr sME). +Qed. + +Definition normalField_cast (x : gal_of E) : gal_of M := gal M x. + +Lemma normalField_cast_eq x : + x \in 'Gal(E / K) -> {in M, normalField_cast x =1 x}. +Proof. +have [sKM sME] := andP sKME; have sKE := subv_trans sKM sME. +rewrite gal_kAut // => /(normalField_kAut sKME nKM). +by rewrite kAutE => /andP[_ /galK]. +Qed. + +Lemma normalField_castM : + {in 'Gal(E / K) &, {morph normalField_cast : x y / (x * y)%g}}. +Proof. +move=> x y galEx galEy /=; apply/eqP/gal_eqP => a Ma. +have Ea: a \in E by have [_ /subvP->] := andP sKME. +rewrite normalField_cast_eq ?groupM ?galM //=. +by rewrite normalField_cast_eq ?memv_gal // normalField_cast_eq. +Qed. +Canonical normalField_cast_morphism := Morphism normalField_castM. + +Lemma normalField_ker : 'ker normalField_cast = 'Gal(E / M). +Proof. +have [sKM sME] := andP sKME. +apply/setP=> x; apply/idP/idP=> [kerMx | galEMx]. + rewrite gal_kHom //; apply/kAHomP=> a Ma. + by rewrite -normalField_cast_eq ?(dom_ker kerMx) // (mker kerMx) gal_id. +have galEM: x \in 'Gal(E / K) := subsetP (galS E sKM) x galEMx. +apply/kerP=> //; apply/eqP/gal_eqP=> a Ma. +by rewrite normalField_cast_eq // gal_id (fixed_gal sME). +Qed. + +Lemma normalField_normal : 'Gal(E / M) <| 'Gal(E / K). +Proof. by rewrite -normalField_ker ker_normal. Qed. + +Lemma normalField_img : normalField_cast @* 'Gal(E / K) = 'Gal(M / K). +Proof. +have [[sKM sME] [sKE _ nKE]] := (andP sKME, and3P galKE). +apply/setP=> x; apply/idP/idP=> [/morphimP[{x}x galEx _ ->] | galMx]. + rewrite gal_kHom //; apply/kAHomP=> a Ka; have Ma := subvP sKM a Ka. + by rewrite normalField_cast_eq // (fixed_gal sKE). +have /(kHom_to_gal sKME nKE)[y galEy eq_xy]: kHom K M x by rewrite -gal_kHom. +apply/morphimP; exists y => //; apply/eqP/gal_eqP => a Ha. +by rewrite normalField_cast_eq // eq_xy. +Qed. + +Lemma normalField_isom : + {f : {morphism ('Gal(E / K) / 'Gal(E / M)) >-> gal_of M} | + isom ('Gal(E / K) / 'Gal (E / M)) 'Gal(M / K) f + & (forall A, f @* (A / 'Gal(E / M)) = normalField_cast @* A) + /\ {in 'Gal(E / K) & M, forall x, f (coset 'Gal (E / M) x) =1 x} }%g. +Proof. +have:= first_isom normalField_cast_morphism; rewrite normalField_ker. +case=> f injf Df; exists f; first by apply/isomP; rewrite Df normalField_img. +split=> [//|x a galEx /normalField_cast_eq<- //]; congr ((_ : gal_of M) a). +apply: set1_inj; rewrite -!morphim_set1 ?mem_quotient ?Df //. +by rewrite (subsetP (normal_norm normalField_normal)). +Qed. + +Lemma normalField_isog : 'Gal(E / K) / 'Gal(E / M) \isog 'Gal(M / K). +Proof. by rewrite -normalField_ker -normalField_img first_isog. Qed. + +End IntermediateField. + +Section IntermediateGroup. + +Variable G : {group gal_of E}. +Hypothesis nsGgalE : G <| 'Gal(E / K). + +Lemma normal_fixedField_galois : galois K (fixedField G). +Proof. +have [[sKE sepKE nKE] [sGgal nGgal]] := (and3P galKE, andP nsGgalE). +rewrite /galois -(galois_connection _ sKE) sGgal. +rewrite (separableSr _ sepKE) ?capvSl //; apply/forall_inP=> f autKf. +rewrite eqEdim limg_dim_eq ?(eqP (AEnd_lker0 _)) ?capv0 // leqnn andbT. +apply/subvP => _ /memv_imgP[a /mem_fixedFieldP[Ea cGa] ->]. +have /kAut_to_gal[x galEx -> //]: kAut K E f. + rewrite /kAut (forall_inP nKE) // andbT; apply/kAHomP. + by move: autKf; rewrite inE kAutfE => /kHomP[]. +apply/fixedFieldP=> [|y Gy]; first exact: memv_gal. +by rewrite -galM // conjgCV galM //= cGa // memJ_norm ?groupV ?(subsetP nGgal). +Qed. + +End IntermediateGroup. + +End FundamentalTheoremOfGaloisTheory. + +End GaloisTheory. + +Notation "''Gal' ( V / U )" := (galoisG V U) : group_scope. +Notation "''Gal' ( V / U )" := (galoisG_group V U) : Group_scope. + +Implicit Arguments fixedFieldP [F L E A a]. +Implicit Arguments normalFieldP [F L K E]. +Implicit Arguments splitting_galoisField [F L K E]. +Implicit Arguments galois_fixedField [F L K E]. diff --git a/mathcomp/field/separable.v b/mathcomp/field/separable.v new file mode 100644 index 0000000..9638500 --- /dev/null +++ b/mathcomp/field/separable.v @@ -0,0 +1,995 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div choice fintype. +Require Import tuple finfun bigop finset prime binomial ssralg poly polydiv. +Require Import fingroup perm morphism quotient gproduct finalg zmodp cyclic. +Require Import matrix mxalgebra mxpoly polyXY vector falgebra fieldext. + +(******************************************************************************) +(* This file provides a theory of separable and inseparable field extensions. *) +(* *) +(* separable_poly p <=> p has no multiple roots in any field extension. *) +(* separable_element K x <=> the minimal polynomial of x over K is separable. *) +(* separable K E <=> every member of E is separable over K. *) +(* separable_generator K E == some x \in E that generates the largest *) +(* subfield K[x] that is separable over K. *) +(* purely_inseparable_element K x <=> there is a [char L].-nat n such that *) +(* x ^+ n \in K. *) +(* purely_inseparable K E <=> every member of E is purely inseparable over K. *) +(* *) +(* Derivations are introduced to prove the adjoin_separableP Lemma: *) +(* Derivation K D <=> the linear operator D satifies the Leibniz *) +(* product rule inside K. *) +(* extendDerivation x D K == given a derivation D on K and a separable *) +(* element x over K, this function returns the *) +(* unique extension of D to K(x). *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Open Local Scope ring_scope. +Import GRing.Theory. + +Section SeparablePoly. + +Variable R : idomainType. +Implicit Types p q d u v : {poly R}. + +Definition separable_poly p := coprimep p p^`(). + +Local Notation separable := separable_poly. +Local Notation lcn_neq0 := (Pdiv.Idomain.lc_expn_scalp_neq0 _). + +Lemma separable_poly_neq0 p : separable p -> p != 0. +Proof. +by apply: contraTneq => ->; rewrite /separable deriv0 coprime0p eqp01. +Qed. + +Lemma poly_square_freeP p : + (forall u v, u * v %| p -> coprimep u v) + <-> (forall u, size u != 1%N -> ~~ (u ^+ 2 %| p)). +Proof. +split=> [sq'p u | sq'p u v dvd_uv_p]. + by apply: contra => /sq'p; rewrite coprimepp. +rewrite coprimep_def (contraLR (sq'p _)) // (dvdp_trans _ dvd_uv_p) //. +by rewrite dvdp_mul ?dvdp_gcdl ?dvdp_gcdr. +Qed. + +Lemma separable_polyP {p} : + reflect [/\ forall u v, u * v %| p -> coprimep u v + & forall u, u %| p -> 1 < size u -> u^`() != 0] + (separable p). +Proof. +apply: (iffP idP) => [sep_p | [sq'p nz_der1p]]. + split=> [u v | u u_dv_p]; last first. + apply: contraTneq => u'0; rewrite -leqNgt -(eqnP sep_p). + rewrite dvdp_leq -?size_poly_eq0 ?(eqnP sep_p) // dvdp_gcd u_dv_p. + have /dvdp_scaler <-: lead_coef u ^+ scalp p u != 0 by rewrite lcn_neq0. + by rewrite -derivZ -Pdiv.Idomain.divpK //= derivM u'0 mulr0 addr0 dvdp_mull. + rewrite Pdiv.Idomain.dvdp_eq mulrCA mulrA; set c := _ ^+ _ => /eqP Dcp. + have nz_c: c != 0 by rewrite lcn_neq0. + move: sep_p; rewrite coprimep_sym -[separable _](coprimep_scalel _ _ nz_c). + rewrite -(coprimep_scaler _ _ nz_c) -derivZ Dcp derivM coprimep_mull. + by rewrite coprimep_addl_mul !coprimep_mulr -andbA => /and4P[]. +rewrite /separable coprimep_def eqn_leq size_poly_gt0; set g := gcdp _ _. +have nz_g: g != 0. + rewrite -dvd0p dvdp_gcd -(mulr0 0); apply/nandP; left. + by have /poly_square_freeP-> := sq'p; rewrite ?size_poly0. +have [g_p]: g %| p /\ g %| p^`() by rewrite dvdp_gcdr ?dvdp_gcdl. +pose c := lead_coef g ^+ scalp p g; have nz_c: c != 0 by rewrite lcn_neq0. +have Dcp: c *: p = p %/ g * g by rewrite Pdiv.Idomain.divpK. +rewrite nz_g andbT leqNgt -(dvdp_scaler _ _ nz_c) -derivZ Dcp derivM. +rewrite dvdp_addr; last by rewrite dvdp_mull. +rewrite Gauss_dvdpr; last by rewrite sq'p // mulrC -Dcp dvdp_scalel. +by apply: contraL => /nz_der1p nz_g'; rewrite gtNdvdp ?nz_g' ?lt_size_deriv. +Qed. + +Lemma separable_coprime p u v : separable p -> u * v %| p -> coprimep u v. +Proof. by move=> /separable_polyP[sq'p _] /sq'p. Qed. + +Lemma separable_nosquare p u k : + separable p -> 1 < k -> size u != 1%N -> (u ^+ k %| p) = false. +Proof. +move=> /separable_polyP[/poly_square_freeP sq'p _] /subnKC <- /sq'p. +by apply: contraNF; apply: dvdp_trans; rewrite exprD dvdp_mulr. +Qed. + +Lemma separable_deriv_eq0 p u : + separable p -> u %| p -> 1 < size u -> (u^`() == 0) = false. +Proof. by move=> /separable_polyP[_ nz_der1p] u_p /nz_der1p/negPf->. Qed. + +Lemma dvdp_separable p q : q %| p -> separable p -> separable q. +Proof. +move=> /(dvdp_trans _)q_dv_p /separable_polyP[sq'p nz_der1p]. +by apply/separable_polyP; split=> [u v /q_dv_p/sq'p | u /q_dv_p/nz_der1p]. +Qed. + +Lemma separable_mul p q : + separable (p * q) = [&& separable p, separable q & coprimep p q]. +Proof. +apply/idP/and3P => [sep_pq | [sep_p seq_q co_pq]]. + rewrite !(dvdp_separable _ sep_pq) ?dvdp_mulIr ?dvdp_mulIl //. + by rewrite (separable_coprime sep_pq). +rewrite /separable derivM coprimep_mull {1}addrC mulrC !coprimep_addl_mul. +by rewrite !coprimep_mulr (coprimep_sym q p) co_pq !andbT; apply/andP. +Qed. + +Lemma eqp_separable p q : p %= q -> separable p = separable q. +Proof. by case/andP=> p_q q_p; apply/idP/idP=> /dvdp_separable->. Qed. + +Lemma separable_root p x : + separable (p * ('X - x%:P)) = separable p && ~~ root p x. +Proof. +rewrite separable_mul; apply: andb_id2l => seq_p. +by rewrite /separable derivXsubC coprimep1 coprimep_XsubC. +Qed. + +Lemma separable_prod_XsubC (r : seq R) : + separable (\prod_(x <- r) ('X - x%:P)) = uniq r. +Proof. +elim: r => [|x r IH]; first by rewrite big_nil /separable_poly coprime1p. +by rewrite big_cons mulrC separable_root IH root_prod_XsubC andbC. +Qed. + +Lemma make_separable p : p != 0 -> separable (p %/ gcdp p p^`()). +Proof. +set g := gcdp p p^`() => nz_p; apply/separable_polyP. +have max_dvd_u (u : {poly R}): 1 < size u -> exists k, ~~ (u ^+ k %| p). + move=> u_gt1; exists (size p); rewrite gtNdvdp // polySpred //. + by rewrite -(ltn_subRL 1) subn1 size_exp leq_pmull // -(subnKC u_gt1). +split=> [|u u_pg u_gt1]; last first. + apply/eqP=> u'0 /=; have [k /negP[]] := max_dvd_u u u_gt1. + elim: k => [|k IHk]; first by rewrite dvd1p. + suffices: u ^+ k.+1 %| (p %/ g) * g. + by rewrite Pdiv.Idomain.divpK ?dvdp_gcdl // dvdp_scaler ?lcn_neq0. + rewrite exprS dvdp_mul // dvdp_gcd IHk //=. + suffices: u ^+ k %| (p %/ u ^+ k * u ^+ k)^`(). + by rewrite Pdiv.Idomain.divpK // derivZ dvdp_scaler ?lcn_neq0. + by rewrite !derivCE u'0 mul0r mul0rn mulr0 addr0 dvdp_mull. +have pg_dv_p: p %/ g %| p by rewrite divp_dvd ?dvdp_gcdl. +apply/poly_square_freeP=> u; rewrite neq_ltn ltnS leqn0 size_poly_eq0. +case/predU1P=> [-> | /max_dvd_u[k]]. + by apply: contra nz_p; rewrite expr0n -dvd0p => /dvdp_trans->. +apply: contra => u2_dv_pg; case: k; [by rewrite dvd1p | elim=> [|n IHn]]. + exact: dvdp_trans (dvdp_mulr _ _) (dvdp_trans u2_dv_pg pg_dv_p). +suff: u ^+ n.+2 %| (p %/ g) * g. + by rewrite Pdiv.Idomain.divpK ?dvdp_gcdl // dvdp_scaler ?lcn_neq0. +rewrite -add2n exprD dvdp_mul // dvdp_gcd. +rewrite (dvdp_trans _ IHn) ?exprS ?dvdp_mull //=. +suff: u ^+ n %| ((p %/ u ^+ n.+1) * u ^+ n.+1)^`(). + by rewrite Pdiv.Idomain.divpK // derivZ dvdp_scaler ?lcn_neq0. +by rewrite !derivCE dvdp_add // -1?mulr_natl ?exprS !dvdp_mull. +Qed. + +End SeparablePoly. + +Implicit Arguments separable_polyP [R p]. + +Lemma separable_map (F : fieldType) (R : idomainType) + (f : {rmorphism F -> R}) (p : {poly F}) : + separable_poly (map_poly f p) = separable_poly p. +Proof. +by rewrite /separable_poly deriv_map /coprimep -gcdp_map size_map_poly. +Qed. + +Section InfinitePrimitiveElementTheorem. + +Local Notation "p ^ f" := (map_poly f p) : ring_scope. + +Variables (F L : fieldType) (iota : {rmorphism F -> L}). +Variables (x y : L) (p : {poly F}). +Hypotheses (nz_p : p != 0) (px_0 : root (p ^ iota) x). + +Let inFz z w := exists q, (q ^ iota).[z] = w. + +Lemma large_field_PET q : + root (q ^ iota) y -> separable_poly q -> + exists2 r, r != 0 + & forall t (z := iota t * y - x), ~~ root r (iota t) -> inFz z x /\ inFz z y. +Proof. +move=> qy_0 sep_q; have nz_q := separable_poly_neq0 sep_q. +have /factor_theorem[q0 Dq] := qy_0. +set p1 := p ^ iota \Po ('X + x%:P); set q1 := q0 \Po ('X + y%:P). +have nz_p1: p1 != 0. + apply: contraNneq nz_p => /(canRL (fun r => comp_polyXaddC_K r _))/eqP. + by rewrite comp_poly0 map_poly_eq0. +have{sep_q} nz_q10: q1.[0] != 0. + move: sep_q; rewrite -(separable_map iota) Dq separable_root => /andP[_]. + by rewrite horner_comp !hornerE. +have nz_q1: q1 != 0 by apply: contraNneq nz_q10 => ->; rewrite horner0. +pose p2 := p1 ^ polyC \Po ('X * 'Y); pose q2 := q1 ^ polyC. +have /Bezout_coprimepP[[u v]]: coprimep p2 q2. + rewrite coprimep_def eqn_leq leqNgt andbC size_poly_gt0 gcdp_eq0 poly_XmY_eq0. + by rewrite map_polyC_eq0 (negPf nz_p1) -resultant_eq0 div_annihilant_neq0. +rewrite -size_poly_eq1 => /size_poly1P[r nzr Dr]; exists r => {nzr}// t z nz_rt. +have [r1 nz_r1 r1z_0]: algebraicOver iota z. + apply/algebraic_sub; last by exists p. + by apply: algebraic_mul; [apply: algebraic_id | exists q]. +pose Fz := subFExtend iota z r1; pose kappa : Fz -> L := subfx_inj. +pose kappa' := inj_subfx iota z r1. +have /eq_map_poly Diota: kappa \o kappa' =1 iota. + by move=> w; rewrite /kappa /= subfx_inj_eval // map_polyC hornerC. +suffices [y3]: exists y3, y = kappa y3. + have [q3 ->] := subfxE y3; rewrite /kappa subfx_inj_eval // => Dy. + split; [exists (t *: q3 - 'X) | by exists q3]. + by rewrite rmorphB linearZ /= map_polyX !hornerE -Dy opprB addrC addrNK. +pose p0 := p ^ iota \Po (iota t *: 'X - z%:P). +have co_p0_q0: coprimep p0 q0. + pose at_t := horner_eval (iota t); have at_t0: at_t 0 = 0 by apply: rmorph0. + have /map_polyK polyCK: cancel polyC at_t by move=> w; apply: hornerC. + have ->: p0 = p2 ^ at_t \Po ('X - y%:P). + rewrite map_comp_poly polyCK // rmorphM /= map_polyC map_polyX /=. + rewrite horner_evalE hornerX. + rewrite -!comp_polyA comp_polyM comp_polyD !comp_polyC !comp_polyX. + by rewrite mulrC mulrBr mul_polyC addrAC -addrA -opprB -rmorphM -rmorphB. + have ->: q0 = q2 ^ at_t \Po ('X - y%:P) by rewrite polyCK ?comp_polyXaddC_K. + apply/coprimep_comp_poly/Bezout_coprimepP; exists (u ^ at_t, v ^ at_t). + by rewrite -!rmorphM -rmorphD Dr /= map_polyC polyC_eqp1. +have{co_p0_q0}: gcdp p0 (q ^ iota) %= 'X - y%:P. + rewrite /eqp Dq (eqp_dvdl _ (Gauss_gcdpr _ _)) // dvdp_gcdr dvdp_gcd. + rewrite dvdp_mull // -root_factor_theorem rootE horner_comp !hornerE. + by rewrite opprB addrC subrK. +have{p0} [p3 ->]: exists p3, p0 = p3 ^ kappa. + exists (p ^ kappa' \Po (kappa' t *: 'X - (subfx_eval iota z r1 'X)%:P)). + rewrite map_comp_poly rmorphB linearZ /= map_polyC map_polyX /=. + rewrite !subfx_inj_eval // map_polyC hornerC map_polyX hornerX. + by rewrite -map_poly_comp Diota. +rewrite -Diota map_poly_comp -gcdp_map /= -/kappa. +move: (gcdp _ _) => r3 /eqpf_eq[c nz_c Dr3]. +exists (- (r3`_0 / r3`_1)); rewrite [kappa _]rmorphN fmorph_div -!coef_map Dr3. +by rewrite !coefZ polyseqXsubC mulr1 mulrC mulKf ?opprK. +Qed. + +Lemma char0_PET (q : {poly F}) : + q != 0 -> root (q ^ iota) y -> [char F] =i pred0 -> + exists n, let z := y *+ n - x in inFz z x /\ inFz z y. +Proof. +move=> nz_q qy_0 /charf0P charF0. +without loss{nz_q} sep_q: q qy_0 / separable_poly q. + move=> IHq; apply: IHq (make_separable nz_q). + have /dvdpP[q1 Dq] := dvdp_gcdl q q^`(). + rewrite {1}Dq mulpK ?gcdp_eq0; last by apply/nandP; left. + have [n [r nz_ry Dr]] := multiplicity_XsubC (q ^ iota) y. + rewrite map_poly_eq0 nz_q /= in nz_ry. + case: n => [|n] in Dr; first by rewrite Dr mulr1 (negPf nz_ry) in qy_0. + have: ('X - y%:P) ^+ n.+1 %| q ^ iota by rewrite Dr dvdp_mulIr. + rewrite Dq rmorphM /= gcdp_map -(eqp_dvdr _ (gcdp_mul2l _ _ _)) -deriv_map Dr. + rewrite dvdp_gcd derivM deriv_exp derivXsubC mul1r !mulrA dvdp_mulIr /=. + rewrite mulrDr mulrA dvdp_addr ?dvdp_mulIr // exprS -scaler_nat -!scalerAr. + rewrite dvdp_scaler -?(rmorph_nat iota) ?fmorph_eq0 ?charF0 //. + rewrite mulrA dvdp_mul2r ?expf_neq0 ?polyXsubC_eq0 //. + by rewrite Gauss_dvdpl ?dvdp_XsubCl // coprimep_sym coprimep_XsubC. +have [r nz_r PETxy] := large_field_PET qy_0 sep_q. +pose ts := mkseq (fun n => iota n%:R) (size r). +have /(max_ring_poly_roots nz_r)/=/implyP: uniq_roots ts. + rewrite uniq_rootsE mkseq_uniq // => m n eq_mn; apply/eqP; rewrite eqn_leq. + wlog suffices: m n eq_mn / m <= n by move=> IHmn; rewrite !IHmn. + move/fmorph_inj/eqP: eq_mn; rewrite -subr_eq0 leqNgt; apply: contraL => lt_mn. + by rewrite -natrB ?(ltnW lt_mn) // charF0 -lt0n subn_gt0. +rewrite size_mkseq ltnn implybF all_map => /allPn[n _ /= /PETxy]. +by rewrite rmorph_nat mulr_natl; exists n. +Qed. + +End InfinitePrimitiveElementTheorem. + +Section Separable. + +Variables (F : fieldType) (L : fieldExtType F). +Implicit Types (U V W : {vspace L}) (E K M : {subfield L}) (D : 'End(L)). + +Section Derivation. + +Variables (K : {vspace L}) (D : 'End(L)). + +(* A deriviation only needs to be additive and satify Lebniz's law, but all *) +(* the deriviations used here are going to be linear, so we only define *) +(* the Derivation predicate for linear endomorphisms. *) +Definition Derivation (s := vbasis K) : bool := + all (fun u => all (fun v => D (u * v) == D u * v + u * D v) s) s. + +Hypothesis derD : Derivation. + +Lemma Derivation_mul : {in K &, forall u v, D (u * v) = D u * v + u * D v}. +Proof. +move=> u v /coord_vbasis-> /coord_vbasis->. +rewrite !(mulr_sumr, linear_sum) -big_split; apply: eq_bigr => /= j _. +rewrite !mulr_suml linear_sum -big_split; apply: eq_bigr => /= i _. +rewrite !(=^~ scalerAl, linearZZ) -!scalerAr linearZZ -!scalerDr !scalerA /=. +by congr (_ *: _); apply/eqP; rewrite (allP (allP derD _ _)) ?memt_nth. +Qed. + +Lemma Derivation_mul_poly (Dp := map_poly D) : + {in polyOver K &, forall p q, Dp (p * q) = Dp p * q + p * Dp q}. +Proof. +move=> p q Kp Kq; apply/polyP=> i; rewrite {}/Dp coefD coef_map /= !coefM. +rewrite linear_sum -big_split; apply: eq_bigr => /= j _. +by rewrite !{1}coef_map Derivation_mul ?(polyOverP _). +Qed. + +End Derivation. + +Lemma DerivationS E K D : (K <= E)%VS -> Derivation E D -> Derivation K D. +Proof. +move/subvP=> sKE derD; apply/allP=> x Kx; apply/allP=> y Ky; apply/eqP. +by rewrite (Derivation_mul derD) ?sKE // vbasis_mem. +Qed. + +Section DerivationAlgebra. + +Variables (E : {subfield L}) (D : 'End(L)). +Hypothesis derD : Derivation E D. + +Lemma Derivation1 : D 1 = 0. +Proof. +apply: (addIr (D (1 * 1))); rewrite add0r {1}mul1r. +by rewrite (Derivation_mul derD) ?mem1v // mulr1 mul1r. +Qed. + +Lemma Derivation_scalar x : x \in 1%VS -> D x = 0. +Proof. by case/vlineP=> y ->; rewrite linearZ /= Derivation1 scaler0. Qed. + +Lemma Derivation_exp x m : x \in E -> D (x ^+ m) = x ^+ m.-1 *+ m * D x. +Proof. +move=> Ex; case: m; first by rewrite expr0 mulr0n mul0r Derivation1. +elim=> [|m IHm]; first by rewrite mul1r. +rewrite exprS (Derivation_mul derD) //; last by apply: rpredX. +by rewrite mulrC IHm mulrA mulrnAr -exprS -mulrDl. +Qed. + +Lemma Derivation_horner p x : + p \is a polyOver E -> x \in E -> + D p.[x] = (map_poly D p).[x] + p^`().[x] * D x. +Proof. +move=> Ep Ex; elim/poly_ind: p Ep => [|p c IHp] /polyOverP EpXc. + by rewrite !(raddf0, horner0) mul0r add0r. +have Ep: p \is a polyOver E. + by apply/polyOverP=> i; have:= EpXc i.+1; rewrite coefD coefMX coefC addr0. +have->: map_poly D (p * 'X + c%:P) = map_poly D p * 'X + (D c)%:P. + apply/polyP=> i; rewrite !(coefD, coefMX, coef_map) /= linearD /= !coefC. + by rewrite !(fun_if D) linear0. +rewrite derivMXaddC !hornerE mulrDl mulrAC addrAC linearD /=; congr (_ + _). +by rewrite addrCA -mulrDl -IHp // addrC (Derivation_mul derD) ?rpred_horner. +Qed. + +End DerivationAlgebra. + +Definition separable_element U x := separable_poly (minPoly U x). + +Section SeparableElement. + +Variables (K : {subfield L}) (x : L). +(* begin hide *) +Let sKxK : (K <= <<K; x>>)%VS := subv_adjoin K x. +Let Kx_x : x \in <<K; x>>%VS := memv_adjoin K x. +(* end hide *) + +Lemma separable_elementP : + reflect (exists f, [/\ f \is a polyOver K, root f x & separable_poly f]) + (separable_element K x). +Proof. +apply: (iffP idP) => [sep_x | [f [Kf /(minPoly_dvdp Kf)/dvdpP[g ->]]]]. + by exists (minPoly K x); rewrite minPolyOver root_minPoly. +by rewrite separable_mul => /and3P[]. +Qed. + +Lemma base_separable : x \in K -> separable_element K x. +Proof. +move=> Kx; apply/separable_elementP; exists ('X - x%:P). +by rewrite polyOverXsubC root_XsubC /separable_poly !derivCE coprimep1. +Qed. + +Lemma separable_nz_der : separable_element K x = ((minPoly K x)^`() != 0). +Proof. +rewrite /separable_element /separable_poly. +apply/idP/idP=> [|nzPx']. + by apply: contraTneq => ->; rewrite coprimep0 -size_poly_eq1 size_minPoly. +have gcdK : gcdp (minPoly K x) (minPoly K x)^`() \in polyOver K. + by rewrite gcdp_polyOver ?polyOver_deriv // minPolyOver. +rewrite -gcdp_eqp1 -size_poly_eq1 -dvdp1. +have /orP[/andP[_]|/andP[]//] := minPoly_irr gcdK (dvdp_gcdl _ _). +rewrite dvdp_gcd dvdpp /= => /(dvdp_leq nzPx')/leq_trans/(_ (size_poly _ _)). +by rewrite size_minPoly ltnn. +Qed. + +Lemma separablePn : + reflect (exists2 p, p \in [char L] & + exists2 g, g \is a polyOver K & minPoly K x = g \Po 'X^p) + (~~ separable_element K x). +Proof. +rewrite separable_nz_der negbK; set f := minPoly K x. +apply: (iffP eqP) => [f'0 | [p Hp [g _ ->]]]; last first. + by rewrite deriv_comp derivXn -scaler_nat (charf0 Hp) scale0r mulr0. +pose n := adjoin_degree K x; have sz_f: size f = n.+1 := size_minPoly K x. +have fn1: f`_n = 1 by rewrite -(monicP (monic_minPoly K x)) lead_coefE sz_f. +have dimKx: (adjoin_degree K x)%:R == 0 :> L. + by rewrite -(coef0 _ n.-1) -f'0 coef_deriv fn1. +have /natf0_char[// | p charLp] := dimKx. +have /dvdnP[r Dn]: (p %| n)%N by rewrite (dvdn_charf charLp). +exists p => //; exists (\poly_(i < r.+1) f`_(i * p)). + by apply: polyOver_poly => i _; rewrite (polyOverP _) ?minPolyOver. +rewrite comp_polyE size_poly_eq -?Dn ?fn1 ?oner_eq0 //. +have pr_p := charf_prime charLp; have p_gt0 := prime_gt0 pr_p. +apply/polyP=> i; rewrite coef_sum. +have [[{i} i ->] | p'i] := altP (@dvdnP p i); last first. + rewrite big1 => [|j _]; last first. + rewrite coefZ -exprM coefXn [_ == _](contraNF _ p'i) ?mulr0 // => /eqP->. + by rewrite dvdn_mulr. + rewrite (dvdn_charf charLp) in p'i; apply: mulfI p'i _ _ _. + by rewrite mulr0 mulr_natl; case: i => // i; rewrite -coef_deriv f'0 coef0. +have [ltri | leir] := leqP r.+1 i. + rewrite nth_default ?sz_f ?Dn ?ltn_pmul2r ?big1 // => j _. + rewrite coefZ -exprM coefXn mulnC gtn_eqF ?mulr0 //. + by rewrite ltn_pmul2l ?(leq_trans _ ltri). +rewrite (bigD1 (Sub i _)) //= big1 ?addr0 => [|j i'j]; last first. + by rewrite coefZ -exprM coefXn mulnC eqn_pmul2l // mulr_natr mulrb ifN_eqC. +by rewrite coef_poly leir coefZ -exprM coefXn mulnC eqxx mulr1. +Qed. + +Lemma separable_root_der : separable_element K x (+) root (minPoly K x)^`() x. +Proof. +have KpKx': _^`() \is a polyOver K := polyOver_deriv (minPolyOver K x). +rewrite separable_nz_der addNb (root_small_adjoin_poly KpKx') ?addbb //. +by rewrite (leq_trans (size_poly _ _)) ?size_minPoly. +Qed. + +Lemma Derivation_separable D : + Derivation <<K; x>> D -> separable_element K x -> + D x = - (map_poly D (minPoly K x)).[x] / (minPoly K x)^`().[x]. +Proof. +move=> derD sepKx; have:= separable_root_der; rewrite {}sepKx -sub0r => nzKx'x. +apply: canRL (mulfK nzKx'x) (canRL (addrK _) _); rewrite mulrC addrC. +rewrite -(Derivation_horner derD) ?minPolyxx ?linear0 //. +exact: polyOverSv sKxK _ (minPolyOver _ _). +Qed. + +Section ExtendDerivation. + +Variable D : 'End(L). + +Let Dx E := - (map_poly D (minPoly E x)).[x] / ((minPoly E x)^`()).[x]. + +Fact extendDerivation_subproof E (adjEx := Fadjoin_poly E x) : + let body y (p := adjEx y) := (map_poly D p).[x] + p^`().[x] * Dx E in + linear body. +Proof. +move: Dx => C /= a u v. +rewrite /adjEx linearP /= -mul_polyC derivD derivM derivC mul0r add0r -/adjEx. +rewrite !hornerE /= -scalerAl mul1r raddfD /=. +have ->: map_poly D (a%:A%:P * adjEx u) = a%:A%:P * map_poly D (adjEx u). + apply/polyP=> i; rewrite !mul_polyC !coef_map !coefZ !mulr_algl /= linearZ. + by rewrite coef_map. +rewrite !hornerE !mulr_algl mulrDl scalerDr -scalerAl -!addrA; congr (_ + _). +by rewrite addrCA. +Qed. + +Definition extendDerivation E : 'End(L) := + linfun (Linear (extendDerivation_subproof E)). + +Hypothesis derD : Derivation K D. + +Lemma extendDerivation_id y : y \in K -> extendDerivation K y = D y. +Proof. +move=> yK; rewrite lfunE /= Fadjoin_polyC // derivC map_polyC hornerC. +by rewrite horner0 mul0r addr0. +Qed. + +Lemma extendDerivation_horner p : + p \is a polyOver K -> separable_element K x -> + extendDerivation K p.[x] = (map_poly D p).[x] + p^`().[x] * Dx K. +Proof. +move=> Kp sepKx; have:= separable_root_der; rewrite {}sepKx /= => nz_pKx'x. +rewrite {-1}(divp_eq p (minPoly K x)) lfunE /= Fadjoin_poly_mod // raddfD /=. +rewrite {1}(Derivation_mul_poly derD) ?divp_polyOver ?minPolyOver //. +rewrite derivD derivM !{1}hornerD !{1}hornerM minPolyxx !{1}mulr0 !{1}add0r. +rewrite mulrDl addrA [_ + (_ * _ * _)]addrC {2}/Dx -mulrA -/Dx. +by rewrite [_ / _]mulrC (mulVKf nz_pKx'x) mulrN addKr. +Qed. + +Lemma extendDerivationP : + separable_element K x -> Derivation <<K; x>> (extendDerivation K). +Proof. +move=> sep; apply/allP=> u /vbasis_mem Hu; apply/allP=> v /vbasis_mem Hv. +apply/eqP. +rewrite -(Fadjoin_poly_eq Hu) -(Fadjoin_poly_eq Hv) -hornerM. +rewrite !{1}extendDerivation_horner ?{1}rpredM ?Fadjoin_polyOver //. +rewrite (Derivation_mul_poly derD) ?Fadjoin_polyOver //. +rewrite derivM !{1}hornerD !{1}hornerM !{1}mulrDl !{1}mulrDr -!addrA. +congr (_ + _); rewrite [Dx K]lock -!{1}mulrA !{1}addrA; congr (_ + _). +by rewrite addrC; congr (_ * _ + _); rewrite mulrC. +Qed. + +End ExtendDerivation. + +(* Reference: +http://www.math.uconn.edu/~kconrad/blurbs/galoistheory/separable2.pdf *) +Lemma Derivation_separableP : + reflect + (forall D, Derivation <<K; x>> D -> K <= lker D -> <<K; x>> <= lker D)%VS + (separable_element K x). +Proof. +apply: (iffP idP) => [sepKx D derD /subvP DK_0 | derKx_0]. + have{DK_0} DK_0 q: q \is a polyOver K -> map_poly D q = 0. + move=> /polyOverP Kq; apply/polyP=> i; apply/eqP. + by rewrite coef0 coef_map -memv_ker DK_0. + apply/subvP=> _ /Fadjoin_polyP[p Kp ->]; rewrite memv_ker. + rewrite (Derivation_horner derD) ?(polyOverSv sKxK) //. + rewrite (Derivation_separable derD sepKx) !DK_0 ?minPolyOver //. + by rewrite horner0 oppr0 mul0r mulr0 addr0. +apply: wlog_neg; rewrite {1}separable_nz_der negbK => /eqP pKx'_0. +have Dlin: linear (fun y => (Fadjoin_poly K x y)^`().[x]). + move=> a u v; rewrite linearP /= -mul_polyC derivD derivM derivC mul0r add0r. + by rewrite hornerD hornerM hornerC -scalerAl mul1r. +pose D := linfun (Linear Dlin); apply: base_separable. +have DK_0: (K <= lker D)%VS. + apply/subvP=> v Kv; rewrite memv_ker lfunE /= Fadjoin_polyC //. + by rewrite derivC horner0. +have Dder: Derivation <<K; x>> D. + apply/allP=> u /vbasis_mem Kx_u; apply/allP=> v /vbasis_mem Kx_v. + rewrite !lfunE /= -{-2}(Fadjoin_poly_eq Kx_u) -{-3}(Fadjoin_poly_eq Kx_v). + rewrite -!hornerM -hornerD -derivM. + rewrite Fadjoin_poly_mod ?rpredM ?Fadjoin_polyOver //. + rewrite {2}(divp_eq (_ * _) (minPoly K x)) derivD derivM pKx'_0 mulr0 addr0. + by rewrite hornerD hornerM minPolyxx mulr0 add0r. +have{Dder DK_0}: x \in lker D by apply: subvP Kx_x; apply: derKx_0. +apply: contraLR => K'x; rewrite memv_ker lfunE /= Fadjoin_polyX //. +by rewrite derivX hornerC oner_eq0. +Qed. + +End SeparableElement. + +Implicit Arguments separable_elementP [K x]. + +Lemma separable_elementS K E x : + (K <= E)%VS -> separable_element K x -> separable_element E x. +Proof. +move=> sKE /separable_elementP[f [fK rootf sepf]]; apply/separable_elementP. +by exists f; rewrite (polyOverSv sKE). +Qed. + +Lemma adjoin_separableP {K x} : + reflect (forall y, y \in <<K; x>>%VS -> separable_element K y) + (separable_element K x). +Proof. +apply: (iffP idP) => [sepKx | -> //]; last exact: memv_adjoin. +move=> _ /Fadjoin_polyP[q Kq ->]; apply/Derivation_separableP=> D derD DK_0. +apply/subvP=> _ /Fadjoin_polyP[p Kp ->]. +rewrite memv_ker -(extendDerivation_id x D (mempx_Fadjoin _ Kp)). +have sepFyx: (separable_element <<K; q.[x]>> x). + by apply: (separable_elementS (subv_adjoin _ _)). +have KyxEqKx: (<< <<K; q.[x]>>; x>> = <<K; x>>)%VS. + apply/eqP; rewrite eqEsubv andbC adjoinSl ?subv_adjoin //=. + apply/FadjoinP/andP; rewrite memv_adjoin andbT. + by apply/FadjoinP/andP; rewrite subv_adjoin mempx_Fadjoin. +have:= extendDerivationP derD sepFyx; rewrite KyxEqKx => derDx. +rewrite -horner_comp (Derivation_horner derDx) ?memv_adjoin //; last first. + by apply: (polyOverSv (subv_adjoin _ _)); apply: polyOver_comp. +set Dx_p := map_poly _; have Dx_p_0 t: t \is a polyOver K -> (Dx_p t).[x] = 0. + move/polyOverP=> Kt; congr (_.[x] = 0): (horner0 x); apply/esym/polyP => i. + have /eqP Dti_0: D t`_i == 0 by rewrite -memv_ker (subvP DK_0) ?Kt. + by rewrite coef0 coef_map /= {1}extendDerivation_id ?subvP_adjoin. +rewrite (Derivation_separable derDx sepKx) -/Dx_p Dx_p_0 ?polyOver_comp //. +by rewrite add0r mulrCA Dx_p_0 ?minPolyOver ?oppr0 ?mul0r. +Qed. + +Lemma separable_exponent K x : + exists n, [char L].-nat n && separable_element K (x ^+ n). +Proof. +pose d := adjoin_degree K x; move: {2}d.+1 (ltnSn d) => n. +elim: n => // n IHn in x @d *; rewrite ltnS => le_d_n. +have [[p charLp]|] := altP (separablePn K x); last by rewrite negbK; exists 1%N. +case=> g Kg defKx; have p_pr := charf_prime charLp. +suffices /IHn[m /andP[charLm sepKxpm]]: adjoin_degree K (x ^+ p) < n. + by exists (p * m)%N; rewrite pnat_mul pnatE // charLp charLm exprM. +apply: leq_trans le_d_n; rewrite -ltnS -!size_minPoly. +have nzKx: minPoly K x != 0 by rewrite monic_neq0 ?monic_minPoly. +have nzg: g != 0 by apply: contra_eqN defKx => /eqP->; rewrite comp_poly0. +apply: leq_ltn_trans (dvdp_leq nzg _) _. + by rewrite minPoly_dvdp // rootE -hornerXn -horner_comp -defKx minPolyxx. +rewrite (polySpred nzKx) ltnS defKx size_comp_poly size_polyXn /=. +suffices g_gt1: 1 < size g by rewrite -(subnKC g_gt1) ltn_Pmulr ?prime_gt1. +apply: contra_eqT (size_minPoly K x); rewrite defKx -leqNgt => /size1_polyC->. +by rewrite comp_polyC size_polyC; case: (_ != 0). +Qed. + +Lemma charf0_separable K : [char L] =i pred0 -> forall x, separable_element K x. +Proof. +move=> charL0 x; have [n /andP[charLn]] := separable_exponent K x. +by rewrite (pnat_1 charLn (sub_in_pnat _ charLn)) // => p _; rewrite charL0. +Qed. + +Lemma charf_p_separable K x e p : + p \in [char L] -> separable_element K x = (x \in <<K; x ^+ (p ^ e.+1)>>%VS). +Proof. +move=> charLp; apply/idP/idP=> [sepKx | /Fadjoin_poly_eq]; last first. + set m := p ^ _;set f := Fadjoin_poly K _ x => Dx; apply/separable_elementP. + have mL0: m%:R = 0 :> L by apply/eqP; rewrite -(dvdn_charf charLp) dvdn_exp. + exists ('X - (f \Po 'X^m)); split. + - by rewrite rpredB ?polyOver_comp ?rpredX ?polyOverX ?Fadjoin_polyOver. + - by rewrite rootE !hornerE horner_comp hornerXn Dx subrr. + rewrite /separable_poly !(derivE, deriv_comp) -mulr_natr -rmorphMn /= mL0. + by rewrite !mulr0 subr0 coprimep1. +without loss{e} ->: e x sepKx / e = 0%N. + move=> IH; elim: {e}e.+1 => [|e]; [exact: memv_adjoin | apply: subvP]. + apply/FadjoinP/andP; rewrite subv_adjoin expnSr exprM (IH 0%N) //. + by have /adjoin_separableP-> := sepKx; rewrite ?rpredX ?memv_adjoin. +set K' := <<K; x ^+ p>>%VS; have sKK': (K <= K')%VS := subv_adjoin _ _. +pose q := minPoly K' x; pose g := 'X^p - (x ^+ p)%:P. +have [K'g]: g \is a polyOver K' /\ q \is a polyOver K'. + by rewrite minPolyOver rpredB ?rpredX ?polyOverX // polyOverC memv_adjoin. +have /dvdpP[c Dq]: 'X - x%:P %| q by rewrite dvdp_XsubCl root_minPoly. +have co_c_g: coprimep c g. + have charPp: p \in [char {poly L}] := rmorph_char (polyC_rmorphism _) charLp. + rewrite /g polyC_exp -!(Frobenius_autE charPp) -rmorphB coprimep_expr //. + have: separable_poly q := separable_elementS sKK' sepKx. + by rewrite Dq separable_mul => /and3P[]. +have{g K'g co_c_g} /size_poly1P[a nz_a Dc]: size c == 1%N. + suffices c_dv_g: c %| g by rewrite -(eqp_size (dvdp_gcd_idl c_dv_g)). + have: q %| g by rewrite minPoly_dvdp // rootE !hornerE hornerXn subrr. + by apply: dvdp_trans; rewrite Dq dvdp_mulIl. +rewrite {q}Dq {c}Dc mulrBr -rmorphM -rmorphN -cons_poly_def qualifE. +by rewrite polyseq_cons !polyseqC nz_a /= rpredN andbCA => /and3P[/fpredMl->]. +Qed. + +Lemma charf_n_separable K x n : + [char L].-nat n -> 1 < n -> separable_element K x = (x \in <<K; x ^+ n>>%VS). +Proof. +rewrite -pi_pdiv; set p := pdiv n => charLn pi_n_p. +have charLp: p \in [char L] := pnatPpi charLn pi_n_p. +have <-: (n`_p)%N = n by rewrite -(eq_partn n (charf_eq charLp)) part_pnat_id. +by rewrite p_part lognE -mem_primes pi_n_p -charf_p_separable. +Qed. + +Definition purely_inseparable_element U x := + x ^+ ex_minn (separable_exponent <<U>> x) \in U. + +Lemma purely_inseparable_elementP {K x} : + reflect (exists2 n, [char L].-nat n & x ^+ n \in K) + (purely_inseparable_element K x). +Proof. +rewrite /purely_inseparable_element. +case: ex_minnP => n /andP[charLn /=]; rewrite subfield_closed => sepKxn min_xn. +apply: (iffP idP) => [Kxn | [m charLm Kxm]]; first by exists n. +have{min_xn}: n <= m by rewrite min_xn ?charLm ?base_separable. +rewrite leq_eqVlt => /predU1P[-> // | ltnm]; pose p := pdiv m. +have m_gt1: 1 < m by have [/leq_ltn_trans->] := andP charLn. +have charLp: p \in [char L] by rewrite (pnatPpi charLm) ?pi_pdiv. +have [/p_natP[em Dm] /p_natP[en Dn]]: p.-nat m /\ p.-nat n. + by rewrite -!(eq_pnat _ (charf_eq charLp)). +rewrite Dn Dm ltn_exp2l ?prime_gt1 ?pdiv_prime // in ltnm. +rewrite -(Fadjoin_idP Kxm) Dm -(subnKC ltnm) addSnnS expnD exprM -Dn. +by rewrite -charf_p_separable. +Qed. + +Lemma separable_inseparable_element K x : + separable_element K x && purely_inseparable_element K x = (x \in K). +Proof. +rewrite /purely_inseparable_element; case: ex_minnP => [[|m]] //=. +rewrite subfield_closed; case: m => /= [-> //| m _ /(_ 1%N)/implyP/= insepKx]. +by rewrite (negPf insepKx) (contraNF (@base_separable K x) insepKx). +Qed. + +Lemma base_inseparable K x : x \in K -> purely_inseparable_element K x. +Proof. by rewrite -separable_inseparable_element => /andP[]. Qed. + +Lemma sub_inseparable K E x : + (K <= E)%VS -> purely_inseparable_element K x -> + purely_inseparable_element E x. +Proof. +move/subvP=> sKE /purely_inseparable_elementP[n charLn /sKE Exn]. +by apply/purely_inseparable_elementP; exists n. +Qed. + +Section PrimitiveElementTheorem. + +Variables (K : {subfield L}) (x y : L). + +Section FiniteCase. + +Variable N : nat. + +Let K_is_large := exists s, [/\ uniq s, {subset s <= K} & N < size s]. + +Let cyclic_or_large (z : L) : z != 0 -> K_is_large \/ exists a, z ^+ a.+1 = 1. +Proof. +move=> nz_z; pose d := adjoin_degree K z. +pose h0 (i : 'I_(N ^ d).+1) (j : 'I_d) := (Fadjoin_poly K z (z ^+ i))`_j. +pose s := undup [seq h0 i j | i <- enum 'I_(N ^ d).+1, j <- enum 'I_d]. +have s_h0 i j: h0 i j \in s. + by rewrite mem_undup; apply/allpairsP; exists (i, j); rewrite !mem_enum. +pose h i := [ffun j => Ordinal (etrans (index_mem _ _) (s_h0 i j))]. +pose h' (f : {ffun 'I_d -> 'I_(size s)}) := \sum_(j < d) s`_(f j) * z ^+ j. +have hK i: h' (h i) = z ^+ i. + have Kz_zi: z ^+ i \in <<K; z>>%VS by rewrite rpredX ?memv_adjoin. + rewrite -(Fadjoin_poly_eq Kz_zi) (horner_coef_wide z (size_poly _ _)) -/d. + by apply: eq_bigr => j _; rewrite ffunE /= nth_index. +have [inj_h | ] := altP (@injectiveP _ _ h). + left; exists s; split=> [|zi_j|]; rewrite ?undup_uniq ?mem_undup //=. + by case/allpairsP=> ij [_ _ ->]; apply/polyOverP/Fadjoin_polyOver. + rewrite -[size s]card_ord -(@ltn_exp2r _ _ d) // -{2}[d]card_ord -card_ffun. + by rewrite -[_.+1]card_ord -(card_image inj_h) max_card. +case/injectivePn=> i1 [i2 i1'2 /(congr1 h')]; rewrite !hK => eq_zi12; right. +without loss{i1'2} lti12: i1 i2 eq_zi12 / i1 < i2. + by move=> IH; move: i1'2; rewrite neq_ltn => /orP[]; apply: IH. +by exists (i2 - i1.+1)%N; rewrite subnSK ?expfB // eq_zi12 divff ?expf_neq0. +Qed. + +Lemma finite_PET : K_is_large \/ exists z, (<< <<K; y>>; x>> = <<K; z>>)%VS. +Proof. +have [-> | /cyclic_or_large[|[a Dxa]]] := eqVneq x 0; first 2 [by left]. + by rewrite addv0 subfield_closed; right; exists y. +have [-> | /cyclic_or_large[|[b Dyb]]] := eqVneq y 0; first 2 [by left]. + by rewrite addv0 subfield_closed; right; exists x. +pose h0 (ij : 'I_a.+1 * 'I_b.+1) := x ^+ ij.1 * y ^+ ij.2. +pose H := <<[set ij | h0 ij == 1%R]>>%G; pose h (u : coset_of H) := h0 (repr u). +have h0M: {morph h0: ij1 ij2 / (ij1 * ij2)%g >-> ij1 * ij2}. + by rewrite /h0 => [] [i1 j1] [i2 j2] /=; rewrite mulrACA -!exprD !expr_mod. +have memH ij: (ij \in H) = (h0 ij == 1). + rewrite /= gen_set_id ?inE //; apply/group_setP; rewrite inE [h0 _]mulr1. + by split=> // ? ?; rewrite !inE h0M => /eqP-> /eqP->; rewrite mulr1. +have nH ij: ij \in 'N(H)%g. + by apply/(subsetP (cent_sub _))/centP=> ij1 _; congr (_, _); rewrite Zp_mulgC. +have hE ij: h (coset H ij) = h0 ij. + rewrite /h val_coset //; case: repr_rcosetP => ij1. + by rewrite memH h0M => /eqP->; rewrite mul1r. +have h1: h 1%g = 1 by rewrite /h repr_coset1 [h0 _]mulr1. +have hM: {morph h: u v / (u * v)%g >-> u * v}. + by do 2![move=> u; have{u} [? _ ->] := cosetP u]; rewrite -morphM // !hE h0M. +have /cyclicP[w defW]: cyclic [set: coset_of H]. + apply: field_mul_group_cyclic (in2W hM) _ => u _; have [ij _ ->] := cosetP u. + by split=> [/eqP | -> //]; rewrite hE -memH => /coset_id. +have Kw_h ij t: h0 ij = t -> t \in <<K; h w>>%VS. + have /cycleP[k Dk]: coset H ij \in <[w]>%g by rewrite -defW inE. + rewrite -hE {}Dk => <-; elim: k => [|k IHk]; first by rewrite h1 rpred1. + by rewrite expgS hM rpredM // memv_adjoin. +right; exists (h w); apply/eqP; rewrite eqEsubv !(sameP FadjoinP andP). +rewrite subv_adjoin (subv_trans (subv_adjoin K y)) ?subv_adjoin //=. +rewrite (Kw_h (0, inZp 1)) 1?(Kw_h (inZp 1, 0)) /h0 ?mulr1 ?mul1r ?expr_mod //=. +by rewrite rpredM ?rpredX ?memv_adjoin // subvP_adjoin ?memv_adjoin. +Qed. + +End FiniteCase. + +Hypothesis sepKy : separable_element K y. + +Lemma Primitive_Element_Theorem : exists z, (<< <<K; y>>; x>> = <<K; z>>)%VS. +Proof. +have /polyOver_subvs[p Dp]: minPoly K x \is a polyOver K := minPolyOver K x. +have nz_pKx: minPoly K x != 0 by rewrite monic_neq0 ?monic_minPoly. +have{nz_pKx} nz_p: p != 0 by rewrite Dp map_poly_eq0 in nz_pKx. +have{Dp} px0: root (map_poly vsval p) x by rewrite -Dp root_minPoly. +have [q0 [Kq0 [q0y0 sepKq0]]] := separable_elementP sepKy. +have /polyOver_subvs[q Dq]: minPoly K y \is a polyOver K := minPolyOver K y. +have qy0: root (map_poly vsval q) y by rewrite -Dq root_minPoly. +have sep_pKy: separable_poly (minPoly K y). + by rewrite (dvdp_separable _ sepKq0) ?minPoly_dvdp. +have{sep_pKy} sep_q: separable_poly q by rewrite Dq separable_map in sep_pKy. +have [r [nz_r PETr]] := large_field_PET nz_p px0 qy0 sep_q. +have [[s [Us Ks /ltnW leNs]] | //] := finite_PET (size r). +have{s Us Ks leNs} /allPn[t /Ks Kt nz_rt]: ~~ all (root r) s. + by apply: contraTN leNs; rewrite -ltnNge => /max_poly_roots->. +have{PETr} [/= [p1 Dx] [q1 Dy]] := PETr (Subvs Kt) nz_rt. +set z := t * y - x in Dx Dy; exists z; apply/eqP. +rewrite eqEsubv !(sameP FadjoinP andP) subv_adjoin. +have Kz_p1z (r1 : {poly subvs_of K}): (map_poly vsval r1).[z] \in <<K; z>>%VS. + rewrite rpred_horner ?memv_adjoin ?(polyOverSv (subv_adjoin K z)) //. + by apply/polyOver_subvs; exists r1. +rewrite -{1}Dx -{1}Dy !{Dx Dy}Kz_p1z /=. +rewrite (subv_trans (subv_adjoin K y)) ?subv_adjoin // rpredB ?memv_adjoin //. +by rewrite subvP_adjoin // rpredM ?memv_adjoin ?subvP_adjoin. +Qed. + +Lemma adjoin_separable : separable_element <<K; y>> x -> separable_element K x. +Proof. +have /Derivation_separableP derKy := sepKy => /Derivation_separableP derKy_x. +have [z defKz] := Primitive_Element_Theorem. +suffices /adjoin_separableP: separable_element K z. + by apply; rewrite -defKz memv_adjoin. +apply/Derivation_separableP=> D; rewrite -defKz => derKxyD DK_0. +suffices derKyD: Derivation <<K; y>>%VS D by rewrite derKy_x // derKy. +by apply: DerivationS derKxyD; apply: subv_adjoin. +Qed. + +End PrimitiveElementTheorem. + +Lemma strong_Primitive_Element_Theorem K x y : + separable_element <<K; x>> y -> + exists2 z : L, (<< <<K; y>>; x>> = <<K; z>>)%VS + & separable_element K x -> separable_element K y. +Proof. +move=> sepKx_y; have [n /andP[charLn sepKyn]] := separable_exponent K y. +have adjK_C z t: (<<<<K; z>>; t>> = <<<<K; t>>; z>>)%VS. + by rewrite !agenv_add_id -!addvA (addvC <[_]>%VS). +have [z defKz] := Primitive_Element_Theorem x sepKyn. +exists z => [|/adjoin_separable->]; rewrite ?sepKx_y // -defKz. +have [|n_gt1|-> //] := ltngtP n 1%N; first by case: (n) charLn. +apply/eqP; rewrite !(adjK_C _ x) eqEsubv; apply/andP. +split; apply/FadjoinP/andP; rewrite subv_adjoin ?rpredX ?memv_adjoin //=. +by rewrite -charf_n_separable ?sepKx_y. +Qed. + +Definition separable U W : bool := + all (separable_element U) (vbasis W). + +Definition purely_inseparable U W : bool := + all (purely_inseparable_element U) (vbasis W). + +Lemma separable_add K x y : + separable_element K x -> separable_element K y -> separable_element K (x + y). +Proof. +move/(separable_elementS (subv_adjoin K y))=> sepKy_x sepKy. +have [z defKz] := Primitive_Element_Theorem x sepKy. +have /(adjoin_separableP _): x + y \in <<K; z>>%VS. + by rewrite -defKz rpredD ?memv_adjoin // subvP_adjoin ?memv_adjoin. +apply; apply: adjoin_separable sepKy (adjoin_separable sepKy_x _). +by rewrite defKz base_separable ?memv_adjoin. +Qed. + +Lemma separable_sum I r (P : pred I) (v_ : I -> L) K : + (forall i, P i -> separable_element K (v_ i)) -> + separable_element K (\sum_(i <- r | P i) v_ i). +Proof. +move=> sepKi. +by elim/big_ind: _; [apply/base_separable/mem0v | apply: separable_add |]. +Qed. + +Lemma inseparable_add K x y : + purely_inseparable_element K x -> purely_inseparable_element K y -> + purely_inseparable_element K (x + y). +Proof. +have insepP := purely_inseparable_elementP. +move=> /insepP[n charLn Kxn] /insepP[m charLm Kym]; apply/insepP. +have charLnm: [char L].-nat (n * m)%N by rewrite pnat_mul charLn. +by exists (n * m)%N; rewrite ?exprDn_char // {2}mulnC !exprM memvD // rpredX. +Qed. + +Lemma inseparable_sum I r (P : pred I) (v_ : I -> L) K : + (forall i, P i -> purely_inseparable_element K (v_ i)) -> + purely_inseparable_element K (\sum_(i <- r | P i) v_ i). +Proof. +move=> insepKi. +by elim/big_ind: _; [apply/base_inseparable/mem0v | apply: inseparable_add |]. +Qed. + +Lemma separableP {K E} : + reflect (forall y, y \in E -> separable_element K y) (separable K E). +Proof. +apply/(iffP idP)=> [/allP|] sepK_E; last by apply/allP=> x /vbasis_mem/sepK_E. +move=> y /coord_vbasis->; apply/separable_sum=> i _. +have: separable_element K (vbasis E)`_i by apply/sepK_E/memt_nth. +by move/adjoin_separableP; apply; rewrite rpredZ ?memv_adjoin. +Qed. + +Lemma purely_inseparableP {K E} : + reflect (forall y, y \in E -> purely_inseparable_element K y) + (purely_inseparable K E). +Proof. +apply/(iffP idP)=> [/allP|] sep'K_E; last by apply/allP=> x /vbasis_mem/sep'K_E. +move=> y /coord_vbasis->; apply/inseparable_sum=> i _. +have: purely_inseparable_element K (vbasis E)`_i by apply/sep'K_E/memt_nth. +case/purely_inseparable_elementP=> n charLn K_Ein. +by apply/purely_inseparable_elementP; exists n; rewrite // exprZn rpredZ. +Qed. + +Lemma adjoin_separable_eq K x : separable_element K x = separable K <<K; x>>%VS. +Proof. exact: sameP adjoin_separableP separableP. Qed. + +Lemma separable_inseparable_decomposition E K : + {x | x \in E /\ separable_element K x & purely_inseparable <<K; x>> E}. +Proof. +without loss sKE: K / (K <= E)%VS. + case/(_ _ (capvSr K E)) => x [Ex sepKEx] /purely_inseparableP sep'KExE. + exists x; first by split; last exact/(separable_elementS _ sepKEx)/capvSl. + apply/purely_inseparableP=> y /sep'KExE; apply: sub_inseparable. + exact/adjoinSl/capvSl. +pose E_ i := (vbasis E)`_i; pose fP i := separable_exponent K (E_ i). +pose f i := E_ i ^+ ex_minn (fP i); pose s := mkseq f (\dim E). +pose K' := <<K & s>>%VS. +have sepKs: all (separable_element K) s. + by rewrite all_map /f; apply/allP=> i _ /=; case: ex_minnP => m /andP[]. +have [x sepKx defKx]: {x | x \in E /\ separable_element K x & K' = <<K; x>>%VS}. + have: all (mem E) s. + rewrite all_map; apply/allP=> i; rewrite mem_iota => ltis /=. + by rewrite rpredX // vbasis_mem // memt_nth. + rewrite {}/K'; elim/last_ind: s sepKs => [|s t IHs]. + by exists 0; [rewrite base_separable mem0v | rewrite adjoin_nil addv0]. + rewrite adjoin_rcons !all_rcons => /andP[sepKt sepKs] /andP[/= Et Es]. + have{IHs sepKs Es} [y [Ey sepKy] ->{s}] := IHs sepKs Es. + have /sig_eqW[x defKx] := Primitive_Element_Theorem t sepKy. + exists x; [split | exact: defKx]. + suffices: (<<K; x>> <= E)%VS by case/FadjoinP. + by rewrite -defKx !(sameP FadjoinP andP) sKE Ey Et. + apply/adjoin_separableP=> z; rewrite -defKx => Kyt_z. + apply: adjoin_separable sepKy _; apply: adjoin_separableP Kyt_z. + exact: separable_elementS (subv_adjoin K y) sepKt. +exists x; rewrite // -defKx; apply/(all_nthP 0)=> i; rewrite size_tuple => ltiE. +apply/purely_inseparable_elementP. +exists (ex_minn (fP i)); first by case: ex_minnP => n /andP[]. +by apply/seqv_sub_adjoin/map_f; rewrite mem_iota. +Qed. + +Definition separable_generator K E : L := + s2val (locked (separable_inseparable_decomposition E K)). + +Lemma separable_generator_mem E K : separable_generator K E \in E. +Proof. by rewrite /separable_generator; case: (locked _) => ? []. Qed. + +Lemma separable_generatorP E K : separable_element K (separable_generator K E). +Proof. by rewrite /separable_generator; case: (locked _) => ? []. Qed. + +Lemma separable_generator_maximal E K : + purely_inseparable <<K; separable_generator K E>> E. +Proof. by rewrite /separable_generator; case: (locked _). Qed. + +Lemma sub_adjoin_separable_generator E K : + separable K E -> (E <= <<K; separable_generator K E>>)%VS. +Proof. +move/separableP=> sepK_E; apply/subvP=> v Ev. +rewrite -separable_inseparable_element. +have /purely_inseparableP-> // := separable_generator_maximal E K. +by rewrite (separable_elementS _ (sepK_E _ Ev)) // subv_adjoin. +Qed. + +Lemma eq_adjoin_separable_generator E K : + separable K E -> (K <= E)%VS -> + E = <<K; separable_generator K E>>%VS :> {vspace _}. +Proof. +move=> sepK_E sKE; apply/eqP; rewrite eqEsubv sub_adjoin_separable_generator //. +by apply/FadjoinP/andP; rewrite sKE separable_generator_mem. +Qed. + +Lemma separable_refl K : separable K K. +Proof. by apply/separableP; apply: base_separable. Qed. + +Lemma separable_trans M K E : separable K M -> separable M E -> separable K E. +Proof. +move/sub_adjoin_separable_generator. +set x := separable_generator K M => sMKx /separableP sepM_E. +apply/separableP => w /sepM_E/(separable_elementS sMKx). +case/strong_Primitive_Element_Theorem => _ _ -> //. +exact: separable_generatorP. +Qed. + +Lemma separableS K1 K2 E2 E1 : + (K1 <= K2)%VS -> (E2 <= E1)%VS -> separable K1 E1 -> separable K2 E2. +Proof. +move=> sK12 /subvP sE21 /separableP sepK1_E1. +by apply/separableP=> y /sE21/sepK1_E1/(separable_elementS sK12). +Qed. + +Lemma separableSl K M E : (K <= M)%VS -> separable K E -> separable M E. +Proof. by move/separableS; apply. Qed. + +Lemma separableSr K M E : (M <= E)%VS -> separable K E -> separable K M. +Proof. exact: separableS. Qed. + +Lemma separable_Fadjoin_seq K rs : + all (separable_element K) rs -> separable K <<K & rs>>. +Proof. +elim/last_ind: rs => [|s x IHs] in K *. + by rewrite adjoin_nil subfield_closed separable_refl. +rewrite all_rcons adjoin_rcons => /andP[sepKx /IHs/separable_trans-> //]. +by rewrite -adjoin_separable_eq (separable_elementS _ sepKx) ?subv_adjoin_seq. +Qed. + +Lemma purely_inseparable_refl K : purely_inseparable K K. +Proof. by apply/purely_inseparableP; apply: base_inseparable. Qed. + +Lemma purely_inseparable_trans M K E : + purely_inseparable K M -> purely_inseparable M E -> purely_inseparable K E. +Proof. +have insepP := purely_inseparableP => /insepP insepK_M /insepP insepM_E. +have insepPe := purely_inseparable_elementP. +apply/insepP=> x /insepM_E/insepPe[n charLn /insepK_M/insepPe[m charLm Kxnm]]. +by apply/insepPe; exists (n * m)%N; rewrite ?exprM // pnat_mul charLn charLm. +Qed. + +End Separable. + +Implicit Arguments separable_elementP [F L K x]. +Implicit Arguments separablePn [F L K x]. +Implicit Arguments Derivation_separableP [F L K x]. +Implicit Arguments adjoin_separableP [F L K x]. +Implicit Arguments purely_inseparable_elementP [F L K x]. +Implicit Arguments separableP [F L K E]. +Implicit Arguments purely_inseparableP [F L K E]. |
