Documentation

SpherePacking.ModularForms.Cauchylems

theorem cc (f : ) (hc : CauchySeq fun (N : ) => mFinset.Icc (-N) N, f m) (hs : ∀ (n : ), f n = f (-n)) :
theorem sum_Icc_eq_sum_Ico_succ {α : Type u_1} [AddCommMonoid α] (f : α) {l u : } (h : l u) :
mFinset.Icc l u, f m = mFinset.Ico l u, f m + f u
theorem auxl2 (a b c : ) :
theorem CauchySeq_Icc_iff_CauchySeq_Ico (f : ) (hs : ∀ (n : ), f n = f (-n)) (hc : CauchySeq fun (N : ) => mFinset.Icc (-N) N, f m) :
CauchySeq fun (N : ) => mFinset.Ico (-N) N, f m
theorem extracted_2 (z : UpperHalfPlane) (b : ) :
CauchySeq fun (N : ) => nFinset.Ico (-N) N, 1 / ((b * z + n) ^ 2 * (b * z + n + 1))
theorem extracted_2_δ (z : UpperHalfPlane) (b : ) :
CauchySeq fun (N : ) => nFinset.Ico (-N) N, (1 / ((b * z + n) ^ 2 * (b * z + n + 1)) + δ b n)
theorem telescope_aux (z : UpperHalfPlane) (m : ) (b : ) :
nFinset.Ico (-b) b, (1 / (m * z + n) - 1 / (m * z + n + 1)) = 1 / (m * z - b) - 1 / (m * z + b)
theorem tendstozero_inv_linear (z : UpperHalfPlane) (b : ) :
Filter.Tendsto (fun (d : ) => 1 / (b * z + d)) Filter.atTop (nhds 0)
theorem tendstozero_inv_linear_neg (z : UpperHalfPlane) (b : ) :
Filter.Tendsto (fun (d : ) => 1 / (b * z - d)) Filter.atTop (nhds 0)
theorem extracted_3 (z : UpperHalfPlane) (b : ) :
CauchySeq fun (N : ) => nFinset.Ico (-N) N, (1 / (b * z + n) - 1 / (b * z + n + 1))
theorem extracted_4 (z : UpperHalfPlane) (b : ) :
CauchySeq fun (N : ) => nFinset.Ico (-N) N, 1 / (b * z + n) ^ 2
theorem extracted_5 (z : UpperHalfPlane) (b : ) :
CauchySeq fun (N : ) => nFinset.Ico (-N) N, 1 / (b * z - n) ^ 2
theorem CauchySeq.congr (f g : ) (hf : f = g) (hh : CauchySeq g) :
theorem cauchy_seq_mul_const (f : ) (c : ) (hc : c 0) :
CauchySeq fCauchySeq (c f)
theorem auxer (a c : ) :
a + 2 * 2 * c - 2 * c = a + 2 * c
noncomputable def summable_term (z : UpperHalfPlane) :
Equations
Instances For
    theorem t8 (z : UpperHalfPlane) :
    (fun (N : ) => mFinset.Icc (-N) N, ∑' (n : ), 1 / (m * z + n) ^ 2) = (fun (x : ) => 2 * riemannZeta 2) + fun (N : ) => mFinset.range N, 2 * (-2 * Real.pi * Complex.I) ^ 2 / (2 - 1).factorial * ∑' (n : ℕ+), n ^ (2 - 1) * Complex.exp (2 * Real.pi * Complex.I * (m + 1) * z * n)
    theorem G2_c_tendsto (z : UpperHalfPlane) :
    Filter.Tendsto (fun (N : ) => xFinset.range N, 2 * (2 * Real.pi * Complex.I) ^ 2 * ∑' (n : ℕ+), n * Complex.exp (2 * Real.pi * Complex.I * (x + 1) * z * n)) Filter.atTop (nhds (-8 * Real.pi ^ 2 * ∑' (n : ℕ+), ((ArithmeticFunction.sigma 1) n) * Complex.exp (2 * Real.pi * Complex.I * n * z)))
    theorem G2_cauchy (z : UpperHalfPlane) :
    CauchySeq fun (N : ) => mFinset.Icc (-N) N, ∑' (n : ), 1 / (m * z + n) ^ 2