import Mathlib.Analysis.Complex.CauchyIntegral --import Mathlib.Analysis.Complex.Module open ComplexConjugate #check DiffContOnCl.circleIntegral_sub_inv_smul open Real theorem CauchyIntegralFormula : ∀ {R : ℝ} -- Radius of the ball {w : ℂ} -- Point in the interior of the ball {f : ℂ → ℂ}, -- Holomorphic function DiffContOnCl ℂ f (Metric.ball 0 R) → w ∈ Metric.ball 0 R → (∮ (z : ℂ) in C(0, R), (z - w)⁻¹ • f z) = (2 * π * Complex.I) • f w := by exact DiffContOnCl.circleIntegral_sub_inv_smul #check CauchyIntegralFormula #check HasDerivAt.continuousAt #check Real.log #check ComplexConjugate.conj #check Complex.exp theorem SimpleCauchyFormula : ∀ {R : ℝ} -- Radius of the ball {w : ℂ} -- Point in the interior of the ball {f : ℂ → ℂ}, -- Holomorphic function Differentiable ℂ f → w ∈ Metric.ball 0 R → (∮ (z : ℂ) in C(0, R), (z - w)⁻¹ • f z) = (2 * Real.pi * Complex.I) • f w := by intro R w f fHyp apply DiffContOnCl.circleIntegral_sub_inv_smul constructor · exact Differentiable.differentiableOn fHyp · suffices Continuous f from by exact Continuous.continuousOn this rw [continuous_iff_continuousAt] intro x exact DifferentiableAt.continuousAt (fHyp x) theorem JensenFormula₂ : ∀ {R : ℝ} -- Radius of the ball {f : ℂ → ℂ}, -- Holomorphic function Differentiable ℂ f → (∀ z ∈ Metric.ball 0 R, f z ≠ 0) → (∫ (θ : ℝ) in Set.Icc 0 (2 * π), Complex.log ‖f (circleMap 0 R θ)‖ ) = 2 * π * Complex.log ‖f 0‖ := by intro r f fHyp₁ fHyp₂ /- We treat the trivial case r = 0 separately. -/ by_cases rHyp : r = 0 rw [rHyp] simp left unfold ENNReal.ofReal simp rw [max_eq_left (mul_nonneg zero_le_two pi_nonneg)] simp /- From hereon, we assume that r ≠ 0. -/ /- Replace the integral over 0 … 2π by a circle integral -/ suffices (∮ (z : ℂ) in C(0, r), -(Complex.I * z⁻¹ * Complex.log ↑(Complex.abs (f z)))) = 2 * ↑π * Complex.log ↑‖f 0‖ from by have : ∫ (θ : ℝ) in Set.Icc 0 (2 * π), Complex.log ↑‖f (circleMap 0 r θ)‖ = (∮ (z : ℂ) in C(0, r), -(Complex.I * z⁻¹ * Complex.log ↑(Complex.abs (f z)))) := by have : (fun θ ↦ Complex.log ‖f (circleMap 0 r θ)‖) = (fun θ ↦ ((deriv (circleMap 0 r) θ)) • ((deriv (circleMap 0 r) θ)⁻¹ • Complex.log ↑‖f (circleMap 0 r θ)‖)) := by funext θ rw [← smul_assoc, smul_eq_mul, smul_eq_mul, mul_inv_cancel, one_mul] simp exact rHyp rw [this] simp let tmp := circleIntegral_def_Icc (fun z ↦ -(Complex.I * z⁻¹ * (Complex.log ↑‖f z‖))) 0 r simp at tmp rw [← tmp] rw [this] have : ∀ z : ℂ, log (Complex.abs z) = 1/2 * Complex.log z + 1/2 * Complex.log (conj z) := by intro z by_cases argHyp : Complex.arg z = π · rw [Complex.log, argHyp, Complex.log] let ZZ := Complex.arg_conj z rw [argHyp] at ZZ simp at ZZ have : conj z = z := by apply? sorry let ZZ := Complex.log_conj z nth_rewrite 1 [Complex.log] simp let φ := Complex.arg z let r := Complex.abs z have : z = r * Complex.exp (φ * Complex.I) := by exact (Complex.abs_mul_exp_arg_mul_I z).symm have : Complex.log z = Complex.log r + r*Complex.I := by apply? sorry simp at XX sorry sorry