This commit is contained in:
Stefan Kebekus 2024-05-30 16:20:06 +02:00
parent ac3cd65bf2
commit b0ab121868
3 changed files with 49 additions and 4 deletions

View File

@ -43,6 +43,35 @@ theorem HarmonicOn_of_locally_HarmonicOn {f : → F} {s : Set } (h : ∀
exact (uHyp.2.2.2) x ⟨xHyp, uHyp.2.1⟩ 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
apply?
sorry
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 sorry
rw [laplace_eventuallyEq this]
exact h₁.2 z hz
theorem harmonic_add_harmonic_is_harmonic {f₁ f₂ : → F} (h₁ : Harmonic f₁) (h₂ : Harmonic f₂) : theorem harmonic_add_harmonic_is_harmonic {f₁ f₂ : → F} (h₁ : Harmonic f₁) (h₂ : Harmonic f₂) :
Harmonic (f₁ + f₂) := by Harmonic (f₁ + f₂) := by
constructor constructor
@ -239,8 +268,8 @@ theorem log_normSq_of_holomorphicOn_is_harmonicOn
· exact Real.pi_nonneg · exact Real.pi_nonneg
exact (AddEquivClass.map_ne_zero_iff starRingAut).mpr (h₂ z hz) exact (AddEquivClass.map_ne_zero_iff starRingAut).mpr (h₂ z hz)
exact h₂ z hz exact h₂ z hz
rw [HarmonicOn_ext this]
rw [this] simp
apply harmonic_add_harmonic_is_harmonic apply harmonic_add_harmonic_is_harmonic
have : Complex.log ∘ ⇑(starRingEnd ) ∘ f = Complex.conjCLE ∘ Complex.log ∘ f := by have : Complex.log ∘ ⇑(starRingEnd ) ∘ f = Complex.conjCLE ∘ Complex.log ∘ f := by

View File

@ -24,6 +24,13 @@ noncomputable def Complex.laplace : ( → F) → ( → F) :=
fun f ↦ partialDeriv 1 (partialDeriv 1 f) + partialDeriv Complex.I (partialDeriv Complex.I f) fun f ↦ partialDeriv 1 (partialDeriv 1 f) + partialDeriv Complex.I (partialDeriv Complex.I f)
theorem laplace_eventuallyEq {f₁ f₂ : → F} {x : } (h : f₁ =ᶠ[nhds x] f₂) : Complex.laplace f₁ x = Complex.laplace f₂ x := by
unfold Complex.laplace
simp
rw [partialDeriv_eventuallyEq (partialDeriv_eventuallyEq' h 1) 1]
rw [partialDeriv_eventuallyEq (partialDeriv_eventuallyEq' h Complex.I) Complex.I]
theorem laplace_add {f₁ f₂ : → F} (h₁ : ContDiff 2 f₁) (h₂ : ContDiff 2 f₂): Complex.laplace (f₁ + f₂) = (Complex.laplace f₁) + (Complex.laplace f₂) := by theorem laplace_add {f₁ f₂ : → F} (h₁ : ContDiff 2 f₁) (h₂ : ContDiff 2 f₂): Complex.laplace (f₁ + f₂) = (Complex.laplace f₁) + (Complex.laplace f₂) := by
unfold Complex.laplace unfold Complex.laplace
rw [partialDeriv_add₂] rw [partialDeriv_add₂]

View File

@ -128,6 +128,15 @@ theorem partialDeriv_eventuallyEq {f₁ f₂ : E → F} {x : E} (h : f₁ =ᶠ[n
exact fun v => rfl exact fun v => rfl
theorem partialDeriv_eventuallyEq' {f₁ f₂ : E → F} {x : E} (h : f₁ =ᶠ[nhds x] f₂) : ∀ v : E, partialDeriv 𝕜 v f₁ =ᶠ[nhds x] partialDeriv 𝕜 v f₂ := by
unfold partialDeriv
intro v
let A : fderiv 𝕜 f₁ =ᶠ[nhds x] fderiv 𝕜 f₂ := Filter.EventuallyEq.fderiv h
apply Filter.EventuallyEq.comp₂
exact A
simp
section restrictScalars section restrictScalars
variable (𝕜 : Type*) [NontriviallyNormedField 𝕜] variable (𝕜 : Type*) [NontriviallyNormedField 𝕜]