Integral over a circle in ℂ #
In this file we define ∮ z in C(c, R), f z to be the integral $\oint_{|z-c|=|R|} f(z)\,dz$ and
prove some properties of this integral. We give definition and prove most lemmas for a function
f : ℂ → E, where E is a complex Banach space. For this reason,
some lemmas use, e.g., (z - c)⁻¹ • f z instead of f z / (z - c).
Main definitions #
- circleMap c R: the exponential map $θ ↦ c + R e^{θi}$;
- CircleIntegrable f c R: a function- f : ℂ → Eis integrable on the circle with center- cand radius- Rif- f ∘ circleMap c Ris integrable on- [0, 2π];
- circleIntegral f c R: the integral $\oint_{|z-c|=|R|} f(z)\,dz$, defined as $\int_{0}^{2π}(c + Re^{θ i})' f(c+Re^{θ i})\,dθ$;
- cauchyPowerSeries f c R: the power series that is equal to $\sum_{n=0}^{\infty} \oint_{|z-c|=R} \left(\frac{w-c}{z - c}\right)^n \frac{1}{z-c}f(z)\,dz$ at- w - c. The coefficients of this power series depend only on- f ∘ circleMap c R, and the power series converges to- f wif- fis differentiable on the closed ball- Metric.closedBall c Rand- wbelongs to the corresponding open ball.
Main statements #
- hasFPowerSeriesOn_cauchy_integral: for any circle integrable function- f, the power series- cauchyPowerSeries f c R,- R > 0, converges to the Cauchy integral- (2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - w)⁻¹ • f zon the open disc- Metric.ball c R;
- circleIntegral.integral_sub_zpow_of_undef,- circleIntegral.integral_sub_zpow_of_ne, and- circleIntegral.integral_sub_inv_of_mem_ball: formulas for- ∮ z in C(c, R), (z - w) ^ n,- n : ℤ. These lemmas cover the following cases:- circleIntegral.integral_sub_zpow_of_undef,- n < 0and- |w - c| = |R|: in this case the function is not integrable, so the integral is equal to its default value (zero);
- circleIntegral.integral_sub_zpow_of_ne,- n ≠ -1: in the cases not covered by the previous lemma, we have- (z - w) ^ n = ((z - w) ^ (n + 1) / (n + 1))', thus the integral equals zero;
- circleIntegral.integral_sub_inv_of_mem_ball,- n = -1,- |w - c| < R: in this case the integral is equal to- 2πi.
 - The case - n = -1,- |w -c| > Ris not covered by these lemmas. While it is possible to construct an explicit primitive, it is easier to apply Cauchy theorem, so we postpone the proof till we have this theorem (see #10000).
Notation #
- ∮ z in C(c, R), f z: notation for the integral $\oint_{|z-c|=|R|} f(z)\,dz$, defined as $\int_{0}^{2π}(c + Re^{θ i})' f(c+Re^{θ i})\,dθ$.
Tags #
integral, circle, Cauchy integral
circleMap is 2π-periodic.
The range of circleMap c R is the circle with center c and radius |R|.
Integrability of a function on a circle #
We say that a function f : ℂ → E is integrable on the circle with center c and radius R if
the function f ∘ circleMap c R is integrable on [0, 2π].
Note that the actual function used in the definition of circleIntegral is
(deriv (circleMap c R) θ) • f (circleMap c R θ). Integrability of this function is equivalent
to integrability of f ∘ circleMap c R whenever R ≠ 0.
Equations
- CircleIntegrable f c R = IntervalIntegrable (fun (θ : ℝ) => f (circleMap c R θ)) MeasureTheory.volume 0 (2 * Real.pi)
Instances For
The function we actually integrate over [0, 2π] in the definition of circleIntegral is
integrable.
The function fun z ↦ (z - w) ^ n, n : ℤ, is circle integrable on the circle with center c
and radius |R| if and only if R = 0 or 0 ≤ n, or w does not belong to this circle.
Definition for $\oint_{|z-c|=R} f(z)\,dz$.
Equations
Instances For
Pretty printer defined by notation3 command.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
If f is continuous on the circle |z - c| = R, R > 0, the ‖f z‖ is less than or equal to
C : ℝ on this circle, and this norm is strictly less than C at some point z of the circle,
then ‖∮ z in C(c, R), f z‖ < 2 * π * R * C.
If f' : ℂ → E is a derivative of a complex differentiable function on the circle
Metric.sphere c |R|, then ∮ z in C(c, R), f' z = 0.
If f' : ℂ → E is a derivative of a complex differentiable function on the circle
Metric.sphere c R, then ∮ z in C(c, R), f' z = 0.
If n < 0 and |w - c| = |R|, then (z - w) ^ n is not circle integrable on the circle with
center c and radius |R|, so the integral ∮ z in C(c, R), (z - w) ^ n is equal to zero.
The power series that is equal to
$\frac{1}{2πi}\sum_{n=0}^{\infty} \oint_{|z-c|=R} \left(\frac{w-c}{z - c}\right)^n \frac{1}{z-c}f(z)\,dz$ at
w - c. The coefficients of this power series depend only on f ∘ circleMap c R, and the power
series converges to f w if f is differentiable on the closed ball Metric.closedBall c R and
w belongs to the corresponding open ball. For any circle integrable function f, this power
series converges to the Cauchy integral for f.
Equations
Instances For
For any circle integrable function f, the power series cauchyPowerSeries f c R multiplied
by 2πI converges to the integral ∮ z in C(c, R), (z - w)⁻¹ • f z on the open disc
Metric.ball c R.
For any circle integrable function f, the power series cauchyPowerSeries f c R, R > 0,
converges to the Cauchy integral (2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - w)⁻¹ • f z on the open
disc Metric.ball c R.
For any circle integrable function f, the power series cauchyPowerSeries f c R, R > 0,
converges to the Cauchy integral (2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - w)⁻¹ • f z on the open
disc Metric.ball c R.
For any circle integrable function f, the power series cauchyPowerSeries f c R, R > 0,
converges to the Cauchy integral (2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - w)⁻¹ • f z on the open
disc Metric.ball c R.