nevanlinna/Nevanlinna/stronglyMeromorphicOn.lean

178 lines
4.9 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
import Mathlib.Algebra.BigOperators.Finprod
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 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 [← StronglyMeromorphicAt.makeStronglyMeromorphic_id]
exact AnalyticAt.stronglyMeromorphicAt (h₂V v hv)
· simp [h₂v]
theorem stronglyMeromorphicOn_ratlPolynomial
{U : Set }
(d : ) :
StronglyMeromorphicOn (∏ᶠ u, fun z ↦ (z - u) ^ d u) U := by
by_cases hd : (Function.mulSupport fun u z => (z - u) ^ d u).Finite
· rw [finprod_eq_prod _ hd]
intro z h₁z
by_cases h₂z : d z = 0
· apply AnalyticAt.stronglyMeromorphicAt
rw [Finset.prod_fn]
apply Finset.analyticAt_prod
intro u hu
by_cases huz : u = z
· rw [← huz] at h₂z
rw [h₂z]
simp
exact analyticAt_const
· apply AnalyticAt.zpow
apply AnalyticAt.sub
apply analyticAt_id
apply analyticAt_const
rwa [sub_ne_zero, ne_comm]
· have : z ∈ hd.toFinset := by
simp
by_contra hCon
have A : 0 ≠ (1 : ) z := by simp
rw [← hCon] at A
simp only [sub_self] at A
rw [ne_comm] at A
rw [zpow_ne_zero_iff] at A
tauto
exact h₂z
rw [← Finset.mul_prod_erase hd.toFinset _ this]
right
use d z
use ∏ x ∈ hd.toFinset.erase z, fun z => (z - x) ^ d x
constructor
· rw [Finset.prod_fn]
apply Finset.analyticAt_prod
intro u hu
apply AnalyticAt.zpow
apply AnalyticAt.sub
apply analyticAt_id
apply analyticAt_const
rw [sub_ne_zero, ne_comm]
by_contra hCon
simp at hu
tauto
· constructor
· simp only [Finset.prod_apply]
rw [Finset.prod_ne_zero_iff]
intro u hu
rw [zpow_ne_zero_iff]
rw [sub_ne_zero]
by_contra hCon
rw [hCon] at hu
let A := Finset.not_mem_erase u hd.toFinset
tauto
--
have : u ∈ hd.toFinset := by
exact Finset.mem_of_mem_erase hu
simp at this
by_contra hCon
rw [hCon] at this
simp at this
tauto
· exact Filter.Eventually.of_forall (congrFun rfl)
· rw [finprod_of_infinite_mulSupport hd]
apply AnalyticOn.stronglyMeromorphicOn
apply analyticOnNhd_const
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₀)