This commit is contained in:
Stefan Kebekus 2024-11-19 16:33:33 +01:00
parent 971b7cc23c
commit 25d0e2086a
2 changed files with 55 additions and 4 deletions

View File

@ -93,11 +93,14 @@ theorem MeromorphicOn.open_of_order_neq_top
· exact h₃t'
theorem MeromorphicOn.order_ne_top
theorem MeromorphicOn.clopen_of_order_eq_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
IsClopen { u : U | (h₁f u.1 u.2).order = } := by
sorry
constructor
· rw [← isOpen_compl_iff]
exact open_of_order_neq_top h₁f
· exact open_of_order_eq_top h₁f

View File

@ -2,6 +2,7 @@ import Mathlib.Analysis.Analytic.Meromorphic
import Nevanlinna.analyticAt
import Nevanlinna.divisor
import Nevanlinna.meromorphicAt
import Nevanlinna.meromorphicOn
import Nevanlinna.meromorphicOn_divisor
import Nevanlinna.stronglyMeromorphicOn
import Nevanlinna.mathlibAddOn
@ -290,3 +291,50 @@ theorem MeromorphicOn.decompose₂
rw [this]
--
simpa
theorem MeromorphicOn.decompose₃
{f : }
{U : Set }
(h₁U : IsCompact U)
(h₂U : IsConnected U)
(h₁f : StronglyMeromorphicOn f U)
(h₂f : ∃ u : U, f u ≠ 0) :
∃ g : , (MeromorphicOn g U)
∧ (AnalyticOn g U)
∧ (∀ u : U, g u ≠ 0)
∧ (f = g * ∏ᶠ u : U, fun z ↦ (z - u.1) ^ (h₁f.meromorphicOn.divisor u.1)) := by
have h₃f : ∀ u : U, (h₁f u u.2).meromorphicAt.order ≠ := by
let A := h₁f.meromorphicOn.clopen_of_order_eq_top
have : PreconnectedSpace U := by
apply isPreconnected_iff_preconnectedSpace.mp
exact IsConnected.isPreconnected h₂U
rw [isClopen_iff] at A
rcases A with h|h
· intro u
have : u ∉ (∅ : Set U) := by exact fun a => a
rw [← h] at this
simp at this
tauto
· obtain ⟨u, hu⟩ := h₂f
let A := (h₁f u u.2).order_eq_zero_iff.2 hu
have : u ∈ (Set.univ : Set U) := by trivial
rw [← h] at this
simp at this
rw [A] at this
tauto
have h₄f : Finite (Function.support h₁f.meromorphicOn.divisor) := by
exact h₁f.meromorphicOn.divisor.finiteSupport h₁U
let P' : Set U := Subtype.val ⁻¹' Function.support h₁f.meromorphicOn.divisor
have : Finite P' := by
unfold P'
refine Finite.of_injective ?f ?H
simp
apply Finite.of_injective
sorry
sorry