Documentation

Mathlib.LinearAlgebra.CliffordAlgebra.Equivs

Other constructions isomorphic to Clifford Algebras #

This file contains isomorphisms showing that other types are equivalent to some CliffordAlgebra.

Rings #

Complex numbers #

We show additionally that this equivalence sends Complex.conj to CliffordAlgebra.involute and vice-versa:

Note that in this algebra CliffordAlgebra.reverse is the identity and so the clifford conjugate is the same as CliffordAlgebra.involute.

Quaternion algebras #

We show additionally that this equivalence sends QuaternionAlgebra.conj to the clifford conjugate and vice-versa:

Dual numbers #

The clifford algebra isomorphic to a ring #

The clifford algebra over a 0-dimensional vector space is isomorphic to its scalars.

Equations
Instances For

    The clifford algebra isomorphic to the complex numbers #

    The quadratic form sending elements to the negation of their square.

    Equations
    Instances For
      @[simp]
      theorem CliffordAlgebraComplex.Q_apply (r : ) :
      Q r = -(r * r)

      The clifford algebras over CliffordAlgebraComplex.Q is isomorphic as an -algebra to .

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        The clifford algebra is commutative since it is isomorphic to the complex numbers.

        TODO: prove this is true for all CliffordAlgebras over a 1-dimensional vector space.

        Equations

        The clifford algebra isomorphic to the quaternions #

        def CliffordAlgebraQuaternion.Q {R : Type u_1} [CommRing R] (c₁ c₂ : R) :

        Q c₁ c₂ is a quadratic form over R × R such that CliffordAlgebra (Q c₁ c₂) is isomorphic as an R-algebra to ℍ[R,c₁,c₂].

        Equations
        Instances For
          @[simp]
          theorem CliffordAlgebraQuaternion.Q_apply {R : Type u_1} [CommRing R] (c₁ c₂ : R) (v : R × R) :
          (Q c₁ c₂) v = c₁ * (v.1 * v.1) + c₂ * (v.2 * v.2)
          def CliffordAlgebraQuaternion.quaternionBasis {R : Type u_1} [CommRing R] (c₁ c₂ : R) :
          QuaternionAlgebra.Basis (CliffordAlgebra (Q c₁ c₂)) c₁ 0 c₂

          The quaternion basis vectors within the algebra.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            @[simp]
            theorem CliffordAlgebraQuaternion.quaternionBasis_k {R : Type u_1} [CommRing R] (c₁ c₂ : R) :
            (quaternionBasis c₁ c₂).k = (CliffordAlgebra.ι (Q c₁ c₂)) (1, 0) * (CliffordAlgebra.ι (Q c₁ c₂)) (0, 1)
            @[simp]
            theorem CliffordAlgebraQuaternion.quaternionBasis_i {R : Type u_1} [CommRing R] (c₁ c₂ : R) :
            (quaternionBasis c₁ c₂).i = (CliffordAlgebra.ι (Q c₁ c₂)) (1, 0)
            @[simp]
            theorem CliffordAlgebraQuaternion.quaternionBasis_j {R : Type u_1} [CommRing R] (c₁ c₂ : R) :
            (quaternionBasis c₁ c₂).j = (CliffordAlgebra.ι (Q c₁ c₂)) (0, 1)
            def CliffordAlgebraQuaternion.toQuaternion {R : Type u_1} [CommRing R] {c₁ c₂ : R} :
            CliffordAlgebra (Q c₁ c₂) →ₐ[R] QuaternionAlgebra R c₁ 0 c₂

            Intermediate result of CliffordAlgebraQuaternion.equiv: clifford algebras over CliffordAlgebraQuaternion.Q can be converted to ℍ[R,c₁,c₂].

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem CliffordAlgebraQuaternion.toQuaternion_ι {R : Type u_1} [CommRing R] {c₁ c₂ : R} (v : R × R) :
              toQuaternion ((CliffordAlgebra.ι (Q c₁ c₂)) v) = { re := 0, imI := v.1, imJ := v.2, imK := 0 }
              theorem CliffordAlgebraQuaternion.toQuaternion_star {R : Type u_1} [CommRing R] {c₁ c₂ : R} (c : CliffordAlgebra (Q c₁ c₂)) :

              The "clifford conjugate" maps to the quaternion conjugate.

              def CliffordAlgebraQuaternion.ofQuaternion {R : Type u_1} [CommRing R] {c₁ c₂ : R} :
              QuaternionAlgebra R c₁ 0 c₂ →ₐ[R] CliffordAlgebra (Q c₁ c₂)

              Map a quaternion into the clifford algebra.

              Equations
              Instances For
                @[simp]
                theorem CliffordAlgebraQuaternion.ofQuaternion_mk {R : Type u_1} [CommRing R] {c₁ c₂ : R} (a₁ a₂ a₃ a₄ : R) :
                ofQuaternion { re := a₁, imI := a₂, imJ := a₃, imK := a₄ } = (algebraMap R (CliffordAlgebra (Q c₁ c₂))) a₁ + a₂ (CliffordAlgebra.ι (Q c₁ c₂)) (1, 0) + a₃ (CliffordAlgebra.ι (Q c₁ c₂)) (0, 1) + a₄ ((CliffordAlgebra.ι (Q c₁ c₂)) (1, 0) * (CliffordAlgebra.ι (Q c₁ c₂)) (0, 1))
                @[simp]
                theorem CliffordAlgebraQuaternion.ofQuaternion_toQuaternion {R : Type u_1} [CommRing R] {c₁ c₂ : R} (c : CliffordAlgebra (Q c₁ c₂)) :
                @[simp]
                theorem CliffordAlgebraQuaternion.toQuaternion_ofQuaternion {R : Type u_1} [CommRing R] {c₁ c₂ : R} (q : QuaternionAlgebra R c₁ 0 c₂) :
                def CliffordAlgebraQuaternion.equiv {R : Type u_1} [CommRing R] {c₁ c₂ : R} :
                CliffordAlgebra (Q c₁ c₂) ≃ₐ[R] QuaternionAlgebra R c₁ 0 c₂

                The clifford algebra over CliffordAlgebraQuaternion.Q c₁ c₂ is isomorphic as an R-algebra to ℍ[R,c₁,c₂].

                Equations
                Instances For
                  @[simp]
                  theorem CliffordAlgebraQuaternion.ofQuaternion_star {R : Type u_1} [CommRing R] {c₁ c₂ : R} (q : QuaternionAlgebra R c₁ 0 c₂) :

                  The quaternion conjugate maps to the "clifford conjugate" (aka star).

                  The clifford algebra isomorphic to the dual numbers #

                  theorem CliffordAlgebraDualNumber.ι_mul_ι {R : Type u_1} [CommRing R] (r₁ r₂ : R) :

                  The clifford algebra over a 1-dimensional vector space with 0 quadratic form is isomorphic to the dual numbers.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For