Isomorphisms with the even subalgebra of a Clifford algebra #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file provides some notable isomorphisms regarding the even subalgebra, clifford_algebra.even.
Main definitions #
- 
clifford_algebra.equiv_even: Every Clifford algebra is isomorphic as an algebra to the even subalgebra of a Clifford algebra with one more dimension.- clifford_algebra.even_equiv.Q': The quadratic form used by this "one-up" algebra.
- clifford_algebra.to_even: The simp-normal form of the forward direction of this isomorphism.
- clifford_algebra.of_even: The simp-normal form of the reverse direction of this isomorphism.
 
- 
clifford_algebra.even_equiv_even_neg: Every even subalgebra is isomorphic to the even subalgebra of the Clifford algebra with negated quadratic form.- clifford_algebra.even_to_neg: The simp-normal form of each direction of this isomorphism.
 
Main results #
- clifford_algebra.coe_to_even_reverse_involute: the behavior of- clifford_algebra.to_evenon the "Clifford conjugate", that is- clifford_algebra.reversecomposed with- clifford_algebra.involute.
Constructions needed for clifford_algebra.equiv_even #
        The quadratic form on the augmented vector space M × R sending v + r•e0 to Q v - r^2.
Equations
The unit vector in the new dimension
Equations
The embedding from the existing vector space
Equations
The embedding from the smaller algebra into the new larger one.
Equations
The embedding from the even subalgebra with an extra dimension into the original algebra.
Equations
- clifford_algebra.of_even Q = let f : M × R →ₗ[R] M × R →ₗ[R] clifford_algebra Q := ((algebra.lmul R (clifford_algebra Q)).to_linear_map.comp ((clifford_algebra.ι Q).comp (linear_map.fst R M R) + (algebra.linear_map R (clifford_algebra Q)).comp (linear_map.snd R M R))).compl₂ ((clifford_algebra.ι Q).comp (linear_map.fst R M R) - (algebra.linear_map R (clifford_algebra Q)).comp (linear_map.snd R M R)) in ⇑(clifford_algebra.even.lift (clifford_algebra.equiv_even.Q' Q)) {bilin := f, contract := _, contract_mid := _}
Any clifford algebra is isomorphic to the even subalgebra of a clifford algebra with an extra
dimension (that is, with vector space M × R), with a quadratic form evaluating to -1 on that new
basis vector.
Equations
The representation of the clifford conjugate (i.e. the reverse of the involute) in the even subalgebra is just the reverse of the representation.
Constructions needed for clifford_algebra.even_equiv_even_neg #
        One direction of clifford_algebra.even_equiv_even_neg
Equations
- clifford_algebra.even_to_neg Q Q' h = ⇑(clifford_algebra.even.lift Q) {bilin := -(clifford_algebra.even.ι Q').bilin, contract := _, contract_mid := _}
The even subalgebras of the algebras with quadratic form Q and -Q are isomorphic.
Stated another way, 𝒞ℓ⁺(p,q,r) and 𝒞ℓ⁺(q,p,r) are isomorphic.
Equations
- clifford_algebra.even_equiv_even_neg Q = alg_equiv.of_alg_hom (clifford_algebra.even_to_neg Q (-Q) _) (clifford_algebra.even_to_neg (-Q) Q _) _ _