Update complexHarmonic.examples.lean

This commit is contained in:
Stefan Kebekus 2024-06-03 17:32:33 +02:00
parent e76e9abaf3
commit 89793b75d8
1 changed files with 77 additions and 0 deletions

View File

@ -11,6 +11,7 @@ import Mathlib.Data.Complex.Order
import Mathlib.Data.Complex.Exponential
import Mathlib.Data.Fin.Tuple.Basic
import Mathlib.Topology.Algebra.InfiniteSum.Module
import Mathlib.Topology.Defs.Filter
import Mathlib.Topology.Instances.RealVectorSpace
import Nevanlinna.cauchyRiemann
import Nevanlinna.laplace
@ -240,6 +241,82 @@ theorem log_normSq_of_holomorphicOn_is_harmonicOn'
exact DifferentiableOn.clog h₁ h₃
theorem log_normSq_of_holomorphicOn_is_harmonicOn
{f : }
{s : Set }
(hs : IsOpen s)
(h₁ : DifferentiableOn f s)
(h₂ : ∀ z ∈ s, f z ≠ 0)
(h₃ : ∀ z ∈ s, f z ∈ Complex.slitPlane) :
HarmonicOn (Real.log ∘ Complex.normSq ∘ f) s := by
let s₁ : Set := { z | f z ∈ Complex.slitPlane} ∩ s
have hs₁ : IsOpen s₁ := by
let A := DifferentiableOn.continuousOn h₁
let B := continuousOn_iff'.1 A
obtain ⟨u, hu₁, hu₂⟩ := B Complex.slitPlane Complex.isOpen_slitPlane
have : u ∩ s = s₁ := by
rw [← hu₂]
tauto
rw [← this]
apply IsOpen.inter hu₁ hs
have harm₁ : HarmonicOn (Real.log ∘ Complex.normSq ∘ f) s₁ := by
apply log_normSq_of_holomorphicOn_is_harmonicOn'
exact hs₁
apply DifferentiableOn.mono h₁ (Set.inter_subset_right {z | f z ∈ Complex.slitPlane} s)
-- ∀ z ∈ s₁, f z ≠ 0
exact fun z hz ↦ h₂ z (Set.mem_of_mem_inter_right hz)
-- ∀ z ∈ s₁, f z ∈ Complex.slitPlane
exact fun z hz ↦ h₃ z (Set.mem_of_mem_inter_right hz)
let s₂ : Set := { z | -f z ∈ Complex.slitPlane} ∩ s
have h₁' : DifferentiableOn (-f) s := by
rw [← differentiableOn_neg_iff]
simp
exact h₁
have hs₂ : IsOpen s₂ := by
let A := DifferentiableOn.continuousOn h₁'
let B := continuousOn_iff'.1 A
obtain ⟨u, hu₁, hu₂⟩ := B Complex.slitPlane Complex.isOpen_slitPlane
have : u ∩ s = s₂ := by
rw [← hu₂]
tauto
rw [← this]
apply IsOpen.inter hu₁ hs
have harm₂ : HarmonicOn (Real.log ∘ Complex.normSq ∘ (-f)) s₂ := by
apply log_normSq_of_holomorphicOn_is_harmonicOn'
exact hs₂
apply DifferentiableOn.mono h₁' (Set.inter_subset_right {z | -f z ∈ Complex.slitPlane} s)
-- ∀ z ∈ s₁, f z ≠ 0
intro z hz
simp
exact h₂ z (Set.mem_of_mem_inter_right hz)
-- ∀ z ∈ s₁, f z ∈ Complex.slitPlane
intro z hz
apply hz.1
apply HarmonicOn_of_locally_HarmonicOn
intro z hz
by_cases hfz : f z ∈ Complex.slitPlane
· use s₁
constructor
· exact hs₁
· constructor
· tauto
· have : s₁ = s ∩ s₁ := by
apply Set.right_eq_inter.mpr
exact Set.inter_subset_right {z | f z ∈ Complex.slitPlane} s
rw [← this]
exact harm₁
· use s₂
sorry
theorem log_normSq_of_holomorphic_is_harmonic
{f : }
(h₁ : Differentiable f)