Working…

This commit is contained in:
Stefan Kebekus 2024-09-10 16:50:29 +02:00
parent 745e614016
commit e901f241cc
2 changed files with 34 additions and 38 deletions

View File

@ -357,8 +357,6 @@ theorem AnalyticOnCompact.eliminateZeros
· exact inter · exact inter
theorem AnalyticOnCompact.eliminateZeros₂ theorem AnalyticOnCompact.eliminateZeros₂
{f : } {f : }
{U : Set } {U : Set }

View File

@ -81,18 +81,14 @@ theorem jensen_case_R_eq_one
(h₂f : f 0 ≠ 0) : (h₂f : f 0 ≠ 0) :
log ‖f 0‖ = -∑ᶠ s, (h'₁f.order s).toNat * log (‖s.1‖⁻¹) + (2 * π)⁻¹ * ∫ (x : ) in (0)..(2 * π), log ‖f (circleMap 0 1 x)‖ := by log ‖f 0‖ = -∑ᶠ s, (h'₁f.order s).toNat * log (‖s.1‖⁻¹) + (2 * π)⁻¹ * ∫ (x : ) in (0)..(2 * π), log ‖f (circleMap 0 1 x)‖ := by
have h₁U : IsPreconnected (Metric.closedBall (0 : ) 1) := by have h₁U : IsPreconnected (Metric.closedBall (0 : ) 1) :=
apply IsConnected.isPreconnected (convex_closedBall (0 : ) 1).isPreconnected
apply Convex.isConnected
exact convex_closedBall 0 1
exact Set.nonempty_of_nonempty_subtype
have h₂U : IsCompact (Metric.closedBall (0 : ) 1) := isCompact_closedBall 0 1 have h₂U : IsCompact (Metric.closedBall (0 : ) 1) :=
isCompact_closedBall 0 1
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 use 0; simp; exact h₂f
simp
exact h₂f
obtain ⟨F, h₁F, h₂F, h₃F⟩ := AnalyticOnCompact.eliminateZeros₂ h₁U h₂U h'₁f h'₂f obtain ⟨F, h₁F, h₂F, h₃F⟩ := AnalyticOnCompact.eliminateZeros₂ h₁U h₂U h'₁f h'₂f
@ -127,24 +123,31 @@ theorem jensen_case_R_eq_one
dsimp [G] dsimp [G]
abel abel
-- ∀ x ∈ (ZeroFinset h₁f).attach, Complex.abs (z - ↑x) ^ order x ≠ 0 -- ∀ x ∈ ⋯.toFinset, Complex.abs (z - ↑x) ^ (h'₁f.order x).toNat ≠ 0
have : ∀ x ∈ (finiteZeros h₁U h₂U h'₁f h'₂f).toFinset, Complex.abs (z - ↑x) ^ (h'₁f.order x).toNat ≠ 0 := by
simp
intro s hs intro s hs
rw [ZeroFinset_mem_iff h₁f s] at hs simp at hs
rw [← hs.2] at h₂z simp
intro h₂s
rw [h₂s] at h₂z
tauto tauto
exact this
-- ∏ x ∈ ⋯.toFinset, Complex.abs (z - ↑x) ^ (h'₁f.order x).toNat ≠ 0
have : ∀ x ∈ (finiteZeros h₁U h₂U h'₁f h'₂f).toFinset, Complex.abs (z - ↑x) ^ (h'₁f.order x).toNat ≠ 0 := by
intro s hs
simp at hs
simp
intro h₂s
rw [h₂s] at h₂z
tauto
rw [Finset.prod_ne_zero_iff]
exact this
-- Complex.abs (F z) ≠ 0 -- Complex.abs (F z) ≠ 0
simp simp
exact h₂F z h₁z exact h₂F z h₁z
-- ∏ I : { x // x ∈ S }, Complex.abs (z - a I) ≠ 0
by_contra C
obtain ⟨s, h₁s, h₂s⟩ := Finset.prod_eq_zero_iff.1 C
simp at h₂s
rw [← ((ZeroFinset_mem_iff h₁f s).1 (Finset.coe_mem s)).2, h₂s.1] at h₂z
tauto
have : ∫ (x : ) in (0)..2 * π, log ‖f (circleMap 0 1 x)‖ = ∫ (x : ) in (0)..2 * π, G (circleMap 0 1 x) := by have : ∫ (x : ) in (0)..2 * π, log ‖f (circleMap 0 1 x)‖ = ∫ (x : ) in (0)..2 * π, G (circleMap 0 1 x) := by
@ -159,29 +162,24 @@ theorem jensen_case_R_eq_one
simp at ha simp at ha
simp simp
by_contra C by_contra C
have : (circleMap 0 1 a) ∈ Metric.closedBall 0 1 := by have : (circleMap 0 1 a) ∈ Metric.closedBall 0 1 :=
sorry circleMap_mem_closedBall 0 (zero_le_one' ) a
exact ha.2 (decompose_f (circleMap 0 1 a) this C) exact ha.2 (decompose_f (circleMap 0 1 a) this C)
apply Set.Countable.mono t₀ apply Set.Countable.mono t₀
apply Set.Countable.preimage_circleMap apply Set.Countable.preimage_circleMap
apply Set.Finite.countable apply Set.Finite.countable
apply finiteZeros let A := finiteZeros h₁U h₂U h'₁f h'₂f
-- IsPreconnected (Metric.closedBall (0 : ) 1) have : (Metric.closedBall 0 1 ∩ f ⁻¹' {0}) = (Metric.closedBall 0 1).restrict f ⁻¹' {0} := by
ext z
apply IsConnected.isPreconnected simp
apply Convex.isConnected tauto
exact convex_closedBall 0 1 rw [this]
exact Set.nonempty_of_nonempty_subtype exact Set.Finite.image Subtype.val A
--
exact isCompact_closedBall 0 1
--
exact h'₁f
use 0
exact ⟨Metric.mem_closedBall_self (zero_le_one' ), h₂f⟩
exact Ne.symm (zero_ne_one' ) exact Ne.symm (zero_ne_one' )
have h₁Gi : ∀ i ∈ (ZeroFinset h₁f h₂f).attach, IntervalIntegrable (fun x => log (Complex.abs (circleMap 0 1 x - ↑i))) MeasureTheory.volume 0 (2 * π) := by have h₁Gi : ∀ i ∈ (ZeroFinset h₁f h₂f).attach, IntervalIntegrable (fun x => log (Complex.abs (circleMap 0 1 x - ↑i))) MeasureTheory.volume 0 (2 * π) := by
-- This is hard. Need to invoke specialFunctions_CircleIntegral_affine. -- This is hard. Need to invoke specialFunctions_CircleIntegral_affine.
sorry sorry