Working…

This commit is contained in:
Stefan Kebekus 2024-11-13 14:31:45 +01:00
parent 15fa18c52f
commit 74ba95926e
3 changed files with 25 additions and 6 deletions

View File

@ -118,10 +118,6 @@ theorem AnalyticAt.supp_order_toNat
intro h₁f intro h₁f
simp [hf.order_eq_zero_iff.2 h₁f] simp [hf.order_eq_zero_iff.2 h₁f]
/-
theorem ContinuousLinearEquiv.analyticAt
( : ≃L[] ) (z₀ : ) : AnalyticAt (⇑ℓ) z₀ := .toContinuousLinearMap.analyticAt z₀
-/
theorem eventually_nhds_comp_composition theorem eventually_nhds_comp_composition
{f₁ f₂ : } {f₁ f₂ : }
@ -260,3 +256,13 @@ theorem AnalyticAt.localIdentity
let _ := Filter.Eventually.exists A let _ := Filter.Eventually.exists A
tauto tauto
exact Filter.eventuallyEq_iff_sub.mpr this exact Filter.eventuallyEq_iff_sub.mpr this
theorem AnalyticAt.mul₁
{f g : }
{z : }
(hf : AnalyticAt f z)
(hg : AnalyticAt g z) :
AnalyticAt (f * g) z := by
rw [(by rfl : f * g = (fun x ↦ f x * g x))]
exact mul hf hg

View File

@ -20,7 +20,7 @@ theorem MeromorphicOn.decompose
(h₂f : ∃ z₀ ∈ U, f z₀ ≠ 0) : (h₂f : ∃ z₀ ∈ U, f z₀ ≠ 0) :
∃ g : , (AnalyticOnNhd g U) ∃ g : , (AnalyticOnNhd g U)
∧ (∀ z ∈ U, g z ≠ 0) ∧ (∀ z ∈ U, g z ≠ 0)
∧ (Set.EqOn h₁f.makeStronglyMeromorphicOn (fun z ↦ ∏ᶠ p, (z - p) ^ (h₁f.divisor p) * g z ) U) := by ∧ (Set.EqOn h₁f.makeStronglyMeromorphicOn ((∏ᶠ p, fun z ↦ (z - p) ^ (h₁f.divisor p)) * g) U) := by
let g₁ : := f * (fun z ↦ ∏ᶠ p, (z - p) ^ (h₁f.divisor p)) let g₁ : := f * (fun z ↦ ∏ᶠ p, (z - p) ^ (h₁f.divisor p))
have h₁g₁ : MeromorphicOn g₁ U := by have h₁g₁ : MeromorphicOn g₁ U := by
@ -65,5 +65,18 @@ theorem MeromorphicOn.decompose
sorry sorry
apply Filter.EventuallyEq.eq_of_nhds apply Filter.EventuallyEq.eq_of_nhds
apply StronglyMeromorphicAt.localIdentity apply StronglyMeromorphicAt.localIdentity
· exact StronglyMeromorphicOn_of_makeStronglyMeromorphic h₁f z hz
· right
use h₁f.divisor z
use (∏ᶠ p ≠ z, (fun x ↦ (x - p) ^ h₁f.divisor p)) * g
constructor
· apply AnalyticAt.mul₁
·
sorry
· apply (h₃g z hz).analytic
rw [h₂g ⟨z, hz⟩]
· constructor
· sorry
· sorry
sorry sorry

View File

@ -81,9 +81,9 @@ theorem makeStronglyMeromorphicOn_changeDiscrete
theorem StronglyMeromorphicOn_of_makeStronglyMeromorphic theorem StronglyMeromorphicOn_of_makeStronglyMeromorphic
{f : } {f : }
{z₀ : }
(hf : MeromorphicOn f U) : (hf : MeromorphicOn f U) :
StronglyMeromorphicOn hf.makeStronglyMeromorphicOn U := by StronglyMeromorphicOn hf.makeStronglyMeromorphicOn U := by
sorry sorry