import Mathlib.Analysis.Analytic.Meromorphic import Nevanlinna.analyticAt import Nevanlinna.divisor open scoped Interval Topology open Real Filter MeasureTheory intervalIntegral theorem MeromorphicAt.eventually_eq_zero_or_eventually_ne_zero {f : ℂ → ℂ} {z₀ : ℂ} (hf : MeromorphicAt f z₀) : (∀ᶠ (z : ℂ) in nhdsWithin z₀ {z₀}ᶜ, f z = 0) ∨ ∀ᶠ (z : ℂ) in nhdsWithin z₀ {z₀}ᶜ, f z ≠ 0 := by obtain ⟨n, h⟩ := hf let A := h.eventually_eq_zero_or_eventually_ne_zero rw [eventually_nhdsWithin_iff] rw [eventually_nhds_iff] rcases A with h₁|h₂ · rw [eventually_nhds_iff] at h₁ obtain ⟨N, h₁N, h₂N, h₃N⟩ := h₁ left use N constructor · intro y h₁y h₂y let A := h₁N y h₁y simp at A rcases A with h₃|h₄ · let B := h₃.1 simp at h₂y let C := sub_eq_zero.1 B tauto · assumption · constructor · exact h₂N · exact h₃N · right rw [eventually_nhdsWithin_iff] rw [eventually_nhds_iff] rw [eventually_nhdsWithin_iff] at h₂ rw [eventually_nhds_iff] at h₂ obtain ⟨N, h₁N, h₂N, h₃N⟩ := h₂ use N constructor · intro y h₁y h₂y by_contra h let A := h₁N y h₁y h₂y rw [h] at A simp at A · constructor · exact h₂N · exact h₃N 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] 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 obtain ⟨N, h₁N, h₂N, h₃N⟩ := h use N constructor · intro y h₁y h₂y by_cases h₃y : y ∈ U · simp [h₃y] right rw [MeromorphicAt.order_eq_top_iff (hf y h₃y)] rw [eventually_nhdsWithin_iff] rw [eventually_nhds_iff] 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 · simp [h₃y] · tauto · 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 use N constructor · intro y h₁y h₂y by_cases h₃y : y ∈ U · simp [h₃y] left 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 · simp [h₃y] · tauto