Update stronglyMeromorphic_JensenFormula.lean

This commit is contained in:
Stefan Kebekus
2024-12-02 16:49:16 +01:00
parent 7fa50d47e9
commit 4320db0533

View File

@@ -29,34 +29,56 @@ theorem jensen_case_R_eq_one
have h'₂f : u : (Metric.closedBall (0 : ) 1), f u 0 := by have h'₂f : u : (Metric.closedBall (0 : ) 1), f u 0 := by
use 0, Metric.mem_closedBall_self (by simp) use 0, Metric.mem_closedBall_self (by simp)
have h₃f : Set.Finite (Function.support h₁f.meromorphicOn.divisor) := by
exact Divisor.finiteSupport h₂U (StronglyMeromorphicOn.meromorphicOn h₁f).divisor
have h₄f: Function.support (fun s (h₁f.meromorphicOn.divisor s) * log (s⁻¹)) h₃f.toFinset := by
intro x
contrapose
simp
intro hx
rw [hx]
simp
rw [finsum_eq_sum_of_support_subset _ h₄f]
obtain F, h₁F, h₂F, h₃F, h₄F := MeromorphicOn.decompose₃' h₂U h₁U h₁f h'₂f obtain F, h₁F, h₂F, h₃F, h₄F := MeromorphicOn.decompose₃' h₂U h₁U h₁f h'₂f
let G := fun z log F z + s (finiteZeros h₁U h₂U h₁f h'₂f).toFinset, (h₁f.order s).toNat * log z - s have h₁F : Function.mulSupport (fun u fun z => (z - u) ^ (h₁f.meromorphicOn.divisor u)) h₃f.toFinset := by
intro u
contrapose
simp
intro hu
rw [hu]
simp
exact rfl
rw [finprod_eq_prod_of_mulSupport_subset _ h₁F] at h₄F
let G := fun z log F z + s, (h₁f.meromorphicOn.divisor s) * log z - s
have h₁G {z : } : Function.support (fun s (h₁f.meromorphicOn.divisor s) * log z - s) h₃f.toFinset := by
intro s
contrapose
simp
intro hs
rw [hs]
simp
have decompose_f : z Metric.closedBall (0 : ) 1, f z 0 log f z = G z := by have decompose_f : z Metric.closedBall (0 : ) 1, f z 0 log f z = G z := by
intro z h₁z h₂z intro z h₁z h₂z
conv => rw [h₄F]
left simp only [Pi.mul_apply, norm_mul]
arg 1 simp only [Finset.prod_apply, norm_prod, norm_zpow]
rw [h₃F]
rw [smul_eq_mul]
rw [norm_mul]
rw [norm_prod]
left
arg 2
intro b
rw [norm_pow]
simp only [Complex.norm_eq_abs, Finset.univ_eq_attach]
rw [Real.log_mul] rw [Real.log_mul]
rw [Real.log_prod] rw [Real.log_prod]
conv => simp_rw [Real.log_zpow]
left dsimp only [G]
left rw [finsum_eq_sum_of_support_subset _ h₁G]
arg 2 --
intro s intro x hx
rw [Real.log_pow] simp at hx
dsimp [G]
abel abel
-- ∀ x ∈ ⋯.toFinset, Complex.abs (z - ↑x) ^ (h'₁f.order x).toNat ≠ 0 -- ∀ x ∈ ⋯.toFinset, Complex.abs (z - ↑x) ^ (h'₁f.order x).toNat ≠ 0