From 1114eeaea1ff5e2c3c64d488248aefd9998cfecd Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 6 Feb 2024 17:48:30 +0000 Subject: [PATCH 001/124] wip save --- .../Complex/UpperHalfPlane/Basic.lean | 2 + .../Complex/UpperHalfPlane/Metric.lean | 47 +++ .../Finset_Decomposition.lean | 90 +++++ .../EisensteinSeries/UniformConvergence.lean | 355 ++++++++++++++++++ 4 files changed, 494 insertions(+) create mode 100644 Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean create mode 100644 Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 05a0516dce3d5..5427b2cc31b28 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -123,6 +123,8 @@ theorem mk_coe (z : ℍ) (h : 0 < (z : ℂ).im := z.2) : mk z h = z := rfl #align upper_half_plane.mk_coe UpperHalfPlane.mk_coe +lemma coe_eq_fst (z : ℍ) : (z : ℂ) = z.1 := rfl + theorem re_add_im (z : ℍ) : (z.re + z.im * Complex.I : ℂ) = z := Complex.re_add_im z #align upper_half_plane.re_add_im UpperHalfPlane.re_add_im diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index f3b40a0249970..fdb49ac066cbf 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -384,4 +384,51 @@ instance : IsometricSMul SL(2, ℝ) ℍ := exact (isometry_real_vadd w).comp (h₀.comp <| (isometry_real_vadd v).comp <| isometry_pos_mul u)⟩ +section slices + +/--The verticel strip of width A and height B-/ +def upperHalfPlaneSlice (A B : ℝ) := + {z : ℍ | Complex.abs z.1.1 ≤ A ∧ Complex.abs z.1.2 ≥ B} + +theorem slice_mem (A B : ℝ) (z : ℍ) : + z ∈ upperHalfPlaneSlice A B ↔ Complex.abs z.1.1 ≤ A ∧ Complex.abs z.1.2 ≥ B := Iff.rfl + +lemma compact_in_some_slice (K : Set ℍ) (hK : IsCompact K) : ∃ A B : ℝ, 0 < B ∧ + K ⊆ upperHalfPlaneSlice A B := by + by_cases hne : Set.Nonempty K + · have hcts : ContinuousOn (fun t => t.im) K := by + apply Continuous.continuousOn UpperHalfPlane.continuous_im + obtain ⟨b, _, HB⟩ := IsCompact.exists_isMinOn hK hne hcts + let t := (⟨Complex.I, by simp⟩ : ℍ) + have ht : UpperHalfPlane.im t = I.im := by rfl + obtain ⟨r, _, hr2⟩ := Bornology.IsBounded.subset_closedBall_lt hK.isBounded 0 t + refine' ⟨Real.sinh (r) + Complex.abs ((UpperHalfPlane.center t r)), b.im, b.2, _⟩ + intro z hz + simp only [I_im, slice_mem, abs_ofReal, ge_iff_le] at * + constructor + have hr3 := hr2 hz + simp only [Metric.mem_closedBall] at hr3 + apply le_trans (abs_re_le_abs z) + have := Complex.abs.sub_le (z : ℂ) (UpperHalfPlane.center t r) 0 + simp only [sub_zero, ge_iff_le] at this + rw [dist_le_iff_dist_coe_center_le] at hr3 + apply le_trans this + have htim : UpperHalfPlane.im t = 1 := by + simp only [ht] + rw [htim] at hr3 + simp only [one_mul, add_le_add_iff_right, ge_iff_le] at * + exact hr3 + have hbz := HB hz + simp only [mem_setOf_eq, ge_iff_le] at * + convert hbz + rw [UpperHalfPlane.im] + apply abs_eq_self.mpr z.2.le + · rw [not_nonempty_iff_eq_empty] at hne + rw [hne] + simp only [empty_subset, and_true, exists_const] + use 1 + linarith + +end slices + end UpperHalfPlane diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean new file mode 100644 index 0000000000000..2809b6e4da6af --- /dev/null +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -0,0 +1,90 @@ +import Mathlib.Data.Complex.Abs +import Mathlib.Data.IsROrC.Basic + +open Complex + +open scoped BigOperators NNReal Classical Filter Matrix + +noncomputable section + +namespace EisensteinSeries + +open Finset + +/-- For `m : ℤ` this is the finset of `ℤ × ℤ` of elements such that the maximum of the +absolute values of the pair is `m` -/ +def square (m : ℤ) : Finset (ℤ × ℤ) := + ((Icc (-m) (m)) ×ˢ (Icc (-m) (m))).filter fun x => max x.1.natAbs x.2.natAbs = m + +-- a re-definition in simp-normal form +lemma square_eq (m : ℤ) : + square m = ((Icc (-m) m) ×ˢ (Icc (-m) m)).filter fun x => max |x.1| |x.2| = m := by + simp [square] + +lemma mem_square_aux {m : ℤ} {i} : i ∈ Icc (-m) m ×ˢ Icc (-m) m ↔ max |i.1| |i.2| ≤ m := by + simp [abs_le] + +lemma square_disj {n : ℤ} : Disjoint (square (n+1)) (Icc (-n) n ×ˢ Icc (-n) n) := by + rw [square_eq] + simp only [Finset.disjoint_left, mem_filter, mem_square_aux] + rintro x ⟨-, hx⟩ + simp [hx] + +lemma square_union {n : ℤ} : + square (n+1) ∪ Icc (-n) n ×ˢ Icc (-n) n = Icc (-(n+1)) (n+1) ×ˢ Icc (-(n+1)) (n+1) := by + ext x + rw [mem_union, square_eq, mem_filter, mem_square_aux, mem_square_aux, + and_iff_right_of_imp le_of_eq, Int.le_add_one_iff, or_comm] + +lemma square_disjunion (n : ℤ) : + (square (n+1)).disjUnion (Icc (-n) n ×ˢ Icc (-n) n) square_disj = + Icc (-(n+1)) (n+1) ×ˢ Icc (-(n+1)) (n+1) := by rw [disjUnion_eq_union, square_union] + +theorem square_size (n : ℕ) : Finset.card (square (n + 1)) = 8 * (n + 1) := by + have : (((square (n+1)).disjUnion (Icc (-n : ℤ) n ×ˢ Icc (-n : ℤ) n) square_disj).card : ℤ) = + (Icc (-(n+1 : ℤ)) (n+1) ×ˢ Icc (-(n+1 : ℤ)) (n+1)).card + · rw [square_disjunion] + rw [card_disjUnion, card_product, Nat.cast_add, Nat.cast_mul, card_product, Nat.cast_mul, + Int.card_Icc, Int.card_Icc, Int.toNat_sub_of_le, Int.toNat_sub_of_le, + ←eq_sub_iff_add_eq] at this + · rw [←Nat.cast_inj (R := ℤ), this, Nat.cast_mul, Nat.cast_ofNat, Nat.cast_add_one] + ring_nf + · linarith + · linarith + +theorem natAbs_le_iff_mem_Icc (a : ℤ) (n : ℕ) : + Int.natAbs a ≤ n ↔ a ∈ Finset.Icc (-(n : ℤ)) n := by + rw [mem_Icc, ← abs_le, Int.abs_eq_natAbs, Int.ofNat_le] + +@[simp] +theorem square_mem (n : ℕ) (x : ℤ × ℤ) : x ∈ square n ↔ max x.1.natAbs x.2.natAbs = n := by + simp_rw [square, Finset.mem_filter, Nat.cast_inj, mem_product, and_iff_right_iff_imp, + ← natAbs_le_iff_mem_Icc] + rintro rfl + simp only [le_max_iff, le_refl, true_or, or_true, and_self] + +theorem square_size' {n : ℕ} (h : n ≠ 0) : Finset.card (square n) = 8 * n := by + obtain ⟨n, rfl⟩ := n.exists_eq_succ_of_ne_zero h + exact mod_cast square_size n + +theorem squares_cover_all (y : ℤ × ℤ) : ∃! i : ℕ, y ∈ square i := by + use max y.1.natAbs y.2.natAbs + simp only [square_mem, and_self_iff, forall_eq'] + +@[simp] +lemma square_zero : square 0 = {(0, 0)} := rfl + +theorem square_zero_card : Finset.card (square 0) = 1 := by + rw [square_zero, card_singleton] + +lemma Complex_abs_eq_of_mem_square (n : ℕ) (x : ℤ × ℤ) (h : x ∈ square n) : + Complex.abs x.1 = n ∨ Complex.abs x.2 = n := by + simp_rw [eq_comm (b := (n : ℝ)), ← int_cast_abs, square_mem] at * + norm_cast + subst n + simp only [Nat.cast_max, Int.coe_natAbs, max_eq_left_iff, max_eq_right_iff] at * + exact Int.le_total |x.2| |x.1| + +lemma Complex_abs_square_left_ne (n : ℕ) (x : ℤ × ℤ) (h : x ∈ square n) + (hx : Complex.abs (x.1) ≠ n) : Complex.abs (x.2) = n := + Complex_abs_eq_of_mem_square n x h |>.resolve_left hx diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean new file mode 100644 index 0000000000000..d96b04c73fb5a --- /dev/null +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -0,0 +1,355 @@ +import Mathlib.Analysis.Complex.UpperHalfPlane.Basic +import Mathlib.NumberTheory.Modular +import Mathlib.Data.Int.Interval +import Mathlib.Analysis.SpecialFunctions.Pow.Complex +import Mathlib.Analysis.SpecialFunctions.Pow.Real +import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic +import Mathlib.Analysis.Complex.UpperHalfPlane.Manifold +import Mathlib.Analysis.Complex.UpperHalfPlane.Metric +import Mathlib.Analysis.NormedSpace.FunctionSeries +import Mathlib.Analysis.PSeries +import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Finset_Decomposition + +noncomputable section + +open Complex Filter UpperHalfPlane Set + +open scoped BigOperators NNReal Classical Filter Matrix UpperHalfPlane Complex + +namespace EisensteinSeries + +/-- Auxilary function used for bounding Eisentein series-/ +def lowerBound1 (z : ℍ) : ℝ := + ((z.1.2 ^ (2 : ℕ)) / (z.1.1 ^ (2 : ℕ) + z.1.2 ^ (2 : ℕ))) + +theorem lowerBound1_pos (z : ℍ) : 0 < lowerBound1 z := by + have H2 : 0 < (z.1.1 ^ (2 : ℕ) + z.1.2 ^ (2 : ℕ)) := by + apply_rules [pow_pos, add_pos_of_nonneg_of_pos, pow_two_nonneg, z.2] + exact div_pos (pow_pos z.im_pos 2) H2 + +/-- This function is used to give an upper bound on Eisenstein series-/ +def r (z : ℍ) : ℝ := min (z.1.2) (Real.sqrt (lowerBound1 z)) + +theorem r_pos (z : ℍ) : 0 < r z := by + simp only [r, lt_min_iff, z.property, Real.sqrt_pos, lowerBound1_pos, and_self] + +theorem r_ne_zero (z : ℍ) : r z ≠ 0 := ne_of_gt (r_pos z) + +lemma r_mul_n_pos (k : ℕ) (z : ℍ) (n : ℕ) (hn : 1 ≤ n) : + 0 < (Complex.abs ((r z : ℂ) ^ (k : ℤ) * (n : ℂ)^ (k : ℤ))) := by + norm_cast + apply _root_.abs_pos.mpr (mul_ne_zero (pow_ne_zero k (ne_of_gt (r_pos z))) ?_) + simp only [Nat.cast_pow, ne_eq, pow_eq_zero_iff', Nat.cast_eq_zero, not_and, not_not] at * + intro _ + linarith + +lemma pow_two_of_nonzero_ge_one (a : ℤ) (ha : a ≠ 0) : 0 ≤ a^2 - 1 := by + simp only [sub_nonneg, one_le_sq_iff_one_le_abs, ge_iff_le] + exact Int.one_le_abs ha + +theorem aux4 (a b : ℝ) (h : 0 < b) : (b ^ 2) / (a ^ 2 + b ^ 2) = 1 / ((a / b) ^ 2 + 1) := by + field_simp + +theorem auxlb (z : ℍ) (δ : ℝ) (n : ℤ) (hn : n ≠ 0) : + (z.1.2 ^ 2 ) / (z.1.1 ^ 2 + z.1.2 ^ 2) ≤ (δ * z.1.1 + n) ^ 2 + (δ * z.1.2) ^ 2 := by + have H1 : (δ * z.1.1 + n) ^ 2 + (δ * z.1.2) ^ 2 = + δ ^ 2 * (z.1.1 ^ 2 + z.1.2 ^ 2) + n * 2 * δ * z.1.1 + n^2 := by ring + have H4 : (δ ^ 2 * (z.1.1 ^ 2 + z.1.2 ^ 2) + n * 2 * δ * z.1.1 + n^2) * (z.1.1 ^ 2 + z.1.2 ^ 2) + - (z.1.2 ^ 2) = (δ * (z.1.1 ^ 2 + z.1.2 ^ 2)+ n * z.1.1)^2 + (n^2 - 1)* (z.1.2)^2 := by ring + rw [H1, div_le_iff, ← sub_nonneg, H4] + · apply add_nonneg (pow_two_nonneg _) ?_ + apply mul_nonneg + norm_cast + apply pow_two_of_nonzero_ge_one n hn + apply pow_two_nonneg + · apply_rules [add_pos_of_nonneg_of_pos, pow_two_nonneg, (pow_pos z.im_pos 2)] + +theorem auxbound1 (z : ℍ) (δ : ℝ) : r z ≤ Complex.abs ((z : ℂ) + δ) := by + rw [r, Complex.abs] + have H1 : (z : ℂ).im ≤ + Real.sqrt (((z : ℂ).re + δ) * ((z : ℂ).re + δ) + (z : ℂ).im * (z : ℂ).im) := by + rw [Real.le_sqrt' ] + nlinarith + exact z.2 + simp only [UpperHalfPlane.coe_im, UpperHalfPlane.coe_re, AbsoluteValue.coe_mk, MulHom.coe_mk, + min_le_iff] at * + left + simp only [normSq_apply, add_re, coe_re, ofReal_re, add_im, coe_im, ofReal_im, add_zero] + exact H1 + +theorem auxbound2 (z : ℍ) (δ : ℝ) (n : ℤ) (h : n ≠ 0) : r z ≤ Complex.abs (δ * (z : ℂ) + n) := by + rw [r, Complex.abs, min_le_iff] + have H1 : + Real.sqrt (lowerBound1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + n) * (δ * (z : ℂ).re + n) + + δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by + rw [lowerBound1, Real.sqrt_le_sqrt_iff, ← pow_two, ← pow_two] + apply auxlb z δ n h + nlinarith + right + simp only [ne_eq, coe_re, coe_im, normSq_apply, AbsoluteValue.coe_mk, MulHom.coe_mk, add_re, + mul_re, ofReal_re, ofReal_im, zero_mul, sub_zero, int_cast_re, add_im, mul_im, add_zero, + int_cast_im] at * + exact H1 + +theorem bound1 (z : ℍ) (x : Fin 2 → ℤ) (k : ℕ) : + (r z ) ^ k ≤ Complex.abs (((z : ℂ) + (x 1 : ℂ) / (x 0 : ℂ)) ^ k) := by + have := auxbound1 z (x 1 / x 0 : ℝ) + simp only [zpow_coe_nat, map_pow, ge_iff_le] + apply pow_le_pow_left (r_pos _).le + simp only [ofReal_div, ofReal_int_cast] at * + exact this + +theorem bound2 (z : ℍ) (x : Fin 2 → ℤ) (k : ℕ) : + (r z ) ^ k ≤ Complex.abs (((x 0 : ℂ) / (x 1 : ℂ) * (z : ℂ) + 1) ^ k) := by + have := auxbound2 z (x 0 / x 1 : ℝ) 1 one_ne_zero + simp only [zpow_coe_nat, map_pow, ge_iff_le] + apply pow_le_pow_left (r_pos _).le + simp only [ofReal_div, ofReal_int_cast, Int.cast_one] at * + exact this + +theorem r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice A B) : + r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by + have hz := z.2 + simp only [slice_mem, abs_ofReal, ge_iff_le] at hz + rw [r, r] + apply min_le_min + · dsimp only + convert hz.2 + have := abs_eq_self.mpr (UpperHalfPlane.im_pos z.1).le + convert this.symm + rw [Real.sqrt_le_sqrt_iff] + have := aux4 (z : ℂ).re (z : ℂ).im (UpperHalfPlane.im_pos z.1) + simp only [coe_eq_fst, div_pow, one_div] at this + simp_rw [lowerBound1] + rw [this, aux4 A B h, one_div, inv_le_inv, add_le_add_iff_right, div_pow] + apply div_le_div (sq_nonneg _) + · simpa [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 + · positivity + · simpa [even_two.pow_abs] using pow_le_pow_left h.le hz.2 2 + · positivity + · positivity + · apply (lowerBound1_pos z).le + + +variable {α : Type*} {β : Type*} {i : α → Set β} + +/--Equivalence between the sigma of a fammily of finsets of `ℤ × ℤ` and `ℤ × ℤ`-/ +def sigmaEquiv (ι : α → Finset (β × β)) (HI : ∀ y : β × β , ∃! i : α, y ∈ ι i) : + (Σ s : α, ((ι s) : Set (β × β))) ≃ (β × β) where + toFun x := x.2 + invFun x := by + refine ⟨(HI x).choose, x, (HI x).choose_spec.1⟩ + left_inv x := by + ext + exact ((HI x.2).choose_spec.2 x.1 x.2.2).symm + repeat {rfl} + right_inv x := by rfl + +theorem summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), 0 ≤ f y) + (ι : ℕ → Finset (ℤ × ℤ)) (HI : ∀ y : ℤ × ℤ, ∃! i : ℕ, y ∈ ι i) : + Summable f ↔ Summable fun n : ℕ => ∑ x in ι n, f ![x.1, x.2] := by + let h2 := Equiv.trans (sigmaEquiv ι HI) (piFinTwoEquiv fun _ => ℤ).symm + have h22 : ∀ y : Σ s : ℕ, (ι s), 0 ≤ (f ∘ h2) y := by + intro y + apply h + have h4 : Summable f ↔ Summable (f ∘ h2) := by rw [Equiv.summable_iff] + rw [h4, summable_sigma_of_nonneg h22] + constructor + · intro H + convert H.2 + rw [←Finset.tsum_subtype] + rfl + · intro H + constructor + · intro x + simp only [Finset.coe_sort_coe, Equiv.coe_trans, Function.comp_apply,sigmaEquiv] + convert (Finset.summable (ι x) (f ∘ (piFinTwoEquiv fun _ => ℤ).symm)) + · convert H + rw [←Finset.tsum_subtype] + rfl + +lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : + Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by + have hk : 1 < (k - 1 : ℝ) := by + have : 1 < (k - 1 : ℤ) := by linarith + norm_cast at * + have riesum := Real.summable_nat_rpow_inv.2 hk + have nze : (8 / (r z) ^ k : ℝ) ≠ 0 := + by + apply div_ne_zero + simp only [ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true] + apply zpow_ne_zero k (ne_of_gt (r_pos z)) + rw [← (summable_mul_left_iff nze).symm] + convert riesum + norm_cast + +lemma summable_over_square (k : ℤ) (z : ℍ) (h : 3 ≤ k): + Summable (fun n : ℕ => ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by + simp only [one_div, Finset.sum_const, nsmul_eq_mul] + apply Summable.congr (summable_r_pow k z h) + intro b + by_cases b0 : b = 0 + · rw [b0] + have hk0 : k ≠ 0 := by linarith + have hk1 : k - 1 ≠ 0 := by linarith + norm_cast + rw [zero_zpow k hk0, zero_zpow (k - 1) hk1] + simp only [inv_zero, mul_zero, square_zero, Finset.card_singleton, Nat.cast_one] + · rw [square_size' b0, zpow_sub_one₀ (a:= ( b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] + simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] + ring_nf + +lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => + (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by + rw [summable_lemma _ _ (fun (n : ℕ) => square n) squares_cover_all] + have : ∀ n : ℕ, ∑ v in square n, (1 / (r z) ^ k) * ((max v.1.natAbs v.2.natAbs: ℝ) ^ k)⁻¹ = + ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ)^k)⁻¹ := by + intro n + apply Finset.sum_congr rfl + intro x hx + simp only [square_mem] at hx + congr + norm_cast + apply Summable.congr (summable_over_square k z h) + intro b + apply (this b).symm + intro y + apply mul_nonneg + simp only [one_div, inv_nonneg] + apply zpow_nonneg (r_pos z).le + simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] + + +lemma Eise_bound_1 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ n) + (C1 : Complex.abs (x 0 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ + (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by + rw [inv_le_inv] + have h0 : (x 0 : ℂ) ≠ 0 := by + intro hx + rw [hx] at C1 + simp only [map_zero, Int.cast_eq_zero] at * + norm_cast at * + simp_rw [← C1] at hn + simp only [Nat.one_ne_zero, le_zero_iff] at hn + have h1 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = + ((x 0 : ℂ) ^ (k : ℤ)) * ((z : ℂ) + (x 1 : ℂ) / x 0) ^ (k : ℤ) := + by + field_simp + ring + rw [h1] + simp_rw [map_mul Complex.abs] + norm_cast at * + have h3 : Complex.abs (n^k : ℕ) = Complex.abs ((x 0)^k : ℤ) := by + have : Complex.abs ((x 0)^k : ℤ) = Complex.abs (x 0)^k := by + simp only [Int.cast_pow, map_pow, Real.rpow_nat_cast] + rw [this, C1] + norm_cast + simp only [Nat.cast_pow, map_pow, abs_natCast] + rw [h3, mul_comm] + apply mul_le_mul_of_nonneg_left _ (Complex.abs.nonneg _) + convert bound1 z x k + apply abs_eq_self.mpr ( pow_nonneg (r_pos z).le k) + norm_cast + have := Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) + apply this + simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] + intro hx + rw [hx] at C1 + simp [Int.cast_zero] at C1 + norm_cast at C1 + rw [← C1] at hn + simp only [Nat.one_ne_zero, le_zero_iff] at hn + apply r_mul_n_pos k z n hn + +lemma Eis_bound_2 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ n) + (C2 : Complex.abs (x 1 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ + (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by + have h0 : (x 1 : ℂ) ≠ 0 := by + intro hx + rw [hx] at C2 + simp at C2 + norm_cast at * + rw [← C2] at hn + simp only [Nat.one_ne_zero, le_zero_iff] at hn + rw [inv_le_inv] + have h1 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = + (x 1 : ℂ)^ (k) * ((x 0 : ℂ) / (x 1 : ℂ) * (z : ℂ) + 1) ^ (k) := by + field_simp + ring_nf + rw [h1] + simp_rw [map_mul Complex.abs] + have h3 : Complex.abs (n^k ) = Complex.abs ((x 1)^k ) := by + simp only [map_pow, abs_natCast, ge_iff_le, Nat.cast_nonneg, map_nonneg, C2] + simp at * + norm_cast at * + rw [h3, mul_comm] + apply mul_le_mul_of_nonneg_left + convert (bound2 z x k) + apply abs_eq_self.mpr ( (r_pos z).le ) + norm_cast + simp [ge_iff_le, map_nonneg, zpow_nonneg] + apply pow_nonneg (Complex.abs.nonneg _) + apply Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) + simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] + intro hx + norm_cast at * + apply r_mul_n_pos k z n hn + +theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) + (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ + (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by + by_cases hn : n = 0 + · rw [hn] at hx + simp only [CharP.cast_eq_zero, square_zero, Finset.mem_singleton, Prod.mk.injEq] at hx + by_cases hk : k = 0 + rw [hk] at * + simp only [coe_eq_fst, pow_zero, map_one, inv_one, mul_one, le_refl] + rw [hx.1, hx.2] + have h1 : (0 : ℝ) ^ (k : ℕ) = 0 := by + simp only [pow_eq_zero_iff', ne_eq, true_and] + exact hk + simp only [Int.cast_zero, coe_eq_fst, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, + CharP.cast_eq_zero, map_mul, abs_ofReal, mul_zero, le_refl] + · have hnn : 1 ≤ n := by exact Nat.one_le_iff_ne_zero.mpr hn + by_cases C1 : Complex.abs (x 0 : ℂ) = n + apply Eise_bound_1 k z n x hnn C1 + have C2 := Complex_abs_square_left_ne n ⟨x 0, x 1⟩ hx C1 + apply Eis_bound_2 k z n x hnn C2 + +lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) + (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => + (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) + ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by + have hk0 : 0 ≤ k := by linarith + lift k to ℕ using hk0 + rw [tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] + simp only [top_eq_univ, subset_univ, eisensteinSeries, forall_true_left] + intros K hK + obtain ⟨A,B,hB, HABK⟩:= compact_in_some_slice K hK + let u := fun x : (gammaSet N a ) => + (1/((r ⟨⟨A, B⟩, hB⟩)^k))* ( (max (x.1 0).natAbs (x.1 1).natAbs : ℝ)^k)⁻¹ + have hu : Summable u := by + have := Summable.subtype (summable_upper_bound k hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a ) + apply this.congr + intro v + simp only [zpow_coe_nat, one_div, Function.comp_apply] + apply tendstoUniformlyOn_tsum hu + intro v x hx + have sq := square_mem (max (v.1 0).natAbs (v.1 1).natAbs ) ⟨(v.1 0), v.1 1⟩ + have := Eis_is_bounded_on_square k x (max (v.1 0).natAbs (v.1 1).natAbs ) v + simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_coe_nat, one_div, coe_eq_fst, map_pow, + map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, + ge_iff_le] at * + apply le_trans (this sq) + rw [mul_comm] + apply mul_le_mul + rw [inv_le_inv] + apply pow_le_pow_left (r_pos _).le + rw [abs_of_pos (r_pos _)] + · exact r_lower_bound_on_slice A B hB ⟨x, HABK hx⟩ + · apply pow_pos (abs_pos.mpr (r_ne_zero x)) + · apply pow_pos (r_pos _) + · rfl + repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg, + inv_nonneg, pow_nonneg (r_pos _).le]} + · simp only [top_eq_univ, isOpen_univ] From e9440a2acc125be3d83ecf192d5537143c5d85a8 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:00:37 +0000 Subject: [PATCH 002/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/Finset_Decomposition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean index 2809b6e4da6af..3c70628de4f18 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -37,7 +37,7 @@ lemma square_union {n : ℤ} : and_iff_right_of_imp le_of_eq, Int.le_add_one_iff, or_comm] lemma square_disjunion (n : ℤ) : - (square (n+1)).disjUnion (Icc (-n) n ×ˢ Icc (-n) n) square_disj = + (square (n+1)).disjUnion (Icc (-n) n ×ˢ Icc (-n) n) square_disj = Icc (-(n+1)) (n+1) ×ˢ Icc (-(n+1)) (n+1) := by rw [disjUnion_eq_union, square_union] theorem square_size (n : ℕ) : Finset.card (square (n + 1)) = 8 * (n + 1) := by From e229b8436e6377c7f564a497beae4eeab5f6f303 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:00:48 +0000 Subject: [PATCH 003/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/Finset_Decomposition.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean index 3c70628de4f18..5eff092ba7ba0 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -46,8 +46,8 @@ theorem square_size (n : ℕ) : Finset.card (square (n + 1)) = 8 * (n + 1) := by · rw [square_disjunion] rw [card_disjUnion, card_product, Nat.cast_add, Nat.cast_mul, card_product, Nat.cast_mul, Int.card_Icc, Int.card_Icc, Int.toNat_sub_of_le, Int.toNat_sub_of_le, - ←eq_sub_iff_add_eq] at this - · rw [←Nat.cast_inj (R := ℤ), this, Nat.cast_mul, Nat.cast_ofNat, Nat.cast_add_one] + ← eq_sub_iff_add_eq] at this + · rw [← Nat.cast_inj (R := ℤ), this, Nat.cast_mul, Nat.cast_ofNat, Nat.cast_add_one] ring_nf · linarith · linarith From 3a0fc108056ae2e7b72e906adee6b3b0b494f563 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:00:55 +0000 Subject: [PATCH 004/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index d96b04c73fb5a..8cd8101849575 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -36,7 +36,7 @@ theorem r_pos (z : ℍ) : 0 < r z := by theorem r_ne_zero (z : ℍ) : r z ≠ 0 := ne_of_gt (r_pos z) lemma r_mul_n_pos (k : ℕ) (z : ℍ) (n : ℕ) (hn : 1 ≤ n) : - 0 < (Complex.abs ((r z : ℂ) ^ (k : ℤ) * (n : ℂ)^ (k : ℤ))) := by + 0 < (Complex.abs ((r z : ℂ) ^ (k : ℤ) * (n : ℂ)^ (k : ℤ))) := by norm_cast apply _root_.abs_pos.mpr (mul_ne_zero (pow_ne_zero k (ne_of_gt (r_pos z))) ?_) simp only [Nat.cast_pow, ne_eq, pow_eq_zero_iff', Nat.cast_eq_zero, not_and, not_not] at * From b74fe5c852a3ddfb63abd6e4614f68abeded314e Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:01:04 +0000 Subject: [PATCH 005/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 8cd8101849575..031555a1bdc96 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -146,7 +146,7 @@ def sigmaEquiv (ι : α → Finset (β × β)) (HI : ∀ y : β × β , ∃! i : right_inv x := by rfl theorem summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), 0 ≤ f y) - (ι : ℕ → Finset (ℤ × ℤ)) (HI : ∀ y : ℤ × ℤ, ∃! i : ℕ, y ∈ ι i) : + (ι : ℕ → Finset (ℤ × ℤ)) (HI : ∀ y : ℤ × ℤ, ∃! i : ℕ, y ∈ ι i) : Summable f ↔ Summable fun n : ℕ => ∑ x in ι n, f ![x.1, x.2] := by let h2 := Equiv.trans (sigmaEquiv ι HI) (piFinTwoEquiv fun _ => ℤ).symm have h22 : ∀ y : Σ s : ℕ, (ι s), 0 ≤ (f ∘ h2) y := by From bcf756c276b3021b3ada097aad26833f1a966728 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:01:12 +0000 Subject: [PATCH 006/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 031555a1bdc96..7483ce5530a53 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -157,7 +157,7 @@ theorem summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ constructor · intro H convert H.2 - rw [←Finset.tsum_subtype] + rw [← Finset.tsum_subtype] rfl · intro H constructor From b241e1564b14f07e76d7ddeb9d9307c2ad2ac8c6 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:01:21 +0000 Subject: [PATCH 007/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 7483ce5530a53..1f361bcb3b55b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -165,7 +165,7 @@ theorem summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ simp only [Finset.coe_sort_coe, Equiv.coe_trans, Function.comp_apply,sigmaEquiv] convert (Finset.summable (ι x) (f ∘ (piFinTwoEquiv fun _ => ℤ).symm)) · convert H - rw [←Finset.tsum_subtype] + rw [← Finset.tsum_subtype] rfl lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : From a456d5bbda08f8d7bcd0cfde7787d45db5e07880 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:01:29 +0000 Subject: [PATCH 008/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 1f361bcb3b55b..a402d5c641fcd 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -169,7 +169,7 @@ theorem summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ rfl lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : - Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by + Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by have hk : 1 < (k - 1 : ℝ) := by have : 1 < (k - 1 : ℤ) := by linarith norm_cast at * From eed710a9c9eca9fe7a2dbba3454f26d8527a1d00 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:01:38 +0000 Subject: [PATCH 009/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index a402d5c641fcd..563a04545a580 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -221,7 +221,7 @@ lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : lemma Eise_bound_1 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ n) - (C1 : Complex.abs (x 0 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ + (C1 : Complex.abs (x 0 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by rw [inv_le_inv] have h0 : (x 0 : ℂ) ≠ 0 := by From 6b00e1ff5f88206795cee679b2cbcef26630e256 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:01:46 +0000 Subject: [PATCH 010/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 563a04545a580..59b2c5a16d20e 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -262,7 +262,7 @@ lemma Eise_bound_1 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ apply r_mul_n_pos k z n hn lemma Eis_bound_2 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ n) - (C2 : Complex.abs (x 1 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ + (C2 : Complex.abs (x 1 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by have h0 : (x 1 : ℂ) ≠ 0 := by intro hx From 0655fb3e898d101b711a41329de951058d24b107 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:01:55 +0000 Subject: [PATCH 011/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 59b2c5a16d20e..d74c776c73f8f 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -296,7 +296,7 @@ lemma Eis_bound_2 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ apply r_mul_n_pos k z n hn theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) - (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ + (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by by_cases hn : n = 0 · rw [hn] at hx From b2a5942f6f9f85ea9ac6afd9cfb76bcd73a05caa Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 6 Feb 2024 19:02:06 +0000 Subject: [PATCH 012/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index d74c776c73f8f..ea7e25fbdc22d 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -317,7 +317,7 @@ theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → apply Eis_bound_2 k z n x hnn C2 lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => + (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by have hk0 : 0 ≤ k := by linarith From 80e3230aa289246cf7eac31ad59b53707d7224c6 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 6 Feb 2024 19:14:38 +0000 Subject: [PATCH 013/124] some linting --- .../Finset_Decomposition.lean | 5 +++ .../EisensteinSeries/UniformConvergence.lean | 44 ++++++++++++------- 2 files changed, 32 insertions(+), 17 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean index 5eff092ba7ba0..f13a4dd66b5e6 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -1,3 +1,8 @@ +/- +Copyright (c) 2024 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck, David Loeffler +-/ import Mathlib.Data.Complex.Abs import Mathlib.Data.IsROrC.Basic diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index ea7e25fbdc22d..8a2e28e062397 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -1,3 +1,8 @@ +/- +Copyright (c) 2024 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck +-/ import Mathlib.Analysis.Complex.UpperHalfPlane.Basic import Mathlib.NumberTheory.Modular import Mathlib.Data.Int.Interval @@ -10,6 +15,12 @@ import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.PSeries import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Finset_Decomposition +/-! +# Uniform convergence of Eisenstein series + +We show that `eis` converges locally uniformly on `ℍ` to the Eisenstein series `E` of weight `k` +-/ + noncomputable section open Complex Filter UpperHalfPlane Set @@ -174,8 +185,7 @@ lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : have : 1 < (k - 1 : ℤ) := by linarith norm_cast at * have riesum := Real.summable_nat_rpow_inv.2 hk - have nze : (8 / (r z) ^ k : ℝ) ≠ 0 := - by + have nze : (8 / (r z) ^ k : ℝ) ≠ 0 := by apply div_ne_zero simp only [ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true] apply zpow_ne_zero k (ne_of_gt (r_pos z)) @@ -185,19 +195,19 @@ lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : lemma summable_over_square (k : ℤ) (z : ℍ) (h : 3 ≤ k): Summable (fun n : ℕ => ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by - simp only [one_div, Finset.sum_const, nsmul_eq_mul] - apply Summable.congr (summable_r_pow k z h) - intro b - by_cases b0 : b = 0 - · rw [b0] - have hk0 : k ≠ 0 := by linarith - have hk1 : k - 1 ≠ 0 := by linarith - norm_cast - rw [zero_zpow k hk0, zero_zpow (k - 1) hk1] - simp only [inv_zero, mul_zero, square_zero, Finset.card_singleton, Nat.cast_one] - · rw [square_size' b0, zpow_sub_one₀ (a:= ( b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] - simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] - ring_nf + simp only [one_div, Finset.sum_const, nsmul_eq_mul] + apply Summable.congr (summable_r_pow k z h) + intro b + by_cases b0 : b = 0 + · rw [b0] + have hk0 : k ≠ 0 := by linarith + have hk1 : k - 1 ≠ 0 := by linarith + norm_cast + rw [zero_zpow k hk0, zero_zpow (k - 1) hk1] + simp only [inv_zero, mul_zero, square_zero, Finset.card_singleton, Nat.cast_one] + · rw [square_size' b0, zpow_sub_one₀ (a:= ( b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] + simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] + ring_nf lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by @@ -297,7 +307,7 @@ lemma Eis_bound_2 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ - (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by + (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by by_cases hn : n = 0 · rw [hn] at hx simp only [CharP.cast_eq_zero, square_zero, Finset.mem_singleton, Prod.mk.injEq] at hx @@ -317,7 +327,7 @@ theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → apply Eis_bound_2 k z n x hnn C2 lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => + (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by have hk0 : 0 ≤ k := by linarith From 0b42af4e2de8b0fb8a0f1e42c78183cec61af142 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 6 Feb 2024 19:15:49 +0000 Subject: [PATCH 014/124] missed one --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 8a2e28e062397..8301a063cc42f 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -242,8 +242,7 @@ lemma Eise_bound_1 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ simp_rw [← C1] at hn simp only [Nat.one_ne_zero, le_zero_iff] at hn have h1 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = - ((x 0 : ℂ) ^ (k : ℤ)) * ((z : ℂ) + (x 1 : ℂ) / x 0) ^ (k : ℤ) := - by + ((x 0 : ℂ) ^ (k : ℤ)) * ((z : ℂ) + (x 1 : ℂ) / x 0) ^ (k : ℤ) := by field_simp ring rw [h1] @@ -327,7 +326,7 @@ theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → apply Eis_bound_2 k z n x hnn C2 lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => + (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by have hk0 : 0 ≤ k := by linarith From 8b3ca2ec2dd13fc263cca7cf802e9c8f254ab6f3 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 6 Feb 2024 19:21:00 +0000 Subject: [PATCH 015/124] missed one --- .../EisensteinSeries/Finset_Decomposition.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean index f13a4dd66b5e6..d8704165ce6c3 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -1,11 +1,17 @@ /- Copyright (c) 2024 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Birkbeck, David Loeffler +Authors: Chris Birkbeck -/ import Mathlib.Data.Complex.Abs import Mathlib.Data.IsROrC.Basic +/-! # Decomposing `ℤ × ℤ` into squares + +We partition `ℤ × ℤ` into squares of the form `Icc (-n) n × Icc (-n) n` for `n : ℕ`. This is useful +for bounding Eisenstein series. +-/ + open Complex open scoped BigOperators NNReal Classical Filter Matrix From 3b6159b27eb752f10ce0de321f11d36d68db7cf1 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 6 Feb 2024 19:23:13 +0000 Subject: [PATCH 016/124] more lint --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 8301a063cc42f..92a82d7db387d 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -19,6 +19,7 @@ import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Finset_Decomposition # Uniform convergence of Eisenstein series We show that `eis` converges locally uniformly on `ℍ` to the Eisenstein series `E` of weight `k` +and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. -/ noncomputable section @@ -327,8 +328,8 @@ theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => - (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) - ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by + (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) + ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by have hk0 : 0 ≤ k := by linarith lift k to ℕ using hk0 rw [tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] From 0fc98a5ddb7c3f16838cf616b19f38ff686d29dc Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 6 Feb 2024 19:25:08 +0000 Subject: [PATCH 017/124] more lint --- .../EisensteinSeries/UniformConvergence.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 92a82d7db387d..201811efea9e8 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -159,7 +159,7 @@ def sigmaEquiv (ι : α → Finset (β × β)) (HI : ∀ y : β × β , ∃! i : theorem summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), 0 ≤ f y) (ι : ℕ → Finset (ℤ × ℤ)) (HI : ∀ y : ℤ × ℤ, ∃! i : ℕ, y ∈ ι i) : - Summable f ↔ Summable fun n : ℕ => ∑ x in ι n, f ![x.1, x.2] := by + Summable f ↔ Summable fun n : ℕ => ∑ x in ι n, f ![x.1, x.2] := by let h2 := Equiv.trans (sigmaEquiv ι HI) (piFinTwoEquiv fun _ => ℤ).symm have h22 : ∀ y : Σ s : ℕ, (ι s), 0 ≤ (f ∘ h2) y := by intro y @@ -233,7 +233,7 @@ lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : lemma Eise_bound_1 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ n) (C1 : Complex.abs (x 0 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ - (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by + (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by rw [inv_le_inv] have h0 : (x 0 : ℂ) ≠ 0 := by intro hx @@ -273,7 +273,7 @@ lemma Eise_bound_1 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ lemma Eis_bound_2 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ n) (C2 : Complex.abs (x 1 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ - (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by + (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by have h0 : (x 1 : ℂ) ≠ 0 := by intro hx rw [hx] at C2 @@ -307,7 +307,7 @@ lemma Eis_bound_2 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ - (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by + (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by by_cases hn : n = 0 · rw [hn] at hx simp only [CharP.cast_eq_zero, square_zero, Finset.mem_singleton, Prod.mk.injEq] at hx @@ -328,8 +328,8 @@ theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => - (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) - ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by + (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) + ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by have hk0 : 0 ≤ k := by linarith lift k to ℕ using hk0 rw [tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] From 79fe9aca155a90c304e8228d75de4144870b1ed4 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 9 Feb 2024 09:33:29 +0000 Subject: [PATCH 018/124] save --- .../ModularForms/EisensteinSeries/Finset_Decomposition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean index d8704165ce6c3..9e4a7cea57a40 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -22,7 +22,7 @@ namespace EisensteinSeries open Finset -/-- For `m : ℤ` this is the finset of `ℤ × ℤ` of elements such that the maximum of the +/-- For `m : ℤ` this is the finset of `ℤ × ℤ` of elements such that the maximum of the absolute values of the pair is `m` -/ def square (m : ℤ) : Finset (ℤ × ℤ) := ((Icc (-m) (m)) ×ˢ (Icc (-m) (m))).filter fun x => max x.1.natAbs x.2.natAbs = m From a68247d568d4382958f40ccbf7a91b5228e7a752 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 9 Feb 2024 09:43:41 +0000 Subject: [PATCH 019/124] add file names --- Mathlib.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib.lean b/Mathlib.lean index 54c835c4ada19..3d17498409160 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2845,6 +2845,8 @@ import Mathlib.NumberTheory.Modular import Mathlib.NumberTheory.ModularForms.Basic import Mathlib.NumberTheory.ModularForms.CongruenceSubgroups import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic +import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Finset_Decomposition +import Mathlib.NumberTheory.ModularForms.EisensteinSeries.UniformConvergence import Mathlib.NumberTheory.ModularForms.JacobiTheta.Manifold import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable From 99469cb17fa3c63555f6247a81664c011a204dbf Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 9 Feb 2024 10:05:20 +0000 Subject: [PATCH 020/124] shake fix --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 201811efea9e8..04e92acf97963 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -4,12 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ import Mathlib.Analysis.Complex.UpperHalfPlane.Basic -import Mathlib.NumberTheory.Modular import Mathlib.Data.Int.Interval import Mathlib.Analysis.SpecialFunctions.Pow.Complex import Mathlib.Analysis.SpecialFunctions.Pow.Real import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic -import Mathlib.Analysis.Complex.UpperHalfPlane.Manifold import Mathlib.Analysis.Complex.UpperHalfPlane.Metric import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.PSeries From 89c9ee8e5a515497c70fd92adf6da543ca12608c Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 9 Feb 2024 14:48:37 +0000 Subject: [PATCH 021/124] save --- .../EisensteinSeries/UniformConvergence.lean | 233 ++++++++---------- 1 file changed, 102 insertions(+), 131 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 04e92acf97963..93b8b693dee9a 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -32,8 +32,12 @@ namespace EisensteinSeries def lowerBound1 (z : ℍ) : ℝ := ((z.1.2 ^ (2 : ℕ)) / (z.1.1 ^ (2 : ℕ) + z.1.2 ^ (2 : ℕ))) +lemma lowerBound1' (z : ℍ) : lowerBound1 z = 1/((z.1.1 / z.1.2) ^ (2 : ℕ) + 1) := by + have := z.2 + field_simp [lowerBound1] + theorem lowerBound1_pos (z : ℍ) : 0 < lowerBound1 z := by - have H2 : 0 < (z.1.1 ^ (2 : ℕ) + z.1.2 ^ (2 : ℕ)) := by + have H2 : 0 < (z.1.1 ^ (2 : ℕ) + z.1.2 ^ (2 : ℕ)) := by apply_rules [pow_pos, add_pos_of_nonneg_of_pos, pow_two_nonneg, z.2] exact div_pos (pow_pos z.im_pos 2) H2 @@ -43,9 +47,9 @@ def r (z : ℍ) : ℝ := min (z.1.2) (Real.sqrt (lowerBound1 z)) theorem r_pos (z : ℍ) : 0 < r z := by simp only [r, lt_min_iff, z.property, Real.sqrt_pos, lowerBound1_pos, and_self] -theorem r_ne_zero (z : ℍ) : r z ≠ 0 := ne_of_gt (r_pos z) +theorem r_ne_zero (z : ℍ) : r z ≠ 0 := ne_of_gt (r_pos z) -lemma r_mul_n_pos (k : ℕ) (z : ℍ) (n : ℕ) (hn : 1 ≤ n) : +lemma r_mul_n_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : 0 < (Complex.abs ((r z : ℂ) ^ (k : ℤ) * (n : ℂ)^ (k : ℤ))) := by norm_cast apply _root_.abs_pos.mpr (mul_ne_zero (pow_ne_zero k (ne_of_gt (r_pos z))) ?_) @@ -53,28 +57,20 @@ lemma r_mul_n_pos (k : ℕ) (z : ℍ) (n : ℕ) (hn : 1 ≤ n) : intro _ linarith -lemma pow_two_of_nonzero_ge_one (a : ℤ) (ha : a ≠ 0) : 0 ≤ a^2 - 1 := by - simp only [sub_nonneg, one_le_sq_iff_one_le_abs, ge_iff_le] - exact Int.one_le_abs ha - -theorem aux4 (a b : ℝ) (h : 0 < b) : (b ^ 2) / (a ^ 2 + b ^ 2) = 1 / ((a / b) ^ 2 + 1) := by - field_simp - -theorem auxlb (z : ℍ) (δ : ℝ) (n : ℤ) (hn : n ≠ 0) : - (z.1.2 ^ 2 ) / (z.1.1 ^ 2 + z.1.2 ^ 2) ≤ (δ * z.1.1 + n) ^ 2 + (δ * z.1.2) ^ 2 := by - have H1 : (δ * z.1.1 + n) ^ 2 + (δ * z.1.2) ^ 2 = - δ ^ 2 * (z.1.1 ^ 2 + z.1.2 ^ 2) + n * 2 * δ * z.1.1 + n^2 := by ring - have H4 : (δ ^ 2 * (z.1.1 ^ 2 + z.1.2 ^ 2) + n * 2 * δ * z.1.1 + n^2) * (z.1.1 ^ 2 + z.1.2 ^ 2) - - (z.1.2 ^ 2) = (δ * (z.1.1 ^ 2 + z.1.2 ^ 2)+ n * z.1.1)^2 + (n^2 - 1)* (z.1.2)^2 := by ring +theorem auxlb (z : ℍ) (δ ε : ℝ) (hε : 0 ≤ ε^2 - 1) : + (z.1.2 ^ 2 ) / (z.1.1 ^ 2 + z.1.2 ^ 2) ≤ (δ * z.1.1 + ε) ^ 2 + (δ * z.1.2) ^ 2 := by + have H1 : (δ * z.1.1 + ε) ^ 2 + (δ * z.1.2) ^ 2 = + δ ^ 2 * (z.1.1 ^ 2 + z.1.2 ^ 2) + ε * 2 * δ * z.1.1 + ε^2 := by ring + have H4 : (δ ^ 2 * (z.1.1 ^ 2 + z.1.2 ^ 2) + ε * 2 * δ * z.1.1 + ε^2) * (z.1.1 ^ 2 + z.1.2 ^ 2) + - (z.1.2 ^ 2) = (δ * (z.1.1 ^ 2 + z.1.2 ^ 2)+ ε * z.1.1)^2 + (ε^2 - 1)* (z.1.2)^2 := by ring rw [H1, div_le_iff, ← sub_nonneg, H4] · apply add_nonneg (pow_two_nonneg _) ?_ apply mul_nonneg norm_cast - apply pow_two_of_nonzero_ge_one n hn apply pow_two_nonneg · apply_rules [add_pos_of_nonneg_of_pos, pow_two_nonneg, (pow_pos z.im_pos 2)] -theorem auxbound1 (z : ℍ) (δ : ℝ) : r z ≤ Complex.abs ((z : ℂ) + δ) := by +theorem auxbound1 (z : ℍ) (δ : ℝ) : r z ≤ Complex.abs ((z : ℂ) + δ) := by rw [r, Complex.abs] have H1 : (z : ℂ).im ≤ Real.sqrt (((z : ℂ).re + δ) * ((z : ℂ).re + δ) + (z : ℂ).im * (z : ℂ).im) := by @@ -87,13 +83,13 @@ theorem auxbound1 (z : ℍ) (δ : ℝ) : r z ≤ Complex.abs ((z : ℂ) + δ) : simp only [normSq_apply, add_re, coe_re, ofReal_re, add_im, coe_im, ofReal_im, add_zero] exact H1 -theorem auxbound2 (z : ℍ) (δ : ℝ) (n : ℤ) (h : n ≠ 0) : r z ≤ Complex.abs (δ * (z : ℂ) + n) := by +theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 0 ≤ ε^2 - 1) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs, min_le_iff] have H1 : - Real.sqrt (lowerBound1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + n) * (δ * (z : ℂ).re + n) + + Real.sqrt (lowerBound1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε ) + δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by rw [lowerBound1, Real.sqrt_le_sqrt_iff, ← pow_two, ← pow_two] - apply auxlb z δ n h + apply auxlb z δ ε hε nlinarith right simp only [ne_eq, coe_re, coe_im, normSq_apply, AbsoluteValue.coe_mk, MulHom.coe_mk, add_re, @@ -111,12 +107,85 @@ theorem bound1 (z : ℍ) (x : Fin 2 → ℤ) (k : ℕ) : theorem bound2 (z : ℍ) (x : Fin 2 → ℤ) (k : ℕ) : (r z ) ^ k ≤ Complex.abs (((x 0 : ℂ) / (x 1 : ℂ) * (z : ℂ) + 1) ^ k) := by - have := auxbound2 z (x 0 / x 1 : ℝ) 1 one_ne_zero + have := auxbound2 z (x 0 / x 1 : ℝ) 1 (by norm_cast) simp only [zpow_coe_nat, map_pow, ge_iff_le] apply pow_le_pow_left (r_pos _).le simp only [ofReal_div, ofReal_int_cast, Int.cast_one] at * exact this +lemma fds (z : ℂ) (hz : 0 < Complex.abs (z)) : z ≠ 0 := by exact nnnorm_pos.mp hz + +lemma eis_bound_1 (k : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (C1 : 0 < Complex.abs (x 0 : ℂ)) : + (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ + (Complex.abs ((r z) ^ (k : ℤ) * (Complex.abs (x 0 : ℂ)) ^ (k : ℤ)))⁻¹ := by + have h0 : (x 0 : ℂ) ≠ 0 := nnnorm_pos.mp C1 + rw [inv_le_inv] + have h1 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = + ((x 0 : ℂ) ^ (k : ℤ)) * ((z : ℂ) + (x 1 : ℂ) / x 0) ^ (k : ℤ) := by + field_simp + ring + rw [h1] + have gd := mul_le_mul_of_nonneg_left (bound1 z x k) (pow_nonneg C1.le k) + simp only [AbsoluteValue.pos_iff, ne_eq, Int.cast_eq_zero, zpow_coe_nat, map_pow, map_mul, + abs_ofReal, Complex.abs_abs, ge_iff_le] at * + rw [ mul_comm] + convert gd + apply abs_eq_self.mpr ((r_pos z).le ) + have := Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) + apply this + simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] + intro hx + norm_cast at * + apply r_mul_n_pos k z (Complex.abs (x 0 : ℂ)) C1 + +lemma eis_bound_2 (k : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (C2 : 0 < Complex.abs (x 1 : ℂ)) : + (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ + (Complex.abs ((r z) ^ (k : ℤ) * (Complex.abs (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ := by + have h0 : (x 1 : ℂ) ≠ 0 := nnnorm_pos.mp C2 + rw [inv_le_inv] + have h1 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = + (x 1 : ℂ)^ (k) * ((x 0 : ℂ) / (x 1 : ℂ) * (z : ℂ) + 1) ^ (k) := by + field_simp + ring_nf + rw [h1] + have gd := mul_le_mul_of_nonneg_left (bound2 z x k) (pow_nonneg C2.le k) + simp only [AbsoluteValue.pos_iff, ne_eq, Int.cast_eq_zero, zpow_coe_nat, map_pow, map_mul, + abs_ofReal, Complex.abs_abs, ge_iff_le] at * + rw [ mul_comm] + convert gd + apply abs_eq_self.mpr ((r_pos z).le ) + have := Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) + apply this + simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] + intro hx + norm_cast at * + apply r_mul_n_pos k z (Complex.abs (x 1 : ℂ)) C2 + +theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) + (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ + (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by + by_cases hn : n = 0 + · rw [hn] at hx + simp only [CharP.cast_eq_zero, square_zero, Finset.mem_singleton, Prod.mk.injEq] at hx + by_cases hk : k = 0 + rw [hk] at * + simp only [coe_eq_fst, pow_zero, map_one, inv_one, mul_one, le_refl] + rw [hx.1, hx.2] + have h1 : (0 : ℝ) ^ (k : ℕ) = 0 := by + simp only [pow_eq_zero_iff', ne_eq, true_and] + exact hk + simp only [Int.cast_zero, coe_eq_fst, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, + CharP.cast_eq_zero, map_mul, abs_ofReal, mul_zero, le_refl] + · have hnn : 0 < (n : ℝ) := by norm_cast; exact Nat.pos_of_ne_zero hn + by_cases C1 : Complex.abs (x 0 : ℂ) = n + · have := eis_bound_1 k z x (by rw [C1]; exact hnn) + rw [C1] at * + exact this + · have C2 := Complex_abs_square_left_ne n ⟨x 0, x 1⟩ hx C1 + have := eis_bound_2 k z x (by rw [C2]; exact hnn) + rw [C2] at this + exact this + theorem r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by have hz := z.2 @@ -128,10 +197,8 @@ theorem r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice have := abs_eq_self.mpr (UpperHalfPlane.im_pos z.1).le convert this.symm rw [Real.sqrt_le_sqrt_iff] - have := aux4 (z : ℂ).re (z : ℂ).im (UpperHalfPlane.im_pos z.1) - simp only [coe_eq_fst, div_pow, one_div] at this - simp_rw [lowerBound1] - rw [this, aux4 A B h, one_div, inv_le_inv, add_le_add_iff_right, div_pow] + simp [lowerBound1'] + rw [inv_le_inv, add_le_add_iff_right] apply div_le_div (sq_nonneg _) · simpa [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 · positivity @@ -143,7 +210,7 @@ theorem r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice variable {α : Type*} {β : Type*} {i : α → Set β} -/--Equivalence between the sigma of a fammily of finsets of `ℤ × ℤ` and `ℤ × ℤ`-/ +/--Equivalence between the sigma of a fammily of finsets of `β × β` and `β × β`-/ def sigmaEquiv (ι : α → Finset (β × β)) (HI : ∀ y : β × β , ∃! i : α, y ∈ ι i) : (Σ s : α, ((ι s) : Set (β × β))) ≃ (β × β) where toFun x := x.2 @@ -178,7 +245,7 @@ theorem summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ rw [← Finset.tsum_subtype] rfl -lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : +lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by have hk : 1 < (k - 1 : ℝ) := by have : 1 < (k - 1 : ℤ) := by linarith @@ -193,18 +260,18 @@ lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : norm_cast lemma summable_over_square (k : ℤ) (z : ℍ) (h : 3 ≤ k): - Summable (fun n : ℕ => ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by + Summable (fun n : ℕ => ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by simp only [one_div, Finset.sum_const, nsmul_eq_mul] apply Summable.congr (summable_r_pow k z h) intro b by_cases b0 : b = 0 · rw [b0] - have hk0 : k ≠ 0 := by linarith - have hk1 : k - 1 ≠ 0 := by linarith + have hk0 : k ≠ 0 := by linarith + have hk1 : k - 1 ≠ 0 := by linarith norm_cast rw [zero_zpow k hk0, zero_zpow (k - 1) hk1] simp only [inv_zero, mul_zero, square_zero, Finset.card_singleton, Nat.cast_one] - · rw [square_size' b0, zpow_sub_one₀ (a:= ( b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] + · rw [square_size' b0, zpow_sub_one₀ (a:= ( b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] ring_nf @@ -228,103 +295,7 @@ lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : apply zpow_nonneg (r_pos z).le simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] - -lemma Eise_bound_1 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ n) - (C1 : Complex.abs (x 0 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ - (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by - rw [inv_le_inv] - have h0 : (x 0 : ℂ) ≠ 0 := by - intro hx - rw [hx] at C1 - simp only [map_zero, Int.cast_eq_zero] at * - norm_cast at * - simp_rw [← C1] at hn - simp only [Nat.one_ne_zero, le_zero_iff] at hn - have h1 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = - ((x 0 : ℂ) ^ (k : ℤ)) * ((z : ℂ) + (x 1 : ℂ) / x 0) ^ (k : ℤ) := by - field_simp - ring - rw [h1] - simp_rw [map_mul Complex.abs] - norm_cast at * - have h3 : Complex.abs (n^k : ℕ) = Complex.abs ((x 0)^k : ℤ) := by - have : Complex.abs ((x 0)^k : ℤ) = Complex.abs (x 0)^k := by - simp only [Int.cast_pow, map_pow, Real.rpow_nat_cast] - rw [this, C1] - norm_cast - simp only [Nat.cast_pow, map_pow, abs_natCast] - rw [h3, mul_comm] - apply mul_le_mul_of_nonneg_left _ (Complex.abs.nonneg _) - convert bound1 z x k - apply abs_eq_self.mpr ( pow_nonneg (r_pos z).le k) - norm_cast - have := Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) - apply this - simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] - intro hx - rw [hx] at C1 - simp [Int.cast_zero] at C1 - norm_cast at C1 - rw [← C1] at hn - simp only [Nat.one_ne_zero, le_zero_iff] at hn - apply r_mul_n_pos k z n hn - -lemma Eis_bound_2 (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hn : 1 ≤ n) - (C2 : Complex.abs (x 1 : ℂ) = n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ - (Complex.abs ((r z) ^ (k : ℤ) * n ^ (k : ℤ)))⁻¹ := by - have h0 : (x 1 : ℂ) ≠ 0 := by - intro hx - rw [hx] at C2 - simp at C2 - norm_cast at * - rw [← C2] at hn - simp only [Nat.one_ne_zero, le_zero_iff] at hn - rw [inv_le_inv] - have h1 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = - (x 1 : ℂ)^ (k) * ((x 0 : ℂ) / (x 1 : ℂ) * (z : ℂ) + 1) ^ (k) := by - field_simp - ring_nf - rw [h1] - simp_rw [map_mul Complex.abs] - have h3 : Complex.abs (n^k ) = Complex.abs ((x 1)^k ) := by - simp only [map_pow, abs_natCast, ge_iff_le, Nat.cast_nonneg, map_nonneg, C2] - simp at * - norm_cast at * - rw [h3, mul_comm] - apply mul_le_mul_of_nonneg_left - convert (bound2 z x k) - apply abs_eq_self.mpr ( (r_pos z).le ) - norm_cast - simp [ge_iff_le, map_nonneg, zpow_nonneg] - apply pow_nonneg (Complex.abs.nonneg _) - apply Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) - simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] - intro hx - norm_cast at * - apply r_mul_n_pos k z n hn - -theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) - (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ - (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by - by_cases hn : n = 0 - · rw [hn] at hx - simp only [CharP.cast_eq_zero, square_zero, Finset.mem_singleton, Prod.mk.injEq] at hx - by_cases hk : k = 0 - rw [hk] at * - simp only [coe_eq_fst, pow_zero, map_one, inv_one, mul_one, le_refl] - rw [hx.1, hx.2] - have h1 : (0 : ℝ) ^ (k : ℕ) = 0 := by - simp only [pow_eq_zero_iff', ne_eq, true_and] - exact hk - simp only [Int.cast_zero, coe_eq_fst, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, - CharP.cast_eq_zero, map_mul, abs_ofReal, mul_zero, le_refl] - · have hnn : 1 ≤ n := by exact Nat.one_le_iff_ne_zero.mpr hn - by_cases C1 : Complex.abs (x 0 : ℂ) = n - apply Eise_bound_1 k z n x hnn C1 - have C2 := Complex_abs_square_left_ne n ⟨x 0, x 1⟩ hx C1 - apply Eis_bound_2 k z n x hnn C2 - -lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) +lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by @@ -334,10 +305,10 @@ lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : simp only [top_eq_univ, subset_univ, eisensteinSeries, forall_true_left] intros K hK obtain ⟨A,B,hB, HABK⟩:= compact_in_some_slice K hK - let u := fun x : (gammaSet N a ) => + let u := fun x : (gammaSet N a ) => (1/((r ⟨⟨A, B⟩, hB⟩)^k))* ( (max (x.1 0).natAbs (x.1 1).natAbs : ℝ)^k)⁻¹ have hu : Summable u := by - have := Summable.subtype (summable_upper_bound k hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a ) + have := Summable.subtype (summable_upper_bound k hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a ) apply this.congr intro v simp only [zpow_coe_nat, one_div, Function.comp_apply] From 8e0f4a83e1f7585feb0168a0941ac4c38d18fc2c Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 9 Feb 2024 14:51:02 +0000 Subject: [PATCH 022/124] sections --- .../EisensteinSeries/UniformConvergence.lean | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 93b8b693dee9a..9b207eaf1c7e7 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -28,6 +28,8 @@ open scoped BigOperators NNReal Classical Filter Matrix UpperHalfPlane Complex namespace EisensteinSeries +section bounding_functions + /-- Auxilary function used for bounding Eisentein series-/ def lowerBound1 (z : ℍ) : ℝ := ((z.1.2 ^ (2 : ℕ)) / (z.1.1 ^ (2 : ℕ) + z.1.2 ^ (2 : ℕ))) @@ -49,7 +51,7 @@ theorem r_pos (z : ℍ) : 0 < r z := by theorem r_ne_zero (z : ℍ) : r z ≠ 0 := ne_of_gt (r_pos z) -lemma r_mul_n_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : +lemma r_mul_pos_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : 0 < (Complex.abs ((r z : ℂ) ^ (k : ℤ) * (n : ℂ)^ (k : ℤ))) := by norm_cast apply _root_.abs_pos.mpr (mul_ne_zero (pow_ne_zero k (ne_of_gt (r_pos z))) ?_) @@ -136,7 +138,7 @@ lemma eis_bound_1 (k : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (C1 : 0 < Complex.abs simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] intro hx norm_cast at * - apply r_mul_n_pos k z (Complex.abs (x 0 : ℂ)) C1 + apply r_mul_pos_pos k z (Complex.abs (x 0 : ℂ)) C1 lemma eis_bound_2 (k : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (C2 : 0 < Complex.abs (x 1 : ℂ)) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ @@ -159,7 +161,7 @@ lemma eis_bound_2 (k : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (C2 : 0 < Complex.abs simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] intro hx norm_cast at * - apply r_mul_n_pos k z (Complex.abs (x 1 : ℂ)) C2 + apply r_mul_pos_pos k z (Complex.abs (x 1 : ℂ)) C2 theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ @@ -207,6 +209,9 @@ theorem r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice · positivity · apply (lowerBound1_pos z).le +end bounding_functions + +section summability variable {α : Type*} {β : Type*} {i : α → Set β} @@ -295,6 +300,8 @@ lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : apply zpow_nonneg (r_pos z).le simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] +end summability + lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) From 3e94fcbaecfd08e5b5c12867a59eccf4eab578c5 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Sat, 10 Feb 2024 12:18:19 +0000 Subject: [PATCH 023/124] simplifications, but golfing still needed --- .../Finset_Decomposition.lean | 18 ++ .../EisensteinSeries/UniformConvergence.lean | 221 +++++++++++------- 2 files changed, 161 insertions(+), 78 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean index 9e4a7cea57a40..7652312f2ccaa 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -99,3 +99,21 @@ lemma Complex_abs_eq_of_mem_square (n : ℕ) (x : ℤ × ℤ) (h : x ∈ square lemma Complex_abs_square_left_ne (n : ℕ) (x : ℤ × ℤ) (h : x ∈ square n) (hx : Complex.abs (x.1) ≠ n) : Complex.abs (x.2) = n := Complex_abs_eq_of_mem_square n x h |>.resolve_left hx + +lemma square_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1 ⟩ ∈ square n) : x ≠ 0 ↔ n ≠ 0 := by + constructor + intro h h0 + rw [h0] at hx + simp at hx + have : x = ![x 0, x 1] := by exact List.ofFn_inj.mp rfl + rw [hx.1, hx.2] at this + rw [this] at h + simp at * + intro hn h + have hx0 : x 0 = 0 := by + simp [h] + have hx1 : x 1 = 0 := by + simp [h] + rw [hx0, hx1] at hx + simp at hx + exact hn (id hx.symm) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 9b207eaf1c7e7..f5a70bcec9200 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -49,8 +49,6 @@ def r (z : ℍ) : ℝ := min (z.1.2) (Real.sqrt (lowerBound1 z)) theorem r_pos (z : ℍ) : 0 < r z := by simp only [r, lt_min_iff, z.property, Real.sqrt_pos, lowerBound1_pos, and_self] -theorem r_ne_zero (z : ℍ) : r z ≠ 0 := ne_of_gt (r_pos z) - lemma r_mul_pos_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : 0 < (Complex.abs ((r z : ℂ) ^ (k : ℤ) * (n : ℂ)^ (k : ℤ))) := by norm_cast @@ -59,7 +57,7 @@ lemma r_mul_pos_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : intro _ linarith -theorem auxlb (z : ℍ) (δ ε : ℝ) (hε : 0 ≤ ε^2 - 1) : +theorem auxlb (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : (z.1.2 ^ 2 ) / (z.1.1 ^ 2 + z.1.2 ^ 2) ≤ (δ * z.1.1 + ε) ^ 2 + (δ * z.1.2) ^ 2 := by have H1 : (δ * z.1.1 + ε) ^ 2 + (δ * z.1.2) ^ 2 = δ ^ 2 * (z.1.1 ^ 2 + z.1.2 ^ 2) + ε * 2 * δ * z.1.1 + ε^2 := by ring @@ -68,27 +66,30 @@ theorem auxlb (z : ℍ) (δ ε : ℝ) (hε : 0 ≤ ε^2 - 1) : rw [H1, div_le_iff, ← sub_nonneg, H4] · apply add_nonneg (pow_two_nonneg _) ?_ apply mul_nonneg - norm_cast + linarith apply pow_two_nonneg · apply_rules [add_pos_of_nonneg_of_pos, pow_two_nonneg, (pow_pos z.im_pos 2)] -theorem auxbound1 (z : ℍ) (δ : ℝ) : r z ≤ Complex.abs ((z : ℂ) + δ) := by +theorem auxbound1 (z : ℍ) (δ ε: ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (ε * (z : ℂ) + δ) := by rw [r, Complex.abs] have H1 : (z : ℂ).im ≤ - Real.sqrt (((z : ℂ).re + δ) * ((z : ℂ).re + δ) + (z : ℂ).im * (z : ℂ).im) := by - rw [Real.le_sqrt' ] + Real.sqrt ((ε * (z : ℂ).re + δ) * (ε * (z : ℂ).re + δ) + (ε * z : ℂ).im * (ε * z : ℂ).im) := by + have h1 : (ε * z : ℂ).im * (ε * z : ℂ).im = ε^2 * (z : ℂ).im * (z : ℂ).im := by + simp only [mul_im, ofReal_re, coe_im, ofReal_im, coe_re, zero_mul, add_zero] + ring + rw [Real.le_sqrt', h1 ] nlinarith exact z.2 simp only [UpperHalfPlane.coe_im, UpperHalfPlane.coe_re, AbsoluteValue.coe_mk, MulHom.coe_mk, min_le_iff] at * left - simp only [normSq_apply, add_re, coe_re, ofReal_re, add_im, coe_im, ofReal_im, add_zero] + simp [normSq_apply, add_re, coe_re, ofReal_re, add_im, coe_im, ofReal_im, add_zero] at * exact H1 -theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 0 ≤ ε^2 - 1) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by + +theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs, min_le_iff] - have H1 : - Real.sqrt (lowerBound1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε ) + + have H1 : Real.sqrt (lowerBound1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε ) + δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by rw [lowerBound1, Real.sqrt_le_sqrt_iff, ← pow_two, ← pow_two] apply auxlb z δ ε hε @@ -99,69 +100,116 @@ theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 0 ≤ ε^2 - 1) : r z ≤ Compl int_cast_im] at * exact H1 -theorem bound1 (z : ℍ) (x : Fin 2 → ℤ) (k : ℕ) : - (r z ) ^ k ≤ Complex.abs (((z : ℂ) + (x 1 : ℂ) / (x 0 : ℂ)) ^ k) := by - have := auxbound1 z (x 1 / x 0 : ℝ) - simp only [zpow_coe_nat, map_pow, ge_iff_le] - apply pow_le_pow_left (r_pos _).le - simp only [ofReal_div, ofReal_int_cast] at * - exact this -theorem bound2 (z : ℍ) (x : Fin 2 → ℤ) (k : ℕ) : - (r z ) ^ k ≤ Complex.abs (((x 0 : ℂ) / (x 1 : ℂ) * (z : ℂ) + 1) ^ k) := by - have := auxbound2 z (x 0 / x 1 : ℝ) 1 (by norm_cast) - simp only [zpow_coe_nat, map_pow, ge_iff_le] - apply pow_le_pow_left (r_pos _).le - simp only [ofReal_div, ofReal_int_cast, Int.cast_one] at * - exact this +lemma asfd (a : ℝ) (ha : a ≠ 0) : 1 ≤ a^2 / |a|^2 := by + simp + rw [le_div_iff'] + simp + exact (sq_pos_iff a).mpr ha -lemma fds (z : ℂ) (hz : 0 < Complex.abs (z)) : z ≠ 0 := by exact nnnorm_pos.mp hz +lemma int_abs_eq_complex_abs (a : ℤ) : Complex.abs a = a.natAbs := by + have : Complex.abs a = Complex.abs (a : ℝ) := by rfl + rw [this, abs_ofReal] + norm_cast + rw [Int.abs_eq_natAbs a] + rfl + +lemma ne_zero_if_max (x : Fin 2 → ℤ) (hx : x ≠ 0) + (h : (max (x 0).natAbs (x 1).natAbs) = (x 0).natAbs) : (x 0) ≠ 0 := by + intro h0 + rw [h0] at h + simp at h + simp at hx + have : x = ![x 0, x 1] := by + exact List.ofFn_inj.mp rfl + rw [h0, h] at this + rw [this] at hx + simp at hx + +lemma ne_zero_if_max' (x : Fin 2 → ℤ) (hx : x ≠ 0) + (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by + intro h0 + rw [h0] at h + simp at h + simp at hx + have : x = ![x 0, x 1] := by + exact List.ofFn_inj.mp rfl + rw [h0, h] at this + rw [this] at hx + simp at hx + + +lemma sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs))^2 ∨ + (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs))^2 := by + have h := max_choice (x 0).natAbs (x 1).natAbs + cases' h with H1 H2 + left + rw [H1] + have : (x 0 : ℝ) ≠ 0 := by + have t := ne_zero_if_max x hx H1 + simpa using t + have h1 := asfd (x 0 : ℝ) this + simp at * + convert h1 + norm_cast + rw [int_abs_eq_complex_abs] + right + rw [H2] + have : (x 1 : ℝ) ≠ 0 := by + have t := ne_zero_if_max' x hx H2 + simpa using t + have h1 := asfd (x 1 : ℝ) this + simp at * + convert h1 + norm_cast + rw [int_abs_eq_complex_abs] -lemma eis_bound_1 (k : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (C1 : 0 < Complex.abs (x 0 : ℂ)) : - (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ - (Complex.abs ((r z) ^ (k : ℤ) * (Complex.abs (x 0 : ℂ)) ^ (k : ℤ)))⁻¹ := by - have h0 : (x 0 : ℂ) ≠ 0 := nnnorm_pos.mp C1 - rw [inv_le_inv] - have h1 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = - ((x 0 : ℂ) ^ (k : ℤ)) * ((z : ℂ) + (x 1 : ℂ) / x 0) ^ (k : ℤ) := by - field_simp - ring - rw [h1] - have gd := mul_le_mul_of_nonneg_left (bound1 z x k) (pow_nonneg C1.le k) - simp only [AbsoluteValue.pos_iff, ne_eq, Int.cast_eq_zero, zpow_coe_nat, map_pow, map_mul, - abs_ofReal, Complex.abs_abs, ge_iff_le] at * - rw [ mul_comm] - convert gd - apply abs_eq_self.mpr ((r_pos z).le ) - have := Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) - apply this - simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] - intro hx - norm_cast at * - apply r_mul_pos_pos k z (Complex.abs (x 0 : ℂ)) C1 -lemma eis_bound_2 (k : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (C2 : 0 < Complex.abs (x 1 : ℂ)) : - (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ ≤ - (Complex.abs ((r z) ^ (k : ℤ) * (Complex.abs (x 1 : ℂ)) ^ (k : ℤ)))⁻¹ := by - have h0 : (x 1 : ℂ) ≠ 0 := nnnorm_pos.mp C2 - rw [inv_le_inv] - have h1 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = - (x 1 : ℂ)^ (k) * ((x 0 : ℂ) / (x 1 : ℂ) * (z : ℂ) + 1) ^ (k) := by +lemma bound3 (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : + ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs)^k ≤ + Complex.abs (((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ)) ^ k) := by + by_cases hk : k ≠ 0 + let n := max (x 0).natAbs (x 1).natAbs + have h1 : ((n : ℝ) : ℂ)^k ≠ 0 := by + rw [pow_ne_zero_iff] + norm_cast + have := EisensteinSeries.square_ne_zero n x ?_ + rw [this] at hx + exact hx + rw [square_mem] + exact hk + have hc : Complex.abs ((n : ℝ)^k : ℂ) = n^k := by + simp + congr + simp + have h11 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = + (((x 0 : ℝ) / (n : ℝ) ) * (z : ℂ) + (x 1 : ℝ) /(n : ℝ)) ^ (k : ℤ) * ((n : ℝ)^ (k : ℤ)) := by + simp only [Nat.cast_max] at h1 field_simp - ring_nf - rw [h1] - have gd := mul_le_mul_of_nonneg_left (bound2 z x k) (pow_nonneg C2.le k) - simp only [AbsoluteValue.pos_iff, ne_eq, Int.cast_eq_zero, zpow_coe_nat, map_pow, map_mul, - abs_ofReal, Complex.abs_abs, ge_iff_le] at * - rw [ mul_comm] - convert gd - apply abs_eq_self.mpr ((r_pos z).le ) - have := Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) - apply this - simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] - intro hx + have h2 : (1 : ℝ) ≤ ((x 0)/n)^2 ∨ (1 : ℝ) ≤ (x 1/ n)^2 := sq_ge_one x hx + cases' h2 with H1 H2 norm_cast at * - apply r_mul_pos_pos k z (Complex.abs (x 1 : ℂ)) C2 + rw [h11] + simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * + apply mul_le_mul ?_ (by rfl) + simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, + pow_nonneg] + apply pow_nonneg (Complex.abs.nonneg _) k + apply pow_le_pow_left (r_pos _).le + apply auxbound1 z ((x 1 /n) : ℝ) (((x 0 : ℝ) / (n : ℝ))) H1 + norm_cast at * + rw [h11] + simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * + apply mul_le_mul ?_ (by rfl) + simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, + pow_nonneg] + apply pow_nonneg (Complex.abs.nonneg _) k + apply pow_le_pow_left (r_pos _).le + apply auxbound2 z ((x 0 /n) : ℝ) (((x 1 : ℝ) / (n : ℝ))) H2 + simp at hk + simp_rw [hk] + simp only [pow_zero, Nat.cast_max, mul_one, map_one, le_refl] + theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ @@ -178,15 +226,32 @@ theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → exact hk simp only [Int.cast_zero, coe_eq_fst, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, CharP.cast_eq_zero, map_mul, abs_ofReal, mul_zero, le_refl] - · have hnn : 0 < (n : ℝ) := by norm_cast; exact Nat.pos_of_ne_zero hn - by_cases C1 : Complex.abs (x 0 : ℂ) = n - · have := eis_bound_1 k z x (by rw [C1]; exact hnn) - rw [C1] at * - exact this - · have C2 := Complex_abs_square_left_ne n ⟨x 0, x 1⟩ hx C1 - have := eis_bound_2 k z x (by rw [C2]; exact hnn) - rw [C2] at this - exact this + · have hx2 : x ≠ 0 := by + rw [EisensteinSeries.square_ne_zero n x hx] + exact hn + simp at hx + rw [inv_le_inv] + have := bound3 z x hx2 k + rw [←hx] + simp + convert this + apply abs_eq_self.mpr ((r_pos z).le ) + simp + simp + have := Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) + apply this + simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, + not_and] at * + intro hg + intro h1 + have : x = ![x 0, x 1] := by + exact List.ofFn_inj.mp rfl + rw [this] at hx2 + rw [hg,h1] at hx2 + simp at * + apply r_mul_pos_pos + norm_cast + exact Nat.pos_of_ne_zero hn theorem r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by @@ -333,7 +398,7 @@ lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : apply pow_le_pow_left (r_pos _).le rw [abs_of_pos (r_pos _)] · exact r_lower_bound_on_slice A B hB ⟨x, HABK hx⟩ - · apply pow_pos (abs_pos.mpr (r_ne_zero x)) + · apply pow_pos (abs_pos.mpr (ne_of_gt (r_pos x))) · apply pow_pos (r_pos _) · rfl repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg, From 6012f67081292417d6de5206fdd3f005a12a4b39 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Sat, 10 Feb 2024 13:58:10 +0000 Subject: [PATCH 024/124] save --- .../Finset_Decomposition.lean | 18 +++--- .../EisensteinSeries/UniformConvergence.lean | 59 ++++++++----------- 2 files changed, 35 insertions(+), 42 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean index 7652312f2ccaa..8a798ff948ff5 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -100,20 +100,22 @@ lemma Complex_abs_square_left_ne (n : ℕ) (x : ℤ × ℤ) (h : x ∈ square n) (hx : Complex.abs (x.1) ≠ n) : Complex.abs (x.2) = n := Complex_abs_eq_of_mem_square n x h |>.resolve_left hx +lemma fun_ne_zero_cases (x : Fin 2 → ℤ) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by + rw [Function.ne_iff] + exact Fin.exists_fin_two + lemma square_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1 ⟩ ∈ square n) : x ≠ 0 ↔ n ≠ 0 := by constructor intro h h0 rw [h0] at hx - simp at hx - have : x = ![x 0, x 1] := by exact List.ofFn_inj.mp rfl - rw [hx.1, hx.2] at this - rw [this] at h - simp at * + simp only [Nat.cast_zero, square_zero, mem_singleton, Prod.mk.injEq] at hx + rw [fun_ne_zero_cases, hx.1, hx.2] at h + simp only [ne_eq, not_true_eq_false, or_self] at * intro hn h have hx0 : x 0 = 0 := by - simp [h] + simp only [h, Pi.zero_apply] have hx1 : x 1 = 0 := by - simp [h] + simp only [h, Pi.zero_apply] rw [hx0, hx1] at hx - simp at hx + simp only [square_mem, Int.natAbs_zero, max_self] at hx exact hn (id hx.symm) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index f5a70bcec9200..054df43ec7c1b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -101,10 +101,9 @@ theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.a exact H1 -lemma asfd (a : ℝ) (ha : a ≠ 0) : 1 ≤ a^2 / |a|^2 := by - simp - rw [le_div_iff'] - simp +lemma one_le_sq_div_abs_sq (a : ℝ) (ha : a ≠ 0) : 1 ≤ a^2 / |a|^2 := by + rw [_root_.sq_abs, le_div_iff'] + simp only [mul_one, le_refl] exact (sq_pos_iff a).mpr ha lemma int_abs_eq_complex_abs (a : ℤ) : Complex.abs a = a.natAbs := by @@ -115,28 +114,23 @@ lemma int_abs_eq_complex_abs (a : ℤ) : Complex.abs a = a.natAbs := by rfl lemma ne_zero_if_max (x : Fin 2 → ℤ) (hx : x ≠ 0) - (h : (max (x 0).natAbs (x 1).natAbs) = (x 0).natAbs) : (x 0) ≠ 0 := by + (h : (max (x 0).natAbs (x 1).natAbs) = (x 0).natAbs) : (x 0) ≠ 0 := by intro h0 rw [h0] at h - simp at h - simp at hx + simp only [ne_eq, Int.natAbs_zero, ge_iff_le, zero_le, max_eq_right, Int.natAbs_eq_zero] at * have : x = ![x 0, x 1] := by exact List.ofFn_inj.mp rfl rw [h0, h] at this rw [this] at hx - simp at hx + simp only [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at hx lemma ne_zero_if_max' (x : Fin 2 → ℤ) (hx : x ≠ 0) - (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by - intro h0 - rw [h0] at h - simp at h - simp at hx - have : x = ![x 0, x 1] := by - exact List.ofFn_inj.mp rfl - rw [h0, h] at this - rw [this] at hx - simp at hx + (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by + apply ne_zero_if_max ![x 1, x 0] ?_ (by simpa using h) + simp only [ne_eq, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_true, not_and] + intro h1 h0 + rw [fun_ne_zero_cases, h1, h0] at hx + simp only [ne_eq, not_true_eq_false, or_self] at * lemma sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs))^2 ∨ @@ -148,8 +142,8 @@ lemma sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max ( have : (x 0 : ℝ) ≠ 0 := by have t := ne_zero_if_max x hx H1 simpa using t - have h1 := asfd (x 0 : ℝ) this - simp at * + have h1 := one_le_sq_div_abs_sq (x 0 : ℝ) this + simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 norm_cast rw [int_abs_eq_complex_abs] @@ -158,14 +152,14 @@ lemma sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max ( have : (x 1 : ℝ) ≠ 0 := by have t := ne_zero_if_max' x hx H2 simpa using t - have h1 := asfd (x 1 : ℝ) this - simp at * + have h1 := one_le_sq_div_abs_sq (x 1 : ℝ) this + simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 norm_cast rw [int_abs_eq_complex_abs] -lemma bound3 (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : +lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs)^k ≤ Complex.abs (((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ)) ^ k) := by by_cases hk : k ≠ 0 @@ -229,26 +223,23 @@ theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → · have hx2 : x ≠ 0 := by rw [EisensteinSeries.square_ne_zero n x hx] exact hn - simp at hx + simp only [square_mem] at hx rw [inv_le_inv] - have := bound3 z x hx2 k - rw [←hx] - simp + have := bound z x hx2 k + simp only [← hx, map_mul, map_pow, abs_ofReal, abs_natCast, Nat.cast_max, ge_iff_le] convert this apply abs_eq_self.mpr ((r_pos z).le ) - simp - simp + simp only [Nat.cast_max] + simp only [map_pow] have := Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) apply this simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] at * - intro hg - intro h1 + intro hg h1 have : x = ![x 0, x 1] := by exact List.ofFn_inj.mp rfl - rw [this] at hx2 - rw [hg,h1] at hx2 - simp at * + rw [this, hg,h1 ] at hx2 + simp only [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at * apply r_mul_pos_pos norm_cast exact Nat.pos_of_ne_zero hn From d63d918ba47748ff3a62a10ec4b4fa291a2b8fbc Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Sat, 10 Feb 2024 13:59:49 +0000 Subject: [PATCH 025/124] save2 --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 054df43ec7c1b..bd255a38e858d 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -205,7 +205,7 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : simp only [pow_zero, Nat.cast_max, mul_one, map_one, le_refl] -theorem Eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) +theorem eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by by_cases hn : n = 0 @@ -378,7 +378,7 @@ lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : apply tendstoUniformlyOn_tsum hu intro v x hx have sq := square_mem (max (v.1 0).natAbs (v.1 1).natAbs ) ⟨(v.1 0), v.1 1⟩ - have := Eis_is_bounded_on_square k x (max (v.1 0).natAbs (v.1 1).natAbs ) v + have := eis_is_bounded_on_square k x (max (v.1 0).natAbs (v.1 1).natAbs ) v simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_coe_nat, one_div, coe_eq_fst, map_pow, map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, ge_iff_le] at * From bc282a0f23b5bf90adde0ca5f71a31ab63b5a9c2 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Sat, 10 Feb 2024 14:23:38 +0000 Subject: [PATCH 026/124] lint fix --- .../EisensteinSeries/UniformConvergence.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index bd255a38e858d..0ce7d0da23149 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -125,7 +125,7 @@ lemma ne_zero_if_max (x : Fin 2 → ℤ) (hx : x ≠ 0) simp only [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at hx lemma ne_zero_if_max' (x : Fin 2 → ℤ) (hx : x ≠ 0) - (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by + (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by apply ne_zero_if_max ![x 1, x 0] ?_ (by simpa using h) simp only [ne_eq, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_true, not_and] intro h1 h0 @@ -140,7 +140,7 @@ lemma sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max ( left rw [H1] have : (x 0 : ℝ) ≠ 0 := by - have t := ne_zero_if_max x hx H1 + have t := ne_zero_if_max x hx H1 simpa using t have h1 := one_le_sq_div_abs_sq (x 0 : ℝ) this simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * @@ -150,7 +150,7 @@ lemma sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max ( right rw [H2] have : (x 1 : ℝ) ≠ 0 := by - have t := ne_zero_if_max' x hx H2 + have t := ne_zero_if_max' x hx H2 simpa using t have h1 := one_le_sq_div_abs_sq (x 1 : ℝ) this simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * @@ -176,11 +176,11 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : simp congr simp - have h11 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = + have h11 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = (((x 0 : ℝ) / (n : ℝ) ) * (z : ℂ) + (x 1 : ℝ) /(n : ℝ)) ^ (k : ℤ) * ((n : ℝ)^ (k : ℤ)) := by simp only [Nat.cast_max] at h1 field_simp - have h2 : (1 : ℝ) ≤ ((x 0)/n)^2 ∨ (1 : ℝ) ≤ (x 1/ n)^2 := sq_ge_one x hx + have h2 : (1 : ℝ) ≤ ((x 0)/n)^2 ∨ (1 : ℝ) ≤ (x 1/ n)^2 := sq_ge_one x hx cases' h2 with H1 H2 norm_cast at * rw [h11] From 49d5d22486c12a6ef8cecea317998ffce0d9b800 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Sat, 10 Feb 2024 16:48:24 +0000 Subject: [PATCH 027/124] golf --- .../EisensteinSeries/UniformConvergence.lean | 147 ++++++++---------- 1 file changed, 66 insertions(+), 81 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 0ce7d0da23149..ddf9b7183a03d 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -70,11 +70,11 @@ theorem auxlb (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : apply pow_two_nonneg · apply_rules [add_pos_of_nonneg_of_pos, pow_two_nonneg, (pow_pos z.im_pos 2)] -theorem auxbound1 (z : ℍ) (δ ε: ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (ε * (z : ℂ) + δ) := by +theorem auxbound1 (z : ℍ) (δ ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs] have H1 : (z : ℂ).im ≤ - Real.sqrt ((ε * (z : ℂ).re + δ) * (ε * (z : ℂ).re + δ) + (ε * z : ℂ).im * (ε * z : ℂ).im) := by - have h1 : (ε * z : ℂ).im * (ε * z : ℂ).im = ε^2 * (z : ℂ).im * (z : ℂ).im := by + Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + (δ * z : ℂ).im * (δ * z : ℂ).im) := by + have h1 : (δ * z : ℂ).im * (δ * z : ℂ).im = δ^2 * (z : ℂ).im * (z : ℂ).im := by simp only [mul_im, ofReal_re, coe_im, ofReal_im, coe_re, zero_mul, add_zero] ring rw [Real.le_sqrt', h1 ] @@ -83,10 +83,10 @@ theorem auxbound1 (z : ℍ) (δ ε: ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.ab simp only [UpperHalfPlane.coe_im, UpperHalfPlane.coe_re, AbsoluteValue.coe_mk, MulHom.coe_mk, min_le_iff] at * left - simp [normSq_apply, add_re, coe_re, ofReal_re, add_im, coe_im, ofReal_im, add_zero] at * + simp only [one_le_sq_iff_one_le_abs, mul_im, ofReal_re, coe_im, ofReal_im, coe_re, zero_mul, + add_zero, normSq_apply, add_re, mul_re, sub_zero, add_im] at * exact H1 - theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs, min_le_iff] have H1 : Real.sqrt (lowerBound1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε ) + @@ -100,7 +100,6 @@ theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.a int_cast_im] at * exact H1 - lemma one_le_sq_div_abs_sq (a : ℝ) (ha : a ≠ 0) : 1 ≤ a^2 / |a|^2 := by rw [_root_.sq_abs, le_div_iff'] simp only [mul_one, le_refl] @@ -132,78 +131,66 @@ lemma ne_zero_if_max' (x : Fin 2 → ℤ) (hx : x ≠ 0) rw [fun_ne_zero_cases, h1, h0] at hx simp only [ne_eq, not_true_eq_false, or_self] at * - lemma sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs))^2 ∨ (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs))^2 := by - have h := max_choice (x 0).natAbs (x 1).natAbs - cases' h with H1 H2 - left - rw [H1] - have : (x 0 : ℝ) ≠ 0 := by - have t := ne_zero_if_max x hx H1 - simpa using t - have h1 := one_le_sq_div_abs_sq (x 0 : ℝ) this - simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * - convert h1 - norm_cast - rw [int_abs_eq_complex_abs] - right - rw [H2] - have : (x 1 : ℝ) ≠ 0 := by - have t := ne_zero_if_max' x hx H2 - simpa using t - have h1 := one_le_sq_div_abs_sq (x 1 : ℝ) this - simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * - convert h1 - norm_cast - rw [int_abs_eq_complex_abs] - + cases' (max_choice (x 0).natAbs (x 1).natAbs) with H1 H2 + · left + rw [H1] + have : (x 0 : ℝ) ≠ 0 := by + simpa using (ne_zero_if_max x hx H1) + have h1 := one_le_sq_div_abs_sq (x 0 : ℝ) this + simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * + convert h1 + norm_cast + rw [int_abs_eq_complex_abs] + · right + rw [H2] + have : (x 1 : ℝ) ≠ 0 := by + simpa using (ne_zero_if_max' x hx H2) + have h1 := one_le_sq_div_abs_sq (x 1 : ℝ) this + simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * + convert h1 + norm_cast + rw [int_abs_eq_complex_abs] lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs)^k ≤ Complex.abs (((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ)) ^ k) := by by_cases hk : k ≠ 0 - let n := max (x 0).natAbs (x 1).natAbs - have h1 : ((n : ℝ) : ℂ)^k ≠ 0 := by - rw [pow_ne_zero_iff] - norm_cast - have := EisensteinSeries.square_ne_zero n x ?_ - rw [this] at hx - exact hx - rw [square_mem] - exact hk - have hc : Complex.abs ((n : ℝ)^k : ℂ) = n^k := by - simp - congr - simp - have h11 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = - (((x 0 : ℝ) / (n : ℝ) ) * (z : ℂ) + (x 1 : ℝ) /(n : ℝ)) ^ (k : ℤ) * ((n : ℝ)^ (k : ℤ)) := by - simp only [Nat.cast_max] at h1 - field_simp - have h2 : (1 : ℝ) ≤ ((x 0)/n)^2 ∨ (1 : ℝ) ≤ (x 1/ n)^2 := sq_ge_one x hx - cases' h2 with H1 H2 - norm_cast at * - rw [h11] - simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * - apply mul_le_mul ?_ (by rfl) - simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, - pow_nonneg] - apply pow_nonneg (Complex.abs.nonneg _) k - apply pow_le_pow_left (r_pos _).le - apply auxbound1 z ((x 1 /n) : ℝ) (((x 0 : ℝ) / (n : ℝ))) H1 - norm_cast at * - rw [h11] - simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * - apply mul_le_mul ?_ (by rfl) - simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, - pow_nonneg] - apply pow_nonneg (Complex.abs.nonneg _) k - apply pow_le_pow_left (r_pos _).le - apply auxbound2 z ((x 0 /n) : ℝ) (((x 1 : ℝ) / (n : ℝ))) H2 - simp at hk - simp_rw [hk] - simp only [pow_zero, Nat.cast_max, mul_one, map_one, le_refl] - + · let n := max (x 0).natAbs (x 1).natAbs + have h1 : ((n : ℝ) : ℂ)^k ≠ 0 := by + rw [pow_ne_zero_iff hk] + norm_cast + rw [EisensteinSeries.square_ne_zero n x (by rw [square_mem])] at hx + exact hx + have hc : Complex.abs ((n : ℝ)^k : ℂ) = n^k := by + simp only [Nat.cast_max, map_pow, abs_ofReal, ge_iff_le, abs_nonneg, le_max_iff, + Nat.cast_nonneg, or_self] + congr + simp only [abs_eq_self, le_max_iff, Nat.cast_nonneg, or_self] + have h11 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = + (((x 0 : ℝ) / (n : ℝ) ) * (z : ℂ) + (x 1 : ℝ) /(n : ℝ)) ^ (k : ℤ) * ((n : ℝ)^ (k : ℤ)) := by + simp only [Nat.cast_max] at h1 + field_simp + cases' (sq_ge_one x hx) with H1 H2 + · norm_cast at * + rw [h11] + simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * + apply mul_le_mul ?_ (by rfl) + simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, + pow_nonneg] + apply pow_nonneg (Complex.abs.nonneg _) k + apply pow_le_pow_left (r_pos _).le (auxbound1 z (x 0 / n) (x 1 / n) H1) + · norm_cast at * + rw [h11] + simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * + apply mul_le_mul ?_ (by rfl) + simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, + pow_nonneg] + apply pow_nonneg (Complex.abs.nonneg _) k + apply pow_le_pow_left (r_pos _).le (auxbound2 z (x 0 / n) (x 1 / n) H2) + · simp only [ne_eq, not_not] at hk + simp only [hk, pow_zero, Nat.cast_max, mul_one, map_one, le_refl] theorem eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ @@ -244,7 +231,7 @@ theorem eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → norm_cast exact Nat.pos_of_ne_zero hn -theorem r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice A B) : +lemma r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by have hz := z.2 simp only [slice_mem, abs_ofReal, ge_iff_le] at hz @@ -283,7 +270,7 @@ def sigmaEquiv (ι : α → Finset (β × β)) (HI : ∀ y : β × β , ∃! i : repeat {rfl} right_inv x := by rfl -theorem summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), 0 ≤ f y) +lemma summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), 0 ≤ f y) (ι : ℕ → Finset (ℤ × ℤ)) (HI : ∀ y : ℤ × ℤ, ∃! i : ℕ, y ∈ ι i) : Summable f ↔ Summable fun n : ℕ => ∑ x in ι n, f ![x.1, x.2] := by let h2 := Equiv.trans (sigmaEquiv ι HI) (piFinTwoEquiv fun _ => ℤ).symm @@ -358,21 +345,19 @@ lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : end summability -lemma eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) +theorem eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => - (fun (z : ℍ) => ∑ x in s, eisSummand k x z ) ) + (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by have hk0 : 0 ≤ k := by linarith lift k to ℕ using hk0 rw [tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] simp only [top_eq_univ, subset_univ, eisensteinSeries, forall_true_left] intros K hK - obtain ⟨A,B,hB, HABK⟩:= compact_in_some_slice K hK - let u := fun x : (gammaSet N a ) => - (1/((r ⟨⟨A, B⟩, hB⟩)^k))* ( (max (x.1 0).natAbs (x.1 1).natAbs : ℝ)^k)⁻¹ - have hu : Summable u := by - have := Summable.subtype (summable_upper_bound k hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a ) - apply this.congr + obtain ⟨A, B, hB, HABK⟩:= compact_in_some_slice K hK + have hu : Summable fun x : (gammaSet N a ) => + (1/(r ⟨⟨A, B⟩, hB⟩) ^ k) * ((max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by + apply (Summable.subtype (summable_upper_bound k hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a )).congr intro v simp only [zpow_coe_nat, one_div, Function.comp_apply] apply tendstoUniformlyOn_tsum hu From edf175241041b66fbcf7e515845457c7bd73b48b Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Sat, 10 Feb 2024 16:58:40 +0000 Subject: [PATCH 028/124] save --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index ddf9b7183a03d..5f635a6ac5144 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -256,14 +256,13 @@ end bounding_functions section summability -variable {α : Type*} {β : Type*} {i : α → Set β} +variable {α : Type*} {β : Type*} /--Equivalence between the sigma of a fammily of finsets of `β × β` and `β × β`-/ def sigmaEquiv (ι : α → Finset (β × β)) (HI : ∀ y : β × β , ∃! i : α, y ∈ ι i) : (Σ s : α, ((ι s) : Set (β × β))) ≃ (β × β) where toFun x := x.2 - invFun x := by - refine ⟨(HI x).choose, x, (HI x).choose_spec.1⟩ + invFun x := ⟨(HI x).choose, x, (HI x).choose_spec.1⟩ left_inv x := by ext exact ((HI x.2).choose_spec.2 x.1 x.2.2).symm From 38ca47f2df438d7728ee23ac378639f5a9c9878e Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Mon, 12 Feb 2024 10:10:14 +0000 Subject: [PATCH 029/124] more golf --- Mathlib/Data/Finset/Basic.lean | 12 +++++ .../EisensteinSeries/UniformConvergence.lean | 49 ++++++------------- 2 files changed, 27 insertions(+), 34 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index f6902ed81d678..2eb3dd1dd454e 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -3962,6 +3962,18 @@ def piFinsetUnion {ι} [DecidableEq ι] (α : ι → Type*) {s t : Finset ι} (h let e := Equiv.Finset.union s t h sumPiEquivProdPi (fun b ↦ α (e b)) |>.symm.trans (.piCongrLeft (fun i : ↥(s ∪ t) ↦ α i) e) +/--Equivalence between the sigma of a fammily of finsets of `β × β` and `β × β`-/ +noncomputable +def sigmaEquiv {α : Type*} {β : Type*} (ι : α → Finset (β × β)) + (HI : ∀ y : β × β , ∃! i : α, y ∈ ι i) : (Σ s : α, ((ι s) : Set (β × β))) ≃ (β × β) where + toFun x := x.2 + invFun x := ⟨(HI x).choose, x, (HI x).choose_spec.1⟩ + left_inv x := by + ext + exact ((HI x.2).choose_spec.2 x.1 x.2.2).symm + repeat {rfl} + right_inv x := by rfl + end Equiv namespace Multiset diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 5f635a6ac5144..21dee3ea99604 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -31,23 +31,21 @@ namespace EisensteinSeries section bounding_functions /-- Auxilary function used for bounding Eisentein series-/ -def lowerBound1 (z : ℍ) : ℝ := - ((z.1.2 ^ (2 : ℕ)) / (z.1.1 ^ (2 : ℕ) + z.1.2 ^ (2 : ℕ))) +def lowerBound1 (z : ℍ) : ℝ := ((z.im ^ (2 : ℕ)) / (z.re ^ (2 : ℕ) + z.im ^ (2 : ℕ))) -lemma lowerBound1' (z : ℍ) : lowerBound1 z = 1/((z.1.1 / z.1.2) ^ (2 : ℕ) + 1) := by - have := z.2 - field_simp [lowerBound1] +lemma lowerBound1' (z : ℍ) : lowerBound1 z = 1/((z.re / z.im) ^ (2 : ℕ) + 1) := by + field_simp [lowerBound1, im_pos z] theorem lowerBound1_pos (z : ℍ) : 0 < lowerBound1 z := by - have H2 : 0 < (z.1.1 ^ (2 : ℕ) + z.1.2 ^ (2 : ℕ)) := by + have H2 : 0 < (z.re ^ (2 : ℕ) + z.im ^ (2 : ℕ)) := by apply_rules [pow_pos, add_pos_of_nonneg_of_pos, pow_two_nonneg, z.2] exact div_pos (pow_pos z.im_pos 2) H2 /-- This function is used to give an upper bound on Eisenstein series-/ -def r (z : ℍ) : ℝ := min (z.1.2) (Real.sqrt (lowerBound1 z)) +def r (z : ℍ) : ℝ := min (z.im) (Real.sqrt (lowerBound1 z)) theorem r_pos (z : ℍ) : 0 < r z := by - simp only [r, lt_min_iff, z.property, Real.sqrt_pos, lowerBound1_pos, and_self] + simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, lowerBound1_pos, and_self] lemma r_mul_pos_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : 0 < (Complex.abs ((r z : ℂ) ^ (k : ℤ) * (n : ℂ)^ (k : ℤ))) := by @@ -58,11 +56,11 @@ lemma r_mul_pos_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : linarith theorem auxlb (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : - (z.1.2 ^ 2 ) / (z.1.1 ^ 2 + z.1.2 ^ 2) ≤ (δ * z.1.1 + ε) ^ 2 + (δ * z.1.2) ^ 2 := by - have H1 : (δ * z.1.1 + ε) ^ 2 + (δ * z.1.2) ^ 2 = - δ ^ 2 * (z.1.1 ^ 2 + z.1.2 ^ 2) + ε * 2 * δ * z.1.1 + ε^2 := by ring - have H4 : (δ ^ 2 * (z.1.1 ^ 2 + z.1.2 ^ 2) + ε * 2 * δ * z.1.1 + ε^2) * (z.1.1 ^ 2 + z.1.2 ^ 2) - - (z.1.2 ^ 2) = (δ * (z.1.1 ^ 2 + z.1.2 ^ 2)+ ε * z.1.1)^2 + (ε^2 - 1)* (z.1.2)^2 := by ring + (z.im ^ 2 ) / (z.re ^ 2 + z.im ^ 2) ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by + have H1 : (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 = + δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2 := by ring + have H4 : (δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2) * (z.re ^ 2 + z.im ^ 2) + - (z.im ^ 2) = (δ * (z.re ^ 2 + z.im ^ 2)+ ε * z.re)^2 + (ε^2 - 1)* (z.im)^2 := by ring rw [H1, div_le_iff, ← sub_nonneg, H4] · apply add_nonneg (pow_two_nonneg _) ?_ apply mul_nonneg @@ -117,11 +115,9 @@ lemma ne_zero_if_max (x : Fin 2 → ℤ) (hx : x ≠ 0) intro h0 rw [h0] at h simp only [ne_eq, Int.natAbs_zero, ge_iff_le, zero_le, max_eq_right, Int.natAbs_eq_zero] at * - have : x = ![x 0, x 1] := by - exact List.ofFn_inj.mp rfl + have : x = ![x 0, x 1] := List.ofFn_inj.mp rfl rw [h0, h] at this - rw [this] at hx - simp only [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at hx + simp only [this, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at hx lemma ne_zero_if_max' (x : Fin 2 → ℤ) (hx : x ≠ 0) (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by @@ -212,14 +208,12 @@ theorem eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → exact hn simp only [square_mem] at hx rw [inv_le_inv] - have := bound z x hx2 k simp only [← hx, map_mul, map_pow, abs_ofReal, abs_natCast, Nat.cast_max, ge_iff_le] - convert this + convert (bound z x hx2 k) apply abs_eq_self.mpr ((r_pos z).le ) simp only [Nat.cast_max] simp only [map_pow] - have := Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) - apply this + apply Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] at * intro hg h1 @@ -256,19 +250,6 @@ end bounding_functions section summability -variable {α : Type*} {β : Type*} - -/--Equivalence between the sigma of a fammily of finsets of `β × β` and `β × β`-/ -def sigmaEquiv (ι : α → Finset (β × β)) (HI : ∀ y : β × β , ∃! i : α, y ∈ ι i) : - (Σ s : α, ((ι s) : Set (β × β))) ≃ (β × β) where - toFun x := x.2 - invFun x := ⟨(HI x).choose, x, (HI x).choose_spec.1⟩ - left_inv x := by - ext - exact ((HI x.2).choose_spec.2 x.1 x.2.2).symm - repeat {rfl} - right_inv x := by rfl - lemma summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), 0 ≤ f y) (ι : ℕ → Finset (ℤ × ℤ)) (HI : ∀ y : ℤ × ℤ, ∃! i : ℕ, y ∈ ι i) : Summable f ↔ Summable fun n : ℕ => ∑ x in ι n, f ![x.1, x.2] := by From b402c36da9b32b4c48d1c2b47c73f3401e46e1fb Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Mon, 12 Feb 2024 10:52:14 +0000 Subject: [PATCH 030/124] build fix --- .../Finset_Decomposition.lean | 26 +++------ .../EisensteinSeries/UniformConvergence.lean | 54 +++++++++---------- 2 files changed, 32 insertions(+), 48 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean index 8a798ff948ff5..db8658b932fcb 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -88,34 +88,20 @@ lemma square_zero : square 0 = {(0, 0)} := rfl theorem square_zero_card : Finset.card (square 0) = 1 := by rw [square_zero, card_singleton] -lemma Complex_abs_eq_of_mem_square (n : ℕ) (x : ℤ × ℤ) (h : x ∈ square n) : - Complex.abs x.1 = n ∨ Complex.abs x.2 = n := by - simp_rw [eq_comm (b := (n : ℝ)), ← int_cast_abs, square_mem] at * - norm_cast - subst n - simp only [Nat.cast_max, Int.coe_natAbs, max_eq_left_iff, max_eq_right_iff] at * - exact Int.le_total |x.2| |x.1| - -lemma Complex_abs_square_left_ne (n : ℕ) (x : ℤ × ℤ) (h : x ∈ square n) - (hx : Complex.abs (x.1) ≠ n) : Complex.abs (x.2) = n := - Complex_abs_eq_of_mem_square n x h |>.resolve_left hx - lemma fun_ne_zero_cases (x : Fin 2 → ℤ) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by rw [Function.ne_iff] exact Fin.exists_fin_two -lemma square_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1 ⟩ ∈ square n) : x ≠ 0 ↔ n ≠ 0 := by +lemma square_mem_ne_zero_iff_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1⟩ ∈ square n) : + x ≠ 0 ↔ n ≠ 0 := by constructor intro h h0 - rw [h0] at hx - simp only [Nat.cast_zero, square_zero, mem_singleton, Prod.mk.injEq] at hx + simp only [h0, Nat.cast_zero, square_zero, mem_singleton, Prod.mk.injEq] at hx rw [fun_ne_zero_cases, hx.1, hx.2] at h simp only [ne_eq, not_true_eq_false, or_self] at * intro hn h - have hx0 : x 0 = 0 := by - simp only [h, Pi.zero_apply] - have hx1 : x 1 = 0 := by - simp only [h, Pi.zero_apply] - rw [hx0, hx1] at hx + have hxx : x 0 = 0 ∧ x 1 = 0 := by + simp only [h, Pi.zero_apply, and_self] + rw [hxx.1, hxx.2] at hx simp only [square_mem, Int.natAbs_zero, max_self] at hx exact hn (id hx.symm) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 21dee3ea99604..3b7a93ad85771 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -157,7 +157,7 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : have h1 : ((n : ℝ) : ℂ)^k ≠ 0 := by rw [pow_ne_zero_iff hk] norm_cast - rw [EisensteinSeries.square_ne_zero n x (by rw [square_mem])] at hx + rw [EisensteinSeries.square_mem_ne_zero_iff_ne_zero n x (by rw [square_mem])] at hx exact hx have hc : Complex.abs ((n : ℝ)^k : ℂ) = n^k := by simp only [Nat.cast_max, map_pow, abs_ofReal, ge_iff_le, abs_nonneg, le_max_iff, @@ -204,7 +204,7 @@ theorem eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → simp only [Int.cast_zero, coe_eq_fst, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, CharP.cast_eq_zero, map_mul, abs_ofReal, mul_zero, le_refl] · have hx2 : x ≠ 0 := by - rw [EisensteinSeries.square_ne_zero n x hx] + rw [EisensteinSeries.square_mem_ne_zero_iff_ne_zero n x hx] exact hn simp only [square_mem] at hx rw [inv_le_inv] @@ -253,7 +253,7 @@ section summability lemma summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), 0 ≤ f y) (ι : ℕ → Finset (ℤ × ℤ)) (HI : ∀ y : ℤ × ℤ, ∃! i : ℕ, y ∈ ι i) : Summable f ↔ Summable fun n : ℕ => ∑ x in ι n, f ![x.1, x.2] := by - let h2 := Equiv.trans (sigmaEquiv ι HI) (piFinTwoEquiv fun _ => ℤ).symm + let h2 := Equiv.trans (Equiv.sigmaEquiv ι HI) (piFinTwoEquiv fun _ => ℤ).symm have h22 : ∀ y : Σ s : ℕ, (ι s), 0 ≤ (f ∘ h2) y := by intro y apply h @@ -267,7 +267,7 @@ lemma summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), · intro H constructor · intro x - simp only [Finset.coe_sort_coe, Equiv.coe_trans, Function.comp_apply,sigmaEquiv] + simp only [Finset.coe_sort_coe, Equiv.coe_trans, Function.comp_apply, Equiv.sigmaEquiv] convert (Finset.summable (ι x) (f ∘ (piFinTwoEquiv fun _ => ℤ).symm)) · convert H rw [← Finset.tsum_subtype] @@ -276,15 +276,14 @@ lemma summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by have hk : 1 < (k - 1 : ℝ) := by - have : 1 < (k - 1 : ℤ) := by linarith norm_cast at * - have riesum := Real.summable_nat_rpow_inv.2 hk + linarith have nze : (8 / (r z) ^ k : ℝ) ≠ 0 := by apply div_ne_zero simp only [ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true] apply zpow_ne_zero k (ne_of_gt (r_pos z)) rw [← (summable_mul_left_iff nze).symm] - convert riesum + convert (Real.summable_nat_rpow_inv.2 hk) norm_cast lemma summable_over_square (k : ℤ) (z : ℍ) (h : 3 ≤ k): @@ -306,22 +305,22 @@ lemma summable_over_square (k : ℤ) (z : ℍ) (h : 3 ≤ k): lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by rw [summable_lemma _ _ (fun (n : ℕ) => square n) squares_cover_all] - have : ∀ n : ℕ, ∑ v in square n, (1 / (r z) ^ k) * ((max v.1.natAbs v.2.natAbs: ℝ) ^ k)⁻¹ = - ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ)^k)⁻¹ := by - intro n - apply Finset.sum_congr rfl - intro x hx - simp only [square_mem] at hx - congr - norm_cast - apply Summable.congr (summable_over_square k z h) - intro b - apply (this b).symm - intro y - apply mul_nonneg - simp only [one_div, inv_nonneg] - apply zpow_nonneg (r_pos z).le - simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] + · have : ∀ n : ℕ, ∑ v in square n, (1 / (r z) ^ k) * ((max v.1.natAbs v.2.natAbs: ℝ) ^ k)⁻¹ = + ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ)^k)⁻¹ := by + intro n + apply Finset.sum_congr rfl + intro x hx + simp only [square_mem] at hx + congr + norm_cast + apply Summable.congr (summable_over_square k z h) + intro b + apply (this b).symm + · intro y + apply mul_nonneg + simp only [one_div, inv_nonneg] + apply zpow_nonneg (r_pos z).le + simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] end summability @@ -333,7 +332,7 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : lift k to ℕ using hk0 rw [tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] simp only [top_eq_univ, subset_univ, eisensteinSeries, forall_true_left] - intros K hK + intro K hK obtain ⟨A, B, hB, HABK⟩:= compact_in_some_slice K hK have hu : Summable fun x : (gammaSet N a ) => (1/(r ⟨⟨A, B⟩, hB⟩) ^ k) * ((max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by @@ -349,14 +348,13 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ge_iff_le] at * apply le_trans (this sq) rw [mul_comm] - apply mul_le_mul + apply mul_le_mul _ (by rfl) + repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg, + inv_nonneg, pow_nonneg (r_pos _).le]} rw [inv_le_inv] apply pow_le_pow_left (r_pos _).le rw [abs_of_pos (r_pos _)] · exact r_lower_bound_on_slice A B hB ⟨x, HABK hx⟩ · apply pow_pos (abs_pos.mpr (ne_of_gt (r_pos x))) · apply pow_pos (r_pos _) - · rfl - repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg, - inv_nonneg, pow_nonneg (r_pos _).le]} · simp only [top_eq_univ, isOpen_univ] From a3c4d2199dee4b8ada869b4f5083190874048338 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Mon, 12 Feb 2024 13:32:17 +0000 Subject: [PATCH 031/124] fix imports --- .../EisensteinSeries/Finset_Decomposition.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean index db8658b932fcb..5f1c12ae1a37c 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean @@ -3,8 +3,11 @@ Copyright (c) 2024 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ -import Mathlib.Data.Complex.Abs -import Mathlib.Data.IsROrC.Basic + +import Mathlib.Data.Int.Interval +import Mathlib.Tactic.Ring.RingNF +import Mathlib.Tactic.Linarith +import Mathlib.Tactic.NormNum.Ineq /-! # Decomposing `ℤ × ℤ` into squares @@ -12,9 +15,7 @@ We partition `ℤ × ℤ` into squares of the form `Icc (-n) n × Icc (-n) n` fo for bounding Eisenstein series. -/ -open Complex - -open scoped BigOperators NNReal Classical Filter Matrix +open scoped BigOperators Classical noncomputable section From 90b37fe3c9cb9f0d9f271e9ff53d5d951b602f57 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Mon, 12 Feb 2024 14:41:30 +0000 Subject: [PATCH 032/124] more cleanup --- .../EisensteinSeries/UniformConvergence.lean | 42 ++++++++++--------- 1 file changed, 23 insertions(+), 19 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 3b7a93ad85771..65f7759e6496b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -31,21 +31,21 @@ namespace EisensteinSeries section bounding_functions /-- Auxilary function used for bounding Eisentein series-/ -def lowerBound1 (z : ℍ) : ℝ := ((z.im ^ (2 : ℕ)) / (z.re ^ (2 : ℕ) + z.im ^ (2 : ℕ))) +def r1 (z : ℍ) : ℝ := ((z.im ^ (2 : ℕ)) / (z.re ^ (2 : ℕ) + z.im ^ (2 : ℕ))) -lemma lowerBound1' (z : ℍ) : lowerBound1 z = 1/((z.re / z.im) ^ (2 : ℕ) + 1) := by - field_simp [lowerBound1, im_pos z] +lemma r1' (z : ℍ) : r1 z = 1/((z.re / z.im) ^ (2 : ℕ) + 1) := by + field_simp [r1, im_pos z] -theorem lowerBound1_pos (z : ℍ) : 0 < lowerBound1 z := by +theorem r1_pos (z : ℍ) : 0 < r1 z := by have H2 : 0 < (z.re ^ (2 : ℕ) + z.im ^ (2 : ℕ)) := by apply_rules [pow_pos, add_pos_of_nonneg_of_pos, pow_two_nonneg, z.2] exact div_pos (pow_pos z.im_pos 2) H2 /-- This function is used to give an upper bound on Eisenstein series-/ -def r (z : ℍ) : ℝ := min (z.im) (Real.sqrt (lowerBound1 z)) +def r (z : ℍ) : ℝ := min (z.im) (Real.sqrt (r1 z)) theorem r_pos (z : ℍ) : 0 < r z := by - simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, lowerBound1_pos, and_self] + simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, r1_pos, and_self] lemma r_mul_pos_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : 0 < (Complex.abs ((r z : ℂ) ^ (k : ℤ) * (n : ℂ)^ (k : ℤ))) := by @@ -55,7 +55,7 @@ lemma r_mul_pos_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : intro _ linarith -theorem auxlb (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : +theorem r1_bound (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : (z.im ^ 2 ) / (z.re ^ 2 + z.im ^ 2) ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by have H1 : (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 = δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2 := by ring @@ -87,10 +87,10 @@ theorem auxbound1 (z : ℍ) (δ ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Complex.a theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs, min_le_iff] - have H1 : Real.sqrt (lowerBound1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε ) + + have H1 : Real.sqrt (r1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε ) + δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by - rw [lowerBound1, Real.sqrt_le_sqrt_iff, ← pow_two, ← pow_two] - apply auxlb z δ ε hε + rw [r1, Real.sqrt_le_sqrt_iff, ← pow_two, ← pow_two] + apply r1_bound z δ ε hε nlinarith right simp only [ne_eq, coe_re, coe_im, normSq_apply, AbsoluteValue.coe_mk, MulHom.coe_mk, add_re, @@ -98,6 +98,7 @@ theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.a int_cast_im] at * exact H1 +--any suggestions as to where to put the lemmas below? lemma one_le_sq_div_abs_sq (a : ℝ) (ha : a ≠ 0) : 1 ≤ a^2 / |a|^2 := by rw [_root_.sq_abs, le_div_iff'] simp only [mul_one, le_refl] @@ -127,8 +128,9 @@ lemma ne_zero_if_max' (x : Fin 2 → ℤ) (hx : x ≠ 0) rw [fun_ne_zero_cases, h1, h0] at hx simp only [ne_eq, not_true_eq_false, or_self] at * -lemma sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs))^2 ∨ - (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs))^2 := by +lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : + (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs))^2 ∨ + (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs))^2 := by cases' (max_choice (x 0).natAbs (x 1).natAbs) with H1 H2 · left rw [H1] @@ -149,6 +151,8 @@ lemma sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max ( norm_cast rw [int_abs_eq_complex_abs] +/-This should work for complex `k` (with a condition on `k.re`) but last I checked we were missing +some lemmas about this -/ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs)^k ≤ Complex.abs (((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ)) ^ k) := by @@ -168,7 +172,7 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : (((x 0 : ℝ) / (n : ℝ) ) * (z : ℂ) + (x 1 : ℝ) /(n : ℝ)) ^ (k : ℤ) * ((n : ℝ)^ (k : ℤ)) := by simp only [Nat.cast_max] at h1 field_simp - cases' (sq_ge_one x hx) with H1 H2 + cases' (div_max_sq_ge_one x hx) with H1 H2 · norm_cast at * rw [h11] simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * @@ -229,14 +233,14 @@ lemma r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice A r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by have hz := z.2 simp only [slice_mem, abs_ofReal, ge_iff_le] at hz - rw [r, r] + simp_rw [r] apply min_le_min · dsimp only convert hz.2 have := abs_eq_self.mpr (UpperHalfPlane.im_pos z.1).le convert this.symm rw [Real.sqrt_le_sqrt_iff] - simp [lowerBound1'] + simp only [r1', div_pow, one_div] rw [inv_le_inv, add_le_add_iff_right] apply div_le_div (sq_nonneg _) · simpa [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 @@ -244,7 +248,7 @@ lemma r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice A · simpa [even_two.pow_abs] using pow_le_pow_left h.le hz.2 2 · positivity · positivity - · apply (lowerBound1_pos z).le + · apply (r1_pos z).le end bounding_functions @@ -318,9 +322,9 @@ lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : apply (this b).symm · intro y apply mul_nonneg - simp only [one_div, inv_nonneg] - apply zpow_nonneg (r_pos z).le - simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] + · simp only [one_div, inv_nonneg] + apply zpow_nonneg (r_pos z).le + · simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] end summability From 83ee7dcaf74541f9375a80da9e60ecebaadc275a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Feb 2024 16:08:08 +0000 Subject: [PATCH 033/124] rename file --- Mathlib.lean | 2 +- .../{Finset_Decomposition.lean => FinsetDecomposition.lean} | 0 .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 3 files changed, 2 insertions(+), 2 deletions(-) rename Mathlib/NumberTheory/ModularForms/EisensteinSeries/{Finset_Decomposition.lean => FinsetDecomposition.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 3d17498409160..902fc1b4c6bc1 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2845,7 +2845,7 @@ import Mathlib.NumberTheory.Modular import Mathlib.NumberTheory.ModularForms.Basic import Mathlib.NumberTheory.ModularForms.CongruenceSubgroups import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic -import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Finset_Decomposition +import Mathlib.NumberTheory.ModularForms.EisensteinSeries.FinsetDecomposition import Mathlib.NumberTheory.ModularForms.EisensteinSeries.UniformConvergence import Mathlib.NumberTheory.ModularForms.JacobiTheta.Manifold import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/FinsetDecomposition.lean similarity index 100% rename from Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean rename to Mathlib/NumberTheory/ModularForms/EisensteinSeries/FinsetDecomposition.lean diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 65f7759e6496b..dc071ea2b4391 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -11,7 +11,7 @@ import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic import Mathlib.Analysis.Complex.UpperHalfPlane.Metric import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.PSeries -import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Finset_Decomposition +import Mathlib.NumberTheory.ModularForms.EisensteinSeries.FinsetDecomposition /-! # Uniform convergence of Eisenstein series From 9ba96ec6b030ce692b900c4ecbf9019134eb5a42 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 16:22:47 +0000 Subject: [PATCH 034/124] rev fixes wip --- .../Complex/UpperHalfPlane/Basic.lean | 2 - .../Complex/UpperHalfPlane/Metric.lean | 12 ++-- Mathlib/Data/Finset/Basic.lean | 6 +- ...position.lean => FinsetDecomposition.lean} | 4 +- .../EisensteinSeries/UniformConvergence.lean | 72 +++++++++---------- 5 files changed, 42 insertions(+), 54 deletions(-) rename Mathlib/NumberTheory/ModularForms/EisensteinSeries/{Finset_Decomposition.lean => FinsetDecomposition.lean} (98%) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 5427b2cc31b28..05a0516dce3d5 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -123,8 +123,6 @@ theorem mk_coe (z : ℍ) (h : 0 < (z : ℂ).im := z.2) : mk z h = z := rfl #align upper_half_plane.mk_coe UpperHalfPlane.mk_coe -lemma coe_eq_fst (z : ℍ) : (z : ℂ) = z.1 := rfl - theorem re_add_im (z : ℍ) : (z.re + z.im * Complex.I : ℂ) = z := Complex.re_add_im z #align upper_half_plane.re_add_im UpperHalfPlane.re_add_im diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index fdb49ac066cbf..6cea2fc629c43 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -390,10 +390,10 @@ section slices def upperHalfPlaneSlice (A B : ℝ) := {z : ℍ | Complex.abs z.1.1 ≤ A ∧ Complex.abs z.1.2 ≥ B} -theorem slice_mem (A B : ℝ) (z : ℍ) : +theorem slice_mem_iff (A B : ℝ) (z : ℍ) : z ∈ upperHalfPlaneSlice A B ↔ Complex.abs z.1.1 ≤ A ∧ Complex.abs z.1.2 ≥ B := Iff.rfl -lemma compact_in_some_slice (K : Set ℍ) (hK : IsCompact K) : ∃ A B : ℝ, 0 < B ∧ +lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : ∃ A B : ℝ, 0 < B ∧ K ⊆ upperHalfPlaneSlice A B := by by_cases hne : Set.Nonempty K · have hcts : ContinuousOn (fun t => t.im) K := by @@ -404,7 +404,7 @@ lemma compact_in_some_slice (K : Set ℍ) (hK : IsCompact K) : ∃ A B : ℝ, 0 obtain ⟨r, _, hr2⟩ := Bornology.IsBounded.subset_closedBall_lt hK.isBounded 0 t refine' ⟨Real.sinh (r) + Complex.abs ((UpperHalfPlane.center t r)), b.im, b.2, _⟩ intro z hz - simp only [I_im, slice_mem, abs_ofReal, ge_iff_le] at * + simp only [I_im, slice_mem_iff, abs_ofReal, ge_iff_le] at * constructor have hr3 := hr2 hz simp only [Metric.mem_closedBall] at hr3 @@ -423,11 +423,7 @@ lemma compact_in_some_slice (K : Set ℍ) (hK : IsCompact K) : ∃ A B : ℝ, 0 convert hbz rw [UpperHalfPlane.im] apply abs_eq_self.mpr z.2.le - · rw [not_nonempty_iff_eq_empty] at hne - rw [hne] - simp only [empty_subset, and_true, exists_const] - use 1 - linarith + · exact ⟨1, 1, Real.zero_lt_one, by simp [not_nonempty_iff_eq_empty.1 hne]⟩ end slices diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 2eb3dd1dd454e..15b0bd19b6e03 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -3962,10 +3962,10 @@ def piFinsetUnion {ι} [DecidableEq ι] (α : ι → Type*) {s t : Finset ι} (h let e := Equiv.Finset.union s t h sumPiEquivProdPi (fun b ↦ α (e b)) |>.symm.trans (.piCongrLeft (fun i : ↥(s ∪ t) ↦ α i) e) -/--Equivalence between the sigma of a fammily of finsets of `β × β` and `β × β`-/ +/--Equivalence between the sigma of a family of finsets of `β × γ` and `β × γ`-/ noncomputable -def sigmaEquiv {α : Type*} {β : Type*} (ι : α → Finset (β × β)) - (HI : ∀ y : β × β , ∃! i : α, y ∈ ι i) : (Σ s : α, ((ι s) : Set (β × β))) ≃ (β × β) where +def sigmaEquiv {α β γ : Type*} (ι : α → Finset (β × γ)) + (HI : ∀ y : β × γ , ∃! i : α, y ∈ ι i) : (Σ s : α, ((ι s) : Set (β × γ))) ≃ (β × γ) where toFun x := x.2 invFun x := ⟨(HI x).choose, x, (HI x).choose_spec.1⟩ left_inv x := by diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/FinsetDecomposition.lean similarity index 98% rename from Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean rename to Mathlib/NumberTheory/ModularForms/EisensteinSeries/FinsetDecomposition.lean index 5f1c12ae1a37c..f6fe1d8d40d9a 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Finset_Decomposition.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/FinsetDecomposition.lean @@ -15,12 +15,10 @@ We partition `ℤ × ℤ` into squares of the form `Icc (-n) n × Icc (-n) n` fo for bounding Eisenstein series. -/ -open scoped BigOperators Classical +open scoped BigOperators noncomputable section -namespace EisensteinSeries - open Finset /-- For `m : ℤ` this is the finset of `ℤ × ℤ` of elements such that the maximum of the diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 65f7759e6496b..06400d20498d8 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -11,7 +11,7 @@ import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic import Mathlib.Analysis.Complex.UpperHalfPlane.Metric import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.PSeries -import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Finset_Decomposition +import Mathlib.NumberTheory.ModularForms.EisensteinSeries.FinsetDecomposition /-! # Uniform convergence of Eisenstein series @@ -47,15 +47,7 @@ def r (z : ℍ) : ℝ := min (z.im) (Real.sqrt (r1 z)) theorem r_pos (z : ℍ) : 0 < r z := by simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, r1_pos, and_self] -lemma r_mul_pos_pos (k : ℕ) (z : ℍ) (n : ℝ) (hn : 0 < n) : - 0 < (Complex.abs ((r z : ℂ) ^ (k : ℤ) * (n : ℂ)^ (k : ℤ))) := by - norm_cast - apply _root_.abs_pos.mpr (mul_ne_zero (pow_ne_zero k (ne_of_gt (r_pos z))) ?_) - simp only [Nat.cast_pow, ne_eq, pow_eq_zero_iff', Nat.cast_eq_zero, not_and, not_not] at * - intro _ - linarith - -theorem r1_bound (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : +theorem r1_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : (z.im ^ 2 ) / (z.re ^ 2 + z.im ^ 2) ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by have H1 : (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 = δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2 := by ring @@ -68,7 +60,7 @@ theorem r1_bound (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : apply pow_two_nonneg · apply_rules [add_pos_of_nonneg_of_pos, pow_two_nonneg, (pow_pos z.im_pos 2)] -theorem auxbound1 (z : ℍ) (δ ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by +theorem auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs] have H1 : (z : ℂ).im ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + (δ * z : ℂ).im * (δ * z : ℂ).im) := by @@ -85,12 +77,12 @@ theorem auxbound1 (z : ℍ) (δ ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Complex.a add_zero, normSq_apply, add_re, mul_re, sub_zero, add_im] at * exact H1 -theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by +theorem auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs, min_le_iff] have H1 : Real.sqrt (r1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε ) + δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by rw [r1, Real.sqrt_le_sqrt_iff, ← pow_two, ← pow_two] - apply r1_bound z δ ε hε + apply r1_bound z δ hε nlinarith right simp only [ne_eq, coe_re, coe_im, normSq_apply, AbsoluteValue.coe_mk, MulHom.coe_mk, add_re, @@ -99,11 +91,14 @@ theorem auxbound2 (z : ℍ) (δ ε : ℝ) (hε : 1 ≤ ε^2) : r z ≤ Complex.a exact H1 --any suggestions as to where to put the lemmas below? + lemma one_le_sq_div_abs_sq (a : ℝ) (ha : a ≠ 0) : 1 ≤ a^2 / |a|^2 := by rw [_root_.sq_abs, le_div_iff'] simp only [mul_one, le_refl] exact (sq_pos_iff a).mpr ha +--#find_home! one_le_sq_div_abs_sq + lemma int_abs_eq_complex_abs (a : ℤ) : Complex.abs a = a.natAbs := by have : Complex.abs a = Complex.abs (a : ℝ) := by rfl rw [this, abs_ofReal] @@ -111,7 +106,7 @@ lemma int_abs_eq_complex_abs (a : ℤ) : Complex.abs a = a.natAbs := by rw [Int.abs_eq_natAbs a] rfl -lemma ne_zero_if_max (x : Fin 2 → ℤ) (hx : x ≠ 0) +lemma ne_zero_if_max {x : Fin 2 → ℤ} (hx : x ≠ 0) (h : (max (x 0).natAbs (x 1).natAbs) = (x 0).natAbs) : (x 0) ≠ 0 := by intro h0 rw [h0] at h @@ -120,9 +115,9 @@ lemma ne_zero_if_max (x : Fin 2 → ℤ) (hx : x ≠ 0) rw [h0, h] at this simp only [this, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at hx -lemma ne_zero_if_max' (x : Fin 2 → ℤ) (hx : x ≠ 0) +lemma ne_zero_if_max' {x : Fin 2 → ℤ} (hx : x ≠ 0) (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by - apply ne_zero_if_max ![x 1, x 0] ?_ (by simpa using h) + apply ne_zero_if_max (x :=![x 1, x 0]) ?_ (by simpa using h) simp only [ne_eq, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_true, not_and] intro h1 h0 rw [fun_ne_zero_cases, h1, h0] at hx @@ -135,7 +130,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : · left rw [H1] have : (x 0 : ℝ) ≠ 0 := by - simpa using (ne_zero_if_max x hx H1) + simpa using (ne_zero_if_max hx H1) have h1 := one_le_sq_div_abs_sq (x 0 : ℝ) this simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 @@ -144,7 +139,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : · right rw [H2] have : (x 1 : ℝ) ≠ 0 := by - simpa using (ne_zero_if_max' x hx H2) + simpa using (ne_zero_if_max' hx H2) have h1 := one_le_sq_div_abs_sq (x 1 : ℝ) this simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 @@ -161,7 +156,7 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : have h1 : ((n : ℝ) : ℂ)^k ≠ 0 := by rw [pow_ne_zero_iff hk] norm_cast - rw [EisensteinSeries.square_mem_ne_zero_iff_ne_zero n x (by rw [square_mem])] at hx + rw [square_mem_ne_zero_iff_ne_zero n x (by rw [square_mem])] at hx exact hx have hc : Complex.abs ((n : ℝ)^k : ℂ) = n^k := by simp only [Nat.cast_max, map_pow, abs_ofReal, ge_iff_le, abs_nonneg, le_max_iff, @@ -180,7 +175,7 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg] apply pow_nonneg (Complex.abs.nonneg _) k - apply pow_le_pow_left (r_pos _).le (auxbound1 z (x 0 / n) (x 1 / n) H1) + apply pow_le_pow_left (r_pos _).le (auxbound1 z (x 1 / n) H1) · norm_cast at * rw [h11] simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * @@ -188,7 +183,7 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg] apply pow_nonneg (Complex.abs.nonneg _) k - apply pow_le_pow_left (r_pos _).le (auxbound2 z (x 0 / n) (x 1 / n) H2) + apply pow_le_pow_left (r_pos _).le (auxbound2 z (x 0 / n) H2) · simp only [ne_eq, not_not] at hk simp only [hk, pow_zero, Nat.cast_max, mul_one, map_one, le_refl] @@ -200,15 +195,15 @@ theorem eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → simp only [CharP.cast_eq_zero, square_zero, Finset.mem_singleton, Prod.mk.injEq] at hx by_cases hk : k = 0 rw [hk] at * - simp only [coe_eq_fst, pow_zero, map_one, inv_one, mul_one, le_refl] + simp only [ pow_zero, map_one, inv_one, mul_one, le_refl] rw [hx.1, hx.2] have h1 : (0 : ℝ) ^ (k : ℕ) = 0 := by simp only [pow_eq_zero_iff', ne_eq, true_and] exact hk - simp only [Int.cast_zero, coe_eq_fst, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, + simp only [Int.cast_zero, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, CharP.cast_eq_zero, map_mul, abs_ofReal, mul_zero, le_refl] · have hx2 : x ≠ 0 := by - rw [EisensteinSeries.square_mem_ne_zero_iff_ne_zero n x hx] + rw [square_mem_ne_zero_iff_ne_zero n x hx] exact hn simp only [square_mem] at hx rw [inv_le_inv] @@ -225,14 +220,15 @@ theorem eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → exact List.ofFn_inj.mp rfl rw [this, hg,h1 ] at hx2 simp only [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at * - apply r_mul_pos_pos + have := r_pos z + apply Complex.abs.pos norm_cast - exact Nat.pos_of_ne_zero hn + positivity -lemma r_lower_bound_on_slice (A B : ℝ) (h : 0 < B) (z : upperHalfPlaneSlice A B) : +lemma r_lower_bound_on_slice {A B : ℝ} (h : 0 < B) (z : upperHalfPlaneSlice A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by have hz := z.2 - simp only [slice_mem, abs_ofReal, ge_iff_le] at hz + simp only [slice_mem_iff, abs_ofReal, ge_iff_le] at hz simp_rw [r] apply min_le_min · dsimp only @@ -277,7 +273,7 @@ lemma summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), rw [← Finset.tsum_subtype] rfl -lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : +lemma summable_r_pow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by have hk : 1 < (k - 1 : ℝ) := by norm_cast at * @@ -290,10 +286,10 @@ lemma summable_r_pow (k : ℤ) (z : ℍ) (h : 3 ≤ k) : convert (Real.summable_nat_rpow_inv.2 hk) norm_cast -lemma summable_over_square (k : ℤ) (z : ℍ) (h : 3 ≤ k): +lemma summable_over_square {k : ℤ} (z : ℍ) (h : 3 ≤ k): Summable (fun n : ℕ => ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by simp only [one_div, Finset.sum_const, nsmul_eq_mul] - apply Summable.congr (summable_r_pow k z h) + apply Summable.congr (summable_r_pow z h) intro b by_cases b0 : b = 0 · rw [b0] @@ -306,7 +302,7 @@ lemma summable_over_square (k : ℤ) (z : ℍ) (h : 3 ≤ k): simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] ring_nf -lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => +lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by rw [summable_lemma _ _ (fun (n : ℕ) => square n) squares_cover_all] · have : ∀ n : ℕ, ∑ v in square n, (1 / (r z) ^ k) * ((max v.1.natAbs v.2.natAbs: ℝ) ^ k)⁻¹ = @@ -317,7 +313,7 @@ lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : simp only [square_mem] at hx congr norm_cast - apply Summable.congr (summable_over_square k z h) + apply Summable.congr (summable_over_square z h) intro b apply (this b).symm · intro y @@ -328,7 +324,7 @@ lemma summable_upper_bound (k : ℤ) (h : 3 ≤ k) (z : ℍ) : Summable fun (x : end summability -theorem eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : ℕ) +theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by @@ -337,17 +333,17 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : rw [tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] simp only [top_eq_univ, subset_univ, eisensteinSeries, forall_true_left] intro K hK - obtain ⟨A, B, hB, HABK⟩:= compact_in_some_slice K hK + obtain ⟨A, B, hB, HABK⟩:= subset_slice_of_isCompact hK have hu : Summable fun x : (gammaSet N a ) => (1/(r ⟨⟨A, B⟩, hB⟩) ^ k) * ((max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by - apply (Summable.subtype (summable_upper_bound k hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a )).congr + apply (Summable.subtype (summable_upper_bound hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a )).congr intro v simp only [zpow_coe_nat, one_div, Function.comp_apply] apply tendstoUniformlyOn_tsum hu intro v x hx have sq := square_mem (max (v.1 0).natAbs (v.1 1).natAbs ) ⟨(v.1 0), v.1 1⟩ have := eis_is_bounded_on_square k x (max (v.1 0).natAbs (v.1 1).natAbs ) v - simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_coe_nat, one_div, coe_eq_fst, map_pow, + simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_coe_nat, one_div, map_pow, map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, ge_iff_le] at * apply le_trans (this sq) @@ -358,7 +354,7 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn (k : ℤ) (hk : 3 ≤ k) (N : rw [inv_le_inv] apply pow_le_pow_left (r_pos _).le rw [abs_of_pos (r_pos _)] - · exact r_lower_bound_on_slice A B hB ⟨x, HABK hx⟩ + · exact r_lower_bound_on_slice hB ⟨x, HABK hx⟩ · apply pow_pos (abs_pos.mpr (ne_of_gt (r_pos x))) · apply pow_pos (r_pos _) · simp only [top_eq_univ, isOpen_univ] From c4bfb5b471bd05bff59c8644e14a3890edb0e2d7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Feb 2024 17:33:10 +0000 Subject: [PATCH 035/124] Box theory --- Mathlib.lean | 4 +- Mathlib/Data/DFinsupp/Interval.lean | 2 +- Mathlib/Data/Fin/Interval.lean | 2 +- Mathlib/Data/Finset/Interval.lean | 2 +- .../Basic.lean} | 0 Mathlib/Data/Finset/LocallyFinite/Box.lean | 89 +++++++++++++++ Mathlib/Data/Finsupp/Interval.lean | 2 +- Mathlib/Data/Int/Interval.lean | 2 +- Mathlib/Data/Multiset/Interval.lean | 2 +- Mathlib/Data/Multiset/LocallyFinite.lean | 2 +- Mathlib/Data/Nat/Interval.lean | 2 +- Mathlib/Data/Pi/Interval.lean | 2 +- .../EisensteinSeries/FinsetDecomposition.lean | 106 ------------------ .../EisensteinSeries/UniformConvergence.lean | 18 ++- Mathlib/Order/Birkhoff.lean | 2 +- Mathlib/Order/Disjointed.lean | 8 +- 16 files changed, 122 insertions(+), 123 deletions(-) rename Mathlib/Data/Finset/{LocallyFinite.lean => LocallyFinite/Basic.lean} (100%) create mode 100644 Mathlib/Data/Finset/LocallyFinite/Box.lean delete mode 100644 Mathlib/NumberTheory/ModularForms/EisensteinSeries/FinsetDecomposition.lean diff --git a/Mathlib.lean b/Mathlib.lean index 902fc1b4c6bc1..d41fc7f493f33 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1561,7 +1561,8 @@ import Mathlib.Data.Finset.Grade import Mathlib.Data.Finset.Image import Mathlib.Data.Finset.Interval import Mathlib.Data.Finset.Lattice -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic +import Mathlib.Data.Finset.LocallyFinite.Box import Mathlib.Data.Finset.MulAntidiagonal import Mathlib.Data.Finset.NAry import Mathlib.Data.Finset.NatAntidiagonal @@ -2845,7 +2846,6 @@ import Mathlib.NumberTheory.Modular import Mathlib.NumberTheory.ModularForms.Basic import Mathlib.NumberTheory.ModularForms.CongruenceSubgroups import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic -import Mathlib.NumberTheory.ModularForms.EisensteinSeries.FinsetDecomposition import Mathlib.NumberTheory.ModularForms.EisensteinSeries.UniformConvergence import Mathlib.NumberTheory.ModularForms.JacobiTheta.Manifold import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable diff --git a/Mathlib/Data/DFinsupp/Interval.lean b/Mathlib/Data/DFinsupp/Interval.lean index e889b3ca8716c..7f01d1776f341 100644 --- a/Mathlib/Data/DFinsupp/Interval.lean +++ b/Mathlib/Data/DFinsupp/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.Finset.Pointwise import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.DFinsupp.Order diff --git a/Mathlib/Data/Fin/Interval.lean b/Mathlib/Data/Fin/Interval.lean index 4759ca985de07..b8c85526b18f1 100644 --- a/Mathlib/Data/Fin/Interval.lean +++ b/Mathlib/Data/Fin/Interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Nat.Interval -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.fin.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Finset/Interval.lean b/Mathlib/Data/Finset/Interval.lean index 9c787cd92a468..6c654fea23e23 100644 --- a/Mathlib/Data/Finset/Interval.lean +++ b/Mathlib/Data/Finset/Interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Finset.Grade -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.finset.interval from "leanprover-community/mathlib"@"98e83c3d541c77cdb7da20d79611a780ff8e7d90" diff --git a/Mathlib/Data/Finset/LocallyFinite.lean b/Mathlib/Data/Finset/LocallyFinite/Basic.lean similarity index 100% rename from Mathlib/Data/Finset/LocallyFinite.lean rename to Mathlib/Data/Finset/LocallyFinite/Basic.lean diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean new file mode 100644 index 0000000000000..69835f2d631af --- /dev/null +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -0,0 +1,89 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Data.Int.Cast.Prod +import Mathlib.Data.Int.Interval +import Mathlib.Order.Disjointed +import Mathlib.Tactic.Ring.RingNF +import Mathlib.Tactic.Linarith +import Mathlib.Tactic.NormNum.Ineq +import Mathlib.Algebra.Order.Group.Prod +-- TODO: Why is the `OrderedRing` instance not under `Algebra.Order`? +import Mathlib.Algebra.Ring.Prod + +/-! +# Decomposing a locally finite ordered ring into boxes + +This file proves that any locally finite ordered ring can be decomposed into "boxes", namely +differences of consecutive intervals. +-/ + +/-! ### General locally finite ordered ring -/ + +namespace Finset +variable {α : Type*} [OrderedRing α] [LocallyFiniteOrder α] [DecidableEq α] {n : ℕ} + +/-- Hollow box centered at `0 : α` going from `-n` to `n`. -/ +def box : ℕ → Finset α := disjointed fun n ↦ Icc (-n : α) n + +private lemma Icc_neg_mono : Monotone fun n : ℕ ↦ Icc (-n : α) n := by + refine fun m n hmn ↦ by apply Icc_subset_Icc <;> simpa using Nat.mono_cast hmn + +@[simp] lemma box_zero : (box 0 : Finset α) = {0} := by simp [box] + +lemma box_succ_eq_sdiff (n : ℕ) : + box (n + 1) = Icc (-n.succ : α) n.succ \ Icc (-n) n := Icc_neg_mono.disjointed_succ _ + +lemma disjoint_box_succ_prod (n : ℕ) : Disjoint (box (n + 1)) (Icc (-n : α) n) := by + rw [box_succ_eq_sdiff]; exact disjoint_sdiff_self_left + +@[simp] lemma box_succ_union_prod (n : ℕ) : + box (n + 1) ∪ Icc (-n : α) n = Icc (-n.succ : α) n.succ := Icc_neg_mono.disjointed_succ_sup _ + +@[simp] lemma box_succ_disjUnion (n : ℕ) : + (box (n + 1)).disjUnion (Icc (-n : α) n) (disjoint_box_succ_prod _) = + Icc (-n.succ : α) n.succ := by rw [disjUnion_eq_union, box_succ_union_prod] + +end Finset + +open Finset + +/-! ### Product of locally finite ordered rings -/ + +namespace Prod +variable {α β : Type*} [OrderedRing α] [OrderedRing β] [LocallyFiniteOrder α] [LocallyFiniteOrder β] + [DecidableEq α] [DecidableEq β] [@DecidableRel (α × β) (· ≤ ·)] + +@[simp] lemma card_box_succ (n : ℕ) : + (box (n + 1) : Finset (α × β)).card = + (Icc (-n.succ : α) n.succ).card * (Icc (-n.succ : β) n.succ).card - + (Icc (-n : α) n).card * (Icc (-n : β) n).card := by + rw [box_succ_eq_sdiff, card_sdiff (Icc_neg_mono n.le_succ), Prod.card_Icc, Prod.card_Icc]; rfl + +end Prod + +/-! ### `ℤ × ℤ` -/ + +namespace Int +variable {n : ℕ} {x : ℤ × ℤ} + +attribute [norm_cast] toNat_ofNat + +lemma card_box : ∀ {n}, n ≠ 0 → (box n : Finset (ℤ × ℤ)).card = 8 * n + | n + 1, _ => by + simp_rw [Prod.card_box_succ, card_Icc, sub_neg_eq_add] + norm_cast + refine tsub_eq_of_eq_add ?_ + zify + ring + +@[simp] lemma mem_box : ∀ {n}, x ∈ box n ↔ max x.1.natAbs x.2.natAbs = n + | 0 => by simp [Prod.ext_iff] + | n + 1 => by simp [box_succ_eq_sdiff, Prod.le_def]; omega + +lemma existsUnique_mem_box (x : ℤ × ℤ) : ∃! n : ℕ, x ∈ box n := by + use max x.1.natAbs x.2.natAbs; simp only [mem_box, and_self_iff, forall_eq'] + +end Int diff --git a/Mathlib/Data/Finsupp/Interval.lean b/Mathlib/Data/Finsupp/Interval.lean index 95383857acad4..6c159b6dce65d 100644 --- a/Mathlib/Data/Finsupp/Interval.lean +++ b/Mathlib/Data/Finsupp/Interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Finset.Finsupp -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.Finsupp.Order #align_import data.finsupp.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index e0f08d5e2b71d..99e021b2ca256 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Order.LocallyFinite -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.int.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Multiset/Interval.lean b/Mathlib/Data/Multiset/Interval.lean index 9d8fafeab1c0f..fd010897f5423 100644 --- a/Mathlib/Data/Multiset/Interval.lean +++ b/Mathlib/Data/Multiset/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.DFinsupp.Interval import Mathlib.Data.DFinsupp.Multiset import Mathlib.Data.Nat.Interval diff --git a/Mathlib/Data/Multiset/LocallyFinite.lean b/Mathlib/Data/Multiset/LocallyFinite.lean index 137df9cdf554e..8bb625e14df1d 100644 --- a/Mathlib/Data/Multiset/LocallyFinite.lean +++ b/Mathlib/Data/Multiset/LocallyFinite.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.multiset.locally_finite from "leanprover-community/mathlib"@"59694bd07f0a39c5beccba34bd9f413a160782bf" diff --git a/Mathlib/Data/Nat/Interval.lean b/Mathlib/Data/Nat/Interval.lean index ec44e6b45a493..ca46a4e1a1a91 100644 --- a/Mathlib/Data/Nat/Interval.lean +++ b/Mathlib/Data/Nat/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.nat.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Pi/Interval.lean b/Mathlib/Data/Pi/Interval.lean index c63f64d6b9887..2334d61926de6 100644 --- a/Mathlib/Data/Pi/Interval.lean +++ b/Mathlib/Data/Pi/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.Fintype.BigOperators #align_import data.pi.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/FinsetDecomposition.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/FinsetDecomposition.lean deleted file mode 100644 index f6fe1d8d40d9a..0000000000000 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/FinsetDecomposition.lean +++ /dev/null @@ -1,106 +0,0 @@ -/- -Copyright (c) 2024 Chris Birkbeck. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Birkbeck --/ - -import Mathlib.Data.Int.Interval -import Mathlib.Tactic.Ring.RingNF -import Mathlib.Tactic.Linarith -import Mathlib.Tactic.NormNum.Ineq - -/-! # Decomposing `ℤ × ℤ` into squares - -We partition `ℤ × ℤ` into squares of the form `Icc (-n) n × Icc (-n) n` for `n : ℕ`. This is useful -for bounding Eisenstein series. --/ - -open scoped BigOperators - -noncomputable section - -open Finset - -/-- For `m : ℤ` this is the finset of `ℤ × ℤ` of elements such that the maximum of the -absolute values of the pair is `m` -/ -def square (m : ℤ) : Finset (ℤ × ℤ) := - ((Icc (-m) (m)) ×ˢ (Icc (-m) (m))).filter fun x => max x.1.natAbs x.2.natAbs = m - --- a re-definition in simp-normal form -lemma square_eq (m : ℤ) : - square m = ((Icc (-m) m) ×ˢ (Icc (-m) m)).filter fun x => max |x.1| |x.2| = m := by - simp [square] - -lemma mem_square_aux {m : ℤ} {i} : i ∈ Icc (-m) m ×ˢ Icc (-m) m ↔ max |i.1| |i.2| ≤ m := by - simp [abs_le] - -lemma square_disj {n : ℤ} : Disjoint (square (n+1)) (Icc (-n) n ×ˢ Icc (-n) n) := by - rw [square_eq] - simp only [Finset.disjoint_left, mem_filter, mem_square_aux] - rintro x ⟨-, hx⟩ - simp [hx] - -lemma square_union {n : ℤ} : - square (n+1) ∪ Icc (-n) n ×ˢ Icc (-n) n = Icc (-(n+1)) (n+1) ×ˢ Icc (-(n+1)) (n+1) := by - ext x - rw [mem_union, square_eq, mem_filter, mem_square_aux, mem_square_aux, - and_iff_right_of_imp le_of_eq, Int.le_add_one_iff, or_comm] - -lemma square_disjunion (n : ℤ) : - (square (n+1)).disjUnion (Icc (-n) n ×ˢ Icc (-n) n) square_disj = - Icc (-(n+1)) (n+1) ×ˢ Icc (-(n+1)) (n+1) := by rw [disjUnion_eq_union, square_union] - -theorem square_size (n : ℕ) : Finset.card (square (n + 1)) = 8 * (n + 1) := by - have : (((square (n+1)).disjUnion (Icc (-n : ℤ) n ×ˢ Icc (-n : ℤ) n) square_disj).card : ℤ) = - (Icc (-(n+1 : ℤ)) (n+1) ×ˢ Icc (-(n+1 : ℤ)) (n+1)).card - · rw [square_disjunion] - rw [card_disjUnion, card_product, Nat.cast_add, Nat.cast_mul, card_product, Nat.cast_mul, - Int.card_Icc, Int.card_Icc, Int.toNat_sub_of_le, Int.toNat_sub_of_le, - ← eq_sub_iff_add_eq] at this - · rw [← Nat.cast_inj (R := ℤ), this, Nat.cast_mul, Nat.cast_ofNat, Nat.cast_add_one] - ring_nf - · linarith - · linarith - -theorem natAbs_le_iff_mem_Icc (a : ℤ) (n : ℕ) : - Int.natAbs a ≤ n ↔ a ∈ Finset.Icc (-(n : ℤ)) n := by - rw [mem_Icc, ← abs_le, Int.abs_eq_natAbs, Int.ofNat_le] - -@[simp] -theorem square_mem (n : ℕ) (x : ℤ × ℤ) : x ∈ square n ↔ max x.1.natAbs x.2.natAbs = n := by - simp_rw [square, Finset.mem_filter, Nat.cast_inj, mem_product, and_iff_right_iff_imp, - ← natAbs_le_iff_mem_Icc] - rintro rfl - simp only [le_max_iff, le_refl, true_or, or_true, and_self] - -theorem square_size' {n : ℕ} (h : n ≠ 0) : Finset.card (square n) = 8 * n := by - obtain ⟨n, rfl⟩ := n.exists_eq_succ_of_ne_zero h - exact mod_cast square_size n - -theorem squares_cover_all (y : ℤ × ℤ) : ∃! i : ℕ, y ∈ square i := by - use max y.1.natAbs y.2.natAbs - simp only [square_mem, and_self_iff, forall_eq'] - -@[simp] -lemma square_zero : square 0 = {(0, 0)} := rfl - -theorem square_zero_card : Finset.card (square 0) = 1 := by - rw [square_zero, card_singleton] - -lemma fun_ne_zero_cases (x : Fin 2 → ℤ) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by - rw [Function.ne_iff] - exact Fin.exists_fin_two - -lemma square_mem_ne_zero_iff_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : ⟨x 0, x 1⟩ ∈ square n) : - x ≠ 0 ↔ n ≠ 0 := by - constructor - intro h h0 - simp only [h0, Nat.cast_zero, square_zero, mem_singleton, Prod.mk.injEq] at hx - rw [fun_ne_zero_cases, hx.1, hx.2] at h - simp only [ne_eq, not_true_eq_false, or_self] at * - intro hn h - have hxx : x 0 = 0 ∧ x 1 = 0 := by - simp only [h, Pi.zero_apply, and_self] - rw [hxx.1, hxx.2] at hx - simp only [square_mem, Int.natAbs_zero, max_self] at hx - exact hn (id hx.symm) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 06400d20498d8..a0326fad0995f 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -4,14 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ import Mathlib.Analysis.Complex.UpperHalfPlane.Basic -import Mathlib.Data.Int.Interval import Mathlib.Analysis.SpecialFunctions.Pow.Complex import Mathlib.Analysis.SpecialFunctions.Pow.Real -import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic import Mathlib.Analysis.Complex.UpperHalfPlane.Metric import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.PSeries -import Mathlib.NumberTheory.ModularForms.EisensteinSeries.FinsetDecomposition +import Mathlib.Data.Finset.LocallyFinite.Box +import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic /-! # Uniform convergence of Eisenstein series @@ -20,6 +19,19 @@ We show that `eis` converges locally uniformly on `ℍ` to the Eisenstein series and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. -/ +lemma fun_ne_zero_cases (x : Fin 2 → ℤ) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by + rw [Function.ne_iff]; exact Fin.exists_fin_two + +lemma mem_box_ne_zero_iff_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : + x ≠ 0 ↔ n ≠ 0 := by + constructor + intro h h0 + simp only [h0, Nat.cast_zero, box_zero, mem_singleton, Prod.ext_iff] at hx + rw [fun_ne_zero_cases, hx.1, hx.2] at h + · simp at h + rintro hn rfl + simp [hn, eq_comm] at hx + noncomputable section open Complex Filter UpperHalfPlane Set diff --git a/Mathlib/Order/Birkhoff.lean b/Mathlib/Order/Birkhoff.lean index 16b2a805764af..475bc1914c7c3 100644 --- a/Mathlib/Order/Birkhoff.lean +++ b/Mathlib/Order/Birkhoff.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.Fintype.Order import Mathlib.Order.Irreducible import Mathlib.Order.UpperLower.Basic diff --git a/Mathlib/Order/Disjointed.lean b/Mathlib/Order/Disjointed.lean index ac59da0e5c3a1..9923580d08336 100644 --- a/Mathlib/Order/Disjointed.lean +++ b/Mathlib/Order/Disjointed.lean @@ -102,9 +102,13 @@ theorem disjointedRec_zero {f : ℕ → α} {p : α → Sort*} (hdiff : ∀ ⦃t #align disjointed_rec_zero disjointedRec_zero -- TODO: Find a useful statement of `disjointedRec_succ`. -theorem Monotone.disjointed_eq {f : ℕ → α} (hf : Monotone f) (n : ℕ) : +protected lemma Monotone.disjointed_succ {f : ℕ → α} (hf : Monotone f) (n : ℕ) : disjointed f (n + 1) = f (n + 1) \ f n := by rw [disjointed_succ, hf.partialSups_eq] -#align monotone.disjointed_eq Monotone.disjointed_eq +#align monotone.disjointed_eq Monotone.disjointed_succ + +protected lemma Monotone.disjointed_succ_sup {f : ℕ → α} (hf : Monotone f) (n : ℕ) : + disjointed f (n + 1) ⊔ f n = f (n + 1) := by + rw [hf.disjointed_succ, sdiff_sup_cancel]; exact hf n.le_succ @[simp] theorem partialSups_disjointed (f : ℕ → α) : partialSups (disjointed f) = partialSups f := by From 224f82e3cc63d9a33d88e3861c9e4fa0ad6883d6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Feb 2024 17:47:21 +0000 Subject: [PATCH 036/124] feat: Boxes in locally finite ordered rings Define the sequence of "hollow boxes" indexed by natural numbers as the successive differences of the "boxes" `Icc (-n) n`. --- Mathlib.lean | 3 +- Mathlib/Data/DFinsupp/Interval.lean | 2 +- Mathlib/Data/Fin/Interval.lean | 2 +- Mathlib/Data/Finset/Interval.lean | 2 +- .../Basic.lean} | 0 Mathlib/Data/Finset/LocallyFinite/Box.lean | 92 +++++++++++++++++++ Mathlib/Data/Finsupp/Interval.lean | 2 +- Mathlib/Data/Int/Interval.lean | 2 +- Mathlib/Data/Multiset/Interval.lean | 2 +- Mathlib/Data/Multiset/LocallyFinite.lean | 2 +- Mathlib/Data/Nat/Interval.lean | 2 +- Mathlib/Data/Pi/Interval.lean | 2 +- Mathlib/Order/Birkhoff.lean | 2 +- Mathlib/Order/Disjointed.lean | 8 +- 14 files changed, 110 insertions(+), 13 deletions(-) rename Mathlib/Data/Finset/{LocallyFinite.lean => LocallyFinite/Basic.lean} (100%) create mode 100644 Mathlib/Data/Finset/LocallyFinite/Box.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8472a22e965cd..eef46c5e4d2fe 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1581,7 +1581,8 @@ import Mathlib.Data.Finset.Grade import Mathlib.Data.Finset.Image import Mathlib.Data.Finset.Interval import Mathlib.Data.Finset.Lattice -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic +import Mathlib.Data.Finset.LocallyFinite.Box import Mathlib.Data.Finset.MulAntidiagonal import Mathlib.Data.Finset.NAry import Mathlib.Data.Finset.NatAntidiagonal diff --git a/Mathlib/Data/DFinsupp/Interval.lean b/Mathlib/Data/DFinsupp/Interval.lean index e889b3ca8716c..7f01d1776f341 100644 --- a/Mathlib/Data/DFinsupp/Interval.lean +++ b/Mathlib/Data/DFinsupp/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.Finset.Pointwise import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.DFinsupp.Order diff --git a/Mathlib/Data/Fin/Interval.lean b/Mathlib/Data/Fin/Interval.lean index 4759ca985de07..b8c85526b18f1 100644 --- a/Mathlib/Data/Fin/Interval.lean +++ b/Mathlib/Data/Fin/Interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Nat.Interval -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.fin.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Finset/Interval.lean b/Mathlib/Data/Finset/Interval.lean index 9c787cd92a468..6c654fea23e23 100644 --- a/Mathlib/Data/Finset/Interval.lean +++ b/Mathlib/Data/Finset/Interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Finset.Grade -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.finset.interval from "leanprover-community/mathlib"@"98e83c3d541c77cdb7da20d79611a780ff8e7d90" diff --git a/Mathlib/Data/Finset/LocallyFinite.lean b/Mathlib/Data/Finset/LocallyFinite/Basic.lean similarity index 100% rename from Mathlib/Data/Finset/LocallyFinite.lean rename to Mathlib/Data/Finset/LocallyFinite/Basic.lean diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean new file mode 100644 index 0000000000000..7f2d91f5b3bc8 --- /dev/null +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -0,0 +1,92 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.Ring.Prod +import Mathlib.Data.Int.Interval +import Mathlib.Order.Disjointed +import Mathlib.Tactic.Ring +import Mathlib.Tactic.Zify + +/-! +# Decomposing a locally finite ordered ring into boxes + +This file proves that any locally finite ordered ring can be decomposed into "boxes", namely +differences of consecutive intervals. + +## Implementation notes + +We don't need the full ring structure, only that there is an embedding `ℕ → ` +-/ + +/-! ### General locally finite ordered ring -/ + +namespace Finset +variable {α : Type*} [OrderedRing α] [LocallyFiniteOrder α] [DecidableEq α] {n : ℕ} + +/-- Hollow box centered at `0 : α` going from `-n` to `n`. -/ +def box : ℕ → Finset α := disjointed fun n ↦ Icc (-n : α) n + +private lemma Icc_neg_mono : Monotone fun n : ℕ ↦ Icc (-n : α) n := by + refine fun m n hmn ↦ by apply Icc_subset_Icc <;> simpa using Nat.mono_cast hmn + +@[simp] lemma box_zero : (box 0 : Finset α) = {0} := by simp [box] + +lemma box_succ_eq_sdiff (n : ℕ) : + box (n + 1) = Icc (-n.succ : α) n.succ \ Icc (-n) n := Icc_neg_mono.disjointed_succ _ + +lemma disjoint_box_succ_prod (n : ℕ) : Disjoint (box (n + 1)) (Icc (-n : α) n) := by + rw [box_succ_eq_sdiff]; exact disjoint_sdiff_self_left + +@[simp] lemma box_succ_union_prod (n : ℕ) : + box (n + 1) ∪ Icc (-n : α) n = Icc (-n.succ : α) n.succ := Icc_neg_mono.disjointed_succ_sup _ + +@[simp] lemma box_succ_disjUnion (n : ℕ) : + (box (n + 1)).disjUnion (Icc (-n : α) n) (disjoint_box_succ_prod _) = + Icc (-n.succ : α) n.succ := by rw [disjUnion_eq_union, box_succ_union_prod] + +@[simp] lemma zero_mem_box : (0 : α) ∈ box n ↔ n = 0 := by cases n <;> simp [box_succ_eq_sdiff] + +end Finset + +open Finset + +/-! ### Product of locally finite ordered rings -/ + +namespace Prod +variable {α β : Type*} [OrderedRing α] [OrderedRing β] [LocallyFiniteOrder α] [LocallyFiniteOrder β] + [DecidableEq α] [DecidableEq β] [@DecidableRel (α × β) (· ≤ ·)] + +@[simp] lemma card_box_succ (n : ℕ) : + (box (n + 1) : Finset (α × β)).card = + (Icc (-n.succ : α) n.succ).card * (Icc (-n.succ : β) n.succ).card - + (Icc (-n : α) n).card * (Icc (-n : β) n).card := by + rw [box_succ_eq_sdiff, card_sdiff (Icc_neg_mono n.le_succ), Prod.card_Icc, Prod.card_Icc]; rfl + +end Prod + +/-! ### `ℤ × ℤ` -/ + +namespace Int +variable {n : ℕ} {x : ℤ × ℤ} + +attribute [norm_cast] toNat_ofNat + +lemma card_box : ∀ {n}, n ≠ 0 → (box n : Finset (ℤ × ℤ)).card = 8 * n + | n + 1, _ => by + simp_rw [Prod.card_box_succ, card_Icc, sub_neg_eq_add] + norm_cast + refine tsub_eq_of_eq_add ?_ + zify + ring + +@[simp] lemma mem_box : ∀ {n}, x ∈ box n ↔ max x.1.natAbs x.2.natAbs = n + | 0 => by simp [Prod.ext_iff] + | n + 1 => by simp [box_succ_eq_sdiff, Prod.le_def]; omega + +-- TODO: Can this be generalised to locally finite archimedean ordered rings? +lemma existsUnique_mem_box (x : ℤ × ℤ) : ∃! n : ℕ, x ∈ box n := by + use max x.1.natAbs x.2.natAbs; simp only [mem_box, and_self_iff, forall_eq'] + +end Int diff --git a/Mathlib/Data/Finsupp/Interval.lean b/Mathlib/Data/Finsupp/Interval.lean index 95383857acad4..6c159b6dce65d 100644 --- a/Mathlib/Data/Finsupp/Interval.lean +++ b/Mathlib/Data/Finsupp/Interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Finset.Finsupp -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.Finsupp.Order #align_import data.finsupp.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index e0f08d5e2b71d..99e021b2ca256 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Order.LocallyFinite -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.int.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Multiset/Interval.lean b/Mathlib/Data/Multiset/Interval.lean index 9d8fafeab1c0f..fd010897f5423 100644 --- a/Mathlib/Data/Multiset/Interval.lean +++ b/Mathlib/Data/Multiset/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.DFinsupp.Interval import Mathlib.Data.DFinsupp.Multiset import Mathlib.Data.Nat.Interval diff --git a/Mathlib/Data/Multiset/LocallyFinite.lean b/Mathlib/Data/Multiset/LocallyFinite.lean index 137df9cdf554e..8bb625e14df1d 100644 --- a/Mathlib/Data/Multiset/LocallyFinite.lean +++ b/Mathlib/Data/Multiset/LocallyFinite.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.multiset.locally_finite from "leanprover-community/mathlib"@"59694bd07f0a39c5beccba34bd9f413a160782bf" diff --git a/Mathlib/Data/Nat/Interval.lean b/Mathlib/Data/Nat/Interval.lean index e5344ad3cbbb4..058b72a4d533d 100644 --- a/Mathlib/Data/Nat/Interval.lean +++ b/Mathlib/Data/Nat/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.nat.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Pi/Interval.lean b/Mathlib/Data/Pi/Interval.lean index c63f64d6b9887..2334d61926de6 100644 --- a/Mathlib/Data/Pi/Interval.lean +++ b/Mathlib/Data/Pi/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.Fintype.BigOperators #align_import data.pi.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Order/Birkhoff.lean b/Mathlib/Order/Birkhoff.lean index 16b2a805764af..475bc1914c7c3 100644 --- a/Mathlib/Order/Birkhoff.lean +++ b/Mathlib/Order/Birkhoff.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic import Mathlib.Data.Fintype.Order import Mathlib.Order.Irreducible import Mathlib.Order.UpperLower.Basic diff --git a/Mathlib/Order/Disjointed.lean b/Mathlib/Order/Disjointed.lean index ac59da0e5c3a1..9923580d08336 100644 --- a/Mathlib/Order/Disjointed.lean +++ b/Mathlib/Order/Disjointed.lean @@ -102,9 +102,13 @@ theorem disjointedRec_zero {f : ℕ → α} {p : α → Sort*} (hdiff : ∀ ⦃t #align disjointed_rec_zero disjointedRec_zero -- TODO: Find a useful statement of `disjointedRec_succ`. -theorem Monotone.disjointed_eq {f : ℕ → α} (hf : Monotone f) (n : ℕ) : +protected lemma Monotone.disjointed_succ {f : ℕ → α} (hf : Monotone f) (n : ℕ) : disjointed f (n + 1) = f (n + 1) \ f n := by rw [disjointed_succ, hf.partialSups_eq] -#align monotone.disjointed_eq Monotone.disjointed_eq +#align monotone.disjointed_eq Monotone.disjointed_succ + +protected lemma Monotone.disjointed_succ_sup {f : ℕ → α} (hf : Monotone f) (n : ℕ) : + disjointed f (n + 1) ⊔ f n = f (n + 1) := by + rw [hf.disjointed_succ, sdiff_sup_cancel]; exact hf n.le_succ @[simp] theorem partialSups_disjointed (f : ℕ → α) : partialSups (disjointed f) = partialSups f := by From 1c0a1e85073e26130ba70088e104575138f5183e Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 17:58:38 +0000 Subject: [PATCH 037/124] moved some lemmas around, while waiting for boxes to be PRd --- Mathlib/Algebra/Order/Field/Power.lean | 5 +++++ Mathlib/Data/Complex/Abs.lean | 7 +++++++ .../EisensteinSeries/UniformConvergence.lean | 20 ++----------------- 3 files changed, 14 insertions(+), 18 deletions(-) diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index a4b5838cf166e..815d85e7023e4 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -119,6 +119,11 @@ section LinearOrderedField variable [LinearOrderedField α] {a b c d : α} {n : ℤ} +lemma one_le_sq_div_abs_sq (ha : a ≠ 0) : 1 ≤ a^2 / |a|^2 := by + rw [_root_.sq_abs, le_div_iff'] + simp only [mul_one, le_refl] + exact (sq_pos_iff a).mpr ha + /-! ### Lemmas about powers to numerals. -/ section bits diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 80b27a5f27f64..918749090090f 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -92,6 +92,13 @@ theorem abs_ofNat (n : ℕ) [n.AtLeastTwo] : Complex.abs (no_index (OfNat.ofNat n : ℂ)) = OfNat.ofNat n := abs_natCast n +lemma complex_abs_of_int_eq_natAbs (a : ℤ) : Complex.abs a = a.natAbs := by + have : Complex.abs a = Complex.abs (a : ℝ) := by rfl + rw [this, Complex.abs_ofReal] + norm_cast + rw [Int.abs_eq_natAbs a] + rfl + theorem mul_self_abs (z : ℂ) : Complex.abs z * Complex.abs z = normSq z := Real.mul_self_sqrt (normSq_nonneg _) #align complex.mul_self_abs Complex.mul_self_abs diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 06400d20498d8..625a94de1b805 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -90,22 +90,6 @@ theorem auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : r z ≤ C int_cast_im] at * exact H1 ---any suggestions as to where to put the lemmas below? - -lemma one_le_sq_div_abs_sq (a : ℝ) (ha : a ≠ 0) : 1 ≤ a^2 / |a|^2 := by - rw [_root_.sq_abs, le_div_iff'] - simp only [mul_one, le_refl] - exact (sq_pos_iff a).mpr ha - ---#find_home! one_le_sq_div_abs_sq - -lemma int_abs_eq_complex_abs (a : ℤ) : Complex.abs a = a.natAbs := by - have : Complex.abs a = Complex.abs (a : ℝ) := by rfl - rw [this, abs_ofReal] - norm_cast - rw [Int.abs_eq_natAbs a] - rfl - lemma ne_zero_if_max {x : Fin 2 → ℤ} (hx : x ≠ 0) (h : (max (x 0).natAbs (x 1).natAbs) = (x 0).natAbs) : (x 0) ≠ 0 := by intro h0 @@ -135,7 +119,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 norm_cast - rw [int_abs_eq_complex_abs] + rw [complex_abs_of_int_eq_natAbs] · right rw [H2] have : (x 1 : ℝ) ≠ 0 := by @@ -144,7 +128,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 norm_cast - rw [int_abs_eq_complex_abs] + rw [complex_abs_of_int_eq_natAbs] /-This should work for complex `k` (with a condition on `k.re`) but last I checked we were missing some lemmas about this -/ From de4faf1df1837e2941672c5a363b5e053d4e0cb9 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 19:06:07 +0000 Subject: [PATCH 038/124] fix updates --- .../EisensteinSeries/UniformConvergence.lean | 47 ++++++++++--------- 1 file changed, 25 insertions(+), 22 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 88b4cf9062d1d..9f0ae2a7eb2b8 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -12,6 +12,8 @@ import Mathlib.Analysis.PSeries import Mathlib.Data.Finset.LocallyFinite.Box import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic +open Finset + /-! # Uniform convergence of Eisenstein series @@ -127,7 +129,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : rw [H1] have : (x 0 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max hx H1) - have h1 := one_le_sq_div_abs_sq (x 0 : ℝ) this + have h1 := one_le_sq_div_abs_sq this simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 norm_cast @@ -136,7 +138,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : rw [H2] have : (x 1 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max' hx H2) - have h1 := one_le_sq_div_abs_sq (x 1 : ℝ) this + have h1 := one_le_sq_div_abs_sq this simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 norm_cast @@ -152,7 +154,7 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : have h1 : ((n : ℝ) : ℂ)^k ≠ 0 := by rw [pow_ne_zero_iff hk] norm_cast - rw [square_mem_ne_zero_iff_ne_zero n x (by rw [square_mem])] at hx + rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx exact hx have hc : Complex.abs ((n : ℝ)^k : ℂ) = n^k := by simp only [Nat.cast_max, map_pow, abs_ofReal, ge_iff_le, abs_nonneg, le_max_iff, @@ -183,12 +185,12 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : · simp only [ne_eq, not_not] at hk simp only [hk, pow_zero, Nat.cast_max, mul_one, map_one, le_refl] -theorem eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) - (hx : ⟨x 0, x 1⟩ ∈ square n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ +theorem eis_is_bounded_on_box (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) + (hx : (⟨x 0, x 1⟩ : ℤ × ℤ) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by by_cases hn : n = 0 · rw [hn] at hx - simp only [CharP.cast_eq_zero, square_zero, Finset.mem_singleton, Prod.mk.injEq] at hx + simp only [box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx by_cases hk : k = 0 rw [hk] at * simp only [ pow_zero, map_one, inv_one, mul_one, le_refl] @@ -199,9 +201,9 @@ theorem eis_is_bounded_on_square (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → simp only [Int.cast_zero, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, CharP.cast_eq_zero, map_mul, abs_ofReal, mul_zero, le_refl] · have hx2 : x ≠ 0 := by - rw [square_mem_ne_zero_iff_ne_zero n x hx] + rw [mem_box_ne_zero_iff_ne_zero n x hx] exact hn - simp only [square_mem] at hx + simp only [Int.mem_box] at hx rw [inv_le_inv] simp only [← hx, map_mul, map_pow, abs_ofReal, abs_natCast, Nat.cast_max, ge_iff_le] convert (bound z x hx2 k) @@ -282,8 +284,8 @@ lemma summable_r_pow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : convert (Real.summable_nat_rpow_inv.2 hk) norm_cast -lemma summable_over_square {k : ℤ} (z : ℍ) (h : 3 ≤ k): - Summable (fun n : ℕ => ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by +lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): + Summable (fun n : ℕ => ∑ v in (box n : Finset (ℤ × ℤ)), (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by simp only [one_div, Finset.sum_const, nsmul_eq_mul] apply Summable.congr (summable_r_pow z h) intro b @@ -293,23 +295,24 @@ lemma summable_over_square {k : ℤ} (z : ℍ) (h : 3 ≤ k): have hk1 : k - 1 ≠ 0 := by linarith norm_cast rw [zero_zpow k hk0, zero_zpow (k - 1) hk1] - simp only [inv_zero, mul_zero, square_zero, Finset.card_singleton, Nat.cast_one] - · rw [square_size' b0, zpow_sub_one₀ (a:= ( b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] + simp only [inv_zero, mul_zero, box_zero, Finset.card_singleton, Nat.cast_one] + · rw [Int.card_box b0, zpow_sub_one₀ (a:= ( b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] ring_nf lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by - rw [summable_lemma _ _ (fun (n : ℕ) => square n) squares_cover_all] - · have : ∀ n : ℕ, ∑ v in square n, (1 / (r z) ^ k) * ((max v.1.natAbs v.2.natAbs: ℝ) ^ k)⁻¹ = - ∑ v in square n, (1 / (r z) ^ k) * ((n : ℝ)^k)⁻¹ := by + rw [summable_lemma _ _ (fun (n : ℕ) => box n) Int.existsUnique_mem_box] + · have : ∀ n : ℕ, ∑ v in (box n : Finset (ℤ × ℤ)), + (1 / (r z) ^ k) * ((max v.1.natAbs v.2.natAbs: ℝ) ^ k)⁻¹ = + ∑ v in box n, (1 / (r z) ^ k) * ((n : ℝ)^k)⁻¹ := by intro n apply Finset.sum_congr rfl intro x hx - simp only [square_mem] at hx + simp only [Int.mem_box] at hx congr norm_cast - apply Summable.congr (summable_over_square z h) + apply Summable.congr (summable_over_box z h) intro b apply (this b).symm · intro y @@ -327,7 +330,7 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : have hk0 : 0 ≤ k := by linarith lift k to ℕ using hk0 rw [tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] - simp only [top_eq_univ, subset_univ, eisensteinSeries, forall_true_left] + simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] intro K hK obtain ⟨A, B, hB, HABK⟩:= subset_slice_of_isCompact hK have hu : Summable fun x : (gammaSet N a ) => @@ -337,12 +340,11 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : simp only [zpow_coe_nat, one_div, Function.comp_apply] apply tendstoUniformlyOn_tsum hu intro v x hx - have sq := square_mem (max (v.1 0).natAbs (v.1 1).natAbs ) ⟨(v.1 0), v.1 1⟩ - have := eis_is_bounded_on_square k x (max (v.1 0).natAbs (v.1 1).natAbs ) v + have := eis_is_bounded_on_box k x (max (v.1 0).natAbs (v.1 1).natAbs ) v simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_coe_nat, one_div, map_pow, map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, ge_iff_le] at * - apply le_trans (this sq) + apply le_trans (this ?_) rw [mul_comm] apply mul_le_mul _ (by rfl) repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg, @@ -353,4 +355,5 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : · exact r_lower_bound_on_slice hB ⟨x, HABK hx⟩ · apply pow_pos (abs_pos.mpr (ne_of_gt (r_pos x))) · apply pow_pos (r_pos _) - · simp only [top_eq_univ, isOpen_univ] + · simp only [Int.mem_box] + · simp From 0788f1b425276eda4a5b85adaec0e8f516b5f442 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 19:07:29 +0000 Subject: [PATCH 039/124] fix simp --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 9f0ae2a7eb2b8..1182f1a72c07e 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -302,7 +302,7 @@ lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by - rw [summable_lemma _ _ (fun (n : ℕ) => box n) Int.existsUnique_mem_box] + rw [summable_lemma _ _ (fun (n : ℕ) => box n) Int.existsUnique_mem_box] · have : ∀ n : ℕ, ∑ v in (box n : Finset (ℤ × ℤ)), (1 / (r z) ^ k) * ((max v.1.natAbs v.2.natAbs: ℝ) ^ k)⁻¹ = ∑ v in box n, (1 / (r z) ^ k) * ((n : ℝ)^k)⁻¹ := by @@ -356,4 +356,4 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : · apply pow_pos (abs_pos.mpr (ne_of_gt (r_pos x))) · apply pow_pos (r_pos _) · simp only [Int.mem_box] - · simp + · simp only [Set.top_eq_univ, isOpen_univ] From 5d9708c67ef5414ec77c5e1636291f5a722d9ab6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Feb 2024 20:07:36 +0100 Subject: [PATCH 040/124] fix simp nf --- Mathlib/Data/Finset/LocallyFinite/Box.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean index 7f2d91f5b3bc8..a4c633839a38e 100644 --- a/Mathlib/Data/Finset/LocallyFinite/Box.lean +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -42,7 +42,7 @@ lemma disjoint_box_succ_prod (n : ℕ) : Disjoint (box (n + 1)) (Icc (-n : α) n @[simp] lemma box_succ_union_prod (n : ℕ) : box (n + 1) ∪ Icc (-n : α) n = Icc (-n.succ : α) n.succ := Icc_neg_mono.disjointed_succ_sup _ -@[simp] lemma box_succ_disjUnion (n : ℕ) : +lemma box_succ_disjUnion (n : ℕ) : (box (n + 1)).disjUnion (Icc (-n : α) n) (disjoint_box_succ_prod _) = Icc (-n.succ : α) n.succ := by rw [disjUnion_eq_union, box_succ_union_prod] From cc886d89e66d22c21a52c5ef455c838a4ac306f7 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 19:10:55 +0000 Subject: [PATCH 041/124] fix doc string --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 1182f1a72c07e..f78ec77c07b1b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -12,8 +12,6 @@ import Mathlib.Analysis.PSeries import Mathlib.Data.Finset.LocallyFinite.Box import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic -open Finset - /-! # Uniform convergence of Eisenstein series @@ -21,6 +19,8 @@ We show that `eis` converges locally uniformly on `ℍ` to the Eisenstein series and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. -/ +open Finset + lemma fun_ne_zero_cases (x : Fin 2 → ℤ) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by rw [Function.ne_iff]; exact Fin.exists_fin_two From 57154cac9a294e0dbedf6299471b9e81d863dd5e Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 19:20:10 +0000 Subject: [PATCH 042/124] save before trying to generalise the bound --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index f78ec77c07b1b..5b9a1e8a57768 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -144,8 +144,8 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : norm_cast rw [complex_abs_of_int_eq_natAbs] -/-This should work for complex `k` (with a condition on `k.re`) but last I checked we were missing -some lemmas about this -/ +/-This should work for complex `k` (with a condition on `k.re`and taking the power out of the abs) +but last I checked we were missing some lemmas about this -/ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs)^k ≤ Complex.abs (((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ)) ^ k) := by @@ -186,7 +186,7 @@ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : simp only [hk, pow_zero, Nat.cast_max, mul_one, map_one, le_refl] theorem eis_is_bounded_on_box (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) - (hx : (⟨x 0, x 1⟩ : ℤ × ℤ) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ + (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by by_cases hn : n = 0 · rw [hn] at hx @@ -326,7 +326,7 @@ end summability theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) - ( fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by + (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by have hk0 : 0 ≤ k := by linarith lift k to ℕ using hk0 rw [tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] From 627b254e0be020f4096f3d3ea5c01ce3f601cc3a Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 21:44:44 +0000 Subject: [PATCH 043/124] fix --- Mathlib/Algebra/Order/Field/Power.lean | 8 ++++---- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 4 ++-- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index 815d85e7023e4..2e0a7110c7498 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -119,10 +119,10 @@ section LinearOrderedField variable [LinearOrderedField α] {a b c d : α} {n : ℤ} -lemma one_le_sq_div_abs_sq (ha : a ≠ 0) : 1 ≤ a^2 / |a|^2 := by - rw [_root_.sq_abs, le_div_iff'] - simp only [mul_one, le_refl] - exact (sq_pos_iff a).mpr ha +lemma one_eq_sq_div_abs_sq (ha : a ≠ 0) : a^2 / |a|^2 = 1 := by + rw [_root_.sq_abs] + apply div_self + exact pow_ne_zero 2 ha /-! ### Lemmas about powers to numerals. -/ diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 5b9a1e8a57768..f7a358e68db3d 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -129,7 +129,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : rw [H1] have : (x 0 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max hx H1) - have h1 := one_le_sq_div_abs_sq this + have h1 := (one_eq_sq_div_abs_sq this).symm.le simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 norm_cast @@ -138,7 +138,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : rw [H2] have : (x 1 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max' hx H2) - have h1 := one_le_sq_div_abs_sq this + have h1 := (one_eq_sq_div_abs_sq this).symm.le simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * convert h1 norm_cast From f6d0b12c6778c0bfd047c68b7e5cb881f4033276 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 21:45:54 +0000 Subject: [PATCH 044/124] fix --- Mathlib/Data/Int/Interval.lean | 2 +- Mathlib/Data/Nat/Interval.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index 99e021b2ca256..e0f08d5e2b71d 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Order.LocallyFinite -import Mathlib.Data.Finset.LocallyFinite.Basic +import Mathlib.Data.Finset.LocallyFinite #align_import data.int.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Nat/Interval.lean b/Mathlib/Data/Nat/Interval.lean index ca46a4e1a1a91..ec44e6b45a493 100644 --- a/Mathlib/Data/Nat/Interval.lean +++ b/Mathlib/Data/Nat/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite.Basic +import Mathlib.Data.Finset.LocallyFinite #align_import data.nat.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" From 682f52555086e488f4a6c55942ca1ca04d0ee327 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 22:15:28 +0000 Subject: [PATCH 045/124] remove junk lemma --- Mathlib/Algebra/Order/Field/Power.lean | 5 ----- .../EisensteinSeries/UniformConvergence.lean | 12 ++++++++---- 2 files changed, 8 insertions(+), 9 deletions(-) diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index 2e0a7110c7498..a4b5838cf166e 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -119,11 +119,6 @@ section LinearOrderedField variable [LinearOrderedField α] {a b c d : α} {n : ℤ} -lemma one_eq_sq_div_abs_sq (ha : a ≠ 0) : a^2 / |a|^2 = 1 := by - rw [_root_.sq_abs] - apply div_self - exact pow_ne_zero 2 ha - /-! ### Lemmas about powers to numerals. -/ section bits diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index f7a358e68db3d..95244dc4c00a1 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -129,18 +129,22 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : rw [H1] have : (x 0 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max hx H1) - have h1 := (one_eq_sq_div_abs_sq this).symm.le + have h1 : (x 0 : ℝ)^2/( _root_.abs (x 0 : ℝ))^2 = 1 := by + simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, + this, div_self] simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * - convert h1 + convert h1.symm.le norm_cast rw [complex_abs_of_int_eq_natAbs] · right rw [H2] have : (x 1 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max' hx H2) - have h1 := (one_eq_sq_div_abs_sq this).symm.le + have h1 : (x 1 : ℝ)^2/( _root_.abs (x 1 : ℝ))^2 = 1 := by + simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, + this, div_self] simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * - convert h1 + convert h1.symm.le norm_cast rw [complex_abs_of_int_eq_natAbs] From d2182d3a187e810c52c5f25dd4664444bd241ae7 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 23:03:39 +0000 Subject: [PATCH 046/124] fix merge issue --- Mathlib/Data/Int/Interval.lean | 2 +- Mathlib/Data/Nat/Interval.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index e0f08d5e2b71d..c10365e708a5a 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.CharZero.Lemmas -import Mathlib.Order.LocallyFinite +import Mathlib.Order.LocallyFinite.Basic import Mathlib.Data.Finset.LocallyFinite #align_import data.int.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" diff --git a/Mathlib/Data/Nat/Interval.lean b/Mathlib/Data/Nat/Interval.lean index ec44e6b45a493..ca46a4e1a1a91 100644 --- a/Mathlib/Data/Nat/Interval.lean +++ b/Mathlib/Data/Nat/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.nat.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" From e653c329bb10aaf369b60bea7ae12198caf610da Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Tue, 13 Feb 2024 23:04:02 +0000 Subject: [PATCH 047/124] fix merge issue --- Mathlib/Data/Int/Interval.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index c10365e708a5a..99e021b2ca256 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.CharZero.Lemmas -import Mathlib.Order.LocallyFinite.Basic -import Mathlib.Data.Finset.LocallyFinite +import Mathlib.Order.LocallyFinite +import Mathlib.Data.Finset.LocallyFinite.Basic #align_import data.int.interval from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" From f82014d65ea18abab5eedcac2d03d609c4da0421 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Wed, 14 Feb 2024 08:55:58 +0000 Subject: [PATCH 048/124] shake fix? --- Mathlib/Data/Finset/LocallyFinite/Box.lean | 2 +- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 1 + 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean index 69835f2d631af..ccc22efdec508 100644 --- a/Mathlib/Data/Finset/LocallyFinite/Box.lean +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -6,12 +6,12 @@ Authors: Yaël Dillies import Mathlib.Data.Int.Cast.Prod import Mathlib.Data.Int.Interval import Mathlib.Order.Disjointed -import Mathlib.Tactic.Ring.RingNF import Mathlib.Tactic.Linarith import Mathlib.Tactic.NormNum.Ineq import Mathlib.Algebra.Order.Group.Prod -- TODO: Why is the `OrderedRing` instance not under `Algebra.Order`? import Mathlib.Algebra.Ring.Prod +import Mathlib.Tactic.Ring.RingNF /-! # Decomposing a locally finite ordered ring into boxes diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 95244dc4c00a1..4725e2ef0684a 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -11,6 +11,7 @@ import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.PSeries import Mathlib.Data.Finset.LocallyFinite.Box import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic +import Mathlib.Tactic.Ring.RingNF /-! # Uniform convergence of Eisenstein series From 034f09db2688f5c7ab1e96656dc124ed0067d9b6 Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Wed, 14 Feb 2024 08:57:39 +0000 Subject: [PATCH 049/124] fix import --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 4725e2ef0684a..590709ddbc730 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -3,15 +3,12 @@ Copyright (c) 2024 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ -import Mathlib.Analysis.Complex.UpperHalfPlane.Basic -import Mathlib.Analysis.SpecialFunctions.Pow.Complex -import Mathlib.Analysis.SpecialFunctions.Pow.Real + import Mathlib.Analysis.Complex.UpperHalfPlane.Metric import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.PSeries import Mathlib.Data.Finset.LocallyFinite.Box import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic -import Mathlib.Tactic.Ring.RingNF /-! # Uniform convergence of Eisenstein series From c6383741c63302f30d657b577021ffec7c1eb3ed Mon Sep 17 00:00:00 2001 From: CBirkbeck Date: Wed, 14 Feb 2024 09:33:28 +0000 Subject: [PATCH 050/124] shake ignore? --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- scripts/noshake.json | 1 + 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 590709ddbc730..d959032ca0d0c 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -146,7 +146,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : norm_cast rw [complex_abs_of_int_eq_natAbs] -/-This should work for complex `k` (with a condition on `k.re`and taking the power out of the abs) +/--This should work for complex `k` (with a condition on `k.re`and taking the power out of the abs) but last I checked we were missing some lemmas about this -/ lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs)^k ≤ diff --git a/scripts/noshake.json b/scripts/noshake.json index 191caa4ac0023..4532931d34e2a 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -302,6 +302,7 @@ "Mathlib.Data.Multiset.Bind": ["Mathlib.GroupTheory.GroupAction.Defs"], "Mathlib.Data.List.EditDistance.Defs": ["Mathlib.Data.Nat.Basic"], "Mathlib.Data.List.Basic": ["Mathlib.Init.Core"], + "Mathlib.Data.Finset.LocallyFinite.Box": ["Mathlib.Tactic.Ring.RingNF"], "Mathlib.Control.Traversable.Instances": ["Mathlib.Control.Applicative"], "Mathlib.Control.Monad.Basic": ["Mathlib.Init.Control.Lawful"], "Mathlib.Control.Basic": ["Mathlib.Init.Function"], From 12d7fdfd4e588b39b63b7194477fc6e6ceee5c10 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 14 Feb 2024 18:37:22 +0000 Subject: [PATCH 051/124] save --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index d959032ca0d0c..28d54148511a4 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -293,10 +293,8 @@ lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): intro b by_cases b0 : b = 0 · rw [b0] - have hk0 : k ≠ 0 := by linarith - have hk1 : k - 1 ≠ 0 := by linarith norm_cast - rw [zero_zpow k hk0, zero_zpow (k - 1) hk1] + rw [zero_zpow k (by linarith), zero_zpow (k - 1) (by linarith)] simp only [inv_zero, mul_zero, box_zero, Finset.card_singleton, Nat.cast_one] · rw [Int.card_box b0, zpow_sub_one₀ (a:= ( b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] From c99c5a9b86cedb930331da687d6d5d57c89fff1a Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 16 Feb 2024 14:45:56 +0000 Subject: [PATCH 052/124] fix merge update --- .../EisensteinSeries/UniformConvergence.lean | 14 ++++++++------ 1 file changed, 8 insertions(+), 6 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 28d54148511a4..87ab04373f0d3 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -119,9 +119,11 @@ lemma ne_zero_if_max' {x : Fin 2 → ℤ} (hx : x ≠ 0) rw [fun_ne_zero_cases, h1, h0] at hx simp only [ne_eq, not_true_eq_false, or_self] at * +example (a : ℤ) : (a.natAbs)^2 = a^2 := by exact Int.natAbs_eq_iff_sq_eq.mp rfl + lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : - (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs))^2 ∨ - (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs))^2 := by + (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 ∨ + (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 := by cases' (max_choice (x 0).natAbs (x 1).natAbs) with H1 H2 · left rw [H1] @@ -130,10 +132,10 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : have h1 : (x 0 : ℝ)^2/( _root_.abs (x 0 : ℝ))^2 = 1 := by simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, this, div_self] - simp only [ne_eq, max_eq_left_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * + rw [div_pow] convert h1.symm.le norm_cast - rw [complex_abs_of_int_eq_natAbs] + exact Int.cast_natAbs (x 0) · right rw [H2] have : (x 1 : ℝ) ≠ 0 := by @@ -141,10 +143,10 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : have h1 : (x 1 : ℝ)^2/( _root_.abs (x 1 : ℝ))^2 = 1 := by simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, this, div_self] - simp only [ne_eq, max_eq_right_iff, Int.cast_eq_zero, int_cast_abs, div_pow, ge_iff_le] at * + rw [div_pow] convert h1.symm.le norm_cast - rw [complex_abs_of_int_eq_natAbs] + exact Int.cast_natAbs (x 1) /--This should work for complex `k` (with a condition on `k.re`and taking the power out of the abs) but last I checked we were missing some lemmas about this -/ From 4b566d28db12dd40071f4fa19e4b279179cfde0d Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 16 Feb 2024 14:50:04 +0000 Subject: [PATCH 053/124] remove junk lemma --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 87ab04373f0d3..af5e08f75d1dc 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -119,8 +119,6 @@ lemma ne_zero_if_max' {x : Fin 2 → ℤ} (hx : x ≠ 0) rw [fun_ne_zero_cases, h1, h0] at hx simp only [ne_eq, not_true_eq_false, or_self] at * -example (a : ℤ) : (a.natAbs)^2 = a^2 := by exact Int.natAbs_eq_iff_sq_eq.mp rfl - lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 ∨ (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 := by From 23246b3afa66d0709b4c3e10a01ac90d0de0135b Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 16 Feb 2024 15:06:29 +0000 Subject: [PATCH 054/124] golf --- .../EisensteinSeries/UniformConvergence.lean | 10 ++-------- 1 file changed, 2 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index af5e08f75d1dc..0fa05bc520371 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -124,27 +124,21 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 := by cases' (max_choice (x 0).natAbs (x 1).natAbs) with H1 H2 · left - rw [H1] + rw [H1, div_pow, Int.cast_natAbs (x 0),Int.cast_abs] have : (x 0 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max hx H1) have h1 : (x 0 : ℝ)^2/( _root_.abs (x 0 : ℝ))^2 = 1 := by simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, this, div_self] - rw [div_pow] convert h1.symm.le - norm_cast - exact Int.cast_natAbs (x 0) · right - rw [H2] + rw [H2,div_pow, Int.cast_natAbs (x 1),Int.cast_abs] have : (x 1 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max' hx H2) have h1 : (x 1 : ℝ)^2/( _root_.abs (x 1 : ℝ))^2 = 1 := by simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, this, div_self] - rw [div_pow] convert h1.symm.le - norm_cast - exact Int.cast_natAbs (x 1) /--This should work for complex `k` (with a condition on `k.re`and taking the power out of the abs) but last I checked we were missing some lemmas about this -/ From c4bf406c542e9ab98950df06a4a8e6b2ea53f55e Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 16 Feb 2024 17:17:50 +0000 Subject: [PATCH 055/124] update bound to real powers --- .../EisensteinSeries/UniformConvergence.lean | 146 +++++++++--------- 1 file changed, 76 insertions(+), 70 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 0fa05bc520371..688d8135d0c45 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -130,7 +130,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : have h1 : (x 0 : ℝ)^2/( _root_.abs (x 0 : ℝ))^2 = 1 := by simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, this, div_self] - convert h1.symm.le + exact h1.symm.le · right rw [H2,div_pow, Int.cast_natAbs (x 1),Int.cast_abs] have : (x 1 : ℝ) ≠ 0 := by @@ -138,86 +138,92 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : have h1 : (x 1 : ℝ)^2/( _root_.abs (x 1 : ℝ))^2 = 1 := by simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, this, div_self] - convert h1.symm.le - -/--This should work for complex `k` (with a condition on `k.re`and taking the power out of the abs) -but last I checked we were missing some lemmas about this -/ -lemma bound (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) (k : ℕ) : - ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs)^k ≤ - Complex.abs (((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ)) ^ k) := by - by_cases hk : k ≠ 0 - · let n := max (x 0).natAbs (x 1).natAbs - have h1 : ((n : ℝ) : ℂ)^k ≠ 0 := by - rw [pow_ne_zero_iff hk] - norm_cast - rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx - exact hx - have hc : Complex.abs ((n : ℝ)^k : ℂ) = n^k := by - simp only [Nat.cast_max, map_pow, abs_ofReal, ge_iff_le, abs_nonneg, le_max_iff, - Nat.cast_nonneg, or_self] - congr - simp only [abs_eq_self, le_max_iff, Nat.cast_nonneg, or_self] - have h11 : ((x 0) * ↑z + (x 1)) ^ (k : ℤ) = - (((x 0 : ℝ) / (n : ℝ) ) * (z : ℂ) + (x 1 : ℝ) /(n : ℝ)) ^ (k : ℤ) * ((n : ℝ)^ (k : ℤ)) := by - simp only [Nat.cast_max] at h1 + exact h1.symm.le + +lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) : + ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k ≤ + (Complex.abs ((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ))) ^ k := by + by_cases hk0 : k ≠ 0 + let n := max (x 0).natAbs (x 1).natAbs + have h11 : ((x 0) * ↑z + (x 1)) = + (((x 0 : ℝ) / (n : ℝ) ) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by field_simp - cases' (div_max_sq_ge_one x hx) with H1 H2 - · norm_cast at * - rw [h11] - simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * - apply mul_le_mul ?_ (by rfl) - simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, - pow_nonneg] - apply pow_nonneg (Complex.abs.nonneg _) k - apply pow_le_pow_left (r_pos _).le (auxbound1 z (x 1 / n) H1) - · norm_cast at * - rw [h11] - simp only [hc, map_pow, map_mul, abs_ofReal, Complex.abs_abs, ge_iff_le, zpow_coe_nat] at * - apply mul_le_mul ?_ (by rfl) - simp only [Nat.cast_pow, Nat.cast_max, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, - pow_nonneg] - apply pow_nonneg (Complex.abs.nonneg _) k - apply pow_le_pow_left (r_pos _).le (auxbound2 z (x 0 / n) H2) - · simp only [ne_eq, not_not] at hk - simp only [hk, pow_zero, Nat.cast_max, mul_one, map_one, le_refl] - -theorem eis_is_bounded_on_box (k : ℕ) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) - (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ - (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by + rw [← mul_div, div_self] + simp + rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx + norm_cast at * + rw [h11] + simp only [Nat.cast_max, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] + rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] + cases' (div_max_sq_ge_one x hx) with H1 H2 + · apply mul_le_mul + simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) + norm_cast + apply Real.rpow_nonneg + rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx + simp only [le_max_iff, Nat.cast_nonneg, or_self] + apply Real.rpow_nonneg (Complex.abs.nonneg _) k + · apply mul_le_mul + simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) + norm_cast + apply Real.rpow_nonneg + rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx + simp only [le_max_iff, Nat.cast_nonneg, or_self] + apply Real.rpow_nonneg (Complex.abs.nonneg _) k + · simp only [ne_eq, not_not] at hk0 + simp only [hk0, Real.rpow_zero, Nat.cast_max, mul_one, map_one, le_refl] + +theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) + (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)))) ^ (-k) ≤ + ( ((r z) ^ (-k) * n ^ (-k))) := by by_cases hn : n = 0 · rw [hn] at hx simp only [box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx - by_cases hk : k = 0 - rw [hk] at * - simp only [ pow_zero, map_one, inv_one, mul_one, le_refl] + by_cases hk0 : k = 0 + rw [hk0] at * + simp [ Real.rpow_zero, map_one, inv_one, mul_one, le_refl] rw [hx.1, hx.2] - have h1 : (0 : ℝ) ^ (k : ℕ) = 0 := by - simp only [pow_eq_zero_iff', ne_eq, true_and] - exact hk + simp only [Int.cast_zero, zero_mul, add_zero, map_zero] + have h1 : (0 : ℝ) ^ (-k) = 0 := by + rw [Real.rpow_eq_zero_iff_of_nonneg (by rfl)] + simp only [ne_eq, neg_eq_zero, hk0, not_false_eq_true, and_self] simp only [Int.cast_zero, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, CharP.cast_eq_zero, map_mul, abs_ofReal, mul_zero, le_refl] · have hx2 : x ≠ 0 := by rw [mem_box_ne_zero_iff_ne_zero n x hx] exact hn simp only [Int.mem_box] at hx - rw [inv_le_inv] + rw [Real.rpow_neg (by apply apply_nonneg), Real.rpow_neg ((r_pos z).le), + Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv] simp only [← hx, map_mul, map_pow, abs_ofReal, abs_natCast, Nat.cast_max, ge_iff_le] - convert (bound z x hx2 k) - apply abs_eq_self.mpr ((r_pos z).le ) - simp only [Nat.cast_max] - simp only [map_pow] - apply Complex.abs.pos (pow_ne_zero k (linear_ne_zero ![x 0, x 1] z ?_)) - simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, - not_and] at * - intro hg h1 - have : x = ![x 0, x 1] := by - exact List.ofFn_inj.mp rfl - rw [this, hg,h1 ] at hx2 - simp only [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at * - have := r_pos z - apply Complex.abs.pos - norm_cast - positivity + convert (rpow_bound hk z x hx2) + · simp only [Nat.cast_max] + · apply Real.rpow_pos_of_pos + apply Complex.abs.pos ( (linear_ne_zero ![x 0, x 1] z ?_)) + simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, + not_and] at * + intro hg h1 + have : x = ![x 0, x 1] := by + exact List.ofFn_inj.mp rfl + rw [this, hg, h1] at hx2 + simp [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at * + · apply mul_pos (Real.rpow_pos_of_pos (r_pos z) _) + apply Real.rpow_pos_of_pos + norm_cast + exact Nat.pos_of_ne_zero hn + +theorem eis_is_bounded_on_box (k n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) + (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ + (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by + have := eis_is_bounded_on_box_rpow (Nat.cast_nonneg k) z n x hx + norm_cast at * + simp_rw [zpow_neg, ← mul_inv] at this + have H : |r z ^ k * ↑(n ^ k)| = r z ^ k * ↑(n ^ k) := by + refine abs_eq_self.mpr ?_ + apply mul_nonneg (pow_nonneg (r_pos z).le _) + simp only [Nat.cast_pow, ge_iff_le, Nat.cast_nonneg, pow_nonneg] + simp only [map_pow, zpow_coe_nat, H] + simpa using this lemma r_lower_bound_on_slice {A B : ℝ} (h : 0 < B) (z : upperHalfPlaneSlice A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by @@ -334,7 +340,7 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : simp only [zpow_coe_nat, one_div, Function.comp_apply] apply tendstoUniformlyOn_tsum hu intro v x hx - have := eis_is_bounded_on_box k x (max (v.1 0).natAbs (v.1 1).natAbs ) v + have := eis_is_bounded_on_box k (max (v.1 0).natAbs (v.1 1).natAbs ) x v simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_coe_nat, one_div, map_pow, map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, ge_iff_le] at * From 27421d7e9bca653ddc14251b5a160cff059c8c7e Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 16 Feb 2024 17:59:02 +0000 Subject: [PATCH 056/124] style fix --- .../EisensteinSeries/UniformConvergence.lean | 28 +++++++++---------- 1 file changed, 14 insertions(+), 14 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 688d8135d0c45..8121de1f167de 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -13,7 +13,7 @@ import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic /-! # Uniform convergence of Eisenstein series -We show that `eis` converges locally uniformly on `ℍ` to the Eisenstein series `E` of weight `k` +We show that `eisSummand` converges locally uniformly on `ℍ` to the Eisenstein series of weight `k` and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. -/ @@ -30,7 +30,7 @@ lemma mem_box_ne_zero_iff_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) rw [fun_ne_zero_cases, hx.1, hx.2] at h · simp at h rintro hn rfl - simp [hn, eq_comm] at hx + simp only [Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self, eq_comm, hn] at hx noncomputable section @@ -60,7 +60,7 @@ theorem r_pos (z : ℍ) : 0 < r z := by simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, r1_pos, and_self] theorem r1_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : - (z.im ^ 2 ) / (z.re ^ 2 + z.im ^ 2) ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by + (z.im ^ 2) / (z.re ^ 2 + z.im ^ 2) ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by have H1 : (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 = δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2 := by ring have H4 : (δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2) * (z.re ^ 2 + z.im ^ 2) @@ -91,7 +91,7 @@ theorem auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ C theorem auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs, min_le_iff] - have H1 : Real.sqrt (r1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε ) + + have H1 : Real.sqrt (r1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by rw [r1, Real.sqrt_le_sqrt_iff, ← pow_two, ← pow_two] apply r1_bound z δ hε @@ -127,7 +127,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : rw [H1, div_pow, Int.cast_natAbs (x 0),Int.cast_abs] have : (x 0 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max hx H1) - have h1 : (x 0 : ℝ)^2/( _root_.abs (x 0 : ℝ))^2 = 1 := by + have h1 : (x 0 : ℝ)^2/(_root_.abs (x 0 : ℝ))^2 = 1 := by simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, this, div_self] exact h1.symm.le @@ -135,7 +135,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : rw [H2,div_pow, Int.cast_natAbs (x 1),Int.cast_abs] have : (x 1 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max' hx H2) - have h1 : (x 1 : ℝ)^2/( _root_.abs (x 1 : ℝ))^2 = 1 := by + have h1 : (x 1 : ℝ)^2/(_root_.abs (x 1 : ℝ))^2 = 1 := by simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, this, div_self] exact h1.symm.le @@ -146,7 +146,7 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x by_cases hk0 : k ≠ 0 let n := max (x 0).natAbs (x 1).natAbs have h11 : ((x 0) * ↑z + (x 1)) = - (((x 0 : ℝ) / (n : ℝ) ) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by + (((x 0 : ℝ) / (n : ℝ)) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by field_simp rw [← mul_div, div_self] simp @@ -175,7 +175,7 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)))) ^ (-k) ≤ - ( ((r z) ^ (-k) * n ^ (-k))) := by + (((r z) ^ (-k) * n ^ (-k))) := by by_cases hn : n = 0 · rw [hn] at hx simp only [box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx @@ -199,7 +199,7 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) convert (rpow_bound hk z x hx2) · simp only [Nat.cast_max] · apply Real.rpow_pos_of_pos - apply Complex.abs.pos ( (linear_ne_zero ![x 0, x 1] z ?_)) + apply Complex.abs.pos (linear_ne_zero ![x 0, x 1] z ?_) simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, not_and] at * intro hg h1 @@ -296,7 +296,7 @@ lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): norm_cast rw [zero_zpow k (by linarith), zero_zpow (k - 1) (by linarith)] simp only [inv_zero, mul_zero, box_zero, Finset.card_singleton, Nat.cast_one] - · rw [Int.card_box b0, zpow_sub_one₀ (a:= ( b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] + · rw [Int.card_box b0, zpow_sub_one₀ (a:= (b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] ring_nf @@ -324,7 +324,7 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : end summability theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : ℕ) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => + (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a)) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by have hk0 : 0 ≤ k := by linarith @@ -333,14 +333,14 @@ theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] intro K hK obtain ⟨A, B, hB, HABK⟩:= subset_slice_of_isCompact hK - have hu : Summable fun x : (gammaSet N a ) => + have hu : Summable fun x : (gammaSet N a) => (1/(r ⟨⟨A, B⟩, hB⟩) ^ k) * ((max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by - apply (Summable.subtype (summable_upper_bound hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a )).congr + apply (Summable.subtype (summable_upper_bound hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a)).congr intro v simp only [zpow_coe_nat, one_div, Function.comp_apply] apply tendstoUniformlyOn_tsum hu intro v x hx - have := eis_is_bounded_on_box k (max (v.1 0).natAbs (v.1 1).natAbs ) x v + have := eis_is_bounded_on_box k (max (v.1 0).natAbs (v.1 1).natAbs) x v simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_coe_nat, one_div, map_pow, map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, ge_iff_le] at * From 22374fc36e87b19d329d9ddb2ab6c1bfd756f4b2 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 16 Feb 2024 18:01:12 +0000 Subject: [PATCH 057/124] remove unused lemma --- Mathlib/Data/Complex/Abs.lean | 7 ------- 1 file changed, 7 deletions(-) diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 0779bb56a50da..a0e4105602899 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -92,13 +92,6 @@ theorem abs_ofNat (n : ℕ) [n.AtLeastTwo] : Complex.abs (no_index (OfNat.ofNat n : ℂ)) = OfNat.ofNat n := abs_natCast n -lemma complex_abs_of_int_eq_natAbs (a : ℤ) : Complex.abs a = a.natAbs := by - have : Complex.abs a = Complex.abs (a : ℝ) := by rfl - rw [this, Complex.abs_ofReal] - norm_cast - rw [Int.abs_eq_natAbs a] - rfl - theorem mul_self_abs (z : ℂ) : Complex.abs z * Complex.abs z = normSq z := Real.mul_self_sqrt (normSq_nonneg _) #align complex.mul_self_abs Complex.mul_self_abs From ebcf5beca030a17c4303e97bc430f4c8d8df5dfa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 20 Feb 2024 09:36:57 +0000 Subject: [PATCH 058/124] cleanup --- Mathlib/Data/Finset/Basic.lean | 12 ---- .../EisensteinSeries/UniformConvergence.lean | 62 ++++++++----------- 2 files changed, 26 insertions(+), 48 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 23aa1b7b58b49..841527432279f 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -3974,18 +3974,6 @@ def piFinsetUnion {ι} [DecidableEq ι] (α : ι → Type*) {s t : Finset ι} (h let e := Equiv.Finset.union s t h sumPiEquivProdPi (fun b ↦ α (e b)) |>.symm.trans (.piCongrLeft (fun i : ↥(s ∪ t) ↦ α i) e) -/--Equivalence between the sigma of a family of finsets of `β × γ` and `β × γ`-/ -noncomputable -def sigmaEquiv {α β γ : Type*} (ι : α → Finset (β × γ)) - (HI : ∀ y : β × γ , ∃! i : α, y ∈ ι i) : (Σ s : α, ((ι s) : Set (β × γ))) ≃ (β × γ) where - toFun x := x.2 - invFun x := ⟨(HI x).choose, x, (HI x).choose_spec.1⟩ - left_inv x := by - ext - exact ((HI x.2).choose_spec.2 x.1 x.2.2).symm - repeat {rfl} - right_inv x := by rfl - end Equiv namespace Multiset diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 8121de1f167de..883559715232c 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -249,29 +249,21 @@ lemma r_lower_bound_on_slice {A B : ℝ} (h : 0 < B) (z : upperHalfPlaneSlice A end bounding_functions section summability - -lemma summable_lemma (f : (Fin 2 → ℤ) → ℝ) (h : ∀ y : (Fin 2 → ℤ), 0 ≤ f y) - (ι : ℕ → Finset (ℤ × ℤ)) (HI : ∀ y : ℤ × ℤ, ∃! i : ℕ, y ∈ ι i) : - Summable f ↔ Summable fun n : ℕ => ∑ x in ι n, f ![x.1, x.2] := by - let h2 := Equiv.trans (Equiv.sigmaEquiv ι HI) (piFinTwoEquiv fun _ => ℤ).symm - have h22 : ∀ y : Σ s : ℕ, (ι s), 0 ≤ (f ∘ h2) y := by - intro y - apply h - have h4 : Summable f ↔ Summable (f ∘ h2) := by rw [Equiv.summable_iff] - rw [h4, summable_sigma_of_nonneg h22] - constructor - · intro H - convert H.2 - rw [← Finset.tsum_subtype] - rfl - · intro H - constructor - · intro x - simp only [Finset.coe_sort_coe, Equiv.coe_trans, Function.comp_apply, Equiv.sigmaEquiv] - convert (Finset.summable (ι x) (f ∘ (piFinTwoEquiv fun _ => ℤ).symm)) - · convert H - rw [← Finset.tsum_subtype] - rfl +variable {ι κ α : Type*} [AddCommMonoid α] [TopologicalSpace α] + +/-- Equivalence between the sigma of a family of finsets of `β` and `β`. -/ +noncomputable def sigmaEquiv {ι κ : Type*} (s : κ → Set ι) (hs : ∀ i, ∃! j, i ∈ s j) : + (Σ j, s j) ≃ ι where + toFun x := x.2 + invFun x := ⟨(hs x).choose, x, (hs x).choose_spec.1⟩ + left_inv x := by ext; exacts [((hs x.2).choose_spec.2 x.1 x.2.2).symm, rfl] + right_inv x := by rfl + +lemma summable_partition {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} (hs : ∀ i, ∃! j, i ∈ s j) : + Summable f ↔ (∀ j, Summable fun i : s j ↦ f i) ∧ Summable fun j ↦ ∑' i : s j, f i := by + rw [← (sigmaEquiv s hs).summable_iff, summable_sigma_of_nonneg] + simp only [coe_sort_coe, sigmaEquiv, Equiv.coe_fn_mk, Function.comp_apply] + exact fun _ ↦ hf _ lemma summable_r_pow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by @@ -302,19 +294,17 @@ lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by - rw [summable_lemma _ _ (fun (n : ℕ) => box n) Int.existsUnique_mem_box] - · have : ∀ n : ℕ, ∑ v in (box n : Finset (ℤ × ℤ)), - (1 / (r z) ^ k) * ((max v.1.natAbs v.2.natAbs: ℝ) ^ k)⁻¹ = - ∑ v in box n, (1 / (r z) ^ k) * ((n : ℝ)^k)⁻¹ := by - intro n - apply Finset.sum_congr rfl - intro x hx - simp only [Int.mem_box] at hx - congr - norm_cast - apply Summable.congr (summable_over_box z h) - intro b - apply (this b).symm + set f := fun x : Fin 2 → ℤ ↦ (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ + rw [← (piFinTwoEquiv _).symm.summable_iff, + summable_partition _ (s := fun n ↦ (box n : Finset (ℤ × ℤ))) Int.existsUnique_mem_box] + · simp_rw [coe_sort_coe, Finset.tsum_subtype] + simp only [one_div, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one] + refine ⟨fun n ↦ ?_, Summable.congr (summable_over_box z h) fun n ↦ Finset.sum_congr rfl + fun x hx ↦ ?_⟩ + · simpa using (box n).summable (f ∘ (piFinTwoEquiv _).symm) + simp only [Int.mem_box] at hx + rw [← hx, one_div] + norm_cast · intro y apply mul_nonneg · simp only [one_div, inv_nonneg] From 4c85191add8983509ecbf3855b8394ee7cbc528f Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 23 Feb 2024 11:08:01 +0000 Subject: [PATCH 059/124] small golf --- .../EisensteinSeries/UniformConvergence.lean | 12 ++++-------- 1 file changed, 4 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 883559715232c..142c2aabdedbe 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -152,8 +152,7 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x simp rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx norm_cast at * - rw [h11] - simp only [Nat.cast_max, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] + simp only [h11, Nat.cast_max, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] cases' (div_max_sq_ge_one x hx) with H1 H2 · apply mul_le_mul @@ -267,9 +266,7 @@ lemma summable_partition {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} (hs lemma summable_r_pow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by - have hk : 1 < (k - 1 : ℝ) := by - norm_cast at * - linarith + have hk : 1 < (k - 1 : ℝ) := by norm_cast; linarith have nze : (8 / (r z) ^ k : ℝ) ≠ 0 := by apply div_ne_zero simp only [ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true] @@ -284,8 +281,7 @@ lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): apply Summable.congr (summable_r_pow z h) intro b by_cases b0 : b = 0 - · rw [b0] - norm_cast + · simp only [b0, CharP.cast_eq_zero, box_zero, Finset.card_singleton, Nat.cast_one, one_mul] rw [zero_zpow k (by linarith), zero_zpow (k - 1) (by linarith)] simp only [inv_zero, mul_zero, box_zero, Finset.card_singleton, Nat.cast_one] · rw [Int.card_box b0, zpow_sub_one₀ (a:= (b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] @@ -298,7 +294,7 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : rw [← (piFinTwoEquiv _).symm.summable_iff, summable_partition _ (s := fun n ↦ (box n : Finset (ℤ × ℤ))) Int.existsUnique_mem_box] · simp_rw [coe_sort_coe, Finset.tsum_subtype] - simp only [one_div, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one] + simp only [coe_sort_coe, Finset.tsum_subtype, one_div, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one] refine ⟨fun n ↦ ?_, Summable.congr (summable_over_box z h) fun n ↦ Finset.sum_congr rfl fun x hx ↦ ?_⟩ · simpa using (box n).summable (f ∘ (piFinTwoEquiv _).symm) From b8ccdadc3e1409baf967872a36316346267179f0 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 23 Feb 2024 11:43:21 +0000 Subject: [PATCH 060/124] lint fix and re-squeeze some simps --- .../EisensteinSeries/UniformConvergence.lean | 24 +++++++++---------- 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 142c2aabdedbe..df3308cc4b44b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -152,7 +152,7 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x simp rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx norm_cast at * - simp only [h11, Nat.cast_max, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] + simp only [Nat.cast_max, h11, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] cases' (div_max_sq_ge_one x hx) with H1 H2 · apply mul_le_mul @@ -170,7 +170,7 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x simp only [le_max_iff, Nat.cast_nonneg, or_self] apply Real.rpow_nonneg (Complex.abs.nonneg _) k · simp only [ne_eq, not_not] at hk0 - simp only [hk0, Real.rpow_zero, Nat.cast_max, mul_one, map_one, le_refl] + simp only [hk0, Real.rpow_zero, Nat.cast_max, mul_one, le_refl] theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)))) ^ (-k) ≤ @@ -180,21 +180,20 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) simp only [box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx by_cases hk0 : k = 0 rw [hk0] at * - simp [ Real.rpow_zero, map_one, inv_one, mul_one, le_refl] + simp only [neg_zero, Real.rpow_zero, mul_one, le_refl] rw [hx.1, hx.2] simp only [Int.cast_zero, zero_mul, add_zero, map_zero] have h1 : (0 : ℝ) ^ (-k) = 0 := by rw [Real.rpow_eq_zero_iff_of_nonneg (by rfl)] simp only [ne_eq, neg_eq_zero, hk0, not_false_eq_true, and_self] - simp only [Int.cast_zero, zero_mul, add_zero, map_pow, map_zero, h1, inv_zero, hn, - CharP.cast_eq_zero, map_mul, abs_ofReal, mul_zero, le_refl] + simp only [h1, hn, CharP.cast_eq_zero, mul_zero, le_refl] · have hx2 : x ≠ 0 := by rw [mem_box_ne_zero_iff_ne_zero n x hx] exact hn simp only [Int.mem_box] at hx rw [Real.rpow_neg (by apply apply_nonneg), Real.rpow_neg ((r_pos z).le), Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv] - simp only [← hx, map_mul, map_pow, abs_ofReal, abs_natCast, Nat.cast_max, ge_iff_le] + simp only [← hx, Nat.cast_max] convert (rpow_bound hk z x hx2) · simp only [Nat.cast_max] · apply Real.rpow_pos_of_pos @@ -205,7 +204,7 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) have : x = ![x 0, x 1] := by exact List.ofFn_inj.mp rfl rw [this, hg, h1] at hx2 - simp [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at * + simp only [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at * · apply mul_pos (Real.rpow_pos_of_pos (r_pos z) _) apply Real.rpow_pos_of_pos norm_cast @@ -261,7 +260,7 @@ noncomputable def sigmaEquiv {ι κ : Type*} (s : κ → Set ι) (hs : ∀ i, lemma summable_partition {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} (hs : ∀ i, ∃! j, i ∈ s j) : Summable f ↔ (∀ j, Summable fun i : s j ↦ f i) ∧ Summable fun j ↦ ∑' i : s j, f i := by rw [← (sigmaEquiv s hs).summable_iff, summable_sigma_of_nonneg] - simp only [coe_sort_coe, sigmaEquiv, Equiv.coe_fn_mk, Function.comp_apply] + simp only [sigmaEquiv, Equiv.coe_fn_mk, Function.comp_apply] exact fun _ ↦ hf _ lemma summable_r_pow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : @@ -277,13 +276,13 @@ lemma summable_r_pow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): Summable (fun n : ℕ => ∑ v in (box n : Finset (ℤ × ℤ)), (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by - simp only [one_div, Finset.sum_const, nsmul_eq_mul] + simp only [one_div, sum_const, nsmul_eq_mul] apply Summable.congr (summable_r_pow z h) intro b by_cases b0 : b = 0 · simp only [b0, CharP.cast_eq_zero, box_zero, Finset.card_singleton, Nat.cast_one, one_mul] rw [zero_zpow k (by linarith), zero_zpow (k - 1) (by linarith)] - simp only [inv_zero, mul_zero, box_zero, Finset.card_singleton, Nat.cast_one] + simp only [inv_zero, mul_zero] · rw [Int.card_box b0, zpow_sub_one₀ (a:= (b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] ring_nf @@ -294,7 +293,7 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : rw [← (piFinTwoEquiv _).symm.summable_iff, summable_partition _ (s := fun n ↦ (box n : Finset (ℤ × ℤ))) Int.existsUnique_mem_box] · simp_rw [coe_sort_coe, Finset.tsum_subtype] - simp only [coe_sort_coe, Finset.tsum_subtype, one_div, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one] + simp only [one_div, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one] refine ⟨fun n ↦ ?_, Summable.congr (summable_over_box z h) fun n ↦ Finset.sum_congr rfl fun x hx ↦ ?_⟩ · simpa using (box n).summable (f ∘ (piFinTwoEquiv _).symm) @@ -305,7 +304,8 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : apply mul_nonneg · simp only [one_div, inv_nonneg] apply zpow_nonneg (r_pos z).le - · simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] + · simp only [piFinTwoEquiv_symm_apply, Fin.cons_zero, Fin.cons_one, inv_nonneg, ge_iff_le, + le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] end summability From 87297e89cd49b16ac23cf159014a8023440b95fc Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 23 Feb 2024 11:58:24 +0000 Subject: [PATCH 061/124] 2 --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index df3308cc4b44b..a808b086d7447 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -310,12 +310,13 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : end summability theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : ℕ) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a)) => + (a : Fin 2 → ZMod N) : TendstoLocallyUniformly (fun (s : Finset (gammaSet N a)) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) - (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop ⊤ := by + (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop := by have hk0 : 0 ≤ k := by linarith lift k to ℕ using hk0 - rw [tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] + rw [←tendstoLocallyUniformlyOn_univ,tendstoLocallyUniformlyOn_iff_forall_isCompact, + eisensteinSeries_SIF] simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] intro K hK obtain ⟨A, B, hB, HABK⟩:= subset_slice_of_isCompact hK From 2ed727de4e62de4562e5b55e6aaf2ecbc834d220 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 23 Feb 2024 11:59:07 +0000 Subject: [PATCH 062/124] fix last thm --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index a808b086d7447..07e955b543633 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -183,7 +183,7 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) simp only [neg_zero, Real.rpow_zero, mul_one, le_refl] rw [hx.1, hx.2] simp only [Int.cast_zero, zero_mul, add_zero, map_zero] - have h1 : (0 : ℝ) ^ (-k) = 0 := by + have h1 : (0 : ℝ) ^ (-k) = 0 := by rw [Real.rpow_eq_zero_iff_of_nonneg (by rfl)] simp only [ne_eq, neg_eq_zero, hk0, not_false_eq_true, and_self] simp only [h1, hn, CharP.cast_eq_zero, mul_zero, le_refl] From 9e4428b6bb54abbe7467ab757dc66a47ea2a97d5 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 23 Feb 2024 11:59:41 +0000 Subject: [PATCH 063/124] fix last thm name --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 07e955b543633..397c2de9b99a2 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -309,7 +309,7 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : end summability -theorem eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : ℕ) +theorem eisensteinSeries_TendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformly (fun (s : Finset (gammaSet N a)) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop := by From 33f9ee36b1795e08d47536f5f8c447b7ef2e3468 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 23 Feb 2024 12:00:55 +0000 Subject: [PATCH 064/124] lint fix --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 397c2de9b99a2..bc3b3061966ec 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -315,7 +315,7 @@ theorem eisensteinSeries_TendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop := by have hk0 : 0 ≤ k := by linarith lift k to ℕ using hk0 - rw [←tendstoLocallyUniformlyOn_univ,tendstoLocallyUniformlyOn_iff_forall_isCompact, + rw [← tendstoLocallyUniformlyOn_univ,tendstoLocallyUniformlyOn_iff_forall_isCompact, eisensteinSeries_SIF] simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] intro K hK From b8878455788453634a73adcdbb8f3c376d926c08 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 23 Feb 2024 15:44:11 +0000 Subject: [PATCH 065/124] generalise fun_ne_zero_cases --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index bc3b3061966ec..52c8b7c756531 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -19,7 +19,7 @@ and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. open Finset -lemma fun_ne_zero_cases (x : Fin 2 → ℤ) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by +lemma fun_ne_zero_cases {G : Type*} [OfNat G 0] (x : Fin 2 → G) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by rw [Function.ne_iff]; exact Fin.exists_fin_two lemma mem_box_ne_zero_iff_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : From 42b6b8adddd36738a85bf68ceea50c9c16c83afd Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 27 Feb 2024 15:05:04 +0000 Subject: [PATCH 066/124] add a complex version of main thm --- .../EisensteinSeries/UniformConvergence.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 52c8b7c756531..f952ba45bb204 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -344,3 +344,17 @@ theorem eisensteinSeries_TendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : · apply pow_pos (r_pos _) · simp only [Int.mem_box] · simp only [Set.top_eq_univ, isOpen_univ] + +local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm + (OpenEmbedding.toPartialHomeomorph UpperHalfPlane.coe openEmbedding_coe)) + +/- A version for the extension to maps `ℂ → ℂ` that is nice to have for holomorphicity later -/ +lemma eisensteinSeries_TendstoLocallyUniformlyOn3 {k : ℤ} (hk : 3 ≤ k) (N : ℕ) + (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => + ↑ₕ(fun (z : ℍ) => ∑ x in s, eisSummand k x z )) (↑ₕ((eisensteinSeries_SIF a k).toFun )) + Filter.atTop (UpperHalfPlane.coe '' ⊤) := by + apply TendstoLocallyUniformlyOn.comp (s := ⊤) + simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ] + apply eisensteinSeries_TendstoLocallyUniformly hk + simp only [Set.top_eq_univ, image_univ, mapsTo_range_iff, Set.mem_univ, forall_const] + apply PartialHomeomorph.continuousOn_symm From 96e63cc89acabb850834be1786c40a27aae7681f Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 27 Feb 2024 15:24:52 +0000 Subject: [PATCH 067/124] fix name --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index f952ba45bb204..809ba38273f26 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -349,7 +349,7 @@ local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm (OpenEmbedding.toPartialHomeomorph UpperHalfPlane.coe openEmbedding_coe)) /- A version for the extension to maps `ℂ → ℂ` that is nice to have for holomorphicity later -/ -lemma eisensteinSeries_TendstoLocallyUniformlyOn3 {k : ℤ} (hk : 3 ≤ k) (N : ℕ) +lemma eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => ↑ₕ(fun (z : ℍ) => ∑ x in s, eisSummand k x z )) (↑ₕ((eisensteinSeries_SIF a k).toFun )) Filter.atTop (UpperHalfPlane.coe '' ⊤) := by From dcd880a14a19cd2f1d5fcd009a06e09e76e8a2fb Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 1 Mar 2024 15:19:23 +0000 Subject: [PATCH 068/124] rev updates --- .../Complex/UpperHalfPlane/Basic.lean | 9 ++++ .../Complex/UpperHalfPlane/Metric.lean | 46 +++++++++---------- Mathlib/Data/Finset/LocallyFinite/Box.lean | 13 ++++++ .../EisensteinSeries/UniformConvergence.lean | 44 ++++++------------ 4 files changed, 58 insertions(+), 54 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index abc8afe3bfc41..fb6db0db6727c 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -137,6 +137,15 @@ theorem ne_zero (z : ℍ) : (z : ℂ) ≠ 0 := mt (congr_arg Complex.im) z.im_ne_zero #align upper_half_plane.ne_zero UpperHalfPlane.ne_zero +/-- Define √-1 as an element on the upper half plane.-/ +def I : ℍ := ⟨Complex.I, by simp⟩ + +@[simp] +lemma I_im : I.im = 1 := rfl + +@[simp] +lemma I_re : I.re = 0 := rfl + end UpperHalfPlane namespace Mathlib.Meta.Positivity diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index 6cea2fc629c43..fd5c3b3e469b7 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -386,44 +386,40 @@ instance : IsometricSMul SL(2, ℝ) ℍ := section slices -/--The verticel strip of width A and height B-/ +/--The vertical strip of width A and height B-/ def upperHalfPlaneSlice (A B : ℝ) := - {z : ℍ | Complex.abs z.1.1 ≤ A ∧ Complex.abs z.1.2 ≥ B} + {z : ℍ | Complex.abs z.1.1 ≤ A ∧ B ≤ Complex.abs z.1.2} theorem slice_mem_iff (A B : ℝ) (z : ℍ) : - z ∈ upperHalfPlaneSlice A B ↔ Complex.abs z.1.1 ≤ A ∧ Complex.abs z.1.2 ≥ B := Iff.rfl - -lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : ∃ A B : ℝ, 0 < B ∧ - K ⊆ upperHalfPlaneSlice A B := by - by_cases hne : Set.Nonempty K - · have hcts : ContinuousOn (fun t => t.im) K := by - apply Continuous.continuousOn UpperHalfPlane.continuous_im - obtain ⟨b, _, HB⟩ := IsCompact.exists_isMinOn hK hne hcts - let t := (⟨Complex.I, by simp⟩ : ℍ) - have ht : UpperHalfPlane.im t = I.im := by rfl - obtain ⟨r, _, hr2⟩ := Bornology.IsBounded.subset_closedBall_lt hK.isBounded 0 t - refine' ⟨Real.sinh (r) + Complex.abs ((UpperHalfPlane.center t r)), b.im, b.2, _⟩ - intro z hz - simp only [I_im, slice_mem_iff, abs_ofReal, ge_iff_le] at * - constructor - have hr3 := hr2 hz + z ∈ upperHalfPlaneSlice A B ↔ Complex.abs z.1.1 ≤ A ∧ B ≤ Complex.abs z.1.2 := Iff.rfl + +lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : ∃ A B : ℝ, 0 < B ∧ + K ⊆ upperHalfPlaneSlice A B := by + obtain rfl | hne := K.eq_empty_or_nonempty + exact ⟨1, 1, Real.zero_lt_one, by simp⟩ + have hcts : ContinuousOn (fun t => t.im) K := by + apply Continuous.continuousOn UpperHalfPlane.continuous_im + obtain ⟨b, _, HB⟩ := IsCompact.exists_isMinOn hK hne hcts + obtain ⟨r, _, hr2⟩ := Bornology.IsBounded.subset_closedBall_lt hK.isBounded 0 UpperHalfPlane.I + refine' ⟨Real.sinh r + Complex.abs ((UpperHalfPlane.center UpperHalfPlane.I r)), b.im, b.2, _⟩ + intro z hz + simp only [I_im, slice_mem_iff, abs_ofReal, ge_iff_le] at * + constructor + · have hr3 := hr2 hz simp only [Metric.mem_closedBall] at hr3 apply le_trans (abs_re_le_abs z) - have := Complex.abs.sub_le (z : ℂ) (UpperHalfPlane.center t r) 0 + have := Complex.abs.sub_le (z : ℂ) (UpperHalfPlane.center UpperHalfPlane.I r) 0 simp only [sub_zero, ge_iff_le] at this rw [dist_le_iff_dist_coe_center_le] at hr3 apply le_trans this - have htim : UpperHalfPlane.im t = 1 := by - simp only [ht] - rw [htim] at hr3 - simp only [one_mul, add_le_add_iff_right, ge_iff_le] at * + simp only [I_im, one_mul, add_le_add_iff_right] at * exact hr3 - have hbz := HB hz + · have hbz := HB hz simp only [mem_setOf_eq, ge_iff_le] at * convert hbz rw [UpperHalfPlane.im] apply abs_eq_self.mpr z.2.le - · exact ⟨1, 1, Real.zero_lt_one, by simp [not_nonempty_iff_eq_empty.1 hne]⟩ + end slices diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean index bf5aad6ff1bc5..d91ba1a795314 100644 --- a/Mathlib/Data/Finset/LocallyFinite/Box.lean +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -90,3 +90,16 @@ lemma existsUnique_mem_box (x : ℤ × ℤ) : ∃! n : ℕ, x ∈ box n := by use max x.1.natAbs x.2.natAbs; simp only [mem_box, and_self_iff, forall_eq'] end Int + +lemma Finset.fun_ne_zero_cases {G : Type*} [OfNat G 0] (x : Fin 2 → G) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by + rw [Function.ne_iff]; exact Fin.exists_fin_two + +lemma Finset.mem_box_ne_zero_iff_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : + x ≠ 0 ↔ n ≠ 0 := by + constructor + intro h h0 + simp only [h0, Nat.cast_zero, box_zero, mem_singleton, Prod.ext_iff] at hx + rw [fun_ne_zero_cases, hx.1, hx.2] at h + · simp at h + rintro hn rfl + simp only [Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self, eq_comm, hn] at hx diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 809ba38273f26..12513455b81e9 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -17,24 +17,10 @@ We show that `eisSummand` converges locally uniformly on `ℍ` to the Eisenstein and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. -/ -open Finset - -lemma fun_ne_zero_cases {G : Type*} [OfNat G 0] (x : Fin 2 → G) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by - rw [Function.ne_iff]; exact Fin.exists_fin_two - -lemma mem_box_ne_zero_iff_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : - x ≠ 0 ↔ n ≠ 0 := by - constructor - intro h h0 - simp only [h0, Nat.cast_zero, box_zero, mem_singleton, Prod.ext_iff] at hx - rw [fun_ne_zero_cases, hx.1, hx.2] at h - · simp at h - rintro hn rfl - simp only [Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self, eq_comm, hn] at hx noncomputable section -open Complex Filter UpperHalfPlane Set +open Complex Filter UpperHalfPlane Set Finset open scoped BigOperators NNReal Classical Filter Matrix UpperHalfPlane Complex @@ -42,14 +28,14 @@ namespace EisensteinSeries section bounding_functions -/-- Auxilary function used for bounding Eisentein series-/ -def r1 (z : ℍ) : ℝ := ((z.im ^ (2 : ℕ)) / (z.re ^ (2 : ℕ) + z.im ^ (2 : ℕ))) +/-- Auxiliary function used for bounding Eisenstein series-/ +def r1 (z : ℍ) : ℝ := ((z.im ^ 2) / (z.re ^ 2 + z.im ^ 2)) -lemma r1' (z : ℍ) : r1 z = 1/((z.re / z.im) ^ (2 : ℕ) + 1) := by +lemma r1' (z : ℍ) : r1 z = 1/((z.re / z.im) ^ 2 + 1) := by field_simp [r1, im_pos z] theorem r1_pos (z : ℍ) : 0 < r1 z := by - have H2 : 0 < (z.re ^ (2 : ℕ) + z.im ^ (2 : ℕ)) := by + have H2 : 0 < (z.re ^ 2 + z.im ^ 2) := by apply_rules [pow_pos, add_pos_of_nonneg_of_pos, pow_two_nonneg, z.2] exact div_pos (pow_pos z.im_pos 2) H2 @@ -63,13 +49,13 @@ theorem r1_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : (z.im ^ 2) / (z.re ^ 2 + z.im ^ 2) ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by have H1 : (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 = δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2 := by ring - have H4 : (δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2) * (z.re ^ 2 + z.im ^ 2) + have H2 : (δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2) * (z.re ^ 2 + z.im ^ 2) - (z.im ^ 2) = (δ * (z.re ^ 2 + z.im ^ 2)+ ε * z.re)^2 + (ε^2 - 1)* (z.im)^2 := by ring - rw [H1, div_le_iff, ← sub_nonneg, H4] + rw [H1, div_le_iff, ← sub_nonneg, H2] · apply add_nonneg (pow_two_nonneg _) ?_ apply mul_nonneg - linarith - apply pow_two_nonneg + · linarith + · apply pow_two_nonneg · apply_rules [add_pos_of_nonneg_of_pos, pow_two_nonneg, (pow_pos z.im_pos 2)] theorem auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by @@ -113,7 +99,7 @@ lemma ne_zero_if_max {x : Fin 2 → ℤ} (hx : x ≠ 0) lemma ne_zero_if_max' {x : Fin 2 → ℤ} (hx : x ≠ 0) (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by - apply ne_zero_if_max (x :=![x 1, x 0]) ?_ (by simpa using h) + apply ne_zero_if_max (x := ![x 1, x 0]) ?_ (by simpa using h) simp only [ne_eq, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_true, not_and] intro h1 h0 rw [fun_ne_zero_cases, h1, h0] at hx @@ -124,7 +110,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 := by cases' (max_choice (x 0).natAbs (x 1).natAbs) with H1 H2 · left - rw [H1, div_pow, Int.cast_natAbs (x 0),Int.cast_abs] + rw [H1, div_pow, Int.cast_natAbs (x 0), Int.cast_abs] have : (x 0 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max hx H1) have h1 : (x 0 : ℝ)^2/(_root_.abs (x 0 : ℝ))^2 = 1 := by @@ -132,7 +118,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : this, div_self] exact h1.symm.le · right - rw [H2,div_pow, Int.cast_natAbs (x 1),Int.cast_abs] + rw [H2,div_pow, Int.cast_natAbs (x 1), Int.cast_abs] have : (x 1 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max' hx H2) have h1 : (x 1 : ℝ)^2/(_root_.abs (x 1 : ℝ))^2 = 1 := by @@ -309,7 +295,7 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : end summability -theorem eisensteinSeries_TendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : ℕ) +theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformly (fun (s : Finset (gammaSet N a)) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop := by @@ -349,12 +335,12 @@ local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm (OpenEmbedding.toPartialHomeomorph UpperHalfPlane.coe openEmbedding_coe)) /- A version for the extension to maps `ℂ → ℂ` that is nice to have for holomorphicity later -/ -lemma eisensteinSeries_TendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : ℕ) +lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : ℕ) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => ↑ₕ(fun (z : ℍ) => ∑ x in s, eisSummand k x z )) (↑ₕ((eisensteinSeries_SIF a k).toFun )) Filter.atTop (UpperHalfPlane.coe '' ⊤) := by apply TendstoLocallyUniformlyOn.comp (s := ⊤) simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ] - apply eisensteinSeries_TendstoLocallyUniformly hk + apply eisensteinSeries_tendstoLocallyUniformly hk simp only [Set.top_eq_univ, image_univ, mapsTo_range_iff, Set.mem_univ, forall_const] apply PartialHomeomorph.continuousOn_symm From 3a09ccc60d293ea12482c44933ff933e7e050912 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 1 Mar 2024 15:45:40 +0000 Subject: [PATCH 069/124] make arg implicit --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 12513455b81e9..1590d97a060d8 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -335,7 +335,7 @@ local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm (OpenEmbedding.toPartialHomeomorph UpperHalfPlane.coe openEmbedding_coe)) /- A version for the extension to maps `ℂ → ℂ` that is nice to have for holomorphicity later -/ -lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} (hk : 3 ≤ k) (N : ℕ) +lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} {N : ℕ} (hk : 3 ≤ k) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => ↑ₕ(fun (z : ℍ) => ∑ x in s, eisSummand k x z )) (↑ₕ((eisensteinSeries_SIF a k).toFun )) Filter.atTop (UpperHalfPlane.coe '' ⊤) := by From 623b79fcf5da1fd29d4dd306e62d0907811ced40 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 1 Mar 2024 16:08:45 +0000 Subject: [PATCH 070/124] lint fix --- Mathlib/Data/Finset/LocallyFinite/Box.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean index d91ba1a795314..1935333813cba 100644 --- a/Mathlib/Data/Finset/LocallyFinite/Box.lean +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -91,7 +91,8 @@ lemma existsUnique_mem_box (x : ℤ × ℤ) : ∃! n : ℕ, x ∈ box n := by end Int -lemma Finset.fun_ne_zero_cases {G : Type*} [OfNat G 0] (x : Fin 2 → G) : x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by +lemma Finset.fun_ne_zero_cases {G : Type*} [OfNat G 0] (x : Fin 2 → G) : + x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by rw [Function.ne_iff]; exact Fin.exists_fin_two lemma Finset.mem_box_ne_zero_iff_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : From 1ad2e3136c0fbfddf8af15103f8562c4dfd20fe3 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 6 Mar 2024 10:10:55 +0000 Subject: [PATCH 071/124] some golf --- .../Complex/UpperHalfPlane/Basic.lean | 3 + .../EisensteinSeries/UniformConvergence.lean | 68 +++++++++---------- 2 files changed, 35 insertions(+), 36 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index fb6db0db6727c..17c8a4a82a24b 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -146,6 +146,9 @@ lemma I_im : I.im = 1 := rfl @[simp] lemma I_re : I.re = 0 := rfl +@[simp] +lemma coe_I : I = Complex.I := rfl + end UpperHalfPlane namespace Mathlib.Meta.Positivity diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 1590d97a060d8..d2b7603fb06e4 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -131,12 +131,14 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x (Complex.abs ((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ))) ^ k := by by_cases hk0 : k ≠ 0 let n := max (x 0).natAbs (x 1).natAbs + have hn0 : n ≠ 0 := by + rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx + exact hx have h11 : ((x 0) * ↑z + (x 1)) = (((x 0 : ℝ) / (n : ℝ)) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by field_simp rw [← mul_div, div_self] - simp - rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx + simp only [mul_one] norm_cast at * simp only [Nat.cast_max, h11, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] @@ -145,14 +147,12 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) norm_cast apply Real.rpow_nonneg - rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx simp only [le_max_iff, Nat.cast_nonneg, or_self] apply Real.rpow_nonneg (Complex.abs.nonneg _) k · apply mul_le_mul simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) norm_cast apply Real.rpow_nonneg - rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx simp only [le_max_iff, Nat.cast_nonneg, or_self] apply Real.rpow_nonneg (Complex.abs.nonneg _) k · simp only [ne_eq, not_not] at hk0 @@ -164,18 +164,16 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) by_cases hn : n = 0 · rw [hn] at hx simp only [box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx + rw [hx.1, hx.2] at * by_cases hk0 : k = 0 - rw [hk0] at * - simp only [neg_zero, Real.rpow_zero, mul_one, le_refl] - rw [hx.1, hx.2] - simp only [Int.cast_zero, zero_mul, add_zero, map_zero] - have h1 : (0 : ℝ) ^ (-k) = 0 := by - rw [Real.rpow_eq_zero_iff_of_nonneg (by rfl)] - simp only [ne_eq, neg_eq_zero, hk0, not_false_eq_true, and_self] - simp only [h1, hn, CharP.cast_eq_zero, mul_zero, le_refl] - · have hx2 : x ≠ 0 := by - rw [mem_box_ne_zero_iff_ne_zero n x hx] - exact hn + . rw [hk0] at * + simp only [neg_zero, Real.rpow_zero, mul_one, le_refl] + · simp only [Int.cast_zero, zero_mul, add_zero, map_zero] + have h1 : (0 : ℝ) ^ (-k) = 0 := by + rw [Real.rpow_eq_zero_iff_of_nonneg (by rfl)] + simp only [ne_eq, neg_eq_zero, hk0, not_false_eq_true, and_self] + simp only [h1, hn, CharP.cast_eq_zero, mul_zero, le_refl] + · have hx2 := (mem_box_ne_zero_iff_ne_zero n x hx).mpr hn simp only [Int.mem_box] at hx rw [Real.rpow_neg (by apply apply_nonneg), Real.rpow_neg ((r_pos z).le), Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv] @@ -219,16 +217,16 @@ lemma r_lower_bound_on_slice {A B : ℝ} (h : 0 < B) (z : upperHalfPlaneSlice A convert hz.2 have := abs_eq_self.mpr (UpperHalfPlane.im_pos z.1).le convert this.symm - rw [Real.sqrt_le_sqrt_iff] - simp only [r1', div_pow, one_div] - rw [inv_le_inv, add_le_add_iff_right] - apply div_le_div (sq_nonneg _) - · simpa [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 - · positivity - · simpa [even_two.pow_abs] using pow_le_pow_left h.le hz.2 2 - · positivity - · positivity - · apply (r1_pos z).le + . rw [Real.sqrt_le_sqrt_iff] + simp only [r1', div_pow, one_div] + rw [inv_le_inv, add_le_add_iff_right] + apply div_le_div (sq_nonneg _) + · simpa [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 + · positivity + · simpa [even_two.pow_abs] using pow_le_pow_left h.le hz.2 2 + · positivity + · positivity + · apply (r1_pos z).le end bounding_functions @@ -283,9 +281,9 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : refine ⟨fun n ↦ ?_, Summable.congr (summable_over_box z h) fun n ↦ Finset.sum_congr rfl fun x hx ↦ ?_⟩ · simpa using (box n).summable (f ∘ (piFinTwoEquiv _).symm) - simp only [Int.mem_box] at hx - rw [← hx, one_div] - norm_cast + · simp only [Int.mem_box] at hx + rw [← hx, one_div] + norm_cast · intro y apply mul_nonneg · simp only [one_div, inv_nonneg] @@ -301,8 +299,8 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop := by have hk0 : 0 ≤ k := by linarith lift k to ℕ using hk0 - rw [← tendstoLocallyUniformlyOn_univ,tendstoLocallyUniformlyOn_iff_forall_isCompact, - eisensteinSeries_SIF] + rw [← tendstoLocallyUniformlyOn_univ,tendstoLocallyUniformlyOn_iff_forall_isCompact + (by simp only [Set.top_eq_univ, isOpen_univ]), eisensteinSeries_SIF] simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] intro K hK obtain ⟨A, B, hB, HABK⟩:= subset_slice_of_isCompact hK @@ -317,19 +315,17 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_coe_nat, one_div, map_pow, map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, ge_iff_le] at * - apply le_trans (this ?_) + apply le_trans (this (by simp only [Int.mem_box])) rw [mul_comm] apply mul_le_mul _ (by rfl) repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg, inv_nonneg, pow_nonneg (r_pos _).le]} rw [inv_le_inv] - apply pow_le_pow_left (r_pos _).le - rw [abs_of_pos (r_pos _)] - · exact r_lower_bound_on_slice hB ⟨x, HABK hx⟩ + . apply pow_le_pow_left (r_pos _).le + rw [abs_of_pos (r_pos _)] + · exact r_lower_bound_on_slice hB ⟨x, HABK hx⟩ · apply pow_pos (abs_pos.mpr (ne_of_gt (r_pos x))) · apply pow_pos (r_pos _) - · simp only [Int.mem_box] - · simp only [Set.top_eq_univ, isOpen_univ] local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm (OpenEmbedding.toPartialHomeomorph UpperHalfPlane.coe openEmbedding_coe)) From 8f1d15e405108ae815ab248f73a028685643c968 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 6 Mar 2024 10:14:59 +0000 Subject: [PATCH 072/124] more golf --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index d2b7603fb06e4..633a6ed00732d 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -303,7 +303,7 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : (by simp only [Set.top_eq_univ, isOpen_univ]), eisensteinSeries_SIF] simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] intro K hK - obtain ⟨A, B, hB, HABK⟩:= subset_slice_of_isCompact hK + obtain ⟨A, B, hB, HABK⟩ := subset_slice_of_isCompact hK have hu : Summable fun x : (gammaSet N a) => (1/(r ⟨⟨A, B⟩, hB⟩) ^ k) * ((max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by apply (Summable.subtype (summable_upper_bound hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a)).congr From bc25990c13db45c052445c6b500235295241e47b Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 6 Mar 2024 10:23:47 +0000 Subject: [PATCH 073/124] more cleanup --- .../EisensteinSeries/UniformConvergence.lean | 56 +++++++++---------- 1 file changed, 26 insertions(+), 30 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 633a6ed00732d..a18a46689c70b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -31,7 +31,7 @@ section bounding_functions /-- Auxiliary function used for bounding Eisenstein series-/ def r1 (z : ℍ) : ℝ := ((z.im ^ 2) / (z.re ^ 2 + z.im ^ 2)) -lemma r1' (z : ℍ) : r1 z = 1/((z.re / z.im) ^ 2 + 1) := by +lemma r1' (z : ℍ) : r1 z = 1 / ((z.re / z.im) ^ 2 + 1) := by field_simp [r1, im_pos z] theorem r1_pos (z : ℍ) : 0 < r1 z := by @@ -42,10 +42,10 @@ theorem r1_pos (z : ℍ) : 0 < r1 z := by /-- This function is used to give an upper bound on Eisenstein series-/ def r (z : ℍ) : ℝ := min (z.im) (Real.sqrt (r1 z)) -theorem r_pos (z : ℍ) : 0 < r z := by +lemma r_pos (z : ℍ) : 0 < r z := by simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, r1_pos, and_self] -theorem r1_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : +lemma r1_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : (z.im ^ 2) / (z.re ^ 2 + z.im ^ 2) ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by have H1 : (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 = δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2 := by ring @@ -58,7 +58,7 @@ theorem r1_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : · apply pow_two_nonneg · apply_rules [add_pos_of_nonneg_of_pos, pow_two_nonneg, (pow_pos z.im_pos 2)] -theorem auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by +lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs] have H1 : (z : ℂ).im ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + (δ * z : ℂ).im * (δ * z : ℂ).im) := by @@ -75,7 +75,7 @@ theorem auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ C add_zero, normSq_apply, add_re, mul_re, sub_zero, add_im] at * exact H1 -theorem auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by +lemma auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs, min_le_iff] have H1 : Real.sqrt (r1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by @@ -130,31 +130,27 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k ≤ (Complex.abs ((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ))) ^ k := by by_cases hk0 : k ≠ 0 - let n := max (x 0).natAbs (x 1).natAbs - have hn0 : n ≠ 0 := by - rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx - exact hx - have h11 : ((x 0) * ↑z + (x 1)) = - (((x 0 : ℝ) / (n : ℝ)) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by - field_simp - rw [← mul_div, div_self] - simp only [mul_one] - norm_cast at * - simp only [Nat.cast_max, h11, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] - rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] - cases' (div_max_sq_ge_one x hx) with H1 H2 - · apply mul_le_mul - simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) - norm_cast - apply Real.rpow_nonneg - simp only [le_max_iff, Nat.cast_nonneg, or_self] - apply Real.rpow_nonneg (Complex.abs.nonneg _) k - · apply mul_le_mul - simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) - norm_cast - apply Real.rpow_nonneg - simp only [le_max_iff, Nat.cast_nonneg, or_self] - apply Real.rpow_nonneg (Complex.abs.nonneg _) k + · let n := max (x 0).natAbs (x 1).natAbs + have hn0 : n ≠ 0 := by + rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx + exact hx + have h11 : ((x 0) * ↑z + (x 1)) = + (((x 0 : ℝ) / (n : ℝ)) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by + field_simp + rw [← mul_div, div_self] + · simp only [mul_one] + · norm_cast at * + simp only [Nat.cast_max, h11, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] + rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] + cases' (div_max_sq_ge_one x hx) with H1 H2 + · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) + · simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) + · apply Real.rpow_nonneg + simp only [le_max_iff, Nat.cast_nonneg, or_self] + · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) + · simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) + · apply Real.rpow_nonneg + simp only [le_max_iff, Nat.cast_nonneg, or_self] · simp only [ne_eq, not_not] at hk0 simp only [hk0, Real.rpow_zero, Nat.cast_max, mul_one, le_refl] From b5b30b6a33192bac1171f95ca6a7f774ebf0fadf Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 6 Mar 2024 11:10:55 +0000 Subject: [PATCH 074/124] move files around --- Mathlib/Logic/Equiv/Set.lean | 8 ++++ .../EisensteinSeries/UniformConvergence.lean | 40 +++++-------------- Mathlib/Topology/Instances/ENNReal.lean | 7 ++++ 3 files changed, 25 insertions(+), 30 deletions(-) diff --git a/Mathlib/Logic/Equiv/Set.lean b/Mathlib/Logic/Equiv/Set.lean index e9e8e7ea0e2f8..4bd032bb2f694 100644 --- a/Mathlib/Logic/Equiv/Set.lean +++ b/Mathlib/Logic/Equiv/Set.lean @@ -175,6 +175,14 @@ def setProdEquivSigma {α β : Type*} (s : Set (α × β)) : right_inv := fun ⟨x, y, h⟩ => rfl #align equiv.set_prod_equiv_sigma Equiv.setProdEquivSigma +/-- Equivalence between the sigma of a family of finsets of `β` and `β`. -/ +noncomputable def sigmaEquiv {ι κ : Type*} (s : κ → Set ι) (hs : ∀ i, ∃! j, i ∈ s j) : + (Σ j, s j) ≃ ι where + toFun x := x.2 + invFun x := ⟨(hs x).choose, x, (hs x).choose_spec.1⟩ + left_inv x := by ext; exacts [((hs x.2).choose_spec.2 x.1 x.2.2).symm, rfl] + right_inv x := by rfl + /-- The subtypes corresponding to equal sets are equivalent. -/ @[simps! apply] def setCongr {α : Type*} {s t : Set α} (h : s = t) : s ≃ t := diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index a18a46689c70b..848b501c239c5 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -17,7 +17,6 @@ We show that `eisSummand` converges locally uniformly on `ℍ` to the Eisenstein and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. -/ - noncomputable section open Complex Filter UpperHalfPlane Set Finset @@ -58,11 +57,11 @@ lemma r1_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : · apply pow_two_nonneg · apply_rules [add_pos_of_nonneg_of_pos, pow_two_nonneg, (pow_pos z.im_pos 2)] -lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by +lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ ^ 2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs] have H1 : (z : ℂ).im ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + (δ * z : ℂ).im * (δ * z : ℂ).im) := by - have h1 : (δ * z : ℂ).im * (δ * z : ℂ).im = δ^2 * (z : ℂ).im * (z : ℂ).im := by + have h1 : (δ * z : ℂ).im * (δ * z : ℂ).im = δ ^ 2 * (z : ℂ).im * (z : ℂ).im := by simp only [mul_im, ofReal_re, coe_im, ofReal_im, coe_re, zero_mul, add_zero] ring rw [Real.le_sqrt', h1 ] @@ -75,7 +74,7 @@ lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ^2) : r z ≤ Com add_zero, normSq_apply, add_re, mul_re, sub_zero, add_im] at * exact H1 -lemma auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by +lemma auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε ^ 2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs, min_le_iff] have H1 : Real.sqrt (r1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by @@ -113,7 +112,7 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : rw [H1, div_pow, Int.cast_natAbs (x 0), Int.cast_abs] have : (x 0 : ℝ) ≠ 0 := by simpa using (ne_zero_if_max hx H1) - have h1 : (x 0 : ℝ)^2/(_root_.abs (x 0 : ℝ))^2 = 1 := by + have h1 : (x 0 : ℝ) ^ 2/(_root_.abs (x 0 : ℝ)) ^ 2 = 1 := by simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, this, div_self] exact h1.symm.le @@ -162,7 +161,7 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) simp only [box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx rw [hx.1, hx.2] at * by_cases hk0 : k = 0 - . rw [hk0] at * + · rw [hk0] at * simp only [neg_zero, Real.rpow_zero, mul_one, le_refl] · simp only [Int.cast_zero, zero_mul, add_zero, map_zero] have h1 : (0 : ℝ) ^ (-k) = 0 := by @@ -213,35 +212,16 @@ lemma r_lower_bound_on_slice {A B : ℝ} (h : 0 < B) (z : upperHalfPlaneSlice A convert hz.2 have := abs_eq_self.mpr (UpperHalfPlane.im_pos z.1).le convert this.symm - . rw [Real.sqrt_le_sqrt_iff] + · rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] simp only [r1', div_pow, one_div] - rw [inv_le_inv, add_le_add_iff_right] - apply div_le_div (sq_nonneg _) - · simpa [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 - · positivity + rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] + apply div_le_div (sq_nonneg _) _ (by positivity) · simpa [even_two.pow_abs] using pow_le_pow_left h.le hz.2 2 - · positivity - · positivity - · apply (r1_pos z).le + · simpa [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 end bounding_functions section summability -variable {ι κ α : Type*} [AddCommMonoid α] [TopologicalSpace α] - -/-- Equivalence between the sigma of a family of finsets of `β` and `β`. -/ -noncomputable def sigmaEquiv {ι κ : Type*} (s : κ → Set ι) (hs : ∀ i, ∃! j, i ∈ s j) : - (Σ j, s j) ≃ ι where - toFun x := x.2 - invFun x := ⟨(hs x).choose, x, (hs x).choose_spec.1⟩ - left_inv x := by ext; exacts [((hs x.2).choose_spec.2 x.1 x.2.2).symm, rfl] - right_inv x := by rfl - -lemma summable_partition {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} (hs : ∀ i, ∃! j, i ∈ s j) : - Summable f ↔ (∀ j, Summable fun i : s j ↦ f i) ∧ Summable fun j ↦ ∑' i : s j, f i := by - rw [← (sigmaEquiv s hs).summable_iff, summable_sigma_of_nonneg] - simp only [sigmaEquiv, Equiv.coe_fn_mk, Function.comp_apply] - exact fun _ ↦ hf _ lemma summable_r_pow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by @@ -317,7 +297,7 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg, inv_nonneg, pow_nonneg (r_pos _).le]} rw [inv_le_inv] - . apply pow_le_pow_left (r_pos _).le + · apply pow_le_pow_left (r_pos _).le rw [abs_of_pos (r_pos _)] · exact r_lower_bound_on_slice hB ⟨x, HABK hx⟩ · apply pow_pos (abs_pos.mpr (ne_of_gt (r_pos x))) diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 925fccf576829..ecf6367297ce3 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -1369,6 +1369,13 @@ theorem summable_sigma_of_nonneg {β : α → Type*} {f : (Σ x, β x) → ℝ} exact mod_cast NNReal.summable_sigma #align summable_sigma_of_nonneg summable_sigma_of_nonneg +lemma summable_partition {κ : Type*} {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} + (hs : ∀ i, ∃! j, i ∈ s j) : Summable f ↔ + (∀ j, Summable fun i : s j ↦ f i) ∧ Summable fun j ↦ ∑' i : s j, f i := by + rw [← (Equiv.sigmaEquiv s hs).summable_iff, summable_sigma_of_nonneg] + simp only [Equiv.sigmaEquiv, Equiv.coe_fn_mk, Function.comp_apply] + exact fun _ ↦ hf _ + theorem summable_prod_of_nonneg {f : (α × β) → ℝ} (hf : 0 ≤ f) : Summable f ↔ (∀ x, Summable fun y ↦ f (x, y)) ∧ Summable fun x ↦ ∑' y, f (x, y) := (Equiv.sigmaEquivProd _ _).summable_iff.symm.trans <| summable_sigma_of_nonneg fun _ ↦ hf _ From eaa2ca50412f804f71cb34d807c97144864d3af1 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 6 Mar 2024 11:12:46 +0000 Subject: [PATCH 075/124] fix --- Mathlib/Topology/Instances/ENNReal.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index ecf6367297ce3..3b4f3c9c99dc1 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -1369,7 +1369,7 @@ theorem summable_sigma_of_nonneg {β : α → Type*} {f : (Σ x, β x) → ℝ} exact mod_cast NNReal.summable_sigma #align summable_sigma_of_nonneg summable_sigma_of_nonneg -lemma summable_partition {κ : Type*} {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} +lemma summable_partition {κ ι: Type*} {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} (hs : ∀ i, ∃! j, i ∈ s j) : Summable f ↔ (∀ j, Summable fun i : s j ↦ f i) ∧ Summable fun j ↦ ∑' i : s j, f i := by rw [← (Equiv.sigmaEquiv s hs).summable_iff, summable_sigma_of_nonneg] From 15a5ff6e63eaa107a9d9c6e8c03632eb58d1b0fd Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 8 Mar 2024 15:29:52 +0000 Subject: [PATCH 076/124] rev fixes wip --- .../Complex/UpperHalfPlane/Basic.lean | 2 +- Mathlib/Data/Finset/LocallyFinite/Box.lean | 25 ++++++++----------- .../EisensteinSeries/UniformConvergence.lean | 13 ++++++---- 3 files changed, 20 insertions(+), 20 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 17c8a4a82a24b..25038ee2d492e 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -146,7 +146,7 @@ lemma I_im : I.im = 1 := rfl @[simp] lemma I_re : I.re = 0 := rfl -@[simp] +@[simp, norm_cast] lemma coe_I : I = Complex.I := rfl end UpperHalfPlane diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean index 1935333813cba..953669c8a47dd 100644 --- a/Mathlib/Data/Finset/LocallyFinite/Box.lean +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -48,6 +48,17 @@ lemma box_succ_disjUnion (n : ℕ) : @[simp] lemma zero_mem_box : (0 : α) ∈ box n ↔ n = 0 := by cases n <;> simp [box_succ_eq_sdiff] +lemma mem_box_eq_zero_iff_eq_zero (x : α) (hx : x ∈ box n) : x = 0 ↔ n = 0 := by + rw [← zero_mem_box (α := α) (n := n)] + constructor + · intro h + rw [h] at hx + exact hx + · intro h + rw [(zero_mem_box (α := α) (n := n)).mp h] at hx + simp only [box_zero, mem_singleton] at hx + exact hx + end Finset open Finset @@ -90,17 +101,3 @@ lemma existsUnique_mem_box (x : ℤ × ℤ) : ∃! n : ℕ, x ∈ box n := by use max x.1.natAbs x.2.natAbs; simp only [mem_box, and_self_iff, forall_eq'] end Int - -lemma Finset.fun_ne_zero_cases {G : Type*} [OfNat G 0] (x : Fin 2 → G) : - x ≠ 0 ↔ x 0 ≠ 0 ∨ x 1 ≠ 0 := by - rw [Function.ne_iff]; exact Fin.exists_fin_two - -lemma Finset.mem_box_ne_zero_iff_ne_zero (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : - x ≠ 0 ↔ n ≠ 0 := by - constructor - intro h h0 - simp only [h0, Nat.cast_zero, box_zero, mem_singleton, Prod.ext_iff] at hx - rw [fun_ne_zero_cases, hx.1, hx.2] at h - · simp at h - rintro hn rfl - simp only [Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self, eq_comm, hn] at hx diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 848b501c239c5..0be41b579eb3b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -101,8 +101,8 @@ lemma ne_zero_if_max' {x : Fin 2 → ℤ} (hx : x ≠ 0) apply ne_zero_if_max (x := ![x 1, x 0]) ?_ (by simpa using h) simp only [ne_eq, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_true, not_and] intro h1 h0 - rw [fun_ne_zero_cases, h1, h0] at hx - simp only [ne_eq, not_true_eq_false, or_self] at * + rw [Function.ne_iff, Fin.exists_fin_two, h1, h0] at hx + simp only [Pi.zero_apply, ne_eq, not_true_eq_false, or_self] at * lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 ∨ @@ -131,8 +131,9 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x by_cases hk0 : k ≠ 0 · let n := max (x 0).natAbs (x 1).natAbs have hn0 : n ≠ 0 := by - rw [mem_box_ne_zero_iff_ne_zero n x (by rw [Int.mem_box])] at hx - exact hx + rw [← Iff.ne ((mem_box_eq_zero_iff_eq_zero (α := ℤ × ℤ)) (x 0, x 1) (by simp)), + ← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun x ↦ ℤ)))] at * + simpa using hx have h11 : ((x 0) * ↑z + (x 1)) = (((x 0 : ℝ) / (n : ℝ)) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by field_simp @@ -168,7 +169,9 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) rw [Real.rpow_eq_zero_iff_of_nonneg (by rfl)] simp only [ne_eq, neg_eq_zero, hk0, not_false_eq_true, and_self] simp only [h1, hn, CharP.cast_eq_zero, mul_zero, le_refl] - · have hx2 := (mem_box_ne_zero_iff_ne_zero n x hx).mpr hn + · have hx2 : x ≠ 0 := by + rw [← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℤ)))] + simpa using (Iff.ne ((mem_box_eq_zero_iff_eq_zero (α := ℤ × ℤ)) (x 0, x 1) hx)).mpr hn simp only [Int.mem_box] at hx rw [Real.rpow_neg (by apply apply_nonneg), Real.rpow_neg ((r_pos z).le), Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv] From 9a7b9130254403232814d5ff12b1c55b07346a91 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 8 Mar 2024 15:38:08 +0000 Subject: [PATCH 077/124] this can probs be golfed more, but I'm stoopid --- Mathlib/Data/Finset/LocallyFinite/Box.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean index 953669c8a47dd..9e02385337eca 100644 --- a/Mathlib/Data/Finset/LocallyFinite/Box.lean +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -49,14 +49,13 @@ lemma box_succ_disjUnion (n : ℕ) : @[simp] lemma zero_mem_box : (0 : α) ∈ box n ↔ n = 0 := by cases n <;> simp [box_succ_eq_sdiff] lemma mem_box_eq_zero_iff_eq_zero (x : α) (hx : x ∈ box n) : x = 0 ↔ n = 0 := by - rw [← zero_mem_box (α := α) (n := n)] + rw [← zero_mem_box (α := α)] constructor · intro h rw [h] at hx exact hx · intro h - rw [(zero_mem_box (α := α) (n := n)).mp h] at hx - simp only [box_zero, mem_singleton] at hx + simp only [(zero_mem_box (α := α) (n := n)).mp h, box_zero, mem_singleton] at hx exact hx end Finset From 3600a74f4056195368d1b4387c25d1603985b025 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 8 Mar 2024 15:56:54 +0000 Subject: [PATCH 078/124] composing non_zero function with inj fun is non_zero --- Mathlib/Algebra/Group/Pi/Basic.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/Mathlib/Algebra/Group/Pi/Basic.lean b/Mathlib/Algebra/Group/Pi/Basic.lean index 87828eef02898..0601dd1509f4d 100644 --- a/Mathlib/Algebra/Group/Pi/Basic.lean +++ b/Mathlib/Algebra/Group/Pi/Basic.lean @@ -548,6 +548,21 @@ theorem bijective_pi_map {F : ∀ i, f i → g i} (hF : ∀ i, Bijective (F i)) ⟨injective_pi_map fun i => (hF i).injective, surjective_pi_map fun i => (hF i).surjective⟩ #align function.bijective_pi_map Function.bijective_pi_map +lemma comp_eq_const_iff {α β γ: Type*} (b : β) (f : α → β) (g : β → γ) + (hg : Injective g) : g ∘ f = Function.const _ (g b) ↔ f = Function.const _ b := + hg.comp_left.eq_iff' rfl + +lemma comp_eq_zero_iff {α β γ: Type*} [OfNat β 0] [ OfNat γ 0] (f : α → β) (g : β → γ) + (hg : Injective g) (hg0 : g 0 = 0) : g ∘ f = 0 ↔ f = 0 := by + have := (comp_eq_const_iff 0 f g hg) + rw [hg0] at this + simp only [const_zero] at this + exact this + +lemma comp_inj_ne_zero {α β γ: Type*} [OfNat β 0] [ OfNat γ 0] (f : α → β) (g : β → γ) + (hg : Injective g) (hg0 : g 0 = 0) : (g ∘ f) ≠ 0 ↔ f ≠ 0 := + (Iff.ne (comp_eq_zero_iff f g hg hg0)) + end Function /-- If the one function is surjective, the codomain is trivial. -/ From 07ec99b4f2a19db0e77fe3e27dd19e0a338098c5 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 8 Mar 2024 16:53:30 +0000 Subject: [PATCH 079/124] use inj func pr --- Mathlib/Algebra/Group/Pi/Basic.lean | 15 +++++++++++++ .../EisensteinSeries/UniformConvergence.lean | 22 ++++++------------- 2 files changed, 22 insertions(+), 15 deletions(-) diff --git a/Mathlib/Algebra/Group/Pi/Basic.lean b/Mathlib/Algebra/Group/Pi/Basic.lean index 5720ae845a1cd..a247f498efb5d 100644 --- a/Mathlib/Algebra/Group/Pi/Basic.lean +++ b/Mathlib/Algebra/Group/Pi/Basic.lean @@ -548,6 +548,21 @@ theorem bijective_pi_map {F : ∀ i, f i → g i} (hF : ∀ i, Bijective (F i)) ⟨injective_pi_map fun i => (hF i).injective, surjective_pi_map fun i => (hF i).surjective⟩ #align function.bijective_pi_map Function.bijective_pi_map +lemma comp_eq_const_iff {α β γ: Type*} (b : β) (f : α → β) (g : β → γ) + (hg : Function.Injective g) : g ∘ f = Function.const _ (g b) ↔ f = Function.const _ b := + hg.comp_left.eq_iff' rfl + +lemma comp_eq_zero_iff {α β γ: Type*} [OfNat β 0] [ OfNat γ 0] (f : α → β) (g : β → γ) + (hg : Function.Injective g) (hg0 : g 0 = 0) : g ∘ f = 0 ↔ f = 0 := by + have := (comp_eq_const_iff 0 f g hg) + rw [hg0] at this + simp only [Function.const_zero] at this + exact this + +lemma comp_inj_ne_zero {α β γ: Type*} [OfNat β 0] [ OfNat γ 0] (f : α → β) (g : β → γ) + (hg : Function.Injective g) (hg0 : g 0 = 0) : (g ∘ f) ≠ 0 ↔ f ≠ 0 := + (Iff.ne (comp_eq_zero_iff f g hg hg0)) + end Function /-- If the one function is surjective, the codomain is trivial. -/ diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 0be41b579eb3b..a4604fee81eb6 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -145,12 +145,10 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x cases' (div_max_sq_ge_one x hx) with H1 H2 · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) · simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) - · apply Real.rpow_nonneg - simp only [le_max_iff, Nat.cast_nonneg, or_self] + · positivity · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) · simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) - · apply Real.rpow_nonneg - simp only [le_max_iff, Nat.cast_nonneg, or_self] + · positivity · simp only [ne_eq, not_not] at hk0 simp only [hk0, Real.rpow_zero, Nat.cast_max, mul_one, le_refl] @@ -158,12 +156,10 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)))) ^ (-k) ≤ (((r z) ^ (-k) * n ^ (-k))) := by by_cases hn : n = 0 - · rw [hn] at hx - simp only [box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx + · simp only [hn, box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx rw [hx.1, hx.2] at * by_cases hk0 : k = 0 - · rw [hk0] at * - simp only [neg_zero, Real.rpow_zero, mul_one, le_refl] + · simp only [hk0, neg_zero, Real.rpow_zero, mul_one, le_refl] · simp only [Int.cast_zero, zero_mul, add_zero, map_zero] have h1 : (0 : ℝ) ^ (-k) = 0 := by rw [Real.rpow_eq_zero_iff_of_nonneg (by rfl)] @@ -180,13 +176,9 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) · simp only [Nat.cast_max] · apply Real.rpow_pos_of_pos apply Complex.abs.pos (linear_ne_zero ![x 0, x 1] z ?_) - simp only [ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, and_true, - not_and] at * - intro hg h1 - have : x = ![x 0, x 1] := by - exact List.ofFn_inj.mp rfl - rw [this, hg, h1] at hx2 - simp only [Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at * + have := (Function.comp_inj_ne_zero x _ Int.cast_injective Int.cast_zero (γ := ℝ)).mpr hx2 + rw [← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℝ)))] at this + simpa using this · apply mul_pos (Real.rpow_pos_of_pos (r_pos z) _) apply Real.rpow_pos_of_pos norm_cast From 2b6b90de1fa77d52a9b1ff524febc539fe636879 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 8 Mar 2024 17:06:27 +0000 Subject: [PATCH 080/124] clean a bit --- .../EisensteinSeries/UniformConvergence.lean | 14 ++++++-------- 1 file changed, 6 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index a4604fee81eb6..aade233fdab80 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -44,7 +44,7 @@ def r (z : ℍ) : ℝ := min (z.im) (Real.sqrt (r1 z)) lemma r_pos (z : ℍ) : 0 < r z := by simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, r1_pos, and_self] -lemma r1_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : +lemma r1_aux_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : (z.im ^ 2) / (z.re ^ 2 + z.im ^ 2) ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by have H1 : (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 = δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2 := by ring @@ -64,7 +64,7 @@ lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ ^ 2) : r z ≤ C have h1 : (δ * z : ℂ).im * (δ * z : ℂ).im = δ ^ 2 * (z : ℂ).im * (z : ℂ).im := by simp only [mul_im, ofReal_re, coe_im, ofReal_im, coe_re, zero_mul, add_zero] ring - rw [Real.le_sqrt', h1 ] + rw [Real.le_sqrt', h1] nlinarith exact z.2 simp only [UpperHalfPlane.coe_im, UpperHalfPlane.coe_re, AbsoluteValue.coe_mk, MulHom.coe_mk, @@ -79,7 +79,7 @@ lemma auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε ^ 2) : r z ≤ C have H1 : Real.sqrt (r1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by rw [r1, Real.sqrt_le_sqrt_iff, ← pow_two, ← pow_two] - apply r1_bound z δ hε + apply r1_aux_bound z δ hε nlinarith right simp only [ne_eq, coe_re, coe_im, normSq_apply, AbsoluteValue.coe_mk, MulHom.coe_mk, add_re, @@ -184,9 +184,8 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) norm_cast exact Nat.pos_of_ne_zero hn -theorem eis_is_bounded_on_box (k n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) - (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ - (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by +theorem eis_is_bounded_on_box (k n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : + (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by have := eis_is_bounded_on_box_rpow (Nat.cast_nonneg k) z n x hx norm_cast at * simp_rw [zpow_neg, ← mul_inv] at this @@ -205,8 +204,7 @@ lemma r_lower_bound_on_slice {A B : ℝ} (h : 0 < B) (z : upperHalfPlaneSlice A apply min_le_min · dsimp only convert hz.2 - have := abs_eq_self.mpr (UpperHalfPlane.im_pos z.1).le - convert this.symm + apply (abs_eq_self.mpr (UpperHalfPlane.im_pos z.1).le).symm · rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] simp only [r1', div_pow, one_div] rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] From 4e2ac144e3605492f22036807ca0670f0918412f Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 25 Mar 2024 10:23:09 +0000 Subject: [PATCH 081/124] fix --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index aade233fdab80..fcab583e155e7 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -144,10 +144,10 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] cases' (div_max_sq_ge_one x hx) with H1 H2 · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) - · simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) + · simpa [n] using (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) · positivity · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) - · simpa using (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) + · simpa [n] using (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) · positivity · simp only [ne_eq, not_not] at hk0 simp only [hk0, Real.rpow_zero, Nat.cast_max, mul_one, le_refl] @@ -176,7 +176,7 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) · simp only [Nat.cast_max] · apply Real.rpow_pos_of_pos apply Complex.abs.pos (linear_ne_zero ![x 0, x 1] z ?_) - have := (Function.comp_inj_ne_zero x _ Int.cast_injective Int.cast_zero (γ := ℝ)).mpr hx2 + have := (Function.comp_ne_zero_iff x Int.cast_injective Int.cast_zero (γ := ℝ)).mpr hx2 rw [← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℝ)))] at this simpa using this · apply mul_pos (Real.rpow_pos_of_pos (r_pos z) _) @@ -252,7 +252,7 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : · simpa using (box n).summable (f ∘ (piFinTwoEquiv _).symm) · simp only [Int.mem_box] at hx rw [← hx, one_div] - norm_cast + simp only [Nat.cast_max, one_div, Fin.isValue, Fin.cons_zero, Fin.cons_one, f] · intro y apply mul_nonneg · simp only [one_div, inv_nonneg] From a0e1e79fb1470b8c03a26be2ac24186490d0b1f4 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 15 Apr 2024 11:14:20 +0100 Subject: [PATCH 082/124] Update Mathlib.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 9205cdaaf3805..903d5f292e8be 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3015,6 +3015,7 @@ import Mathlib.NumberTheory.ModularForms.Basic import Mathlib.NumberTheory.ModularForms.CongruenceSubgroups import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic import Mathlib.NumberTheory.ModularForms.EisensteinSeries.UniformConvergence +import Mathlib.NumberTheory.ModularForms.JacobiTheta.Bounds import Mathlib.NumberTheory.ModularForms.JacobiTheta.Manifold import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable From 77bafcb42fc951e43f3acedb1d8fa8d069c30b3c Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 15 Apr 2024 12:15:26 +0100 Subject: [PATCH 083/124] fix --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index fcab583e155e7..4916dfb4095df 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -193,7 +193,7 @@ theorem eis_is_bounded_on_box (k n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hx : (x refine abs_eq_self.mpr ?_ apply mul_nonneg (pow_nonneg (r_pos z).le _) simp only [Nat.cast_pow, ge_iff_le, Nat.cast_nonneg, pow_nonneg] - simp only [map_pow, zpow_coe_nat, H] + simp only [map_pow, zpow_natCast, H] simpa using this lemma r_lower_bound_on_slice {A B : ℝ} (h : 0 < B) (z : upperHalfPlaneSlice A B) : @@ -277,11 +277,11 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : (1/(r ⟨⟨A, B⟩, hB⟩) ^ k) * ((max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by apply (Summable.subtype (summable_upper_bound hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a)).congr intro v - simp only [zpow_coe_nat, one_div, Function.comp_apply] + simp only [zpow_natCast, one_div, Function.comp_apply] apply tendstoUniformlyOn_tsum hu intro v x hx have := eis_is_bounded_on_box k (max (v.1 0).natAbs (v.1 1).natAbs) x v - simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_coe_nat, one_div, map_pow, + simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_natCast, one_div, map_pow, map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, ge_iff_le] at * apply le_trans (this (by simp only [Int.mem_box])) From e981f3d240c5ebea495363610c604874fa97eb3e Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 15 Apr 2024 17:51:11 +0100 Subject: [PATCH 084/124] rev fixes --- .../Complex/UpperHalfPlane/Metric.lean | 11 ++-- Mathlib/Data/Finset/LocallyFinite/Box.lean | 4 +- .../EisensteinSeries/UniformConvergence.lean | 52 ++++++------------- 3 files changed, 24 insertions(+), 43 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index 234148c8ec3a2..0479358f18c05 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -394,15 +394,14 @@ def upperHalfPlaneSlice (A B : ℝ) := theorem slice_mem_iff (A B : ℝ) (z : ℍ) : z ∈ upperHalfPlaneSlice A B ↔ Complex.abs z.1.1 ≤ A ∧ B ≤ Complex.abs z.1.2 := Iff.rfl -lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : ∃ A B : ℝ, 0 < B ∧ - K ⊆ upperHalfPlaneSlice A B := by +lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : + ∃ A B : ℝ, 0 < B ∧ K ⊆ upperHalfPlaneSlice A B := by obtain rfl | hne := K.eq_empty_or_nonempty - exact ⟨1, 1, Real.zero_lt_one, by simp⟩ - have hcts : ContinuousOn (fun t => t.im) K := by - apply Continuous.continuousOn UpperHalfPlane.continuous_im + · exact ⟨1, 1, Real.zero_lt_one, by simp⟩ + have hcts : ContinuousOn (fun t => t.im) K := UpperHalfPlane.continuous_im.continuousOn obtain ⟨b, _, HB⟩ := IsCompact.exists_isMinOn hK hne hcts obtain ⟨r, _, hr2⟩ := Bornology.IsBounded.subset_closedBall_lt hK.isBounded 0 UpperHalfPlane.I - refine' ⟨Real.sinh r + Complex.abs ((UpperHalfPlane.center UpperHalfPlane.I r)), b.im, b.2, _⟩ + refine ⟨Real.sinh r + Complex.abs (UpperHalfPlane.center UpperHalfPlane.I r), b.im, b.2, ?_⟩ intro z hz simp only [I_im, slice_mem_iff, abs_ofReal, ge_iff_le] at * constructor diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean index 388e66526a1ef..0c28f657ebb39 100644 --- a/Mathlib/Data/Finset/LocallyFinite/Box.lean +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -48,14 +48,14 @@ lemma box_succ_disjUnion (n : ℕ) : @[simp] lemma zero_mem_box : (0 : α) ∈ box n ↔ n = 0 := by cases n <;> simp [box_succ_eq_sdiff] -lemma mem_box_eq_zero_iff_eq_zero (x : α) (hx : x ∈ box n) : x = 0 ↔ n = 0 := by +lemma eq_zero_iff_eq_zero_of_mem_box {x : α} (hx : x ∈ box n) : x = 0 ↔ n = 0 := by rw [← zero_mem_box (α := α)] constructor · intro h rw [h] at hx exact hx · intro h - simp only [(zero_mem_box (α := α) (n := n)).mp h, box_zero, mem_singleton] at hx + simp only [zero_mem_box.mp h, box_zero, mem_singleton] at hx exact hx end Finset diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 4916dfb4095df..20a9fa0326b59 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -28,18 +28,17 @@ namespace EisensteinSeries section bounding_functions /-- Auxiliary function used for bounding Eisenstein series-/ -def r1 (z : ℍ) : ℝ := ((z.im ^ 2) / (z.re ^ 2 + z.im ^ 2)) +def r1 (z : ℍ) : ℝ := z.im ^ 2 / (z.re ^ 2 + z.im ^ 2) lemma r1' (z : ℍ) : r1 z = 1 / ((z.re / z.im) ^ 2 + 1) := by field_simp [r1, im_pos z] theorem r1_pos (z : ℍ) : 0 < r1 z := by - have H2 : 0 < (z.re ^ 2 + z.im ^ 2) := by - apply_rules [pow_pos, add_pos_of_nonneg_of_pos, pow_two_nonneg, z.2] - exact div_pos (pow_pos z.im_pos 2) H2 + dsimp only [r1] + positivity /-- This function is used to give an upper bound on Eisenstein series-/ -def r (z : ℍ) : ℝ := min (z.im) (Real.sqrt (r1 z)) +def r (z : ℍ) : ℝ := min z.im (Real.sqrt (r1 z)) lemma r_pos (z : ℍ) : 0 < r z := by simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, r1_pos, and_self] @@ -55,7 +54,7 @@ lemma r1_aux_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : apply mul_nonneg · linarith · apply pow_two_nonneg - · apply_rules [add_pos_of_nonneg_of_pos, pow_two_nonneg, (pow_pos z.im_pos 2)] + · positivity lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ ^ 2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by rw [r, Complex.abs] @@ -65,8 +64,8 @@ lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ ^ 2) : r z ≤ C simp only [mul_im, ofReal_re, coe_im, ofReal_im, coe_re, zero_mul, add_zero] ring rw [Real.le_sqrt', h1] - nlinarith - exact z.2 + · nlinarith + · exact z.2 simp only [UpperHalfPlane.coe_im, UpperHalfPlane.coe_re, AbsoluteValue.coe_mk, MulHom.coe_mk, min_le_iff] at * left @@ -88,13 +87,10 @@ lemma auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε ^ 2) : r z ≤ C exact H1 lemma ne_zero_if_max {x : Fin 2 → ℤ} (hx : x ≠ 0) - (h : (max (x 0).natAbs (x 1).natAbs) = (x 0).natAbs) : (x 0) ≠ 0 := by - intro h0 - rw [h0] at h - simp only [ne_eq, Int.natAbs_zero, ge_iff_le, zero_le, max_eq_right, Int.natAbs_eq_zero] at * - have : x = ![x 0, x 1] := List.ofFn_inj.mp rfl - rw [h0, h] at this - simp only [this, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_self, not_true_eq_false] at hx + (h : max (x 0).natAbs (x 1).natAbs = (x 0).natAbs) : x 0 ≠ 0 := by + contrapose! hx + ext x + fin_cases x <;> aesop lemma ne_zero_if_max' {x : Fin 2 → ℤ} (hx : x ≠ 0) (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by @@ -107,23 +103,9 @@ lemma ne_zero_if_max' {x : Fin 2 → ℤ} (hx : x ≠ 0) lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 ∨ (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 := by - cases' (max_choice (x 0).natAbs (x 1).natAbs) with H1 H2 - · left - rw [H1, div_pow, Int.cast_natAbs (x 0), Int.cast_abs] - have : (x 0 : ℝ) ≠ 0 := by - simpa using (ne_zero_if_max hx H1) - have h1 : (x 0 : ℝ) ^ 2/(_root_.abs (x 0 : ℝ)) ^ 2 = 1 := by - simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, - this, div_self] - exact h1.symm.le - · right - rw [H2,div_pow, Int.cast_natAbs (x 1), Int.cast_abs] - have : (x 1 : ℝ) ≠ 0 := by - simpa using (ne_zero_if_max' hx H2) - have h1 : (x 1 : ℝ)^2/(_root_.abs (x 1 : ℝ))^2 = 1 := by - simp only [_root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, - this, div_self] - exact h1.symm.le + refine (max_choice (x 0).natAbs (x 1).natAbs).imp (fun H1 => ?_) (fun H2 => ?_) + · simp [H1, Int.cast_natAbs, ne_zero_if_max hx H1] + · simp [H2, Int.cast_natAbs, ne_zero_if_max' hx H2] lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) : ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k ≤ @@ -131,7 +113,7 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x by_cases hk0 : k ≠ 0 · let n := max (x 0).natAbs (x 1).natAbs have hn0 : n ≠ 0 := by - rw [← Iff.ne ((mem_box_eq_zero_iff_eq_zero (α := ℤ × ℤ)) (x 0, x 1) (by simp)), + rw [← Iff.ne ((eq_zero_iff_eq_zero_of_mem_box (α := ℤ × ℤ) (x := ((x 0, x 1)))) (by simp)), ← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun x ↦ ℤ)))] at * simpa using hx have h11 : ((x 0) * ↑z + (x 1)) = @@ -167,7 +149,7 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) simp only [h1, hn, CharP.cast_eq_zero, mul_zero, le_refl] · have hx2 : x ≠ 0 := by rw [← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℤ)))] - simpa using (Iff.ne ((mem_box_eq_zero_iff_eq_zero (α := ℤ × ℤ)) (x 0, x 1) hx)).mpr hn + simpa using (Iff.ne ((eq_zero_iff_eq_zero_of_mem_box (x := (x 0, x 1))) hx)).mpr hn simp only [Int.mem_box] at hx rw [Real.rpow_neg (by apply apply_nonneg), Real.rpow_neg ((r_pos z).le), Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv] @@ -281,7 +263,7 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : apply tendstoUniformlyOn_tsum hu intro v x hx have := eis_is_bounded_on_box k (max (v.1 0).natAbs (v.1 1).natAbs) x v - simp only [Nat.cast_max, Int.coe_natAbs, iff_true, zpow_natCast, one_div, map_pow, + simp only [Nat.cast_max,Int.natCast_natAbs, iff_true, zpow_natCast, one_div, map_pow, map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, ge_iff_le] at * apply le_trans (this (by simp only [Int.mem_box])) From f6f8b083c7dbee1795942876ca7ac1320c90b8f3 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 15 Apr 2024 18:06:58 +0100 Subject: [PATCH 085/124] missed one --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 20a9fa0326b59..fa9fb9db4990d 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -118,10 +118,8 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x simpa using hx have h11 : ((x 0) * ↑z + (x 1)) = (((x 0 : ℝ) / (n : ℝ)) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by - field_simp - rw [← mul_div, div_self] - · simp only [mul_one] - · norm_cast at * + have : (n : ℂ) ≠ 0 := by norm_cast + field_simp [this] simp only [Nat.cast_max, h11, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] cases' (div_max_sq_ge_one x hx) with H1 H2 From 123dd623b329aace54f8d16b219d123c254e2b35 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 09:09:24 +0100 Subject: [PATCH 086/124] merge fix/update --- Mathlib/Data/Finset/LocallyFinite/Box.lean | 11 ++--------- 1 file changed, 2 insertions(+), 9 deletions(-) diff --git a/Mathlib/Data/Finset/LocallyFinite/Box.lean b/Mathlib/Data/Finset/LocallyFinite/Box.lean index 0c28f657ebb39..53b8b7c8c315b 100644 --- a/Mathlib/Data/Finset/LocallyFinite/Box.lean +++ b/Mathlib/Data/Finset/LocallyFinite/Box.lean @@ -48,15 +48,8 @@ lemma box_succ_disjUnion (n : ℕ) : @[simp] lemma zero_mem_box : (0 : α) ∈ box n ↔ n = 0 := by cases n <;> simp [box_succ_eq_sdiff] -lemma eq_zero_iff_eq_zero_of_mem_box {x : α} (hx : x ∈ box n) : x = 0 ↔ n = 0 := by - rw [← zero_mem_box (α := α)] - constructor - · intro h - rw [h] at hx - exact hx - · intro h - simp only [zero_mem_box.mp h, box_zero, mem_singleton] at hx - exact hx +lemma eq_zero_iff_eq_zero_of_mem_box {x : α} (hx : x ∈ box n) : x = 0 ↔ n = 0 := + ⟨zero_mem_box.mp ∘ (· ▸ hx), fun hn ↦ by rwa [hn, box_zero, mem_singleton] at hx⟩ end Finset From dc0a5b12b75bd2c83a741ca24f3fcbd3dd40c207 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 09:14:48 +0100 Subject: [PATCH 087/124] fix --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index fa9fb9db4990d..cef971aaa95d9 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -7,7 +7,7 @@ Authors: Chris Birkbeck import Mathlib.Analysis.Complex.UpperHalfPlane.Metric import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.PSeries -import Mathlib.Data.Finset.LocallyFinite.Box +import Mathlib.Order.Interval.Finset.Box import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic /-! From fdeeff6f645065a2e648b0a80e27fc5f9466b4ff Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 09:27:20 +0100 Subject: [PATCH 088/124] feat(Complex/UpperHalfPlane): add vertical strips needed for mod forms and make i as an element of the upper half plane --- .../Analysis/Complex/UpperHalfPlane/Basic.lean | 11 +++++++++++ .../Complex/UpperHalfPlane/Metric.lean | 18 ++++++++++++++++++ 2 files changed, 29 insertions(+) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 58baf4669137d..52456d152a314 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -137,6 +137,17 @@ theorem ne_zero (z : ℍ) : (z : ℂ) ≠ 0 := mt (congr_arg Complex.im) z.im_ne_zero #align upper_half_plane.ne_zero UpperHalfPlane.ne_zero +def I : ℍ := ⟨Complex.I, by simp⟩ + +@[simp] +lemma I_im : I.im = 1 := rfl + +@[simp] +lemma I_re : I.re = 0 := rfl + +@[simp, norm_cast] +lemma coe_I : I = Complex.I := rfl + end UpperHalfPlane namespace Mathlib.Meta.Positivity diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index ce8903b16fbc2..961e40fbec3c9 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -385,4 +385,22 @@ instance : IsometricSMul SL(2, ℝ) ℍ := exact (isometry_real_vadd w).comp (h₀.comp <| (isometry_real_vadd v).comp <| isometry_pos_mul u)⟩ +section slices + +/-- The vertical strip of width A and height B. -/ +def verticalStrip (A B : ℝ) := {z : ℍ | |z.re| ≤ A ∧ B ≤ z.im} + +theorem slice_mem_iff (A B : ℝ) (z : ℍ) : z ∈ verticalStrip A B ↔ |z.re| ≤ A ∧ B ≤ z.im := + Iff.rfl + +lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : + ∃ A B : ℝ, 0 < B ∧ K ⊆ verticalStrip A B := by + rcases K.eq_empty_or_nonempty with rfl | hne + · exact ⟨1, 1, Real.zero_lt_one, empty_subset _⟩ + obtain ⟨u, _, hu⟩ := hK.exists_isMinOn hne continuous_im.continuousOn + obtain ⟨v, _, hv⟩ := hK.exists_isMaxOn hne (_root_.continuous_abs.comp continuous_re).continuousOn + exact ⟨|re v|, im u, u.im_pos, fun k hk ↦ ⟨isMaxOn_iff.mp hv _ hk, isMinOn_iff.mp hu _ hk⟩⟩ + +end slices + end UpperHalfPlane From a243ce69b070ec0bf912344a486dd8c72d11e496 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 09:45:27 +0100 Subject: [PATCH 089/124] doc string --- Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 52456d152a314..8f608e4ac41bf 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -137,6 +137,7 @@ theorem ne_zero (z : ℍ) : (z : ℂ) ≠ 0 := mt (congr_arg Complex.im) z.im_ne_zero #align upper_half_plane.ne_zero UpperHalfPlane.ne_zero +/-- Define √-1 as an element on the upper half plane.-/ def I : ℍ := ⟨Complex.I, by simp⟩ @[simp] From f18cfea10b4c222fd7f317797a40b05d3e841f70 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 10:20:08 +0100 Subject: [PATCH 090/124] move strips into topology --- .../Complex/UpperHalfPlane/Metric.lean | 18 ------------------ .../Complex/UpperHalfPlane/Topology.lean | 18 ++++++++++++++++++ 2 files changed, 18 insertions(+), 18 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index 961e40fbec3c9..ce8903b16fbc2 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -385,22 +385,4 @@ instance : IsometricSMul SL(2, ℝ) ℍ := exact (isometry_real_vadd w).comp (h₀.comp <| (isometry_real_vadd v).comp <| isometry_pos_mul u)⟩ -section slices - -/-- The vertical strip of width A and height B. -/ -def verticalStrip (A B : ℝ) := {z : ℍ | |z.re| ≤ A ∧ B ≤ z.im} - -theorem slice_mem_iff (A B : ℝ) (z : ℍ) : z ∈ verticalStrip A B ↔ |z.re| ≤ A ∧ B ≤ z.im := - Iff.rfl - -lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : - ∃ A B : ℝ, 0 < B ∧ K ⊆ verticalStrip A B := by - rcases K.eq_empty_or_nonempty with rfl | hne - · exact ⟨1, 1, Real.zero_lt_one, empty_subset _⟩ - obtain ⟨u, _, hu⟩ := hK.exists_isMinOn hne continuous_im.continuousOn - obtain ⟨v, _, hv⟩ := hK.exists_isMaxOn hne (_root_.continuous_abs.comp continuous_re).continuousOn - exact ⟨|re v|, im u, u.im_pos, fun k hk ↦ ⟨isMaxOn_iff.mp hv _ hk, isMinOn_iff.mp hu _ hk⟩⟩ - -end slices - end UpperHalfPlane diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index e6f35bacdfa58..fac32503f7e7b 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -74,4 +74,22 @@ instance : NoncompactSpace ℍ := by instance : LocallyCompactSpace ℍ := openEmbedding_coe.locallyCompactSpace +section slices + +/-- The vertical strip of width A and height B. -/ +def verticalStrip (A B : ℝ) := {z : ℍ | |z.re| ≤ A ∧ B ≤ z.im} + +theorem slice_mem_iff (A B : ℝ) (z : ℍ) : z ∈ verticalStrip A B ↔ |z.re| ≤ A ∧ B ≤ z.im := + Iff.rfl + +lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : + ∃ A B : ℝ, 0 < B ∧ K ⊆ verticalStrip A B := by + rcases K.eq_empty_or_nonempty with rfl | hne + · exact ⟨1, 1, Real.zero_lt_one, empty_subset _⟩ + obtain ⟨u, _, hu⟩ := hK.exists_isMinOn hne continuous_im.continuousOn + obtain ⟨v, _, hv⟩ := hK.exists_isMaxOn hne (_root_.continuous_abs.comp continuous_re).continuousOn + exact ⟨|re v|, im u, u.im_pos, fun k hk ↦ ⟨isMaxOn_iff.mp hv _ hk, isMinOn_iff.mp hu _ hk⟩⟩ + +end slices + end UpperHalfPlane From 61b00563977f09fca3b258dee0767ba1454fcd30 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 11:06:44 +0100 Subject: [PATCH 091/124] summable lems move test --- .../Analysis/Normed/Field/InfiniteSum.lean | 55 +++++++++++++++++++ Mathlib/Data/Set/Lattice.lean | 8 +++ Mathlib/Topology/Instances/ENNReal.lean | 47 ---------------- 3 files changed, 63 insertions(+), 47 deletions(-) diff --git a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean index b76c8f580f3f5..b47f56223a038 100644 --- a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean +++ b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean @@ -30,6 +30,61 @@ open Finset /-! ### Arbitrary index types -/ +theorem not_summable_iff_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : + ¬Summable f ↔ Tendsto (fun n : ℕ => ∑ i in Finset.range n, f i) atTop atTop := by + lift f to ℕ → ℝ≥0 using hf + exact mod_cast NNReal.not_summable_iff_tendsto_nat_atTop +#align not_summable_iff_tendsto_nat_at_top_of_nonneg not_summable_iff_tendsto_nat_atTop_of_nonneg + +theorem summable_iff_not_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : + Summable f ↔ ¬Tendsto (fun n : ℕ => ∑ i in Finset.range n, f i) atTop atTop := by + rw [← not_iff_not, Classical.not_not, not_summable_iff_tendsto_nat_atTop_of_nonneg hf] +#align summable_iff_not_tendsto_nat_at_top_of_nonneg summable_iff_not_tendsto_nat_atTop_of_nonneg + +theorem summable_sigma_of_nonneg {β : α → Type*} {f : (Σ x, β x) → ℝ} (hf : ∀ x, 0 ≤ f x) : + Summable f ↔ (∀ x, Summable fun y => f ⟨x, y⟩) ∧ Summable fun x => ∑' y, f ⟨x, y⟩ := by + lift f to (Σx, β x) → ℝ≥0 using hf + exact mod_cast NNReal.summable_sigma +#align summable_sigma_of_nonneg summable_sigma_of_nonneg + +lemma summable_partition {κ ι: Type*} {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} + (hs : ∀ i, ∃! j, i ∈ s j) : Summable f ↔ + (∀ j, Summable fun i : s j ↦ f i) ∧ Summable fun j ↦ ∑' i : s j, f i := by + rw [← (Set.sigmaEquiv s hs).summable_iff, summable_sigma_of_nonneg] + simp only [Set.sigmaEquiv, Equiv.coe_fn_mk, Function.comp_apply] + exact fun _ ↦ hf _ + +theorem summable_prod_of_nonneg {f : (α × β) → ℝ} (hf : 0 ≤ f) : + Summable f ↔ (∀ x, Summable fun y ↦ f (x, y)) ∧ Summable fun x ↦ ∑' y, f (x, y) := + (Equiv.sigmaEquivProd _ _).summable_iff.symm.trans <| summable_sigma_of_nonneg fun _ ↦ hf _ + +theorem summable_of_sum_le {ι : Type*} {f : ι → ℝ} {c : ℝ} (hf : 0 ≤ f) + (h : ∀ u : Finset ι, ∑ x in u, f x ≤ c) : Summable f := + ⟨⨆ u : Finset ι, ∑ x in u, f x, + tendsto_atTop_ciSup (Finset.sum_mono_set_of_nonneg hf) ⟨c, fun _ ⟨u, hu⟩ => hu ▸ h u⟩⟩ +#align summable_of_sum_le summable_of_sum_le + +theorem summable_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) + (h : ∀ n, ∑ i in Finset.range n, f i ≤ c) : Summable f := by + refine (summable_iff_not_tendsto_nat_atTop_of_nonneg hf).2 fun H => ?_ + rcases exists_lt_of_tendsto_atTop H 0 c with ⟨n, -, hn⟩ + exact lt_irrefl _ (hn.trans_le (h n)) +#align summable_of_sum_range_le summable_of_sum_range_le + +theorem Real.tsum_le_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) + (h : ∀ n, ∑ i in Finset.range n, f i ≤ c) : ∑' n, f n ≤ c := + _root_.tsum_le_of_sum_range_le (summable_of_sum_range_le hf h) h +#align real.tsum_le_of_sum_range_le Real.tsum_le_of_sum_range_le + +/-- If a sequence `f` with non-negative terms is dominated by a sequence `g` with summable +series and at least one term of `f` is strictly smaller than the corresponding term in `g`, +then the series of `f` is strictly smaller than the series of `g`. -/ +theorem tsum_lt_tsum_of_nonneg {i : ℕ} {f g : ℕ → ℝ} (h0 : ∀ b : ℕ, 0 ≤ f b) + (h : ∀ b : ℕ, f b ≤ g b) (hi : f i < g i) (hg : Summable g) : ∑' n, f n < ∑' n, g n := + tsum_lt_tsum h hi (.of_nonneg_of_le h0 h hg) hg +#align tsum_lt_tsum_of_nonneg tsum_lt_tsum_of_nonneg + + theorem Summable.mul_of_nonneg {f : ι → ℝ} {g : ι' → ℝ} (hf : Summable f) (hg : Summable g) (hf' : 0 ≤ f) (hg' : 0 ≤ g) : Summable fun x : ι × ι' => f x.1 * g x.2 := (summable_prod_of_nonneg fun _ ↦ mul_nonneg (hf' _) (hg' _)).2 ⟨fun x ↦ hg.mul_left (f x), diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index b895eb939e20d..429cd1efaefe9 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -2197,6 +2197,14 @@ theorem sigmaToiUnion_bijective (h : Pairwise fun i j => Disjoint (t i) (t j)) : ⟨sigmaToiUnion_injective t h, sigmaToiUnion_surjective t⟩ #align set.sigma_to_Union_bijective Set.sigmaToiUnion_bijective +/-- Equivalence between the sigma of a family of sets of `β` and `β`. -/ +noncomputable def sigmaEquiv {ι κ : Type*} (s : κ → Set ι) (hs : ∀ i, ∃! j, i ∈ s j) : + (Σ j, s j) ≃ ι where + toFun x := x.2 + invFun x := ⟨(hs x).choose, x, (hs x).choose_spec.1⟩ + left_inv x := by ext; exacts [((hs x.2).choose_spec.2 x.1 x.2.2).symm, rfl] + right_inv x := by rfl + /-- Equivalence between a disjoint union and a dependent sum. -/ noncomputable def unionEqSigmaOfDisjoint {t : α → Set β} (h : Pairwise fun i j => Disjoint (t i) (t j)) : diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 85cbf97b780a6..a29f41c272fca 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -1352,53 +1352,6 @@ theorem ENNReal.ofReal_tsum_of_nonneg {f : α → ℝ} (hf_nonneg : ∀ n, 0 ≤ simp_rw [ENNReal.ofReal, ENNReal.tsum_coe_eq (NNReal.hasSum_real_toNNReal_of_nonneg hf_nonneg hf)] #align ennreal.of_real_tsum_of_nonneg ENNReal.ofReal_tsum_of_nonneg -theorem not_summable_iff_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : - ¬Summable f ↔ Tendsto (fun n : ℕ => ∑ i in Finset.range n, f i) atTop atTop := by - lift f to ℕ → ℝ≥0 using hf - exact mod_cast NNReal.not_summable_iff_tendsto_nat_atTop -#align not_summable_iff_tendsto_nat_at_top_of_nonneg not_summable_iff_tendsto_nat_atTop_of_nonneg - -theorem summable_iff_not_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : - Summable f ↔ ¬Tendsto (fun n : ℕ => ∑ i in Finset.range n, f i) atTop atTop := by - rw [← not_iff_not, Classical.not_not, not_summable_iff_tendsto_nat_atTop_of_nonneg hf] -#align summable_iff_not_tendsto_nat_at_top_of_nonneg summable_iff_not_tendsto_nat_atTop_of_nonneg - -theorem summable_sigma_of_nonneg {β : α → Type*} {f : (Σ x, β x) → ℝ} (hf : ∀ x, 0 ≤ f x) : - Summable f ↔ (∀ x, Summable fun y => f ⟨x, y⟩) ∧ Summable fun x => ∑' y, f ⟨x, y⟩ := by - lift f to (Σx, β x) → ℝ≥0 using hf - exact mod_cast NNReal.summable_sigma -#align summable_sigma_of_nonneg summable_sigma_of_nonneg - -theorem summable_prod_of_nonneg {f : (α × β) → ℝ} (hf : 0 ≤ f) : - Summable f ↔ (∀ x, Summable fun y ↦ f (x, y)) ∧ Summable fun x ↦ ∑' y, f (x, y) := - (Equiv.sigmaEquivProd _ _).summable_iff.symm.trans <| summable_sigma_of_nonneg fun _ ↦ hf _ - -theorem summable_of_sum_le {ι : Type*} {f : ι → ℝ} {c : ℝ} (hf : 0 ≤ f) - (h : ∀ u : Finset ι, ∑ x in u, f x ≤ c) : Summable f := - ⟨⨆ u : Finset ι, ∑ x in u, f x, - tendsto_atTop_ciSup (Finset.sum_mono_set_of_nonneg hf) ⟨c, fun _ ⟨u, hu⟩ => hu ▸ h u⟩⟩ -#align summable_of_sum_le summable_of_sum_le - -theorem summable_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) - (h : ∀ n, ∑ i in Finset.range n, f i ≤ c) : Summable f := by - refine (summable_iff_not_tendsto_nat_atTop_of_nonneg hf).2 fun H => ?_ - rcases exists_lt_of_tendsto_atTop H 0 c with ⟨n, -, hn⟩ - exact lt_irrefl _ (hn.trans_le (h n)) -#align summable_of_sum_range_le summable_of_sum_range_le - -theorem Real.tsum_le_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) - (h : ∀ n, ∑ i in Finset.range n, f i ≤ c) : ∑' n, f n ≤ c := - _root_.tsum_le_of_sum_range_le (summable_of_sum_range_le hf h) h -#align real.tsum_le_of_sum_range_le Real.tsum_le_of_sum_range_le - -/-- If a sequence `f` with non-negative terms is dominated by a sequence `g` with summable -series and at least one term of `f` is strictly smaller than the corresponding term in `g`, -then the series of `f` is strictly smaller than the series of `g`. -/ -theorem tsum_lt_tsum_of_nonneg {i : ℕ} {f g : ℕ → ℝ} (h0 : ∀ b : ℕ, 0 ≤ f b) - (h : ∀ b : ℕ, f b ≤ g b) (hi : f i < g i) (hg : Summable g) : ∑' n, f n < ∑' n, g n := - tsum_lt_tsum h hi (.of_nonneg_of_le h0 h hg) hg -#align tsum_lt_tsum_of_nonneg tsum_lt_tsum_of_nonneg - section variable [EMetricSpace β] From e5e78c026d24215e31b4ff550f1d62fdac0ba35b Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 11:27:54 +0100 Subject: [PATCH 092/124] test2 --- .../Analysis/Normed/Field/InfiniteSum.lean | 58 +------------ .../Topology/Algebra/InfiniteSum/Real.lean | 87 +++++++++++++------ Mathlib/Topology/Instances/ENNReal.lean | 22 ++++- 3 files changed, 84 insertions(+), 83 deletions(-) diff --git a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean index b47f56223a038..28a4922181ec9 100644 --- a/Mathlib/Analysis/Normed/Field/InfiniteSum.lean +++ b/Mathlib/Analysis/Normed/Field/InfiniteSum.lean @@ -5,7 +5,7 @@ Authors: Anatole Dedecker -/ import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Analysis.Normed.Group.InfiniteSum -import Mathlib.Topology.Instances.ENNReal +import Mathlib.Topology.Algebra.InfiniteSum.Real #align_import analysis.normed.field.infinite_sum from "leanprover-community/mathlib"@"008205aa645b3f194c1da47025c5f110c8406eab" @@ -30,61 +30,6 @@ open Finset /-! ### Arbitrary index types -/ -theorem not_summable_iff_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : - ¬Summable f ↔ Tendsto (fun n : ℕ => ∑ i in Finset.range n, f i) atTop atTop := by - lift f to ℕ → ℝ≥0 using hf - exact mod_cast NNReal.not_summable_iff_tendsto_nat_atTop -#align not_summable_iff_tendsto_nat_at_top_of_nonneg not_summable_iff_tendsto_nat_atTop_of_nonneg - -theorem summable_iff_not_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : - Summable f ↔ ¬Tendsto (fun n : ℕ => ∑ i in Finset.range n, f i) atTop atTop := by - rw [← not_iff_not, Classical.not_not, not_summable_iff_tendsto_nat_atTop_of_nonneg hf] -#align summable_iff_not_tendsto_nat_at_top_of_nonneg summable_iff_not_tendsto_nat_atTop_of_nonneg - -theorem summable_sigma_of_nonneg {β : α → Type*} {f : (Σ x, β x) → ℝ} (hf : ∀ x, 0 ≤ f x) : - Summable f ↔ (∀ x, Summable fun y => f ⟨x, y⟩) ∧ Summable fun x => ∑' y, f ⟨x, y⟩ := by - lift f to (Σx, β x) → ℝ≥0 using hf - exact mod_cast NNReal.summable_sigma -#align summable_sigma_of_nonneg summable_sigma_of_nonneg - -lemma summable_partition {κ ι: Type*} {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} - (hs : ∀ i, ∃! j, i ∈ s j) : Summable f ↔ - (∀ j, Summable fun i : s j ↦ f i) ∧ Summable fun j ↦ ∑' i : s j, f i := by - rw [← (Set.sigmaEquiv s hs).summable_iff, summable_sigma_of_nonneg] - simp only [Set.sigmaEquiv, Equiv.coe_fn_mk, Function.comp_apply] - exact fun _ ↦ hf _ - -theorem summable_prod_of_nonneg {f : (α × β) → ℝ} (hf : 0 ≤ f) : - Summable f ↔ (∀ x, Summable fun y ↦ f (x, y)) ∧ Summable fun x ↦ ∑' y, f (x, y) := - (Equiv.sigmaEquivProd _ _).summable_iff.symm.trans <| summable_sigma_of_nonneg fun _ ↦ hf _ - -theorem summable_of_sum_le {ι : Type*} {f : ι → ℝ} {c : ℝ} (hf : 0 ≤ f) - (h : ∀ u : Finset ι, ∑ x in u, f x ≤ c) : Summable f := - ⟨⨆ u : Finset ι, ∑ x in u, f x, - tendsto_atTop_ciSup (Finset.sum_mono_set_of_nonneg hf) ⟨c, fun _ ⟨u, hu⟩ => hu ▸ h u⟩⟩ -#align summable_of_sum_le summable_of_sum_le - -theorem summable_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) - (h : ∀ n, ∑ i in Finset.range n, f i ≤ c) : Summable f := by - refine (summable_iff_not_tendsto_nat_atTop_of_nonneg hf).2 fun H => ?_ - rcases exists_lt_of_tendsto_atTop H 0 c with ⟨n, -, hn⟩ - exact lt_irrefl _ (hn.trans_le (h n)) -#align summable_of_sum_range_le summable_of_sum_range_le - -theorem Real.tsum_le_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) - (h : ∀ n, ∑ i in Finset.range n, f i ≤ c) : ∑' n, f n ≤ c := - _root_.tsum_le_of_sum_range_le (summable_of_sum_range_le hf h) h -#align real.tsum_le_of_sum_range_le Real.tsum_le_of_sum_range_le - -/-- If a sequence `f` with non-negative terms is dominated by a sequence `g` with summable -series and at least one term of `f` is strictly smaller than the corresponding term in `g`, -then the series of `f` is strictly smaller than the series of `g`. -/ -theorem tsum_lt_tsum_of_nonneg {i : ℕ} {f g : ℕ → ℝ} (h0 : ∀ b : ℕ, 0 ≤ f b) - (h : ∀ b : ℕ, f b ≤ g b) (hi : f i < g i) (hg : Summable g) : ∑' n, f n < ∑' n, g n := - tsum_lt_tsum h hi (.of_nonneg_of_le h0 h hg) hg -#align tsum_lt_tsum_of_nonneg tsum_lt_tsum_of_nonneg - - theorem Summable.mul_of_nonneg {f : ι → ℝ} {g : ι' → ℝ} (hf : Summable f) (hg : Summable g) (hf' : 0 ≤ f) (hg' : 0 ≤ g) : Summable fun x : ι × ι' => f x.1 * g x.2 := (summable_prod_of_nonneg fun _ ↦ mul_nonneg (hf' _) (hg' _)).2 ⟨fun x ↦ hg.mul_left (f x), @@ -122,7 +67,6 @@ In order to avoid `Nat` subtraction, we also provide where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the `Finset` `Finset.antidiagonal n`. -/ - section Nat open Finset.Nat diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Real.lean b/Mathlib/Topology/Algebra/InfiniteSum/Real.lean index c14bce1e4ea00..77ad02c2490e9 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Real.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Real.lean @@ -6,41 +6,20 @@ Authors: Sébastien Gouëzel, Yury Kudryashov import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Topology.Algebra.InfiniteSum.Order import Mathlib.Topology.Instances.Real +import Mathlib.Topology.Instances.ENNReal #align_import topology.algebra.infinite_sum.real from "leanprover-community/mathlib"@"9a59dcb7a2d06bf55da57b9030169219980660cd" /-! # Infinite sum in the reals -This file provides lemmas about Cauchy sequences in terms of infinite sums. +This file provides lemmas about Cauchy sequences in terms of infinite sums and infinite sums valued +in the reals. -/ open Filter Finset BigOperators NNReal Topology -variable {α : Type*} - -/-- If the extended distance between consecutive points of a sequence is estimated -by a summable series of `NNReal`s, then the original sequence is a Cauchy sequence. -/ -theorem cauchySeq_of_edist_le_of_summable [PseudoEMetricSpace α] {f : ℕ → α} (d : ℕ → ℝ≥0) - (hf : ∀ n, edist (f n) (f n.succ) ≤ d n) (hd : Summable d) : CauchySeq f := by - refine EMetric.cauchySeq_iff_NNReal.2 fun ε εpos ↦ ?_ - -- Actually we need partial sums of `d` to be a Cauchy sequence - replace hd : CauchySeq fun n : ℕ ↦ ∑ x in range n, d x := - let ⟨_, H⟩ := hd - H.tendsto_sum_nat.cauchySeq - -- Now we take the same `N` as in one of the definitions of a Cauchy sequence - refine (Metric.cauchySeq_iff'.1 hd ε (NNReal.coe_pos.2 εpos)).imp fun N hN n hn ↦ ?_ - specialize hN n hn - -- We simplify the known inequality - rw [dist_nndist, NNReal.nndist_eq, ← sum_range_add_sum_Ico _ hn, add_tsub_cancel_left, - NNReal.coe_lt_coe, max_lt_iff] at hN - rw [edist_comm] - -- Then use `hf` to simplify the goal to the same form - refine lt_of_le_of_lt (edist_le_Ico_sum_of_edist_le hn fun _ _ ↦ hf _) ?_ - exact mod_cast hN.1 -#align cauchy_seq_of_edist_le_of_summable cauchySeq_of_edist_le_of_summable - -variable [PseudoMetricSpace α] {f : ℕ → α} {a : α} +variable {α β : Type*} [PseudoMetricSpace α] {f : ℕ → α} {a : α} /-- If the distance between consecutive points of a sequence is estimated by a summable series, then the original sequence is a Cauchy sequence. -/ @@ -94,3 +73,61 @@ theorem dist_le_tsum_dist_of_tendsto₀ (h : Summable fun n ↦ dist (f n) (f n. (ha : Tendsto f atTop (𝓝 a)) : dist (f 0) a ≤ ∑' n, dist (f n) (f n.succ) := by simpa only [zero_add] using dist_le_tsum_dist_of_tendsto h ha 0 #align dist_le_tsum_dist_of_tendsto₀ dist_le_tsum_dist_of_tendsto₀ + +section summable + +theorem not_summable_iff_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : + ¬Summable f ↔ Tendsto (fun n : ℕ => ∑ i in Finset.range n, f i) atTop atTop := by + lift f to ℕ → ℝ≥0 using hf + exact mod_cast NNReal.not_summable_iff_tendsto_nat_atTop +#align not_summable_iff_tendsto_nat_at_top_of_nonneg not_summable_iff_tendsto_nat_atTop_of_nonneg + +theorem summable_iff_not_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : + Summable f ↔ ¬Tendsto (fun n : ℕ => ∑ i in Finset.range n, f i) atTop atTop := by + rw [← not_iff_not, Classical.not_not, not_summable_iff_tendsto_nat_atTop_of_nonneg hf] +#align summable_iff_not_tendsto_nat_at_top_of_nonneg summable_iff_not_tendsto_nat_atTop_of_nonneg + +theorem summable_sigma_of_nonneg {β : α → Type*} {f : (Σ x, β x) → ℝ} (hf : ∀ x, 0 ≤ f x) : + Summable f ↔ (∀ x, Summable fun y => f ⟨x, y⟩) ∧ Summable fun x => ∑' y, f ⟨x, y⟩ := by + lift f to (Σx, β x) → ℝ≥0 using hf + exact mod_cast NNReal.summable_sigma +#align summable_sigma_of_nonneg summable_sigma_of_nonneg + +lemma summable_partition {κ ι: Type*} {f : ι → ℝ} (hf : 0 ≤ f) {s : κ → Set ι} + (hs : ∀ i, ∃! j, i ∈ s j) : Summable f ↔ + (∀ j, Summable fun i : s j ↦ f i) ∧ Summable fun j ↦ ∑' i : s j, f i := by + rw [← (Set.sigmaEquiv s hs).summable_iff, summable_sigma_of_nonneg] + simp only [Set.sigmaEquiv, Equiv.coe_fn_mk, Function.comp_apply] + exact fun _ ↦ hf _ + +theorem summable_prod_of_nonneg {f : (α × β) → ℝ} (hf : 0 ≤ f) : + Summable f ↔ (∀ x, Summable fun y ↦ f (x, y)) ∧ Summable fun x ↦ ∑' y, f (x, y) := + (Equiv.sigmaEquivProd _ _).summable_iff.symm.trans <| summable_sigma_of_nonneg fun _ ↦ hf _ + +theorem summable_of_sum_le {ι : Type*} {f : ι → ℝ} {c : ℝ} (hf : 0 ≤ f) + (h : ∀ u : Finset ι, ∑ x in u, f x ≤ c) : Summable f := + ⟨⨆ u : Finset ι, ∑ x in u, f x, + tendsto_atTop_ciSup (Finset.sum_mono_set_of_nonneg hf) ⟨c, fun _ ⟨u, hu⟩ => hu ▸ h u⟩⟩ +#align summable_of_sum_le summable_of_sum_le + +theorem summable_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) + (h : ∀ n, ∑ i in Finset.range n, f i ≤ c) : Summable f := by + refine (summable_iff_not_tendsto_nat_atTop_of_nonneg hf).2 fun H => ?_ + rcases exists_lt_of_tendsto_atTop H 0 c with ⟨n, -, hn⟩ + exact lt_irrefl _ (hn.trans_le (h n)) +#align summable_of_sum_range_le summable_of_sum_range_le + +theorem Real.tsum_le_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) + (h : ∀ n, ∑ i in Finset.range n, f i ≤ c) : ∑' n, f n ≤ c := + _root_.tsum_le_of_sum_range_le (summable_of_sum_range_le hf h) h +#align real.tsum_le_of_sum_range_le Real.tsum_le_of_sum_range_le + +/-- If a sequence `f` with non-negative terms is dominated by a sequence `g` with summable +series and at least one term of `f` is strictly smaller than the corresponding term in `g`, +then the series of `f` is strictly smaller than the series of `g`. -/ +theorem tsum_lt_tsum_of_nonneg {i : ℕ} {f g : ℕ → ℝ} (h0 : ∀ b : ℕ, 0 ≤ f b) + (h : ∀ b : ℕ, f b ≤ g b) (hi : f i < g i) (hg : Summable g) : ∑' n, f n < ∑' n, g n := + tsum_lt_tsum h hi (.of_nonneg_of_le h0 h hg) hg +#align tsum_lt_tsum_of_nonneg tsum_lt_tsum_of_nonneg + +end summable diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index a29f41c272fca..e0778fd720e54 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Topology.Order.MonotoneContinuity -import Mathlib.Topology.Algebra.InfiniteSum.Real import Mathlib.Topology.Algebra.Order.LiminfLimsup import Mathlib.Topology.Instances.NNReal import Mathlib.Topology.EMetricSpace.Lipschitz @@ -1450,6 +1449,27 @@ theorem Filter.Tendsto.edist {f g : β → α} {x : Filter β} {a b : α} (hf : (continuous_edist.tendsto (a, b)).comp (hf.prod_mk_nhds hg) #align filter.tendsto.edist Filter.Tendsto.edist +/-- If the extended distance between consecutive points of a sequence is estimated +by a summable series of `NNReal`s, then the original sequence is a Cauchy sequence. -/ +theorem cauchySeq_of_edist_le_of_summable [PseudoEMetricSpace α] {f : ℕ → α} (d : ℕ → ℝ≥0) + (hf : ∀ n, edist (f n) (f n.succ) ≤ d n) (hd : Summable d) : CauchySeq f := by + refine EMetric.cauchySeq_iff_NNReal.2 fun ε εpos ↦ ?_ + -- Actually we need partial sums of `d` to be a Cauchy sequence + replace hd : CauchySeq fun n : ℕ ↦ ∑ x in Finset.range n, d x := + let ⟨_, H⟩ := hd + H.tendsto_sum_nat.cauchySeq + -- Now we take the same `N` as in one of the definitions of a Cauchy sequence + refine (Metric.cauchySeq_iff'.1 hd ε (NNReal.coe_pos.2 εpos)).imp fun N hN n hn ↦ ?_ + specialize hN n hn + -- We simplify the known inequality + rw [dist_nndist, NNReal.nndist_eq, ← Finset.sum_range_add_sum_Ico _ hn, add_tsub_cancel_left, + NNReal.coe_lt_coe, max_lt_iff] at hN + rw [edist_comm] + -- Then use `hf` to simplify the goal to the same form + refine lt_of_le_of_lt (edist_le_Ico_sum_of_edist_le hn fun _ _ ↦ hf _) ?_ + exact mod_cast hN.1 +#align cauchy_seq_of_edist_le_of_summable cauchySeq_of_edist_le_of_summable + theorem cauchySeq_of_edist_le_of_tsum_ne_top {f : ℕ → α} (d : ℕ → ℝ≥0∞) (hf : ∀ n, edist (f n) (f n.succ) ≤ d n) (hd : tsum d ≠ ∞) : CauchySeq f := by lift d to ℕ → NNReal using fun i => ENNReal.ne_top_of_tsum_ne_top hd i From a8380aebf7a771111f7a1b13d6bbd05311022254 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 11:40:15 +0100 Subject: [PATCH 093/124] do todo from port --- .../Topology/Algebra/InfiniteSum/Real.lean | 22 +++++-------------- 1 file changed, 5 insertions(+), 17 deletions(-) diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Real.lean b/Mathlib/Topology/Algebra/InfiniteSum/Real.lean index 77ad02c2490e9..92d6293d370d6 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Real.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Real.lean @@ -25,23 +25,11 @@ variable {α β : Type*} [PseudoMetricSpace α] {f : ℕ → α} {a : α} then the original sequence is a Cauchy sequence. -/ theorem cauchySeq_of_dist_le_of_summable (d : ℕ → ℝ) (hf : ∀ n, dist (f n) (f n.succ) ≤ d n) (hd : Summable d) : CauchySeq f := by - -- Porting note (#11215): TODO: with `Topology.Instances.NNReal` we can use this: - -- lift d to ℕ → ℝ≥0 using fun n ↦ dist_nonneg.trans (hf n) - -- refine cauchySeq_of_edist_le_of_summable d ?_ ?_ - -- · exact_mod_cast hf - -- · exact_mod_cast hd - refine' Metric.cauchySeq_iff'.2 fun ε εpos ↦ _ - replace hd : CauchySeq fun n : ℕ ↦ ∑ x in range n, d x := - let ⟨_, H⟩ := hd - H.tendsto_sum_nat.cauchySeq - refine' (Metric.cauchySeq_iff'.1 hd ε εpos).imp fun N hN n hn ↦ _ - have hsum := hN n hn - rw [Real.dist_eq, ← sum_Ico_eq_sub _ hn] at hsum - calc - dist (f n) (f N) = dist (f N) (f n) := dist_comm _ _ - _ ≤ ∑ x in Ico N n, d x := dist_le_Ico_sum_of_dist_le hn fun _ _ ↦ hf _ - _ ≤ |∑ x in Ico N n, d x| := le_abs_self _ - _ < ε := hsum + lift d to ℕ → ℝ≥0 using fun n ↦ dist_nonneg.trans (hf n) + apply cauchySeq_of_edist_le_of_summable d (α := α) (f := f) + · exact_mod_cast hf + · exact_mod_cast hd + #align cauchy_seq_of_dist_le_of_summable cauchySeq_of_dist_le_of_summable theorem cauchySeq_of_summable_dist (h : Summable fun n ↦ dist (f n) (f n.succ)) : CauchySeq f := From 95f199f02cc680aa1f4140827d0253d5b5fd1928 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 12:14:24 +0100 Subject: [PATCH 094/124] build fix --- Mathlib/Analysis/SpecificLimits/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 50df1409ebdf7..7f03cf063746c 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel, Johannes Hölzl, Yury G. Kudryashov, Patrick Masso import Mathlib.Algebra.GeomSum import Mathlib.Order.Filter.Archimedean import Mathlib.Order.Iterate -import Mathlib.Topology.Instances.ENNReal +import Mathlib.Topology.Algebra.InfiniteSum.Real import Mathlib.Topology.Algebra.Algebra #align_import analysis.specific_limits.basic from "leanprover-community/mathlib"@"57ac39bd365c2f80589a700f9fbb664d3a1a30c2" From 0472942653fe3639268eb1e5e2c8625aa00943ea Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 13:09:58 +0100 Subject: [PATCH 095/124] fix --- Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean | 2 +- Mathlib/Analysis/SpecificLimits/Basic.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean index 093702ce769fb..fc500d57c253f 100644 --- a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean +++ b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Manuel Candales. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Manuel Candales -/ -import Mathlib.Topology.Instances.ENNReal +import Mathlib.Topology.Algebra.InfiniteSum.Real import Mathlib.Data.Nat.Squarefree #align_import wiedijk_100_theorems.sum_of_prime_reciprocals_diverges from "leanprover-community/mathlib"@"5563b1b49e86e135e8c7b556da5ad2f5ff881cad" diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 50df1409ebdf7..7f03cf063746c 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel, Johannes Hölzl, Yury G. Kudryashov, Patrick Masso import Mathlib.Algebra.GeomSum import Mathlib.Order.Filter.Archimedean import Mathlib.Order.Iterate -import Mathlib.Topology.Instances.ENNReal +import Mathlib.Topology.Algebra.InfiniteSum.Real import Mathlib.Topology.Algebra.Algebra #align_import analysis.specific_limits.basic from "leanprover-community/mathlib"@"57ac39bd365c2f80589a700f9fbb664d3a1a30c2" From ac145a12800d8aa30bea7d1c4aab1b2c4533585f Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 15:35:10 +0100 Subject: [PATCH 096/124] merge fix --- .../Complex/UpperHalfPlane/Metric.lean | 36 ------------------- .../Complex/UpperHalfPlane/Topology.lean | 8 ++--- .../EisensteinSeries/UniformConvergence.lean | 6 ++-- 3 files changed, 7 insertions(+), 43 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index 0479358f18c05..88b19a7e71485 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -387,40 +387,4 @@ instance : IsometricSMul SL(2, ℝ) ℍ := section slices -/--The vertical strip of width A and height B-/ -def upperHalfPlaneSlice (A B : ℝ) := - {z : ℍ | Complex.abs z.1.1 ≤ A ∧ B ≤ Complex.abs z.1.2} - -theorem slice_mem_iff (A B : ℝ) (z : ℍ) : - z ∈ upperHalfPlaneSlice A B ↔ Complex.abs z.1.1 ≤ A ∧ B ≤ Complex.abs z.1.2 := Iff.rfl - -lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : - ∃ A B : ℝ, 0 < B ∧ K ⊆ upperHalfPlaneSlice A B := by - obtain rfl | hne := K.eq_empty_or_nonempty - · exact ⟨1, 1, Real.zero_lt_one, by simp⟩ - have hcts : ContinuousOn (fun t => t.im) K := UpperHalfPlane.continuous_im.continuousOn - obtain ⟨b, _, HB⟩ := IsCompact.exists_isMinOn hK hne hcts - obtain ⟨r, _, hr2⟩ := Bornology.IsBounded.subset_closedBall_lt hK.isBounded 0 UpperHalfPlane.I - refine ⟨Real.sinh r + Complex.abs (UpperHalfPlane.center UpperHalfPlane.I r), b.im, b.2, ?_⟩ - intro z hz - simp only [I_im, slice_mem_iff, abs_ofReal, ge_iff_le] at * - constructor - · have hr3 := hr2 hz - simp only [Metric.mem_closedBall] at hr3 - apply le_trans (abs_re_le_abs z) - have := Complex.abs.sub_le (z : ℂ) (UpperHalfPlane.center UpperHalfPlane.I r) 0 - simp only [sub_zero, ge_iff_le] at this - rw [dist_le_iff_dist_coe_center_le] at hr3 - apply le_trans this - simp only [I_im, one_mul, add_le_add_iff_right] at * - exact hr3 - · have hbz := HB hz - simp only [mem_setOf_eq, ge_iff_le] at * - convert hbz - rw [UpperHalfPlane.im] - apply abs_eq_self.mpr z.2.le - - -end slices - end UpperHalfPlane diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index fac32503f7e7b..381fdcefb9058 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -74,15 +74,15 @@ instance : NoncompactSpace ℍ := by instance : LocallyCompactSpace ℍ := openEmbedding_coe.locallyCompactSpace -section slices +section strips /-- The vertical strip of width A and height B. -/ def verticalStrip (A B : ℝ) := {z : ℍ | |z.re| ≤ A ∧ B ≤ z.im} -theorem slice_mem_iff (A B : ℝ) (z : ℍ) : z ∈ verticalStrip A B ↔ |z.re| ≤ A ∧ B ≤ z.im := +theorem strip_mem_iff (A B : ℝ) (z : ℍ) : z ∈ verticalStrip A B ↔ |z.re| ≤ A ∧ B ≤ z.im := Iff.rfl -lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : +lemma subset_strip_of_isCompact {K : Set ℍ} (hK : IsCompact K) : ∃ A B : ℝ, 0 < B ∧ K ⊆ verticalStrip A B := by rcases K.eq_empty_or_nonempty with rfl | hne · exact ⟨1, 1, Real.zero_lt_one, empty_subset _⟩ @@ -90,6 +90,6 @@ lemma subset_slice_of_isCompact {K : Set ℍ} (hK : IsCompact K) : obtain ⟨v, _, hv⟩ := hK.exists_isMaxOn hne (_root_.continuous_abs.comp continuous_re).continuousOn exact ⟨|re v|, im u, u.im_pos, fun k hk ↦ ⟨isMaxOn_iff.mp hv _ hk, isMinOn_iff.mp hu _ hk⟩⟩ -end slices +end strips end UpperHalfPlane diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index cef971aaa95d9..17ec74f145a43 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -176,7 +176,7 @@ theorem eis_is_bounded_on_box (k n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hx : (x simp only [map_pow, zpow_natCast, H] simpa using this -lemma r_lower_bound_on_slice {A B : ℝ} (h : 0 < B) (z : upperHalfPlaneSlice A B) : +lemma r_lower_bound_on_strip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by have hz := z.2 simp only [slice_mem_iff, abs_ofReal, ge_iff_le] at hz @@ -252,7 +252,7 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : (by simp only [Set.top_eq_univ, isOpen_univ]), eisensteinSeries_SIF] simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] intro K hK - obtain ⟨A, B, hB, HABK⟩ := subset_slice_of_isCompact hK + obtain ⟨A, B, hB, HABK⟩ := subset_strip_of_isCompact hK have hu : Summable fun x : (gammaSet N a) => (1/(r ⟨⟨A, B⟩, hB⟩) ^ k) * ((max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by apply (Summable.subtype (summable_upper_bound hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a)).congr @@ -272,7 +272,7 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : rw [inv_le_inv] · apply pow_le_pow_left (r_pos _).le rw [abs_of_pos (r_pos _)] - · exact r_lower_bound_on_slice hB ⟨x, HABK hx⟩ + · exact r_lower_bound_on_strip hB ⟨x, HABK hx⟩ · apply pow_pos (abs_pos.mpr (ne_of_gt (r_pos x))) · apply pow_pos (r_pos _) From 01a93573156aff510bf7cd6a753c042ccd95cb0b Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 15:37:57 +0100 Subject: [PATCH 097/124] more fixes --- Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean | 2 -- Mathlib/Logic/Equiv/Set.lean | 8 -------- 2 files changed, 10 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index 88b19a7e71485..ce8903b16fbc2 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -385,6 +385,4 @@ instance : IsometricSMul SL(2, ℝ) ℍ := exact (isometry_real_vadd w).comp (h₀.comp <| (isometry_real_vadd v).comp <| isometry_pos_mul u)⟩ -section slices - end UpperHalfPlane diff --git a/Mathlib/Logic/Equiv/Set.lean b/Mathlib/Logic/Equiv/Set.lean index 212c95b70ee28..f6406fb978388 100644 --- a/Mathlib/Logic/Equiv/Set.lean +++ b/Mathlib/Logic/Equiv/Set.lean @@ -179,14 +179,6 @@ def setProdEquivSigma {α β : Type*} (s : Set (α × β)) : right_inv := fun ⟨x, y, h⟩ => rfl #align equiv.set_prod_equiv_sigma Equiv.setProdEquivSigma -/-- Equivalence between the sigma of a family of finsets of `β` and `β`. -/ -noncomputable def sigmaEquiv {ι κ : Type*} (s : κ → Set ι) (hs : ∀ i, ∃! j, i ∈ s j) : - (Σ j, s j) ≃ ι where - toFun x := x.2 - invFun x := ⟨(hs x).choose, x, (hs x).choose_spec.1⟩ - left_inv x := by ext; exacts [((hs x.2).choose_spec.2 x.1 x.2.2).symm, rfl] - right_inv x := by rfl - /-- The subtypes corresponding to equal sets are equivalent. -/ @[simps! apply] def setCongr {α : Type*} {s t : Set α} (h : s = t) : s ≃ t := From 3a77b2fe69a80a71d4b537d2f8b24c98c28babc2 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 15:39:56 +0100 Subject: [PATCH 098/124] fix docstring --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 17ec74f145a43..d9dadf72e8ab6 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -13,8 +13,8 @@ import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic /-! # Uniform convergence of Eisenstein series -We show that `eisSummand` converges locally uniformly on `ℍ` to the Eisenstein series of weight `k` -and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. +We show that the sum of `eisSummand` converges locally uniformly on `ℍ` to the Eisenstein series +of weight `k` and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. -/ noncomputable section From 10365c4ba37a8d9d561ce1295fe29d301465e34e Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 15:49:18 +0100 Subject: [PATCH 099/124] build fixes --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index d9dadf72e8ab6..0ba7c8fa3762b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -82,8 +82,7 @@ lemma auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε ^ 2) : r z ≤ C nlinarith right simp only [ne_eq, coe_re, coe_im, normSq_apply, AbsoluteValue.coe_mk, MulHom.coe_mk, add_re, - mul_re, ofReal_re, ofReal_im, zero_mul, sub_zero, int_cast_re, add_im, mul_im, add_zero, - int_cast_im] at * + mul_re, ofReal_re, ofReal_im, zero_mul, sub_zero, add_im, mul_im, add_zero] at * exact H1 lemma ne_zero_if_max {x : Fin 2 → ℤ} (hx : x ≠ 0) @@ -120,7 +119,7 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x (((x 0 : ℝ) / (n : ℝ)) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by have : (n : ℂ) ≠ 0 := by norm_cast field_simp [this] - simp only [Nat.cast_max, h11, ofReal_int_cast, map_mul, abs_ofReal, ge_iff_le] + simp only [Nat.cast_max, h11, map_mul, abs_ofReal, ge_iff_le] rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] cases' (div_max_sq_ge_one x hx) with H1 H2 · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) @@ -179,12 +178,11 @@ theorem eis_is_bounded_on_box (k n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hx : (x lemma r_lower_bound_on_strip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by have hz := z.2 - simp only [slice_mem_iff, abs_ofReal, ge_iff_le] at hz + simp only [strip_mem_iff, abs_ofReal, ge_iff_le] at hz simp_rw [r] apply min_le_min · dsimp only convert hz.2 - apply (abs_eq_self.mpr (UpperHalfPlane.im_pos z.1).le).symm · rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] simp only [r1', div_pow, one_div] rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] From 105e5d5e324256132e8094514a14b074c7a7bbb0 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 26 Apr 2024 16:03:34 +0100 Subject: [PATCH 100/124] fix import --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 0ba7c8fa3762b..a904636898b9c 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ -import Mathlib.Analysis.Complex.UpperHalfPlane.Metric +import Mathlib.Analysis.Complex.UpperHalfPlane.Topology import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.PSeries import Mathlib.Order.Interval.Finset.Box From 8891f5ae7b10dc41b3335900fa370c27dbd48eac Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 30 Apr 2024 13:33:33 +0100 Subject: [PATCH 101/124] remove junk complex.abs --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index a904636898b9c..903ac28586bd1 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -164,7 +164,7 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) exact Nat.pos_of_ne_zero hn theorem eis_is_bounded_on_box (k n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : - (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (Complex.abs ((r z) ^ k * n ^ k))⁻¹ := by + (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (((r z) ^ k * n ^ k))⁻¹ := by have := eis_is_bounded_on_box_rpow (Nat.cast_nonneg k) z n x hx norm_cast at * simp_rw [zpow_neg, ← mul_inv] at this @@ -269,9 +269,8 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : inv_nonneg, pow_nonneg (r_pos _).le]} rw [inv_le_inv] · apply pow_le_pow_left (r_pos _).le - rw [abs_of_pos (r_pos _)] · exact r_lower_bound_on_strip hB ⟨x, HABK hx⟩ - · apply pow_pos (abs_pos.mpr (ne_of_gt (r_pos x))) + · apply pow_pos ( (r_pos x)) · apply pow_pos (r_pos _) local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm @@ -280,7 +279,7 @@ local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm /- A version for the extension to maps `ℂ → ℂ` that is nice to have for holomorphicity later -/ lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} {N : ℕ} (hk : 3 ≤ k) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => - ↑ₕ(fun (z : ℍ) => ∑ x in s, eisSummand k x z )) (↑ₕ((eisensteinSeries_SIF a k).toFun )) + ↑ₕ(fun (z : ℍ) => ∑ x in s, eisSummand k x z )) (↑ₕ(eisensteinSeries_SIF a k).toFun) Filter.atTop (UpperHalfPlane.coe '' ⊤) := by apply TendstoLocallyUniformlyOn.comp (s := ⊤) simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ] From 9c1c1d2791935f4bd714cc527ace08f75ab0b37c Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 30 Apr 2024 14:01:30 +0100 Subject: [PATCH 102/124] fix --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 6 +----- 1 file changed, 1 insertion(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 903ac28586bd1..2693669c007f5 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -168,11 +168,7 @@ theorem eis_is_bounded_on_box (k n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hx : (x have := eis_is_bounded_on_box_rpow (Nat.cast_nonneg k) z n x hx norm_cast at * simp_rw [zpow_neg, ← mul_inv] at this - have H : |r z ^ k * ↑(n ^ k)| = r z ^ k * ↑(n ^ k) := by - refine abs_eq_self.mpr ?_ - apply mul_nonneg (pow_nonneg (r_pos z).le _) - simp only [Nat.cast_pow, ge_iff_le, Nat.cast_nonneg, pow_nonneg] - simp only [map_pow, zpow_natCast, H] + simp only [map_pow, zpow_natCast] simpa using this lemma r_lower_bound_on_strip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : From e6b0cdfcbb91b4b14705b27290db9a523898698c Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 1 May 2024 19:51:58 +0100 Subject: [PATCH 103/124] updates --- .../EisensteinSeries/UniformConvergence.lean | 23 +++++++++---------- 1 file changed, 11 insertions(+), 12 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 2693669c007f5..7142cd064f35a 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -163,12 +163,12 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) norm_cast exact Nat.pos_of_ne_zero hn -theorem eis_is_bounded_on_box (k n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : +/-This is a special case of the above, but one that we use more.-/ +theorem eis_is_bounded_on_box {k: ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (((r z) ^ k * n ^ k))⁻¹ := by - have := eis_is_bounded_on_box_rpow (Nat.cast_nonneg k) z n x hx + have := eis_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx norm_cast at * - simp_rw [zpow_neg, ← mul_inv] at this - simp only [map_pow, zpow_natCast] + simp_rw [zpow_neg, ← mul_inv] at * simpa using this lemma r_lower_bound_on_strip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : @@ -240,8 +240,6 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : (a : Fin 2 → ZMod N) : TendstoLocallyUniformly (fun (s : Finset (gammaSet N a)) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop := by - have hk0 : 0 ≤ k := by linarith - lift k to ℕ using hk0 rw [← tendstoLocallyUniformlyOn_univ,tendstoLocallyUniformlyOn_iff_forall_isCompact (by simp only [Set.top_eq_univ, isOpen_univ]), eisensteinSeries_SIF] simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] @@ -254,18 +252,19 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : simp only [zpow_natCast, one_div, Function.comp_apply] apply tendstoUniformlyOn_tsum hu intro v x hx - have := eis_is_bounded_on_box k (max (v.1 0).natAbs (v.1 1).natAbs) x v + have := eis_is_bounded_on_box (k := k) (max (v.1 0).natAbs (v.1 1).natAbs) x v (by linarith) simp only [Nat.cast_max,Int.natCast_natAbs, iff_true, zpow_natCast, one_div, map_pow, map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, ge_iff_le] at * apply le_trans (this (by simp only [Int.mem_box])) rw [mul_comm] apply mul_le_mul _ (by rfl) - repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, pow_nonneg, - inv_nonneg, pow_nonneg (r_pos _).le]} - rw [inv_le_inv] - · apply pow_le_pow_left (r_pos _).le - · exact r_lower_bound_on_strip hB ⟨x, HABK hx⟩ + repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg, + inv_nonneg, zpow_nonneg (r_pos _).le]} + have hk0 : 0 ≤ k := by linarith + lift k to ℕ using hk0 + rw [inv_le_inv ] + · apply pow_le_pow_left (r_pos _).le (r_lower_bound_on_strip hB ⟨x, HABK hx⟩) · apply pow_pos ( (r_pos x)) · apply pow_pos (r_pos _) From 1e48883b34e11ddfef2f789378d91fc508fdb7ce Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 1 May 2024 19:55:48 +0100 Subject: [PATCH 104/124] fix2 --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 7142cd064f35a..48a351c140b9e 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -174,7 +174,7 @@ theorem eis_is_bounded_on_box {k: ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) lemma r_lower_bound_on_strip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by have hz := z.2 - simp only [strip_mem_iff, abs_ofReal, ge_iff_le] at hz + simp only [mem_verticalStrip_iff, abs_ofReal, ge_iff_le] at hz simp_rw [r] apply min_le_min · dsimp only @@ -244,7 +244,7 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : (by simp only [Set.top_eq_univ, isOpen_univ]), eisensteinSeries_SIF] simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] intro K hK - obtain ⟨A, B, hB, HABK⟩ := subset_strip_of_isCompact hK + obtain ⟨A, B, hB, HABK⟩ := subset_verticalStrip_of_isCompact hK have hu : Summable fun x : (gammaSet N a) => (1/(r ⟨⟨A, B⟩, hB⟩) ^ k) * ((max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by apply (Summable.subtype (summable_upper_bound hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a)).congr From aeb0b434d6fd1af0e75ce4e0f1a997519aae72f8 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Thu, 2 May 2024 09:15:28 +0100 Subject: [PATCH 105/124] fix lintr --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 48a351c140b9e..fac176d3875ab 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -164,7 +164,8 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) exact Nat.pos_of_ne_zero hn /-This is a special case of the above, but one that we use more.-/ -theorem eis_is_bounded_on_box {k: ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) (hx : (x 0, x 1) ∈ box n) : +theorem eis_is_bounded_on_box {k: ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) + (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (((r z) ^ k * n ^ k))⁻¹ := by have := eis_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx norm_cast at * From 1c0523895b30b45d86811ff9081e3fde8c92d6aa Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Thu, 2 May 2024 15:47:23 +0100 Subject: [PATCH 106/124] updates --- .../EisensteinSeries/UniformConvergence.lean | 119 +++++++++--------- 1 file changed, 63 insertions(+), 56 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index fac176d3875ab..f5fc74b7c6f59 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -19,15 +19,15 @@ of weight `k` and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N` noncomputable section -open Complex Filter UpperHalfPlane Set Finset +open Complex UpperHalfPlane Set Finset -open scoped BigOperators NNReal Classical Filter Matrix UpperHalfPlane Complex +open scoped BigOperators Matrix UpperHalfPlane Complex namespace EisensteinSeries section bounding_functions -/-- Auxiliary function used for bounding Eisenstein series-/ +/-- Auxiliary function used for bounding Eisenstein series. -/ def r1 (z : ℍ) : ℝ := z.im ^ 2 / (z.re ^ 2 + z.im ^ 2) lemma r1' (z : ℍ) : r1 z = 1 / ((z.re / z.im) ^ 2 + 1) := by @@ -37,7 +37,7 @@ theorem r1_pos (z : ℍ) : 0 < r1 z := by dsimp only [r1] positivity -/-- This function is used to give an upper bound on Eisenstein series-/ +/-- This function is used to give an upper bound on Eisenstein series. -/ def r (z : ℍ) : ℝ := min z.im (Real.sqrt (r1 z)) lemma r_pos (z : ℍ) : 0 < r z := by @@ -91,9 +91,11 @@ lemma ne_zero_if_max {x : Fin 2 → ℤ} (hx : x ≠ 0) ext x fin_cases x <;> aesop +/-This proof is faster than the one above, but longer. -/ lemma ne_zero_if_max' {x : Fin 2 → ℤ} (hx : x ≠ 0) (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by - apply ne_zero_if_max (x := ![x 1, x 0]) ?_ (by simpa using h) + apply ne_zero_if_max (x := ![x 1, x 0]) ?_ (by simpa only [Fin.isValue, Matrix.cons_val_zero, + Matrix.cons_val_one, Matrix.head_cons, max_eq_left_iff, max_eq_right_iff] using h) simp only [ne_eq, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_true, not_and] intro h1 h0 rw [Function.ne_iff, Fin.exists_fin_two, h1, h0] at hx @@ -103,8 +105,12 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 ∨ (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 := by refine (max_choice (x 0).natAbs (x 1).natAbs).imp (fun H1 => ?_) (fun H2 => ?_) - · simp [H1, Int.cast_natAbs, ne_zero_if_max hx H1] - · simp [H2, Int.cast_natAbs, ne_zero_if_max' hx H2] + · simp only [Fin.isValue, H1, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, + OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, ne_zero_if_max hx H1, + div_self, le_refl] + · simp only [Fin.isValue, H2, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, + OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, + ne_zero_if_max' hx H2, div_self, le_refl] lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) : ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k ≤ @@ -112,9 +118,10 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x by_cases hk0 : k ≠ 0 · let n := max (x 0).natAbs (x 1).natAbs have hn0 : n ≠ 0 := by - rw [← Iff.ne ((eq_zero_iff_eq_zero_of_mem_box (α := ℤ × ℤ) (x := ((x 0, x 1)))) (by simp)), + rw [← Iff.ne ((eq_zero_iff_eq_zero_of_mem_box (α := ℤ × ℤ) (x := ((x 0, x 1)))) (by simp)), ← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun x ↦ ℤ)))] at * - simpa using hx + simpa only [Fin.isValue, ne_eq, Prod.mk_eq_zero, not_and, piFinTwoEquiv_apply, Pi.zero_apply, + Prod.mk.injEq] using hx have h11 : ((x 0) * ↑z + (x 1)) = (((x 0 : ℝ) / (n : ℝ)) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by have : (n : ℂ) ≠ 0 := by norm_cast @@ -123,17 +130,19 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] cases' (div_max_sq_ge_one x hx) with H1 H2 · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) - · simpa [n] using (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) + · simpa only [Fin.isValue, ofReal_intCast, Nat.cast_max, ofReal_div, n] using + (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) · positivity · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) - · simpa [n] using (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) + · simpa only [Fin.isValue, ofReal_intCast, Nat.cast_max, ofReal_div, n] using + (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) · positivity · simp only [ne_eq, not_not] at hk0 simp only [hk0, Real.rpow_zero, Nat.cast_max, mul_one, le_refl] -theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) +theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)))) ^ (-k) ≤ - (((r z) ^ (-k) * n ^ (-k))) := by + (r z) ^ (-k) * n ^ (-k) := by by_cases hn : n = 0 · simp only [hn, box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx rw [hx.1, hx.2] at * @@ -146,7 +155,9 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) simp only [h1, hn, CharP.cast_eq_zero, mul_zero, le_refl] · have hx2 : x ≠ 0 := by rw [← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℤ)))] - simpa using (Iff.ne ((eq_zero_iff_eq_zero_of_mem_box (x := (x 0, x 1))) hx)).mpr hn + simpa only [piFinTwoEquiv_apply, Fin.isValue, Pi.zero_apply, ne_eq, Prod.mk.injEq, not_and, + Prod.mk_eq_zero] using + (Iff.ne ((eq_zero_iff_eq_zero_of_mem_box (x := (x 0, x 1))) hx)).mpr hn simp only [Int.mem_box] at hx rw [Real.rpow_neg (by apply apply_nonneg), Real.rpow_neg ((r_pos z).le), Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv] @@ -155,37 +166,37 @@ theorem eis_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) · simp only [Nat.cast_max] · apply Real.rpow_pos_of_pos apply Complex.abs.pos (linear_ne_zero ![x 0, x 1] z ?_) - have := (Function.comp_ne_zero_iff x Int.cast_injective Int.cast_zero (γ := ℝ)).mpr hx2 + have := (Function.comp_ne_zero_iff x Int.cast_injective Int.cast_zero (γ := ℝ)).mpr hx2 rw [← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℝ)))] at this - simpa using this + simpa only [Fin.isValue, ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, + and_true, not_and, piFinTwoEquiv_apply, Function.comp_apply, Pi.zero_apply, + Prod.mk.injEq] using this · apply mul_pos (Real.rpow_pos_of_pos (r_pos z) _) apply Real.rpow_pos_of_pos norm_cast exact Nat.pos_of_ne_zero hn -/-This is a special case of the above, but one that we use more.-/ -theorem eis_is_bounded_on_box {k: ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) - (hx : (x 0, x 1) ∈ box n) : - (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)) ^ k))⁻¹ ≤ (((r z) ^ k * n ^ k))⁻¹ := by - have := eis_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx +/-This is a special case of the above, but one that we use more. -/ +theorem eisSummand_is_bounded_on_box {k: ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) + (hx : (x 0, x 1) ∈ box n) : Complex.abs (eisSummand k x z) ≤ (((r z) ^ k * n ^ k))⁻¹ := by + have := summand_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx norm_cast at * - simp_rw [zpow_neg, ← mul_inv] at * - simpa using this + rw [zpow_neg, mul_inv, eisSummand] at * + simpa only [Fin.isValue, one_div, map_inv₀, map_zpow₀, ge_iff_le, zpow_neg] using this -lemma r_lower_bound_on_strip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : +lemma r_lower_bound_on_verticalStrip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by have hz := z.2 simp only [mem_verticalStrip_iff, abs_ofReal, ge_iff_le] at hz - simp_rw [r] + rw [r] apply min_le_min - · dsimp only - convert hz.2 + · convert hz.2 · rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] simp only [r1', div_pow, one_div] rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] apply div_le_div (sq_nonneg _) _ (by positivity) - · simpa [even_two.pow_abs] using pow_le_pow_left h.le hz.2 2 - · simpa [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 + · simpa only using pow_le_pow_left h.le hz.2 2 + · simpa only [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 end bounding_functions @@ -216,8 +227,8 @@ lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): ring_nf lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => - (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by - set f := fun x : Fin 2 → ℤ ↦ (1 / (r z) ^ k) * ((max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ + (((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by + set f := fun x : Fin 2 → ℤ ↦ (((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ rw [← (piFinTwoEquiv _).symm.summable_iff, summable_partition _ (s := fun n ↦ (box n : Finset (ℤ × ℤ))) Int.existsUnique_mem_box] · simp_rw [coe_sort_coe, Finset.tsum_subtype] @@ -225,15 +236,17 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : refine ⟨fun n ↦ ?_, Summable.congr (summable_over_box z h) fun n ↦ Finset.sum_congr rfl fun x hx ↦ ?_⟩ · simpa using (box n).summable (f ∘ (piFinTwoEquiv _).symm) - · simp only [Int.mem_box] at hx - rw [← hx, one_div] - simp only [Nat.cast_max, one_div, Fin.isValue, Fin.cons_zero, Fin.cons_one, f] + · simp only [Int.mem_box, one_div] at * + simp only [← hx, Nat.cast_max, Fin.isValue, mul_inv_rev, Fin.cons_zero, Fin.cons_one, f] + rw [mul_comm] · intro y + simp only [Pi.zero_apply, Fin.isValue, mul_inv_rev, piFinTwoEquiv_symm_apply, + Function.comp_apply, Fin.cons_zero, Fin.cons_one, f] apply mul_nonneg - · simp only [one_div, inv_nonneg] - apply zpow_nonneg (r_pos z).le · simp only [piFinTwoEquiv_symm_apply, Fin.cons_zero, Fin.cons_one, inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] + · simp only [one_div, inv_nonneg] + apply zpow_nonneg (r_pos z).le end summability @@ -247,38 +260,32 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : intro K hK obtain ⟨A, B, hB, HABK⟩ := subset_verticalStrip_of_isCompact hK have hu : Summable fun x : (gammaSet N a) => - (1/(r ⟨⟨A, B⟩, hB⟩) ^ k) * ((max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by + (((r ⟨⟨A, B⟩, hB⟩) ^ k) * (max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by apply (Summable.subtype (summable_upper_bound hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a)).congr intro v simp only [zpow_natCast, one_div, Function.comp_apply] apply tendstoUniformlyOn_tsum hu intro v x hx - have := eis_is_bounded_on_box (k := k) (max (v.1 0).natAbs (v.1 1).natAbs) x v (by linarith) - simp only [Nat.cast_max,Int.natCast_natAbs, iff_true, zpow_natCast, one_div, map_pow, - map_mul, abs_ofReal, abs_natCast, mul_inv_rev, eisSummand, norm_inv, norm_pow, norm_eq_abs, - ge_iff_le] at * - apply le_trans (this (by simp only [Int.mem_box])) - rw [mul_comm] - apply mul_le_mul _ (by rfl) - repeat {simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg, - inv_nonneg, zpow_nonneg (r_pos _).le]} + apply le_trans (eisSummand_is_bounded_on_box (k := k) (max (v.1 0).natAbs (v.1 1).natAbs) x v + (by linarith) (by simp only [Int.mem_box])) + simp [Fin.isValue, Nat.cast_max, mul_inv_rev, one_div] have hk0 : 0 ≤ k := by linarith lift k to ℕ using hk0 - rw [inv_le_inv ] - · apply pow_le_pow_left (r_pos _).le (r_lower_bound_on_strip hB ⟨x, HABK hx⟩) - · apply pow_pos ( (r_pos x)) + gcongr · apply pow_pos (r_pos _) + · apply (r_pos _).le + · apply (r_lower_bound_on_verticalStrip hB ⟨x, HABK hx⟩) local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm (OpenEmbedding.toPartialHomeomorph UpperHalfPlane.coe openEmbedding_coe)) -/- A version for the extension to maps `ℂ → ℂ` that is nice to have for holomorphicity later -/ -lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} {N : ℕ} (hk : 3 ≤ k) +/- A version for the extension to maps `ℂ → ℂ` that is nice to have for holomorphicity later. -/ +lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} {N : ℕ} (hk : 3 ≤ k) (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => ↑ₕ(fun (z : ℍ) => ∑ x in s, eisSummand k x z )) (↑ₕ(eisensteinSeries_SIF a k).toFun) Filter.atTop (UpperHalfPlane.coe '' ⊤) := by - apply TendstoLocallyUniformlyOn.comp (s := ⊤) - simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ] - apply eisensteinSeries_tendstoLocallyUniformly hk - simp only [Set.top_eq_univ, image_univ, mapsTo_range_iff, Set.mem_univ, forall_const] - apply PartialHomeomorph.continuousOn_symm + apply TendstoLocallyUniformlyOn.comp (s := ⊤) _ _ _ (PartialHomeomorph.continuousOn_symm _) + · simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ] + apply eisensteinSeries_tendstoLocallyUniformly hk + · simp only [OpenEmbedding.toPartialHomeomorph_target, Set.top_eq_univ, mapsTo_range_iff, + Set.mem_univ, forall_const] From 760665569ddfaccc34ab14ccf18499e310844eaf Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 7 May 2024 11:53:32 +0100 Subject: [PATCH 107/124] lintr fix --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index f5fc74b7c6f59..fd230140d10a1 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -119,7 +119,7 @@ lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x · let n := max (x 0).natAbs (x 1).natAbs have hn0 : n ≠ 0 := by rw [← Iff.ne ((eq_zero_iff_eq_zero_of_mem_box (α := ℤ × ℤ) (x := ((x 0, x 1)))) (by simp)), - ← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun x ↦ ℤ)))] at * + ← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℤ)))] at * simpa only [Fin.isValue, ne_eq, Prod.mk_eq_zero, not_and, piFinTwoEquiv_apply, Pi.zero_apply, Prod.mk.injEq] using hx have h11 : ((x 0) * ↑z + (x 1)) = From 5852fe8d43e1a844899bf01e7e0763b27689a24c Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 10 May 2024 10:54:52 +0100 Subject: [PATCH 108/124] space --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index fd230140d10a1..b0bc31b6f7610 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -177,7 +177,7 @@ theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : exact Nat.pos_of_ne_zero hn /-This is a special case of the above, but one that we use more. -/ -theorem eisSummand_is_bounded_on_box {k: ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) +theorem eisSummand_is_bounded_on_box {k : ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) (hx : (x 0, x 1) ∈ box n) : Complex.abs (eisSummand k x z) ≤ (((r z) ^ k * n ^ k))⁻¹ := by have := summand_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx norm_cast at * From 7fb014d35a9176d37353bfe9bb967139b24b96a5 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 10 May 2024 17:23:27 +0100 Subject: [PATCH 109/124] rev changes wip --- .../EisensteinSeries/UniformConvergence.lean | 235 +++++++----------- 1 file changed, 96 insertions(+), 139 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index b0bc31b6f7610..8058c08614c47 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -27,145 +27,105 @@ namespace EisensteinSeries section bounding_functions -/-- Auxiliary function used for bounding Eisenstein series. -/ +/-- Auxiliary function used for bounding Eisenstein series, defined as + `z.im ^ 2 / (z.re ^ 2 + z.im ^ 2)`. -/ def r1 (z : ℍ) : ℝ := z.im ^ 2 / (z.re ^ 2 + z.im ^ 2) -lemma r1' (z : ℍ) : r1 z = 1 / ((z.re / z.im) ^ 2 + 1) := by - field_simp [r1, im_pos z] +lemma r1_eq (z : ℍ) : r1 z = 1 / ((z.re / z.im) ^ 2 + 1) := by + rw [div_pow, div_add_one (by positivity), one_div_div, r1] theorem r1_pos (z : ℍ) : 0 < r1 z := by dsimp only [r1] positivity /-- This function is used to give an upper bound on Eisenstein series. -/ -def r (z : ℍ) : ℝ := min z.im (Real.sqrt (r1 z)) +def r (z : ℍ) : ℝ := min z.im √(r1 z) lemma r_pos (z : ℍ) : 0 < r z := by simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, r1_pos, and_self] -lemma r1_aux_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε^2) : - (z.im ^ 2) / (z.re ^ 2 + z.im ^ 2) ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by +lemma r1_aux_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε ^ 2) : + r1 z ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by have H1 : (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 = - δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2 := by ring - have H2 : (δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε^2) * (z.re ^ 2 + z.im ^ 2) - - (z.im ^ 2) = (δ * (z.re ^ 2 + z.im ^ 2)+ ε * z.re)^2 + (ε^2 - 1)* (z.im)^2 := by ring - rw [H1, div_le_iff, ← sub_nonneg, H2] - · apply add_nonneg (pow_two_nonneg _) ?_ - apply mul_nonneg - · linarith - · apply pow_two_nonneg - · positivity - -lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ ^ 2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by - rw [r, Complex.abs] - have H1 : (z : ℂ).im ≤ - Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + (δ * z : ℂ).im * (δ * z : ℂ).im) := by - have h1 : (δ * z : ℂ).im * (δ * z : ℂ).im = δ ^ 2 * (z : ℂ).im * (z : ℂ).im := by - simp only [mul_im, ofReal_re, coe_im, ofReal_im, coe_re, zero_mul, add_zero] - ring - rw [Real.le_sqrt', h1] - · nlinarith - · exact z.2 - simp only [UpperHalfPlane.coe_im, UpperHalfPlane.coe_re, AbsoluteValue.coe_mk, MulHom.coe_mk, - min_le_iff] at * - left - simp only [one_le_sq_iff_one_le_abs, mul_im, ofReal_re, coe_im, ofReal_im, coe_re, zero_mul, - add_zero, normSq_apply, add_re, mul_re, sub_zero, add_im] at * - exact H1 - -lemma auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε ^ 2) : r z ≤ Complex.abs (δ * (z : ℂ) + ε) := by - rw [r, Complex.abs, min_le_iff] - have H1 : Real.sqrt (r1 z) ≤ Real.sqrt ((δ * (z : ℂ).re + ε) * (δ * (z : ℂ).re + ε) + - δ * (z : ℂ).im * (δ * (z : ℂ).im)) := by - rw [r1, Real.sqrt_le_sqrt_iff, ← pow_two, ← pow_two] - apply r1_aux_bound z δ hε - nlinarith - right - simp only [ne_eq, coe_re, coe_im, normSq_apply, AbsoluteValue.coe_mk, MulHom.coe_mk, add_re, - mul_re, ofReal_re, ofReal_im, zero_mul, sub_zero, add_im, mul_im, add_zero] at * - exact H1 - -lemma ne_zero_if_max {x : Fin 2 → ℤ} (hx : x ≠ 0) + δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε ^ 2 := by ring + have H2 : (δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε ^ 2) * (z.re ^ 2 + z.im ^ 2) + - z.im ^ 2 = (δ * (z.re ^ 2 + z.im ^ 2) + ε * z.re) ^ 2 + (ε ^ 2 - 1) * z.im ^ 2 := by ring + rw [r1, H1, div_le_iff (by positivity), ← sub_nonneg, H2] + exact add_nonneg (sq_nonneg _) (mul_nonneg (sub_nonneg.mpr hε) (sq_nonneg _)) + +lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ ^ 2) : r z ≤ Complex.abs (δ * z + ε) := by + rcases z with ⟨z, hz⟩ + have H1 : z.im ≤ √((δ * z.re + ε) ^ 2 + (δ * z).im ^ 2) := by + rw [Real.le_sqrt' hz, im_ofReal_mul, mul_pow] + exact (le_mul_of_one_le_left (sq_nonneg _) hδ).trans <| le_add_of_nonneg_left (sq_nonneg _) + simpa only [r, abs_apply, normSq_apply, add_re, re_ofReal_mul, coe_re, ← pow_two, add_im, mul_im, + coe_im, ofReal_im, zero_mul, add_zero, min_le_iff] using Or.inl H1 + +lemma auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε ^ 2) : r z ≤ Complex.abs (δ * z + ε) := by + have H1 : √(r1 z) ≤ √((δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2) := + (Real.sqrt_le_sqrt_iff (by positivity)).mpr (r1_aux_bound _ _ hε) + simpa only [r, abs_apply, normSq_apply, add_re, re_ofReal_mul, coe_re, ofReal_re, ← pow_two, + add_im, im_ofReal_mul, coe_im, ofReal_im, add_zero, min_le_iff] using Or.inr H1 + +lemma left_ne_zero_of_max_eq {x : Fin 2 → ℤ} (hx : x ≠ 0) (h : max (x 0).natAbs (x 1).natAbs = (x 0).natAbs) : x 0 ≠ 0 := by contrapose! hx - ext x - fin_cases x <;> aesop - -/-This proof is faster than the one above, but longer. -/ -lemma ne_zero_if_max' {x : Fin 2 → ℤ} (hx : x ≠ 0) - (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : (x 1) ≠ 0 := by - apply ne_zero_if_max (x := ![x 1, x 0]) ?_ (by simpa only [Fin.isValue, Matrix.cons_val_zero, - Matrix.cons_val_one, Matrix.head_cons, max_eq_left_iff, max_eq_right_iff] using h) - simp only [ne_eq, Matrix.cons_eq_zero_iff, Matrix.zero_empty, and_true, not_and] - intro h1 h0 - rw [Function.ne_iff, Fin.exists_fin_two, h1, h0] at hx - simp only [Pi.zero_apply, ne_eq, not_true_eq_false, or_self] at * + rw [hx, Int.natAbs_zero, max_eq_left_iff, Nat.le_zero, Int.natAbs_eq_zero] at h + ext1 j + fin_cases j <;> assumption + +lemma right_ne_zero_of_max_eq {x : Fin 2 → ℤ} (hx : x ≠ 0) + (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : x 1 ≠ 0 := by + contrapose! hx + rw [hx, Int.natAbs_zero, max_eq_right_iff, Nat.le_zero, Int.natAbs_eq_zero] at h + ext1 j + fin_cases j <;> assumption lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 ∨ (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 := by refine (max_choice (x 0).natAbs (x 1).natAbs).imp (fun H1 => ?_) (fun H2 => ?_) · simp only [Fin.isValue, H1, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, - OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, ne_zero_if_max hx H1, - div_self, le_refl] + OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, + left_ne_zero_of_max_eq hx H1, div_self, le_refl] · simp only [Fin.isValue, H2, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, - ne_zero_if_max' hx H2, div_self, le_refl] + right_ne_zero_of_max_eq hx H2, div_self, le_refl] lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) : - ((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k ≤ - (Complex.abs ((x 0 : ℂ) * (z : ℂ) + (x 1 : ℂ))) ^ k := by - by_cases hk0 : k ≠ 0 - · let n := max (x 0).natAbs (x 1).natAbs - have hn0 : n ≠ 0 := by - rw [← Iff.ne ((eq_zero_iff_eq_zero_of_mem_box (α := ℤ × ℤ) (x := ((x 0, x 1)))) (by simp)), - ← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℤ)))] at * - simpa only [Fin.isValue, ne_eq, Prod.mk_eq_zero, not_and, piFinTwoEquiv_apply, Pi.zero_apply, - Prod.mk.injEq] using hx - have h11 : ((x 0) * ↑z + (x 1)) = - (((x 0 : ℝ) / (n : ℝ)) * (z : ℂ) + (x 1 : ℝ) / (n : ℝ)) * ((n : ℝ)) := by - have : (n : ℂ) ≠ 0 := by norm_cast - field_simp [this] - simp only [Nat.cast_max, h11, map_mul, abs_ofReal, ge_iff_le] - rw [Real.mul_rpow (by apply apply_nonneg) (by apply abs_nonneg)] - cases' (div_max_sq_ge_one x hx) with H1 H2 - · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) - · simpa only [Fin.isValue, ofReal_intCast, Nat.cast_max, ofReal_div, n] using - (Real.rpow_le_rpow (r_pos _).le (auxbound1 z (x 1 / n) H1) hk) - · positivity - · apply mul_le_mul _ (by norm_cast) _ (by apply Real.rpow_nonneg (Complex.abs.nonneg _) k) - · simpa only [Fin.isValue, ofReal_intCast, Nat.cast_max, ofReal_div, n] using - (Real.rpow_le_rpow (r_pos _).le (auxbound2 z (x 0 / n) H2) hk) - · positivity - · simp only [ne_eq, not_not] at hk0 - simp only [hk0, Real.rpow_zero, Nat.cast_max, mul_one, le_refl] + (r z) ^ k * (max (x 0).natAbs (x 1).natAbs) ^ k ≤ Complex.abs (x 0 * z + x 1) ^ k := by + rw [← Real.mul_rpow (r_pos _).le (Nat.cast_nonneg _)] + refine Real.rpow_le_rpow (mul_nonneg (r_pos _).le (Nat.cast_nonneg _)) ?_ hk + let n := max (x 0).natAbs (x 1).natAbs + have hn0 : n ≠ 0 := by + contrapose! hx + rw [Nat.max_eq_zero_iff, Int.natAbs_eq_zero, Int.natAbs_eq_zero] at hx + exact funext fun j ↦ by fin_cases j <;> tauto + have h11 : x 0 * (z : ℂ) + x 1 = (x 0 / n * z + x 1 / n) * n := by + push_cast + rw [div_mul_eq_mul_div, ← add_div, div_mul_cancel₀ _ (Nat.cast_ne_zero.mpr hn0)] + rw [Nat.cast_max, h11, map_mul, abs_natCast] + gcongr + · rcases div_max_sq_ge_one x hx with H1 | H2 + · simpa only [Fin.isValue, ofReal_div, ofReal_intCast, n] using auxbound1 z (x 1 / n) H1 + · simpa only [Fin.isValue, ofReal_div, ofReal_intCast, n] using auxbound2 z (x 0 / n) H2 + · simp only [Fin.isValue, Nat.cast_max, le_refl, n] theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) - (hx : (x 0, x 1) ∈ box n) : (Complex.abs (((x 0 : ℂ) * z + (x 1 : ℂ)))) ^ (-k) ≤ - (r z) ^ (-k) * n ^ (-k) := by + (hx : (x 0, x 1) ∈ box n) : Complex.abs (x 0 * z + x 1) ^ (-k) ≤ (r z) ^ (-k) * n ^ (-k) := by by_cases hn : n = 0 · simp only [hn, box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx - rw [hx.1, hx.2] at * - by_cases hk0 : k = 0 - · simp only [hk0, neg_zero, Real.rpow_zero, mul_one, le_refl] - · simp only [Int.cast_zero, zero_mul, add_zero, map_zero] - have h1 : (0 : ℝ) ^ (-k) = 0 := by - rw [Real.rpow_eq_zero_iff_of_nonneg (by rfl)] - simp only [ne_eq, neg_eq_zero, hk0, not_false_eq_true, and_self] - simp only [h1, hn, CharP.cast_eq_zero, mul_zero, le_refl] + rw [hx.1, hx.2, hn, ← Real.mul_rpow (r_pos z).le (Nat.cast_nonneg 0)] + simp only [Int.cast_zero, zero_mul, add_zero, map_zero, CharP.cast_eq_zero, mul_zero, le_refl] · have hx2 : x ≠ 0 := by - rw [← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℤ)))] - simpa only [piFinTwoEquiv_apply, Fin.isValue, Pi.zero_apply, ne_eq, Prod.mk.injEq, not_and, - Prod.mk_eq_zero] using - (Iff.ne ((eq_zero_iff_eq_zero_of_mem_box (x := (x 0, x 1))) hx)).mpr hn - simp only [Int.mem_box] at hx + contrapose! hn + simp only [hn, Fin.isValue, Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self] at hx + exact hx.symm + rw [Int.mem_box] at hx rw [Real.rpow_neg (by apply apply_nonneg), Real.rpow_neg ((r_pos z).le), - Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv] - simp only [← hx, Nat.cast_max] - convert (rpow_bound hk z x hx2) - · simp only [Nat.cast_max] - · apply Real.rpow_pos_of_pos - apply Complex.abs.pos (linear_ne_zero ![x 0, x 1] z ?_) + Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv, ← hx, Nat.cast_max] + simpa only [Fin.isValue, Nat.cast_max] using (rpow_bound hk z x hx2) + · apply Real.rpow_pos_of_pos (Complex.abs.pos (linear_ne_zero ![x 0, x 1] z ?_)) have := (Function.comp_ne_zero_iff x Int.cast_injective Int.cast_zero (γ := ℝ)).mpr hx2 rw [← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℝ)))] at this simpa only [Fin.isValue, ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, @@ -173,72 +133,69 @@ theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : Prod.mk.injEq] using this · apply mul_pos (Real.rpow_pos_of_pos (r_pos z) _) apply Real.rpow_pos_of_pos - norm_cast - exact Nat.pos_of_ne_zero hn + exact Nat.cast_pos.mpr (Nat.pos_of_ne_zero hn) /-This is a special case of the above, but one that we use more. -/ theorem eisSummand_is_bounded_on_box {k : ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) (hx : (x 0, x 1) ∈ box n) : Complex.abs (eisSummand k x z) ≤ (((r z) ^ k * n ^ k))⁻¹ := by have := summand_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx - norm_cast at * - rw [zpow_neg, mul_inv, eisSummand] at * - simpa only [Fin.isValue, one_div, map_inv₀, map_zpow₀, ge_iff_le, zpow_neg] using this + simp_rw [← Int.cast_neg k, Real.rpow_intCast, zpow_neg] at this + rw [mul_inv, eisSummand, one_div, map_inv₀, map_zpow₀] + exact this lemma r_lower_bound_on_verticalStrip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by - have hz := z.2 - simp only [mem_verticalStrip_iff, abs_ofReal, ge_iff_le] at hz - rw [r] - apply min_le_min - · convert hz.2 + rcases z with ⟨z, hz⟩ + simp only [mem_verticalStrip_iff, abs_ofReal] at hz + apply min_le_min hz.2 · rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] - simp only [r1', div_pow, one_div] + simp only [r1_eq, div_pow, one_div] rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] apply div_le_div (sq_nonneg _) _ (by positivity) - · simpa only using pow_le_pow_left h.le hz.2 2 + · exact pow_le_pow_left h.le hz.2 2 · simpa only [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 end bounding_functions section summability -lemma summable_r_pow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : +lemma summable_r_zpow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by - have hk : 1 < (k - 1 : ℝ) := by norm_cast; linarith + have hk : 1 < (k - 1 : ℝ) := by norm_cast; exact Int.lt_sub_left_of_add_lt h have nze : (8 / (r z) ^ k : ℝ) ≠ 0 := by - apply div_ne_zero - simp only [ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true] - apply zpow_ne_zero k (ne_of_gt (r_pos z)) + exact div_ne_zero (OfNat.ofNat_ne_zero 8) (zpow_ne_zero k (ne_of_gt (r_pos z))) rw [← (summable_mul_left_iff nze).symm] - convert (Real.summable_nat_rpow_inv.2 hk) + apply (Real.summable_nat_rpow_inv.2 hk).congr + intro b norm_cast lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): Summable (fun n : ℕ => ∑ v in (box n : Finset (ℤ × ℤ)), (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by simp only [one_div, sum_const, nsmul_eq_mul] - apply Summable.congr (summable_r_pow z h) + apply (summable_r_zpow z h).congr intro b by_cases b0 : b = 0 · simp only [b0, CharP.cast_eq_zero, box_zero, Finset.card_singleton, Nat.cast_one, one_mul] rw [zero_zpow k (by linarith), zero_zpow (k - 1) (by linarith)] simp only [inv_zero, mul_zero] - · rw [Int.card_box b0, zpow_sub_one₀ (a:= (b: ℝ)) (Nat.cast_ne_zero.mpr b0) k] + · rw [Int.card_box b0, zpow_sub_one₀ (a:= (b : ℝ)) (Nat.cast_ne_zero.mpr b0) k] simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] ring_nf lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => - (((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ := by - set f := fun x : Fin 2 → ℤ ↦ (((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs : ℝ) ^ k)⁻¹ + (((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k)⁻¹ := by + set f := fun x : Fin 2 → ℤ ↦ (((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k)⁻¹ rw [← (piFinTwoEquiv _).symm.summable_iff, summable_partition _ (s := fun n ↦ (box n : Finset (ℤ × ℤ))) Int.existsUnique_mem_box] · simp_rw [coe_sort_coe, Finset.tsum_subtype] simp only [one_div, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one] - refine ⟨fun n ↦ ?_, Summable.congr (summable_over_box z h) fun n ↦ Finset.sum_congr rfl + refine ⟨fun n ↦ ?_, (summable_over_box z h).congr fun n ↦ Finset.sum_congr rfl fun x hx ↦ ?_⟩ - · simpa using (box n).summable (f ∘ (piFinTwoEquiv _).symm) - · simp only [Int.mem_box, one_div] at * - simp only [← hx, Nat.cast_max, Fin.isValue, mul_inv_rev, Fin.cons_zero, Fin.cons_one, f] - rw [mul_comm] + · simpa only [coe_sort_coe, piFinTwoEquiv_symm_apply] using + (box n).summable (f ∘ (piFinTwoEquiv _).symm) + · rw [Int.mem_box] at hx + simp only [one_div, ← hx, Nat.cast_max, Fin.isValue, mul_inv_rev, mul_comm, Fin.cons_zero, + Fin.cons_one, f] · intro y simp only [Pi.zero_apply, Fin.isValue, mul_inv_rev, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one, f] @@ -260,16 +217,16 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : intro K hK obtain ⟨A, B, hB, HABK⟩ := subset_verticalStrip_of_isCompact hK have hu : Summable fun x : (gammaSet N a) => - (((r ⟨⟨A, B⟩, hB⟩) ^ k) * (max (x.1 0).natAbs (x.1 1).natAbs : ℝ) ^ k)⁻¹ := by - apply (Summable.subtype (summable_upper_bound hk ⟨⟨A, B⟩, hB⟩) (gammaSet N a)).congr + (((r ⟨⟨A, B⟩, hB⟩) ^ k) * (max (x.1 0).natAbs (x.1 1).natAbs) ^ k)⁻¹ := by + apply ((summable_upper_bound hk ⟨⟨A, B⟩, hB⟩).subtype (gammaSet N a)).congr intro v simp only [zpow_natCast, one_div, Function.comp_apply] apply tendstoUniformlyOn_tsum hu intro v x hx apply le_trans (eisSummand_is_bounded_on_box (k := k) (max (v.1 0).natAbs (v.1 1).natAbs) x v (by linarith) (by simp only [Int.mem_box])) - simp [Fin.isValue, Nat.cast_max, mul_inv_rev, one_div] - have hk0 : 0 ≤ k := by linarith + simp only [Fin.isValue, Nat.cast_max, mul_inv_rev] + have hk0 : 0 ≤ k := by exact le_trans (Int.nonneg_of_normalize_eq_self rfl) hk lift k to ℕ using hk0 gcongr · apply pow_pos (r_pos _) From 6823211f2492b94b0fd4dbeddef0ffca3e3ed8c6 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Sun, 12 May 2024 22:32:51 +0100 Subject: [PATCH 110/124] small golf --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 8058c08614c47..886e97aba63d3 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -132,8 +132,7 @@ theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : and_true, not_and, piFinTwoEquiv_apply, Function.comp_apply, Pi.zero_apply, Prod.mk.injEq] using this · apply mul_pos (Real.rpow_pos_of_pos (r_pos z) _) - apply Real.rpow_pos_of_pos - exact Nat.cast_pos.mpr (Nat.pos_of_ne_zero hn) + apply Real.rpow_pos_of_pos (Nat.cast_pos.mpr (Nat.pos_of_ne_zero hn)) /-This is a special case of the above, but one that we use more. -/ theorem eisSummand_is_bounded_on_box {k : ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) @@ -188,7 +187,7 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : rw [← (piFinTwoEquiv _).symm.summable_iff, summable_partition _ (s := fun n ↦ (box n : Finset (ℤ × ℤ))) Int.existsUnique_mem_box] · simp_rw [coe_sort_coe, Finset.tsum_subtype] - simp only [one_div, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one] + simp only [piFinTwoEquiv_symm_apply, Function.comp_apply] refine ⟨fun n ↦ ?_, (summable_over_box z h).congr fun n ↦ Finset.sum_congr rfl fun x hx ↦ ?_⟩ · simpa only [coe_sort_coe, piFinTwoEquiv_symm_apply] using From 26d952427727cf6b93bba0c4c8c3ff9fd45586c1 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 13 May 2024 19:57:31 +0100 Subject: [PATCH 111/124] more speedups --- .../EisensteinSeries/UniformConvergence.lean | 57 ++++++++----------- 1 file changed, 24 insertions(+), 33 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 886e97aba63d3..ac055187a66ec 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -117,11 +117,11 @@ theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : · simp only [hn, box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx rw [hx.1, hx.2, hn, ← Real.mul_rpow (r_pos z).le (Nat.cast_nonneg 0)] simp only [Int.cast_zero, zero_mul, add_zero, map_zero, CharP.cast_eq_zero, mul_zero, le_refl] - · have hx2 : x ≠ 0 := by + · simp only [ Fin.isValue, Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self] at hx + have hx2 : x ≠ 0 := by contrapose! hn - simp only [hn, Fin.isValue, Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self] at hx + simp only [hn, Fin.isValue, Pi.zero_apply, Int.natAbs_zero, max_self] at hx exact hx.symm - rw [Int.mem_box] at hx rw [Real.rpow_neg (by apply apply_nonneg), Real.rpow_neg ((r_pos z).le), Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv, ← hx, Nat.cast_max] simpa only [Fin.isValue, Nat.cast_max] using (rpow_bound hk z x hx2) @@ -134,7 +134,7 @@ theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : · apply mul_pos (Real.rpow_pos_of_pos (r_pos z) _) apply Real.rpow_pos_of_pos (Nat.cast_pos.mpr (Nat.pos_of_ne_zero hn)) -/-This is a special case of the above, but one that we use more. -/ +/- This is a special case of the above, but one that we use more. -/ theorem eisSummand_is_bounded_on_box {k : ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) (hx : (x 0, x 1) ∈ box n) : Complex.abs (eisSummand k x z) ≤ (((r z) ^ k * n ^ k))⁻¹ := by have := summand_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx @@ -158,46 +158,39 @@ end bounding_functions section summability -lemma summable_r_zpow {k : ℤ} (z : ℍ) (h : 3 ≤ k) : - Summable fun n : ℕ => 8 / (r z) ^ k * ((n : ℝ) ^ (k - 1))⁻¹ := by +/-An auxilary summable lemma needed later. Added the set_option as replacing `v` with + `_` didn't work. -/ +set_option linter.unusedVariables false +lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): + Summable (fun n : ℕ => ∑ v in (box n : Finset (ℤ × ℤ)), ((r z) ^ k * (n : ℝ) ^ k)⁻¹) := by + simp only [sum_const, nsmul_eq_mul] have hk : 1 < (k - 1 : ℝ) := by norm_cast; exact Int.lt_sub_left_of_add_lt h - have nze : (8 / (r z) ^ k : ℝ) ≠ 0 := by + have nze : (8 * ((r z) ^ k)⁻¹ : ℝ) ≠ 0 := by exact div_ne_zero (OfNat.ofNat_ne_zero 8) (zpow_ne_zero k (ne_of_gt (r_pos z))) - rw [← (summable_mul_left_iff nze).symm] - apply (Real.summable_nat_rpow_inv.2 hk).congr + apply ((summable_mul_left_iff nze).mpr (Real.summable_nat_rpow_inv.2 hk)).congr intro b norm_cast - -lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): - Summable (fun n : ℕ => ∑ v in (box n : Finset (ℤ × ℤ)), (1 / (r z) ^ k) * ((n : ℝ) ^ k)⁻¹) := by - simp only [one_div, sum_const, nsmul_eq_mul] - apply (summable_r_zpow z h).congr - intro b by_cases b0 : b = 0 - · simp only [b0, CharP.cast_eq_zero, box_zero, Finset.card_singleton, Nat.cast_one, one_mul] - rw [zero_zpow k (by linarith), zero_zpow (k - 1) (by linarith)] - simp only [inv_zero, mul_zero] + · simp only [b0, CharP.cast_eq_zero, zero_zpow (k - 1) (by omega), inv_zero, mul_zero, + box_zero, Finset.card_singleton, Nat.cast_one, zero_zpow k (by omega)] · rw [Int.card_box b0, zpow_sub_one₀ (a:= (b : ℝ)) (Nat.cast_ne_zero.mpr b0) k] simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] ring_nf lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => - (((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k)⁻¹ := by - set f := fun x : Fin 2 → ℤ ↦ (((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k)⁻¹ + ((r z) ^ k * (max (x 0).natAbs (x 1).natAbs) ^ k)⁻¹ := by rw [← (piFinTwoEquiv _).symm.summable_iff, summable_partition _ (s := fun n ↦ (box n : Finset (ℤ × ℤ))) Int.existsUnique_mem_box] - · simp_rw [coe_sort_coe, Finset.tsum_subtype] - simp only [piFinTwoEquiv_symm_apply, Function.comp_apply] + · simp_rw [coe_sort_coe, Finset.tsum_subtype, piFinTwoEquiv_symm_apply, Function.comp_apply] refine ⟨fun n ↦ ?_, (summable_over_box z h).congr fun n ↦ Finset.sum_congr rfl fun x hx ↦ ?_⟩ - · simpa only [coe_sort_coe, piFinTwoEquiv_symm_apply] using - (box n).summable (f ∘ (piFinTwoEquiv _).symm) - · rw [Int.mem_box] at hx - simp only [one_div, ← hx, Nat.cast_max, Fin.isValue, mul_inv_rev, mul_comm, Fin.cons_zero, - Fin.cons_one, f] + · apply (box n).summable ((fun x : Fin 2 → ℤ ↦ (((r z) ^ k) * + (max (x 0).natAbs (x 1).natAbs) ^ k)⁻¹) ∘ (piFinTwoEquiv _).symm) + · simp only [one_div, Int.mem_box.mp hx, Nat.cast_max, Fin.isValue, mul_inv_rev, mul_comm, Fin.cons_zero, + Fin.cons_one] · intro y simp only [Pi.zero_apply, Fin.isValue, mul_inv_rev, piFinTwoEquiv_symm_apply, - Function.comp_apply, Fin.cons_zero, Fin.cons_one, f] + Function.comp_apply, Fin.cons_zero, Fin.cons_one] apply mul_nonneg · simp only [piFinTwoEquiv_symm_apply, Fin.cons_zero, Fin.cons_one, inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] @@ -216,14 +209,12 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : intro K hK obtain ⟨A, B, hB, HABK⟩ := subset_verticalStrip_of_isCompact hK have hu : Summable fun x : (gammaSet N a) => - (((r ⟨⟨A, B⟩, hB⟩) ^ k) * (max (x.1 0).natAbs (x.1 1).natAbs) ^ k)⁻¹ := by - apply ((summable_upper_bound hk ⟨⟨A, B⟩, hB⟩).subtype (gammaSet N a)).congr - intro v - simp only [zpow_natCast, one_div, Function.comp_apply] + ((r ⟨⟨A, B⟩, hB⟩) ^ k * (max (x.1 0).natAbs (x.1 1).natAbs) ^ k)⁻¹ := by + apply ((summable_upper_bound hk ⟨⟨A, B⟩, hB⟩).subtype (gammaSet N a)).congr (fun _ => rfl) apply tendstoUniformlyOn_tsum hu intro v x hx apply le_trans (eisSummand_is_bounded_on_box (k := k) (max (v.1 0).natAbs (v.1 1).natAbs) x v - (by linarith) (by simp only [Int.mem_box])) + (by omega) (by simp only [Int.mem_box])) simp only [Fin.isValue, Nat.cast_max, mul_inv_rev] have hk0 : 0 ≤ k := by exact le_trans (Int.nonneg_of_normalize_eq_self rfl) hk lift k to ℕ using hk0 From 5b3cfb87d8e89e22c2aecbd0c7501d107ac5f675 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 13 May 2024 19:58:17 +0100 Subject: [PATCH 112/124] more speedups 2 --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index ac055187a66ec..819c48d340438 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -216,7 +216,7 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : apply le_trans (eisSummand_is_bounded_on_box (k := k) (max (v.1 0).natAbs (v.1 1).natAbs) x v (by omega) (by simp only [Int.mem_box])) simp only [Fin.isValue, Nat.cast_max, mul_inv_rev] - have hk0 : 0 ≤ k := by exact le_trans (Int.nonneg_of_normalize_eq_self rfl) hk + have hk0 : 0 ≤ k := by omega lift k to ℕ using hk0 gcongr · apply pow_pos (r_pos _) From 6c6ddf3f41706c073b901af2d5ba9551a68c51b7 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 13 May 2024 20:26:54 +0100 Subject: [PATCH 113/124] lintr fix --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 819c48d340438..1d86553fe3f40 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -186,8 +186,8 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : fun x hx ↦ ?_⟩ · apply (box n).summable ((fun x : Fin 2 → ℤ ↦ (((r z) ^ k) * (max (x 0).natAbs (x 1).natAbs) ^ k)⁻¹) ∘ (piFinTwoEquiv _).symm) - · simp only [one_div, Int.mem_box.mp hx, Nat.cast_max, Fin.isValue, mul_inv_rev, mul_comm, Fin.cons_zero, - Fin.cons_one] + · simp only [mul_inv_rev, mul_comm, Fin.isValue, Fin.cons_zero, Fin.cons_one, + Int.mem_box.mp hx] · intro y simp only [Pi.zero_apply, Fin.isValue, mul_inv_rev, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one] From 5201ffbf43d5c79239ea4995a8a98ae385ce7b9a Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 14 May 2024 09:43:00 +0100 Subject: [PATCH 114/124] replace rpow_bound --- .../EisensteinSeries/UniformConvergence.lean | 35 +++++++------------ 1 file changed, 12 insertions(+), 23 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 1d86553fe3f40..45e17539544d8 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -92,10 +92,8 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, right_ne_zero_of_max_eq hx H2, div_self, le_refl] -lemma rpow_bound {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (x : Fin 2 → ℤ) (hx : x ≠ 0) : - (r z) ^ k * (max (x 0).natAbs (x 1).natAbs) ^ k ≤ Complex.abs (x 0 * z + x 1) ^ k := by - rw [← Real.mul_rpow (r_pos _).le (Nat.cast_nonneg _)] - refine Real.rpow_le_rpow (mul_nonneg (r_pos _).le (Nat.cast_nonneg _)) ?_ hk +lemma r_mul_max_le (z : ℍ) {x : Fin 2 → ℤ} (hx : x ≠ 0) : + r z * (max (x 0).natAbs (x 1).natAbs) ≤ Complex.abs (x 0 * z + x 1) := by let n := max (x 0).natAbs (x 1).natAbs have hn0 : n ≠ 0 := by contrapose! hx @@ -117,26 +115,17 @@ theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : · simp only [hn, box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx rw [hx.1, hx.2, hn, ← Real.mul_rpow (r_pos z).le (Nat.cast_nonneg 0)] simp only [Int.cast_zero, zero_mul, add_zero, map_zero, CharP.cast_eq_zero, mul_zero, le_refl] - · simp only [ Fin.isValue, Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self] at hx + · simp only [Fin.isValue, Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self] at hx have hx2 : x ≠ 0 := by contrapose! hn - simp only [hn, Fin.isValue, Pi.zero_apply, Int.natAbs_zero, max_self] at hx - exact hx.symm - rw [Real.rpow_neg (by apply apply_nonneg), Real.rpow_neg ((r_pos z).le), - Real.rpow_neg (Nat.cast_nonneg n), ← mul_inv, inv_le_inv, ← hx, Nat.cast_max] - simpa only [Fin.isValue, Nat.cast_max] using (rpow_bound hk z x hx2) - · apply Real.rpow_pos_of_pos (Complex.abs.pos (linear_ne_zero ![x 0, x 1] z ?_)) - have := (Function.comp_ne_zero_iff x Int.cast_injective Int.cast_zero (γ := ℝ)).mpr hx2 - rw [← Iff.ne (Function.Injective.eq_iff (Equiv.injective (piFinTwoEquiv fun _ ↦ ℝ)))] at this - simpa only [Fin.isValue, ne_eq, Matrix.cons_eq_zero_iff, Int.cast_eq_zero, Matrix.zero_empty, - and_true, not_and, piFinTwoEquiv_apply, Function.comp_apply, Pi.zero_apply, - Prod.mk.injEq] using this - · apply mul_pos (Real.rpow_pos_of_pos (r_pos z) _) - apply Real.rpow_pos_of_pos (Nat.cast_pos.mpr (Nat.pos_of_ne_zero hn)) + simpa only [hn, Fin.isValue, Pi.zero_apply, Int.natAbs_zero, max_self] using hx.symm + rw [← Real.mul_rpow (r_pos _).le (Nat.cast_nonneg _)] + exact Real.rpow_le_rpow_of_nonpos (mul_pos (r_pos _) (by positivity)) (hx ▸ r_mul_max_le z hx2) + (neg_nonpos.mpr hk) /- This is a special case of the above, but one that we use more. -/ -theorem eisSummand_is_bounded_on_box {k : ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) - (hx : (x 0, x 1) ∈ box n) : Complex.abs (eisSummand k x z) ≤ (((r z) ^ k * n ^ k))⁻¹ := by +theorem eisSummand_is_bounded_on_box_zpow {k : ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) + (hx : (x 0, x 1) ∈ box n) : Complex.abs (eisSummand k x z) ≤ ((r z) ^ k * n ^ k)⁻¹ := by have := summand_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx simp_rw [← Int.cast_neg k, Real.rpow_intCast, zpow_neg] at this rw [mul_inv, eisSummand, one_div, map_inv₀, map_zpow₀] @@ -213,16 +202,16 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : apply ((summable_upper_bound hk ⟨⟨A, B⟩, hB⟩).subtype (gammaSet N a)).congr (fun _ => rfl) apply tendstoUniformlyOn_tsum hu intro v x hx - apply le_trans (eisSummand_is_bounded_on_box (k := k) (max (v.1 0).natAbs (v.1 1).natAbs) x v + apply le_trans (eisSummand_is_bounded_on_box_zpow (k := k) (max (v.1 0).natAbs (v.1 1).natAbs) x v (by omega) (by simp only [Int.mem_box])) simp only [Fin.isValue, Nat.cast_max, mul_inv_rev] - have hk0 : 0 ≤ k := by omega - lift k to ℕ using hk0 + lift k to ℕ using (by omega) gcongr · apply pow_pos (r_pos _) · apply (r_pos _).le · apply (r_lower_bound_on_verticalStrip hB ⟨x, HABK hx⟩) +/-This is essentially extend by zero outside the upper half plane.-/ local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm (OpenEmbedding.toPartialHomeomorph UpperHalfPlane.coe openEmbedding_coe)) From fec390807355cbc2e316331ebaaae4a9b950c09d Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 14 May 2024 09:51:21 +0100 Subject: [PATCH 115/124] some style fixes, wip --- .../EisensteinSeries/UniformConvergence.lean | 13 ++++++------- 1 file changed, 6 insertions(+), 7 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 45e17539544d8..60fa982b39d3e 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -136,12 +136,11 @@ lemma r_lower_bound_on_verticalStrip {A B : ℝ} (h : 0 < B) (z : verticalStrip rcases z with ⟨z, hz⟩ simp only [mem_verticalStrip_iff, abs_ofReal] at hz apply min_le_min hz.2 - · rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] - simp only [r1_eq, div_pow, one_div] - rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] - apply div_le_div (sq_nonneg _) _ (by positivity) - · exact pow_le_pow_left h.le hz.2 2 - · simpa only [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 + rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] + simp only [r1_eq, div_pow, one_div] + rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] + apply div_le_div (sq_nonneg _) _ (by positivity) (pow_le_pow_left h.le hz.2 2) + simpa only [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 end bounding_functions @@ -154,7 +153,7 @@ lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): Summable (fun n : ℕ => ∑ v in (box n : Finset (ℤ × ℤ)), ((r z) ^ k * (n : ℝ) ^ k)⁻¹) := by simp only [sum_const, nsmul_eq_mul] have hk : 1 < (k - 1 : ℝ) := by norm_cast; exact Int.lt_sub_left_of_add_lt h - have nze : (8 * ((r z) ^ k)⁻¹ : ℝ) ≠ 0 := by + have nze : 8 * ((r z) ^ k)⁻¹ ≠ 0 := by exact div_ne_zero (OfNat.ofNat_ne_zero 8) (zpow_ne_zero k (ne_of_gt (r_pos z))) apply ((summable_mul_left_iff nze).mpr (Real.summable_nat_rpow_inv.2 hk)).congr intro b From bbab1d5784b477f685b18406a2b19de41bbc5609 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Tue, 14 May 2024 19:06:36 +0100 Subject: [PATCH 116/124] more golf --- .../EisensteinSeries/UniformConvergence.lean | 34 ++++++++----------- 1 file changed, 15 insertions(+), 19 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 60fa982b39d3e..dd27de0c522df 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -38,7 +38,8 @@ theorem r1_pos (z : ℍ) : 0 < r1 z := by dsimp only [r1] positivity -/-- This function is used to give an upper bound on Eisenstein series. -/ +/-- This function is used to give an upper bound on Eisenstein series, defined as +the minimun of `z.im` and `√(z.im ^ 2 / (z.re ^ 2 + z.im ^ 2))`. -/ def r (z : ℍ) : ℝ := min z.im √(r1 z) lemma r_pos (z : ℍ) : 0 < r z := by @@ -123,10 +124,10 @@ theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : exact Real.rpow_le_rpow_of_nonpos (mul_pos (r_pos _) (by positivity)) (hx ▸ r_mul_max_le z hx2) (neg_nonpos.mpr hk) -/- This is a special case of the above, but one that we use more. -/ +/- This is a special case of the above, but one that we use more for Eisenstein series. -/ theorem eisSummand_is_bounded_on_box_zpow {k : ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) (hx : (x 0, x 1) ∈ box n) : Complex.abs (eisSummand k x z) ≤ ((r z) ^ k * n ^ k)⁻¹ := by - have := summand_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx + have := summand_is_bounded_on_box_rpow (Int.cast_nonneg.mpr hk) z n x hx simp_rw [← Int.cast_neg k, Real.rpow_intCast, zpow_neg] at this rw [mul_inv, eisSummand, one_div, map_inv₀, map_zpow₀] exact this @@ -134,12 +135,10 @@ theorem eisSummand_is_bounded_on_box_zpow {k : ℤ} (n : ℕ) (z : ℍ) (x : Fin lemma r_lower_bound_on_verticalStrip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by rcases z with ⟨z, hz⟩ - simp only [mem_verticalStrip_iff, abs_ofReal] at hz - apply min_le_min hz.2 - rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] - simp only [r1_eq, div_pow, one_div] + apply min_le_min ((mem_verticalStrip_iff A B z).mp hz).2 + simp_rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le), r1_eq, div_pow, one_div] rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] - apply div_le_div (sq_nonneg _) _ (by positivity) (pow_le_pow_left h.le hz.2 2) + apply div_le_div (sq_nonneg _) ?_ (by positivity) (pow_le_pow_left h.le hz.2 2) simpa only [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 end bounding_functions @@ -155,14 +154,14 @@ lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): have hk : 1 < (k - 1 : ℝ) := by norm_cast; exact Int.lt_sub_left_of_add_lt h have nze : 8 * ((r z) ^ k)⁻¹ ≠ 0 := by exact div_ne_zero (OfNat.ofNat_ne_zero 8) (zpow_ne_zero k (ne_of_gt (r_pos z))) - apply ((summable_mul_left_iff nze).mpr (Real.summable_nat_rpow_inv.2 hk)).congr + apply ((summable_mul_left_iff nze).mpr (Real.summable_nat_rpow_inv.mpr hk)).congr intro b norm_cast by_cases b0 : b = 0 · simp only [b0, CharP.cast_eq_zero, zero_zpow (k - 1) (by omega), inv_zero, mul_zero, box_zero, Finset.card_singleton, Nat.cast_one, zero_zpow k (by omega)] - · rw [Int.card_box b0, zpow_sub_one₀ (a:= (b : ℝ)) (Nat.cast_ne_zero.mpr b0) k] - simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] + · rw [Int.card_box b0, zpow_sub_one₀ (a:= (b : ℝ)) (Nat.cast_ne_zero.mpr b0) k, + mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] ring_nf lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => @@ -177,18 +176,15 @@ lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : · simp only [mul_inv_rev, mul_comm, Fin.isValue, Fin.cons_zero, Fin.cons_one, Int.mem_box.mp hx] · intro y - simp only [Pi.zero_apply, Fin.isValue, mul_inv_rev, piFinTwoEquiv_symm_apply, + simp only [Pi.zero_apply, Fin.isValue, Nat.cast_max, mul_inv_rev, piFinTwoEquiv_symm_apply, Function.comp_apply, Fin.cons_zero, Fin.cons_one] - apply mul_nonneg - · simp only [piFinTwoEquiv_symm_apply, Fin.cons_zero, Fin.cons_one, inv_nonneg, ge_iff_le, - le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] - · simp only [one_div, inv_nonneg] - apply zpow_nonneg (r_pos z).le + apply mul_nonneg ?_ (inv_nonneg.mpr ((zpow_nonneg (r_pos z).le) k)) + simp only [inv_nonneg, ge_iff_le, le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] end summability -theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : ℕ) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformly (fun (s : Finset (gammaSet N a)) => +theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} {N : ℕ} (hk : 3 ≤ k) (a : Fin 2 → ZMod N) : + TendstoLocallyUniformly (fun (s : Finset (gammaSet N a)) => (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop := by rw [← tendstoLocallyUniformlyOn_univ,tendstoLocallyUniformlyOn_iff_forall_isCompact From f748460aaccf4acf931e48fd9e25119acfff1c79 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 14 May 2024 21:18:35 +0200 Subject: [PATCH 117/124] refactor using norm, etc --- .../EisensteinSeries/UniformConvergence.lean | 288 ++++++++---------- 1 file changed, 133 insertions(+), 155 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 60fa982b39d3e..1e6b85a3817de 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -15,6 +15,15 @@ import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic We show that the sum of `eisSummand` converges locally uniformly on `ℍ` to the Eisenstein series of weight `k` and level `Γ(N)` with congruence condition `a : Fin 2 → ZMod N`. + +## Outline of argument + +The key lemma `r_mul_max_le` shows that, for `z ∈ ℍ` and `c, d ∈ ℤ` (not both zero), +`|c z + d|` is bounded below by `r z * max (|c|, |d|)`, where `r z` is an explicit function of `z` +(independent of `c, d`) satisfying `0 < r z < 1` for all `z`. + +We then show in `summable_one_div_rpow_max` that the sum of `max (|c|, |d|) ^ (-k)` over +`(c, d) ∈ ℤ × ℤ` is convergent for `3 ≤ k`. This uses -/ noncomputable section @@ -23,204 +32,173 @@ open Complex UpperHalfPlane Set Finset open scoped BigOperators Matrix UpperHalfPlane Complex +/-- move this somewhere! -/ +lemma tendstoLocallyUniformly_iff_forall_isCompact {α β ι : Type*} [TopologicalSpace α] + [LocallyCompactSpace α] [UniformSpace β] {F : ι → α → β} {f : α → β} {p : Filter ι} : + TendstoLocallyUniformly F f p ↔ ∀ K : Set α, IsCompact K → TendstoUniformlyOn F f p K := by + simp only [← tendstoLocallyUniformlyOn_univ, + tendstoLocallyUniformlyOn_iff_forall_isCompact isOpen_univ, Set.subset_univ, forall_true_left] + +variable (z : ℍ) + namespace EisensteinSeries +lemma norm_def (x : Fin 2 → ℤ) : ‖x‖ = max (x 0).natAbs (x 1).natAbs := by + rw [← coe_nnnorm, ← NNReal.coe_natCast, NNReal.coe_inj, Nat.cast_max] + refine eq_of_forall_ge_iff fun c ↦ ?_ + simp only [pi_nnnorm_le_iff, Fin.forall_fin_two, max_le_iff, NNReal.natCast_natAbs] + section bounding_functions /-- Auxiliary function used for bounding Eisenstein series, defined as `z.im ^ 2 / (z.re ^ 2 + z.im ^ 2)`. -/ -def r1 (z : ℍ) : ℝ := z.im ^ 2 / (z.re ^ 2 + z.im ^ 2) +def r1 : ℝ := z.im ^ 2 / (z.re ^ 2 + z.im ^ 2) -lemma r1_eq (z : ℍ) : r1 z = 1 / ((z.re / z.im) ^ 2 + 1) := by +lemma r1_eq : r1 z = 1 / ((z.re / z.im) ^ 2 + 1) := by rw [div_pow, div_add_one (by positivity), one_div_div, r1] -theorem r1_pos (z : ℍ) : 0 < r1 z := by +theorem r1_pos : 0 < r1 z := by dsimp only [r1] positivity -/-- This function is used to give an upper bound on Eisenstein series. -/ -def r (z : ℍ) : ℝ := min z.im √(r1 z) +/-- For `c, d ∈ ℝ` with `1 ≤ d ^ 2`, we have `r1 z ≤ |c * z + d| ^ 2`. -/ +lemma r1_aux_bound (c : ℝ) {d : ℝ} (hd : 1 ≤ d ^ 2) : + r1 z ≤ (c * z.re + d) ^ 2 + (c * z.im) ^ 2 := by + have H1 : (c * z.re + d) ^ 2 + (c * z.im) ^ 2 = + c ^ 2 * (z.re ^ 2 + z.im ^ 2) + d * 2 * c * z.re + d ^ 2 := by ring + have H2 : (c ^ 2 * (z.re ^ 2 + z.im ^ 2) + d * 2 * c * z.re + d ^ 2) * (z.re ^ 2 + z.im ^ 2) + - z.im ^ 2 = (c * (z.re ^ 2 + z.im ^ 2) + d * z.re) ^ 2 + (d ^ 2 - 1) * z.im ^ 2 := by ring + rw [r1, H1, div_le_iff (by positivity), ← sub_nonneg, H2] + exact add_nonneg (sq_nonneg _) (mul_nonneg (sub_nonneg.mpr hd) (sq_nonneg _)) + +/-- This function is used to give an upper bound on the summands in Eisenstein series; it is +defined by `z ↦ min z.im √(z.im ^ 2 / (z.re ^ 2 + z.im ^ 2))`. -/ +def r : ℝ := min z.im √(r1 z) -lemma r_pos (z : ℍ) : 0 < r z := by +lemma r_pos : 0 < r z := by simp only [r, lt_min_iff, im_pos, Real.sqrt_pos, r1_pos, and_self] -lemma r1_aux_bound (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε ^ 2) : - r1 z ≤ (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 := by - have H1 : (δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2 = - δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε ^ 2 := by ring - have H2 : (δ ^ 2 * (z.re ^ 2 + z.im ^ 2) + ε * 2 * δ * z.re + ε ^ 2) * (z.re ^ 2 + z.im ^ 2) - - z.im ^ 2 = (δ * (z.re ^ 2 + z.im ^ 2) + ε * z.re) ^ 2 + (ε ^ 2 - 1) * z.im ^ 2 := by ring - rw [r1, H1, div_le_iff (by positivity), ← sub_nonneg, H2] - exact add_nonneg (sq_nonneg _) (mul_nonneg (sub_nonneg.mpr hε) (sq_nonneg _)) +lemma r_lower_bound_on_verticalStrip {A B : ℝ} (h : 0 < B) (hz : z ∈ verticalStrip A B) : + r ⟨⟨A, B⟩, h⟩ ≤ r z := by + apply min_le_min hz.2 + rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] + simp only [r1_eq, div_pow, one_div] + rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] + apply div_le_div (sq_nonneg _) _ (by positivity) (pow_le_pow_left h.le hz.2 2) + simpa only [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 -lemma auxbound1 (z : ℍ) {δ : ℝ} (ε : ℝ) (hδ : 1 ≤ δ ^ 2) : r z ≤ Complex.abs (δ * z + ε) := by +lemma auxbound1 {c : ℝ} (d : ℝ) (hc : 1 ≤ c ^ 2) : r z ≤ Complex.abs (c * z + d) := by rcases z with ⟨z, hz⟩ - have H1 : z.im ≤ √((δ * z.re + ε) ^ 2 + (δ * z).im ^ 2) := by + have H1 : z.im ≤ √((c * z.re + d) ^ 2 + (c * z).im ^ 2) := by rw [Real.le_sqrt' hz, im_ofReal_mul, mul_pow] - exact (le_mul_of_one_le_left (sq_nonneg _) hδ).trans <| le_add_of_nonneg_left (sq_nonneg _) + exact (le_mul_of_one_le_left (sq_nonneg _) hc).trans <| le_add_of_nonneg_left (sq_nonneg _) simpa only [r, abs_apply, normSq_apply, add_re, re_ofReal_mul, coe_re, ← pow_two, add_im, mul_im, coe_im, ofReal_im, zero_mul, add_zero, min_le_iff] using Or.inl H1 -lemma auxbound2 (z : ℍ) (δ : ℝ) {ε : ℝ} (hε : 1 ≤ ε ^ 2) : r z ≤ Complex.abs (δ * z + ε) := by - have H1 : √(r1 z) ≤ √((δ * z.re + ε) ^ 2 + (δ * z.im) ^ 2) := - (Real.sqrt_le_sqrt_iff (by positivity)).mpr (r1_aux_bound _ _ hε) +lemma auxbound2 (c : ℝ) {d : ℝ} (hd : 1 ≤ d ^ 2) : r z ≤ Complex.abs (c * z + d) := by + have H1 : √(r1 z) ≤ √((c * z.re + d) ^ 2 + (c * z.im) ^ 2) := + (Real.sqrt_le_sqrt_iff (by positivity)).mpr (r1_aux_bound _ _ hd) simpa only [r, abs_apply, normSq_apply, add_re, re_ofReal_mul, coe_re, ofReal_re, ← pow_two, add_im, im_ofReal_mul, coe_im, ofReal_im, add_zero, min_le_iff] using Or.inr H1 -lemma left_ne_zero_of_max_eq {x : Fin 2 → ℤ} (hx : x ≠ 0) - (h : max (x 0).natAbs (x 1).natAbs = (x 0).natAbs) : x 0 ≠ 0 := by - contrapose! hx - rw [hx, Int.natAbs_zero, max_eq_left_iff, Nat.le_zero, Int.natAbs_eq_zero] at h - ext1 j - fin_cases j <;> assumption - -lemma right_ne_zero_of_max_eq {x : Fin 2 → ℤ} (hx : x ≠ 0) - (h : (max (x 0).natAbs (x 1).natAbs) = (x 1).natAbs) : x 1 ≠ 0 := by - contrapose! hx - rw [hx, Int.natAbs_zero, max_eq_right_iff, Nat.le_zero, Int.natAbs_eq_zero] at h - ext1 j - fin_cases j <;> assumption - lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : - (1 : ℝ) ≤ (x 0 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 ∨ - (1 : ℝ) ≤ (x 1 / (max (x 0).natAbs (x 1).natAbs)) ^ 2 := by - refine (max_choice (x 0).natAbs (x 1).natAbs).imp (fun H1 => ?_) (fun H2 => ?_) - · simp only [Fin.isValue, H1, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, - OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, - left_ne_zero_of_max_eq hx H1, div_self, le_refl] - · simp only [Fin.isValue, H2, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, - OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, - right_ne_zero_of_max_eq hx H2, div_self, le_refl] - -lemma r_mul_max_le (z : ℍ) {x : Fin 2 → ℤ} (hx : x ≠ 0) : - r z * (max (x 0).natAbs (x 1).natAbs) ≤ Complex.abs (x 0 * z + x 1) := by - let n := max (x 0).natAbs (x 1).natAbs - have hn0 : n ≠ 0 := by - contrapose! hx - rw [Nat.max_eq_zero_iff, Int.natAbs_eq_zero, Int.natAbs_eq_zero] at hx - exact funext fun j ↦ by fin_cases j <;> tauto + 1 ≤ (x 0 / ‖x‖) ^ 2 ∨ 1 ≤ (x 1 / ‖x‖) ^ 2 := by + refine (max_choice (x 0).natAbs (x 1).natAbs).imp (fun H0 ↦ ?_) (fun H1 ↦ ?_) + · have : x 0 ≠ 0 := by + rwa [← norm_ne_zero_iff, norm_def, H0, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx + simp only [norm_def, H0, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, + OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, this, div_self, + le_refl] + · have : x 1 ≠ 0 := by + rwa [← norm_ne_zero_iff, norm_def, H1, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx + simp only [norm_def, H1, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, + OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, this, div_self, + le_refl] + +lemma r_mul_max_le {x : Fin 2 → ℤ} (hx : x ≠ 0) : r z * ‖x‖ ≤ Complex.abs (x 0 * z + x 1) := by + let n := ‖x‖ + have hn0 : n ≠ 0 := by rwa [norm_ne_zero_iff] have h11 : x 0 * (z : ℂ) + x 1 = (x 0 / n * z + x 1 / n) * n := by - push_cast - rw [div_mul_eq_mul_div, ← add_div, div_mul_cancel₀ _ (Nat.cast_ne_zero.mpr hn0)] - rw [Nat.cast_max, h11, map_mul, abs_natCast] + rw [div_mul_eq_mul_div, ← add_div, div_mul_cancel₀ _ (mod_cast hn0)] + rw [norm_def, h11, map_mul, Complex.abs_ofReal, abs_norm, norm_def] gcongr · rcases div_max_sq_ge_one x hx with H1 | H2 - · simpa only [Fin.isValue, ofReal_div, ofReal_intCast, n] using auxbound1 z (x 1 / n) H1 - · simpa only [Fin.isValue, ofReal_div, ofReal_intCast, n] using auxbound2 z (x 0 / n) H2 - · simp only [Fin.isValue, Nat.cast_max, le_refl, n] - -theorem summand_is_bounded_on_box_rpow {k : ℝ} (hk : 0 ≤ k) (z : ℍ) (n : ℕ) (x : Fin 2 → ℤ) - (hx : (x 0, x 1) ∈ box n) : Complex.abs (x 0 * z + x 1) ^ (-k) ≤ (r z) ^ (-k) * n ^ (-k) := by - by_cases hn : n = 0 - · simp only [hn, box_zero, Finset.mem_singleton, Prod.mk_eq_zero] at hx - rw [hx.1, hx.2, hn, ← Real.mul_rpow (r_pos z).le (Nat.cast_nonneg 0)] - simp only [Int.cast_zero, zero_mul, add_zero, map_zero, CharP.cast_eq_zero, mul_zero, le_refl] - · simp only [Fin.isValue, Pi.zero_apply, Int.mem_box, Int.natAbs_zero, max_self] at hx - have hx2 : x ≠ 0 := by - contrapose! hn - simpa only [hn, Fin.isValue, Pi.zero_apply, Int.natAbs_zero, max_self] using hx.symm - rw [← Real.mul_rpow (r_pos _).le (Nat.cast_nonneg _)] - exact Real.rpow_le_rpow_of_nonpos (mul_pos (r_pos _) (by positivity)) (hx ▸ r_mul_max_le z hx2) + · simpa only [norm_def, ofReal_div, ofReal_intCast, n] using auxbound1 z (x 1 / n) H1 + · simpa only [norm_def, ofReal_div, ofReal_intCast, n] using auxbound2 z (x 0 / n) H2 + +/-- Upper bound for the summand `|c * z + d| ^ (-k)`, as a product of a function of `z` and a +function of `c, d`. -/ +lemma summand_bound {k : ℝ} (hk : 0 ≤ k) (x : Fin 2 → ℤ) : + Complex.abs (x 0 * z + x 1) ^ (-k) ≤ (r z) ^ (-k) * ‖x‖ ^ (-k) := by + by_cases hx : x = 0 + · simp only [hx, Pi.zero_apply, Int.cast_zero, zero_mul, add_zero, ← norm_eq_abs, norm_zero] + by_cases h : -k = 0 + · rw [h, Real.rpow_zero, Real.rpow_zero, one_mul] + · rw [Real.zero_rpow h, mul_zero] + · rw [← Real.mul_rpow (r_pos _).le (norm_nonneg _)] + exact Real.rpow_le_rpow_of_nonpos (mul_pos (r_pos _) (norm_pos_iff.mpr hx)) (r_mul_max_le z hx) (neg_nonpos.mpr hk) -/- This is a special case of the above, but one that we use more. -/ -theorem eisSummand_is_bounded_on_box_zpow {k : ℤ} (n : ℕ) (z : ℍ) (x : Fin 2 → ℤ) (hk : 0 ≤ k) - (hx : (x 0, x 1) ∈ box n) : Complex.abs (eisSummand k x z) ≤ ((r z) ^ k * n ^ k)⁻¹ := by - have := summand_is_bounded_on_box_rpow (Int.cast_nonneg.2 hk) z n x hx - simp_rw [← Int.cast_neg k, Real.rpow_intCast, zpow_neg] at this - rw [mul_inv, eisSummand, one_div, map_inv₀, map_zpow₀] - exact this - -lemma r_lower_bound_on_verticalStrip {A B : ℝ} (h : 0 < B) (z : verticalStrip A B) : - r ⟨⟨A, B⟩, h⟩ ≤ r z.1 := by - rcases z with ⟨z, hz⟩ - simp only [mem_verticalStrip_iff, abs_ofReal] at hz - apply min_le_min hz.2 - rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] - simp only [r1_eq, div_pow, one_div] - rw [inv_le_inv (by positivity) (by positivity), add_le_add_iff_right] - apply div_le_div (sq_nonneg _) _ (by positivity) (pow_le_pow_left h.le hz.2 2) - simpa only [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 +variable {z} in +lemma summand_bound_of_mem_verticalStrip {k : ℝ} (hk : 0 ≤ k) (x : Fin 2 → ℤ) + {A B : ℝ} (hB : 0 < B) (hz : z ∈ verticalStrip A B) : + Complex.abs (x 0 * z + x 1) ^ (-k) ≤ r ⟨⟨A, B⟩, hB⟩ ^ (-k) * ‖x‖ ^ (-k) := by + refine (summand_bound z hk x).trans (mul_le_mul_of_nonneg_right ?_ (by positivity)) + exact Real.rpow_le_rpow_of_nonpos (r_pos _) (r_lower_bound_on_verticalStrip z hB hz) + (neg_nonpos.mpr hk) end bounding_functions section summability -/-An auxilary summable lemma needed later. Added the set_option as replacing `v` with - `_` didn't work. -/ -set_option linter.unusedVariables false -lemma summable_over_box {k : ℤ} (z : ℍ) (h : 3 ≤ k): - Summable (fun n : ℕ => ∑ v in (box n : Finset (ℤ × ℤ)), ((r z) ^ k * (n : ℝ) ^ k)⁻¹) := by - simp only [sum_const, nsmul_eq_mul] - have hk : 1 < (k - 1 : ℝ) := by norm_cast; exact Int.lt_sub_left_of_add_lt h - have nze : 8 * ((r z) ^ k)⁻¹ ≠ 0 := by - exact div_ne_zero (OfNat.ofNat_ne_zero 8) (zpow_ne_zero k (ne_of_gt (r_pos z))) - apply ((summable_mul_left_iff nze).mpr (Real.summable_nat_rpow_inv.2 hk)).congr - intro b - norm_cast - by_cases b0 : b = 0 - · simp only [b0, CharP.cast_eq_zero, zero_zpow (k - 1) (by omega), inv_zero, mul_zero, - box_zero, Finset.card_singleton, Nat.cast_one, zero_zpow k (by omega)] - · rw [Int.card_box b0, zpow_sub_one₀ (a:= (b : ℝ)) (Nat.cast_ne_zero.mpr b0) k] - simp only [mul_inv_rev, inv_inv, Nat.cast_mul, Nat.cast_ofNat] - ring_nf - -lemma summable_upper_bound {k : ℤ} (h : 3 ≤ k) (z : ℍ) : Summable fun (x : Fin 2 → ℤ) => - ((r z) ^ k * (max (x 0).natAbs (x 1).natAbs) ^ k)⁻¹ := by - rw [← (piFinTwoEquiv _).symm.summable_iff, - summable_partition _ (s := fun n ↦ (box n : Finset (ℤ × ℤ))) Int.existsUnique_mem_box] - · simp_rw [coe_sort_coe, Finset.tsum_subtype, piFinTwoEquiv_symm_apply, Function.comp_apply] - refine ⟨fun n ↦ ?_, (summable_over_box z h).congr fun n ↦ Finset.sum_congr rfl - fun x hx ↦ ?_⟩ - · apply (box n).summable ((fun x : Fin 2 → ℤ ↦ (((r z) ^ k) * - (max (x 0).natAbs (x 1).natAbs) ^ k)⁻¹) ∘ (piFinTwoEquiv _).symm) - · simp only [mul_inv_rev, mul_comm, Fin.isValue, Fin.cons_zero, Fin.cons_one, - Int.mem_box.mp hx] - · intro y - simp only [Pi.zero_apply, Fin.isValue, mul_inv_rev, piFinTwoEquiv_symm_apply, - Function.comp_apply, Fin.cons_zero, Fin.cons_one] - apply mul_nonneg - · simp only [piFinTwoEquiv_symm_apply, Fin.cons_zero, Fin.cons_one, inv_nonneg, ge_iff_le, - le_max_iff, Nat.cast_nonneg, or_self, zpow_nonneg] - · simp only [one_div, inv_nonneg] - apply zpow_nonneg (r_pos z).le - -end summability - -theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) (N : ℕ) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformly (fun (s : Finset (gammaSet N a)) => - (fun (z : ℍ) => ∑ x in s, eisSummand k x z)) - (fun (z : ℍ) => (eisensteinSeries_SIF a k).1 z) Filter.atTop := by - rw [← tendstoLocallyUniformlyOn_univ,tendstoLocallyUniformlyOn_iff_forall_isCompact - (by simp only [Set.top_eq_univ, isOpen_univ]), eisensteinSeries_SIF] - simp only [Set.top_eq_univ, Set.subset_univ, eisensteinSeries, forall_true_left] +/-- The function `ℤ ^ 2 → ℝ` given by `x ↦ ‖x‖ ^ (-k)` is summable if `2 < k`. We prove this by +splitting into boxes using `Finset.box`. -/ +lemma summable_one_div_norm_rpow {k : ℝ} (hk : 2 < k) : + Summable fun (x : Fin 2 → ℤ) ↦ ‖x‖ ^ (-k) := by + rw [← (finTwoArrowEquiv _).symm.summable_iff, summable_partition _ Int.existsUnique_mem_box] + · simp only [finTwoArrowEquiv_symm_apply, Function.comp_def] + refine ⟨fun n ↦ (hasSum_fintype (β := box (α := ℤ × ℤ) n) _).summable, ?_⟩ + suffices Summable fun n : ℕ ↦ ∑' (_ : box (α := ℤ × ℤ) n), (n : ℝ) ^ (-k) by + refine this.congr fun n ↦ tsum_congr fun p ↦ ?_ + simp only [← Int.mem_box.mp p.2, Nat.cast_max, norm_def, Matrix.cons_val_zero, + Matrix.cons_val_one, Matrix.head_cons] + simp only [tsum_fintype, univ_eq_attach, sum_const, card_attach, nsmul_eq_mul] + apply ((Real.summable_nat_rpow.mpr (by linarith : 1 - k < -1)).mul_left + 8).of_norm_bounded_eventually_nat + filter_upwards [Filter.eventually_gt_atTop 0] with n hn + rw [Int.card_box hn.ne', Real.norm_of_nonneg (by positivity), sub_eq_add_neg, + Real.rpow_add (Nat.cast_pos.mpr hn), Real.rpow_one, Nat.cast_mul, Nat.cast_ofNat, mul_assoc] + · exact fun n ↦ Real.rpow_nonneg (norm_nonneg _) _ + +/-- The sum defining the Eisenstein series (of weight `k` and level `Γ(N)` with congruence +condition `a : Fin 2 → ZMod N`) converges locally uniformly on `ℍ`. -/ +theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) {N : ℕ} (a : Fin 2 → ZMod N) : + TendstoLocallyUniformly (fun (s : Finset (gammaSet N a)) ↦ (∑ x in s, eisSummand k x ·)) + (eisensteinSeries a k ·) Filter.atTop := by + have hk' : (2 : ℝ) < k := by norm_cast + have p_sum : Summable fun x : gammaSet N a ↦ ‖x.val‖ ^ (-k) := + mod_cast (summable_one_div_norm_rpow hk').subtype (gammaSet N a) + simp only [tendstoLocallyUniformly_iff_forall_isCompact, eisensteinSeries] intro K hK obtain ⟨A, B, hB, HABK⟩ := subset_verticalStrip_of_isCompact hK - have hu : Summable fun x : (gammaSet N a) => - ((r ⟨⟨A, B⟩, hB⟩) ^ k * (max (x.1 0).natAbs (x.1 1).natAbs) ^ k)⁻¹ := by - apply ((summable_upper_bound hk ⟨⟨A, B⟩, hB⟩).subtype (gammaSet N a)).congr (fun _ => rfl) - apply tendstoUniformlyOn_tsum hu - intro v x hx - apply le_trans (eisSummand_is_bounded_on_box_zpow (k := k) (max (v.1 0).natAbs (v.1 1).natAbs) x v - (by omega) (by simp only [Int.mem_box])) - simp only [Fin.isValue, Nat.cast_max, mul_inv_rev] - lift k to ℕ using (by omega) - gcongr - · apply pow_pos (r_pos _) - · apply (r_pos _).le - · apply (r_lower_bound_on_verticalStrip hB ⟨x, HABK hx⟩) + refine (tendstoUniformlyOn_tsum (hu := p_sum.mul_left <| r ⟨⟨A, B⟩, hB⟩ ^ (-k : ℝ)) + (fun p z hz ↦ ?_)).mono HABK + simpa only [eisSummand, one_div, ← zpow_neg, norm_eq_abs, abs_zpow, ← Real.rpow_intCast, + Int.cast_neg] using summand_bound_of_mem_verticalStrip (by positivity) p hB hz -/-This is essentially extend by zero outside the upper half plane.-/ -local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm - (OpenEmbedding.toPartialHomeomorph UpperHalfPlane.coe openEmbedding_coe)) +/-- Extend a function on `ℍ` arbitrarily to a function on all of `ℂ`. -/ +local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm (openEmbedding_coe.toPartialHomeomorph _)) -/- A version for the extension to maps `ℂ → ℂ` that is nice to have for holomorphicity later. -/ +/-- Variant of `eisensteinSeries_tendstoLocallyUniformly` formulated with maps `ℂ → ℂ`, which is +nice to have for holomorphicity later. -/ lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} {N : ℕ} (hk : 3 ≤ k) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) => - ↑ₕ(fun (z : ℍ) => ∑ x in s, eisSummand k x z )) (↑ₕ(eisensteinSeries_SIF a k).toFun) + (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) ↦ + ↑ₕ(fun (z : ℍ) ↦ ∑ x in s, eisSummand k x z )) (↑ₕ(eisensteinSeries_SIF a k).toFun) Filter.atTop (UpperHalfPlane.coe '' ⊤) := by apply TendstoLocallyUniformlyOn.comp (s := ⊤) _ _ _ (PartialHomeomorph.continuousOn_symm _) · simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ] apply eisensteinSeries_tendstoLocallyUniformly hk · simp only [OpenEmbedding.toPartialHomeomorph_target, Set.top_eq_univ, mapsTo_range_iff, Set.mem_univ, forall_const] + +end summability From 2bb8c9b613b78195769c08d84e6ee22bfd7ef18c Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 14 May 2024 21:40:17 +0200 Subject: [PATCH 118/124] fix truncated docstring --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 1e6b85a3817de..a55c9b123779b 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -23,7 +23,8 @@ The key lemma `r_mul_max_le` shows that, for `z ∈ ℍ` and `c, d ∈ ℤ` (not (independent of `c, d`) satisfying `0 < r z < 1` for all `z`. We then show in `summable_one_div_rpow_max` that the sum of `max (|c|, |d|) ^ (-k)` over -`(c, d) ∈ ℤ × ℤ` is convergent for `3 ≤ k`. This uses +`(c, d) ∈ ℤ × ℤ` is convergent for `2 < k`. This is proved by decomposing `ℤ × ℤ` using the +`Finset.box` lemmas. -/ noncomputable section From 2cc89845ed6c51931a89a44b7701794502bb21c1 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 14 May 2024 22:24:07 +0200 Subject: [PATCH 119/124] move lemma --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 7 ------- Mathlib/Topology/UniformSpace/UniformConvergence.lean | 5 +++++ 2 files changed, 5 insertions(+), 7 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index a55c9b123779b..2e03056d26bea 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -33,13 +33,6 @@ open Complex UpperHalfPlane Set Finset open scoped BigOperators Matrix UpperHalfPlane Complex -/-- move this somewhere! -/ -lemma tendstoLocallyUniformly_iff_forall_isCompact {α β ι : Type*} [TopologicalSpace α] - [LocallyCompactSpace α] [UniformSpace β] {F : ι → α → β} {f : α → β} {p : Filter ι} : - TendstoLocallyUniformly F f p ↔ ∀ K : Set α, IsCompact K → TendstoUniformlyOn F f p K := by - simp only [← tendstoLocallyUniformlyOn_univ, - tendstoLocallyUniformlyOn_iff_forall_isCompact isOpen_univ, Set.subset_univ, forall_true_left] - variable (z : ℍ) namespace EisensteinSeries diff --git a/Mathlib/Topology/UniformSpace/UniformConvergence.lean b/Mathlib/Topology/UniformSpace/UniformConvergence.lean index 05974bd8d082e..a6ff670b7a722 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergence.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergence.lean @@ -757,6 +757,11 @@ theorem tendstoLocallyUniformlyOn_iff_forall_isCompact [LocallyCompactSpace α] (tendstoLocallyUniformlyOn_TFAE F f p hs).out 0 1 #align tendsto_locally_uniformly_on_iff_forall_is_compact tendstoLocallyUniformlyOn_iff_forall_isCompact +lemma tendstoLocallyUniformly_iff_forall_isCompact [LocallyCompactSpace α] : + TendstoLocallyUniformly F f p ↔ ∀ K : Set α, IsCompact K → TendstoUniformlyOn F f p K := by + simp only [← tendstoLocallyUniformlyOn_univ, + tendstoLocallyUniformlyOn_iff_forall_isCompact isOpen_univ, Set.subset_univ, forall_true_left] + theorem tendstoLocallyUniformlyOn_iff_filter : TendstoLocallyUniformlyOn F f p s ↔ ∀ x ∈ s, TendstoUniformlyOnFilter F f p (𝓝[s] x) := by simp only [TendstoUniformlyOnFilter, eventually_prod_iff] From 81c1c0c1ec2f8bec3eccb9bea57a28e0f636ce9d Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 14 May 2024 22:25:18 +0200 Subject: [PATCH 120/124] prune `open` statements --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 2e03056d26bea..3dd71cb5dcee3 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -31,7 +31,7 @@ noncomputable section open Complex UpperHalfPlane Set Finset -open scoped BigOperators Matrix UpperHalfPlane Complex +open scoped BigOperators variable (z : ℍ) From 80b223bb23e610451785fa87b15637f351a81968 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 15 May 2024 09:54:19 +0100 Subject: [PATCH 121/124] rename norm_def to norm_eq_max_natAbs --- .../EisensteinSeries/UniformConvergence.lean | 29 +++++++++---------- 1 file changed, 14 insertions(+), 15 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 3dd71cb5dcee3..0f35aac0aee75 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -37,7 +37,7 @@ variable (z : ℍ) namespace EisensteinSeries -lemma norm_def (x : Fin 2 → ℤ) : ‖x‖ = max (x 0).natAbs (x 1).natAbs := by +lemma norm_eq_max_abs (x : Fin 2 → ℤ) : ‖x‖ = max (x 0).natAbs (x 1).natAbs := by rw [← coe_nnnorm, ← NNReal.coe_natCast, NNReal.coe_inj, Nat.cast_max] refine eq_of_forall_ge_iff fun c ↦ ?_ simp only [pi_nnnorm_le_iff, Fin.forall_fin_two, max_le_iff, NNReal.natCast_natAbs] @@ -51,7 +51,7 @@ def r1 : ℝ := z.im ^ 2 / (z.re ^ 2 + z.im ^ 2) lemma r1_eq : r1 z = 1 / ((z.re / z.im) ^ 2 + 1) := by rw [div_pow, div_add_one (by positivity), one_div_div, r1] -theorem r1_pos : 0 < r1 z := by +lemma r1_pos : 0 < r1 z := by dsimp only [r1] positivity @@ -99,33 +99,32 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : 1 ≤ (x 0 / ‖x‖) ^ 2 ∨ 1 ≤ (x 1 / ‖x‖) ^ 2 := by refine (max_choice (x 0).natAbs (x 1).natAbs).imp (fun H0 ↦ ?_) (fun H1 ↦ ?_) · have : x 0 ≠ 0 := by - rwa [← norm_ne_zero_iff, norm_def, H0, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx - simp only [norm_def, H0, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, + rwa [← norm_ne_zero_iff, norm_eq_max_abs, H0, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx + simp only [norm_eq_max_abs, H0, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, this, div_self, le_refl] · have : x 1 ≠ 0 := by - rwa [← norm_ne_zero_iff, norm_def, H1, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx - simp only [norm_def, H1, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, + rwa [← norm_ne_zero_iff, norm_eq_max_abs, H1, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx + simp only [norm_eq_max_abs, H1, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, this, div_self, le_refl] lemma r_mul_max_le {x : Fin 2 → ℤ} (hx : x ≠ 0) : r z * ‖x‖ ≤ Complex.abs (x 0 * z + x 1) := by - let n := ‖x‖ - have hn0 : n ≠ 0 := by rwa [norm_ne_zero_iff] - have h11 : x 0 * (z : ℂ) + x 1 = (x 0 / n * z + x 1 / n) * n := by + have hn0 : ‖x‖ ≠ 0 := by rwa [norm_ne_zero_iff] + have h11 : x 0 * (z : ℂ) + x 1 = (x 0 / ‖x‖ * z + x 1 / ‖x‖) * ‖x‖ := by rw [div_mul_eq_mul_div, ← add_div, div_mul_cancel₀ _ (mod_cast hn0)] - rw [norm_def, h11, map_mul, Complex.abs_ofReal, abs_norm, norm_def] + rw [norm_eq_max_abs, h11, map_mul, Complex.abs_ofReal, abs_norm, norm_eq_max_abs] gcongr · rcases div_max_sq_ge_one x hx with H1 | H2 - · simpa only [norm_def, ofReal_div, ofReal_intCast, n] using auxbound1 z (x 1 / n) H1 - · simpa only [norm_def, ofReal_div, ofReal_intCast, n] using auxbound2 z (x 0 / n) H2 + · simpa only [norm_eq_max_abs, ofReal_div, ofReal_intCast] using auxbound1 z (x 1 / ‖x‖) H1 + · simpa only [norm_eq_max_abs, ofReal_div, ofReal_intCast] using auxbound2 z (x 0 / ‖x‖) H2 /-- Upper bound for the summand `|c * z + d| ^ (-k)`, as a product of a function of `z` and a function of `c, d`. -/ lemma summand_bound {k : ℝ} (hk : 0 ≤ k) (x : Fin 2 → ℤ) : Complex.abs (x 0 * z + x 1) ^ (-k) ≤ (r z) ^ (-k) * ‖x‖ ^ (-k) := by by_cases hx : x = 0 - · simp only [hx, Pi.zero_apply, Int.cast_zero, zero_mul, add_zero, ← norm_eq_abs, norm_zero] + · simp only [hx, Pi.zero_apply, Int.cast_zero, zero_mul, add_zero, ← norm_eq_max_abs_abs, norm_zero] by_cases h : -k = 0 · rw [h, Real.rpow_zero, Real.rpow_zero, one_mul] · rw [Real.zero_rpow h, mul_zero] @@ -154,7 +153,7 @@ lemma summable_one_div_norm_rpow {k : ℝ} (hk : 2 < k) : refine ⟨fun n ↦ (hasSum_fintype (β := box (α := ℤ × ℤ) n) _).summable, ?_⟩ suffices Summable fun n : ℕ ↦ ∑' (_ : box (α := ℤ × ℤ) n), (n : ℝ) ^ (-k) by refine this.congr fun n ↦ tsum_congr fun p ↦ ?_ - simp only [← Int.mem_box.mp p.2, Nat.cast_max, norm_def, Matrix.cons_val_zero, + simp only [← Int.mem_box.mp p.2, Nat.cast_max, norm_eq_max_abs, Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons] simp only [tsum_fintype, univ_eq_attach, sum_const, card_attach, nsmul_eq_mul] apply ((Real.summable_nat_rpow.mpr (by linarith : 1 - k < -1)).mul_left @@ -177,7 +176,7 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) {N : obtain ⟨A, B, hB, HABK⟩ := subset_verticalStrip_of_isCompact hK refine (tendstoUniformlyOn_tsum (hu := p_sum.mul_left <| r ⟨⟨A, B⟩, hB⟩ ^ (-k : ℝ)) (fun p z hz ↦ ?_)).mono HABK - simpa only [eisSummand, one_div, ← zpow_neg, norm_eq_abs, abs_zpow, ← Real.rpow_intCast, + simpa only [eisSummand, one_div, ← zpow_neg, norm_eq_max_abs_abs, abs_zpow, ← Real.rpow_intCast, Int.cast_neg] using summand_bound_of_mem_verticalStrip (by positivity) p hB hz /-- Extend a function on `ℍ` arbitrarily to a function on all of `ℂ`. -/ From 346bbd031133c50b21375f659b3e0ec344c61f3e Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 15 May 2024 09:54:27 +0100 Subject: [PATCH 122/124] rename norm_def to norm_eq_max_natAbs --- .../EisensteinSeries/UniformConvergence.lean | 22 +++++++++---------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 0f35aac0aee75..733e92308e638 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -37,7 +37,7 @@ variable (z : ℍ) namespace EisensteinSeries -lemma norm_eq_max_abs (x : Fin 2 → ℤ) : ‖x‖ = max (x 0).natAbs (x 1).natAbs := by +lemma norm_eq_max_natAbs (x : Fin 2 → ℤ) : ‖x‖ = max (x 0).natAbs (x 1).natAbs := by rw [← coe_nnnorm, ← NNReal.coe_natCast, NNReal.coe_inj, Nat.cast_max] refine eq_of_forall_ge_iff fun c ↦ ?_ simp only [pi_nnnorm_le_iff, Fin.forall_fin_two, max_le_iff, NNReal.natCast_natAbs] @@ -99,13 +99,13 @@ lemma div_max_sq_ge_one (x : Fin 2 → ℤ) (hx : x ≠ 0) : 1 ≤ (x 0 / ‖x‖) ^ 2 ∨ 1 ≤ (x 1 / ‖x‖) ^ 2 := by refine (max_choice (x 0).natAbs (x 1).natAbs).imp (fun H0 ↦ ?_) (fun H1 ↦ ?_) · have : x 0 ≠ 0 := by - rwa [← norm_ne_zero_iff, norm_eq_max_abs, H0, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx - simp only [norm_eq_max_abs, H0, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, + rwa [← norm_ne_zero_iff, norm_eq_max_natAbs, H0, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx + simp only [norm_eq_max_natAbs, H0, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, this, div_self, le_refl] · have : x 1 ≠ 0 := by - rwa [← norm_ne_zero_iff, norm_eq_max_abs, H1, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx - simp only [norm_eq_max_abs, H1, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, + rwa [← norm_ne_zero_iff, norm_eq_max_natAbs, H1, Nat.cast_ne_zero, Int.natAbs_ne_zero] at hx + simp only [norm_eq_max_natAbs, H1, Int.cast_natAbs, Int.cast_abs, div_pow, _root_.sq_abs, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, Int.cast_eq_zero, this, div_self, le_refl] @@ -113,18 +113,18 @@ lemma r_mul_max_le {x : Fin 2 → ℤ} (hx : x ≠ 0) : r z * ‖x‖ ≤ Comple have hn0 : ‖x‖ ≠ 0 := by rwa [norm_ne_zero_iff] have h11 : x 0 * (z : ℂ) + x 1 = (x 0 / ‖x‖ * z + x 1 / ‖x‖) * ‖x‖ := by rw [div_mul_eq_mul_div, ← add_div, div_mul_cancel₀ _ (mod_cast hn0)] - rw [norm_eq_max_abs, h11, map_mul, Complex.abs_ofReal, abs_norm, norm_eq_max_abs] + rw [norm_eq_max_natAbs, h11, map_mul, Complex.abs_ofReal, abs_norm, norm_eq_max_natAbs] gcongr · rcases div_max_sq_ge_one x hx with H1 | H2 - · simpa only [norm_eq_max_abs, ofReal_div, ofReal_intCast] using auxbound1 z (x 1 / ‖x‖) H1 - · simpa only [norm_eq_max_abs, ofReal_div, ofReal_intCast] using auxbound2 z (x 0 / ‖x‖) H2 + · simpa only [norm_eq_max_natAbs, ofReal_div, ofReal_intCast] using auxbound1 z (x 1 / ‖x‖) H1 + · simpa only [norm_eq_max_natAbs, ofReal_div, ofReal_intCast] using auxbound2 z (x 0 / ‖x‖) H2 /-- Upper bound for the summand `|c * z + d| ^ (-k)`, as a product of a function of `z` and a function of `c, d`. -/ lemma summand_bound {k : ℝ} (hk : 0 ≤ k) (x : Fin 2 → ℤ) : Complex.abs (x 0 * z + x 1) ^ (-k) ≤ (r z) ^ (-k) * ‖x‖ ^ (-k) := by by_cases hx : x = 0 - · simp only [hx, Pi.zero_apply, Int.cast_zero, zero_mul, add_zero, ← norm_eq_max_abs_abs, norm_zero] + · simp only [hx, Pi.zero_apply, Int.cast_zero, zero_mul, add_zero, ← norm_eq_abs, norm_zero] by_cases h : -k = 0 · rw [h, Real.rpow_zero, Real.rpow_zero, one_mul] · rw [Real.zero_rpow h, mul_zero] @@ -153,7 +153,7 @@ lemma summable_one_div_norm_rpow {k : ℝ} (hk : 2 < k) : refine ⟨fun n ↦ (hasSum_fintype (β := box (α := ℤ × ℤ) n) _).summable, ?_⟩ suffices Summable fun n : ℕ ↦ ∑' (_ : box (α := ℤ × ℤ) n), (n : ℝ) ^ (-k) by refine this.congr fun n ↦ tsum_congr fun p ↦ ?_ - simp only [← Int.mem_box.mp p.2, Nat.cast_max, norm_eq_max_abs, Matrix.cons_val_zero, + simp only [← Int.mem_box.mp p.2, Nat.cast_max, norm_eq_max_natAbs, Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons] simp only [tsum_fintype, univ_eq_attach, sum_const, card_attach, nsmul_eq_mul] apply ((Real.summable_nat_rpow.mpr (by linarith : 1 - k < -1)).mul_left @@ -176,7 +176,7 @@ theorem eisensteinSeries_tendstoLocallyUniformly {k : ℤ} (hk : 3 ≤ k) {N : obtain ⟨A, B, hB, HABK⟩ := subset_verticalStrip_of_isCompact hK refine (tendstoUniformlyOn_tsum (hu := p_sum.mul_left <| r ⟨⟨A, B⟩, hB⟩ ^ (-k : ℝ)) (fun p z hz ↦ ?_)).mono HABK - simpa only [eisSummand, one_div, ← zpow_neg, norm_eq_max_abs_abs, abs_zpow, ← Real.rpow_intCast, + simpa only [eisSummand, one_div, ← zpow_neg, norm_eq_abs, abs_zpow, ← Real.rpow_intCast, Int.cast_neg] using summand_bound_of_mem_verticalStrip (by positivity) p hB hz /-- Extend a function on `ℍ` arbitrarily to a function on all of `ℂ`. -/ From 1bb0494df6314b6e727a05692cb82f56d2cce7a4 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 15 May 2024 12:14:52 +0100 Subject: [PATCH 123/124] Update Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean Co-authored-by: Johan Commelin --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 733e92308e638..65f70755a6ae4 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -185,8 +185,8 @@ local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm (openEmbedding_coe.to /-- Variant of `eisensteinSeries_tendstoLocallyUniformly` formulated with maps `ℂ → ℂ`, which is nice to have for holomorphicity later. -/ lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} {N : ℕ} (hk : 3 ≤ k) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) ↦ - ↑ₕ(fun (z : ℍ) ↦ ∑ x in s, eisSummand k x z )) (↑ₕ(eisensteinSeries_SIF a k).toFun) + (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a)) ↦ + ↑ₕ(fun (z : ℍ) ↦ ∑ x in s, eisSummand k x z)) (↑ₕ(eisensteinSeries_SIF a k).toFun) Filter.atTop (UpperHalfPlane.coe '' ⊤) := by apply TendstoLocallyUniformlyOn.comp (s := ⊤) _ _ _ (PartialHomeomorph.continuousOn_symm _) · simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ] From 911192d322c19742f67eb87d3071b59f2d7cf350 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 15 May 2024 12:15:55 +0100 Subject: [PATCH 124/124] add David as Author --- .../ModularForms/EisensteinSeries/UniformConvergence.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index 65f70755a6ae4..cac947cc40afe 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Birkbeck +Authors: Chris Birkbeck, David Loeffler -/ import Mathlib.Analysis.Complex.UpperHalfPlane.Topology @@ -185,8 +185,8 @@ local notation "↑ₕ" f => f ∘ (PartialHomeomorph.symm (openEmbedding_coe.to /-- Variant of `eisensteinSeries_tendstoLocallyUniformly` formulated with maps `ℂ → ℂ`, which is nice to have for holomorphicity later. -/ lemma eisensteinSeries_tendstoLocallyUniformlyOn {k : ℤ} {N : ℕ} (hk : 3 ≤ k) - (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a)) ↦ - ↑ₕ(fun (z : ℍ) ↦ ∑ x in s, eisSummand k x z)) (↑ₕ(eisensteinSeries_SIF a k).toFun) + (a : Fin 2 → ZMod N) : TendstoLocallyUniformlyOn (fun (s : Finset (gammaSet N a )) ↦ + ↑ₕ(fun (z : ℍ) ↦ ∑ x in s, eisSummand k x z )) (↑ₕ(eisensteinSeries_SIF a k).toFun) Filter.atTop (UpperHalfPlane.coe '' ⊤) := by apply TendstoLocallyUniformlyOn.comp (s := ⊤) _ _ _ (PartialHomeomorph.continuousOn_symm _) · simp only [SlashInvariantForm.toFun_eq_coe, Set.top_eq_univ, tendstoLocallyUniformlyOn_univ]