Working…

This commit is contained in:
Stefan Kebekus 2024-12-19 21:24:13 +01:00
parent 4cc853a5d9
commit 1e28302c17
3 changed files with 115 additions and 19 deletions

View File

@ -77,7 +77,7 @@ theorem Nevanlinna_proximity₀
{f : } {f : }
{r : } {r : }
(h₁f : MeromorphicOn f ) (h₁f : MeromorphicOn f )
(hr : 0 < r ) : (hr : 0 < r) :
(2 * π)⁻¹ * ∫ x in (0)..(2 * π), log ‖f (circleMap 0 r x)‖ = (h₁f.m_infty r) - (h₁f.inv.m_infty r) := by (2 * π)⁻¹ * ∫ x in (0)..(2 * π), log ‖f (circleMap 0 r x)‖ = (h₁f.m_infty r) - (h₁f.inv.m_infty r) := by
unfold MeromorphicOn.m_infty unfold MeromorphicOn.m_infty
@ -88,10 +88,13 @@ theorem Nevanlinna_proximity₀
exact Eq.symm (IsAbsoluteValue.abv_inv Norm.norm (f (circleMap 0 r x))) exact Eq.symm (IsAbsoluteValue.abv_inv Norm.norm (f (circleMap 0 r x)))
-- --
apply MeromorphicOn.integrable_logpos_abs_f hr apply MeromorphicOn.integrable_logpos_abs_f hr
intro z hx
exact h₁f z trivial
--
apply MeromorphicOn.integrable_logpos_abs_f hr
exact MeromorphicOn.inv_iff.mpr fun x a => h₁f x trivial
sorry
theorem Nevanlinna_proximity theorem Nevanlinna_proximity
{f : } {f : }
{r : } {r : }
@ -105,9 +108,8 @@ theorem Nevanlinna_proximity
rw [← mul_sub]; congr rw [← mul_sub]; congr
exact loglogpos exact loglogpos
by_cases h₂r : 0 < r
· exact Nevanlinna_proximity₀ h₁f h₂r
have hr : 0 < r := by sorry
unfold MeromorphicOn.m_infty unfold MeromorphicOn.m_infty
rw [← mul_sub]; congr rw [← mul_sub]; congr
@ -116,8 +118,10 @@ theorem Nevanlinna_proximity
simp_rw [loglogpos]; congr simp_rw [loglogpos]; congr
exact Eq.symm (IsAbsoluteValue.abv_inv Norm.norm (f (circleMap 0 r x))) exact Eq.symm (IsAbsoluteValue.abv_inv Norm.norm (f (circleMap 0 r x)))
-- --
apply MeromorphicOn.integrable_logpos_abs_f hr apply MeromorphicOn.integrable_logpos_abs_f
sorry
sorry sorry
noncomputable def MeromorphicOn.T_infty noncomputable def MeromorphicOn.T_infty

View File

@ -1,13 +1,5 @@
import Mathlib.Analysis.Analytic.Meromorphic
import Mathlib.MeasureTheory.Integral.CircleIntegral import Mathlib.MeasureTheory.Integral.CircleIntegral
import Mathlib.MeasureTheory.Integral.IntervalIntegral import Nevanlinna.periodic_integrability
import Nevanlinna.analyticAt
import Nevanlinna.divisor
import Nevanlinna.meromorphicAt
import Nevanlinna.meromorphicOn_divisor
import Nevanlinna.stronglyMeromorphicOn
import Nevanlinna.stronglyMeromorphicOn_eliminate
import Nevanlinna.mathlibAddOn
open scoped Interval Topology open scoped Interval Topology
open Real Filter MeasureTheory intervalIntegral open Real Filter MeasureTheory intervalIntegral
@ -96,3 +88,61 @@ theorem integral_congr_changeDiscrete
constructor constructor
· apply Set.Countable.measure_zero (d hr hU hf) · apply Set.Countable.measure_zero (d hr hU hf)
· tauto · tauto
lemma circleMap_neg
{r x : } :
circleMap 0 (-r) x = circleMap 0 r (x + π) := by
unfold circleMap
simp
rw [add_mul]
rw [Complex.exp_add]
simp
theorem integrability_congr_negRadius
{f : }
{r : } :
IntervalIntegrable (fun (θ : ) ↦ f (circleMap 0 r θ)) MeasureTheory.volume 0 (2 * π) →
IntervalIntegrable (fun (θ : ) ↦ f (circleMap 0 (-r) θ)) MeasureTheory.volume 0 (2 * π) := by
intro h
simp_rw [circleMap_neg]
let A := IntervalIntegrable.comp_add_right h π
have t₀ : Function.Periodic (fun (θ : ) ↦ f (circleMap 0 r θ)) (2 * π) := by
intro x
simp
congr 1
exact periodic_circleMap 0 r x
rw [← zero_add (2 * π)] at h
have B := periodic_integrability4 (a₁ := π) (a₂ := 3 * π) t₀ two_pi_pos h
let A := IntervalIntegrable.comp_add_right B π
simp at A
have : 3 * π - π = 2 * π := by
ring
rw [this] at A
exact A
theorem integrabl_congr_negRadius
{f : }
{r : } :
∫ (x : ) in (0)..(2 * π), f (circleMap 0 r x) = ∫ (x : ) in (0)..(2 * π), f (circleMap 0 (-r) x) := by
simp_rw [circleMap_neg]
have t₀ : Function.Periodic (fun (θ : ) ↦ f (circleMap 0 r θ)) (2 * π) := by
intro x
simp
congr 1
exact periodic_circleMap 0 r x
--rw [← zero_add (2 * π)] at h
--have B := periodic_integrability4 (a₁ := π) (a₂ := 3 * π) t₀ two_pi_pos h
have B := intervalIntegral.integral_comp_add_right (a := 0) (b := 2 * π) (fun θ => f (circleMap 0 r θ)) π
rw [B]
let X := t₀.intervalIntegral_add_eq 0 (0 + π)
simp at X
rw [X]
simp
rw [add_comm]

View File

@ -16,9 +16,10 @@ open scoped Interval Topology
open Real Filter MeasureTheory intervalIntegral open Real Filter MeasureTheory intervalIntegral
theorem MeromorphicOn.integrable_log_abs_f theorem MeromorphicOn.integrable_log_abs_f
{f : } {f : }
{r : } {r : }
-- WARNING: Not optimal. This needs to go
(hr : 0 < r) (hr : 0 < r)
-- WARNING: Not optimal. It suffices to be meromorphic on the Sphere -- WARNING: Not optimal. It suffices to be meromorphic on the Sphere
(h₁f : MeromorphicOn f (Metric.closedBall (0 : ) r)) : (h₁f : MeromorphicOn f (Metric.closedBall (0 : ) r)) :
@ -111,10 +112,51 @@ theorem MeromorphicOn.integrable_log_abs_f
rw [integrability_congr_changeDiscrete hU this] rw [integrability_congr_changeDiscrete hU this]
have : ∀ x ∈ Metric.closedBall 0 r, F x = 0 := by have : ∀ x ∈ Metric.closedBall 0 r, F x = 0 := by
sorry intro x hx
let A := h₂f ⟨x, hx⟩
rw [← makeStronglyMeromorphicOn_changeOrder h₁f hx] at A
let B := ((stronglyMeromorphicOn_of_makeStronglyMeromorphicOn h₁f) x hx).order_eq_zero_iff.not.1
simp [A] at B
assumption
have : (fun z => log ‖F z‖) ∘ circleMap 0 r = 0 := by
funext x
simp
left
apply this
simp [abs_of_pos hr]
simp_rw [this]
apply intervalIntegral.intervalIntegrable_const
theorem MeromorphicOn.integrable_log_abs_f
{f : }
{r : }
-- WARNING: Not optimal. It suffices to be meromorphic on the Sphere
(h₁f : MeromorphicOn f (Metric.closedBall (0 : ) |r|)) :
IntervalIntegrable (fun z ↦ log ‖f (circleMap 0 r z)‖) MeasureTheory.volume 0 (2 * π) := by
by_cases h₁r : r = 0
· rw [h₁r]
simp
· by_cases h₂r : 0 < r
· have : |r| = r := by
exact abs_of_pos h₂r
rw [this] at h₁f
exact MeromorphicOn.integrable_log_abs_f₀ h₂r h₁f
· have t₀ : 0 < -r := by
sorry sorry
have : |r| = -r := by
apply abs_of_neg
exact Left.neg_pos_iff.mp t₀
rw [this] at h₁f
let A := MeromorphicOn.integrable_log_abs_f₀ t₀ h₁f
let B := integrability_congr_negRadius (f := fun z => log ‖f z‖) (r := -r)
let C := B A
simp at C
simpa
theorem MeromorphicOn.integrable_logpos_abs_f theorem MeromorphicOn.integrable_logpos_abs_f
{f : } {f : }