Working…

This commit is contained in:
Stefan Kebekus 2024-06-04 10:27:46 +02:00
parent 4d3332b15d
commit 6eea56e788
3 changed files with 55 additions and 49 deletions

View File

@ -249,6 +249,18 @@ theorem log_normSq_of_holomorphicOn_is_harmonicOn
(h₂ : ∀ z ∈ s, f z ≠ 0) : (h₂ : ∀ z ∈ s, f z ≠ 0) :
HarmonicOn (Real.log ∘ Complex.normSq ∘ f) s := by HarmonicOn (Real.log ∘ Complex.normSq ∘ f) s := by
have slitPlaneLemma {z : } (hz : z ≠ 0) : z ∈ Complex.slitPlane -z ∈ Complex.slitPlane := by
rw [Complex.mem_slitPlane_iff]
rw [Complex.mem_slitPlane_iff]
simp at hz
rw [Complex.ext_iff] at hz
push_neg at hz
simp at hz
simp
by_contra contra
push_neg at contra
exact hz (le_antisymm contra.1.1 contra.2.1) contra.1.2
let s₁ : Set := { z | f z ∈ Complex.slitPlane} ∩ s let s₁ : Set := { z | f z ∈ Complex.slitPlane} ∩ s
have hs₁ : IsOpen s₁ := by have hs₁ : IsOpen s₁ := by
@ -318,10 +330,8 @@ theorem log_normSq_of_holomorphicOn_is_harmonicOn
· exact hs₂ · exact hs₂
· constructor · constructor
· constructor · constructor
· simp · apply Or.resolve_left (slitPlaneLemma (h₂ z hz)) hfz
rw [Complex.mem_slitPlane_iff] · exact hz
rw [Complex.mem_slitPlane_iff] at hfz
simp at hfz
· have : s₂ = s ∩ s₂ := by · have : s₂ = s ∩ s₂ := by
apply Set.right_eq_inter.mpr apply Set.right_eq_inter.mpr
exact Set.inter_subset_right {z | -f z ∈ Complex.slitPlane} s exact Set.inter_subset_right {z | -f z ∈ Complex.slitPlane} s

View File

@ -23,11 +23,12 @@ variable {G₁ : Type*} [NormedAddCommGroup G₁] [NormedSpace G₁] [Comple
def Harmonic (f : → F) : Prop := def Harmonic (f : → F) : Prop :=
(ContDiff 2 f) ∧ (∀ z, Complex.laplace f z = 0) (ContDiff 2 f) ∧ (∀ z, Δ f z = 0)
def HarmonicOn (f : → F) (s : Set ) : Prop := def HarmonicOn (f : → F) (s : Set ) : Prop :=
(ContDiffOn 2 f s) ∧ (∀ z ∈ s, Complex.laplace f z = 0) (ContDiffOn 2 f s) ∧ (∀ z ∈ s, Δ 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)) : theorem HarmonicOn_of_locally_HarmonicOn {f : → F} {s : Set } (h : ∀ x ∈ s, ∃ (u : Set ), IsOpen u ∧ x ∈ u ∧ HarmonicOn f (s ∩ u)) :
@ -135,7 +136,7 @@ theorem harmonic_comp_CLM_is_harmonic {f : → F₁} {l : F₁ →L[] G}
apply ContDiff.comp apply ContDiff.comp
exact ContinuousLinearMap.contDiff l exact ContinuousLinearMap.contDiff l
exact h.1 exact h.1
· rw [laplace_compContLin] · rw [laplace_compCLM]
simp simp
intro z intro z
rw [h.2 z] rw [h.2 z]
@ -152,7 +153,7 @@ theorem harmonicOn_comp_CLM_is_harmonicOn {f : → F₁} {s : Set } {l :
exact h.1 exact h.1
· -- Vanishing of Laplace · -- Vanishing of Laplace
intro z zHyp intro z zHyp
rw [laplace_compContLinAt] rw [laplace_compCLMAt]
simp simp
rw [h.2 z] rw [h.2 z]
simp simp

View File

@ -20,18 +20,25 @@ variable {F : Type*} [NormedAddCommGroup F] [NormedSpace F]
variable {G : Type*} [NormedAddCommGroup G] [NormedSpace G] variable {G : Type*} [NormedAddCommGroup G] [NormedSpace G]
noncomputable def Complex.laplace : ( → F) → ( → F) := noncomputable def Complex.laplace (f : → F) : → F :=
fun f ↦ partialDeriv 1 (partialDeriv 1 f) + partialDeriv Complex.I (partialDeriv Complex.I f) partialDeriv 1 (partialDeriv 1 f) + partialDeriv Complex.I (partialDeriv Complex.I f)
notation "Δ" => Complex.laplace
theorem laplace_eventuallyEq {f₁ f₂ : → F} {x : } (h : f₁ =ᶠ[nhds x] f₂) : Complex.laplace f₁ x = Complex.laplace f₂ x := by theorem laplace_eventuallyEq {f₁ f₂ : → F} {x : } (h : f₁ =ᶠ[nhds x] f₂) : Δ f₁ x = Δ f₂ x := by
unfold Complex.laplace unfold Complex.laplace
simp simp
rw [partialDeriv_eventuallyEq (partialDeriv_eventuallyEq' h 1) 1] rw [partialDeriv_eventuallyEq (partialDeriv_eventuallyEq' h 1) 1]
rw [partialDeriv_eventuallyEq (partialDeriv_eventuallyEq' h Complex.I) Complex.I] 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₂) :
Δ (f₁ + f₂) = (Δ f₁) + (Δ f₂) := by
unfold Complex.laplace unfold Complex.laplace
rw [partialDeriv_add₂] rw [partialDeriv_add₂]
rw [partialDeriv_add₂] rw [partialDeriv_add₂]
@ -59,7 +66,7 @@ theorem laplace_add_ContDiffOn
(hs : IsOpen s) (hs : IsOpen s)
(h₁ : ContDiffOn 2 f₁ s) (h₁ : ContDiffOn 2 f₁ s)
(h₂ : ContDiffOn 2 f₂ s) : (h₂ : ContDiffOn 2 f₂ s) :
∀ x ∈ s, Complex.laplace (f₁ + f₂) x = (Complex.laplace f₁) x + (Complex.laplace f₂) x := by ∀ x ∈ s, Δ (f₁ + f₂) x = (Δ f₁) x + (Δ f₂) x := by
unfold Complex.laplace unfold Complex.laplace
simp simp
@ -124,39 +131,21 @@ theorem laplace_add_ContDiffOn
rw [add_comm] rw [add_comm]
theorem laplace_smul {f : → F} : ∀ v : , Complex.laplace (v • f) = v • (Complex.laplace f) := by theorem laplace_smul {f : → F} : ∀ v : , Δ (v • f) = v • (Δ f) := by
intro v intro v
unfold Complex.laplace unfold Complex.laplace
rw [partialDeriv_smul₂] repeat
rw [partialDeriv_smul₂]
rw [partialDeriv_smul₂]
rw [partialDeriv_smul₂] rw [partialDeriv_smul₂]
simp simp
theorem laplace_compContLin {f : → F} {l : F →L[] G} (h : ContDiff 2 f) : theorem laplace_compCLMAt {f : → F} {l : F →L[] G} {x : } (h : ContDiffAt 2 f x) :
Complex.laplace (l ∘ f) = l ∘ (Complex.laplace f) := by Δ (l ∘ f) x = (l ∘ (Δ f)) x := by
unfold Complex.laplace
rw [partialDeriv_compContLin]
rw [partialDeriv_compContLin]
rw [partialDeriv_compContLin]
rw [partialDeriv_compContLin]
simp
exact (partialDeriv_contDiff h Complex.I).differentiable le_rfl obtain ⟨v, hv₁, hv₂, hv₃, hv₄⟩ : ∃ v ∈ nhds x, (IsOpen v) ∧ (x ∈ v) ∧ (ContDiffOn 2 f v) := by
exact h.differentiable one_le_two obtain ⟨u, hu₁, hu₂⟩ : ∃ u ∈ nhds x, ContDiffOn 2 f u := by
exact (partialDeriv_contDiff h 1).differentiable le_rfl
exact h.differentiable one_le_two
theorem laplace_compContLinAt {f : → F} {l : F →L[] G} {x : } (h : ContDiffAt 2 f x) :
Complex.laplace (l ∘ f) x = (l ∘ (Complex.laplace f)) x := by
have A₂ : ∃ v ∈ nhds x, (IsOpen v) ∧ (x ∈ v) ∧ (ContDiffOn 2 f v) := by
have : ∃ u ∈ nhds x, ContDiffOn 2 f u := by
apply ContDiffAt.contDiffOn h apply ContDiffAt.contDiffOn h
rfl rfl
obtain ⟨u, hu₁, hu₂⟩ := this
obtain ⟨v, hv₁, hv₂, hv₃⟩ := mem_nhds_iff.1 hu₁ obtain ⟨v, hv₁, hv₂, hv₃⟩ := mem_nhds_iff.1 hu₁
use v use v
constructor constructor
@ -166,7 +155,6 @@ theorem laplace_compContLinAt {f : → F} {l : F →L[] G} {x : } (h :
constructor constructor
· exact hv₃ · exact hv₃
· exact ContDiffOn.congr_mono hu₂ (fun x => congrFun rfl) hv₁ · exact ContDiffOn.congr_mono hu₂ (fun x => congrFun rfl) hv₁
obtain ⟨v, hv₁, hv₂, hv₃, hv₄⟩ := A₂
have D : ∀ w : , partialDeriv w (l ∘ f) =ᶠ[nhds x] l ∘ partialDeriv w (f) := by have D : ∀ w : , partialDeriv w (l ∘ f) =ᶠ[nhds x] l ∘ partialDeriv w (f) := by
intro w intro w
@ -192,17 +180,24 @@ theorem laplace_compContLinAt {f : → F} {l : F →L[] G} {x : } (h :
simp simp
-- DifferentiableAt (partialDeriv Complex.I f) x -- DifferentiableAt (partialDeriv Complex.I f) x
apply ContDiffAt.differentiableAt (partialDeriv_contDiffAt (ContDiffOn.contDiffAt hv₄ hv₁) Complex.I) apply ContDiffAt.differentiableAt (partialDeriv_contDiffAt (ContDiffOn.contDiffAt hv₄ hv₁) Complex.I) le_rfl
rfl
-- DifferentiableAt (partialDeriv 1 f) x -- DifferentiableAt (partialDeriv 1 f) x
apply ContDiffAt.differentiableAt (partialDeriv_contDiffAt (ContDiffOn.contDiffAt hv₄ hv₁) 1) apply ContDiffAt.differentiableAt (partialDeriv_contDiffAt (ContDiffOn.contDiffAt hv₄ hv₁) 1) le_rfl
rfl
theorem laplace_compCLE {f : → F} {l : F ≃L[] G} (h : ContDiff 2 f) : theorem laplace_compCLM
Complex.laplace (l ∘ f) = l ∘ (Complex.laplace f) := by {f : → F}
let l' := (l : F →L[] G) {l : F →L[] G}
have : Complex.laplace (l' ∘ f) = l' ∘ (Complex.laplace f) := by (h : ContDiff 2 f) :
exact laplace_compContLin h Δ (l ∘ f) = l ∘ (Δ f) := by
exact this funext z
exact laplace_compCLMAt h.contDiffAt
theorem laplace_compCLE {f : → F} {l : F ≃L[] G} :
Δ (l ∘ f) = l ∘ (Δ f) := by
unfold Complex.laplace
repeat
rw [partialDeriv_compCLE]
simp