Working…
This commit is contained in:
parent
449de2e42a
commit
30ad49b90d
|
@ -1,6 +1,3 @@
|
|||
import Mathlib.Analysis.SpecialFunctions.Integrals
|
||||
import Mathlib.Analysis.SpecialFunctions.Log.NegMulLog
|
||||
import Mathlib.Analysis.Convex.SpecificFunctions.Deriv
|
||||
import Nevanlinna.analyticAt
|
||||
import Nevanlinna.divisor
|
||||
|
||||
|
@ -8,8 +5,6 @@ open scoped Interval Topology
|
|||
open Real Filter MeasureTheory intervalIntegral
|
||||
|
||||
|
||||
|
||||
|
||||
noncomputable def AnalyticOnNhd.zeroDivisor
|
||||
{f : ℂ → ℂ}
|
||||
{U : Set ℂ}
|
||||
|
@ -25,99 +20,40 @@ noncomputable def AnalyticOnNhd.zeroDivisor
|
|||
|
||||
supportInU := by
|
||||
intro z hz
|
||||
simp only [Function.mem_support] at hz
|
||||
simp only [Function.mem_support, ne_eq, dite_eq_else, Nat.cast_eq_zero, ENat.toNat_eq_zero, not_forall, not_or] at hz
|
||||
simp at hz
|
||||
by_contra h₂z
|
||||
simp [h₂z] at hz
|
||||
|
||||
locallyFiniteInU := by
|
||||
intro z hz
|
||||
|
||||
discreteSupport := by
|
||||
simp_rw [← singletons_open_iff_discrete]
|
||||
simp_rw [Metric.isOpen_singleton_iff]
|
||||
simp
|
||||
apply eventually_nhdsWithin_iff.2
|
||||
rw [eventually_nhds_iff]
|
||||
|
||||
-- simp only [dite_eq_ite, gt_iff_lt, Subtype.forall, Function.mem_support, ne_eq, ite_eq_else, Classical.not_imp, not_or, Subtype.mk.injEq]
|
||||
|
||||
intro a ha ⟨h₁a, h₂a⟩
|
||||
obtain ⟨g, h₁g, h₂g, h₃g⟩ := (AnalyticAt.order_neq_top_iff (hf a ha)).1 h₂a
|
||||
rw [Metric.eventually_nhds_iff_ball] at h₃g
|
||||
|
||||
have : ∃ ε > 0, ∀ y ∈ Metric.ball (↑a) ε, g y ≠ 0 := by
|
||||
have h₄g : ContinuousAt g a :=
|
||||
AnalyticAt.continuousAt h₁g
|
||||
have : {0}ᶜ ∈ nhds (g a) := by
|
||||
exact compl_singleton_mem_nhds_iff.mpr h₂g
|
||||
let F := h₄g.preimage_mem_nhds this
|
||||
rw [Metric.mem_nhds_iff] at F
|
||||
obtain ⟨ε, h₁ε, h₂ε⟩ := F
|
||||
use ε
|
||||
constructor; exact h₁ε
|
||||
intro y hy
|
||||
let G := h₂ε hy
|
||||
simp at G
|
||||
exact G
|
||||
|
||||
obtain ⟨ε₁, h₁ε₁⟩ := this
|
||||
obtain ⟨ε₂, h₁ε₂, h₂ε₂⟩ := h₃g
|
||||
use min ε₁ ε₂
|
||||
rcases AnalyticAt.eventually_eq_zero_or_eventually_ne_zero (hf z hz) with h|h
|
||||
· rw [eventually_nhds_iff] at h
|
||||
obtain ⟨N, h₁N, h₂N, h₃N⟩ := h
|
||||
use N
|
||||
constructor
|
||||
· have : 0 < min ε₁ ε₂ := by
|
||||
rw [lt_min_iff]
|
||||
exact And.imp_right (fun _ => h₁ε₂) h₁ε₁
|
||||
exact this
|
||||
· intro y h₁y _
|
||||
by_cases h₃y : y ∈ U
|
||||
· simp [h₃y]
|
||||
right
|
||||
rw [AnalyticAt.order_eq_top_iff (hf y h₃y)]
|
||||
rw [eventually_nhds_iff]
|
||||
use N
|
||||
· simp [h₃y]
|
||||
· tauto
|
||||
|
||||
intro y hy ⟨h₁y, h₂y⟩ h₃
|
||||
|
||||
have h'₂y : ↑y ∈ Metric.ball (↑a) ε₂ := by
|
||||
simp
|
||||
calc dist y a
|
||||
_ < min ε₁ ε₂ := by assumption
|
||||
_ ≤ ε₂ := by exact min_le_right ε₁ ε₂
|
||||
|
||||
have h₃y : ↑y ∈ Metric.ball (↑a) ε₁ := by
|
||||
simp
|
||||
calc dist y a
|
||||
_ < min ε₁ ε₂ := by assumption
|
||||
_ ≤ ε₁ := by exact min_le_left ε₁ ε₂
|
||||
|
||||
have F := h₂ε₂ y h'₂y
|
||||
have : f y = 0 := by
|
||||
rw [AnalyticAt.order_eq_zero_iff (hf y hy)] at h₁y
|
||||
tauto
|
||||
rw [this] at F
|
||||
simp at F
|
||||
|
||||
have : g y ≠ 0 := by
|
||||
exact h₁ε₁.2 y h₃y
|
||||
simp [this] at F
|
||||
rw [sub_eq_zero] at F
|
||||
tauto
|
||||
|
||||
|
||||
theorem AnalyticOnNhd.support_of_zeroDivisor
|
||||
{f : ℂ → ℂ}
|
||||
{U : Set ℂ}
|
||||
(hf : AnalyticOnNhd ℂ f U) :
|
||||
Function.support hf.zeroDivisor ⊆ U := by
|
||||
|
||||
intro z
|
||||
contrapose
|
||||
intro hz
|
||||
dsimp [AnalyticOnNhd.zeroDivisor]
|
||||
simp
|
||||
tauto
|
||||
|
||||
|
||||
theorem AnalyticOnNhd.support_of_zeroDivisor₂
|
||||
{f : ℂ → ℂ}
|
||||
{U : Set ℂ}
|
||||
(hf : AnalyticOnNhd ℂ f U) :
|
||||
Function.support hf.zeroDivisor ⊆ f⁻¹' {0} := by
|
||||
|
||||
intro z hz
|
||||
dsimp [AnalyticOnNhd.zeroDivisor] at hz
|
||||
have t₀ := hf.support_of_zeroDivisor hz
|
||||
simp [hf.support_of_zeroDivisor hz] at hz
|
||||
let A := hz.1
|
||||
let C := (hf z t₀).order_eq_zero_iff
|
||||
simp
|
||||
rw [C] at A
|
||||
tauto
|
||||
· rw [eventually_nhdsWithin_iff, eventually_nhds_iff] at h
|
||||
obtain ⟨N, h₁N, h₂N, h₃N⟩ := h
|
||||
use N
|
||||
constructor
|
||||
· intro y h₁y h₂y
|
||||
by_cases h₃y : y ∈ U
|
||||
· simp [h₃y]
|
||||
left
|
||||
rw [AnalyticAt.order_eq_zero_iff (hf y h₃y)]
|
||||
exact h₁N y h₁y h₂y
|
||||
· simp [h₃y]
|
||||
· tauto
|
||||
|
|
|
@ -12,7 +12,7 @@ structure Divisor
|
|||
where
|
||||
toFun : ℂ → ℤ
|
||||
supportInU : toFun.support ⊆ U
|
||||
locallyFiniteInU : ∀ x ∈ U, ∃ N ∈ 𝓝 x, (N ∩ toFun.support).Finite
|
||||
locallyFiniteInU : ∀ x ∈ U, toFun =ᶠ[𝓝[≠] x] 0
|
||||
|
||||
instance
|
||||
(U : Set ℂ) :
|
||||
|
|
Loading…
Reference in New Issue