2024-10-31 17:09:09 +01:00
|
|
|
|
import Mathlib.Analysis.Analytic.Meromorphic
|
|
|
|
|
import Nevanlinna.analyticAt
|
|
|
|
|
import Nevanlinna.divisor
|
2024-11-07 09:53:34 +01:00
|
|
|
|
import Nevanlinna.meromorphicAt
|
2024-10-31 17:09:09 +01:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
open scoped Interval Topology
|
|
|
|
|
open Real Filter MeasureTheory intervalIntegral
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
noncomputable def MeromorphicOn.divisor
|
|
|
|
|
{f : ℂ → ℂ}
|
|
|
|
|
{U : Set ℂ}
|
|
|
|
|
(hf : MeromorphicOn f U) :
|
|
|
|
|
Divisor U where
|
|
|
|
|
|
|
|
|
|
toFun := by
|
|
|
|
|
intro z
|
|
|
|
|
if hz : z ∈ U then
|
|
|
|
|
exact ((hf z hz).order.untop' 0 : ℤ)
|
|
|
|
|
else
|
|
|
|
|
exact 0
|
|
|
|
|
|
|
|
|
|
supportInU := by
|
|
|
|
|
intro z hz
|
|
|
|
|
simp at hz
|
|
|
|
|
by_contra h₂z
|
|
|
|
|
simp [h₂z] at hz
|
|
|
|
|
|
|
|
|
|
locallyFiniteInU := by
|
|
|
|
|
intro z hz
|
|
|
|
|
|
|
|
|
|
apply eventually_nhdsWithin_iff.2
|
|
|
|
|
rw [eventually_nhds_iff]
|
|
|
|
|
|
2024-11-04 13:22:12 +01:00
|
|
|
|
rcases MeromorphicAt.eventually_eq_zero_or_eventually_ne_zero (hf z hz) with h|h
|
|
|
|
|
· rw [eventually_nhdsWithin_iff] at h
|
|
|
|
|
rw [eventually_nhds_iff] at h
|
2024-10-31 17:09:09 +01:00
|
|
|
|
obtain ⟨N, h₁N, h₂N, h₃N⟩ := h
|
|
|
|
|
use N
|
|
|
|
|
constructor
|
2024-11-04 13:22:12 +01:00
|
|
|
|
· intro y h₁y h₂y
|
2024-10-31 17:09:09 +01:00
|
|
|
|
by_cases h₃y : y ∈ U
|
|
|
|
|
· simp [h₃y]
|
|
|
|
|
right
|
2024-11-04 13:22:12 +01:00
|
|
|
|
rw [MeromorphicAt.order_eq_top_iff (hf y h₃y)]
|
|
|
|
|
rw [eventually_nhdsWithin_iff]
|
2024-10-31 17:09:09 +01:00
|
|
|
|
rw [eventually_nhds_iff]
|
2024-11-04 13:22:12 +01:00
|
|
|
|
use N ∩ {z}ᶜ
|
|
|
|
|
constructor
|
|
|
|
|
· intro x h₁x _
|
|
|
|
|
exact h₁N x h₁x.1 h₁x.2
|
|
|
|
|
· constructor
|
|
|
|
|
· exact IsOpen.inter h₂N isOpen_compl_singleton
|
|
|
|
|
· exact Set.mem_inter h₁y h₂y
|
2024-10-31 17:09:09 +01:00
|
|
|
|
· simp [h₃y]
|
|
|
|
|
· tauto
|
|
|
|
|
|
2024-11-04 13:22:12 +01:00
|
|
|
|
· let A := (hf z hz).eventually_analyticAt
|
|
|
|
|
let B := Filter.eventually_and.2 ⟨h, A⟩
|
|
|
|
|
rw [eventually_nhdsWithin_iff, eventually_nhds_iff] at B
|
|
|
|
|
obtain ⟨N, h₁N, h₂N, h₃N⟩ := B
|
2024-10-31 17:09:09 +01:00
|
|
|
|
use N
|
|
|
|
|
constructor
|
|
|
|
|
· intro y h₁y h₂y
|
|
|
|
|
by_cases h₃y : y ∈ U
|
|
|
|
|
· simp [h₃y]
|
|
|
|
|
left
|
2024-11-04 13:22:12 +01:00
|
|
|
|
rw [(h₁N y h₁y h₂y).2.meromorphicAt_order]
|
|
|
|
|
let D := (h₁N y h₁y h₂y).2.order_eq_zero_iff.2
|
|
|
|
|
let C := (h₁N y h₁y h₂y).1
|
|
|
|
|
let E := D C
|
|
|
|
|
rw [E]
|
|
|
|
|
simp
|
2024-10-31 17:09:09 +01:00
|
|
|
|
· simp [h₃y]
|
|
|
|
|
· tauto
|
2024-11-19 10:07:20 +01:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
theorem MeromorphicOn.divisor_def₁
|
|
|
|
|
{f : ℂ → ℂ}
|
|
|
|
|
{U : Set ℂ}
|
|
|
|
|
{z : ℂ}
|
|
|
|
|
(hf : MeromorphicOn f U)
|
|
|
|
|
(hz : z ∈ U) :
|
|
|
|
|
hf.divisor z = ((hf z hz).order.untop' 0 : ℤ) := by
|
|
|
|
|
unfold MeromorphicOn.divisor
|
|
|
|
|
simp [hz]
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
theorem MeromorphicOn.divisor_def₂
|
|
|
|
|
{f : ℂ → ℂ}
|
|
|
|
|
{U : Set ℂ}
|
|
|
|
|
{z : ℂ}
|
|
|
|
|
(hf : MeromorphicOn f U)
|
|
|
|
|
(hz : z ∈ U)
|
|
|
|
|
(h₂f : (hf z hz).order ≠ ⊤) :
|
|
|
|
|
hf.divisor z = (hf z hz).order.untop h₂f := by
|
|
|
|
|
unfold MeromorphicOn.divisor
|
|
|
|
|
simp [hz]
|
|
|
|
|
rw [WithTop.untop'_eq_iff]
|
|
|
|
|
left
|
|
|
|
|
exact Eq.symm (WithTop.coe_untop (hf z hz).order h₂f)
|
2024-11-20 08:12:12 +01:00
|
|
|
|
|
|
|
|
|
|
2024-11-20 11:43:09 +01:00
|
|
|
|
theorem MeromorphicOn.divisor_mul₀
|
|
|
|
|
{f₁ f₂ : ℂ → ℂ}
|
|
|
|
|
{U : Set ℂ}
|
|
|
|
|
{z : ℂ}
|
|
|
|
|
(hz : z ∈ U)
|
|
|
|
|
(h₁f₁ : MeromorphicOn f₁ U)
|
|
|
|
|
(h₂f₁ : (h₁f₁ z hz).order ≠ ⊤)
|
|
|
|
|
(h₁f₂ : MeromorphicOn f₂ U)
|
|
|
|
|
(h₂f₂ : (h₁f₂ z hz).order ≠ ⊤) :
|
|
|
|
|
(h₁f₁.mul h₁f₂).divisor.toFun z = h₁f₁.divisor.toFun z + h₁f₂.divisor.toFun z := by
|
|
|
|
|
by_cases h₁z : z ∈ U
|
|
|
|
|
· rw [MeromorphicOn.divisor_def₂ h₁f₁ hz h₂f₁]
|
|
|
|
|
rw [MeromorphicOn.divisor_def₂ h₁f₂ hz h₂f₂]
|
|
|
|
|
have B : ((h₁f₁.mul h₁f₂) z hz).order ≠ ⊤ := by
|
|
|
|
|
rw [MeromorphicAt.order_mul (h₁f₁ z hz) (h₁f₂ z hz)]
|
|
|
|
|
simp; tauto
|
|
|
|
|
rw [MeromorphicOn.divisor_def₂ (h₁f₁.mul h₁f₂) hz B]
|
|
|
|
|
simp_rw [MeromorphicAt.order_mul (h₁f₁ z hz) (h₁f₂ z hz)]
|
|
|
|
|
rw [untop_add]
|
|
|
|
|
· unfold MeromorphicOn.divisor
|
|
|
|
|
simp [h₁z]
|
|
|
|
|
|
|
|
|
|
|
2024-11-20 08:12:12 +01:00
|
|
|
|
theorem MeromorphicOn.divisor_mul
|
|
|
|
|
{f₁ f₂ : ℂ → ℂ}
|
|
|
|
|
{U : Set ℂ}
|
|
|
|
|
(h₁f₁ : MeromorphicOn f₁ U)
|
|
|
|
|
(h₂f₁ : ∀ z, (hz : z ∈ U) → (h₁f₁ z hz).order ≠ ⊤)
|
|
|
|
|
(h₁f₂ : MeromorphicOn f₂ U)
|
|
|
|
|
(h₂f₂ : ∀ z, (hz : z ∈ U) → (h₁f₂ z hz).order ≠ ⊤) :
|
|
|
|
|
(h₁f₁.mul h₁f₂).divisor.toFun = h₁f₁.divisor.toFun + h₁f₂.divisor.toFun := by
|
|
|
|
|
funext z
|
|
|
|
|
by_cases hz : z ∈ U
|
2024-11-20 11:43:09 +01:00
|
|
|
|
· rw [MeromorphicOn.divisor_mul₀ hz h₁f₁ (h₂f₁ z hz) h₁f₂ (h₂f₂ z hz)]
|
|
|
|
|
simp
|
|
|
|
|
· simp
|
|
|
|
|
rw [Function.nmem_support.mp (fun a => hz (h₁f₁.divisor.supportInU a))]
|
|
|
|
|
rw [Function.nmem_support.mp (fun a => hz (h₁f₂.divisor.supportInU a))]
|
|
|
|
|
rw [Function.nmem_support.mp (fun a => hz ((h₁f₁.mul h₁f₂).divisor.supportInU a))]
|
|
|
|
|
simp
|