2024-12-11 12:36:44 +01:00
|
|
|
|
import Mathlib.Analysis.Analytic.Meromorphic
|
2024-12-13 16:03:00 +01:00
|
|
|
|
import Mathlib.MeasureTheory.Integral.CircleIntegral
|
|
|
|
|
import Mathlib.MeasureTheory.Integral.IntervalIntegral
|
2024-12-11 12:36:44 +01:00
|
|
|
|
import Nevanlinna.analyticAt
|
|
|
|
|
import Nevanlinna.divisor
|
|
|
|
|
import Nevanlinna.meromorphicAt
|
|
|
|
|
import Nevanlinna.meromorphicOn_divisor
|
|
|
|
|
import Nevanlinna.stronglyMeromorphicOn
|
2024-12-16 17:17:43 +01:00
|
|
|
|
import Nevanlinna.stronglyMeromorphicOn_eliminate
|
2024-12-11 12:36:44 +01:00
|
|
|
|
import Nevanlinna.mathlibAddOn
|
|
|
|
|
|
|
|
|
|
open scoped Interval Topology
|
|
|
|
|
open Real Filter MeasureTheory intervalIntegral
|
|
|
|
|
|
|
|
|
|
|
2024-12-16 07:36:01 +01:00
|
|
|
|
/- Integral and Integrability up to changes on codiscrete sets -/
|
|
|
|
|
|
2024-12-13 12:26:05 +01:00
|
|
|
|
theorem d
|
|
|
|
|
{U S : Set ℂ}
|
2024-12-13 16:03:00 +01:00
|
|
|
|
{c : ℂ}
|
2024-12-13 12:26:05 +01:00
|
|
|
|
{r : ℝ}
|
2024-12-13 16:03:00 +01:00
|
|
|
|
(hr : r ≠ 0)
|
|
|
|
|
(hU : Metric.sphere c |r| ⊆ U)
|
2024-12-13 12:26:05 +01:00
|
|
|
|
(hS : S ∈ Filter.codiscreteWithin U) :
|
2024-12-13 16:03:00 +01:00
|
|
|
|
Countable ((circleMap c r)⁻¹' Sᶜ) := by
|
2024-12-13 12:26:05 +01:00
|
|
|
|
|
2024-12-13 16:03:00 +01:00
|
|
|
|
have : (circleMap c r)⁻¹' (S ∪ Uᶜ)ᶜ = (circleMap c r)⁻¹' Sᶜ := by
|
|
|
|
|
simp [(by simpa : (circleMap c r)⁻¹' U = ⊤)]
|
|
|
|
|
rw [← this]
|
2024-12-13 12:26:05 +01:00
|
|
|
|
|
2024-12-13 16:03:00 +01:00
|
|
|
|
apply Set.Countable.preimage_circleMap _ c hr
|
|
|
|
|
have : DiscreteTopology ((S ∪ Uᶜ)ᶜ : Set ℂ) := by
|
|
|
|
|
rw [discreteTopology_subtype_iff]
|
|
|
|
|
rw [mem_codiscreteWithin] at hS; simp at hS
|
2024-12-13 12:26:05 +01:00
|
|
|
|
|
2024-12-13 16:03:00 +01:00
|
|
|
|
intro x hx
|
|
|
|
|
rw [← mem_iff_inf_principal_compl, (by ext z; simp; tauto : S ∪ Uᶜ = (U \ S)ᶜ)]
|
|
|
|
|
rw [Set.compl_union, compl_compl] at hx
|
|
|
|
|
exact hS x hx.2
|
2024-12-13 12:26:05 +01:00
|
|
|
|
|
2024-12-13 16:03:00 +01:00
|
|
|
|
apply TopologicalSpace.separableSpace_iff_countable.1
|
|
|
|
|
exact TopologicalSpace.SecondCountableTopology.to_separableSpace
|
2024-12-13 12:26:05 +01:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
theorem integrability_congr_changeDiscrete₀
|
2024-12-16 17:17:43 +01:00
|
|
|
|
{f₁ f₂ : ℂ → ℝ}
|
2024-12-13 07:58:40 +01:00
|
|
|
|
{U : Set ℂ}
|
2024-12-11 12:36:44 +01:00
|
|
|
|
{r : ℝ}
|
2024-12-13 12:26:05 +01:00
|
|
|
|
(hU : Metric.sphere 0 |r| ⊆ U)
|
2024-12-13 07:58:40 +01:00
|
|
|
|
(hf : f₁ =ᶠ[Filter.codiscreteWithin U] f₂) :
|
2024-12-11 12:36:44 +01:00
|
|
|
|
IntervalIntegrable (f₁ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) → IntervalIntegrable (f₂ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) := by
|
|
|
|
|
intro hf₁
|
|
|
|
|
|
2024-12-13 12:26:05 +01:00
|
|
|
|
by_cases hr : r = 0
|
|
|
|
|
· unfold circleMap
|
|
|
|
|
rw [hr]
|
|
|
|
|
simp
|
|
|
|
|
have : f₂ ∘ (fun (θ : ℝ) ↦ 0) = (fun r ↦ f₂ 0) := by
|
2024-12-11 12:36:44 +01:00
|
|
|
|
exact rfl
|
|
|
|
|
rw [this]
|
2024-12-13 07:58:40 +01:00
|
|
|
|
simp
|
2024-12-13 12:26:05 +01:00
|
|
|
|
|
|
|
|
|
· apply IntervalIntegrable.congr hf₁
|
|
|
|
|
rw [Filter.eventuallyEq_iff_exists_mem]
|
2024-12-13 16:03:00 +01:00
|
|
|
|
use (circleMap 0 r)⁻¹' {z | f₁ z = f₂ z}
|
2024-12-13 12:26:05 +01:00
|
|
|
|
constructor
|
2024-12-13 16:03:00 +01:00
|
|
|
|
· apply Set.Countable.measure_zero (d hr hU hf)
|
|
|
|
|
· tauto
|
|
|
|
|
|
2024-12-13 12:26:05 +01:00
|
|
|
|
|
|
|
|
|
theorem integrability_congr_changeDiscrete
|
2024-12-16 17:17:43 +01:00
|
|
|
|
{f₁ f₂ : ℂ → ℝ}
|
2024-12-13 12:26:05 +01:00
|
|
|
|
{U : Set ℂ}
|
|
|
|
|
{r : ℝ}
|
2024-12-16 17:17:43 +01:00
|
|
|
|
(hU : Metric.sphere (0 : ℂ) |r| ⊆ U)
|
2024-12-13 12:26:05 +01:00
|
|
|
|
(hf : f₁ =ᶠ[Filter.codiscreteWithin U] f₂) :
|
|
|
|
|
IntervalIntegrable (f₁ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) ↔ IntervalIntegrable (f₂ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) := by
|
|
|
|
|
|
|
|
|
|
constructor
|
|
|
|
|
· exact integrability_congr_changeDiscrete₀ hU hf
|
|
|
|
|
· exact integrability_congr_changeDiscrete₀ hU (EventuallyEq.symm hf)
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
theorem integral_congr_changeDiscrete
|
2024-12-16 17:17:43 +01:00
|
|
|
|
{f₁ f₂ : ℂ → ℝ}
|
2024-12-13 12:26:05 +01:00
|
|
|
|
{U : Set ℂ}
|
|
|
|
|
{r : ℝ}
|
2024-12-13 16:03:00 +01:00
|
|
|
|
(hr : r ≠ 0)
|
2024-12-13 12:26:05 +01:00
|
|
|
|
(hU : Metric.sphere 0 |r| ⊆ U)
|
|
|
|
|
(hf : f₁ =ᶠ[Filter.codiscreteWithin U] f₂) :
|
|
|
|
|
∫ (x : ℝ) in (0)..(2 * π), f₁ (circleMap 0 r x) = ∫ (x : ℝ) in (0)..(2 * π), f₂ (circleMap 0 r x) := by
|
|
|
|
|
|
|
|
|
|
apply intervalIntegral.integral_congr_ae
|
|
|
|
|
rw [eventually_iff_exists_mem]
|
2024-12-13 16:03:00 +01:00
|
|
|
|
use (circleMap 0 r)⁻¹' {z | f₁ z = f₂ z}
|
|
|
|
|
constructor
|
|
|
|
|
· apply Set.Countable.measure_zero (d hr hU hf)
|
|
|
|
|
· tauto
|
2024-12-16 17:17:43 +01:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
theorem MeromorphicOn.integrable_log_abs_f
|
|
|
|
|
{f : ℂ → ℂ}
|
|
|
|
|
{r : ℝ}
|
|
|
|
|
(hr : 0 < r)
|
|
|
|
|
(h₁f : MeromorphicOn f (Metric.closedBall (0 : ℂ) r))
|
|
|
|
|
(h₂f : ∃ u : (Metric.closedBall (0 : ℂ) r), (h₁f u u.2).order ≠ ⊤) :
|
|
|
|
|
IntervalIntegrable (fun z ↦ log ‖f (circleMap 0 r z)‖) MeasureTheory.volume 0 (2 * π) := by
|
|
|
|
|
|
|
|
|
|
have h₁U : IsCompact (Metric.closedBall (0 : ℂ) r) := by sorry
|
|
|
|
|
have h₂U : IsConnected (Metric.closedBall (0 : ℂ) r) := by sorry
|
|
|
|
|
have h₃U : interior (Metric.closedBall (0 : ℂ) r) ≠ ∅ := by sorry
|
|
|
|
|
|
|
|
|
|
obtain ⟨g, h₁g, h₂g, h₃g⟩ := MeromorphicOn.decompose_log h₁U h₂U h₃U h₁f h₂f
|
|
|
|
|
have : (fun z ↦ log ‖f (circleMap 0 r z)‖) = (fun z ↦ log ‖f z‖) ∘ (circleMap 0 r) := by
|
|
|
|
|
rfl
|
|
|
|
|
rw [this]
|
|
|
|
|
have : Metric.sphere (0 : ℂ) |r| ⊆ Metric.closedBall (0 : ℂ) r := by
|
|
|
|
|
sorry
|
|
|
|
|
rw [integrability_congr_changeDiscrete this h₃g]
|
|
|
|
|
|
|
|
|
|
apply IntervalIntegrable.add
|
|
|
|
|
sorry
|
|
|
|
|
sorry
|