nevanlinna/Nevanlinna/meromorphicOn_divisor.lean

77 lines
2.0 KiB
Plaintext
Raw Normal View History

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