nevanlinna/nevanlinna/test.lean

56 lines
1.5 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import Mathlib.Analysis.Complex.CauchyIntegral
#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 Complex.log
#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
{w : } -- Point in the interior of the ball
{f : }, -- Holomorphic function
Differentiable f
→ ∀ z ∈ Metric.ball 0 R, f z ≠ 0
→ (∮ (z : ) in C(0, R), Complex.log ‖f z‖ ) = 2 * π * R * log ‖f 0‖ := by
sorry