diff --git a/Nevanlinna/laplace.lean b/Nevanlinna/laplace.lean index 313a1d3..405e220 100644 --- a/Nevanlinna/laplace.lean +++ b/Nevanlinna/laplace.lean @@ -33,10 +33,10 @@ theorem laplace_eventuallyEq {f₁ f₂ : ℂ → F} {x : ℂ} (h : f₁ =ᶠ[nh rw [partialDeriv_eventuallyEq ℝ (partialDeriv_eventuallyEq' ℝ h Complex.I) Complex.I] -theorem laplace_add +theorem laplace_add {f₁ f₂ : ℂ → F} (h₁ : ContDiff ℝ 2 f₁) - (h₂ : ContDiff ℝ 2 f₂) : + (h₂ : ContDiff ℝ 2 f₂) : Δ (f₁ + f₂) = (Δ f₁) + (Δ f₂) := by unfold Complex.laplace @@ -141,55 +141,11 @@ theorem laplace_add_ContDiffAt unfold Complex.laplace simp - have hf₁ : ∀ z ∈ s, DifferentiableAt ℝ f₁ z := by - intro z hz - convert DifferentiableOn.differentiableAt _ (IsOpen.mem_nhds hs hz) - apply ContDiffOn.differentiableOn h₁ one_le_two - - have hf₂ : ∀ z ∈ s, DifferentiableAt ℝ f₂ z := by - intro z hz - convert DifferentiableOn.differentiableAt _ (IsOpen.mem_nhds hs hz) - apply ContDiffOn.differentiableOn h₂ one_le_two - - have : partialDeriv ℝ 1 (f₁ + f₂) =ᶠ[nhds x] (partialDeriv ℝ 1 f₁) + (partialDeriv ℝ 1 f₂) := by - apply Filter.eventuallyEq_iff_exists_mem.2 - use s - constructor - · exact IsOpen.mem_nhds hs hx - · intro z hz - apply partialDeriv_add₂_differentiableAt - exact hf₁ z hz - exact hf₂ z hz - rw [partialDeriv_eventuallyEq ℝ this] - have t₁ : DifferentiableAt ℝ (partialDeriv ℝ 1 f₁) x := by - let B₀ := (h₁ x hx).contDiffAt (IsOpen.mem_nhds hs hx) - let A₀ := partialDeriv_contDiffAt ℝ B₀ 1 - exact A₀.differentiableAt (Submonoid.oneLE.proof_2 ℕ∞) - have t₂ : DifferentiableAt ℝ (partialDeriv ℝ 1 f₂) x := by - let B₀ := (h₂ x hx).contDiffAt (IsOpen.mem_nhds hs hx) - let A₀ := partialDeriv_contDiffAt ℝ B₀ 1 - exact A₀.differentiableAt (Submonoid.oneLE.proof_2 ℕ∞) - rw [partialDeriv_add₂_differentiableAt ℝ t₁ t₂] - - have : partialDeriv ℝ Complex.I (f₁ + f₂) =ᶠ[nhds x] (partialDeriv ℝ Complex.I f₁) + (partialDeriv ℝ Complex.I f₂) := by - apply Filter.eventuallyEq_iff_exists_mem.2 - use s - constructor - · exact IsOpen.mem_nhds hs hx - · intro z hz - apply partialDeriv_add₂_differentiableAt - exact hf₁ z hz - exact hf₂ z hz - rw [partialDeriv_eventuallyEq ℝ this] - have t₃ : DifferentiableAt ℝ (partialDeriv ℝ Complex.I f₁) x := by - let B₀ := (h₁ x hx).contDiffAt (IsOpen.mem_nhds hs hx) - let A₀ := partialDeriv_contDiffAt ℝ B₀ Complex.I - exact A₀.differentiableAt (Submonoid.oneLE.proof_2 ℕ∞) - have t₄ : DifferentiableAt ℝ (partialDeriv ℝ Complex.I f₂) x := by - let B₀ := (h₂ x hx).contDiffAt (IsOpen.mem_nhds hs hx) - let A₀ := partialDeriv_contDiffAt ℝ B₀ Complex.I - exact A₀.differentiableAt (Submonoid.oneLE.proof_2 ℕ∞) - rw [partialDeriv_add₂_differentiableAt ℝ t₃ t₄] + have h₁₁ : ContDiffAt ℝ 1 f₁ x := h₁.of_le one_le_two + have h₂₁ : ContDiffAt ℝ 1 f₂ x := h₂.of_le one_le_two + repeat + rw [partialDeriv_eventuallyEq ℝ (partialDeriv_add₂_contDiffAt ℝ h₁₁ h₂₁)] + rw [partialDeriv_add₂_differentiableAt] -- 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. @@ -201,6 +157,10 @@ theorem laplace_add_ContDiffAt rw [add_right_inj (partialDeriv ℝ Complex.I (partialDeriv ℝ Complex.I f₁) x)] rw [add_comm] + repeat + apply fun v ↦ (partialDeriv_contDiffAt ℝ h₁ v).differentiableAt le_rfl + apply fun v ↦ (partialDeriv_contDiffAt ℝ h₂ v).differentiableAt le_rfl + theorem laplace_smul {f : ℂ → F} : ∀ v : ℝ, Δ (v • f) = v • (Δ f) := by intro v @@ -257,13 +217,13 @@ theorem laplace_compCLMAt {f : ℂ → F} {l : F →L[ℝ] G} {x : ℂ} (h : Con apply ContDiffAt.differentiableAt (partialDeriv_contDiffAt ℝ (ContDiffOn.contDiffAt hv₄ hv₁) 1) le_rfl -theorem laplace_compCLM - {f : ℂ → F} - {l : F →L[ℝ] G} +theorem laplace_compCLM + {f : ℂ → F} + {l : F →L[ℝ] G} (h : ContDiff ℝ 2 f) : Δ (l ∘ f) = l ∘ (Δ f) := by funext z - exact laplace_compCLMAt h.contDiffAt + exact laplace_compCLMAt h.contDiffAt theorem laplace_compCLE {f : ℂ → F} {l : F ≃L[ℝ] G} : diff --git a/Nevanlinna/partialDeriv.lean b/Nevanlinna/partialDeriv.lean index 245597f..d9813ce 100644 --- a/Nevanlinna/partialDeriv.lean +++ b/Nevanlinna/partialDeriv.lean @@ -66,9 +66,9 @@ theorem partialDeriv_smul₂ {f : E → F} {a : 𝕜} {v : E} : partialDeriv simp -theorem partialDeriv_add₂ {f₁ f₂ : E → F} {v : E} (h₁ : Differentiable 𝕜 f₁) (h₂ : Differentiable 𝕜 f₂) : partialDeriv 𝕜 v (f₁ + f₂) = (partialDeriv 𝕜 v f₁) + (partialDeriv 𝕜 v f₂) := by +theorem partialDeriv_add₂ {f₁ f₂ : E → F} (h₁ : Differentiable 𝕜 f₁) (h₂ : Differentiable 𝕜 f₂) : ∀ v : E, partialDeriv 𝕜 v (f₁ + f₂) = (partialDeriv 𝕜 v f₁) + (partialDeriv 𝕜 v f₂) := by unfold partialDeriv - + intro v have : f₁ + f₂ = fun y ↦ f₁ y + f₂ y := by rfl rw [this] conv => @@ -110,16 +110,16 @@ theorem partialDeriv_add₂_contDiffAt · exact Filter.inter_mem hu₁ hu₂ · intro x hx simp - apply partialDeriv_add₂_differentiableAt 𝕜 + apply partialDeriv_add₂_differentiableAt 𝕜 exact (hf₁' x (Set.mem_of_mem_inter_left hx)).differentiableAt exact (hf₂' x (Set.mem_of_mem_inter_right hx)).differentiableAt - -theorem partialDeriv_compContLin - {f : E → F} - {l : F →L[𝕜] G} - {v : E} - (h : Differentiable 𝕜 f) : + +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