import Mathlib.Analysis.Meromorphic.Basic import Mathlib.Analysis.Meromorphic.Order open MeromorphicOn Metric Real Set Classical variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] /-- Derivatives of meromorphic functions are meromorphic. -/ @[fun_prop] theorem meromorphicAt_deriv {f : 𝕜 → 𝕜} {x : 𝕜} (h : MeromorphicAt f x) (h₁ : h.order ≠ ⊤) : MeromorphicAt (deriv f) x := by obtain ⟨g, h₁g, h₂g, h₃⟩ := h.order_ne_top_iff.1 h₁ lift h.order to ℤ using h₁ with n hn have : (n : WithTop ℤ).untop₀ = n := by sorry simp_all [this] have : deriv f =ᶠ[nhdsWithin x {x}ᶜ] deriv (fun z ↦ (z - x) ^ n * g z) := by sorry have : deriv f =ᶠ[nhdsWithin x {x}ᶜ] fun z ↦ n * (z - x) ^ (n - 1) * g z + (z - x) ^ n * deriv g z := by sorry apply MeromorphicAt.congr _ this.symm sorry