77 lines
2.6 KiB
Plaintext
77 lines
2.6 KiB
Plaintext
import Mathlib.Analysis.Analytic.IsolatedZeros
|
||
import Mathlib.Analysis.Complex.Basic
|
||
|
||
|
||
theorem AnalyticAt.order_mul
|
||
{f₁ f₂ : ℂ → ℂ}
|
||
{z₀ : ℂ}
|
||
(hf₁ : AnalyticAt ℂ f₁ z₀)
|
||
(hf₂ : AnalyticAt ℂ f₂ z₀) :
|
||
(AnalyticAt.mul hf₁ hf₂).order = hf₁.order + hf₂.order := by
|
||
by_cases h₂f₁ : hf₁.order = ⊤
|
||
· simp [h₂f₁]
|
||
rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff]
|
||
rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff] at h₂f₁
|
||
obtain ⟨t, h₁t, h₂t, h₃t⟩ := h₂f₁
|
||
use t
|
||
constructor
|
||
· intro y hy
|
||
rw [h₁t y hy]
|
||
ring
|
||
· exact ⟨h₂t, h₃t⟩
|
||
· by_cases h₂f₂ : hf₂.order = ⊤
|
||
· simp [h₂f₂]
|
||
rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff]
|
||
rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff] at h₂f₂
|
||
obtain ⟨t, h₁t, h₂t, h₃t⟩ := h₂f₂
|
||
use t
|
||
constructor
|
||
· intro y hy
|
||
rw [h₁t y hy]
|
||
ring
|
||
· exact ⟨h₂t, h₃t⟩
|
||
· obtain ⟨g₁, h₁g₁, h₂g₁, h₃g₁⟩ := (AnalyticAt.order_eq_nat_iff hf₁ ↑hf₁.order.toNat).1 (eq_comm.1 (ENat.coe_toNat h₂f₁))
|
||
obtain ⟨g₂, h₁g₂, h₂g₂, h₃g₂⟩ := (AnalyticAt.order_eq_nat_iff hf₂ ↑hf₂.order.toNat).1 (eq_comm.1 (ENat.coe_toNat h₂f₂))
|
||
rw [← ENat.coe_toNat h₂f₁, ← ENat.coe_toNat h₂f₂, ← ENat.coe_add]
|
||
rw [AnalyticAt.order_eq_nat_iff (AnalyticAt.mul hf₁ hf₂) ↑(hf₁.order.toNat + hf₂.order.toNat)]
|
||
use g₁ * g₂
|
||
constructor
|
||
· exact AnalyticAt.mul h₁g₁ h₁g₂
|
||
· constructor
|
||
· simp; tauto
|
||
· obtain ⟨t₁, h₁t₁, h₂t₁, h₃t₁⟩ := eventually_nhds_iff.1 h₃g₁
|
||
obtain ⟨t₂, h₁t₂, h₂t₂, h₃t₂⟩ := eventually_nhds_iff.1 h₃g₂
|
||
rw [eventually_nhds_iff]
|
||
use t₁ ∩ t₂
|
||
constructor
|
||
· intro y hy
|
||
rw [h₁t₁ y hy.1, h₁t₂ y hy.2]
|
||
simp; ring
|
||
· constructor
|
||
· exact IsOpen.inter h₂t₁ h₂t₂
|
||
· exact Set.mem_inter h₃t₁ h₃t₂
|
||
|
||
|
||
theorem AnalyticAt.order_eq_zero_iff
|
||
{f : ℂ → ℂ}
|
||
{z₀ : ℂ}
|
||
(hf : AnalyticAt ℂ f z₀) :
|
||
hf.order = 0 ↔ f z₀ ≠ 0 := by
|
||
|
||
have : (0 : ENat) = (0 : Nat) := by rfl
|
||
rw [this, AnalyticAt.order_eq_nat_iff hf 0]
|
||
|
||
constructor
|
||
· intro hz
|
||
obtain ⟨g, _, h₂g, h₃g⟩ := hz
|
||
simp at h₃g
|
||
rw [Filter.Eventually.self_of_nhds h₃g]
|
||
tauto
|
||
· intro hz
|
||
use f
|
||
constructor
|
||
· exact hf
|
||
· constructor
|
||
· exact hz
|
||
· simp
|