Documentation

Mathlib.Analysis.Complex.Circle

The circle #

This file defines circle to be the metric sphere (Metric.sphere) in β„‚ centred at 0 of radius 1. We equip it with the following structure:

We furthermore define Circle.exp to be the natural map fun t ↦ exp (t * I) from ℝ to circle, and show that this map is a group homomorphism.

We define two additive characters onto the circle:

Implementation notes #

Because later (in Geometry.Manifold.Instances.Sphere) one wants to equip the circle with a smooth manifold structure borrowed from Metric.sphere, the underlying set is {z : β„‚ | abs (z - 0) = 1}. This prevents certain algebraic facts from working definitionally -- for example, the circle is not defeq to {z : β„‚ | abs z = 1}, which is the kernel of Complex.abs considered as a homomorphism from β„‚ to ℝ, nor is it defeq to {z : β„‚ | normSq z = 1}, which is the kernel of the homomorphism Complex.normSq from β„‚ to ℝ.

The unit circle in β„‚.

Equations
theorem Circle.ext {x y : Circle} :
↑x = ↑y β†’ x = y
theorem Circle.ext_iff {x y : Circle} :
x = y ↔ ↑x = ↑y
theorem Circle.coe_inj {x y : Circle} :
↑x = ↑y ↔ x = y
theorem Circle.norm_coe (z : Circle) :
‖↑zβ€– = 1
@[deprecated Circle.norm_coe (since := "2025-02-16")]
theorem Circle.abs_coe (z : Circle) :
‖↑zβ€– = 1

Alias of Circle.norm_coe.

@[simp]
theorem Circle.normSq_coe (z : Circle) :
Complex.normSq ↑z = 1
@[simp]
theorem Circle.coe_ne_zero (z : Circle) :
↑z β‰  0
@[simp]
theorem Circle.coe_one :
↑1 = 1
theorem Circle.coe_eq_one {x : Circle} :
↑x = 1 ↔ x = 1
@[simp]
theorem Circle.coe_mul (z w : Circle) :
↑(z * w) = ↑z * ↑w
@[simp]
theorem Circle.coe_inv (z : Circle) :
↑z⁻¹ = (↑z)⁻¹
@[simp]
theorem Circle.coe_div (z w : Circle) :
↑(z / w) = ↑z / ↑w

The coercion Circle β†’ β„‚ as a monoid homomorphism.

Equations
@[simp]
theorem Circle.coeHom_apply (self : β†₯(Submonoid.unitSphere β„‚)) :
coeHom self = ↑self

The elements of the circle embed into the units.

Equations
@[simp]
theorem Circle.toUnits_apply (z : Circle) :
toUnits z = Units.mk0 ↑z β‹―

If z is a nonzero complex number, then conj z / z belongs to the unit circle.

Equations
@[simp]
theorem Circle.ofConjDivSelf_coe (z : β„‚) (hz : z β‰  0) :
↑(ofConjDivSelf z hz) = (starRingEnd β„‚) z / z

The map fun t => exp (t * I) from ℝ to the unit circle in β„‚.

Equations
@[simp]
theorem Circle.coe_exp (t : ℝ) :
↑(exp t) = Complex.exp (↑t * Complex.I)
@[simp]
theorem Circle.exp_zero :
exp 0 = 1
@[simp]
theorem Circle.exp_add (x y : ℝ) :
exp (x + y) = exp x * exp y

The map fun t => exp (t * I) from ℝ to the unit circle in β„‚, considered as a homomorphism of groups.

Equations
@[simp]
theorem Circle.expHom_apply (a✝ : ℝ) :
expHom a✝ = (⇑Additive.ofMul ∘ ⇑exp) a✝
@[simp]
theorem Circle.exp_sub (x y : ℝ) :
exp (x - y) = exp x / exp y
@[simp]
theorem Circle.exp_neg (x : ℝ) :
@[simp]
theorem Circle.star_addChar {e : AddChar ℝ Circle} (x : ℝ) :
star ↑(e x) = ↑(e (-x))
@[simp]
theorem Circle.starRingEnd_addChar {e : AddChar ℝ Circle} (x : ℝ) :
(starRingEnd β„‚) ↑(e x) = ↑(e (-x))
instance Circle.instSMulCommClass_left {Ξ± : Type u_1} {Ξ² : Type u_2} [SMul β„‚ Ξ²] [SMul Ξ± Ξ²] [SMulCommClass β„‚ Ξ± Ξ²] :
instance Circle.instSMulCommClass_right {Ξ± : Type u_1} {Ξ² : Type u_2} [SMul β„‚ Ξ²] [SMul Ξ± Ξ²] [SMulCommClass Ξ± β„‚ Ξ²] :
instance Circle.instIsScalarTower {Ξ± : Type u_1} {Ξ² : Type u_2} [SMul β„‚ Ξ±] [SMul β„‚ Ξ²] [SMul Ξ± Ξ²] [IsScalarTower β„‚ Ξ± Ξ²] :
theorem Circle.smul_def {Ξ± : Type u_1} [SMul β„‚ Ξ±] (z : Circle) (a : Ξ±) :
z β€’ a = ↑z β€’ a

The additive character from ℝ onto the circle, given by fun x ↦ exp (2 * Ο€ * x * I). Denoted as 𝐞 within the Real.FourierTransform namespace. This uses the analyst convention that there is a 2 * Ο€ in the exponent.

Equations

The additive character from ℝ onto the circle, given by fun x ↦ exp (2 * Ο€ * x * I). Denoted as 𝐞 within the Real.FourierTransform namespace. This uses the analyst convention that there is a 2 * Ο€ in the exponent.

Equations

The additive character from ℝ onto the circle, given by fun x ↦ exp (x * I). This uses the probabilist convention that there is no 2 * Ο€ in the exponent.

Equations
theorem Real.probChar_apply (x : ℝ) :
↑(probChar x) = Complex.exp (↑x * Complex.I)