diff --git a/Nevanlinna/holomorphic_examples.lean b/Nevanlinna/holomorphic_examples.lean index 58689cb..e3cedd2 100644 --- a/Nevanlinna/holomorphic_examples.lean +++ b/Nevanlinna/holomorphic_examples.lean @@ -1,6 +1,6 @@ import Nevanlinna.complexHarmonic import Nevanlinna.holomorphicAt -import Nevanlinna.holomorphic_primitive +import Nevanlinna.holomorphic_primitive2 import Nevanlinna.mathlibAddOn @@ -108,103 +108,145 @@ function. theorem harmonic_is_realOfHolomorphic {f : ℂ → ℝ} - (hf : ∀ z, HarmonicAt f z) : - ∃ F : ℂ → ℂ, (∀ z, HolomorphicAt F z) ∧ (Complex.reCLM ∘ F = f) := by + {z : ℂ} + {R : ℝ} + (hR : 0 < R) + (hf : ∀ x ∈ Metric.ball z R, HarmonicAt f x) : + ∃ F : ℂ → ℂ, (∀ z ∈ Metric.ball z R, HolomorphicAt F z) ∧ (Set.EqOn (Complex.reCLM ∘ F) f (Metric.ball z R)) := by let f_1 : ℂ → ℂ := Complex.ofRealCLM ∘ (partialDeriv ℝ 1 f) let f_I : ℂ → ℂ := Complex.ofRealCLM ∘ (partialDeriv ℝ Complex.I f) let g : ℂ → ℂ := f_1 - Complex.I • f_I - have reg₂f : ContDiff ℝ 2 f := by - apply contDiff_iff_contDiffAt.mpr - intro z - exact (hf z).1 + have contDiffOn_if_contDiffAt + {f' : ℂ → ℝ} + {z' : ℂ} + {R' : ℝ} + {n' : ℕ} + (hf' : ∀ x ∈ Metric.ball z' R', ContDiffAt ℝ n' f' x) : + ContDiffOn ℝ n' f' (Metric.ball z' R') := by + intro z hz + apply ContDiffAt.contDiffWithinAt + exact hf' z hz - have reg₁f_1 : ContDiff ℝ 1 f_1 := by - apply contDiff_iff_contDiffAt.mpr - intro z + have reg₂f : ContDiffOn ℝ 2 f (Metric.ball z R) := by + apply contDiffOn_if_contDiffAt + intro x hx + exact (hf x hx).1 + + have contDiffOn_if_contDiffAt' + {f' : ℂ → ℂ} + {z' : ℂ} + {R' : ℝ} + {n' : ℕ} + (hf' : ∀ x ∈ Metric.ball z' R', ContDiffAt ℝ n' f' x) : + ContDiffOn ℝ n' f' (Metric.ball z' R') := by + intro z hz + apply ContDiffAt.contDiffWithinAt + exact hf' z hz + + have reg₁f_1 : ContDiffOn ℝ 1 f_1 (Metric.ball z R) := by + apply contDiffOn_if_contDiffAt' + intro z hz dsimp [f_1] apply ContDiffAt.continuousLinearMap_comp - exact partialDeriv_contDiffAt ℝ (hf z).1 1 + exact partialDeriv_contDiffAt ℝ (hf z hz).1 1 - have reg₁f_I : ContDiff ℝ 1 f_I := by - apply contDiff_iff_contDiffAt.mpr - intro z + have reg₁f_I : ContDiffOn ℝ 1 f_I (Metric.ball z R) := by + apply contDiffOn_if_contDiffAt' + intro z hz dsimp [f_I] apply ContDiffAt.continuousLinearMap_comp - exact partialDeriv_contDiffAt ℝ (hf z).1 Complex.I + exact partialDeriv_contDiffAt ℝ (hf z hz).1 Complex.I - have reg₁g : ContDiff ℝ 1 g := by + have reg₁g : ContDiffOn ℝ 1 g (Metric.ball z R) := by dsimp [g] - apply ContDiff.sub + apply ContDiffOn.sub exact reg₁f_1 - apply ContDiff.const_smul' + have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := by rfl + rw [this] + apply ContDiffOn.const_smul exact reg₁f_I - have reg₁ : Differentiable ℂ g := by - intro z + have reg₁ : DifferentiableOn ℂ g (Metric.ball z R) := by + intro x hx + apply DifferentiableAt.differentiableWithinAt apply CauchyRiemann₇.2 constructor - · apply Differentiable.differentiableAt - apply ContDiff.differentiable - exact reg₁g - rfl + · apply DifferentiableWithinAt.differentiableAt (reg₁g.differentiableOn le_rfl x hx) + apply IsOpen.mem_nhds Metric.isOpen_ball hx · dsimp [g] - rw [partialDeriv_sub₂, partialDeriv_sub₂] - simp + rw [partialDeriv_sub₂_differentiableAt, partialDeriv_sub₂_differentiableAt] dsimp [f_1, f_I] rw [partialDeriv_smul'₂, partialDeriv_smul'₂] - rw [partialDeriv_compContLin, partialDeriv_compContLin, partialDeriv_compContLin, partialDeriv_compContLin] + rw [partialDeriv_compContLinAt, partialDeriv_compContLinAt] + simp + rw [partialDeriv_compContLinAt, partialDeriv_compContLinAt] rw [mul_sub] simp rw [← mul_assoc] simp rw [add_comm, sub_eq_add_neg] congr 1 - · rw [partialDeriv_comm reg₂f Complex.I 1] - · let A := Filter.EventuallyEq.eq_of_nhds (hf z).2 + · rw [partialDeriv_commOn _ reg₂f Complex.I 1] + exact hx + exact Metric.isOpen_ball + · let A := Filter.EventuallyEq.eq_of_nhds (hf x hx).2 simp at A unfold Complex.laplace at A conv => right right - rw [← sub_zero (partialDeriv ℝ 1 (partialDeriv ℝ 1 f) z)] + rw [← sub_zero (partialDeriv ℝ 1 (partialDeriv ℝ 1 f) x)] rw [← A] simp - --Differentiable ℝ (partialDeriv ℝ _ f) + --DifferentiableAt ℝ (partialDeriv ℝ _ f) repeat - apply ContDiff.differentiable - apply contDiff_iff_contDiffAt.mpr - exact fun w ↦ partialDeriv_contDiffAt ℝ (hf w).1 _ + apply ContDiffAt.differentiableAt + apply partialDeriv_contDiffAt ℝ (hf x hx).1 apply le_rfl - -- Differentiable ℝ f_1 - exact reg₁f_1.differentiable le_rfl - -- Differentiable ℝ (Complex.I • f_I) - apply Differentiable.const_smul' - exact reg₁f_I.differentiable le_rfl - -- Differentiable ℝ f_1 - exact reg₁f_1.differentiable le_rfl - -- Differentiable ℝ (Complex.I • f_I) - apply Differentiable.const_smul' - exact reg₁f_I.differentiable le_rfl - let F := fun z ↦ (primitive 0 g) z + f 0 + -- DifferentiableAt ℝ f_1 x + apply (reg₁f_1.differentiableOn le_rfl).differentiableAt + apply IsOpen.mem_nhds Metric.isOpen_ball hx - have regF : Differentiable ℂ F := by - apply Differentiable.add - apply primitive_differentiable reg₁ + -- DifferentiableAt ℝ (Complex.I • f_I) + have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := by rfl + rw [this] + apply DifferentiableAt.const_smul + apply (reg₁f_I.differentiableOn le_rfl).differentiableAt + apply IsOpen.mem_nhds Metric.isOpen_ball hx + + -- Differentiable ℝ f_1 + apply (reg₁f_1.differentiableOn le_rfl).differentiableAt + apply IsOpen.mem_nhds Metric.isOpen_ball hx + + -- Differentiable ℝ (Complex.I • f_I) + have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := by rfl + rw [this] + apply DifferentiableAt.const_smul + apply (reg₁f_I.differentiableOn le_rfl).differentiableAt + apply IsOpen.mem_nhds Metric.isOpen_ball hx + + let F := fun z' ↦ (primitive z g) z' + f z + + have regF : DifferentiableOn ℂ F (Metric.ball z R) := by + apply DifferentiableOn.add + apply primitive_differentiableOn reg₁ simp - have pF'' : ∀ x, (fderiv ℝ F x) = ContinuousLinearMap.lsmul ℝ ℂ (g x) := by - intro x - rw [DifferentiableAt.fderiv_restrictScalars ℝ (regF x)] + have pF'' : ∀ x ∈ Metric.ball z R, (fderiv ℝ F x) = ContinuousLinearMap.lsmul ℝ ℂ (g x) := by + intro x hx + have : DifferentiableAt ℂ F x := by + apply (regF x hx).differentiableAt + apply IsOpen.mem_nhds Metric.isOpen_ball hx + rw [DifferentiableAt.fderiv_restrictScalars ℝ this] dsimp [F] rw [fderiv_add_const] - rw [primitive_fderiv'] + rw [primitive_fderiv' reg₁ x hx] exact rfl - exact reg₁ use F diff --git a/Nevanlinna/holomorphic_primitive2.lean b/Nevanlinna/holomorphic_primitive2.lean index 61ee4b3..640ee43 100644 --- a/Nevanlinna/holomorphic_primitive2.lean +++ b/Nevanlinna/holomorphic_primitive2.lean @@ -781,12 +781,11 @@ theorem primitive_hasDerivAt apply hasDerivAt_const - -theorem primitive_differentiable +theorem primitive_differentiableOn {E : Type u} [NormedAddCommGroup E] [NormedSpace ℂ E] [CompleteSpace E] {f : ℂ → E} - (z₀ : ℂ) - (R : ℝ) + {z₀ : ℂ} + {R : ℝ} (hf : DifferentiableOn ℂ f (Metric.ball z₀ R)) : DifferentiableOn ℂ (primitive z₀ f) (Metric.ball z₀ R) := by diff --git a/Nevanlinna/partialDeriv.lean b/Nevanlinna/partialDeriv.lean index a2d819e..2b3e1c9 100644 --- a/Nevanlinna/partialDeriv.lean +++ b/Nevanlinna/partialDeriv.lean @@ -38,10 +38,10 @@ theorem partialDeriv_eventuallyEq exact fun v => rfl -theorem partialDeriv_smul₁ +theorem partialDeriv_smul₁ {f : E → F} {a : 𝕜} - {v : E} : + {v : E} : partialDeriv 𝕜 (a • v) f = a • partialDeriv 𝕜 v f := by unfold partialDeriv conv => @@ -102,20 +102,6 @@ theorem partialDeriv_add₂ {f₁ f₂ : E → F} (h₁ : Differentiable 𝕜 f simp -theorem partialDeriv_sub₂ {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 => - left - intro w - left - rw [fderiv_sub (h₁ w) (h₂ w)] - funext w - simp - - theorem partialDeriv_add₂_differentiableAt {f₁ f₂ : E → F} {v : E} @@ -153,6 +139,57 @@ theorem partialDeriv_add₂_contDiffAt exact (hf₂' x (Set.mem_of_mem_inter_right hx)).differentiableAt +theorem partialDeriv_sub₂ {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 => + left + intro w + left + rw [fderiv_sub (h₁ w) (h₂ w)] + funext w + simp + + +theorem partialDeriv_sub₂_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_sub h₁ h₂] + rfl + + +theorem partialDeriv_sub₂_contDiffAt + {f₁ f₂ : E → F} + {v : E} + {x : E} + (h₁ : ContDiffAt 𝕜 1 f₁ x) + (h₂ : ContDiffAt 𝕜 1 f₂ x) : + partialDeriv 𝕜 v (f₁ - f₂) =ᶠ[nhds x] (partialDeriv 𝕜 v f₁) - (partialDeriv 𝕜 v f₂) := by + + obtain ⟨f₁', u₁, hu₁, _, hf₁'⟩ := contDiffAt_one_iff.1 h₁ + obtain ⟨f₂', u₂, hu₂, _, hf₂'⟩ := contDiffAt_one_iff.1 h₂ + + apply Filter.eventuallyEq_iff_exists_mem.2 + use u₁ ∩ u₂ + constructor + · exact Filter.inter_mem hu₁ hu₂ + · intro x hx + simp + apply partialDeriv_sub₂_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}