diff --git a/Mathlib.lean b/Mathlib.lean index dd909bcc0de0bc..46e81d66fbbac5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1784,6 +1784,7 @@ public import Mathlib.Analysis.Distribution.ContDiffMapSupportedIn public import Mathlib.Analysis.Distribution.DerivNotation public import Mathlib.Analysis.Distribution.Distribution public import Mathlib.Analysis.Distribution.FourierSchwartz +public import Mathlib.Analysis.Distribution.Laplacian public import Mathlib.Analysis.Distribution.SchwartzSpace public import Mathlib.Analysis.Distribution.TemperateGrowth public import Mathlib.Analysis.Distribution.TemperedDistribution diff --git a/Mathlib/Analysis/Distribution/DerivNotation.lean b/Mathlib/Analysis/Distribution/DerivNotation.lean index e5344004c3a6e0..2157f7820cdcfe 100644 --- a/Mathlib/Analysis/Distribution/DerivNotation.lean +++ b/Mathlib/Analysis/Distribution/DerivNotation.lean @@ -8,6 +8,7 @@ module public import Mathlib.Algebra.Module.Equiv.Defs public import Mathlib.Data.Fin.Tuple.Basic public import Mathlib.Topology.Algebra.Module.LinearMap +public import Mathlib.Analysis.InnerProductSpace.CanonicalTensor /-! # Type classes for derivatives @@ -21,11 +22,12 @@ test functions, distributions, tempered distributions, and Sobolev spaces (and o function spaces). -/ -@[expose] public section +@[expose] public noncomputable section universe u' u v w -variable {R V E F : Type*} +variable {ฮน ฮน' ๐•œ R V E F Fโ‚ Fโ‚‚ Fโ‚ƒ Vโ‚ Vโ‚‚ Vโ‚ƒ : Type*} + open Fin @@ -93,13 +95,15 @@ end LineDeriv open LineDeriv /-- -The line derivative is additive, `โˆ‚_{v} (x + y) = โˆ‚_{v} x + โˆ‚_{v} y` for all `x y : E`. +The line derivative is additive, `โˆ‚_{v} (x + y) = โˆ‚_{v} x + โˆ‚_{v} y` for all `x y : E` +and `โˆ‚_{v + w} x = โˆ‚_{v} x + โˆ‚_{w} y` for all `v w : V`. Note that `lineDeriv` on functions is not additive. -/ class LineDerivAdd (V : Type u) (E : Type v) (F : outParam (Type w)) - [AddCommGroup E] [AddCommGroup F] [LineDeriv V E F] where + [AddCommGroup V] [AddCommGroup E] [AddCommGroup F] [LineDeriv V E F] where lineDerivOp_add (v : V) (x y : E) : โˆ‚_{v} (x + y) = โˆ‚_{v} x + โˆ‚_{v} y + lineDerivOp_left_add (v w : V) (x : E) : โˆ‚_{v + w} x = โˆ‚_{v} x + โˆ‚_{w} x /-- The line derivative commutes with scalar multiplication, `โˆ‚_{v} (r โ€ข x) = r โ€ข โˆ‚_{v} x` for all @@ -109,6 +113,14 @@ class LineDerivSMul (R : Type*) (V : Type u) (E : Type v) (F : outParam (Type w) [SMul R E] [SMul R F] [LineDeriv V E F] where lineDerivOp_smul (v : V) (r : R) (x : E) : โˆ‚_{v} (r โ€ข x) = r โ€ข โˆ‚_{v} x +/-- +The line derivative commutes with scalar multiplication, `โˆ‚_{r โ€ข v} x = r โ€ข โˆ‚_{v} x` for all +`r : R` and `v : V`. +-/ +class LineDerivLeftSMul (R : Type*) (V : Type u) (E : Type v) (F : outParam (Type w)) + [SMul R V] [SMul R F] [LineDeriv V E F] where + lineDerivOp_left_smul (r : R) (v : V) (x : E) : โˆ‚_{r โ€ข v} x = r โ€ข โˆ‚_{v} x + /-- The line derivative is continuous. -/ @@ -121,13 +133,15 @@ attribute [fun_prop] ContinuousLineDeriv.continuous_lineDerivOp namespace LineDeriv export LineDerivAdd (lineDerivOp_add) +export LineDerivAdd (lineDerivOp_left_add) export LineDerivSMul (lineDerivOp_smul) +export LineDerivLeftSMul (lineDerivOp_left_smul) export ContinuousLineDeriv (continuous_lineDerivOp) section lineDerivOpCLM variable [Ring R] [AddCommGroup E] [Module R E] [AddCommGroup F] [Module R F] - [TopologicalSpace E] [TopologicalSpace F] + [TopologicalSpace E] [TopologicalSpace F] [AddCommGroup V] [LineDeriv V E F] [LineDerivAdd V E F] [LineDerivSMul R V E F] [ContinuousLineDeriv V E F] variable (R E) in @@ -149,7 +163,7 @@ section iteratedLineDerivOp variable [LineDeriv V E E] variable {n : โ„•} (m : Fin n โ†’ V) -theorem iteratedLineDerivOp_add [AddCommGroup E] [LineDerivAdd V E E] (x y : E) : +theorem iteratedLineDerivOp_add [AddCommGroup V] [AddCommGroup E] [LineDerivAdd V E E] (x y : E) : โˆ‚^{m} (x + y) = โˆ‚^{m} x + โˆ‚^{m} y := by induction n with | zero => @@ -176,7 +190,7 @@ theorem continuous_iteratedLineDerivOp [ContinuousLineDeriv V E E] {n : โ„•} (m | succ n IH => exact (continuous_lineDerivOp _).comp (IH _) -variable [Ring R] [AddCommGroup E] [Module R E] +variable [Ring R] [AddCommGroup V] [AddCommGroup E] [Module R E] [LineDerivAdd V E E] [LineDerivSMul R V E E] [ContinuousLineDeriv V E E] variable (R E) in @@ -194,3 +208,108 @@ theorem iteratedLineDerivOpCLM_apply {n : โ„•} (m : Fin n โ†’ V) (x : E) : end iteratedLineDerivOp end LineDeriv + +/-- +The notation typeclass for the Laplace operator. +-/ +class Laplacian (E : Type v) (F : outParam (Type w)) where + /-- `ฮ” f` is the Laplacian of `f`. The meaning of this notation is type-dependent. -/ + laplacian : E โ†’ F + +namespace Laplacian + +@[inherit_doc] scoped notation "ฮ”" => Laplacian.laplacian + +end Laplacian + +namespace LineDeriv + +variable [LineDeriv E Vโ‚ Vโ‚‚] [LineDeriv E Vโ‚‚ Vโ‚ƒ] + [AddCommGroup Vโ‚] [AddCommGroup Vโ‚‚] [AddCommGroup Vโ‚ƒ] + +/-! ## Laplacian of `LineDeriv` -/ + +section TensorProduct + +variable [CommRing R] [AddCommGroup E] [Module R E] + [Module R Vโ‚‚] [Module R Vโ‚ƒ] + [LineDerivAdd E Vโ‚‚ Vโ‚ƒ] [LineDerivAdd E Vโ‚ Vโ‚‚] + [LineDerivSMul R E Vโ‚‚ Vโ‚ƒ] [LineDerivLeftSMul R E Vโ‚ Vโ‚‚] [LineDerivLeftSMul R E Vโ‚‚ Vโ‚ƒ] + +open InnerProductSpace TensorProduct + +variable (R) in +/-- The second derivative as a bilinear map. + +Mainly used to give an abstract definition of the Laplacian. -/ +def bilinearLineDerivTwo (f : Vโ‚) : E โ†’โ‚—[R] E โ†’โ‚—[R] Vโ‚ƒ := + LinearMap.mkโ‚‚ R (โˆ‚_{ยท} <| โˆ‚_{ยท} f) (by simp [lineDerivOp_left_add]) + (by simp [lineDerivOp_left_smul]) (by simp [lineDerivOp_left_add, lineDerivOp_add]) + (by simp [lineDerivOp_left_smul, lineDerivOp_smul]) + +variable (R) in +/-- The second derivative as a linear map from the tensor product. + +Mainly used to give an abstract definition of the Laplacian. -/ +def tensorLineDerivTwo (f : Vโ‚) : E โŠ—[R] E โ†’โ‚—[R] Vโ‚ƒ := + lift (bilinearLineDerivTwo R f) + +lemma tensorLineDerivTwo_eq_lineDerivOp_lineDerivOp (f : Vโ‚) (v w : E) : + tensorLineDerivTwo R f (v โŠ—โ‚œ[R] w) = โˆ‚_{v} (โˆ‚_{w} f) := lift.tmul _ _ + +end TensorProduct + +section InnerProductSpace + +variable [NormedAddCommGroup E] [InnerProductSpace โ„ E] [FiniteDimensional โ„ E] + +section LinearMap + +variable [CommRing R] + [Module R E] [Module โ„ Vโ‚‚] [Module โ„ Vโ‚ƒ] + [LineDerivAdd E Vโ‚ Vโ‚‚] [LineDerivAdd E Vโ‚‚ Vโ‚ƒ] + [LineDerivSMul โ„ E Vโ‚‚ Vโ‚ƒ] [LineDerivLeftSMul โ„ E Vโ‚ Vโ‚‚] [LineDerivLeftSMul โ„ E Vโ‚‚ Vโ‚ƒ] + +open TensorProduct InnerProductSpace + +theorem tensorLineDerivTwo_canonicalCovariantTensor_eq_sum [Fintype ฮน] (v : OrthonormalBasis ฮน โ„ E) + (f : Vโ‚) : tensorLineDerivTwo โ„ f (canonicalCovariantTensor E) = โˆ‘ i, โˆ‚_{v i} (โˆ‚_{v i} f) := by + simp [InnerProductSpace.canonicalCovariantTensor_eq_sum E v, + tensorLineDerivTwo_eq_lineDerivOp_lineDerivOp] + +end LinearMap + +section ContinuousLinearMap + +section definition + +variable [CommRing R] + [NormedAddCommGroup E] [InnerProductSpace โ„ E] [FiniteDimensional โ„ E] + [Module R Vโ‚] [Module R Vโ‚‚] [Module R Vโ‚ƒ] + [TopologicalSpace Vโ‚] [TopologicalSpace Vโ‚‚] [TopologicalSpace Vโ‚ƒ] [IsTopologicalAddGroup Vโ‚ƒ] + [LineDerivAdd E Vโ‚ Vโ‚‚] [LineDerivSMul R E Vโ‚ Vโ‚‚] [ContinuousLineDeriv E Vโ‚ Vโ‚‚] + [LineDerivAdd E Vโ‚‚ Vโ‚ƒ] [LineDerivSMul R E Vโ‚‚ Vโ‚ƒ] [ContinuousLineDeriv E Vโ‚‚ Vโ‚ƒ] + +variable (R E Vโ‚) in +/-- The Laplacian defined by iterated `lineDerivOp` as a continuous linear map. -/ +def laplacianCLM : Vโ‚ โ†’L[R] Vโ‚ƒ := + โˆ‘ i, lineDerivOpCLM R Vโ‚‚ (stdOrthonormalBasis โ„ E i) โˆ˜L + lineDerivOpCLM R Vโ‚ (stdOrthonormalBasis โ„ E i) + +end definition + +variable [Module โ„ Vโ‚] [Module โ„ Vโ‚‚] [Module โ„ Vโ‚ƒ] + [TopologicalSpace Vโ‚] [TopologicalSpace Vโ‚‚] [TopologicalSpace Vโ‚ƒ] [IsTopologicalAddGroup Vโ‚ƒ] + [LineDerivAdd E Vโ‚ Vโ‚‚] [LineDerivSMul โ„ E Vโ‚ Vโ‚‚] [ContinuousLineDeriv E Vโ‚ Vโ‚‚] + [LineDerivAdd E Vโ‚‚ Vโ‚ƒ] [LineDerivSMul โ„ E Vโ‚‚ Vโ‚ƒ] [ContinuousLineDeriv E Vโ‚‚ Vโ‚ƒ] + [LineDerivLeftSMul โ„ E Vโ‚ Vโ‚‚] [LineDerivLeftSMul โ„ E Vโ‚‚ Vโ‚ƒ] + +theorem laplacianCLM_eq_sum [Fintype ฮน] (v : OrthonormalBasis ฮน โ„ E) (f : Vโ‚) : + laplacianCLM โ„ E Vโ‚ f = โˆ‘ i, โˆ‚_{v i} (โˆ‚_{v i} f) := by + simp [laplacianCLM, โ† tensorLineDerivTwo_canonicalCovariantTensor_eq_sum] + +end ContinuousLinearMap + +end InnerProductSpace + +end LineDeriv diff --git a/Mathlib/Analysis/Distribution/Laplacian.lean b/Mathlib/Analysis/Distribution/Laplacian.lean new file mode 100644 index 00000000000000..4e6e19206807bc --- /dev/null +++ b/Mathlib/Analysis/Distribution/Laplacian.lean @@ -0,0 +1,133 @@ +/- +Copyright (c) 2025 Moritz Doll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Doll +-/ +module + +public import Mathlib.Analysis.Distribution.TemperedDistribution +public import Mathlib.Analysis.InnerProductSpace.Laplacian + +/-! # The Laplacian on Schwartz functions + +We define the Laplacian on Schwartz functions. + +## Main definitions + +* `LineDeriv.laplacianCLM`: The abstract definition of a Laplacian as a sum over the second + derivatives. +* `SchwartzMap.instLaplacian`: The Laplacian for `๐“ข(E, F)` as an instance of the notation type-class + `Laplacian`. + +## Main statements +* `SchwartzMap.laplacian_eq_sum`: The Laplacian is equal to the sum of second derivatives in any + orthonormal basis. +* `SchwartzMap.integral_bilinear_laplacian_right_eq_left`: Integration by parts for the Laplacian. + +## Implementation notes +The abstract definition `LineDeriv.laplacianCLM` does not provide an instance of `Laplacian` because +the type-class system is not able to infer the inner product space `E`. In order to avoid duplicated +definitions, we do not define `LineDeriv.laplacian` and subsequently every concrete instance of +`LineDeriv` has to provide an instance for `Laplacian` and a proof that +`LineDeriv.laplacianCLM _ _ _ f = ฮ” f`, for example see `SchwartzMap.laplacianCLM_eq'` and +`SchwartzMap.laplacian_eq_sum'` below. + +We also note that since `LineDeriv` merely notation and not tied to `fderiv`, it is not possible to +prove the independence of the basis in the definition of the Laplacian in the abstract setting. +In the case of sufficiently smooth functions, this follows from an equality of `lineDerivOp` and +`fderiv`, see for example `SchwartzMap.coe_laplacian_eq_sum`, and in the case of distributions, this +follows from duality. Therefore, when implementing `Laplacian` using `LineDeriv.laplacianCLM`, you +should prove a version of `SchwartzMap.laplacian_eq_sum`. + +-/ + +@[expose] public noncomputable section + +variable {ฮน ฮน' ๐•œ R E F Fโ‚ Fโ‚‚ Fโ‚ƒ Vโ‚ Vโ‚‚ Vโ‚ƒ : Type*} + +variable [NormedAddCommGroup E] [InnerProductSpace โ„ E] [FiniteDimensional โ„ E] + [NormedAddCommGroup F] + +namespace SchwartzMap + +/-! ## Laplacian on `๐“ข(E, F)` -/ + +variable [NormedSpace โ„ F] + +open Laplacian LineDeriv + +instance instLaplacian : Laplacian ๐“ข(E, F) ๐“ข(E, F) where + laplacian := laplacianCLM โ„ E ๐“ข(E, F) + +theorem laplacianCLM_eq' (f : ๐“ข(E, F)) : laplacianCLM โ„ E ๐“ข(E, F) f = ฮ” f := rfl + +theorem laplacian_eq_sum [Fintype ฮน] (b : OrthonormalBasis ฮน โ„ E) (f : ๐“ข(E, F)) : + ฮ” f = โˆ‘ i, โˆ‚_{b i} (โˆ‚_{b i} f) := + LineDeriv.laplacianCLM_eq_sum b f + +variable (๐•œ) in +@[simp] +theorem laplacianCLM_eq [RCLike ๐•œ] [NormedSpace ๐•œ F] (f : ๐“ข(E, F)) : + laplacianCLM ๐•œ E ๐“ข(E, F) f = ฮ” f := by + simp [laplacianCLM, laplacian_eq_sum (stdOrthonormalBasis โ„ E)] + +theorem laplacian_apply (f : ๐“ข(E, F)) (x : E) : ฮ” f x = ฮ” (f : E โ†’ F) x := by + rw [laplacian_eq_sum (stdOrthonormalBasis โ„ E)] + simp only [InnerProductSpace.laplacian_eq_iteratedFDeriv_orthonormalBasis f + (stdOrthonormalBasis โ„ E), sum_apply] + congr 1 + ext i + rw [โ† iteratedLineDerivOp_eq_iteratedFDeriv] + rfl + +open MeasureTheory + +/-! ### Integration by parts -/ + +variable + [NormedAddCommGroup Fโ‚] [NormedSpace โ„ Fโ‚] + [NormedAddCommGroup Fโ‚‚] [NormedSpace โ„ Fโ‚‚] + [NormedAddCommGroup Fโ‚ƒ] [NormedSpace โ„ Fโ‚ƒ] + [MeasurableSpace E] {ฮผ : Measure E} [BorelSpace E] [ฮผ.IsAddHaarMeasure] + +/-- Integration by parts of Schwartz functions for the Laplacian. + +Version for a general bilinear map. -/ +theorem integral_bilinear_laplacian_right_eq_left (f : ๐“ข(E, Fโ‚)) (g : ๐“ข(E, Fโ‚‚)) + (L : Fโ‚ โ†’L[โ„] Fโ‚‚ โ†’L[โ„] Fโ‚ƒ) : + โˆซ x, L (f x) (ฮ” g x) โˆ‚ฮผ = โˆซ x, L (ฮ” f x) (g x) โˆ‚ฮผ := by + simp_rw [laplacian_eq_sum (stdOrthonormalBasis โ„ E), sum_apply, map_sum, + ContinuousLinearMap.coe_sum', Finset.sum_apply] + rw [MeasureTheory.integral_finset_sum, MeasureTheory.integral_finset_sum] + ยท simp [integral_bilinear_lineDerivOp_right_eq_neg_left] + ยท exact fun _ _ โ†ฆ (pairing L (โˆ‚_{_} <| โˆ‚_{_} f) g).integrable + ยท exact fun _ _ โ†ฆ (pairing L f (โˆ‚_{_} <| โˆ‚_{_} g)).integrable + +variable [NormedRing ๐•œ] [NormedSpace โ„ ๐•œ] [IsScalarTower โ„ ๐•œ ๐•œ] [SMulCommClass โ„ ๐•œ ๐•œ] in +/-- Integration by parts of Schwartz functions for the Laplacian. + +Version for multiplication of scalar-valued Schwartz functions. -/ +theorem integral_mul_laplacian_right_eq_left (f : ๐“ข(E, ๐•œ)) (g : ๐“ข(E, ๐•œ)) : + โˆซ x, f x * ฮ” g x โˆ‚ฮผ = โˆซ x, ฮ” f x * g x โˆ‚ฮผ := + integral_bilinear_laplacian_right_eq_left f g (ContinuousLinearMap.mul โ„ ๐•œ) + +variable [RCLike ๐•œ] [NormedSpace ๐•œ F] + +/-- Integration by parts of Schwartz functions for the Laplacian. + +Version for scalar multiplication. -/ +theorem integral_smul_laplacian_right_eq_left (f : ๐“ข(E, ๐•œ)) (g : ๐“ข(E, F)) : + โˆซ x, f x โ€ข ฮ” g x โˆ‚ฮผ = โˆซ x, ฮ” f x โ€ข g x โˆ‚ฮผ := + integral_bilinear_laplacian_right_eq_left f g (ContinuousLinearMap.lsmul โ„ ๐•œ) + +variable [NormedSpace ๐•œ Fโ‚] [NormedSpace ๐•œ Fโ‚‚] + +/-- Integration by parts of Schwartz functions for the Laplacian. + +Version for a Schwartz function with values in continuous linear maps. -/ +theorem integral_clm_comp_laplacian_right_eq_left (f : ๐“ข(E, Fโ‚ โ†’L[๐•œ] Fโ‚‚)) (g : ๐“ข(E, Fโ‚)) : + โˆซ x, f x (ฮ” g x) โˆ‚ฮผ = โˆซ x, ฮ” f x (g x) โˆ‚ฮผ := + integral_bilinear_laplacian_right_eq_left f g + ((ContinuousLinearMap.id ๐•œ (Fโ‚ โ†’L[๐•œ] Fโ‚‚)).bilinearRestrictScalars โ„) + +end SchwartzMap diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 46928e072c7fba..469044c529e6bd 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -1033,17 +1033,29 @@ continuous linear map on Schwartz space. -/ instance instLineDeriv : LineDeriv E ๐“ข(E, F) ๐“ข(E, F) where lineDerivOp m f := (SchwartzMap.evalCLM m).comp (fderivCLM โ„ E F) f +theorem lineDerivOp_apply (m : E) (f : ๐“ข(E, F)) (x : E) : โˆ‚_{m} f x = lineDeriv โ„ f x m := + f.differentiableAt.lineDeriv_eq_fderiv.symm + +theorem lineDerivOp_apply_eq_fderiv (m : E) (f : ๐“ข(E, F)) (x : E) : + โˆ‚_{m} f x = fderiv โ„ f x m := rfl + instance instLineDerivAdd : LineDerivAdd E ๐“ข(E, F) ๐“ข(E, F) where lineDerivOp_add m := ((SchwartzMap.evalCLM m).comp (fderivCLM โ„ E F)).map_add + lineDerivOp_left_add v w f := by + ext x + simp [lineDerivOp_apply_eq_fderiv] instance instLineDerivSMul : LineDerivSMul ๐•œ E ๐“ข(E, F) ๐“ข(E, F) where lineDerivOp_smul m := ((SchwartzMap.evalCLM m).comp (fderivCLM ๐•œ E F)).map_smul +instance instLineDerivLeftSMul : LineDerivLeftSMul โ„ E ๐“ข(E, F) ๐“ข(E, F) where + lineDerivOp_left_smul r y f := by + ext x + simp [lineDerivOp_apply_eq_fderiv] + instance instContinuousLineDeriv : ContinuousLineDeriv E ๐“ข(E, F) ๐“ข(E, F) where continuous_lineDerivOp m := ((SchwartzMap.evalCLM m).comp (fderivCLM โ„ E F)).continuous -open LineDeriv - theorem lineDerivOpCLM_eq (m : E) : lineDerivOpCLM ๐•œ ๐“ข(E, F) m = (SchwartzMap.evalCLM m).comp (fderivCLM ๐•œ E F) := rfl @@ -1053,12 +1065,6 @@ alias pderivCLM := lineDerivOpCLM @[deprecated (since := "2025-11-25")] alias pderivCLM_apply := LineDeriv.lineDerivOpCLM_apply -theorem lineDerivOp_apply (m : E) (f : ๐“ข(E, F)) (x : E) : โˆ‚_{m} f x = lineDeriv โ„ f x m := - f.differentiableAt.lineDeriv_eq_fderiv.symm - -theorem lineDerivOp_apply_eq_fderiv (m : E) (f : ๐“ข(E, F)) (x : E) : - โˆ‚_{m} f x = fderiv โ„ f x m := rfl - @[deprecated (since := "2025-11-25")] alias iteratedPDeriv := LineDeriv.iteratedLineDerivOpCLM @@ -1490,5 +1496,6 @@ theorem integral_clm_comp_lineDerivOp_right_eq_neg_left (f : ๐“ข(D, F โ†’L[๐•œ end integration_by_parts - end SchwartzMap + +set_option linter.style.longFile 1700 diff --git a/Mathlib/Analysis/Distribution/TemperedDistribution.lean b/Mathlib/Analysis/Distribution/TemperedDistribution.lean index e1c97d09b603c2..5882b61266fb14 100644 --- a/Mathlib/Analysis/Distribution/TemperedDistribution.lean +++ b/Mathlib/Analysis/Distribution/TemperedDistribution.lean @@ -313,12 +313,23 @@ continuous linear map on tempered distributions. -/ instance instLineDeriv : LineDeriv E ๐“ข'(E, F) ๐“ข'(E, F) where lineDerivOp m f := PointwiseConvergenceCLM.precomp F (-lineDerivOpCLM โ„‚ ๐“ข(E, โ„‚) m) f +@[simp] +theorem lineDerivOp_apply_apply (f : ๐“ข'(E, F)) (g : ๐“ข(E, โ„‚)) (m : E) : + โˆ‚_{m} f g = f (- โˆ‚_{m} g) := rfl + instance instLineDerivAdd : LineDerivAdd E ๐“ข'(E, F) ๐“ข'(E, F) where lineDerivOp_add m := (PointwiseConvergenceCLM.precomp F (-lineDerivOpCLM โ„‚ ๐“ข(E, โ„‚) m)).map_add + lineDerivOp_left_add x y f := by + ext u + simp [lineDerivOp_left_add, UniformConvergenceCLM.add_apply, add_comm] instance instLineDerivSMul : LineDerivSMul โ„‚ E ๐“ข'(E, F) ๐“ข'(E, F) where lineDerivOp_smul m := (PointwiseConvergenceCLM.precomp F (-lineDerivOpCLM โ„‚ ๐“ข(E, โ„‚) m)).map_smul +instance instLineDerivSMul' : LineDerivSMul โ„ E ๐“ข'(E, F) ๐“ข'(E, F) where + lineDerivOp_smul m := + (PointwiseConvergenceCLM.precomp F (-lineDerivOpCLM โ„‚ ๐“ข(E, โ„‚) m)).map_smul_of_tower + instance instContinuousLineDeriv : ContinuousLineDeriv E ๐“ข'(E, F) ๐“ข'(E, F) where continuous_lineDerivOp m := (PointwiseConvergenceCLM.precomp F (-lineDerivOpCLM โ„‚ ๐“ข(E, โ„‚) m)).continuous @@ -326,10 +337,6 @@ instance instContinuousLineDeriv : ContinuousLineDeriv E ๐“ข'(E, F) ๐“ข'(E, F) theorem lineDerivOpCLM_eq (m : E) : lineDerivOpCLM โ„‚ ๐“ข'(E, F) m = PointwiseConvergenceCLM.precomp F (-lineDerivOpCLM โ„‚ ๐“ข(E, โ„‚) m) := rfl -@[simp] -theorem lineDerivOp_apply_apply (f : ๐“ข'(E, F)) (g : ๐“ข(E, โ„‚)) (m : E) : - โˆ‚_{m} f g = f (- โˆ‚_{m} g) := rfl - variable [MeasurableSpace E] [BorelSpace E] [SecondCountableTopology E] [FiniteDimensional โ„ E] {ฮผ : Measure E} [ฮผ.IsAddHaarMeasure] diff --git a/Mathlib/Analysis/InnerProductSpace/Harmonic/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Harmonic/Basic.lean index 9457ba761d40ed..82c09185b758f5 100644 --- a/Mathlib/Analysis/InnerProductSpace/Harmonic/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Harmonic/Basic.lean @@ -22,7 +22,7 @@ variable {f fโ‚ fโ‚‚ : E โ†’ F} {x : E} {s t : Set E} {c : โ„} -open Topology +open Topology Laplacian namespace InnerProductSpace diff --git a/Mathlib/Analysis/InnerProductSpace/Laplacian.lean b/Mathlib/Analysis/InnerProductSpace/Laplacian.lean index 64154c1ac76c8d..47dc9f208859c6 100644 --- a/Mathlib/Analysis/InnerProductSpace/Laplacian.lean +++ b/Mathlib/Analysis/InnerProductSpace/Laplacian.lean @@ -8,6 +8,7 @@ module public import Mathlib.Analysis.Calculus.ContDiff.Basic public import Mathlib.Analysis.Calculus.ContDiff.Operations public import Mathlib.Analysis.Calculus.IteratedDeriv.Defs +public import Mathlib.Analysis.Distribution.DerivNotation public import Mathlib.Analysis.InnerProductSpace.CanonicalTensor /-! @@ -133,16 +134,14 @@ noncomputable def laplacianWithin : E โ†’ F := @[inherit_doc] scoped[InnerProductSpace] notation "ฮ”[" s "] " f:60 => laplacianWithin f s -variable (f) in -/-- -Laplacian for functions on real inner product spaces. Use `open InnerProductSpace` to access the -notation `ฮ”` for `InnerProductSpace.Laplacian`. --/ -noncomputable def laplacian : E โ†’ F := - fun x โ†ฆ tensorIteratedFDerivTwo โ„ f x (InnerProductSpace.canonicalCovariantTensor E) +noncomputable +instance instLaplacian : Laplacian (E โ†’ F) (E โ†’ F) where + laplacian f x := tensorIteratedFDerivTwo โ„ f x (InnerProductSpace.canonicalCovariantTensor E) -@[inherit_doc] -scoped[InnerProductSpace] notation "ฮ”" => laplacian +@[deprecated (since := "2025-12-31")] +alias InnerProduct.laplacian := _root_.Laplacian.laplacian + +open Laplacian /-- The Laplacian equals the Laplacian with respect to `Set.univ`. @@ -176,8 +175,7 @@ theorem laplacian_eq_iteratedFDeriv_orthonormalBasis {ฮน : Type*} [Fintype ฮน] (v : OrthonormalBasis ฮน โ„ E) : ฮ” f = fun x โ†ฆ โˆ‘ i, iteratedFDeriv โ„ 2 f x ![v i, v i] := by ext x - simp [InnerProductSpace.laplacian, canonicalCovariantTensor_eq_sum E v, - tensorIteratedFDerivTwo_eq_iteratedFDeriv] + simp [laplacian, canonicalCovariantTensor_eq_sum E v, tensorIteratedFDerivTwo_eq_iteratedFDeriv] variable (f) in /--