Done with elimination

This commit is contained in:
Stefan Kebekus
2024-11-29 13:24:20 +01:00
parent eec4cd1ffa
commit 580ea61f96
3 changed files with 146 additions and 78 deletions

View File

@@ -45,6 +45,19 @@ theorem AnalyticOn.stronglyMeromorphicOn
exact h₁f z hz exact h₁f z hz
theorem stronglyMeromorphicOn_of_mul_analytic'
{f g : }
{U : Set }
(h₁g : AnalyticOnNhd g U)
(h₂g : u : U, g u 0)
(h₁f : StronglyMeromorphicOn f U) :
StronglyMeromorphicOn (g * f) U := by
intro z hz
rw [mul_comm]
apply (stronglyMeromorphicAt_of_mul_analytic (h₁g z hz) ?h₂g).mp (h₁f z hz)
exact h₂g z, hz
/- Make strongly MeromorphicOn -/ /- Make strongly MeromorphicOn -/
noncomputable def MeromorphicOn.makeStronglyMeromorphicOn noncomputable def MeromorphicOn.makeStronglyMeromorphicOn
{f : } {f : }

View File

@@ -293,62 +293,6 @@ theorem MeromorphicOn.decompose₂
simpa 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, fun z (z - u) ^ (h₁f.meromorphicOn.divisor u)) := by
have h₃f : u : U, (h₁f u u.2).meromorphicAt.order := MeromorphicOn.order_ne_top h₂U h₁f h₂f
have h₄f : Set.Finite (Function.support h₁f.meromorphicOn.divisor) := h₁f.meromorphicOn.divisor.finiteSupport h₁U
let P' : Set U := Subtype.val ⁻¹' Function.support h₁f.meromorphicOn.divisor
let P := (h₄f.preimage Set.injOn_subtype_val : Set.Finite P').toFinset
have hP : p P, (h₁f p p.2).meromorphicAt.order := by
intro p hp
apply h₃f
obtain g, h₁g, h₂g, h₃g, h₄g := MeromorphicOn.decompose₂ h₁f (P := P) hP
let h := p P, fun z => (z - p.1) ^ h₁f.meromorphicOn.divisor p.1
have h₅g : AnalyticOn g U := by
intro z hz
by_cases h₂z : z, hz P
· apply AnalyticAt.analyticWithinAt
exact h₂g z, hz, h₂z
· apply AnalyticAt.analyticWithinAt
rw [analyticAt_of_mul_analytic]
sorry
have h₂h : StronglyMeromorphicOn h U := by
sorry
have h₁h : MeromorphicOn h U := by
exact StronglyMeromorphicOn.meromorphicOn h₂h
have h₃h : h₁h.divisor = h₁f.meromorphicOn.divisor := by
sorry
use g
constructor
· exact h₁g
· constructor
· sorry
· constructor
· sorry
· conv =>
left
rw [h₄g]
congr
simp
sorry
theorem MeromorphicOn.decompose₃' theorem MeromorphicOn.decompose₃'
{f : } {f : }
@@ -438,4 +382,111 @@ theorem MeromorphicOn.decompose₃'
· exact AnalyticOnNhd.analyticOn h₃g · exact AnalyticOnNhd.analyticOn h₃g
· constructor · constructor
· exact h₄g · exact h₄g
· sorry · have t₀ : StronglyMeromorphicOn (g * (u : ), fun z => (z - u) ^ (h₁f.meromorphicOn.divisor u)) U := by
apply stronglyMeromorphicOn_of_mul_analytic' h₃g h₄g
apply stronglyMeromorphicOn_ratlPolynomial₃U
funext z
by_cases hz : z U
· apply Filter.EventuallyEq.eq_of_nhds
apply StronglyMeromorphicAt.localIdentity (h₁f z hz) (t₀ z hz)
have h₅g : g =[𝓝[] z] g' := makeStronglyMeromorphicOn_changeDiscrete h₁g' hz
have Y' : (g' * (u : ), fun z => (z - u) ^ (h₁f.meromorphicOn.divisor u)) =[𝓝[] z] g * (u : ), fun z => (z - u) ^ (h₁f.meromorphicOn.divisor u) := by
apply Filter.EventuallyEq.symm
exact Filter.EventuallyEq.mul h₅g (by simp)
apply Filter.EventuallyEq.trans _ Y'
unfold g'
unfold h₁
let A := h₁f.meromorphicOn.divisor.locallyFiniteInU z hz
let B := eventually_nhdsWithin_iff.1 A
obtain t, h₁t, h₂t, h₃t := eventually_nhds_iff.1 B
apply eventually_nhdsWithin_iff.2
rw [eventually_nhds_iff]
use t
constructor
· intro y h₁y h₂y
let C := h₁t y h₁y h₂y
rw [mul_assoc]
simp
have : (finprod (fun u z => (z - u) ^ d u) y * finprod (fun u z => (z - u) ^ (h₁f.meromorphicOn.divisor u)) y) = 1 := by
have t₀ : (Function.mulSupport fun u z => (z - u) ^ d u).Finite := by
rwa [ratlPoly_mulsupport, h₁d]
rw [finprod_eq_prod _ t₀]
have t₁ : (Function.mulSupport fun u z => (z - u) ^ h₁f.meromorphicOn.divisor u).Finite := by
rwa [ratlPoly_mulsupport]
rw [finprod_eq_prod _ t₁]
have : (Function.mulSupport fun u z => (z - u) ^ d u) = (Function.mulSupport fun u z => (z - u) ^ h₁f.meromorphicOn.divisor u) := by
rw [ratlPoly_mulsupport]
rw [ratlPoly_mulsupport]
unfold d
simp
have : t₀.toFinset = t₁.toFinset := by congr
rw [this]
simp
rw [ Finset.prod_mul_distrib]
apply Finset.prod_eq_one
intro x hx
apply zpow_neg_mul_zpow_self
have : y t₁.toFinset := by
simp
simp at C
rw [C]
simp
tauto
by_contra H
rw [sub_eq_zero] at H
rw [H] at this
tauto
rw [this]
simp
· exact h₂t, h₃t
· simp
have : g z = g' z := by
unfold g
unfold MeromorphicOn.makeStronglyMeromorphicOn
simp [hz]
rw [this]
unfold g'
unfold h₁
simp
rw [mul_assoc]
nth_rw 1 [ mul_one (f z)]
congr
have t₀ : (Function.mulSupport fun u z => (z - u) ^ d u).Finite := by
rwa [ratlPoly_mulsupport, h₁d]
rw [finprod_eq_prod _ t₀]
have t₁ : (Function.mulSupport fun u z => (z - u) ^ h₁f.meromorphicOn.divisor u).Finite := by
rwa [ratlPoly_mulsupport]
rw [finprod_eq_prod _ t₁]
have : (Function.mulSupport fun u z => (z - u) ^ d u) = (Function.mulSupport fun u z => (z - u) ^ h₁f.meromorphicOn.divisor u) := by
rw [ratlPoly_mulsupport]
rw [ratlPoly_mulsupport]
unfold d
simp
have : t₀.toFinset = t₁.toFinset := by congr
rw [this]
simp
rw [ Finset.prod_mul_distrib]
rw [eq_comm]
apply Finset.prod_eq_one
intro x hx
apply zpow_neg_mul_zpow_self
have : z t₁.toFinset := by
simp
have : h₁f.meromorphicOn.divisor z = 0 := by
let A := h₁f.meromorphicOn.divisor.supportInU
simp at A
by_contra H
let B := A z H
tauto
rw [this]
simp
rfl
by_contra H
rw [sub_eq_zero] at H
rw [H] at this
tauto

View File

@@ -71,31 +71,34 @@ theorem stronglyMeromorphicOn_ratlPolynomial₃U
exact stronglyMeromorphicOn_ratlPolynomial₃ d z (trivial) exact stronglyMeromorphicOn_ratlPolynomial₃ d z (trivial)
theorem ratlPoly_mulsupport
(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 := ne_zero_of_eq_one A
rw [zpow_ne_zero_iff h] at t₁
tauto
theorem stronglyMeromorphicOn_divisor_ratlPolynomial₁ theorem stronglyMeromorphicOn_divisor_ratlPolynomial₁
{z : } {z : }
(d : ) (d : )
(h₁d : Set.Finite d.support) : (h₁d : Set.Finite d.support) :
(((stronglyMeromorphicOn_ratlPolynomial₃ d).meromorphicOn) z trivial).order = d z := by (((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 := ne_zero_of_eq_one A
rw [zpow_ne_zero_iff h] at t₁
tauto
rw [MeromorphicAt.order_eq_int_iff] rw [MeromorphicAt.order_eq_int_iff]
use x h₁d.toFinset.erase z, fun z => (z - x) ^ d x use x h₁d.toFinset.erase z, fun z => (z - x) ^ d x
constructor constructor
@@ -114,11 +117,12 @@ theorem stronglyMeromorphicOn_divisor_ratlPolynomial₁
· apply Filter.Eventually.of_forall · apply Filter.Eventually.of_forall
intro x intro x
have t₀ : (Function.mulSupport fun u z => (z - u) ^ d u).Finite := by have t₀ : (Function.mulSupport fun u z => (z - u) ^ d u).Finite := by
rwa [h₂d] rwa [ratlPoly_mulsupport d]
rw [finprod_eq_prod _ t₀] rw [finprod_eq_prod _ t₀]
have t₁ : h₁d.toFinset = t₀.toFinset := by have t₁ : h₁d.toFinset = t₀.toFinset := by
simp simp
rwa [eq_comm] rw [eq_comm]
exact ratlPoly_mulsupport d
rw [t₁] rw [t₁]
simp simp
rw [eq_comm] rw [eq_comm]