Documentation

Mathlib.Analysis.Complex.HasPrimitives

Primitives of Holomorphic Functions #

In this file, we give conditions under which holomorphic functions have primitives. The main goal is to prove that holomorphic functions on simply connected domains have primitives. As a first step, we prove that holomorphic functions on disks have primitives. The approach is based on Morera's theorem, that a continuous function (on a disk) whose integral round a rectangle vanishes on all rectangles contained in the disk has a primitive. (Coupled with the fact that holomorphic functions satisfy this property.) To prove Morera's theorem, we first define the Complex.wedgeIntegral, which is the integral of a function over a "wedge" (a horizontal segment followed by a vertical segment in the disk), and compute its derivative.

Main results #

TODO: Extend to holomorphic functions on simply connected domains.

def Complex.wedgeIntegral {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] (z w : ) (f : E) :
E

The (z, w)-wedge-integral of f, is the integral of f over two sides of the rectangle determined by z and w.

Equations
Instances For
    theorem Complex.wedgeIntegral_add_wedgeIntegral_eq {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] (z w : ) (f : E) :
    z.wedgeIntegral w f + w.wedgeIntegral z f = ((( (x : ) in z.re..w.re, f (x + z.im * I)) - (x : ) in z.re..w.re, f (x + w.im * I)) + I (y : ) in z.im..w.im, f (w.re + y * I)) - I (y : ) in z.im..w.im, f (z.re + y * I)
    def Complex.IsConservativeOn {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] (f : E) (U : Set ) :

    A function f IsConservativeOn in U if, for any rectangle contained in U the integral of f over the rectangle is zero.

    Equations
    Instances For
      def Complex.IsExactOn {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] (f : E) (U : Set ) :

      A function f IsExactOn in U if it is the complex derivative of a function on U.

      In complex variable theory, this is also referred to as "having a primitive".

      Equations
      Instances For

        If a function f IsConservativeOn on a disk of center c, then for points z in this disk, the wedge integral from c to z is additive under a detour through a nearby point w.

        theorem Complex.IsConservativeOn.hasDerivAt_wedgeIntegral {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {c : } {r : } {f : E} [CompleteSpace E] (f_cont : ContinuousOn f (Metric.ball c r)) {z : } (hz : z Metric.ball c r) (h : IsConservativeOn f (Metric.ball c r)) :
        HasDerivAt (fun (w : ) => c.wedgeIntegral w f) (f z) z

        The wedgeIntegral has derivative at z equal to f z.

        Morera's theorem for a disk On a disk, a continuous function whose integrals on rectangles vanish, has primitives.

        Morera's theorem for a disk On a disk, a holomorphic function has primitives.