Finite fields
This file contains basic results about finite fields.
Throughout most of this file, K
denotes a finite field
and q
is notation for the cardinality of K
.
Main results
- Every finite integral domain is a field (
field_of_integral_domain
). - The unit group of a finite field is a cyclic group of order
q - 1
. (finite_field.is_cyclic
andcard_units
) sum_pow_units
: The sum ofx^i
, wherex
ranges over the units ofK
, isq-1
ifq-1 ∣ i
0
otherwise
finite_field.card
: The cardinalityq
is a power of the characteristic ofK
. Seecard'
for a variant.
Notation
Throughout most of this file, K
denotes a finite field
and q
is notation for the cardinality of K
.
theorem
finite_field.card_image_polynomial_eval
{R : Type u_2}
[integral_domain R]
[decidable_eq R]
[fintype R]
{p : polynomial R}
(hp : 0 < p.degree) :
fintype.card R ≤ (p.nat_degree) * (finset.image (λ (x : R), polynomial.eval x p) finset.univ).card
The cardinality of a field is at most n
times the cardinality of the image of a degree n
polynomial
theorem
finite_field.exists_root_sum_quadratic
{R : Type u_2}
[integral_domain R]
[fintype R]
{f g : polynomial R}
(hf2 : f.degree = 2)
(hg2 : g.degree = 2)
(hR : fintype.card R % 2 = 1) :
∃ (a b : R), polynomial.eval a f + polynomial.eval b g = 0
If f
and g
are quadratic polynomials, then the f.eval a + g.eval b = 0
has a solution.
theorem
finite_field.card_units
{K : Type u_1}
[field K]
[fintype K] :
fintype.card (units K) = fintype.card K - 1
theorem
finite_field.pow_card_sub_one_eq_one
{K : Type u_1}
[field K]
[fintype K]
(a : K)
(ha : a ≠ 0) :
a ^ (fintype.card K - 1) = 1
@[simp]
theorem
finite_field.sum_pow_lt_card_sub_one
{K : Type u_1}
[field K]
[fintype K]
(i : ℕ)
(h : i < fintype.card K - 1) :
The sum of x ^ i
as x
ranges over a finite field of cardinality q
is equal to 0
if i < q - 1
.
theorem
finite_field.expand_card
{K : Type u_1}
[field K]
[fintype K]
(f : polynomial K) :
⇑(polynomial.expand K (fintype.card K)) f = f ^ fintype.card K
The Fermat-Euler totient theorem. zmod.pow_totient
is an alternative statement
of the same theorem.
@[simp]
A variation on Fermat's little theorem. See zmod.pow_card_sub_one_eq_one
@[simp]
theorem
zmod.frobenius_zmod
(p : ℕ)
[hp : fact (nat.prime p)] :
frobenius (zmod p) p = ring_hom.id (zmod p)
@[simp]
theorem
zmod.expand_card
{p : ℕ}
[fact (nat.prime p)]
(f : polynomial (zmod p)) :
⇑(polynomial.expand (zmod p) p) f = f ^ p