import Mathlib.Analysis.Analytic.Meromorphic import Nevanlinna.analyticAt import Nevanlinna.divisor import Nevanlinna.meromorphicAt import Nevanlinna.meromorphicOn_divisor import Nevanlinna.stronglyMeromorphicOn import Nevanlinna.mathlibAddOn open scoped Interval Topology open Real Filter MeasureTheory intervalIntegral theorem MeromorphicOn.open_of_order_eq_top {f : ℂ → ℂ} {U : Set ℂ} (h₁f : MeromorphicOn f U) : IsOpen { u : U | (h₁f u.1 u.2).order = ⊤ } := by apply isOpen_iff_forall_mem_open.mpr intro z hz simp at hz have h₁z := hz rw [MeromorphicAt.order_eq_top_iff] at hz rw [eventually_nhdsWithin_iff] at hz rw [eventually_nhds_iff] at hz obtain ⟨t', h₁t', h₂t', h₃t'⟩ := hz let t : Set U := Subtype.val ⁻¹' t' use t constructor · intro w hw simp by_cases h₁w : w = z · rwa [h₁w] · -- rw [MeromorphicAt.order_eq_top_iff] rw [eventually_nhdsWithin_iff] rw [eventually_nhds_iff] use t' \ {z.1} constructor · intro y h₁y h₂y apply h₁t' exact Set.mem_of_mem_diff h₁y exact Set.mem_of_mem_inter_right h₁y · constructor · apply IsOpen.sdiff h₂t' exact isClosed_singleton · rw [Set.mem_diff] constructor · exact hw · simp exact Subtype.coe_ne_coe.mpr h₁w · constructor · exact isOpen_induced h₂t' · exact h₃t' theorem MeromorphicOn.open_of_order_neq_top {f : ℂ → ℂ} {U : Set ℂ} (h₁f : MeromorphicOn f U) : IsOpen { u : U | (h₁f u.1 u.2).order ≠ ⊤ } := by apply isOpen_iff_forall_mem_open.mpr intro z hz simp at hz let A := (h₁f z.1 z.2).eventually_eq_zero_or_eventually_ne_zero rcases A with h|h · rw [← (h₁f z.1 z.2).order_eq_top_iff] at h tauto · let A := (h₁f z.1 z.2).eventually_analyticAt let B := Filter.Eventually.and h A rw [eventually_nhdsWithin_iff] at B rw [eventually_nhds_iff] at B obtain ⟨t', h₁t', h₂t', h₃t'⟩ := B let t : Set U := Subtype.val ⁻¹' t' use t constructor · intro w hw simp by_cases h₁w : w = z · rwa [h₁w] · let B := h₁t' w hw simp at B have : (w : ℂ) ≠ (z : ℂ) := by exact Subtype.coe_ne_coe.mpr h₁w let C := B this let D := C.2.order_eq_zero_iff.2 C.1 rw [C.2.meromorphicAt_order, D] simp · constructor · exact isOpen_induced h₂t' · exact h₃t' theorem MeromorphicOn.order_ne_top {f : ℂ → ℂ} {U : Set ℂ} (h₁U : IsConnected U) (h₁f : MeromorphicOn f U) : (∃ z₀ : U, (h₁f z₀.1 z₀.2).order = ⊤) ↔ (∀ z : U, (h₁f z.1 z.2).order = ⊤) := by constructor · intro h obtain ⟨h₁z₀, h₂z₀⟩ := h intro hz sorry · intro h obtain ⟨w, hw⟩ := h₁U.nonempty use ⟨w, hw⟩ exact h ⟨w, hw⟩