mathlib3 documentation

number_theory.legendre_symbol.quadratic_char.basic

Quadratic characters of finite fields #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

This file defines the quadratic character on a finite field F and proves some basic statements about it.

Tags #

quadratic character

Definition of the quadratic character #

We define the quadratic character of a finite field F with values in ℤ.

Define the quadratic character with values in ℤ on a monoid with zero α. It takes the value zero at zero; for non-zero argument a : α, it is 1 if a is a square, otherwise it is -1.

This only deserves the name "character" when it is multiplicative, e.g., when α is a finite field. See quadratic_char_fun_mul.

We will later define quadratic_char to be a multiplicative character of type mul_char F ℤ, when the domain is a finite field F.

Equations

Basic properties of the quadratic character #

We prove some properties of the quadratic character. We work with a finite field F here. The interesting case is when the characteristic of F is odd.

theorem quadratic_char_fun_eq_zero_iff {F : Type u_1} [field F] [fintype F] [decidable_eq F] {a : F} :

Some basic API lemmas

@[simp]
theorem quadratic_char_fun_zero {F : Type u_1} [field F] [fintype F] [decidable_eq F] :
@[simp]
theorem quadratic_char_fun_one {F : Type u_1} [field F] [fintype F] [decidable_eq F] :
theorem quadratic_char_fun_eq_one_of_char_two {F : Type u_1} [field F] [fintype F] [decidable_eq F] (hF : ring_char F = 2) {a : F} (ha : a 0) :

If ring_char F = 2, then quadratic_char_fun F takes the value 1 on nonzero elements.

theorem quadratic_char_fun_eq_pow_of_char_ne_two {F : Type u_1} [field F] [fintype F] [decidable_eq F] (hF : ring_char F 2) {a : F} (ha : a 0) :
quadratic_char_fun F a = ite (a ^ (fintype.card F / 2) = 1) 1 (-1)

If ring_char F is odd, then quadratic_char_fun F a can be computed in terms of a ^ (fintype.card F / 2).

The quadratic character is multiplicative.

def quadratic_char (F : Type u_1) [field F] [fintype F] [decidable_eq F] :

The quadratic character as a multiplicative character.

Equations
@[simp]
theorem quadratic_char_apply (F : Type u_1) [field F] [fintype F] [decidable_eq F] (a : F) :
theorem quadratic_char_eq_zero_iff {F : Type u_1} [field F] [fintype F] [decidable_eq F] {a : F} :
(quadratic_char F) a = 0 a = 0

The value of the quadratic character on a is zero iff a = 0.

@[simp]
theorem quadratic_char_zero {F : Type u_1} [field F] [fintype F] [decidable_eq F] :
theorem quadratic_char_one_iff_is_square {F : Type u_1} [field F] [fintype F] [decidable_eq F] {a : F} (ha : a 0) :

For nonzero a : F, quadratic_char F a = 1 ↔ is_square a.

theorem quadratic_char_sq_one' {F : Type u_1} [field F] [fintype F] [decidable_eq F] {a : F} (ha : a 0) :
(quadratic_char F) (a ^ 2) = 1

The quadratic character takes the value 1 on nonzero squares.

theorem quadratic_char_sq_one {F : Type u_1} [field F] [fintype F] [decidable_eq F] {a : F} (ha : a 0) :

The square of the quadratic character on nonzero arguments is 1.

theorem quadratic_char_dichotomy {F : Type u_1} [field F] [fintype F] [decidable_eq F] {a : F} (ha : a 0) :

The quadratic character is 1 or -1 on nonzero arguments.

theorem quadratic_char_eq_neg_one_iff_not_one {F : Type u_1} [field F] [fintype F] [decidable_eq F] {a : F} (ha : a 0) :

The quadratic character is 1 or -1 on nonzero arguments.

For a : F, quadratic_char F a = -1 ↔ ¬ is_square a.

theorem quadratic_char_exists_neg_one {F : Type u_1} [field F] [fintype F] [decidable_eq F] (hF : ring_char F 2) :
(a : F), (quadratic_char F) a = -1

If F has odd characteristic, then quadratic_char F takes the value -1.

theorem quadratic_char_eq_one_of_char_two {F : Type u_1} [field F] [fintype F] [decidable_eq F] (hF : ring_char F = 2) {a : F} (ha : a 0) :

If ring_char F = 2, then quadratic_char F takes the value 1 on nonzero elements.

theorem quadratic_char_eq_pow_of_char_ne_two {F : Type u_1} [field F] [fintype F] [decidable_eq F] (hF : ring_char F 2) {a : F} (ha : a 0) :
(quadratic_char F) a = ite (a ^ (fintype.card F / 2) = 1) 1 (-1)

If ring_char F is odd, then quadratic_char F a can be computed in terms of a ^ (fintype.card F / 2).

theorem quadratic_char_eq_pow_of_char_ne_two' {F : Type u_1} [field F] [fintype F] [decidable_eq F] (hF : ring_char F 2) (a : F) :

The quadratic character is quadratic as a multiplicative character.

The quadratic character is nontrivial as a multiplicative character when the domain has odd characteristic.

theorem quadratic_char_card_sqrts {F : Type u_1} [field F] [fintype F] [decidable_eq F] (hF : ring_char F 2) (a : F) :
({x : F | x ^ 2 = a}.to_finset.card) = (quadratic_char F) a + 1

The number of solutions to x^2 = a is determined by the quadratic character.

theorem quadratic_char_sum_zero {F : Type u_1} [field F] [fintype F] [decidable_eq F] (hF : ring_char F 2) :
finset.univ.sum (λ (a : F), (quadratic_char F) a) = 0

The sum over the values of the quadratic character is zero when the characteristic is odd.

Special values of the quadratic character #

We express quadratic_char F (-1) in terms of χ₄.

The value of the quadratic character at -1

-1 is a square in F iff #F is not congruent to 3 mod 4.