|
| 1 | +/- |
| 2 | +Copyright (c) 2022 Chris Birkbeck. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Chris Birkbeck, David Loeffler |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module analysis.complex.upper_half_plane.functions_bounded_at_infty |
| 7 | +! leanprover-community/mathlib commit f2ce6086713c78a7f880485f7917ea547a215982 |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.Algebra.Module.Submodule.Basic |
| 12 | +import Mathlib.Analysis.Complex.UpperHalfPlane.Basic |
| 13 | +import Mathlib.Order.Filter.ZeroAndBoundedAtFilter |
| 14 | + |
| 15 | +/-! |
| 16 | +# Bounded at infinity |
| 17 | +
|
| 18 | +For complex valued functions on the upper half plane, this file defines the filter `atImInfty` |
| 19 | +required for defining when functions are bounded at infinity and zero at infinity. |
| 20 | +Both of which are relevant for defining modular forms. |
| 21 | +
|
| 22 | +-/ |
| 23 | + |
| 24 | + |
| 25 | +open Complex Filter |
| 26 | + |
| 27 | +open scoped Topology UpperHalfPlane |
| 28 | + |
| 29 | +noncomputable section |
| 30 | + |
| 31 | +namespace UpperHalfPlane |
| 32 | + |
| 33 | +/-- Filter for approaching `i∞`. -/ |
| 34 | +def atImInfty := |
| 35 | + Filter.atTop.comap UpperHalfPlane.im |
| 36 | +#align upper_half_plane.at_im_infty UpperHalfPlane.atImInfty |
| 37 | + |
| 38 | +theorem atImInfty_basis : atImInfty.HasBasis (fun _ => True) fun i : ℝ => im ⁻¹' Set.Ici i := |
| 39 | + Filter.HasBasis.comap UpperHalfPlane.im Filter.atTop_basis |
| 40 | +#align upper_half_plane.at_im_infty_basis UpperHalfPlane.atImInfty_basis |
| 41 | + |
| 42 | +theorem atImInfty_mem (S : Set ℍ) : S ∈ atImInfty ↔ ∃ A : ℝ, ∀ z : ℍ, A ≤ im z → z ∈ S := by |
| 43 | + simp only [atImInfty, Filter.mem_comap', Filter.mem_atTop_sets, ge_iff_le, Set.mem_setOf_eq, |
| 44 | + UpperHalfPlane.coe_im] |
| 45 | + refine' ⟨fun ⟨a, h⟩ => ⟨a, fun z hz => h (im z) hz rfl⟩, _⟩ |
| 46 | + rintro ⟨A, h⟩ |
| 47 | + refine' ⟨A, fun b hb x hx => h x _⟩ |
| 48 | + rwa [hx] |
| 49 | +#align upper_half_plane.at_im_infty_mem UpperHalfPlane.atImInfty_mem |
| 50 | + |
| 51 | +/-- A function ` f : ℍ → α` is bounded at infinity if it is bounded along `atImInfty`. -/ |
| 52 | +def IsBoundedAtImInfty {α : Type _} [Norm α] (f : ℍ → α) : Prop := |
| 53 | + BoundedAtFilter atImInfty f |
| 54 | +#align upper_half_plane.is_bounded_at_im_infty UpperHalfPlane.IsBoundedAtImInfty |
| 55 | + |
| 56 | +/-- A function ` f : ℍ → α` is zero at infinity it is zero along `atImInfty`. -/ |
| 57 | +def IsZeroAtImInfty {α : Type _} [Zero α] [TopologicalSpace α] (f : ℍ → α) : Prop := |
| 58 | + ZeroAtFilter atImInfty f |
| 59 | +#align upper_half_plane.is_zero_at_im_infty UpperHalfPlane.IsZeroAtImInfty |
| 60 | + |
| 61 | +theorem zero_form_isBoundedAtImInfty {α : Type _} [NormedField α] : |
| 62 | + IsBoundedAtImInfty (0 : ℍ → α) := |
| 63 | + const_boundedAtFilter atImInfty (0 : α) |
| 64 | +#align upper_half_plane.zero_form_is_bounded_at_im_infty UpperHalfPlane.zero_form_isBoundedAtImInfty |
| 65 | + |
| 66 | +/-- Module of functions that are zero at infinity. -/ |
| 67 | +def zeroAtImInftySubmodule (α : Type _) [NormedField α] : Submodule α (ℍ → α) := |
| 68 | + zeroAtFilterSubmodule atImInfty |
| 69 | +#align upper_half_plane.zero_at_im_infty_submodule UpperHalfPlane.zeroAtImInftySubmodule |
| 70 | + |
| 71 | +/-- Subalgebra of functions that are bounded at infinity. -/ |
| 72 | +def boundedAtImInftySubalgebra (α : Type _) [NormedField α] : Subalgebra α (ℍ → α) := |
| 73 | + boundedFilterSubalgebra atImInfty |
| 74 | +#align upper_half_plane.bounded_at_im_infty_subalgebra UpperHalfPlane.boundedAtImInftySubalgebra |
| 75 | + |
| 76 | +nonrec theorem IsBoundedAtImInfty.mul {f g : ℍ → ℂ} (hf : IsBoundedAtImInfty f) |
| 77 | + (hg : IsBoundedAtImInfty g) : IsBoundedAtImInfty (f * g) := by |
| 78 | + simpa only [Pi.one_apply, mul_one, norm_eq_abs] using hf.mul hg |
| 79 | +#align upper_half_plane.is_bounded_at_im_infty.mul UpperHalfPlane.IsBoundedAtImInfty.mul |
| 80 | + |
| 81 | +theorem bounded_mem (f : ℍ → ℂ) : |
| 82 | + IsBoundedAtImInfty f ↔ ∃ M A : ℝ, ∀ z : ℍ, A ≤ im z → abs (f z) ≤ M := by |
| 83 | + simp [IsBoundedAtImInfty, BoundedAtFilter, Asymptotics.isBigO_iff, Filter.Eventually, |
| 84 | + atImInfty_mem] |
| 85 | +#align upper_half_plane.bounded_mem UpperHalfPlane.bounded_mem |
| 86 | + |
| 87 | +theorem zero_at_im_infty (f : ℍ → ℂ) : |
| 88 | + IsZeroAtImInfty f ↔ ∀ ε : ℝ, 0 < ε → ∃ A : ℝ, ∀ z : ℍ, A ≤ im z → abs (f z) ≤ ε := by |
| 89 | + rw [IsZeroAtImInfty, ZeroAtFilter, tendsto_iff_forall_eventually_mem] |
| 90 | + constructor |
| 91 | + · simp_rw [Filter.Eventually, atImInfty_mem] |
| 92 | + intro h ε hε |
| 93 | + simpa using h (Metric.closedBall (0 : ℂ) ε) (Metric.closedBall_mem_nhds (0 : ℂ) hε) |
| 94 | + · simp_rw [Metric.mem_nhds_iff] |
| 95 | + intro h s hs |
| 96 | + simp_rw [Filter.Eventually, atImInfty_mem] |
| 97 | + obtain ⟨ε, h1, h2⟩ := hs |
| 98 | + have h11 : 0 < ε / 2 := by linarith |
| 99 | + obtain ⟨A, hA⟩ := h (ε / 2) h11 |
| 100 | + use A |
| 101 | + intro z hz |
| 102 | + have hzs : f z ∈ s := by |
| 103 | + apply h2 |
| 104 | + simp only [mem_ball_zero_iff, norm_eq_abs] |
| 105 | + apply lt_of_le_of_lt (hA z hz) |
| 106 | + linarith |
| 107 | + apply hzs |
| 108 | +#align upper_half_plane.zero_at_im_infty UpperHalfPlane.zero_at_im_infty |
| 109 | + |
| 110 | +end UpperHalfPlane |
0 commit comments