diff --git a/Nevanlinna/complexHarmonic.lean b/Nevanlinna/complexHarmonic.lean index fb0aefa..f573576 100644 --- a/Nevanlinna/complexHarmonic.lean +++ b/Nevanlinna/complexHarmonic.lean @@ -37,7 +37,7 @@ theorem HarmonicOn_of_locally_HarmonicOn {f : ℂ → F} {s : Set ℂ} (h : ∀ intro x xHyp obtain ⟨u, uHyp⟩ := h x xHyp use u - exact ⟨ uHyp.1, ⟨uHyp.2.1, uHyp.2.2.1⟩⟩ + 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⟩ @@ -96,7 +96,9 @@ theorem harmonicOn_comp_CLM_is_harmonicOn {f : ℂ → F₁} {s : Set ℂ} {l : · -- Continuous differentiability apply ContDiffOn.continuousLinearMap_comp exact h.1 - · rw [laplace_compContLin] + · -- Vanishing of Laplace + + rw [laplace_compContLin] simp intro z zHyp rw [h.2 z] @@ -104,6 +106,8 @@ theorem harmonicOn_comp_CLM_is_harmonicOn {f : ℂ → F₁} {s : Set ℂ} {l : assumption + + theorem harmonic_iff_comp_CLE_is_harmonic {f : ℂ → F₁} {l : F₁ ≃L[ℝ] G₁} : Harmonic f ↔ Harmonic (l ∘ f) := by diff --git a/Nevanlinna/laplace.lean b/Nevanlinna/laplace.lean index 9d96a3c..ccc4b8a 100644 --- a/Nevanlinna/laplace.lean +++ b/Nevanlinna/laplace.lean @@ -9,6 +9,7 @@ import Mathlib.Data.Complex.Module import Mathlib.Data.Complex.Order import Mathlib.Data.Complex.Exponential import Mathlib.Analysis.RCLike.Basic +import Mathlib.Order.Filter.Basic import Mathlib.Topology.Algebra.InfiniteSum.Module import Mathlib.Topology.Instances.RealVectorSpace import Nevanlinna.cauchyRiemann @@ -74,6 +75,46 @@ theorem laplace_compContLin {f : ℂ → F} {l : F →L[ℝ] G} (h : ContDiff 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 + sorry + obtain ⟨v, hv₁, hv₂, hv₃, hv₄⟩ := A₂ + + have D : ∀ w : ℂ, partialDeriv ℝ w (l ∘ f) =ᶠ[nhds x] l ∘ partialDeriv ℝ w (f) := by + intro w + apply Filter.eventuallyEq_iff_exists_mem.2 + use v + constructor + · exact IsOpen.mem_nhds hv₂ hv₃ + · intro y hy + apply partialDeriv_compContLinAt + let V := ContDiffOn.differentiableOn hv₄ one_le_two + apply DifferentiableOn.differentiableAt V + apply IsOpen.mem_nhds + assumption + assumption + + unfold Complex.laplace + simp + + rw [partialDeriv_eventuallyEq ℝ (D 1) 1] + rw [partialDeriv_compContLinAt] + rw [partialDeriv_eventuallyEq ℝ (D Complex.I) Complex.I] + rw [partialDeriv_compContLinAt] + simp + + -- DifferentiableAt ℝ (partialDeriv ℝ Complex.I f) x + unfold partialDeriv + + sorry + + -- DifferentiableAt ℝ (partialDeriv ℝ 1 f) x + + sorry + + theorem laplace_compCLE {f : ℂ → F} {l : F ≃L[ℝ] G} (h : ContDiff ℝ 2 f) : Complex.laplace (l ∘ f) = l ∘ (Complex.laplace f) := by let l' := (l : F →L[ℝ] G) diff --git a/Nevanlinna/partialDeriv.lean b/Nevanlinna/partialDeriv.lean index 6e73723..ce41acc 100644 --- a/Nevanlinna/partialDeriv.lean +++ b/Nevanlinna/partialDeriv.lean @@ -66,6 +66,12 @@ theorem partialDeriv_compContLin {f : E → F} {l : F →L[𝕜] G} {v : E} (h : rfl +theorem partialDeriv_compContLinAt {f : E → F} {l : F →L[𝕜] G} {v : E} {x : E} (h : DifferentiableAt 𝕜 f x) : (partialDeriv 𝕜 v (l ∘ f)) x = (l ∘ partialDeriv 𝕜 v f) x:= by + unfold partialDeriv + rw [fderiv.comp x (ContinuousLinearMap.differentiableAt l) h] + simp + + theorem partialDeriv_contDiff {n : ℕ} {f : E → F} (h : ContDiff 𝕜 (n + 1) f) : ∀ v : E, ContDiff 𝕜 n (partialDeriv 𝕜 v f) := by unfold partialDeriv intro v @@ -84,6 +90,33 @@ theorem partialDeriv_contDiff {n : ℕ} {f : E → F} (h : ContDiff 𝕜 (n + 1) exact contDiff_const +theorem partialDeriv_contDiffAt {n : ℕ} {f : E → F} {x : E} (h : ContDiffAt 𝕜 (n + 1) f x) : ∀ v : E, ContDiffAt 𝕜 n (partialDeriv 𝕜 v f) x := by + unfold partialDeriv + intro v + + let A' := (contDiffAt_succ_iff_hasFDerivAt.1 h) + obtain ⟨f', ⟨u, hu₁, hu₂⟩ , hf'⟩ := A' + + have : (fun w => (fderiv 𝕜 f w) v) = (fun f => f v) ∘ (fun w => (fderiv 𝕜 f w)) := by + rfl + rw [this] + apply ContDiffAt.comp + apply fderiv_clm_apply + + let A := (contDiffAt_succ_iff_fderiv.1 h).right + simp at A + + have : (fun w => (fderiv 𝕜 f w) v) = (fun f => f v) ∘ (fun w => (fderiv 𝕜 f w)) := by + rfl + + rw [this] + refine ContDiff.comp ?hg A + refine ContDiff.of_succ ?hg.h + refine ContDiff.clm_apply ?hg.h.hf ?hg.h.hg + exact contDiff_id + exact contDiff_const + + lemma partialDeriv_fderiv {f : E → F} (hf : ContDiff 𝕜 2 f) (z a b : E) : fderiv 𝕜 (fderiv 𝕜 f) z b a = partialDeriv 𝕜 b (partialDeriv 𝕜 a f) z := by @@ -94,6 +127,12 @@ lemma partialDeriv_fderiv {f : E → F} (hf : ContDiff 𝕜 2 f) (z a b : E) : · simp +theorem partialDeriv_eventuallyEq {f₁ f₂ : E → F} {x : E} (h : f₁ =ᶠ[nhds x] f₂) : ∀ v : E, partialDeriv 𝕜 v f₁ x = partialDeriv 𝕜 v f₂ x := by + unfold partialDeriv + rw [Filter.EventuallyEq.fderiv_eq h] + exact fun v => rfl + + section restrictScalars variable (𝕜 : Type*) [NontriviallyNormedField 𝕜]