nevanlinna/Nevanlinna/complexHarmonic.lean

192 lines
5.7 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import Mathlib.Analysis.Complex.Basic
import Mathlib.Analysis.Complex.TaylorSeries
import Mathlib.Analysis.Calculus.LineDeriv.Basic
import Mathlib.Analysis.Calculus.ContDiff.Defs
import Mathlib.Analysis.Calculus.FDeriv.Basic
import Mathlib.Analysis.Calculus.FDeriv.Symmetric
import Mathlib.Analysis.RCLike.Basic
import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv
import Mathlib.Data.Complex.Module
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.Instances.RealVectorSpace
import Nevanlinna.cauchyRiemann
import Nevanlinna.laplace
import Nevanlinna.partialDeriv
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace F]
variable {F₁ : Type*} [NormedAddCommGroup F₁] [NormedSpace F₁] [CompleteSpace F₁]
variable {G : Type*} [NormedAddCommGroup G] [NormedSpace G]
variable {G₁ : Type*} [NormedAddCommGroup G₁] [NormedSpace G₁] [CompleteSpace G₁]
def Harmonic (f : → F) : Prop :=
(ContDiff 2 f) ∧ (∀ z, Complex.laplace f z = 0)
def HarmonicOn (f : → F) (s : Set ) : Prop :=
(ContDiffOn 2 f s) ∧ (∀ z ∈ s, Complex.laplace f z = 0)
theorem HarmonicOn_of_locally_HarmonicOn {f : → F} {s : Set } (h : ∀ x ∈ s, ∃ (u : Set ), IsOpen u ∧ x ∈ u ∧ HarmonicOn f (s ∩ u)) :
HarmonicOn f s := by
constructor
· apply contDiffOn_of_locally_contDiffOn
intro x xHyp
obtain ⟨u, uHyp⟩ := h x xHyp
use u
exact ⟨ uHyp.1, ⟨uHyp.2.1, uHyp.2.2.1⟩⟩
· intro x xHyp
obtain ⟨u, uHyp⟩ := h x xHyp
exact (uHyp.2.2.2) x ⟨xHyp, uHyp.2.1⟩
theorem HarmonicOn_congr {f₁ f₂ : → F} {s : Set } (hs : IsOpen s) (hf₁₂ : ∀ x ∈ s, f₁ x = f₂ x) :
HarmonicOn f₁ s ↔ HarmonicOn f₂ s := by
constructor
· intro h₁
constructor
· apply ContDiffOn.congr h₁.1
intro x hx
rw [eq_comm]
exact hf₁₂ x hx
· intro z hz
have : f₁ =ᶠ[nhds z] f₂ := by
unfold Filter.EventuallyEq
unfold Filter.Eventually
simp
refine mem_nhds_iff.mpr ?_
use s
constructor
· exact hf₁₂
· constructor
· exact hs
· exact hz
rw [← laplace_eventuallyEq this]
exact h₁.2 z hz
· intro h₁
constructor
· apply ContDiffOn.congr h₁.1
intro x hx
exact hf₁₂ x hx
· intro z hz
have : f₁ =ᶠ[nhds z] f₂ := by
unfold Filter.EventuallyEq
unfold Filter.Eventually
simp
refine mem_nhds_iff.mpr ?_
use s
constructor
· exact hf₁₂
· constructor
· exact hs
· exact hz
rw [laplace_eventuallyEq this]
exact h₁.2 z hz
theorem harmonic_add_harmonic_is_harmonic {f₁ f₂ : → F} (h₁ : Harmonic f₁) (h₂ : Harmonic f₂) :
Harmonic (f₁ + f₂) := by
constructor
· exact ContDiff.add h₁.1 h₂.1
· rw [laplace_add h₁.1 h₂.1]
simp
intro z
rw [h₁.2 z, h₂.2 z]
simp
theorem harmonicOn_add_harmonicOn_is_harmonicOn {f₁ f₂ : → F} {s : Set } (hs : IsOpen s) (h₁ : HarmonicOn f₁ s) (h₂ : HarmonicOn f₂ s) :
HarmonicOn (f₁ + f₂) s := by
constructor
· exact ContDiffOn.add h₁.1 h₂.1
· intro z hz
rw [laplace_add_ContDiffOn hs h₁.1 h₂.1 z hz]
rw [h₁.2 z hz, h₂.2 z hz]
simp
theorem harmonic_smul_const_is_harmonic {f : → F} {c : } (h : Harmonic f) :
Harmonic (c • f) := by
constructor
· exact ContDiff.const_smul c h.1
· rw [laplace_smul]
dsimp
intro z
rw [h.2 z]
simp
theorem harmonic_iff_smul_const_is_harmonic {f : → F} {c : } (hc : c ≠ 0) :
Harmonic f ↔ Harmonic (c • f) := by
constructor
· exact harmonic_smul_const_is_harmonic
· nth_rewrite 2 [((eq_inv_smul_iff₀ hc).mpr rfl : f = c⁻¹ • c • f)]
exact fun a => harmonic_smul_const_is_harmonic a
theorem harmonic_comp_CLM_is_harmonic {f : → F₁} {l : F₁ →L[] G} (h : Harmonic f) :
Harmonic (l ∘ f) := by
constructor
· -- Continuous differentiability
apply ContDiff.comp
exact ContinuousLinearMap.contDiff l
exact h.1
· rw [laplace_compContLin]
simp
intro z
rw [h.2 z]
simp
exact ContDiff.restrict_scalars h.1
theorem harmonicOn_comp_CLM_is_harmonicOn {f : → F₁} {s : Set } {l : F₁ →L[] G} (hs : IsOpen s) (h : HarmonicOn f s) :
HarmonicOn (l ∘ f) s := by
constructor
· -- Continuous differentiability
apply ContDiffOn.continuousLinearMap_comp
exact h.1
· -- Vanishing of Laplace
intro z zHyp
rw [laplace_compContLinAt]
simp
rw [h.2 z]
simp
assumption
apply ContDiffOn.contDiffAt h.1
exact IsOpen.mem_nhds hs zHyp
theorem harmonic_iff_comp_CLE_is_harmonic {f : → F₁} {l : F₁ ≃L[] G₁} :
Harmonic f ↔ Harmonic (l ∘ f) := by
constructor
· have : l ∘ f = (l : F₁ →L[] G₁) ∘ f := by rfl
rw [this]
exact harmonic_comp_CLM_is_harmonic
· have : f = (l.symm : G₁ →L[] F₁) ∘ l ∘ f := by
funext z
unfold Function.comp
simp
nth_rewrite 2 [this]
exact harmonic_comp_CLM_is_harmonic
theorem harmonicOn_iff_comp_CLE_is_harmonicOn {f : → F₁} {s : Set } {l : F₁ ≃L[] G₁} (hs : IsOpen s) :
HarmonicOn f s ↔ HarmonicOn (l ∘ f) s := by
constructor
· have : l ∘ f = (l : F₁ →L[] G₁) ∘ f := by rfl
rw [this]
exact harmonicOn_comp_CLM_is_harmonicOn hs
· have : f = (l.symm : G₁ →L[] F₁) ∘ l ∘ f := by
funext z
unfold Function.comp
simp
nth_rewrite 2 [this]
exact harmonicOn_comp_CLM_is_harmonicOn hs