2024-10-09 12:13:22 +02:00
|
|
|
|
import Mathlib.Analysis.Analytic.Meromorphic
|
2024-07-31 09:40:35 +02:00
|
|
|
|
import Mathlib.Analysis.Calculus.ContDiff.Basic
|
|
|
|
|
import Mathlib.Analysis.Calculus.FDeriv.Add
|
|
|
|
|
|
|
|
|
|
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
|
2024-10-09 12:13:22 +02:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
-- 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)
|