Compare commits

...

2 Commits

Author SHA1 Message Date
Stefan Kebekus 7a1359308e done for today 2024-05-31 10:21:27 +02:00
Stefan Kebekus a7b0790675 working 2024-05-31 09:21:35 +02:00
4 changed files with 173 additions and 45 deletions

View File

@ -62,3 +62,21 @@ theorem CauchyRiemann₄ {F : Type*} [NormedAddCommGroup F] [NormedSpace F]
right
intro w
rw [DifferentiableAt.fderiv_restrictScalars (h w)]
theorem CauchyRiemann₅ {F : Type*} [NormedAddCommGroup F] [NormedSpace F] {f : → F} {z : } : (DifferentiableAt f z)
→ partialDeriv Complex.I f z = Complex.I • partialDeriv 1 f z := by
intro h
unfold partialDeriv
conv =>
left
rw [DifferentiableAt.fderiv_restrictScalars h]
simp
rw [← mul_one Complex.I]
rw [← smul_eq_mul]
rw [ContinuousLinearMap.map_smul_of_tower (fderiv f z) Complex.I 1]
conv =>
right
right
rw [DifferentiableAt.fderiv_restrictScalars h]

View File

@ -53,13 +53,13 @@ theorem HarmonicOn_congr {f₁ f₂ : → F} {s : Set } (hs : IsOpen s) (
rw [eq_comm]
exact hf₁₂ x hx
· intro z hz
have : f₁ =ᶠ[nhds z] f₂ := by
have : f₁ =ᶠ[nhds z] f₂ := by
unfold Filter.EventuallyEq
unfold Filter.Eventually
simp
refine mem_nhds_iff.mpr ?_
use s
constructor
constructor
· exact hf₁₂
· constructor
· exact hs
@ -72,13 +72,13 @@ theorem HarmonicOn_congr {f₁ f₂ : → F} {s : Set } (hs : IsOpen s) (
intro x hx
exact hf₁₂ x hx
· intro z hz
have : f₁ =ᶠ[nhds z] f₂ := by
have : f₁ =ᶠ[nhds z] f₂ := by
unfold Filter.EventuallyEq
unfold Filter.Eventually
simp
refine mem_nhds_iff.mpr ?_
use s
constructor
constructor
· exact hf₁₂
· constructor
· exact hs
@ -102,10 +102,9 @@ theorem harmonicOn_add_harmonicOn_is_harmonicOn {f₁ f₂ : → F} {s : Set
HarmonicOn (f₁ + f₂) s := by
constructor
· exact ContDiffOn.add h₁.1 h₂.1
· rw [laplace_add h₁.1 h₂.1]
simp
intro z
rw [h₁.2 z, h₂.2 z]
· intro z hz
rw [laplace_add_ContDiffOn hs h₁.1 h₂.1 z hz]
rw [h₁.2 z hz, h₂.2 z hz]
simp
@ -177,6 +176,21 @@ theorem harmonic_iff_comp_CLE_is_harmonic {f : → F₁} {l : F₁ ≃L[]
exact harmonic_comp_CLM_is_harmonic
theorem harmonicOn_iff_comp_CLE_is_harmonicOn {f : → F₁} {s : Set } {l : F₁ ≃L[] G₁} (hs : IsOpen s) :
HarmonicOn f s ↔ HarmonicOn (l ∘ f) s := by
constructor
· have : l ∘ f = (l : F₁ →L[] G₁) ∘ f := by rfl
rw [this]
exact harmonicOn_comp_CLM_is_harmonicOn hs
· have : f = (l.symm : G₁ →L[] F₁) ∘ l ∘ f := by
funext z
unfold Function.comp
simp
nth_rewrite 2 [this]
exact harmonicOn_comp_CLM_is_harmonicOn hs
theorem holomorphic_is_harmonic {f : → F₁} (h : Differentiable f) :
Harmonic f := by
@ -233,6 +247,71 @@ theorem holomorphic_is_harmonic {f : → F₁} (h : Differentiable f) :
exact fI_is_real_differentiable
theorem holomorphicOn_is_harmonicOn {f : → F₁} {s : Set } (hs : IsOpen s) (h : DifferentiableOn f s) :
HarmonicOn f s := by
-- f is real C²
have f_is_real_C2 : ContDiffOn 2 f s :=
ContDiffOn.restrict_scalars (DifferentiableOn.contDiffOn h hs)
have fI_is_real_differentiable : DifferentiableOn (partialDeriv 1 f) s := by
intro z hz
apply DifferentiableAt.differentiableWithinAt
let ZZ := (f_is_real_C2 z hz).contDiffAt (IsOpen.mem_nhds hs hz)
let AA := partialDeriv_contDiffAt ZZ 1
exact AA.differentiableAt (by rfl)
constructor
· -- f is two times real continuously differentiable
exact f_is_real_C2
· -- Laplace of f is zero
unfold Complex.laplace
intro z hz
simp
have : DifferentiableAt f z := by
sorry
let ZZ := h z hz
rw [CauchyRiemann₅ this]
-- This lemma says that partial derivatives commute with complex scalar
-- multiplication. This is a consequence of partialDeriv_compContLin once we
-- note that complex scalar multiplication is continuous -linear.
have : ∀ v, ∀ s : , ∀ g : → F₁, Differentiable g → partialDeriv v (s • g) = s • (partialDeriv v g) := by
intro v s g hg
-- Present scalar multiplication as a continuous -linear map. This is
-- horrible, there must be better ways to do that.
let sMuls : F₁ →L[] F₁ :=
{
toFun := fun x ↦ s • x
map_add' := by exact fun x y => DistribSMul.smul_add s x y
map_smul' := by exact fun m x => (smul_comm ((RingHom.id ) m) s x).symm
cont := continuous_const_smul s
}
-- Bring the goal into a form that is recognized by
-- partialDeriv_compContLin.
have : s • g = sMuls ∘ g := by rfl
rw [this]
rw [partialDeriv_compContLin hg]
rfl
rw [this]
rw [partialDeriv_comm f_is_real_C2 Complex.I 1]
rw [CauchyRiemann₄ h]
rw [this]
rw [← smul_assoc]
simp
-- Subgoals coming from the application of 'this'
-- Differentiable (Real.partialDeriv 1 f)
exact fI_is_real_differentiable
-- Differentiable (Real.partialDeriv 1 f)
exact fI_is_real_differentiable
theorem re_of_holomorphic_is_harmonic {f : } (h : Differentiable f) :
Harmonic (Complex.reCLM ∘ f) := by
apply harmonic_comp_CLM_is_harmonic
@ -294,26 +373,31 @@ theorem log_normSq_of_holomorphicOn_is_harmonicOn
· exact Real.pi_nonneg
exact (AddEquivClass.map_ne_zero_iff starRingAut).mpr (h₂ z hz)
exact h₂ z hz
rw [HarmonicOn_congr hs this]
simp
apply harmonic_add_harmonic_is_harmonic
have : Complex.log ∘ ⇑(starRingEnd ) ∘ f = Complex.conjCLE ∘ Complex.log ∘ f := by
funext z
apply harmonicOn_add_harmonicOn_is_harmonicOn
exact hs
have : (fun x => Complex.log ((starRingEnd ) (f x))) = (Complex.log ∘ ⇑(starRingEnd ) ∘ f) := by
rfl
rw [this]
have : ∀ z ∈ s, (Complex.log ∘ ⇑(starRingEnd ) ∘ f) z = (Complex.conjCLE ∘ Complex.log ∘ f) z := by
intro z hz
unfold Function.comp
rw [Complex.log_conj]
rfl
exact Complex.slitPlane_arg_ne_pi (h₃ z)
rw [this]
rw [← harmonic_iff_comp_CLE_is_harmonic]
exact Complex.slitPlane_arg_ne_pi (h₃ z hz)
rw [HarmonicOn_congr hs this]
rw [← harmonicOn_iff_comp_CLE_is_harmonicOn]
apply holomorphicOn_is_harmonicOn
intro z
apply DifferentiableAt.comp
exact Complex.differentiableAt_log (h₃ z)
exact h₁ z
repeat
apply holomorphic_is_harmonic
intro z
apply DifferentiableAt.comp
exact Complex.differentiableAt_log (h₃ z)
exact h₁ z
theorem log_normSq_of_holomorphic_is_harmonic

View File

@ -53,35 +53,45 @@ theorem laplace_add {f₁ f₂ : → F} (h₁ : ContDiff 2 f₁) (h₂ :
exact h₂.differentiable one_le_two
theorem laplace_add_ContDiffOn {f₁ f₂ : → F} {s : Set } (hs : IsOpen 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
theorem laplace_add_ContDiffOn
{f₁ f₂ : → F}
{s : Set }
(hs : IsOpen 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
unfold Complex.laplace
simp
intro x hx
have : partialDeriv 1 (f₁ + f₂) =ᶠ[nhds x] (partialDeriv 1 f₁) + (partialDeriv 1 f₂) := by
sorry
rw [partialDeriv_eventuallyEq this]
have t₁ : DifferentiableAt (partialDeriv 1 f₁) x := by
sorry
have t₂ : DifferentiableAt (partialDeriv 1 f₂) x := by
sorry
rw [partialDeriv_add₂_differentiableAt t₁ t₂]
rw [partialDeriv_add₂]
have : partialDeriv Complex.I (f₁ + f₂) =ᶠ[nhds x] (partialDeriv Complex.I f₁) + (partialDeriv Complex.I f₂) := by
sorry
rw [partialDeriv_eventuallyEq this]
have t₃ : DifferentiableAt (partialDeriv Complex.I f₁) x := by
sorry
have t₄ : DifferentiableAt (partialDeriv Complex.I f₂) x := by
sorry
rw [partialDeriv_add₂_differentiableAt t₃ t₄]
rw [partialDeriv_add₂]
rw [partialDeriv_add₂]
rw [partialDeriv_add₂]
exact
add_add_add_comm (partialDeriv 1 (partialDeriv 1 f₁))
(partialDeriv 1 (partialDeriv 1 f₂))
(partialDeriv Complex.I (partialDeriv Complex.I f₁))
(partialDeriv Complex.I (partialDeriv Complex.I f₂))
exact (partialDeriv_contDiff h₁ Complex.I).differentiable le_rfl
exact (partialDeriv_contDiff h₂ Complex.I).differentiable le_rfl
exact h₁.differentiable one_le_two
exact h₂.differentiable one_le_two
exact (partialDeriv_contDiff h₁ 1).differentiable le_rfl
exact (partialDeriv_contDiff h₂ 1).differentiable le_rfl
exact h₁.differentiable one_le_two
exact h₂.differentiable one_le_two
-- I am super confused at this point because the tactic 'ring' does not work.
-- I do not understand why. So, I need to do things by hand.
rw [add_assoc]
rw [add_assoc]
rw [add_right_inj (partialDeriv 1 (partialDeriv 1 f₁) x)]
rw [add_comm]
rw [add_assoc]
rw [add_right_inj (partialDeriv Complex.I (partialDeriv Complex.I f₁) x)]
rw [add_comm]
theorem laplace_smul {f : → F} (h : ContDiff 2 f) : ∀ v : , Complex.laplace (v • f) = v • (Complex.laplace f) := by
@ -119,7 +129,7 @@ theorem laplace_compContLinAt {f : → F} {l : F →L[] G} {x : } (h :
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
obtain ⟨u, hu₁, hu₂⟩ := this
obtain ⟨v, hv₁, hv₂, hv₃⟩ := mem_nhds_iff.1 hu₁
@ -130,7 +140,7 @@ theorem laplace_compContLinAt {f : → F} {l : F →L[] G} {x : } (h :
exact hv₂
constructor
· 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

View File

@ -54,6 +54,22 @@ theorem partialDeriv_add₂ {f₁ f₂ : E → F} {v : E} (h₁ : Differentiable
rw [fderiv_add (h₁ w) (h₂ w)]
theorem partialDeriv_add₂_differentiableAt
{f₁ f₂ : E → F}
{v : E}
{x : E}
(h₁ : DifferentiableAt 𝕜 f₁ x)
(h₂ : DifferentiableAt 𝕜 f₂ x)
:
partialDeriv 𝕜 v (f₁ + f₂) x = (partialDeriv 𝕜 v f₁) x + (partialDeriv 𝕜 v f₂) x := by
unfold partialDeriv
have : f₁ + f₂ = fun y ↦ f₁ y + f₂ y := by rfl
rw [this]
rw [fderiv_add h₁ h₂]
rfl
theorem partialDeriv_compContLin {f : E → F} {l : F →L[𝕜] G} {v : E} (h : Differentiable 𝕜 f) : partialDeriv 𝕜 v (l ∘ f) = l ∘ partialDeriv 𝕜 v f := by
unfold partialDeriv
@ -95,7 +111,7 @@ theorem partialDeriv_contDiffAt {n : } {f : E → F} {x : E} (h : ContDiffAt
unfold partialDeriv
intro v
let eval_at_v : (E →L[𝕜] F) →L[𝕜] F :=
let eval_at_v : (E →L[𝕜] F) →L[𝕜] F :=
{
toFun := fun l ↦ l v
map_add' := by simp
@ -132,7 +148,7 @@ theorem partialDeriv_eventuallyEq' {f₁ f₂ : E → F} {x : E} (h : f₁ =ᶠ[
unfold partialDeriv
intro v
let A : fderiv 𝕜 f₁ =ᶠ[nhds x] fderiv 𝕜 f₂ := Filter.EventuallyEq.fderiv h
apply Filter.EventuallyEq.comp₂
apply Filter.EventuallyEq.comp₂
exact A
simp