Library mathcomp.field.finfield
(* (c) Copyright 2006-2016 Microsoft Corporation and Inria.
Distributed under the terms of CeCILL-B. *)
From HB Require Import structures.
From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice.
From mathcomp Require Import fintype div tuple bigop prime finset fingroup.
From mathcomp Require Import ssralg poly polydiv morphism action countalg.
From mathcomp Require Import finalg zmodp cyclic center pgroup abelian matrix.
From mathcomp Require Import mxpoly vector falgebra fieldext separable galois.
From mathcomp Require ssrnum ssrint archimedean algC cyclotomic.
Distributed under the terms of CeCILL-B. *)
From HB Require Import structures.
From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice.
From mathcomp Require Import fintype div tuple bigop prime finset fingroup.
From mathcomp Require Import ssralg poly polydiv morphism action countalg.
From mathcomp Require Import finalg zmodp cyclic center pgroup abelian matrix.
From mathcomp Require Import mxpoly vector falgebra fieldext separable galois.
From mathcomp Require ssrnum ssrint archimedean 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
finvect_type vT == alias of vT : vecType R equipped with
canonical instances for finType, finNzRing,
etc structures (including FinFieldExtType
above) for abstract vectType, falgType and
fieldExtType over a finFieldType
pPrimeCharType pcharRp == the carrier of a nzRingType R such that
pcharRp : p \in [pchar R] holds. This type has
canonical nzRingType, ..., 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
structure if R is a finFieldType
FinSplittingFieldFor nz_p == sigma-pair whose sval is a splittingFieldType
that is the splitting field for p : {poly F}
over F : finFieldType, given nz_p : p != 0
PrimePowerField pr_p k_gt0 == sigma2-triple whose s2val is a finFieldType
of characteristic p and order m = p ^ k,
given pr_p : prime p and k_gt0 : k > 0
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 pcharRp == a splittingFieldType structure
that repackages the two constructions above
Set Implicit Arguments.
Import GroupScope GRing.Theory FinRing.Theory.
Local Open Scope ring_scope.
Section FinNzRing.
Variable R : finNzRingType.
Lemma finNzRing_nontrivial : [set: R] != 1%g.
Lemma finNzRing_gt1 : 1 < #|R|.
End FinNzRing.
#[deprecated(since="mathcomp 2.4.0",
note="Use finNzRing_nontrivial instead.")]
Notation finRing_nontrivial := (finNzRing_nontrivial) (only parsing).
#[deprecated(since="mathcomp 2.4.0",
note="Use finNzRing_gt1 instead.")]
Notation finRing_gt1 := (finNzRing_gt1) (only parsing).
Section FinField.
Variable F : finFieldType.
Lemma card_finField_unit : #|[set: {unit F}]| = #|F|.-1.
Definition finField_unit x (nz_x : x != 0) :=
FinRing.unit F (etrans (unitfE x) nz_x).
Lemma expf_card x : x ^+ #|F| = x :> F.
Lemma finField_genPoly : 'X^#|F| - 'X = \prod_x ('X - x%:P) :> {poly F}.
Lemma finPcharP : {p | prime p & p \in [pchar F]}.
Lemma finField_is_abelem : is_abelem [set: F].
Lemma card_finPcharP p n : #|F| = (p ^ n)%N → prime p → p \in [pchar F].
End FinField.
#[deprecated(since="mathcomp 2.4.0", note="Use finPcharP instead.")]
Notation finCharP := (finPcharP) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use card_finPcharP instead.")]
Notation card_finCharP := (card_finPcharP) (only parsing).
Section CardVspace.
Variables (F : finFieldType) (T : finType).
Section Vector.
Variable cvT : Vector F T.
Let vT := Vector.Pack cvT.
Lemma card_vspace (V : {vspace vT}) : #|V| = (#|F| ^ \dim V)%N.
Lemma card_vspacef : #|{: vT}%VS| = #|T|.
End Vector.
Variable caT : Falgebra F T.
Let aT := Falgebra.Pack caT.
Lemma card_vspace1 : #|(1%VS : {vspace aT})| = #|F|.
End CardVspace.
Definition finvect_type (vT : Type) : predArgType := vT.
Section FinVector.
Variables (R : finNzRingType) (vT : vectType R).
Local Notation fvT := (finvect_type vT).
End FinVector.
Section FinFieldExt.
Variables (F : finFieldType) (fT : fieldExtType F).
Local Notation ffT := (finvect_type fT).
Lemma ffT_splitting_subproof : SplittingField.axiom ffT.
End FinFieldExt.
Definition FinSplittingFieldType (F : finFieldType) (fT : fieldExtType F) :=
HB.pack_for (splittingFieldType F) fT (SplittingField.on (finvect_type fT)).
Definition FinFieldExtType (F : finFieldType) (fT : fieldExtType F) :=
HB.pack_for finFieldType (FinSplittingFieldType fT)
(FinRing.Field.on (finvect_type fT)).
Arguments FinSplittingFieldType : clear implicits.
Section PrimeChar.
Variable p : nat.
Section PrimeCharRing.
Variable R0 : nzRingType.
Definition pPrimeCharType of p \in [pchar R0] : predArgType := R0.
Hypothesis pcharRp : p \in [pchar R0].
Local Notation R := (pPrimeCharType pcharRp).
Implicit Types (a b : 'F_p) (x y : R).
Definition pprimeChar_scale a x := a%:R × x.
Local Infix "*p':" := pprimeChar_scale (at level 40).
Let natrFp n : (inZp n : 'F_p)%:R = n%:R :> R.
Lemma pprimeChar_scaleA a b x : a ×p': (b ×p': x) = (a × b) ×p': x.
Lemma pprimeChar_scale1 : left_id 1 pprimeChar_scale.
Lemma pprimeChar_scaleDr : right_distributive pprimeChar_scale +%R.
Lemma pprimeChar_scaleDl x : {morph pprimeChar_scale^~ x: a b / a + b}.
Lemma pprimeChar_scaleAl (a : 'F_p) (u v : R) : a *: (u × v) = (a *: u) × v.
Lemma pprimeChar_scaleAr (a : 'F_p) (x y : R) : a *: (x × y) = x × (a *: y).
End PrimeCharRing.
Local Notation type := @pPrimeCharType.
TODO: automatize parameter inference to do all of these
Section FinNzRing.
Variables (R0 : finNzRingType) (pcharRp : p \in [pchar R0]).
Local Notation R := (type _ pcharRp).
Let pr_p : prime p.
Lemma pprimeChar_abelem : p.-abelem [set: R].
Lemma pprimeChar_pgroup : p.-group [set: R].
Lemma order_pprimeChar x : x != 0 :> R → #[x]%g = p.
Let n := logn p #|R|.
Lemma card_pprimeChar : #|R| = (p ^ n)%N.
Lemma pprimeChar_vectAxiom : Vector.axiom n R.
Lemma pprimeChar_dimf : \dim {: R : vectType 'F_p } = n.
End FinNzRing.
#[warning="-HB.no-new-instance"]
HB.instance Definition _ (R : finUnitRingType) pcharRp :=
FinRing.UnitAlgebra.on (type R pcharRp).
#[warning="-HB.no-new-instance"]
HB.instance Definition _ (R : finUnitRingType) pcharRp :=
Falgebra.on (type R pcharRp).
Section FinField.
Variables (F0 : finFieldType) (pcharFp : p \in [pchar F0]).
Local Notation F := (type _ pcharFp).
End FinField.
End PrimeChar.
#[deprecated(since="mathcomp 2.4.0", note="Use pPrimeCharType instead.")]
Notation PrimeCharType := (pPrimeCharType) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_scale instead.")]
Notation primeChar_scale := (pprimeChar_scale) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_scaleA instead.")]
Notation primeChar_scaleA := (pprimeChar_scaleA) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_scale1 instead.")]
Notation primeChar_scale1 := (pprimeChar_scale1) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_scaleDr instead.")]
Notation primeChar_scaleDr := (pprimeChar_scaleDr) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_scaleDl instead.")]
Notation primeChar_scaleDl := (pprimeChar_scaleDl) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_scaleAl instead.")]
Notation primeChar_scaleAl := (pprimeChar_scaleAl) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_scaleAr instead.")]
Notation primeChar_scaleAr := (pprimeChar_scaleAr) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_abelem instead.")]
Notation primeChar_abelem := (pprimeChar_abelem) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_pgroup instead.")]
Notation primeChar_pgroup := (pprimeChar_pgroup) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use order_pprimeChar instead.")]
Notation order_primeChar := (order_pprimeChar) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use card_pprimeChar instead.")]
Notation card_primeChar := (card_pprimeChar) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_vectAxiom instead.")]
Notation primeChar_vectAxiom := (pprimeChar_vectAxiom) (only parsing).
#[deprecated(since="mathcomp 2.4.0", note="Use pprimeChar_dimf instead.")]
Notation primeChar_dimf := (pprimeChar_dimf) (only parsing).
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}.
Fact galLgen K :
{alpha | generator 'Gal({:L} / K) alpha & ∀ x, alpha x = x ^+ order K}.
Lemma finField_galois K E : (K ≤ E)%VS → galois K E.
Lemma finField_galois_generator K E :
(K ≤ E)%VS →
{alpha | generator 'Gal(E / K) alpha
& {in E, ∀ x, alpha x = x ^+ order K}}.
End FinGalois.
Lemma Fermat's_little_theorem (L : fieldExtType F) (K : {subfield L}) a :
(a \in K) = (a ^+ order K == a).
End FinSplittingField.
Section FinFieldExists.
Section FinGalois.
Variable L : splittingFieldType F.
Implicit Types (a b : F) (x y : L) (K E : {subfield L}).
Let galL K : galois K {:L}.
Fact galLgen K :
{alpha | generator 'Gal({:L} / K) alpha & ∀ x, alpha x = x ^+ order K}.
Lemma finField_galois K E : (K ≤ E)%VS → galois K E.
Lemma finField_galois_generator K E :
(K ≤ E)%VS →
{alpha | generator 'Gal(E / K) alpha
& {in E, ∀ x, alpha x = x ^+ order K}}.
End FinGalois.
Lemma Fermat's_little_theorem (L : fieldExtType F) (K : {subfield L}) a :
(a \in K) = (a ^+ order K == a).
End FinSplittingField.
Section FinFieldExists.
While the existence of finite splitting fields and of finite fields of
arbitrary prime power order is mathematically straightforward, it is
technically challenging to formalize in Coq. The Coq typechecker performs
poorly for some of the deeply nested dependent types used in the
construction, such as polynomials over extensions of extensions of finite
fields. Any conversion in a nested structure parameter incurs a huge
overhead as it is shared across term comparison by call-by-need evalution.
The proof of FinSplittingFieldFor is contrived to mitigate this effect:
the abbreviation map_poly_extField alone divides by 3 the proof checking
time, by reducing the number of occurrences of field(Ext)Type structures
in the subgoals; the successive, apparently redundant 'suffices' localize
some of the conversions to smaller subgoals, yielding a further 8-fold
time gain. In particular, we construct the splitting field as a subtype
of a recursive construction rather than prove that the latter yields
precisely a splitting field.
The apparently redundant type annotation reduces checking time by 30%.
Let map_poly_extField (F : fieldType) (L : fieldExtType F) :=
map_poly (in_alg L) : {poly F} → {poly L}.
Local Notation "p ^%:A" := (map_poly_extField _ p)
(at level 2, format "p ^%:A") : ring_scope.
Lemma FinSplittingFieldFor (F : finFieldType) (p : {poly F}) :
p != 0 → {L : splittingFieldType F | splittingFieldFor 1 p^%:A {:L}}.
Lemma pPrimePowerField p k (m := (p ^ k)%N) :
prime p → 0 < k → {Fm : finFieldType | p \in [pchar Fm] & #|Fm| = m}.
End FinFieldExists.
#[deprecated(since="mathcomp 2.4.0", note="Use pPrimePowerField instead.")]
Notation PrimePowerField := (pPrimePowerField) (only parsing).
Section FinDomain.
Import order ssrnum ssrint archimedean algC cyclotomic Order.TTheory Num.Theory.
Local Infix "%|" := dvdn. (* Hide polynomial divisibility. *)
Variable R : finUnitRingType.
Hypothesis domR : GRing.integral_domain_axiom R.
Implicit Types x y : R.
Let lregR x : x != 0 → GRing.lreg x.
Lemma finDomain_field : GRing.field_axiom R.
map_poly (in_alg L) : {poly F} → {poly L}.
Local Notation "p ^%:A" := (map_poly_extField _ p)
(at level 2, format "p ^%:A") : ring_scope.
Lemma FinSplittingFieldFor (F : finFieldType) (p : {poly F}) :
p != 0 → {L : splittingFieldType F | splittingFieldFor 1 p^%:A {:L}}.
Lemma pPrimePowerField p k (m := (p ^ k)%N) :
prime p → 0 < k → {Fm : finFieldType | p \in [pchar Fm] & #|Fm| = m}.
End FinFieldExists.
#[deprecated(since="mathcomp 2.4.0", note="Use pPrimePowerField instead.")]
Notation PrimePowerField := (pPrimePowerField) (only parsing).
Section FinDomain.
Import order ssrnum ssrint archimedean algC cyclotomic Order.TTheory Num.Theory.
Local Infix "%|" := dvdn. (* Hide polynomial divisibility. *)
Variable R : finUnitRingType.
Hypothesis domR : GRing.integral_domain_axiom R.
Implicit Types x y : R.
Let lregR x : x != 0 → GRing.lreg x.
Lemma finDomain_field : GRing.field_axiom R.
This is Witt's proof of Wedderburn's little theorem.
Theorem finDomain_mulrC : @commutative R R *%R.
Definition FinDomainFieldType : finFieldType :=
let cC := GRing.PzRing_hasCommutativeMul.Build R finDomain_mulrC in
let cR : comUnitRingType := HB.pack R cC in
let iC := GRing.ComUnitRing_isIntegral.Build cR domR in
let iR : finIdomainType := HB.pack cR iC in
let fC := GRing.UnitRing_isField.Build iR finDomain_field in
HB.pack iR fC.
Definition FinDomainSplittingFieldType_pchar p (pcharRp : p \in [pchar R]) :=
SplittingField.clone 'F_p R (@pPrimeCharType p FinDomainFieldType pcharRp).
End FinDomain.
#[deprecated(since="mathcomp 2.4.0", note="Use FinDomainSplittingFieldType_pchar instead.")]
Notation FinDomainSplittingFieldType := (FinDomainSplittingFieldType_pchar) (only parsing).
Definition FinDomainFieldType : finFieldType :=
let cC := GRing.PzRing_hasCommutativeMul.Build R finDomain_mulrC in
let cR : comUnitRingType := HB.pack R cC in
let iC := GRing.ComUnitRing_isIntegral.Build cR domR in
let iR : finIdomainType := HB.pack cR iC in
let fC := GRing.UnitRing_isField.Build iR finDomain_field in
HB.pack iR fC.
Definition FinDomainSplittingFieldType_pchar p (pcharRp : p \in [pchar R]) :=
SplittingField.clone 'F_p R (@pPrimeCharType p FinDomainFieldType pcharRp).
End FinDomain.
#[deprecated(since="mathcomp 2.4.0", note="Use FinDomainSplittingFieldType_pchar instead.")]
Notation FinDomainSplittingFieldType := (FinDomainSplittingFieldType_pchar) (only parsing).