Documentation

Mathlib.Algebra.Quaternion

Quaternions #

In this file we define quaternions ℍ[R] over a commutative ring R, and define some algebraic structures on ℍ[R].

Main definitions #

We also define the following algebraic structures on ℍ[R]:

Notation #

The following notation is available with open Quaternion or open scoped Quaternion.

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

structure QuaternionAlgebra (R : Type u_1) (a b c : R) :
Type u_1

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
    theorem QuaternionAlgebra.ext {R : Type u_1} {a b c : R} {x y : QuaternionAlgebra R a b c} (re : x.re = y.re) (imI : x.imI = y.imI) (imJ : x.imJ = y.imJ) (imK : x.imK = y.imK) :
    x = y

    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.

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

      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.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        def QuaternionAlgebra.equivProd {R : Type u_1} (c₁ c₂ c₃ : R) :
        QuaternionAlgebra R c₁ c₂ c₃ R × R × R × R

        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
          @[simp]
          theorem QuaternionAlgebra.equivProd_symm_apply_imI {R : Type u_1} (c₁ c₂ c₃ : R) (a : R × R × R × R) :
          ((equivProd c₁ c₂ c₃).symm a).imI = a.2.1
          @[simp]
          theorem QuaternionAlgebra.equivProd_apply {R : Type u_1} (c₁ c₂ c₃ : R) (a : QuaternionAlgebra R c₁ c₂ c₃) :
          (equivProd c₁ c₂ c₃) a = (a.re, a.imI, a.imJ, a.imK)
          @[simp]
          theorem QuaternionAlgebra.equivProd_symm_apply_imJ {R : Type u_1} (c₁ c₂ c₃ : R) (a : R × R × R × R) :
          ((equivProd c₁ c₂ c₃).symm a).imJ = a.2.2.1
          @[simp]
          theorem QuaternionAlgebra.equivProd_symm_apply_re {R : Type u_1} (c₁ c₂ c₃ : R) (a : R × R × R × R) :
          ((equivProd c₁ c₂ c₃).symm a).re = a.1
          @[simp]
          theorem QuaternionAlgebra.equivProd_symm_apply_imK {R : Type u_1} (c₁ c₂ c₃ : R) (a : R × R × R × R) :
          ((equivProd c₁ c₂ c₃).symm a).imK = a.2.2.2
          def QuaternionAlgebra.equivTuple {R : Type u_1} (c₁ c₂ c₃ : R) :
          QuaternionAlgebra R c₁ c₂ c₃ (Fin 4R)

          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
            @[simp]
            theorem QuaternionAlgebra.equivTuple_symm_apply {R : Type u_1} (c₁ c₂ c₃ : R) (a : Fin 4R) :
            (equivTuple c₁ c₂ c₃).symm a = { re := a 0, imI := a 1, imJ := a 2, imK := a 3 }
            @[simp]
            theorem QuaternionAlgebra.equivTuple_apply {R : Type u_1} (c₁ c₂ c₃ : R) (x : QuaternionAlgebra R c₁ c₂ c₃) :
            (equivTuple c₁ c₂ c₃) x = ![x.re, x.imI, x.imJ, x.imK]
            @[simp]
            theorem QuaternionAlgebra.mk.eta {R : Type u_1} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) :
            { re := a.re, imI := a.imI, imJ := a.imJ, imK := a.imK } = a
            instance QuaternionAlgebra.instSubsingleton {R : Type u_3} {c₁ c₂ c₃ : R} [Subsingleton R] :
            Subsingleton (QuaternionAlgebra R c₁ c₂ c₃)
            instance QuaternionAlgebra.instNontrivial {R : Type u_3} {c₁ c₂ c₃ : R} [Nontrivial R] :
            Nontrivial (QuaternionAlgebra R c₁ c₂ c₃)
            def QuaternionAlgebra.im {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] (x : QuaternionAlgebra R c₁ c₂ c₃) :
            QuaternionAlgebra R c₁ c₂ c₃

            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.

            Equations
            • x.im = { re := 0, imI := x.imI, imJ := x.imJ, imK := x.imK }
            Instances For
              @[simp]
              theorem QuaternionAlgebra.im_re {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [Zero R] :
              a.im.re = 0
              @[simp]
              theorem QuaternionAlgebra.im_imI {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [Zero R] :
              a.im.imI = a.imI
              @[simp]
              theorem QuaternionAlgebra.im_imJ {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [Zero R] :
              a.im.imJ = a.imJ
              @[simp]
              theorem QuaternionAlgebra.im_imK {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [Zero R] :
              a.im.imK = a.imK
              @[simp]
              theorem QuaternionAlgebra.im_idem {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [Zero R] :
              a.im.im = a.im
              def QuaternionAlgebra.coe {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] (x : R) :
              QuaternionAlgebra R c₁ c₂ c₃

              Coercion R → ℍ[R,c₁,c₂,c₃].

              Equations
              • x = { re := x, imI := 0, imJ := 0, imK := 0 }
              Instances For
                instance QuaternionAlgebra.instCoeTC {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                CoeTC R (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                @[simp]
                theorem QuaternionAlgebra.coe_re {R : Type u_3} {c₁ c₂ c₃ : R} (x : R) [Zero R] :
                (↑x).re = x
                @[simp]
                theorem QuaternionAlgebra.coe_imI {R : Type u_3} {c₁ c₂ c₃ : R} (x : R) [Zero R] :
                (↑x).imI = 0
                @[simp]
                theorem QuaternionAlgebra.coe_imJ {R : Type u_3} {c₁ c₂ c₃ : R} (x : R) [Zero R] :
                (↑x).imJ = 0
                @[simp]
                theorem QuaternionAlgebra.coe_imK {R : Type u_3} {c₁ c₂ c₃ : R} (x : R) [Zero R] :
                (↑x).imK = 0
                theorem QuaternionAlgebra.coe_injective {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                @[simp]
                theorem QuaternionAlgebra.coe_inj {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] {x y : R} :
                x = y x = y
                instance QuaternionAlgebra.instZero {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                Zero (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                theorem QuaternionAlgebra.zero_re {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                re 0 = 0
                theorem QuaternionAlgebra.zero_imI {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                imI 0 = 0
                theorem QuaternionAlgebra.zero_imJ {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                imJ 0 = 0
                theorem QuaternionAlgebra.zero_imK {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                imK 0 = 0
                theorem QuaternionAlgebra.zero_im {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                im 0 = 0
                @[simp]
                theorem QuaternionAlgebra.coe_zero {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                0 = 0
                instance QuaternionAlgebra.instInhabited {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] :
                Inhabited (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                instance QuaternionAlgebra.instOne {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] [One R] :
                One (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                theorem QuaternionAlgebra.one_re {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] [One R] :
                re 1 = 1
                theorem QuaternionAlgebra.one_imI {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] [One R] :
                imI 1 = 0
                theorem QuaternionAlgebra.one_imJ {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] [One R] :
                imJ 1 = 0
                theorem QuaternionAlgebra.one_imK {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] [One R] :
                imK 1 = 0
                theorem QuaternionAlgebra.one_im {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] [One R] :
                im 1 = 0
                @[simp]
                theorem QuaternionAlgebra.coe_one {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] [One R] :
                1 = 1
                instance QuaternionAlgebra.instAdd {R : Type u_3} {c₁ c₂ c₃ : R} [Add R] :
                Add (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                @[simp]
                theorem QuaternionAlgebra.add_re {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [Add R] :
                (a + b).re = a.re + b.re
                @[simp]
                theorem QuaternionAlgebra.add_imI {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [Add R] :
                (a + b).imI = a.imI + b.imI
                @[simp]
                theorem QuaternionAlgebra.add_imJ {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [Add R] :
                (a + b).imJ = a.imJ + b.imJ
                @[simp]
                theorem QuaternionAlgebra.add_imK {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [Add R] :
                (a + b).imK = a.imK + b.imK
                @[simp]
                theorem QuaternionAlgebra.mk_add_mk {R : Type u_3} {c₁ c₂ c₃ : R} [Add R] (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) :
                { re := a₁, imI := a₂, imJ := a₃, imK := a₄ } + { re := b₁, imI := b₂, imJ := b₃, imK := b₄ } = { re := a₁ + b₁, imI := a₂ + b₂, imJ := a₃ + b₃, imK := a₄ + b₄ }
                @[simp]
                theorem QuaternionAlgebra.add_im {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [AddZeroClass R] :
                (a + b).im = a.im + b.im
                @[simp]
                theorem QuaternionAlgebra.coe_add {R : Type u_3} {c₁ c₂ c₃ : R} (x y : R) [AddZeroClass R] :
                ↑(x + y) = x + y
                instance QuaternionAlgebra.instNeg {R : Type u_3} {c₁ c₂ c₃ : R} [Neg R] :
                Neg (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                @[simp]
                theorem QuaternionAlgebra.neg_re {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [Neg R] :
                (-a).re = -a.re
                @[simp]
                theorem QuaternionAlgebra.neg_imI {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [Neg R] :
                (-a).imI = -a.imI
                @[simp]
                theorem QuaternionAlgebra.neg_imJ {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [Neg R] :
                (-a).imJ = -a.imJ
                @[simp]
                theorem QuaternionAlgebra.neg_imK {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [Neg R] :
                (-a).imK = -a.imK
                @[simp]
                theorem QuaternionAlgebra.neg_mk {R : Type u_3} {c₁ c₂ c₃ : R} [Neg R] (a₁ a₂ a₃ a₄ : R) :
                -{ re := a₁, imI := a₂, imJ := a₃, imK := a₄ } = { re := -a₁, imI := -a₂, imJ := -a₃, imK := -a₄ }
                @[simp]
                theorem QuaternionAlgebra.neg_im {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [AddGroup R] :
                (-a).im = -a.im
                @[simp]
                theorem QuaternionAlgebra.coe_neg {R : Type u_3} {c₁ c₂ c₃ : R} (x : R) [AddGroup R] :
                ↑(-x) = -x
                instance QuaternionAlgebra.instSub {R : Type u_3} {c₁ c₂ c₃ : R} [AddGroup R] :
                Sub (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                @[simp]
                theorem QuaternionAlgebra.sub_re {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [AddGroup R] :
                (a - b).re = a.re - b.re
                @[simp]
                theorem QuaternionAlgebra.sub_imI {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [AddGroup R] :
                (a - b).imI = a.imI - b.imI
                @[simp]
                theorem QuaternionAlgebra.sub_imJ {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [AddGroup R] :
                (a - b).imJ = a.imJ - b.imJ
                @[simp]
                theorem QuaternionAlgebra.sub_imK {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [AddGroup R] :
                (a - b).imK = a.imK - b.imK
                @[simp]
                theorem QuaternionAlgebra.sub_im {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [AddGroup R] :
                (a - b).im = a.im - b.im
                @[simp]
                theorem QuaternionAlgebra.mk_sub_mk {R : Type u_3} {c₁ c₂ c₃ : R} [AddGroup R] (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) :
                { re := a₁, imI := a₂, imJ := a₃, imK := a₄ } - { re := b₁, imI := b₂, imJ := b₃, imK := b₄ } = { re := a₁ - b₁, imI := a₂ - b₂, imJ := a₃ - b₃, imK := a₄ - b₄ }
                @[simp]
                theorem QuaternionAlgebra.coe_im {R : Type u_3} {c₁ c₂ c₃ : R} (x : R) [AddGroup R] :
                (↑x).im = 0
                @[simp]
                theorem QuaternionAlgebra.re_add_im {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [AddGroup R] :
                a.re + a.im = a
                @[simp]
                theorem QuaternionAlgebra.sub_self_im {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [AddGroup R] :
                a - a.im = a.re
                @[simp]
                theorem QuaternionAlgebra.sub_self_re {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [AddGroup R] :
                a - a.re = a.im
                instance QuaternionAlgebra.instMul {R : Type u_3} {c₁ c₂ c₃ : R} [Ring R] :
                Mul (QuaternionAlgebra R c₁ c₂ c₃)

                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.
                @[simp]
                theorem QuaternionAlgebra.mul_re {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [Ring R] :
                (a * b).re = a.re * b.re + c₁ * a.imI * b.imI + c₃ * a.imJ * b.imJ + c₂ * c₃ * a.imJ * b.imK - c₁ * c₃ * a.imK * b.imK
                @[simp]
                theorem QuaternionAlgebra.mul_imI {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [Ring R] :
                (a * b).imI = a.re * b.imI + a.imI * b.re + c₂ * a.imI * b.imI - c₃ * a.imJ * b.imK + c₃ * a.imK * b.imJ
                @[simp]
                theorem QuaternionAlgebra.mul_imJ {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [Ring R] :
                (a * b).imJ = a.re * b.imJ + c₁ * a.imI * b.imK + a.imJ * b.re + c₂ * a.imJ * b.imI - c₁ * a.imK * b.imI
                @[simp]
                theorem QuaternionAlgebra.mul_imK {R : Type u_3} {c₁ c₂ c₃ : R} (a b : QuaternionAlgebra R c₁ c₂ c₃) [Ring R] :
                (a * b).imK = a.re * b.imK + a.imI * b.imJ + c₂ * a.imI * b.imK - a.imJ * b.imI + a.imK * b.re
                @[simp]
                theorem QuaternionAlgebra.mk_mul_mk {R : Type u_3} {c₁ c₂ c₃ : R} [Ring R] (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) :
                { re := a₁, imI := a₂, imJ := a₃, imK := a₄ } * { re := b₁, imI := b₂, imJ := b₃, imK := b₄ } = { re := a₁ * b₁ + c₁ * a₂ * b₂ + c₃ * a₃ * b₃ + c₂ * c₃ * a₃ * b₄ - c₁ * c₃ * a₄ * b₄, imI := a₁ * b₂ + a₂ * b₁ + c₂ * a₂ * b₂ - c₃ * a₃ * b₄ + c₃ * a₄ * b₃, imJ := a₁ * b₃ + c₁ * a₂ * b₄ + a₃ * b₁ + c₂ * a₃ * b₂ - c₁ * a₄ * b₂, imK := a₁ * b₄ + a₂ * b₃ + c₂ * a₂ * b₄ - a₃ * b₂ + a₄ * b₁ }
                instance QuaternionAlgebra.instSMul {S : Type u_1} {R : Type u_3} {c₁ c₂ c₃ : R} [SMul S R] :
                SMul S (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                instance QuaternionAlgebra.instIsScalarTower {S : Type u_1} {T : Type u_2} {R : Type u_3} {c₁ c₂ c₃ : R} [SMul S R] [SMul T R] [SMul S T] [IsScalarTower S T R] :
                IsScalarTower S T (QuaternionAlgebra R c₁ c₂ c₃)
                instance QuaternionAlgebra.instSMulCommClass {S : Type u_1} {T : Type u_2} {R : Type u_3} {c₁ c₂ c₃ : R} [SMul S R] [SMul T R] [SMulCommClass S T R] :
                SMulCommClass S T (QuaternionAlgebra R c₁ c₂ c₃)
                @[simp]
                theorem QuaternionAlgebra.smul_re {S : Type u_1} {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [SMul S R] (s : S) :
                (s a).re = s a.re
                @[simp]
                theorem QuaternionAlgebra.smul_imI {S : Type u_1} {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [SMul S R] (s : S) :
                (s a).imI = s a.imI
                @[simp]
                theorem QuaternionAlgebra.smul_imJ {S : Type u_1} {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [SMul S R] (s : S) :
                (s a).imJ = s a.imJ
                @[simp]
                theorem QuaternionAlgebra.smul_imK {S : Type u_1} {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [SMul S R] (s : S) :
                (s a).imK = s a.imK
                @[simp]
                theorem QuaternionAlgebra.smul_im {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) {S : Type u_4} [CommRing R] [SMulZeroClass S R] (s : S) :
                (s a).im = s a.im
                @[simp]
                theorem QuaternionAlgebra.smul_mk {S : Type u_1} {R : Type u_3} {c₁ c₂ c₃ : R} [SMul S R] (s : S) (re im_i im_j im_k : R) :
                s { re := re, imI := im_i, imJ := im_j, imK := im_k } = { re := s re, imI := s im_i, imJ := s im_j, imK := s im_k }
                @[simp]
                theorem QuaternionAlgebra.coe_smul {S : Type u_1} {R : Type u_3} {c₁ c₂ c₃ : R} [Zero R] [SMulZeroClass S R] (s : S) (r : R) :
                ↑(s r) = s r
                instance QuaternionAlgebra.instAddCommGroup {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroup R] :
                AddCommGroup (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                @[simp]
                theorem QuaternionAlgebra.natCast_re {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) :
                (↑n).re = n
                @[simp]
                theorem QuaternionAlgebra.natCast_imI {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) :
                (↑n).imI = 0
                @[simp]
                theorem QuaternionAlgebra.natCast_imJ {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) :
                (↑n).imJ = 0
                @[simp]
                theorem QuaternionAlgebra.natCast_imK {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) :
                (↑n).imK = 0
                @[simp]
                theorem QuaternionAlgebra.natCast_im {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) :
                (↑n).im = 0
                theorem QuaternionAlgebra.coe_natCast {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) :
                n = n
                @[simp]
                theorem QuaternionAlgebra.intCast_re {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (z : ) :
                (↑z).re = z
                theorem QuaternionAlgebra.ofNat_re {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) [n.AtLeastTwo] :
                theorem QuaternionAlgebra.ofNat_imI {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) [n.AtLeastTwo] :
                theorem QuaternionAlgebra.ofNat_imJ {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) [n.AtLeastTwo] :
                theorem QuaternionAlgebra.ofNat_imK {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) [n.AtLeastTwo] :
                theorem QuaternionAlgebra.ofNat_im {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (n : ) [n.AtLeastTwo] :
                @[simp]
                theorem QuaternionAlgebra.intCast_imI {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (z : ) :
                (↑z).imI = 0
                @[simp]
                theorem QuaternionAlgebra.intCast_imJ {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (z : ) :
                (↑z).imJ = 0
                @[simp]
                theorem QuaternionAlgebra.intCast_imK {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (z : ) :
                (↑z).imK = 0
                @[simp]
                theorem QuaternionAlgebra.intCast_im {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (z : ) :
                (↑z).im = 0
                theorem QuaternionAlgebra.coe_intCast {R : Type u_3} {c₁ c₂ c₃ : R} [AddCommGroupWithOne R] (z : ) :
                z = z
                instance QuaternionAlgebra.instRing {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] :
                Ring (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                @[simp]
                theorem QuaternionAlgebra.coe_mul {R : Type u_3} {c₁ c₂ c₃ : R} (x y : R) [CommRing R] :
                ↑(x * y) = x * y
                @[simp]
                theorem QuaternionAlgebra.coe_ofNat {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] {n : } [n.AtLeastTwo] :
                instance QuaternionAlgebra.instAlgebra {S : Type u_1} {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] [CommSemiring S] [Algebra S R] :
                Algebra S (QuaternionAlgebra R c₁ c₂ c₃)
                Equations
                theorem QuaternionAlgebra.algebraMap_eq {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] (r : R) :
                (algebraMap R (QuaternionAlgebra R c₁ c₂ c₃)) r = { re := r, imI := 0, imJ := 0, imK := 0 }
                theorem QuaternionAlgebra.algebraMap_injective {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] :
                def QuaternionAlgebra.reₗ {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                QuaternionAlgebra R c₁ c₂ c₃ →ₗ[R] R

                QuaternionAlgebra.re as a LinearMap

                Equations
                Instances For
                  @[simp]
                  theorem QuaternionAlgebra.reₗ_apply {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] (self : QuaternionAlgebra R c₁ c₂ c₃) :
                  (reₗ c₁ c₂ c₃) self = self.re
                  def QuaternionAlgebra.imIₗ {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                  QuaternionAlgebra R c₁ c₂ c₃ →ₗ[R] R

                  QuaternionAlgebra.imI as a LinearMap

                  Equations
                  Instances For
                    @[simp]
                    theorem QuaternionAlgebra.imIₗ_apply {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] (self : QuaternionAlgebra R c₁ c₂ c₃) :
                    (imIₗ c₁ c₂ c₃) self = self.imI
                    def QuaternionAlgebra.imJₗ {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                    QuaternionAlgebra R c₁ c₂ c₃ →ₗ[R] R

                    QuaternionAlgebra.imJ as a LinearMap

                    Equations
                    Instances For
                      @[simp]
                      theorem QuaternionAlgebra.imJₗ_apply {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] (self : QuaternionAlgebra R c₁ c₂ c₃) :
                      (imJₗ c₁ c₂ c₃) self = self.imJ
                      def QuaternionAlgebra.imKₗ {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                      QuaternionAlgebra R c₁ c₂ c₃ →ₗ[R] R

                      QuaternionAlgebra.imK as a LinearMap

                      Equations
                      Instances For
                        @[simp]
                        theorem QuaternionAlgebra.imKₗ_apply {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] (self : QuaternionAlgebra R c₁ c₂ c₃) :
                        (imKₗ c₁ c₂ c₃) self = self.imK
                        def QuaternionAlgebra.linearEquivTuple {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                        QuaternionAlgebra R c₁ c₂ c₃ ≃ₗ[R] Fin 4R

                        QuaternionAlgebra.equivTuple as a linear equivalence.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem QuaternionAlgebra.coe_linearEquivTuple {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                          (linearEquivTuple c₁ c₂ c₃) = (equivTuple c₁ c₂ c₃)
                          @[simp]
                          theorem QuaternionAlgebra.coe_linearEquivTuple_symm {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                          (linearEquivTuple c₁ c₂ c₃).symm = (equivTuple c₁ c₂ c₃).symm
                          noncomputable def QuaternionAlgebra.basisOneIJK {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                          Basis (Fin 4) R (QuaternionAlgebra R c₁ c₂ c₃)

                          ℍ[R, c₁, c₂, c₃] has a basis over R given by 1, i, j, and k.

                          Equations
                          Instances For
                            @[simp]
                            theorem QuaternionAlgebra.coe_basisOneIJK_repr {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] (q : QuaternionAlgebra R c₁ c₂ c₃) :
                            ((basisOneIJK c₁ c₂ c₃).repr q) = ![q.re, q.imI, q.imJ, q.imK]
                            instance QuaternionAlgebra.instFinite {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                            Module.Finite R (QuaternionAlgebra R c₁ c₂ c₃)
                            instance QuaternionAlgebra.instFree {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] :
                            Module.Free R (QuaternionAlgebra R c₁ c₂ c₃)
                            theorem QuaternionAlgebra.rank_eq_four {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] [StrongRankCondition R] :
                            Module.rank R (QuaternionAlgebra R c₁ c₂ c₃) = 4
                            theorem QuaternionAlgebra.finrank_eq_four {R : Type u_3} (c₁ c₂ c₃ : R) [CommRing R] [StrongRankCondition R] :
                            Module.finrank R (QuaternionAlgebra R c₁ c₂ c₃) = 4
                            def QuaternionAlgebra.swapEquiv {R : Type u_3} (c₁ c₃ : R) [CommRing R] :
                            QuaternionAlgebra R c₁ 0 c₃ ≃ₐ[R] QuaternionAlgebra R c₃ 0 c₁

                            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
                              @[simp]
                              theorem QuaternionAlgebra.swapEquiv_apply_imK {R : Type u_3} (c₁ c₃ : R) [CommRing R] (t : QuaternionAlgebra R c₁ 0 c₃) :
                              ((swapEquiv c₁ c₃) t).imK = -t.imK
                              @[simp]
                              theorem QuaternionAlgebra.swapEquiv_symm_apply_imJ {R : Type u_3} (c₁ c₃ : R) [CommRing R] (t : QuaternionAlgebra R c₃ 0 c₁) :
                              ((swapEquiv c₁ c₃).symm t).imJ = t.imI
                              @[simp]
                              theorem QuaternionAlgebra.swapEquiv_apply_imJ {R : Type u_3} (c₁ c₃ : R) [CommRing R] (t : QuaternionAlgebra R c₁ 0 c₃) :
                              ((swapEquiv c₁ c₃) t).imJ = t.imI
                              @[simp]
                              theorem QuaternionAlgebra.swapEquiv_symm_apply_imI {R : Type u_3} (c₁ c₃ : R) [CommRing R] (t : QuaternionAlgebra R c₃ 0 c₁) :
                              ((swapEquiv c₁ c₃).symm t).imI = t.imJ
                              @[simp]
                              theorem QuaternionAlgebra.swapEquiv_symm_apply_imK {R : Type u_3} (c₁ c₃ : R) [CommRing R] (t : QuaternionAlgebra R c₃ 0 c₁) :
                              ((swapEquiv c₁ c₃).symm t).imK = -t.imK
                              @[simp]
                              theorem QuaternionAlgebra.swapEquiv_apply_re {R : Type u_3} (c₁ c₃ : R) [CommRing R] (t : QuaternionAlgebra R c₁ 0 c₃) :
                              ((swapEquiv c₁ c₃) t).re = t.re
                              @[simp]
                              theorem QuaternionAlgebra.swapEquiv_symm_apply_re {R : Type u_3} (c₁ c₃ : R) [CommRing R] (t : QuaternionAlgebra R c₃ 0 c₁) :
                              ((swapEquiv c₁ c₃).symm t).re = t.re
                              @[simp]
                              theorem QuaternionAlgebra.swapEquiv_apply_imI {R : Type u_3} (c₁ c₃ : R) [CommRing R] (t : QuaternionAlgebra R c₁ 0 c₃) :
                              ((swapEquiv c₁ c₃) t).imI = t.imJ
                              @[simp]
                              theorem QuaternionAlgebra.coe_sub {R : Type u_3} {c₁ c₂ c₃ : R} (x y : R) [CommRing R] :
                              ↑(x - y) = x - y
                              @[simp]
                              theorem QuaternionAlgebra.coe_pow {R : Type u_3} {c₁ c₂ c₃ : R} (x : R) [CommRing R] (n : ) :
                              ↑(x ^ n) = x ^ n
                              theorem QuaternionAlgebra.coe_commutes {R : Type u_3} {c₁ c₂ c₃ : R} (r : R) (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              r * a = a * r
                              theorem QuaternionAlgebra.coe_commute {R : Type u_3} {c₁ c₂ c₃ : R} (r : R) (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              Commute (↑r) a
                              theorem QuaternionAlgebra.coe_mul_eq_smul {R : Type u_3} {c₁ c₂ c₃ : R} (r : R) (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              r * a = r a
                              theorem QuaternionAlgebra.mul_coe_eq_smul {R : Type u_3} {c₁ c₂ c₃ : R} (r : R) (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              a * r = r a
                              @[simp]
                              theorem QuaternionAlgebra.coe_algebraMap {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] :
                              (algebraMap R (QuaternionAlgebra R c₁ c₂ c₃)) = coe
                              theorem QuaternionAlgebra.smul_coe {R : Type u_3} {c₁ c₂ c₃ : R} (x y : R) [CommRing R] :
                              x y = ↑(x * y)
                              instance QuaternionAlgebra.instStarQuaternionAlgebra {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] :
                              Star (QuaternionAlgebra R c₁ c₂ c₃)

                              Quaternion conjugate.

                              Equations
                              @[simp]
                              theorem QuaternionAlgebra.re_star {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              (star a).re = a.re + c₂ * a.imI
                              @[simp]
                              theorem QuaternionAlgebra.imI_star {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              (star a).imI = -a.imI
                              @[simp]
                              theorem QuaternionAlgebra.imJ_star {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              (star a).imJ = -a.imJ
                              @[simp]
                              theorem QuaternionAlgebra.imK_star {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              (star a).imK = -a.imK
                              @[simp]
                              theorem QuaternionAlgebra.im_star {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              (star a).im = -a.im
                              @[simp]
                              theorem QuaternionAlgebra.star_mk {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] (a₁ a₂ a₃ a₄ : R) :
                              star { re := a₁, imI := a₂, imJ := a₃, imK := a₄ } = { re := a₁ + c₂ * a₂, imI := -a₂, imJ := -a₃, imK := -a₄ }
                              instance QuaternionAlgebra.instStarRing {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] :
                              StarRing (QuaternionAlgebra R c₁ c₂ c₃)
                              Equations
                              theorem QuaternionAlgebra.self_add_star' {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              a + star a = ↑(2 * a.re + c₂ * a.imI)
                              theorem QuaternionAlgebra.self_add_star {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              a + star a = 2 * a.re + c₂ * a.imI
                              theorem QuaternionAlgebra.star_add_self' {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              star a + a = ↑(2 * a.re + c₂ * a.imI)
                              theorem QuaternionAlgebra.star_add_self {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              star a + a = 2 * a.re + c₂ * a.imI
                              theorem QuaternionAlgebra.star_eq_two_re_sub {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              star a = ↑(2 * a.re + c₂ * a.imI) - a
                              theorem QuaternionAlgebra.comm {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] (r : R) (x : QuaternionAlgebra R c₁ c₂ c₃) :
                              r * x = x * r
                              instance QuaternionAlgebra.instIsStarNormal {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              @[simp]
                              theorem QuaternionAlgebra.star_coe {R : Type u_3} {c₁ c₂ c₃ : R} (x : R) [CommRing R] :
                              star x = x
                              @[simp]
                              theorem QuaternionAlgebra.star_im {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              star a.im = -a.im + c₂ * a.imI
                              @[simp]
                              theorem QuaternionAlgebra.star_smul {S : Type u_1} {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] [Monoid S] [DistribMulAction S R] [SMulCommClass S R R] (s : S) (a : QuaternionAlgebra R c₁ c₂ c₃) :
                              star (s a) = s star a
                              theorem QuaternionAlgebra.star_smul' {S : Type u_1} {R : Type u_3} {c₁ c₃ : R} [CommRing R] [Monoid S] [DistribMulAction S R] (s : S) (a : QuaternionAlgebra R c₁ 0 c₃) :
                              star (s a) = s star a

                              A version of star_smul for the special case when c₂ = 0, without SMulCommClass S R R.

                              theorem QuaternionAlgebra.eq_re_of_eq_coe {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] {a : QuaternionAlgebra R c₁ c₂ c₃} {x : R} (h : a = x) :
                              a = a.re
                              theorem QuaternionAlgebra.eq_re_iff_mem_range_coe {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] {a : QuaternionAlgebra R c₁ c₂ c₃} :
                              a = a.re a Set.range coe
                              @[simp]
                              theorem QuaternionAlgebra.star_eq_self {R : Type u_3} {c₃ : R} [CommRing R] [NoZeroDivisors R] [CharZero R] {c₁ c₂ : R} {a : QuaternionAlgebra R c₁ c₂ c₃} :
                              star a = a a = a.re
                              theorem QuaternionAlgebra.star_eq_neg {R : Type u_3} {c₃ : R} [CommRing R] [NoZeroDivisors R] [CharZero R] {c₁ : R} {a : QuaternionAlgebra R c₁ 0 c₃} :
                              star a = -a a.re = 0
                              theorem QuaternionAlgebra.star_mul_eq_coe {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              star a * a = (star a * a).re
                              theorem QuaternionAlgebra.mul_star_eq_coe {R : Type u_3} {c₁ c₂ c₃ : R} (a : QuaternionAlgebra R c₁ c₂ c₃) [CommRing R] :
                              a * star a = (a * star a).re
                              def QuaternionAlgebra.starAe {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] :
                              QuaternionAlgebra R c₁ c₂ c₃ ≃ₐ[R] (QuaternionAlgebra R c₁ c₂ c₃)ᵐᵒᵖ

                              Quaternion conjugate as an AlgEquiv to the opposite ring.

                              Equations
                              Instances For
                                @[simp]
                                theorem QuaternionAlgebra.coe_starAe {R : Type u_3} {c₁ c₂ c₃ : R} [CommRing R] :
                                def Quaternion (R : Type u_1) [Zero R] [One R] [Neg R] :
                                Type u_1

                                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
                                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
                                    def Quaternion.equivProd (R : Type u_1) [Zero R] [One R] [Neg R] :
                                    Quaternion R R × R × R × R

                                    The equivalence between the quaternions over R and R × R × R × R.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem Quaternion.equivProd_symm_apply_imI (R : Type u_1) [Zero R] [One R] [Neg R] (a : R × R × R × R) :
                                      ((equivProd R).symm a).imI = a.2.1
                                      @[simp]
                                      theorem Quaternion.equivProd_apply (R : Type u_1) [Zero R] [One R] [Neg R] (a : QuaternionAlgebra R (-1) 0 (-1)) :
                                      (equivProd R) a = (a.re, a.imI, a.imJ, a.imK)
                                      @[simp]
                                      theorem Quaternion.equivProd_symm_apply_re (R : Type u_1) [Zero R] [One R] [Neg R] (a : R × R × R × R) :
                                      ((equivProd R).symm a).re = a.1
                                      @[simp]
                                      theorem Quaternion.equivProd_symm_apply_imK (R : Type u_1) [Zero R] [One R] [Neg R] (a : R × R × R × R) :
                                      ((equivProd R).symm a).imK = a.2.2.2
                                      @[simp]
                                      theorem Quaternion.equivProd_symm_apply_imJ (R : Type u_1) [Zero R] [One R] [Neg R] (a : R × R × R × R) :
                                      ((equivProd R).symm a).imJ = a.2.2.1
                                      def Quaternion.equivTuple (R : Type u_1) [Zero R] [One R] [Neg R] :
                                      Quaternion R (Fin 4R)

                                      The equivalence between the quaternions over R and Fin 4 → R.

                                      Equations
                                      Instances For
                                        @[simp]
                                        theorem Quaternion.equivTuple_symm_apply (R : Type u_1) [Zero R] [One R] [Neg R] (a : Fin 4R) :
                                        (equivTuple R).symm a = { re := a 0, imI := a 1, imJ := a 2, imK := a 3 }
                                        @[simp]
                                        theorem Quaternion.equivTuple_apply (R : Type u_1) [Zero R] [One R] [Neg R] (x : Quaternion R) :
                                        (equivTuple R) x = ![x.re, x.imI, x.imJ, x.imK]
                                        def Quaternion.coe {R : Type u_3} [CommRing R] :
                                        RQuaternion R

                                        Coercion R → ℍ[R].

                                        Equations
                                        Instances For
                                          instance Quaternion.instSMul {S : Type u_1} {R : Type u_3} [CommRing R] [SMul S R] :
                                          Equations
                                          instance Quaternion.instIsScalarTower {S : Type u_1} {T : Type u_2} {R : Type u_3} [CommRing R] [SMul S T] [SMul S R] [SMul T R] [IsScalarTower S T R] :
                                          instance Quaternion.instSMulCommClass {S : Type u_1} {T : Type u_2} {R : Type u_3} [CommRing R] [SMul S R] [SMul T R] [SMulCommClass S T R] :
                                          instance Quaternion.algebra {S : Type u_1} {R : Type u_3} [CommRing R] [CommSemiring S] [Algebra S R] :
                                          Equations
                                          theorem Quaternion.ext {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                          a.re = b.rea.imI = b.imIa.imJ = b.imJa.imK = b.imKa = b
                                          def Quaternion.im {R : Type u_3} [CommRing R] (x : Quaternion R) :

                                          The imaginary part of a quaternion.

                                          Equations
                                          Instances For
                                            @[simp]
                                            theorem Quaternion.im_re {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a.im.re = 0
                                            @[simp]
                                            theorem Quaternion.im_imI {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a.im.imI = a.imI
                                            @[simp]
                                            theorem Quaternion.im_imJ {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a.im.imJ = a.imJ
                                            @[simp]
                                            theorem Quaternion.im_imK {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a.im.imK = a.imK
                                            @[simp]
                                            theorem Quaternion.im_idem {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a.im.im = a.im
                                            @[simp]
                                            theorem Quaternion.re_add_im {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a.re + a.im = a
                                            @[simp]
                                            theorem Quaternion.sub_self_im {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a - a.im = a.re
                                            @[simp]
                                            theorem Quaternion.sub_self_re {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a - a.re = a.im
                                            @[simp]
                                            theorem Quaternion.coe_re {R : Type u_3} [CommRing R] (x : R) :
                                            (↑x).re = x
                                            @[simp]
                                            theorem Quaternion.coe_imI {R : Type u_3} [CommRing R] (x : R) :
                                            (↑x).imI = 0
                                            @[simp]
                                            theorem Quaternion.coe_imJ {R : Type u_3} [CommRing R] (x : R) :
                                            (↑x).imJ = 0
                                            @[simp]
                                            theorem Quaternion.coe_imK {R : Type u_3} [CommRing R] (x : R) :
                                            (↑x).imK = 0
                                            @[simp]
                                            theorem Quaternion.coe_im {R : Type u_3} [CommRing R] (x : R) :
                                            (↑x).im = 0
                                            theorem Quaternion.zero_im {R : Type u_3} [CommRing R] :
                                            im 0 = 0
                                            @[simp]
                                            theorem Quaternion.coe_zero {R : Type u_3} [CommRing R] :
                                            0 = 0
                                            theorem Quaternion.one_im {R : Type u_3} [CommRing R] :
                                            im 1 = 0
                                            @[simp]
                                            theorem Quaternion.coe_one {R : Type u_3} [CommRing R] :
                                            1 = 1
                                            @[simp]
                                            theorem Quaternion.add_re {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a + b).re = a.re + b.re
                                            @[simp]
                                            theorem Quaternion.add_imI {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a + b).imI = a.imI + b.imI
                                            @[simp]
                                            theorem Quaternion.add_imJ {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a + b).imJ = a.imJ + b.imJ
                                            @[simp]
                                            theorem Quaternion.add_imK {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a + b).imK = a.imK + b.imK
                                            @[simp]
                                            theorem Quaternion.add_im {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a + b).im = a.im + b.im
                                            @[simp]
                                            theorem Quaternion.coe_add {R : Type u_3} [CommRing R] (x y : R) :
                                            ↑(x + y) = x + y
                                            @[simp]
                                            theorem Quaternion.neg_re {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (-a).re = -a.re
                                            @[simp]
                                            theorem Quaternion.neg_imI {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (-a).imI = -a.imI
                                            @[simp]
                                            theorem Quaternion.neg_imJ {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (-a).imJ = -a.imJ
                                            @[simp]
                                            theorem Quaternion.neg_imK {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (-a).imK = -a.imK
                                            @[simp]
                                            theorem Quaternion.neg_im {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (-a).im = -a.im
                                            @[simp]
                                            theorem Quaternion.coe_neg {R : Type u_3} [CommRing R] (x : R) :
                                            ↑(-x) = -x
                                            @[simp]
                                            theorem Quaternion.sub_re {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a - b).re = a.re - b.re
                                            @[simp]
                                            theorem Quaternion.sub_imI {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a - b).imI = a.imI - b.imI
                                            @[simp]
                                            theorem Quaternion.sub_imJ {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a - b).imJ = a.imJ - b.imJ
                                            @[simp]
                                            theorem Quaternion.sub_imK {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a - b).imK = a.imK - b.imK
                                            @[simp]
                                            theorem Quaternion.sub_im {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a - b).im = a.im - b.im
                                            @[simp]
                                            theorem Quaternion.coe_sub {R : Type u_3} [CommRing R] (x y : R) :
                                            ↑(x - y) = x - y
                                            @[simp]
                                            theorem Quaternion.mul_re {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a * b).re = a.re * b.re - a.imI * b.imI - a.imJ * b.imJ - a.imK * b.imK
                                            @[simp]
                                            theorem Quaternion.mul_imI {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a * b).imI = a.re * b.imI + a.imI * b.re + a.imJ * b.imK - a.imK * b.imJ
                                            @[simp]
                                            theorem Quaternion.mul_imJ {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a * b).imJ = a.re * b.imJ - a.imI * b.imK + a.imJ * b.re + a.imK * b.imI
                                            @[simp]
                                            theorem Quaternion.mul_imK {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                            (a * b).imK = a.re * b.imK + a.imI * b.imJ - a.imJ * b.imI + a.imK * b.re
                                            @[simp]
                                            theorem Quaternion.coe_mul {R : Type u_3} [CommRing R] (x y : R) :
                                            ↑(x * y) = x * y
                                            @[simp]
                                            theorem Quaternion.coe_pow {R : Type u_3} [CommRing R] (x : R) (n : ) :
                                            ↑(x ^ n) = x ^ n
                                            @[simp]
                                            theorem Quaternion.natCast_re {R : Type u_3} [CommRing R] (n : ) :
                                            (↑n).re = n
                                            @[simp]
                                            theorem Quaternion.natCast_imI {R : Type u_3} [CommRing R] (n : ) :
                                            (↑n).imI = 0
                                            @[simp]
                                            theorem Quaternion.natCast_imJ {R : Type u_3} [CommRing R] (n : ) :
                                            (↑n).imJ = 0
                                            @[simp]
                                            theorem Quaternion.natCast_imK {R : Type u_3} [CommRing R] (n : ) :
                                            (↑n).imK = 0
                                            @[simp]
                                            theorem Quaternion.natCast_im {R : Type u_3} [CommRing R] (n : ) :
                                            (↑n).im = 0
                                            theorem Quaternion.coe_natCast {R : Type u_3} [CommRing R] (n : ) :
                                            n = n
                                            @[simp]
                                            theorem Quaternion.intCast_re {R : Type u_3} [CommRing R] (z : ) :
                                            (↑z).re = z
                                            @[simp]
                                            theorem Quaternion.intCast_imI {R : Type u_3} [CommRing R] (z : ) :
                                            (↑z).imI = 0
                                            @[simp]
                                            theorem Quaternion.intCast_imJ {R : Type u_3} [CommRing R] (z : ) :
                                            (↑z).imJ = 0
                                            @[simp]
                                            theorem Quaternion.intCast_imK {R : Type u_3} [CommRing R] (z : ) :
                                            (↑z).imK = 0
                                            @[simp]
                                            theorem Quaternion.intCast_im {R : Type u_3} [CommRing R] (z : ) :
                                            (↑z).im = 0
                                            theorem Quaternion.coe_intCast {R : Type u_3} [CommRing R] (z : ) :
                                            z = z
                                            @[simp]
                                            theorem Quaternion.coe_inj {R : Type u_3} [CommRing R] {x y : R} :
                                            x = y x = y
                                            @[simp]
                                            theorem Quaternion.smul_re {S : Type u_1} {R : Type u_3} [CommRing R] (a : Quaternion R) [SMul S R] (s : S) :
                                            (s a).re = s a.re
                                            @[simp]
                                            theorem Quaternion.smul_imI {S : Type u_1} {R : Type u_3} [CommRing R] (a : Quaternion R) [SMul S R] (s : S) :
                                            (s a).imI = s a.imI
                                            @[simp]
                                            theorem Quaternion.smul_imJ {S : Type u_1} {R : Type u_3} [CommRing R] (a : Quaternion R) [SMul S R] (s : S) :
                                            (s a).imJ = s a.imJ
                                            @[simp]
                                            theorem Quaternion.smul_imK {S : Type u_1} {R : Type u_3} [CommRing R] (a : Quaternion R) [SMul S R] (s : S) :
                                            (s a).imK = s a.imK
                                            @[simp]
                                            theorem Quaternion.smul_im {S : Type u_1} {R : Type u_3} [CommRing R] (a : Quaternion R) [SMulZeroClass S R] (s : S) :
                                            (s a).im = s a.im
                                            @[simp]
                                            theorem Quaternion.coe_smul {S : Type u_1} {R : Type u_3} [CommRing R] [SMulZeroClass S R] (s : S) (r : R) :
                                            ↑(s r) = s r
                                            theorem Quaternion.coe_commutes {R : Type u_3} [CommRing R] (r : R) (a : Quaternion R) :
                                            r * a = a * r
                                            theorem Quaternion.coe_commute {R : Type u_3} [CommRing R] (r : R) (a : Quaternion R) :
                                            Commute (↑r) a
                                            theorem Quaternion.coe_mul_eq_smul {R : Type u_3} [CommRing R] (r : R) (a : Quaternion R) :
                                            r * a = r a
                                            theorem Quaternion.mul_coe_eq_smul {R : Type u_3} [CommRing R] (r : R) (a : Quaternion R) :
                                            a * r = r a
                                            @[simp]
                                            theorem Quaternion.smul_coe {R : Type u_3} [CommRing R] (x y : R) :
                                            x y = ↑(x * y)
                                            @[simp]
                                            theorem Quaternion.star_re {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (star a).re = a.re
                                            @[simp]
                                            theorem Quaternion.star_imI {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (star a).imI = -a.imI
                                            @[simp]
                                            theorem Quaternion.star_imJ {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (star a).imJ = -a.imJ
                                            @[simp]
                                            theorem Quaternion.star_imK {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (star a).imK = -a.imK
                                            @[simp]
                                            theorem Quaternion.star_im {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            (star a).im = -a.im
                                            theorem Quaternion.self_add_star' {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a + star a = ↑(2 * a.re)
                                            theorem Quaternion.self_add_star {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a + star a = 2 * a.re
                                            theorem Quaternion.star_add_self' {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            star a + a = ↑(2 * a.re)
                                            theorem Quaternion.star_add_self {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            star a + a = 2 * a.re
                                            theorem Quaternion.star_eq_two_re_sub {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            star a = ↑(2 * a.re) - a
                                            @[simp]
                                            theorem Quaternion.star_coe {R : Type u_3} [CommRing R] (x : R) :
                                            star x = x
                                            @[simp]
                                            theorem Quaternion.im_star {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            star a.im = -a.im
                                            @[simp]
                                            theorem Quaternion.star_smul {S : Type u_1} {R : Type u_3} [CommRing R] [Monoid S] [DistribMulAction S R] (s : S) (a : Quaternion R) :
                                            star (s a) = s star a
                                            theorem Quaternion.eq_re_of_eq_coe {R : Type u_3} [CommRing R] {a : Quaternion R} {x : R} (h : a = x) :
                                            a = a.re
                                            @[simp]
                                            theorem Quaternion.star_eq_self {R : Type u_3} [CommRing R] [NoZeroDivisors R] [CharZero R] {a : Quaternion R} :
                                            star a = a a = a.re
                                            @[simp]
                                            theorem Quaternion.star_eq_neg {R : Type u_3} [CommRing R] [NoZeroDivisors R] [CharZero R] {a : Quaternion R} :
                                            star a = -a a.re = 0
                                            theorem Quaternion.star_mul_eq_coe {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            star a * a = (star a * a).re
                                            theorem Quaternion.mul_star_eq_coe {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                            a * star a = (a * star a).re

                                            Quaternion conjugate as an AlgEquiv to the opposite ring.

                                            Equations
                                            Instances For

                                              Square of the norm.

                                              Equations
                                              Instances For
                                                theorem Quaternion.normSq_def {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                                normSq a = (a * star a).re
                                                theorem Quaternion.normSq_def' {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                                normSq a = a.re ^ 2 + a.imI ^ 2 + a.imJ ^ 2 + a.imK ^ 2
                                                theorem Quaternion.normSq_coe {R : Type u_3} [CommRing R] (x : R) :
                                                normSq x = x ^ 2
                                                @[simp]
                                                theorem Quaternion.normSq_star {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                                theorem Quaternion.normSq_natCast {R : Type u_3} [CommRing R] (n : ) :
                                                normSq n = n ^ 2
                                                theorem Quaternion.normSq_intCast {R : Type u_3} [CommRing R] (z : ) :
                                                normSq z = z ^ 2
                                                @[simp]
                                                theorem Quaternion.normSq_neg {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                                theorem Quaternion.self_mul_star {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                                a * star a = (normSq a)
                                                theorem Quaternion.star_mul_self {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                                star a * a = (normSq a)
                                                theorem Quaternion.im_sq {R : Type u_3} [CommRing R] (a : Quaternion R) :
                                                a.im ^ 2 = -(normSq a.im)
                                                theorem Quaternion.coe_normSq_add {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                                (normSq (a + b)) = (normSq a) + a * star b + b * star a + (normSq b)
                                                theorem Quaternion.normSq_smul {R : Type u_3} [CommRing R] (r : R) (q : Quaternion R) :
                                                normSq (r q) = r ^ 2 * normSq q
                                                theorem Quaternion.normSq_add {R : Type u_3} [CommRing R] (a b : Quaternion R) :
                                                normSq (a + b) = normSq a + normSq b + 2 * (a * star b).re
                                                @[simp]
                                                @[simp]
                                                theorem Quaternion.sq_eq_normSq {R : Type u_1} [LinearOrderedCommRing R] {a : Quaternion R} :
                                                a ^ 2 = (normSq a) a = a.re
                                                Equations
                                                @[simp]
                                                theorem Quaternion.coe_inv {R : Type u_1} [LinearOrderedField R] (x : R) :
                                                x⁻¹ = (↑x)⁻¹
                                                @[simp]
                                                theorem Quaternion.coe_div {R : Type u_1} [LinearOrderedField R] (x y : R) :
                                                ↑(x / y) = x / y
                                                @[simp]
                                                theorem Quaternion.coe_zpow {R : Type u_1} [LinearOrderedField R] (x : R) (z : ) :
                                                ↑(x ^ z) = x ^ z
                                                Equations
                                                Equations
                                                @[simp]
                                                theorem Quaternion.re_nnratCast {R : Type u_1} [LinearOrderedField R] (q : ℚ≥0) :
                                                (↑q).re = q
                                                @[simp]
                                                theorem Quaternion.im_nnratCast {R : Type u_1} [LinearOrderedField R] (q : ℚ≥0) :
                                                (↑q).im = 0
                                                @[simp]
                                                theorem Quaternion.imI_nnratCast {R : Type u_1} [LinearOrderedField R] (q : ℚ≥0) :
                                                (↑q).imI = 0
                                                @[simp]
                                                theorem Quaternion.imJ_nnratCast {R : Type u_1} [LinearOrderedField R] (q : ℚ≥0) :
                                                (↑q).imJ = 0
                                                @[simp]
                                                theorem Quaternion.imK_nnratCast {R : Type u_1} [LinearOrderedField R] (q : ℚ≥0) :
                                                (↑q).imK = 0
                                                @[simp]
                                                theorem Quaternion.ratCast_re {R : Type u_1} [LinearOrderedField R] (q : ) :
                                                (↑q).re = q
                                                @[simp]
                                                theorem Quaternion.ratCast_im {R : Type u_1} [LinearOrderedField R] (q : ) :
                                                (↑q).im = 0
                                                @[simp]
                                                theorem Quaternion.ratCast_imI {R : Type u_1} [LinearOrderedField R] (q : ) :
                                                (↑q).imI = 0
                                                @[simp]
                                                theorem Quaternion.ratCast_imJ {R : Type u_1} [LinearOrderedField R] (q : ) :
                                                (↑q).imJ = 0
                                                @[simp]
                                                theorem Quaternion.ratCast_imK {R : Type u_1} [LinearOrderedField R] (q : ) :
                                                (↑q).imK = 0
                                                theorem Quaternion.coe_nnratCast {R : Type u_1} [LinearOrderedField R] (q : ℚ≥0) :
                                                q = q
                                                theorem Quaternion.coe_ratCast {R : Type u_1} [LinearOrderedField R] (q : ) :
                                                q = q
                                                Equations
                                                theorem Quaternion.normSq_zpow {R : Type u_1} [LinearOrderedField R] (a : Quaternion R) (z : ) :
                                                normSq (a ^ z) = normSq a ^ z
                                                theorem Quaternion.normSq_ratCast {R : Type u_1} [LinearOrderedField R] (q : ) :
                                                (normSq q) = q ^ 2
                                                theorem Cardinal.mk_quaternionAlgebra {R : Type u_1} (c₁ c₂ c₃ : R) :
                                                mk (QuaternionAlgebra R c₁ c₂ c₃) = mk R ^ 4

                                                The cardinality of a quaternion algebra, as a type.

                                                @[simp]
                                                theorem Cardinal.mk_quaternionAlgebra_of_infinite {R : Type u_1} (c₁ c₂ c₃ : R) [Infinite R] :
                                                mk (QuaternionAlgebra R c₁ c₂ c₃) = mk R
                                                theorem Cardinal.mk_univ_quaternionAlgebra {R : Type u_1} (c₁ c₂ c₃ : R) :
                                                mk Set.univ = mk R ^ 4

                                                The cardinality of a quaternion algebra, as a set.

                                                theorem Cardinal.mk_univ_quaternionAlgebra_of_infinite {R : Type u_1} (c₁ c₂ c₃ : R) [Infinite R] :
                                                instance Cardinal.instReprQuaternionAlgebra {R : Type u_1} [Repr R] {a b c : R} :

                                                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.
                                                @[simp]
                                                theorem Cardinal.mk_quaternion (R : Type u_1) [Zero R] [One R] [Neg R] :
                                                mk (Quaternion R) = mk R ^ 4

                                                The cardinality of the quaternions, as a type.

                                                theorem Cardinal.mk_univ_quaternion (R : Type u_1) [Zero R] [One R] [Neg R] :
                                                mk Set.univ = mk R ^ 4

                                                The cardinality of the quaternions, as a set.