Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat: bounds for Dirichlet characters #8449

Closed
wants to merge 8 commits into from
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2659,6 +2659,7 @@ import Mathlib.NumberTheory.Cyclotomic.Rat
import Mathlib.NumberTheory.Dioph
import Mathlib.NumberTheory.DiophantineApproximation
import Mathlib.NumberTheory.DirichletCharacter.Basic
import Mathlib.NumberTheory.DirichletCharacter.Bounds
import Mathlib.NumberTheory.Divisors
import Mathlib.NumberTheory.EulerProduct.Basic
import Mathlib.NumberTheory.FLT.Basic
Expand Down
6 changes: 6 additions & 0 deletions Mathlib/Algebra/GroupPower/Order.lean
Original file line number Diff line number Diff line change
Expand Up @@ -223,6 +223,12 @@ theorem one_le_pow_iff_of_nonneg {a : R} (ha : 0 ≤ a) {n : ℕ} (hn : n ≠ 0)
exact mt fun h => pow_lt_one ha h hn
#align one_le_pow_iff_of_nonneg one_le_pow_iff_of_nonneg

lemma pow_eq_one_iff_of_nonneg {x : R} (hx : 0 ≤ x)
MichaelStollBayreuth marked this conversation as resolved.
Show resolved Hide resolved
{n : ℕ} (hn : n ≠ 0) : x ^ n = 1 ↔ x = 1 :=
⟨fun h ↦ le_antisymm ((pow_le_one_iff_of_nonneg hx hn).mp h.le)
((one_le_pow_iff_of_nonneg hx hn).mp h.ge),
fun h ↦ by rw [h]; exact one_pow _⟩

theorem one_lt_pow_iff_of_nonneg {a : R} (ha : 0 ≤ a) {n : ℕ} (hn : n ≠ 0) : 1 < a ^ n ↔ 1 < a :=
lt_iff_lt_of_le_iff_le (pow_le_one_iff_of_nonneg ha hn)
#align one_lt_pow_iff_of_nonneg one_lt_pow_iff_of_nonneg
Expand Down
5 changes: 5 additions & 0 deletions Mathlib/FieldTheory/Finite/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -479,6 +479,11 @@ theorem Nat.ModEq.pow_totient {x n : ℕ} (h : Nat.Coprime x n) : x ^ φ n ≡ 1
coe_unitOfCoprime, Units.val_pow_eq_pow_val]
#align nat.modeq.pow_totient Nat.ModEq.pow_totient

/-- For each `n ≥ 0`, the unit group of `ZMod n` is finite. -/
instance finiteZModUnits : (n : ℕ) → Finite (ZMod n)ˣ
MichaelStollBayreuth marked this conversation as resolved.
Show resolved Hide resolved
| 0 => Finite.of_fintype ℤˣ
| _ + 1 => instFiniteUnits

ocfnash marked this conversation as resolved.
Show resolved Hide resolved
section

variable {V : Type*} [Fintype K] [DivisionRing K] [AddCommGroup V] [Module K V]
Expand Down
36 changes: 36 additions & 0 deletions Mathlib/NumberTheory/DirichletCharacter/Bounds.lean
MichaelStollBayreuth marked this conversation as resolved.
Show resolved Hide resolved
Original file line number Diff line number Diff line change
@@ -0,0 +1,36 @@
/-
Copyright (c) 2023 Michael Stoll. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Michael Stoll
-/
import Mathlib.Analysis.Normed.Field.Basic
import Mathlib.FieldTheory.Finite.Basic
import Mathlib.NumberTheory.DirichletCharacter.Basic

/-!
# Bounds for values of Dirichlet characters

We consider Dirichlet characters `χ` with values in a normed field `F`.

We show that `‖χ a‖ = 1` if `a` is a unit and `‖χ a‖ ≤ 1` in general.
-/

variable {F : Type*} [NormedField F] {n : ℕ} (χ : DirichletCharacter F n)

namespace DirichletCharacter

/-- The value at a unit of a Dirichlet character with target a normed field has norm `1`. -/
lemma unit_norm_eq_one (a : (ZMod n)ˣ) : ‖χ a‖ = 1 := by
MichaelStollBayreuth marked this conversation as resolved.
Show resolved Hide resolved
refine (pow_eq_one_iff_of_nonneg (norm_nonneg _) (Nat.card_pos (α := (ZMod n)ˣ)).ne').mp ?_
rw [← norm_pow, ← map_pow, ← Units.val_pow_eq_pow_val, pow_card_eq_one', Units.val_one, map_one,
norm_one]

/-- The values of a Dirichlet character with target a normed field have norm bounded by `1`. -/
lemma norm_le_one (a : ZMod n) : ‖χ a‖ ≤ 1 := by
by_cases h : IsUnit a
· exact (χ.unit_norm_eq_one h.unit).le
· rw [χ.map_nonunit h, norm_zero]
exact zero_le_one


end DirichletCharacter