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(NumberTheory/LSeries/Dirichlet): new file, material on specific L-series #11712

Closed
wants to merge 41 commits into from
Closed
Show file tree
Hide file tree
Changes from 30 commits
Commits
Show all changes
41 commits
Select commit Hold shift + click to select a range
fbd32b2
add file NumberTheory/LSeries/Dirichlet
MichaelStollBayreuth Mar 26, 2024
5542cbb
copy over stuff
MichaelStollBayreuth Mar 26, 2024
71cfdd0
move lemma to NumberTheory.Divisors
MichaelStollBayreuth Mar 26, 2024
2c128a1
golf in Convolution.lean
MichaelStollBayreuth Mar 26, 2024
8188b37
move zero and delta to Basic.lean
MichaelStollBayreuth Mar 26, 2024
653fe40
rename →
MichaelStollBayreuth Mar 26, 2024
b4cf80e
move stuff from Basic to Dirichlet and golf
MichaelStollBayreuth Mar 26, 2024
f89b4f5
change name, golf
MichaelStollBayreuth Mar 26, 2024
c2853c7
add Tags to module docstring, remove TODO in Basic
MichaelStollBayreuth Mar 26, 2024
6fe09de
fix imports
MichaelStollBayreuth Mar 26, 2024
ee2abb2
add versions for ↗χ * ↗Λ, plus some API
MichaelStollBayreuth Mar 27, 2024
74b9706
add versions for ↗χ * ↗Λ, plus some API
MichaelStollBayreuth Mar 27, 2024
4f2d21f
fix typo
MichaelStollBayreuth Mar 27, 2024
c843710
format and golf
MichaelStollBayreuth Mar 27, 2024
ca1c189
sort imports
MichaelStollBayreuth Mar 29, 2024
052937c
suggestions from review
MichaelStollBayreuth Mar 29, 2024
9d1808f
move mul_delta up
MichaelStollBayreuth Mar 29, 2024
33face5
zeta_LSeriesSummable_iff -> LSeriesSummable_zeta_iff
MichaelStollBayreuth Mar 29, 2024
fb311d6
..._ne_zero -> ..._ne_zero_of_one_lt_re
MichaelStollBayreuth Mar 29, 2024
33e9109
CharP.cast_eq_zero -> cast_zero
MichaelStollBayreuth Mar 29, 2024
47bc8b1
add deprecation
MichaelStollBayreuth Mar 29, 2024
b4b3a7d
add {left|right}_ne_zero_of_mem_divisorsAntidiagonal
MichaelStollBayreuth Mar 29, 2024
424b667
appease linter
MichaelStollBayreuth Mar 29, 2024
c5ebcaa
two more CharP.cast_eq_zero -> cast_zero
MichaelStollBayreuth Mar 29, 2024
4e42661
suggestion from review, golf
MichaelStollBayreuth Mar 29, 2024
f8e8608
golf one_convolution_eq_zeta_convolution and companion
MichaelStollBayreuth Mar 29, 2024
3dbeb3e
new lemma const_one_eq_zeta; use to golf
MichaelStollBayreuth Mar 29, 2024
a5e6b7b
change congr arguments to {n} (_ : n ≠ 0)
MichaelStollBayreuth Mar 29, 2024
5d9c1bf
appease linter
MichaelStollBayreuth Mar 29, 2024
ada4ece
minor fixes in docstrings
MichaelStollBayreuth Mar 29, 2024
db0e91e
open Nat in Convolution and remove 'Nat.'
MichaelStollBayreuth Apr 1, 2024
67776c4
suggestions from code review
MichaelStollBayreuth Apr 1, 2024
a36f73c
suggestions from code review
MichaelStollBayreuth Apr 1, 2024
64bd5a7
one more
MichaelStollBayreuth Apr 1, 2024
fa1493a
Merge remote-tracking branch 'origin/master' into MS_LSeries_Dirichlet
MichaelStollBayreuth Apr 2, 2024
9866260
golf using 'abs_moebius_le_one'
MichaelStollBayreuth Apr 2, 2024
78af1a4
introduce 'DirichletCharacter.LSeriesSummable_of_one_lt_re' and use i…
MichaelStollBayreuth Apr 2, 2024
3f3eeee
make docstrings more precise
MichaelStollBayreuth Apr 2, 2024
713c687
fix typo; spacing
MichaelStollBayreuth Apr 2, 2024
7548c8b
suggestion from review
MichaelStollBayreuth Apr 3, 2024
a831f03
more of the same
MichaelStollBayreuth Apr 3, 2024
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2972,6 +2972,7 @@ import Mathlib.NumberTheory.LSeries.Basic
import Mathlib.NumberTheory.LSeries.Convergence
import Mathlib.NumberTheory.LSeries.Convolution
import Mathlib.NumberTheory.LSeries.Deriv
import Mathlib.NumberTheory.LSeries.Dirichlet
import Mathlib.NumberTheory.LSeries.Linearity
import Mathlib.NumberTheory.LegendreSymbol.AddCharacter
import Mathlib.NumberTheory.LegendreSymbol.Basic
Expand Down
13 changes: 13 additions & 0 deletions Mathlib/NumberTheory/Divisors.lean
Original file line number Diff line number Diff line change
Expand Up @@ -132,6 +132,19 @@ theorem mem_divisorsAntidiagonal {x : ℕ × ℕ} :
Nat.le_mul_of_pos_left _ (Nat.pos_of_ne_zero h.1)⟩
#align nat.mem_divisors_antidiagonal Nat.mem_divisorsAntidiagonal

lemma ne_zero_of_mem_divisorsAntidiagonal {p : ℕ × ℕ} (hp : p ∈ n.divisorsAntidiagonal) :
MichaelStollBayreuth marked this conversation as resolved.
Show resolved Hide resolved
p.1 ≠ 0 ∧ p.2 ≠ 0 := by
obtain ⟨hp₁, hp₂⟩ := Nat.mem_divisorsAntidiagonal.mp hp
exact mul_ne_zero_iff.mp (hp₁.symm ▸ hp₂)

lemma left_ne_zero_of_mem_divisorsAntidiagonal {p : ℕ × ℕ} (hp : p ∈ n.divisorsAntidiagonal) :
p.1 ≠ 0 :=
(ne_zero_of_mem_divisorsAntidiagonal hp).1

lemma right_ne_zero_of_mem_divisorsAntidiagonal {p : ℕ × ℕ} (hp : p ∈ n.divisorsAntidiagonal) :
p.2 ≠ 0 :=
(ne_zero_of_mem_divisorsAntidiagonal hp).2

-- Porting note: Redundant binder annotation update
-- variable {n}

Expand Down
91 changes: 58 additions & 33 deletions Mathlib/NumberTheory/LSeries/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Aaron Anderson, Michael Stoll
-/
import Mathlib.Analysis.PSeries
import Mathlib.NumberTheory.ArithmeticFunction
import Mathlib.Analysis.NormedSpace.FiniteDimension

#align_import number_theory.l_series from "leanprover-community/mathlib"@"32253a1a1071173b33dc7d6a218cf722c6feb514"
Expand Down Expand Up @@ -46,11 +45,6 @@ to `N` and `R` coerces to `ℂ`) as arguments to `LSeries` etc.
## Tags

L-series

## TODO

* Move `LSeriesSummable.one_iff_one_lt_re` and `zeta_LSeriesSummable_iff_one_lt_r`
to a new file on L-series of specific functions
-/

open scoped BigOperators
Expand Down Expand Up @@ -86,7 +80,7 @@ lemma term_of_ne_zero {n : ℕ} (hn : n ≠ 0) (f : ℕ → ℂ) (s : ℂ) :
term f s n = f n / n ^ s :=
if_neg hn

lemma term_congr {f g : ℕ → ℂ} (h : ∀ n ≠ 0, f n = g n) (s : ℂ) (n : ℕ) :
lemma term_congr {f g : ℕ → ℂ} (h : ∀ {n} (_ : n ≠ 0), f n = g n) (s : ℂ) (n : ℕ) :
MichaelStollBayreuth marked this conversation as resolved.
Show resolved Hide resolved
term f s n = term g s n := by
rcases eq_or_ne n 0 with hn | hn <;> simp [hn, h]

Expand Down Expand Up @@ -130,7 +124,7 @@ def LSeries (f : ℕ → ℂ) (s : ℂ) : ℂ :=
∑' n, term f s n
#align nat.arithmetic_function.l_series LSeries

lemma LSeries_congr {f g : ℕ → ℂ} (s : ℂ) (h : ∀ n ≠ 0, f n = g n) :
lemma LSeries_congr {f g : ℕ → ℂ} (s : ℂ) (h : ∀ {n} (_ : n ≠ 0), f n = g n) :
LSeries f s = LSeries g s :=
tsum_congr <| term_congr h s

Expand All @@ -139,7 +133,7 @@ def LSeriesSummable (f : ℕ → ℂ) (s : ℂ) : Prop :=
Summable (term f s)
#align nat.arithmetic_function.l_series_summable LSeriesSummable

lemma LSeriesSummable_congr {f g : ℕ → ℂ} (s : ℂ) (h : ∀ n ≠ 0, f n = g n) :
lemma LSeriesSummable_congr {f g : ℕ → ℂ} (s : ℂ) (h : ∀ {n} (_ : n ≠ 0), f n = g n) :
LSeriesSummable f s ↔ LSeriesSummable g s :=
summable_congr <| term_congr h s

Expand Down Expand Up @@ -197,7 +191,7 @@ lemma LSeriesHasSum_iff {f : ℕ → ℂ} {s a : ℂ} :
LSeriesHasSum f s a ↔ LSeriesSummable f s ∧ LSeries f s = a :=
⟨fun H ↦ ⟨H.LSeriesSummable, H.LSeries_eq⟩, fun ⟨H₁, H₂⟩ ↦ H₂ ▸ H₁.LSeriesHasSum⟩

lemma LSeriesHasSum_congr {f g : ℕ → ℂ} (s a : ℂ) (h : ∀ n ≠ 0, f n = g n) :
lemma LSeriesHasSum_congr {f g : ℕ → ℂ} (s a : ℂ) (h : ∀ {n} (_ : n ≠ 0), f n = g n) :
LSeriesHasSum f s a ↔ LSeriesHasSum g s a := by
simp only [LSeriesHasSum_iff, LSeriesSummable_congr s h, LSeries_congr s h]

Expand All @@ -211,6 +205,9 @@ theorem LSeriesSummable_iff_of_re_eq_re {f : ℕ → ℂ} {s s' : ℂ} (h : s.re
⟨fun H ↦ H.of_re_le_re h.le, fun H ↦ H.of_re_le_re h.symm.le⟩
#align nat.arithmetic_function.l_series_summable_iff_of_re_eq_re LSeriesSummable_iff_of_re_eq_re

/-- The indicator function of `{1} ⊆ ℕ` with values in `ℂ`. -/
def LSeries.delta (n : ℕ) : ℂ :=
if n = 1 then 1 else 0

/-!
### Notation
Expand All @@ -226,6 +223,57 @@ or `↗f` when `f : ArithmeticFunction R` or simply `f : N → R` with a coercio
as an argument to `LSeries`, `LSeriesHasSum`, `LSeriesSummable` etc. -/
scoped[LSeries.notation] notation:max "↗" f:max => fun n : ℕ ↦ (f n : ℂ)

@[inherit_doc]
scoped[LSeries.notation] notation "δ" => delta

/-!
### LSeries of 0 and δ
-/

@[simp]
lemma LSeries_zero : LSeries 0 = 0 := by
ext
simp only [LSeries, LSeries.term, Pi.zero_apply, zero_div, ite_self, tsum_zero]

section delta

open scoped LSeries.notation

namespace LSeries

open Nat Complex

lemma term_delta (s : ℂ) (n : ℕ) : term δ s n = if n = 1 then 1 else 0 := by
rcases eq_or_ne n 0 with rfl | hn
· simp only [term_zero, zero_ne_one, ↓reduceIte]
· simp only [ne_eq, hn, not_false_eq_true, term_of_ne_zero, delta]
rcases eq_or_ne n 1 with rfl | hn'
· simp only [↓reduceIte, cast_one, one_cpow, ne_eq, one_ne_zero, not_false_eq_true, div_self]
· simp only [hn', ↓reduceIte, zero_div]

lemma mul_delta_eq_smul_delta {f : ℕ → ℂ} : f * δ = f 1 • δ := by
ext n
simp only [Pi.mul_apply, delta, mul_ite, mul_one, mul_zero, Pi.smul_apply, smul_eq_mul]
split_ifs with hn <;> simp only [hn]

lemma mul_delta {f : ℕ → ℂ} (h : f 1 = 1) : f * δ = δ := by
rw [mul_delta_eq_smul_delta, h, one_smul]

lemma delta_mul_eq_smul_delta {f : ℕ → ℂ} : δ * f = f 1 • δ :=
mul_comm δ f ▸ mul_delta_eq_smul_delta

lemma delta_mul {f : ℕ → ℂ} (h : f 1 = 1) : δ * f = δ :=
mul_comm δ f ▸ mul_delta h

end LSeries

/-- The L-series of `δ` is the constant function `1`. -/
lemma LSeries_delta : LSeries δ = 1 := by
ext
simp only [LSeries, LSeries.term_delta, tsum_ite_eq, Pi.one_apply]

end delta


/-!
### Criteria for and consequences of summability of L-series
Expand Down Expand Up @@ -322,26 +370,3 @@ theorem LSeriesSummable_of_bounded_of_one_lt_real {f : ℕ → ℂ} {m : ℝ}
LSeriesSummable f s :=
LSeriesSummable_of_bounded_of_one_lt_re h <| by simp only [ofReal_re, hs]
#align nat.arithmetic_function.l_series_summable_of_bounded_of_one_lt_real LSeriesSummable_of_bounded_of_one_lt_real

-- TODO: Move this to a separate file on concrete L-series

open Set in
/-- The `LSeries` with all coefficients `1` converges at `s` if and only if `re s > 1`. -/
theorem LSeriesSummable.one_iff_one_lt_re {s : ℂ} : LSeriesSummable 1 s ↔ 1 < s.re := by
rw [← LSeriesSummable_iff_of_re_eq_re (Complex.ofReal_re s.re), LSeriesSummable,
← summable_norm_iff, ← Real.summable_one_div_nat_rpow]
simp_rw [← Finite.summable_compl_iff (finite_singleton 0), summable_subtype_iff_indicator]
refine summable_congr fun n ↦ ?_
by_cases hn : n ∈ ({0}ᶜ :Set ℕ)
· simp only [indicator_of_mem hn, norm_term_eq]
simp only [show n ≠ 0 from hn, ↓reduceIte, Pi.one_apply, norm_one, ofReal_re]
· simp only [indicator_of_not_mem hn]

open scoped ArithmeticFunction in
/-- The `LSeries` associated to the arithmetic function `ζ` converges at `s` if and only if
`re s > 1`. -/
theorem zeta_LSeriesSummable_iff_one_lt_re {s : ℂ} : LSeriesSummable (ζ ·) s ↔ 1 < s.re := by
have (n : ℕ) (hn : n ≠ 0) : ζ n = (1 : ℕ → ℂ) n := by
simp only [ArithmeticFunction.zeta_apply, hn, ↓reduceIte, Nat.cast_one, Pi.one_apply]
exact (LSeriesSummable_congr s this).trans <| LSeriesSummable.one_iff_one_lt_re
#align nat.arithmetic_function.zeta_l_series_summable_iff_one_lt_re zeta_LSeriesSummable_iff_one_lt_re
2 changes: 1 addition & 1 deletion Mathlib/NumberTheory/LSeries/Convergence.lean
Original file line number Diff line number Diff line change
Expand Up @@ -26,7 +26,7 @@ when `re s < x`. -/
noncomputable def LSeries.abscissaOfAbsConv (f : ℕ → ℂ) : EReal :=
sInf <| Real.toEReal '' {x : ℝ | LSeriesSummable f x}

lemma LSeries.abscissaOfAbsConv_congr {f g : ℕ → ℂ} (h : ∀ n ≠ 0, f n = g n) :
lemma LSeries.abscissaOfAbsConv_congr {f g : ℕ → ℂ} (h : ∀ {n} (_ : n ≠ 0), f n = g n) :
abscissaOfAbsConv f = abscissaOfAbsConv g :=
congr_arg sInf <| congr_arg _ <| Set.ext fun x ↦ LSeriesSummable_congr x h

Expand Down
27 changes: 20 additions & 7 deletions Mathlib/NumberTheory/LSeries/Convolution.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Michael Stoll
-/
import Mathlib.Analysis.InnerProductSpace.Basic
import Mathlib.NumberTheory.LSeries.Convergence
import Mathlib.Analysis.Normed.Field.InfiniteSum
import Mathlib.NumberTheory.ArithmeticFunction
import Mathlib.NumberTheory.LSeries.Convergence

/-!
# Dirichlet convolution of sequences and products of L-series
Expand Down Expand Up @@ -36,6 +37,14 @@ def toArithmeticFunction {R : Type*} [Zero R] (f : ℕ → R) : ArithmeticFuncti
toFun n := if n = 0 then 0 else f n
map_zero' := rfl

lemma toArithmeticFunction_congr {R : Type*} [Zero R] {f f' : ℕ → R}
(h : ∀ {n} (_ : n ≠ 0), f n = f' n) :
toArithmeticFunction f = toArithmeticFunction f' := by
ext ⟨- | _⟩
· simp only [Nat.zero_eq, ArithmeticFunction.map_zero]
· simp only [toArithmeticFunction, ArithmeticFunction.coe_mk, Nat.succ_ne_zero, ↓reduceIte,
ne_eq, not_false_eq_true, h]

/-- If we consider an arithmetic function just as a function and turn it back into an
arithmetic function, it is the same as before. -/
@[simp]
Expand All @@ -54,6 +63,11 @@ noncomputable def LSeries.convolution {R : Type*} [Semiring R] (f g : ℕ → R)
@[inherit_doc]
scoped[LSeries.notation] infixl:70 " ⍟ " => LSeries.convolution

lemma LSeries.convolution_congr {R : Type*} [Semiring R] {f f' g g' : ℕ → R}
(hf : ∀ {n} (_ : n ≠ 0), f n = f' n) (hg : ∀ {n} (_ : n ≠ 0), g n = g' n) :
f ⍟ g = f' ⍟ g' := by
simp only [convolution, toArithmeticFunction_congr hf, toArithmeticFunction_congr hg]

/-- The product of two arithmetic functions defines the same function as the Dirichlet convolution
of the functions defined by them. -/
lemma ArithmeticFunction.coe_mul {R : Type*} [Semiring R] (f g : ArithmeticFunction R) :
Expand All @@ -68,8 +82,7 @@ lemma convolution_def {R : Type*} [Semiring R] (f g : ℕ → R) :
simp only [convolution, toArithmeticFunction, ArithmeticFunction.mul_apply,
ArithmeticFunction.coe_mk, mul_ite, mul_zero, ite_mul, zero_mul]
refine Finset.sum_congr rfl fun p hp ↦ ?_
obtain ⟨hp₁, hp₂⟩ := Nat.mem_divisorsAntidiagonal.mp hp
obtain ⟨h₁, h₂⟩ := mul_ne_zero_iff.mp (hp₁.symm ▸ hp₂)
obtain ⟨h₁, h₂⟩ := Nat.ne_zero_of_mem_divisorsAntidiagonal hp
simp only [h₂, ↓reduceIte, h₁]

@[simp]
Expand All @@ -81,6 +94,7 @@ lemma convolution_map_zero {R : Type*} [Semiring R] (f g : ℕ → R) : (f ⍟ g
### Multiplication of L-series
-/

open Nat in
/-- We give an expression of the `LSeries.term` of the convolution of two functions
in terms of a sum over `Nat.divisorsAntidiagonal`. -/
lemma term_convolution (f g : ℕ → ℂ) (s : ℂ) (n : ℕ) :
Expand All @@ -90,10 +104,9 @@ lemma term_convolution (f g : ℕ → ℂ) (s : ℂ) (n : ℕ) :
-- now `n ≠ 0`
rw [term_of_ne_zero hn, convolution_def, Finset.sum_div]
refine Finset.sum_congr rfl fun p hp ↦ ?_
obtain ⟨hp, hn₀⟩ := Nat.mem_divisorsAntidiagonal.mp hp
have ⟨hp₁, hp₂⟩ := mul_ne_zero_iff.mp <| hp.symm ▸ hn₀
rw [term_of_ne_zero hp₁ f s, term_of_ne_zero hp₂ g s, mul_comm_div, div_div,
← mul_div_assoc, ← natCast_mul_natCast_cpow, ← Nat.cast_mul, mul_comm p.2, hp]
have ⟨hp₁, hp₂⟩ := Nat.ne_zero_of_mem_divisorsAntidiagonal hp
MichaelStollBayreuth marked this conversation as resolved.
Show resolved Hide resolved
rw [term_of_ne_zero hp₁ f s, term_of_ne_zero hp₂ g s, mul_comm_div, div_div, ← mul_div_assoc,
← natCast_mul_natCast_cpow, ← cast_mul, mul_comm p.2, (mem_divisorsAntidiagonal.mp hp).1]

open Set in
/-- We give an expression of the `LSeries.term` of the convolution of two functions
Expand Down
Loading
Loading