From 9410087ddf1b35fb5a77c363e41551fee9de1afc Mon Sep 17 00:00:00 2001 From: Stefan Kebekus Date: Fri, 16 Aug 2024 10:43:57 +0200 Subject: [PATCH] Update holomorphic_zero.lean --- Nevanlinna/holomorphic_zero.lean | 75 ++++++++++++++++++++++++++++++++ 1 file changed, 75 insertions(+) diff --git a/Nevanlinna/holomorphic_zero.lean b/Nevanlinna/holomorphic_zero.lean index 282eed3..8362bf5 100644 --- a/Nevanlinna/holomorphic_zero.lean +++ b/Nevanlinna/holomorphic_zero.lean @@ -12,9 +12,84 @@ noncomputable def zeroDivisor exact 0 +theorem analyticAtZeroDivisorSupport + {f : ℂ → ℂ} + {z : ℂ} + (h : z ∈ Function.support (zeroDivisor f)) : + AnalyticAt ℂ f z := by + by_contra h₁f + simp at h + dsimp [zeroDivisor] at h + simp [h₁f] at h + + +lemma toNatEqSelf_iff {n : ℕ∞} : n.toNat = n ↔ ∃ m : ℕ, m = n := by + constructor + · intro H₁ + rw [← ENat.some_eq_coe, ← WithTop.ne_top_iff_exists] + by_contra H₂ + rw [H₂] at H₁ + simp at H₁ + · intro H + obtain ⟨m, hm⟩ := H + rw [← hm] + simp + + theorem discreteZeros {f : ℂ → ℂ} : DiscreteTopology (Function.support (zeroDivisor f)) := by + apply singletons_open_iff_discrete.mp + intro z + + let A := analyticAtZeroDivisorSupport z.2 + let c : WithTop ℕ := A.order + let B := AnalyticAt.order_eq_nat_iff A + let n := zeroDivisor f z.1 + + + + have : ∃ a : ℕ, a = A.order := by + rw [← ENat.some_eq_coe] + rw [← WithTop.ne_top_iff_exists] + by_contra H + rw [AnalyticAt.order_eq_top_iff] at H + + + + dsimp [n, zeroDivisor] + simp [A] + + + + + + sorry + let C := (B n).1 this + + apply Metric.isOpen_singleton_iff.mpr + + + + /- + +Try this: refine Metric.isOpen_singleton_iff.mpr ?_ +Remaining subgoals: +⊢ ∃ ε > 0, ∀ (y : ↑(Function.support (zeroDivisor f))), dist y z < ε → y = z +Suggestions +Try this: refine isClosed_compl_iff.mp ?_ +Remaining subgoals: +⊢ IsClosed {z}ᶜ +Suggestions +Try this: refine disjoint_frontier_iff_isOpen.mp ?_ +Remaining subgoals: +⊢ Disjoint (frontier {z}) {z} +Suggestions +Try this: refine isOpen_iff_forall_mem_open.mpr ?_ +Remaining subgoals: +⊢ ∀ x ∈ {z}, ∃ t ⊆ {z}, IsOpen t ∧ x ∈ t + + -/ sorry