nevanlinna/Nevanlinna/stronglyMeromorphicOn.lean

131 lines
3.6 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import Nevanlinna.stronglyMeromorphicAt
open Topology
/- Strongly MeromorphicOn -/
def StronglyMeromorphicOn
(f : )
(U : Set ) :=
∀ z ∈ U, StronglyMeromorphicAt f z
/- Strongly MeromorphicAt is Meromorphic -/
theorem StronglyMeromorphicOn.meromorphicOn
{f : }
{U : Set }
(hf : StronglyMeromorphicOn f U) :
MeromorphicOn f U := by
intro z hz
exact StronglyMeromorphicAt.meromorphicAt (hf z hz)
/- Strongly MeromorphicOn of non-negative order is analytic -/
theorem StronglyMeromorphicOn.analytic
{f : }
{U : Set }
(h₁f : StronglyMeromorphicOn f U)
(h₂f : ∀ x, (hx : x ∈ U) → 0 ≤ (h₁f x hx).meromorphicAt.order):
∀ z ∈ U, AnalyticAt f z := by
intro z hz
apply StronglyMeromorphicAt.analytic
exact h₂f z hz
exact h₁f z hz
/- Analytic functions are strongly meromorphic -/
theorem AnalyticOn.stronglyMeromorphicOn
{f : }
{U : Set }
(h₁f : AnalyticOnNhd f U) :
StronglyMeromorphicOn f U := by
intro z hz
apply AnalyticAt.stronglyMeromorphicAt
exact h₁f z hz
/- Strongly meromorphic functions on compact, preconnected sets are quotients of analytic functions -/
theorem StronglyMeromorphicOn_finite
{f : }
{U : Set }
(h₁U : IsCompact U)
(h₂U : IsPreconnected U)
(h₁f : StronglyMeromorphicOn f U)
(h₂f : ∃ z ∈ U, f z ≠ 0) :
Set.Finite {z ∈ U | f z = 0} := by
sorry
/- Strongly meromorphic functions on compact, preconnected sets are quotients of analytic functions -/
theorem StronglyMeromorphicOn_quotient
{f : }
{U : Set }
(h₁U : IsCompact U)
(h₂U : IsPreconnected U)
(h₁f : StronglyMeromorphicOn f U)
(h₂f : ∃ z ∈ U, f z ≠ 0) :
∃ a b : , (AnalyticOnNhd a U) ∧ (AnalyticOnNhd b U) ∧ (∀ z ∈ U, a z ≠ 0 b z ≠ 0) ∧ f = a / b := by
sorry
/- Make strongly MeromorphicAt -/
noncomputable def MeromorphicOn.makeStronglyMeromorphicOn
{f : }
{U : Set }
(hf : MeromorphicOn f U) :
:= by
intro z
by_cases hz : z ∈ U
· exact (hf z hz).makeStronglyMeromorphicAt z
· exact f z
theorem makeStronglyMeromorphicOn_changeDiscrete
{f : }
{U : Set }
{z₀ : }
(hf : MeromorphicOn f U)
(hz₀ : z₀ ∈ U) :
hf.makeStronglyMeromorphicOn =ᶠ[𝓝[≠] z₀] f := by
apply Filter.eventually_iff_exists_mem.2
let A := (hf z₀ hz₀).eventually_analyticAt
obtain ⟨V, h₁V, h₂V⟩ := Filter.eventually_iff_exists_mem.1 A
use V
constructor
· assumption
· intro v hv
unfold MeromorphicOn.makeStronglyMeromorphicOn
by_cases h₂v : v ∈ U
· simp [h₂v]
rw [← makeStronglyMeromorphic_id]
exact AnalyticAt.stronglyMeromorphicAt (h₂V v hv)
· simp [h₂v]
theorem makeStronglyMeromorphicOn_changeDiscrete'
{f : }
{U : Set }
{z₀ : }
(hf : MeromorphicOn f U)
(hz₀ : z₀ ∈ U) :
hf.makeStronglyMeromorphicOn =ᶠ[𝓝 z₀] (hf z₀ hz₀).makeStronglyMeromorphicAt := by
apply Mnhds
let A := makeStronglyMeromorphicOn_changeDiscrete hf hz₀
apply Filter.EventuallyEq.trans A
exact m₂ (hf z₀ hz₀)
unfold MeromorphicOn.makeStronglyMeromorphicOn
simp [hz₀]
theorem StronglyMeromorphicOn_of_makeStronglyMeromorphicOn
{f : }
{U : Set }
(hf : MeromorphicOn f U) :
StronglyMeromorphicOn hf.makeStronglyMeromorphicOn U := by
intro z₀ hz₀
rw [stronglyMeromorphicAt_congr (makeStronglyMeromorphicOn_changeDiscrete' hf hz₀)]
exact StronglyMeromorphicAt_of_makeStronglyMeromorphic (hf z₀ hz₀)