2024-04-26 21:18:06 +02:00
|
|
|
|
import Mathlib.Analysis.Calculus.LineDeriv.Basic
|
2024-04-28 21:09:38 +02:00
|
|
|
|
import Mathlib.Analysis.Calculus.ContDiff.Defs
|
2024-04-26 21:18:06 +02:00
|
|
|
|
|
2024-04-28 21:09:38 +02:00
|
|
|
|
noncomputable def Real.laplace : (ℝ × ℝ → ℝ) → (ℝ × ℝ → ℝ) := by
|
2024-04-26 21:18:06 +02:00
|
|
|
|
intro f
|
|
|
|
|
|
2024-04-28 21:09:38 +02:00
|
|
|
|
let f₁ := fun x ↦ lineDeriv ℝ f x ⟨1,0⟩
|
|
|
|
|
let f₁₁ := fun x ↦ lineDeriv ℝ f₁ x ⟨1,0⟩
|
|
|
|
|
let f₂ := fun x ↦ lineDeriv ℝ f x ⟨0,1⟩
|
|
|
|
|
let f₂₂ := fun x ↦ lineDeriv ℝ f₂ x ⟨0,1⟩
|
|
|
|
|
|
2024-04-26 21:18:06 +02:00
|
|
|
|
exact f₁₁ + f₂₂
|
2024-04-28 21:09:38 +02:00
|
|
|
|
|
2024-04-28 21:17:45 +02:00
|
|
|
|
|
|
|
|
|
def Harmonic (f : ℝ × ℝ → ℝ) : Prop :=
|
|
|
|
|
(ContDiff ℝ 2 f) ∧ (∀ x, Real.laplace f x = 0)
|