import Mathlib.Analysis.Analytic.Meromorphic import Mathlib.Analysis.Calculus.ContDiff.Basic import Mathlib.Analysis.Calculus.FDeriv.Add import Nevanlinna.analyticAt variable {๐•œ : Type*} [NontriviallyNormedField ๐•œ] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ๐•œ E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ๐•œ F] variable {G : Type*} [NormedAddCommGroup G] [NormedSpace ๐•œ G] variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace ๐•œ G'] variable {f fโ‚€ fโ‚ g : E โ†’ F} variable {f' fโ‚€' fโ‚' g' : E โ†’L[๐•œ] F} variable (e : E โ†’L[๐•œ] F) variable {x : E} variable {s t : Set E} variable {L Lโ‚ Lโ‚‚ : Filter E} variable {R : Type*} [Semiring R] [Module R F] [SMulCommClass ๐•œ R F] [ContinuousConstSMul R F] -- import Mathlib.Analysis.Calculus.FDeriv.Add @[fun_prop] theorem Differentiable.const_smul' (h : Differentiable ๐•œ f) (c : R) : Differentiable ๐•œ (c โ€ข f) := by have : c โ€ข f = fun x โ†ฆ c โ€ข f x := rfl rw [this] exact Differentiable.const_smul h c -- Mathlib.Analysis.Calculus.ContDiff.Basic theorem ContDiff.const_smul' {f : E โ†’ F} (c : R) (hf : ContDiff ๐•œ n f) : ContDiff ๐•œ n (c โ€ข f) := by have : c โ€ข f = fun x โ†ฆ c โ€ข f x := rfl rw [this] exact ContDiff.const_smul c hf -- Mathlib.Analysis.Analytic.Meromorphic theorem meromorphicAt_congr {๐•œ : Type u_1} [NontriviallyNormedField ๐•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace ๐•œ E] {f : ๐•œ โ†’ E} {g : ๐•œ โ†’ E} {x : ๐•œ} (h : f =แถ [nhdsWithin x {x}แถœ] g) : MeromorphicAt f x โ†” MeromorphicAt g x := โŸจfun hf โ†ฆ hf.congr h, fun hg โ†ฆ hg.congr h.symmโŸฉ theorem meromorphicAt_congr' {๐•œ : Type u_1} [NontriviallyNormedField ๐•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace ๐•œ E] {f : ๐•œ โ†’ E} {g : ๐•œ โ†’ E} {x : ๐•œ} (h : f =แถ [nhds x] g) : MeromorphicAt f x โ†” MeromorphicAt g x := meromorphicAt_congr (Filter.EventuallyEq.filter_mono h nhdsWithin_le_nhds) open Topology Filter lemma Mnhds {ฮฑ : Type} {f g : โ„‚ โ†’ ฮฑ} {zโ‚€ : โ„‚} (hโ‚ : f =แถ [๐“[โ‰ ] zโ‚€] g) (hโ‚‚ : f zโ‚€ = g zโ‚€) : f =แถ [๐“ zโ‚€] g := by apply eventually_nhds_iff.2 obtain โŸจt, hโ‚t, hโ‚‚tโŸฉ := eventually_nhds_iff.1 (eventually_nhdsWithin_iff.1 hโ‚) use t constructor ยท intro y hy by_cases hโ‚‚y : y โˆˆ ({zโ‚€}แถœ : Set โ„‚) ยท exact hโ‚t y hy hโ‚‚y ยท simp at hโ‚‚y rwa [hโ‚‚y] ยท exact hโ‚‚t