import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv import Nevanlinna.complexHarmonic import Nevanlinna.complexHarmonic import Nevanlinna.holomorphicAt theorem CauchyRiemann₆ {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] {F : Type*} [NormedAddCommGroup F] [NormedSpace ℂ F] {f : E → F} {z : E} : (DifferentiableAt ℂ f z) ↔ (DifferentiableAt ℝ f z) ∧ ∀ e, partialDeriv ℝ (Complex.I • e) f z = Complex.I • partialDeriv ℝ e f z := by constructor · -- Direction "→" intro h constructor · exact DifferentiableAt.restrictScalars ℝ h · unfold partialDeriv conv => intro e left rw [DifferentiableAt.fderiv_restrictScalars ℝ h] simp rw [← mul_one Complex.I] rw [← smul_eq_mul] conv => intro e right right rw [DifferentiableAt.fderiv_restrictScalars ℝ h] simp · -- Direction "←" intro ⟨h₁, h₂⟩ apply (differentiableAt_iff_restrictScalars ℝ h₁).2 use { toFun := fderiv ℝ f z map_add' := fun x y => ContinuousLinearMap.map_add (fderiv ℝ f z) x y map_smul' := by simp intro m x have : m = m.re + m.im • Complex.I := by simp rw [this, add_smul, add_smul, ContinuousLinearMap.map_add] congr simp rw [smul_assoc, smul_assoc, ContinuousLinearMap.map_smul (fderiv ℝ f z) m.2] congr exact h₂ x } rfl theorem CauchyRiemann₇ {F : Type*} [NormedAddCommGroup F] [NormedSpace ℂ F] {f : ℂ → F} {z : ℂ} : (DifferentiableAt ℂ f z) ↔ (DifferentiableAt ℝ f z) ∧ partialDeriv ℝ Complex.I f z = Complex.I • partialDeriv ℝ 1 f z := by constructor · intro hf constructor · exact (CauchyRiemann₆.1 hf).1 · let A := (CauchyRiemann₆.1 hf).2 1 simp at A exact A · intro ⟨h₁, h₂⟩ apply CauchyRiemann₆.2 constructor · exact h₁ · intro e have : Complex.I • e = e • Complex.I := by rw [smul_eq_mul, smul_eq_mul] exact CommMonoid.mul_comm Complex.I e rw [this] have : e = e.re + e.im • Complex.I := by simp rw [this, add_smul, partialDeriv_add₁, partialDeriv_add₁] simp rw [← smul_eq_mul] have : partialDeriv ℝ ((e.re : ℝ) • Complex.I) f = partialDeriv ℝ ((e.re : ℂ) • Complex.I) f := by rfl rw [← this, partialDeriv_smul₁ ℝ] have : (e.re : ℂ) = (e.re : ℝ) • (1 : ℂ) := by simp rw [this, partialDeriv_smul₁ ℝ] have : partialDeriv ℝ ((e.im : ℂ) * Complex.I) f = partialDeriv ℝ ((e.im : ℝ) • Complex.I) f := by rfl rw [this, partialDeriv_smul₁ ℝ] simp rw [h₂] rw [smul_comm] congr rw [mul_assoc] simp nth_rw 2 [smul_comm] rw [← smul_assoc] simp have : - (e.im : ℂ) = (-e.im : ℝ) • (1 : ℂ) := by simp rw [this, partialDeriv_smul₁ ℝ] simp /- A harmonic, real-valued function on ℂ is the real part of a suitable holomorphic function. -/ theorem harmonic_is_realOfHolomorphic {f : ℂ → ℝ} (hf : ∀ z, HarmonicAt f z) : ∃ F : ℂ → ℂ, ∀ z, (HolomorphicAt F z ∧ ((F z).re = f z)) := 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₀ : Differentiable ℝ g := by let smulICLM : ℂ ≃L[ℝ] ℂ := { toFun := fun x ↦ Complex.I • x map_add' := fun x y => DistribSMul.smul_add Complex.I x y map_smul' := fun m x => (smul_comm ((RingHom.id ℝ) m) Complex.I x).symm invFun := fun x ↦ (Complex.I)⁻¹ • x left_inv := by intro x simp rw [← mul_assoc, mul_comm] simp right_inv := by intro x simp rw [← mul_assoc] simp continuous_toFun := continuous_const_smul Complex.I continuous_invFun := continuous_const_smul (Complex.I)⁻¹ } apply Differentiable.sub apply Differentiable.comp exact ContinuousLinearMap.differentiable Complex.ofRealCLM intro z sorry apply Differentiable.comp sorry sorry have reg₁ : Differentiable ℂ g := by intro z apply CauchyRiemann₇.2 constructor · exact reg₀ z · dsimp [g] have : f_1 - Complex.I • f_I = f_1 + (- Complex.I • f_I) := by rw [sub_eq_add_neg] simp rw [this, partialDeriv_add₂, partialDeriv_add₂] simp dsimp [f_1, f_I] sorry sorry sorry sorry sorry sorry