Documentation

Mathlib.Analysis.Complex.Conformal

Conformal maps between complex vector spaces #

We prove the sufficient and necessary conditions for a real-linear map between complex vector spaces to be conformal.

Main results #

Warning #

Antiholomorphic functions such as the complex conjugate are considered as conformal functions in this file.

TODO #

theorem isConformalMap_conj :
IsConformalMap { toLinearEquiv := Complex.conjLIE.toLinearEquiv, continuous_toFun := , continuous_invFun := }

A real continuous linear map on the complex plane is conformal if and only if the map or its conjugate is complex linear, and the map is nonvanishing.

Conformality of real-differentiable complex maps #

theorem DifferentiableAt.conformalAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {z : } {f : E} (h : DifferentiableAt f z) (hf' : deriv f z 0) :

A real differentiable function of the complex plane into some complex normed space E is conformal at a point z if it is holomorphic at that point with a nonvanishing differential. This is a version of the Cauchy-Riemann equations.

A complex function is conformal if and only if the function is holomorphic or antiholomorphic with a nonvanishing differential.

The Cauchy-Riemann Equation for Complex-Differentiable Functions #

theorem real_linearMap_map_smul_complex {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] { : →ₗ[] E} (h : Complex.I = Complex.I 1) (a b : ) :
(a b) = a b

A real linear map ℓ : ℂ →ₗ[ℝ] E respects complex scalar multiplication if it maps I to I • ℓ 1.

Construct a complex-linear map from a real-linear map that maps I to I • ℓ 1.

Equations
Instances For
    @[simp]
    theorem LinearMap.coe_complexOfReal {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] { : →ₗ[] E} (h : Complex.I = Complex.I 1) :
    (.complexOfReal h) =

    Construct a continuous complex-linear map from a continuous real-linear map that maps I to I • ℓ 1.

    Equations
    Instances For
      @[simp]
      theorem ContinuousLinearMap.coe_complexOfReal {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] { : →L[] E} (h : Complex.I = Complex.I 1) :
      (.complexOfReal h) =

      The Cauchy-Riemann Equation: A real-differentiable function f on is complex-differentiable at x within s iff the derivative fderivWithin ℝ f s x maps I to I • (fderivWithin ℝ f s x) 1.

      theorem HasFDerivWithinAt.complexOfReal {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {f : E} {x : } {s : Set } {f' : →L[] E} (h₁ : HasFDerivWithinAt f f' s x) (h₂ : f' Complex.I = Complex.I f' 1) :

      In cases where the Cauchy-Riemann Equation guarantees complex differentiability at x, the complex derivative equals ContinuousLinearMap.complexOfReal of the real derivative.

      theorem complexOfReal_fderivWithin {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {f : E} {x : } {s : Set } (h₁ : DifferentiableWithinAt f s x) (h₂ : (fderivWithin f s x) Complex.I = Complex.I (fderivWithin f s x) 1) (hs : UniqueDiffWithinAt s x) :

      In cases where the Cauchy-Riemann Equation guarantees complex differentiability at x, the complex derivative equals ContinuousLinearMap.complexOfReal of the real derivative.

      theorem complexOfReal_hasDerivWithinAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {f : E} {x : } {s : Set } (h₁ : DifferentiableWithinAt f s x) (h₂ : (fderivWithin f s x) Complex.I = Complex.I (fderivWithin f s x) 1) :

      In cases where the Cauchy-Riemann Equation guarantees complex differentiability at x, the complex derivative equals ContinuousLinearMap.complexOfReal of the real derivative.

      theorem complexOfReal_derivWithin {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {f : E} {x : } {s : Set } (h₁ : DifferentiableWithinAt f s x) (h₂ : (fderivWithin f s x) Complex.I = Complex.I (fderivWithin f s x) 1) (hs : UniqueDiffWithinAt s x) :
      derivWithin f s x = (fderivWithin f s x) 1

      In cases where the Cauchy-Riemann Equation guarantees complex differentiability at x, the complex derivative equals the real derivative.

      The Cauchy-Riemann Equation: A real-differentiable function f on is complex-differentiable at x if and only if the derivative fderiv ℝ f x maps I to I • (fderiv ℝ f x) 1.

      theorem HasFDerivAt.complexOfReal_hasFDerivAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {f : E} {x : } {f' : →L[] E} (h₁ : HasFDerivAt f f' x) (h₂ : f' Complex.I = Complex.I f' 1) :

      In cases where the Cauchy-Riemann Equation guarantees complex differentiability at x, the complex derivative equals ContinuousLinearMap.complexOfReal of the real derivative.

      theorem complexOfReal_hasDerivAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {f : E} {x : } (h₁ : DifferentiableAt f x) (h₂ : (fderiv f x) Complex.I = Complex.I (fderiv f x) 1) :
      HasDerivAt f (((fderiv f x).complexOfReal h₂) 1) x

      In cases where the Cauchy-Riemann Equation guarantees complex differentiability at x, the complex derivative equals ContinuousLinearMap.complexOfReal of the real derivative.

      theorem complexOfReal_deriv {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {f : E} {x : } (h₁ : DifferentiableAt f x) (h₂ : (fderiv f x) Complex.I = Complex.I (fderiv f x) 1) :
      deriv f x = (fderiv f x) 1

      In cases where the Cauchy-Riemann Equation guarantees complex differentiability at x, the complex derivative equals the real derivative.

      theorem complexOfReal_fderiv {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {f : E} {x : } (h₁ : DifferentiableAt f x) (h₂ : (fderiv f x) Complex.I = Complex.I (fderiv f x) 1) :

      In cases where the Cauchy-Riemann Equation guarantees complex differentiability at x, the complex derivative equals ContinuousLinearMap.complexOfReal of the real derivative.