Quaternions #
In this file we define quaternions ℍ[R]
over a commutative ring R
, and define some
algebraic structures on ℍ[R]
.
Main definitions #
QuaternionAlgebra R a b c
,ℍ[R, a, b, c]
: Bourbaki, Algebra I with coefficientsa
,b
,c
(Many other references such as Wikipedia assume $\operatorname{char} R ≠ 2$ therefore one can complete the square and WLOG assume $b = 0$.)Quaternion R
,ℍ[R]
: the space of quaternions, a.k.a.QuaternionAlgebra R (-1) (0) (-1)
;Quaternion.normSq
: square of the norm of a quaternion;
We also define the following algebraic structures on ℍ[R]
:
Ring ℍ[R, a, b, c]
,StarRing ℍ[R, a, b, c]
, andAlgebra R ℍ[R, a, b, c]
: for any commutative ringR
;Ring ℍ[R]
,StarRing ℍ[R]
, andAlgebra R ℍ[R]
: for any commutative ringR
;IsDomain ℍ[R]
: for a linear ordered commutative ringR
;DivisionRing ℍ[R]
: for a linear ordered fieldR
.
Notation #
The following notation is available with open Quaternion
or open scoped Quaternion
.
ℍ[R, c₁, c₂, c₃]
:QuaternionAlgebra R c₁ c₂ c₃
ℍ[R, c₁, c₂]
:QuaternionAlgebra R c₁ 0 c₂
ℍ[R]
: quaternions overR
.
Implementation notes #
We define quaternions over any ring R
, not just ℝ
to be able to deal with, e.g., integer
or rational quaternions without using real numbers. In particular, all definitions in this file
are computable.
Tags #
quaternion
Quaternion algebra over a type with fixed coefficients where $i^2 = a + bi$ and $j^2 = c$,
denoted as ℍ[R,a,b]
.
Implemented as a structure with four fields: re
, imI
, imJ
, and imK
.
- re : R
Real part of a quaternion.
- imI : R
First imaginary part (i) of a quaternion.
- imJ : R
Second imaginary part (j) of a quaternion.
- imK : R
Third imaginary part (k) of a quaternion.
Instances For
The equivalence between a quaternion algebra over R
and R × R × R × R
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The equivalence between a quaternion algebra over R
and Fin 4 → R
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The imaginary part of a quaternion.
Note that unless c₂ = 0
, this definition is not particularly well-behaved;
for instance, QuaternionAlgebra.star_im
only says that the star of an imaginary quaternion
is imaginary under this condition.
Instances For
Coercion R → ℍ[R,c₁,c₂,c₃]
.
Equations
- ↑x = { re := x, imI := 0, imJ := 0, imK := 0 }
Instances For
Equations
- QuaternionAlgebra.instCoeTC = { coe := QuaternionAlgebra.coe }
Equations
- QuaternionAlgebra.instZero = { zero := { re := 0, imI := 0, imJ := 0, imK := 0 } }
Equations
- QuaternionAlgebra.instInhabited = { default := 0 }
Equations
- QuaternionAlgebra.instOne = { one := { re := 1, imI := 0, imJ := 0, imK := 0 } }
Multiplication is given by
1 * x = x * 1 = x
;i * i = c₁ + c₂ * i
;j * j = c₃
;i * j = k
,j * i = c₂ * j - k
;k * k = - c₁ * c₃
;i * k = c₁ * j + c₂ * k
,k * i = -c₁ * j
;j * k = c₂ * c₃ - c₃ * i
,k * j = c₃ * i
.
Equations
- One or more equations did not get rendered due to their size.
Equations
- QuaternionAlgebra.instAddCommGroup = Function.Injective.addCommGroup ⇑(QuaternionAlgebra.equivProd c₁ c₂ c₃) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
Equations
- QuaternionAlgebra.instRing = Ring.mk ⋯ SubNegMonoid.zsmul ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- QuaternionAlgebra.instAlgebra = Algebra.mk { toFun := fun (s : S) => ↑((algebraMap S R) s), map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ } ⋯ ⋯
QuaternionAlgebra.re
as a LinearMap
Equations
- QuaternionAlgebra.reₗ c₁ c₂ c₃ = { toFun := QuaternionAlgebra.re, map_add' := ⋯, map_smul' := ⋯ }
Instances For
QuaternionAlgebra.imI
as a LinearMap
Equations
- QuaternionAlgebra.imIₗ c₁ c₂ c₃ = { toFun := QuaternionAlgebra.imI, map_add' := ⋯, map_smul' := ⋯ }
Instances For
QuaternionAlgebra.imJ
as a LinearMap
Equations
- QuaternionAlgebra.imJₗ c₁ c₂ c₃ = { toFun := QuaternionAlgebra.imJ, map_add' := ⋯, map_smul' := ⋯ }
Instances For
QuaternionAlgebra.imK
as a LinearMap
Equations
- QuaternionAlgebra.imKₗ c₁ c₂ c₃ = { toFun := QuaternionAlgebra.imK, map_add' := ⋯, map_smul' := ⋯ }
Instances For
QuaternionAlgebra.equivTuple
as a linear equivalence.
Equations
- One or more equations did not get rendered due to their size.
Instances For
ℍ[R, c₁, c₂, c₃]
has a basis over R
given by 1
, i
, j
, and k
.
Equations
- QuaternionAlgebra.basisOneIJK c₁ c₂ c₃ = Basis.ofEquivFun (QuaternionAlgebra.linearEquivTuple c₁ c₂ c₃)
Instances For
There is a natural equivalence when swapping the first and third coefficients of a
quaternion algebra if c₂
is 0.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Quaternion conjugate.
Equations
A version of star_smul
for the special case when c₂ = 0
, without SMulCommClass S R R
.
Quaternion conjugate as an AlgEquiv
to the opposite ring.
Equations
- QuaternionAlgebra.starAe = { toFun := MulOpposite.op ∘ star, invFun := star ∘ MulOpposite.unop, left_inv := ⋯, right_inv := ⋯, map_mul' := ⋯, map_add' := ⋯, commutes' := ⋯ }
Instances For
Space of quaternions over a type, denoted as ℍ[R]
.
Implemented as a structure with four fields: re
, im_i
, im_j
, and im_k
.
Equations
- Quaternion R = QuaternionAlgebra R (-1) 0 (-1)
Instances For
Space of quaternions over a type, denoted as ℍ[R]
.
Implemented as a structure with four fields: re
, im_i
, im_j
, and im_k
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The equivalence between the quaternions over R
and R × R × R × R
.
Equations
- Quaternion.equivProd R = QuaternionAlgebra.equivProd (-1) 0 (-1)
Instances For
The equivalence between the quaternions over R
and Fin 4 → R
.
Equations
- Quaternion.equivTuple R = QuaternionAlgebra.equivTuple (-1) 0 (-1)
Instances For
Coercion R → ℍ[R]
.
Equations
Instances For
Equations
- Quaternion.instCoeTC = { coe := Quaternion.coe }
Equations
- Quaternion.instInhabited = inferInstanceAs (Inhabited (QuaternionAlgebra R (-1) 0 (-1)))
Equations
- Quaternion.instSMul = inferInstanceAs (SMul S (QuaternionAlgebra R (-1) 0 (-1)))
Equations
- Quaternion.algebra = inferInstanceAs (Algebra S (QuaternionAlgebra R (-1) 0 (-1)))
Square of the norm.
Equations
- Quaternion.normSq = { toFun := fun (a : Quaternion R) => (a * star a).re, map_zero' := ⋯, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Equations
- Quaternion.instInv = { inv := fun (a : Quaternion R) => (Quaternion.normSq a)⁻¹ • star a }
Equations
- Quaternion.instGroupWithZero = GroupWithZero.mk ⋯ zpowRec ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- Quaternion.instNNRatCast = { nnratCast := fun (q : ℚ≥0) => ↑↑q }
Equations
- Quaternion.instRatCast = { ratCast := fun (q : ℚ) => ↑↑q }
Equations
- Quaternion.instDivisionRing = DivisionRing.mk ⋯ GroupWithZero.zpow ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ (fun (x1 : ℚ≥0) (x2 : Quaternion R) => x1 • x2) ⋯ ⋯ (fun (x1 : ℚ) (x2 : Quaternion R) => x1 • x2) ⋯
The cardinality of a quaternion algebra, as a type.
Show the quaternion ⟨w, x, y, z⟩ as a string "{ re := w, imI := x, imJ := y, imK := z }".
For the typical case of quaternions over ℝ, each component will show as a Cauchy sequence due to the way Real numbers are represented.
Equations
- One or more equations did not get rendered due to their size.