Documentation

Mathlib.Analysis.CStarAlgebra.Unitary.Connected

The unitary group in a unital C⋆-algebra is locally path connected #

When A is a unital C⋆-algebra and u : unitary A is a unitary element whose distance to 1 is less that 2, the spectrum of u is contained in the slit plane, so the principal branch of the logarithm is continuous on the spectrum of u (or equivalently, Complex.arg is continuous on the spectrum). The continuous functional calculus can then be used to define a selfadjoint element x such that u = exp (I • x). Moreover, there is a relatively nice relationship between the norm of x and the norm of u - 1, namely ‖u - 1‖ ^ 2 = 2 * (1 - cos ‖x‖). In fact, these maps u ↦ x and x ↦ u establish a partial homeomorphism between ball (1 : unitary A) 2 and ball (0 : selfAdjoint A) π.

The map t ↦ exp (t • (I • x)) constitutes a path from 1 to u, showing that unitary elements sufficiently close (i.e., within a distance 2) to 1 : unitary A are path connected to 1. This property can be translated around the unitary group to show that if u v : unitary A are unitary elements with ‖u - v‖ < 2, then there is a path joining them. In fact, this path has the property that it lies within closedBall u ‖u - v‖, and consequently any ball of radius δ < 2 in unitary A is path connected. Therefore, the unitary group is locally path connected.

Finally, we provide the standard characterization of the path component of 1 : unitary A as finite products of exponential unitaries.

Main results #

theorem Unitary.two_mul_one_sub_le_norm_sub_one_sq {A : Type u_1} [CStarAlgebra A] {u : A} (hu : u unitary A) {z : } (hz : z spectrum u) :
2 * (1 - z.re) u - 1 ^ 2
@[deprecated Unitary.two_mul_one_sub_le_norm_sub_one_sq (since := "2025-10-29")]
theorem unitary.two_mul_one_sub_le_norm_sub_one_sq {A : Type u_1} [CStarAlgebra A] {u : A} (hu : u unitary A) {z : } (hz : z spectrum u) :
2 * (1 - z.re) u - 1 ^ 2

Alias of Unitary.two_mul_one_sub_le_norm_sub_one_sq.

theorem Unitary.norm_sub_one_sq_eq {A : Type u_1} [CStarAlgebra A] {u : A} (hu : u unitary A) {x : } (hz : IsLeast (Complex.re '' spectrum u) x) :
u - 1 ^ 2 = 2 * (1 - x)
@[deprecated Unitary.norm_sub_one_sq_eq (since := "2025-10-29")]
theorem unitary.norm_sub_one_sq_eq {A : Type u_1} [CStarAlgebra A] {u : A} (hu : u unitary A) {x : } (hz : IsLeast (Complex.re '' spectrum u) x) :
u - 1 ^ 2 = 2 * (1 - x)

Alias of Unitary.norm_sub_one_sq_eq.

theorem Unitary.norm_sub_one_lt_two_iff {A : Type u_1} [CStarAlgebra A] {u : A} (hu : u unitary A) :
u - 1 < 2 -1spectrum u
@[deprecated Unitary.norm_sub_one_lt_two_iff (since := "2025-10-29")]
theorem unitary.norm_sub_one_lt_two_iff {A : Type u_1} [CStarAlgebra A] {u : A} (hu : u unitary A) :
u - 1 < 2 -1spectrum u

Alias of Unitary.norm_sub_one_lt_two_iff.

@[deprecated Unitary.spectrum_subset_slitPlane_iff_norm_lt_two (since := "2025-10-29")]

Alias of Unitary.spectrum_subset_slitPlane_iff_norm_lt_two.

noncomputable def Unitary.argSelfAdjoint {A : Type u_1} [CStarAlgebra A] (u : (unitary A)) :

The selfadjoint element obtained by taking the argument (using the principal branch and the continuous functional calculus) of a unitary whose spectrum does not contain -1. This returns 0 if the principal branch of the logarithm is not continuous on the spectrum of the unitary element.

Equations
Instances For
    @[simp]
    theorem Unitary.argSelfAdjoint_coe {A : Type u_1} [CStarAlgebra A] (u : (unitary A)) :
    (argSelfAdjoint u) = cfc (fun (x : ) => x.arg) u
    @[deprecated Unitary.argSelfAdjoint (since := "2025-10-29")]
    def unitary.argSelfAdjoint {A : Type u_1} [CStarAlgebra A] (u : (unitary A)) :

    Alias of Unitary.argSelfAdjoint.


    The selfadjoint element obtained by taking the argument (using the principal branch and the continuous functional calculus) of a unitary whose spectrum does not contain -1. This returns 0 if the principal branch of the logarithm is not continuous on the spectrum of the unitary element.

    Equations
    Instances For
      @[deprecated Unitary.norm_argSelfAdjoint_le_pi (since := "2025-10-29")]

      Alias of Unitary.norm_argSelfAdjoint_le_pi.

      @[deprecated Unitary.two_mul_one_sub_cos_norm_argSelfAdjoint (since := "2025-10-29")]

      Alias of Unitary.two_mul_one_sub_cos_norm_argSelfAdjoint.

      theorem Unitary.norm_argSelfAdjoint {A : Type u_1} [CStarAlgebra A] {u : (unitary A)} (hu : u - 1 < 2) :
      @[deprecated Unitary.norm_argSelfAdjoint (since := "2025-10-29")]
      theorem unitary.norm_argSelfAdjoint {A : Type u_1} [CStarAlgebra A] {u : (unitary A)} (hu : u - 1 < 2) :

      Alias of Unitary.norm_argSelfAdjoint.

      @[deprecated Unitary.norm_expUnitary_smul_argSelfAdjoint_sub_one_le (since := "2025-10-29")]

      Alias of Unitary.norm_expUnitary_smul_argSelfAdjoint_sub_one_le.

      @[deprecated Unitary.continuousOn_argSelfAdjoint (since := "2025-10-29")]

      Alias of Unitary.continuousOn_argSelfAdjoint.

      the maps unitary.argSelfAdjoint and selfAdjoint.expUnitary form a partial homeomorphism between ball (1 : unitary A) 2 and ball (0 : selfAdjoint A) π.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[deprecated Unitary.openPartialHomeomorph (since := "2025-10-29")]

        Alias of Unitary.openPartialHomeomorph.


        the maps unitary.argSelfAdjoint and selfAdjoint.expUnitary form a partial homeomorphism between ball (1 : unitary A) 2 and ball (0 : selfAdjoint A) π.

        Equations
        Instances For
          theorem Unitary.norm_sub_eq {A : Type u_1} [CStarAlgebra A] (u v : (unitary A)) :
          u - v = ↑(u * star v) - 1
          @[deprecated Unitary.norm_sub_eq (since := "2025-10-29")]
          theorem unitary.norm_sub_eq {A : Type u_1} [CStarAlgebra A] (u v : (unitary A)) :
          u - v = ↑(u * star v) - 1

          Alias of Unitary.norm_sub_eq.

          theorem Unitary.expUnitary_eq_mul_inv {A : Type u_1} [CStarAlgebra A] (u v : (unitary A)) (huv : u - v < 2) :
          @[deprecated Unitary.expUnitary_eq_mul_inv (since := "2025-10-29")]
          theorem unitary.expUnitary_eq_mul_inv {A : Type u_1} [CStarAlgebra A] (u v : (unitary A)) (huv : u - v < 2) :

          Alias of Unitary.expUnitary_eq_mul_inv.

          noncomputable def selfAdjoint.expUnitaryPathToOne {A : Type u_1} [CStarAlgebra A] (x : (selfAdjoint A)) :

          For a selfadjoint element x in a C⋆-algebra, this is the path from 1 to expUnitary x given by t ↦ expUnitary (t • x).

          Equations
          Instances For
            @[simp]
            noncomputable def Unitary.path {A : Type u_1} [CStarAlgebra A] (u v : (unitary A)) (huv : v - u < 2) :
            Path u v

            The path t ↦ expUnitary (t • argSelfAdjoint (v * star u)) * u from u : unitary A to v when ‖v - u‖ < 2.

            Equations
            Instances For
              @[simp]
              theorem Unitary.path_apply {A : Type u_1} [CStarAlgebra A] (u v : (unitary A)) (huv : v - u < 2) (t : unitInterval) :
              (path u v huv) t = selfAdjoint.expUnitary (t argSelfAdjoint (v * star u)) * u
              @[deprecated Unitary.path (since := "2025-10-29")]
              def unitary.path {A : Type u_1} [CStarAlgebra A] (u v : (unitary A)) (huv : v - u < 2) :
              Path u v

              Alias of Unitary.path.


              The path t ↦ expUnitary (t • argSelfAdjoint (v * star u)) * u from u : unitary A to v when ‖v - u‖ < 2.

              Equations
              Instances For
                theorem Unitary.joined {A : Type u_1} [CStarAlgebra A] (u v : (unitary A)) (huv : v - u < 2) :
                Joined u v

                Two unitary elements u and v in a unital C⋆-algebra are joined by a path if the distance between them is less than 2.

                @[deprecated Unitary.joined (since := "2025-10-29")]
                theorem unitary.joined {A : Type u_1} [CStarAlgebra A] (u v : (unitary A)) (huv : v - u < 2) :
                Joined u v

                Alias of Unitary.joined.


                Two unitary elements u and v in a unital C⋆-algebra are joined by a path if the distance between them is less than 2.

                theorem Unitary.isPathConnected_ball {A : Type u_1} [CStarAlgebra A] (u : (unitary A)) (δ : ) (hδ₀ : 0 < δ) (hδ₂ : δ < 2) :

                Any ball of radius δ < 2 in the unitary group of a unital C⋆-algebra is path connected.

                @[deprecated Unitary.isPathConnected_ball (since := "2025-10-29")]
                theorem unitary.isPathConnected_ball {A : Type u_1} [CStarAlgebra A] (u : (unitary A)) (δ : ) (hδ₀ : 0 < δ) (hδ₂ : δ < 2) :

                Alias of Unitary.isPathConnected_ball.


                Any ball of radius δ < 2 in the unitary group of a unital C⋆-algebra is path connected.

                The unitary group in a C⋆-algebra is locally path connected.

                The path component of the identity in the unitary group of a C⋆-algebra is the set of unitaries that can be expressed as a product of exponential unitaries.

                @[deprecated Unitary.mem_pathComponentOne_iff (since := "2025-10-29")]

                Alias of Unitary.mem_pathComponentOne_iff.


                The path component of the identity in the unitary group of a C⋆-algebra is the set of unitaries that can be expressed as a product of exponential unitaries.