nevanlinna/Nevanlinna/stronglyMeromorphicOn.lean

66 lines
1.5 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
/- 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 StronglyMeromorphicOn_of_makeStronglyMeromorphicOn
{f : }
{U : Set }
(hf : MeromorphicOn f U) :
StronglyMeromorphicOn hf.makeStronglyMeromorphicOn U := by
sorry