236 lines
7.2 KiB
Plaintext
236 lines
7.2 KiB
Plaintext
import Mathlib.Analysis.Analytic.Constructions
|
||
import Mathlib.Analysis.Analytic.IsolatedZeros
|
||
import Mathlib.Analysis.Complex.Basic
|
||
|
||
|
||
theorem AnalyticOn.order_eq_nat_iff
|
||
{f : ℂ → ℂ}
|
||
{U : Set ℂ}
|
||
{z₀ : ℂ}
|
||
(hf : AnalyticOn ℂ f U)
|
||
(hz₀ : z₀ ∈ U)
|
||
(n : ℕ) :
|
||
(hf z₀ hz₀).order = ↑n ↔ ∃ (g : ℂ → ℂ), AnalyticOn ℂ g U ∧ g z₀ ≠ 0 ∧ ∀ z, f z = (z - z₀) ^ n • g z := by
|
||
|
||
constructor
|
||
-- Direction →
|
||
intro hn
|
||
obtain ⟨gloc, h₁gloc, h₂gloc, h₃gloc⟩ := (AnalyticAt.order_eq_nat_iff (hf z₀ hz₀) n).1 hn
|
||
|
||
-- Define a candidate function; this is (f z) / (z - z₀) ^ n with the
|
||
-- removable singularity removed
|
||
let g : ℂ → ℂ := fun z ↦ if z = z₀ then gloc z₀ else (f z) / (z - z₀) ^ n
|
||
|
||
-- Describe g near z₀
|
||
have g_near_z₀ : ∀ᶠ (z : ℂ) in nhds z₀, g z = gloc z := by
|
||
rw [eventually_nhds_iff]
|
||
obtain ⟨t, h₁t, h₂t, h₃t⟩ := eventually_nhds_iff.1 h₃gloc
|
||
use t
|
||
constructor
|
||
· intro y h₁y
|
||
by_cases h₂y : y = z₀
|
||
· dsimp [g]; simp [h₂y]
|
||
· dsimp [g]; simp [h₂y]
|
||
rw [div_eq_iff_mul_eq, eq_comm, mul_comm]
|
||
exact h₁t y h₁y
|
||
norm_num
|
||
rw [sub_eq_zero]
|
||
tauto
|
||
· constructor
|
||
· assumption
|
||
· assumption
|
||
|
||
-- Describe g near points z₁ that are different from z₀
|
||
have g_near_z₁ {z₁ : ℂ} : z₁ ≠ z₀ → ∀ᶠ (z : ℂ) in nhds z₁, g z = f z / (z - z₀) ^ n := by
|
||
intro hz₁
|
||
rw [eventually_nhds_iff]
|
||
use {z₀}ᶜ
|
||
constructor
|
||
· intro y hy
|
||
simp at hy
|
||
simp [g, hy]
|
||
· exact ⟨isOpen_compl_singleton, hz₁⟩
|
||
|
||
-- Use g and show that it has all required properties
|
||
use g
|
||
constructor
|
||
· -- AnalyticOn ℂ g U
|
||
intro z h₁z
|
||
by_cases h₂z : z = z₀
|
||
· rw [h₂z]
|
||
apply AnalyticAt.congr h₁gloc
|
||
exact Filter.EventuallyEq.symm g_near_z₀
|
||
· simp_rw [eq_comm] at g_near_z₁
|
||
apply AnalyticAt.congr _ (g_near_z₁ h₂z)
|
||
apply AnalyticAt.div
|
||
exact hf z h₁z
|
||
apply AnalyticAt.pow
|
||
apply AnalyticAt.sub
|
||
apply analyticAt_id
|
||
apply analyticAt_const
|
||
simp
|
||
rw [sub_eq_zero]
|
||
tauto
|
||
· constructor
|
||
· simp [g]; tauto
|
||
· intro z
|
||
by_cases h₂z : z = z₀
|
||
· rw [h₂z, g_near_z₀.self_of_nhds]
|
||
exact h₃gloc.self_of_nhds
|
||
· rw [(g_near_z₁ h₂z).self_of_nhds]
|
||
simp [h₂z]
|
||
rw [div_eq_mul_inv, mul_comm, mul_assoc, inv_mul_cancel]
|
||
simp; norm_num
|
||
rw [sub_eq_zero]
|
||
tauto
|
||
|
||
-- direction ←
|
||
intro h
|
||
obtain ⟨g, h₁g, h₂g, h₃g⟩ := h
|
||
rw [AnalyticAt.order_eq_nat_iff]
|
||
use g
|
||
exact ⟨h₁g z₀ hz₀, ⟨h₂g, Filter.eventually_of_forall h₃g⟩⟩
|
||
|
||
|
||
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 AnalyticOn.eliminateZeros
|
||
{f : ℂ → ℂ}
|
||
{U : Set ℂ}
|
||
{A : Finset U}
|
||
(hf : AnalyticOn ℂ f U)
|
||
(n : ℂ → ℕ) :
|
||
(∀ a ∈ A, (hf a.1 a.2).order = n a) → ∃ (g : ℂ → ℂ), AnalyticOn ℂ g U ∧ (∀ a ∈ A, g a ≠ 0) ∧ ∀ z, f z = (∏ a ∈ A, (z - a) ^ (n a)) • g z := by
|
||
|
||
apply Finset.induction (α := U) (p := fun A ↦ (∀ a ∈ A, (hf a.1 a.2).order = n a) → ∃ (g : ℂ → ℂ), AnalyticOn ℂ g U ∧ (∀ a ∈ A, g a ≠ 0) ∧ ∀ z, f z = (∏ a ∈ A, (z - a) ^ (n a)) • g z)
|
||
|
||
-- case empty
|
||
simp
|
||
use f
|
||
simp
|
||
exact hf
|
||
|
||
-- case insert
|
||
intro b₀ B hb iHyp
|
||
intro hBinsert
|
||
obtain ⟨g₀, h₁g₀, h₂g₀, h₃g₀⟩ := iHyp (fun a ha ↦ hBinsert a (Finset.mem_insert_of_mem ha))
|
||
|
||
have : (h₁g₀ b₀ b₀.2).order = n b₀ := by
|
||
|
||
rw [← hBinsert b₀ (Finset.mem_insert_self b₀ B)]
|
||
|
||
let φ := fun z ↦ (∏ a ∈ B, (z - a.1) ^ n a.1)
|
||
|
||
have : f = fun z ↦ φ z * g₀ z := by
|
||
funext z
|
||
rw [h₃g₀ z]
|
||
rfl
|
||
simp_rw [this]
|
||
|
||
have h₁φ : AnalyticAt ℂ φ b₀ := by
|
||
dsimp [φ]
|
||
apply Finset.analyticAt_prod
|
||
intro b _
|
||
apply AnalyticAt.pow
|
||
apply AnalyticAt.sub
|
||
apply analyticAt_id ℂ
|
||
exact analyticAt_const
|
||
|
||
have h₂φ : h₁φ.order = (0 : ℕ) := by
|
||
rw [AnalyticAt.order_eq_nat_iff h₁φ 0]
|
||
use φ
|
||
constructor
|
||
· assumption
|
||
· constructor
|
||
· dsimp [φ]
|
||
push_neg
|
||
rw [Finset.prod_ne_zero_iff]
|
||
intro a ha
|
||
simp
|
||
have : ¬ (b₀.1 - a.1 = 0) := by
|
||
by_contra C
|
||
rw [sub_eq_zero] at C
|
||
rw [SetCoe.ext C] at hb
|
||
tauto
|
||
tauto
|
||
· simp
|
||
|
||
rw [AnalyticAt.order_mul h₁φ (h₁g₀ b₀ b₀.2)]
|
||
|
||
rw [h₂φ]
|
||
simp
|
||
|
||
|
||
obtain ⟨g₁, h₁g₁, h₂g₁, h₃g₁⟩ := (AnalyticOn.order_eq_nat_iff h₁g₀ b₀.2 (n b₀)).1 this
|
||
|
||
use g₁
|
||
constructor
|
||
· exact h₁g₁
|
||
· constructor
|
||
· intro a h₁a
|
||
by_cases h₂a : a = b₀
|
||
· rwa [h₂a]
|
||
· let A' := Finset.mem_of_mem_insert_of_ne h₁a h₂a
|
||
let B' := h₃g₁ a
|
||
let C' := h₂g₀ a A'
|
||
rw [B'] at C'
|
||
exact right_ne_zero_of_smul C'
|
||
· intro z
|
||
let A' := h₃g₀ z
|
||
rw [h₃g₁ z] at A'
|
||
rw [A']
|
||
rw [← smul_assoc]
|
||
congr
|
||
simp
|
||
rw [Finset.prod_insert]
|
||
ring
|
||
exact hb
|