Documentation

Mathlib.LinearAlgebra.CliffordAlgebra.Even

The universal property of the even subalgebra #

Main definitions #

Implementation notes #

The approach here is outlined in "Computing with the universal properties of the Clifford algebra and the even subalgebra" (to appear).

The broad summary is that we have two tricks available to us for implementing complex recursors on top of CliffordAlgebra.lift: the first is to use morphisms as the output type, such as A = Module.End R N which is how we obtained CliffordAlgebra.foldr; and the second is to use N = (N', S) where N' is the value we wish to compute, and S is some auxiliary state passed between one recursor invocation and the next. For the universal property of the even subalgebra, we apply a variant of the first trick again by choosing S to itself be a submodule of morphisms.

def CliffordAlgebra.even {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (Q : QuadraticForm R M) :

The even submodule CliffordAlgebra.evenOdd Q 0 is also a subalgebra.

Equations
Instances For
    structure CliffordAlgebra.EvenHom {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (Q : QuadraticForm R M) (A : Type u_3) [Ring A] [Algebra R A] :
    Type (max u_2 u_3)

    The type of bilinear maps which are accepted by CliffordAlgebra.even.lift.

    Instances For
      theorem CliffordAlgebra.EvenHom.ext {R : Type u_1} {M : Type u_2} {inst✝ : CommRing R} {inst✝¹ : AddCommGroup M} {inst✝² : Module R M} {Q : QuadraticForm R M} {A : Type u_3} {inst✝³ : Ring A} {inst✝⁴ : Algebra R A} {x y : EvenHom Q A} (bilin : x.bilin = y.bilin) :
      x = y
      theorem CliffordAlgebra.EvenHom.ext_iff {R : Type u_1} {M : Type u_2} {inst✝ : CommRing R} {inst✝¹ : AddCommGroup M} {inst✝² : Module R M} {Q : QuadraticForm R M} {A : Type u_3} {inst✝³ : Ring A} {inst✝⁴ : Algebra R A} {x y : EvenHom Q A} :
      x = y x.bilin = y.bilin
      def CliffordAlgebra.EvenHom.compr₂ {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} {A : Type u_3} {B : Type u_4} [Ring A] [Ring B] [Algebra R A] [Algebra R B] (g : EvenHom Q A) (f : A →ₐ[R] B) :

      Compose an EvenHom with an AlgHom on the output.

      Equations
      Instances For
        @[simp]
        theorem CliffordAlgebra.EvenHom.compr₂_bilin {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} {A : Type u_3} {B : Type u_4} [Ring A] [Ring B] [Algebra R A] [Algebra R B] (g : EvenHom Q A) (f : A →ₐ[R] B) :
        def CliffordAlgebra.even.ι {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (Q : QuadraticForm R M) :
        EvenHom Q (even Q)

        The embedding of pairs of vectors into the even subalgebra, as a bilinear map.

        Equations
        Instances For
          theorem CliffordAlgebra.even.algHom_ext {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (Q : QuadraticForm R M) {A : Type u_3} [Ring A] [Algebra R A] f g : (even Q) →ₐ[R] A (h : (ι Q).compr₂ f = (ι Q).compr₂ g) :
          f = g

          Two algebra morphisms from the even subalgebra are equal if they agree on pairs of generators.

          See note [partially-applied ext lemmas].

          theorem CliffordAlgebra.even.algHom_ext_iff {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} {A : Type u_3} [Ring A] [Algebra R A] {f g : (even Q) →ₐ[R] A} :
          f = g (ι Q).compr₂ f = (ι Q).compr₂ g
          def CliffordAlgebra.even.lift.aux {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} {A : Type u_3} [Ring A] [Algebra R A] (f : EvenHom Q A) :
          (even Q) →ₗ[R] A

          The final auxiliary construction for CliffordAlgebra.even.lift. This map is the forwards direction of that equivalence, but not in the fully-bundled form.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            theorem CliffordAlgebra.even.lift.aux_apply {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} {A : Type u_3} [Ring A] [Algebra R A] (f : EvenHom Q A) (a✝ : (even Q)) :
            (aux f) a✝ = (((foldr Q (CliffordAlgebra.even.lift.fFold✝ f) ) (1, 0)) a✝).1
            @[simp]
            theorem CliffordAlgebra.even.lift.aux_one {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} {A : Type u_3} [Ring A] [Algebra R A] (f : EvenHom Q A) :
            (aux f) 1 = 1
            @[simp]
            theorem CliffordAlgebra.even.lift.aux_ι {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} {A : Type u_3} [Ring A] [Algebra R A] (f : EvenHom Q A) (m₁ m₂ : M) :
            (aux f) (((ι Q).bilin m₁) m₂) = (f.bilin m₁) m₂
            @[simp]
            theorem CliffordAlgebra.even.lift.aux_algebraMap {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} {A : Type u_3} [Ring A] [Algebra R A] (f : EvenHom Q A) (r : R) (hr : (algebraMap R (CliffordAlgebra Q)) r even Q) :
            (aux f) (algebraMap R (CliffordAlgebra Q)) r, hr = (algebraMap R A) r
            @[simp]
            theorem CliffordAlgebra.even.lift.aux_mul {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} {A : Type u_3} [Ring A] [Algebra R A] (f : EvenHom Q A) (x y : (even Q)) :
            (aux f) (x * y) = (aux f) x * (aux f) y
            def CliffordAlgebra.even.lift {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (Q : QuadraticForm R M) {A : Type u_3} [Ring A] [Algebra R A] :
            EvenHom Q A ((even Q) →ₐ[R] A)

            Every algebra morphism from the even subalgebra is in one-to-one correspondence with a bilinear map that sends duplicate arguments to the quadratic form, and contracts across multiplication.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem CliffordAlgebra.even.lift_symm_apply_bilin {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (Q : QuadraticForm R M) {A : Type u_3} [Ring A] [Algebra R A] (F : (even Q) →ₐ[R] A) :
              @[simp]
              theorem CliffordAlgebra.even.lift_ι {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (Q : QuadraticForm R M) {A : Type u_3} [Ring A] [Algebra R A] (f : EvenHom Q A) (m₁ m₂ : M) :
              ((lift Q) f) (((ι Q).bilin m₁) m₂) = (f.bilin m₁) m₂