Documentation

Mathlib.Analysis.SpecialFunctions.PolarCoord

Polar coordinates #

We define polar coordinates, as a partial homeomorphism in ℝ^2 between ℝ^2 - (-∞, 0] and (0, +∞) × (-π, π). Its inverse is given by (r, θ) ↦ (r cos θ, r sin θ).

It satisfies the following change of variables formula (see integral_comp_polarCoord_symm): ∫ p in polarCoord.target, p.1 • f (polarCoord.symm p) = ∫ p, f p

The polar coordinates partial homeomorphism in ℝ^2, mapping (r cos θ, r sin θ) to (r, θ). It is a homeomorphism between ℝ^2 - (-∞, 0] and (0, +∞) × (-π, π).

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[simp]
    theorem polarCoord_apply (q : × ) :
    @[simp]
    @[simp]
    theorem polarCoord_symm_apply (p : × ) :
    polarCoord.symm p = (p.1 * Real.cos p.2, p.1 * Real.sin p.2)

    The polar coordinates partial homeomorphism in , mapping r (cos θ + I * sin θ) to (r, θ). It is a homeomorphism between ℂ - ℝ≤0 and (0, +∞) × (-π, π).

    Equations
    Instances For
      @[simp]
      theorem Complex.polarCoord_symm_apply (p : × ) :
      Complex.polarCoord.symm p = p.1 * ((Real.cos p.2) + (Real.sin p.2) * I)
      noncomputable def fderivPiPolarCoordSymm {ι : Type u_1} (p : ι × ) :
      (ι × ) →L[] ι ×

      The derivative of polarCoord.symm on ι → ℝ × ℝ, see hasFDerivAt_pi_polarCoord_symm.

      Equations
      Instances For
        theorem injOn_pi_polarCoord_symm {ι : Type u_1} :
        Set.InjOn (fun (p : ι × ) (i : ι) => polarCoord.symm (p i)) (Set.univ.pi fun (x : ι) => polarCoord.target)
        theorem abs_fst_of_mem_pi_polarCoord_target {ι : Type u_1} {p : ι × } (hp : p Set.univ.pi fun (x : ι) => polarCoord.target) (i : ι) :
        |(p i).1| = (p i).1
        theorem hasFDerivAt_pi_polarCoord_symm {ι : Type u_1} [Fintype ι] (p : ι × ) :
        HasFDerivAt (fun (x : ι × ) (i : ι) => polarCoord.symm (x i)) (fderivPiPolarCoordSymm p) p
        theorem det_fderivPiPolarCoordSymm {ι : Type u_1} [Fintype ι] (p : ι × ) :
        (fderivPiPolarCoordSymm p).det = i : ι, (p i).1
        theorem integral_comp_pi_polarCoord_symm {ι : Type u_1} [Fintype ι] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace E] (f : (ι × )E) :
        ( (p : ι × ) in Set.univ.pi fun (x : ι) => polarCoord.target, (∏ i : ι, (p i).1) f fun (i : ι) => polarCoord.symm (p i)) = (p : ι × ), f p
        theorem Complex.integral_comp_pi_polarCoord_symm {ι : Type u_1} [Fintype ι] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace E] (f : (ι)E) :
        ( (p : ι × ) in Set.univ.pi fun (x : ι) => Complex.polarCoord.target, (∏ i : ι, (p i).1) f fun (i : ι) => Complex.polarCoord.symm (p i)) = (p : ι), f p
        theorem lintegral_comp_pi_polarCoord_symm {ι : Type u_1} [Fintype ι] (f : (ι × )ENNReal) :
        (∫⁻ (p : ι × ) in Set.univ.pi fun (x : ι) => polarCoord.target, (∏ i : ι, ENNReal.ofReal (p i).1) * f fun (i : ι) => polarCoord.symm (p i)) = ∫⁻ (p : ι × ), f p
        theorem Complex.lintegral_comp_pi_polarCoord_symm {ι : Type u_1} [Fintype ι] (f : (ι)ENNReal) :
        (∫⁻ (p : ι × ) in Set.univ.pi fun (x : ι) => Complex.polarCoord.target, (∏ i : ι, ENNReal.ofReal (p i).1) * f fun (i : ι) => Complex.polarCoord.symm (p i)) = ∫⁻ (p : ι), f p