nevanlinna/Nevanlinna/stronglyMeromorphicOn_ratlPolynomial.lean
2024-11-25 15:33:44 +01:00

168 lines
4.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.stronglyMeromorphicOn
import Nevanlinna.mathlibAddOn
open scoped Interval Topology
theorem analyticAt_ratlPolynomial₁
{z : }
(d : )
(P : Finset ) :
z ∉ P → AnalyticAt (∏ u ∈ P, fun z ↦ (z - u) ^ d u) z := by
intro hz
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]
exact ne_of_mem_of_not_mem hu hz
theorem stronglyMeromorphicOn_ratlPolynomial₂
(d : )
(P : Finset ) :
StronglyMeromorphicOn (∏ u ∈ P, fun z ↦ (z - u) ^ d u) := by
intro z hz
by_cases h₂z : z ∈ P
· rw [← Finset.mul_prod_erase P _ h₂z]
right
use d z
use ∏ x ∈ P.erase z, fun z => (z - x) ^ d x
constructor
· have : z ∉ P.erase z := Finset.not_mem_erase z P
apply analyticAt_ratlPolynomial₁ d (P.erase z) this
· constructor
· simp only [Finset.prod_apply]
rw [Finset.prod_ne_zero_iff]
intro u hu
apply zpow_ne_zero
rw [sub_ne_zero]
by_contra hCon
rw [hCon] at hu
let A := Finset.not_mem_erase u P
tauto
· exact Filter.Eventually.of_forall (congrFun rfl)
· apply AnalyticAt.stronglyMeromorphicAt
exact analyticAt_ratlPolynomial₁ d P (z := z) h₂z
theorem stronglyMeromorphicOn_ratlPolynomial₃
(d : ) :
StronglyMeromorphicOn (∏ᶠ u, fun z ↦ (z - u) ^ d u) := by
by_cases hd : (Function.mulSupport fun u z => (z - u) ^ d u).Finite
· rw [finprod_eq_prod _ hd]
apply stronglyMeromorphicOn_ratlPolynomial₂ d hd.toFinset
· rw [finprod_of_infinite_mulSupport hd]
apply AnalyticOn.stronglyMeromorphicOn
apply analyticOnNhd_const
theorem stronglyMeromorphicOn_divisor_ratlPolynomial₁
{z : }
(d : )
(h₁d : Set.Finite d.support) :
(((stronglyMeromorphicOn_ratlPolynomial₃ d).meromorphicOn) z trivial).order = d z := by
have h₂d : (Function.mulSupport fun u z ↦ (z - u) ^ d u) = d.support := by
ext u
constructor
· intro h
simp at h
simp
by_contra hCon
rw [hCon] at h
simp at h
tauto
· intro h
simp
by_contra hCon
let A := congrFun hCon u
simp at A
have t₁ : (0 : ) ^ d u ≠ 0 := by
exact ne_zero_of_eq_one A
rw [zpow_ne_zero_iff h] at t₁
tauto
rw [MeromorphicAt.order_eq_int_iff]
use ∏ x ∈ h₁d.toFinset.erase z, fun z => (z - x) ^ d x
constructor
· have : z ∉ h₁d.toFinset.erase z := Finset.not_mem_erase z h₁d.toFinset
apply analyticAt_ratlPolynomial₁ d (h₁d.toFinset.erase z) this
· constructor
· simp only [Finset.prod_apply]
rw [Finset.prod_ne_zero_iff]
intro u hu
apply zpow_ne_zero
rw [sub_ne_zero]
by_contra hCon
rw [hCon] at hu
let A := Finset.not_mem_erase u h₁d.toFinset
tauto
· apply Filter.Eventually.of_forall
intro x
simp
rw [← Finset.mul_prod_erase]
sorry
by_cases hz : z ∈ d.support
· sorry
· use (∏ᶠ u, fun z ↦ (z - u) ^ d u)
constructor
·
apply?
simp [hz]
sorry
· rw [Function.nmem_support.mp fun a => hz (h₂d a)]
rw [MeromorphicOn.divisor]
simp [hz]
theorem stronglyMeromorphicOn_divisor_ratlPolynomial
{U : Set }
(d : )
(h₁d : Set.Finite d.support)
(h₂d : d.support ⊆ U) :
(stronglyMeromorphicOn_ratlPolynomial₃ d (U := U)).meromorphicOn.divisor = d := by
funext z
by_cases hz : z ∈ U
· rw [MeromorphicOn.divisor]
simp [hz]
sorry
· rw [Function.nmem_support.mp fun a => hz (h₂d a)]
rw [MeromorphicOn.divisor]
simp [hz]
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₀)