|
| 1 | +/- |
| 2 | +Copyright (c) 2021 Eric Wieser. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Eric Wieser |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module algebra.quaternion_basis |
| 7 | +! leanprover-community/mathlib commit 3aa5b8a9ed7a7cabd36e6e1d022c9858ab8a8c2d |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.Algebra.Quaternion |
| 12 | +import Mathlib.Tactic.Ring |
| 13 | + |
| 14 | +/-! |
| 15 | +# Basis on a quaternion-like algebra |
| 16 | +
|
| 17 | +## Main definitions |
| 18 | +
|
| 19 | +* `QuaternionAlgebra.Basis A c₁ c₂`: a basis for a subspace of an `R`-algebra `A` that has the |
| 20 | + same algebra structure as `ℍ[R,c₁,c₂]`. |
| 21 | +* `QuaternionAlgebra.Basis.self R`: the canonical basis for `ℍ[R,c₁,c₂]`. |
| 22 | +* `QuaternionAlgebra.Basis.compHom b f`: transform a basis `b` by an AlgHom `f`. |
| 23 | +* `QuaternionAlgebra.lift`: Define an `AlgHom` out of `ℍ[R,c₁,c₂]` by its action on the basis |
| 24 | + elements `i`, `j`, and `k`. In essence, this is a universal property. Analogous to `Complex.lift`, |
| 25 | + but takes a bundled `QuaternionAlgebra.Basis` instead of just a `Subtype` as the amount of |
| 26 | + data / proves is non-negligible. |
| 27 | +-/ |
| 28 | + |
| 29 | + |
| 30 | +open Quaternion |
| 31 | + |
| 32 | +namespace QuaternionAlgebra |
| 33 | + |
| 34 | +/-- A quaternion basis contains the information both sufficient and necessary to construct an |
| 35 | +`R`-algebra homomorphism from `ℍ[R,c₁,c₂]` to `A`; or equivalently, a surjective |
| 36 | +`R`-algebra homomorphism from `ℍ[R,c₁,c₂]` to an `R`-subalgebra of `A`. |
| 37 | +
|
| 38 | +Note that for definitional convenience, `k` is provided as a field even though `i_mul_j` fully |
| 39 | +determines it. -/ |
| 40 | +structure Basis {R : Type _} (A : Type _) [CommRing R] [Ring A] [Algebra R A] (c₁ c₂ : R) where |
| 41 | + (i j k : A) |
| 42 | + i_mul_i : i * i = c₁ • (1 : A) |
| 43 | + j_mul_j : j * j = c₂ • (1 : A) |
| 44 | + i_mul_j : i * j = k |
| 45 | + j_mul_i : j * i = -k |
| 46 | +#align quaternion_algebra.basis QuaternionAlgebra.Basis |
| 47 | + |
| 48 | +variable {R : Type _} {A B : Type _} [CommRing R] [Ring A] [Ring B] [Algebra R A] [Algebra R B] |
| 49 | + |
| 50 | +variable {c₁ c₂ : R} |
| 51 | + |
| 52 | +namespace Basis |
| 53 | + |
| 54 | +/-- Since `k` is redundant, it is not necessary to show `q₁.k = q₂.k` when showing `q₁ = q₂`. -/ |
| 55 | +@[ext] |
| 56 | +protected theorem ext ⦃q₁ q₂ : Basis A c₁ c₂⦄ (hi : q₁.i = q₂.i) (hj : q₁.j = q₂.j) : q₁ = q₂ := by |
| 57 | + cases q₁; rename_i q₁_i_mul_j _ |
| 58 | + cases q₂; rename_i q₂_i_mul_j _ |
| 59 | + congr |
| 60 | + rw [← q₁_i_mul_j, ← q₂_i_mul_j] |
| 61 | + congr |
| 62 | +#align quaternion_algebra.basis.ext QuaternionAlgebra.Basis.ext |
| 63 | + |
| 64 | +variable (R) |
| 65 | + |
| 66 | +/-- There is a natural quaternionic basis for the `QuaternionAlgebra`. -/ |
| 67 | +@[simps i j k] |
| 68 | +protected def self : Basis ℍ[R,c₁,c₂] c₁ c₂ where |
| 69 | + i := ⟨0, 1, 0, 0⟩ |
| 70 | + i_mul_i := by ext <;> simp |
| 71 | + j := ⟨0, 0, 1, 0⟩ |
| 72 | + j_mul_j := by ext <;> simp |
| 73 | + k := ⟨0, 0, 0, 1⟩ |
| 74 | + i_mul_j := by ext <;> simp |
| 75 | + j_mul_i := by ext <;> simp |
| 76 | +#align quaternion_algebra.basis.self QuaternionAlgebra.Basis.self |
| 77 | + |
| 78 | +variable {R} |
| 79 | + |
| 80 | +instance : Inhabited (Basis ℍ[R,c₁,c₂] c₁ c₂) := |
| 81 | + ⟨Basis.self R⟩ |
| 82 | + |
| 83 | +variable (q : Basis A c₁ c₂) |
| 84 | + |
| 85 | +attribute [simp] i_mul_i j_mul_j i_mul_j j_mul_i |
| 86 | + |
| 87 | +@[simp] |
| 88 | +theorem i_mul_k : q.i * q.k = c₁ • q.j := by |
| 89 | + rw [← i_mul_j, ← mul_assoc, i_mul_i, smul_mul_assoc, one_mul] |
| 90 | +#align quaternion_algebra.basis.i_mul_k QuaternionAlgebra.Basis.i_mul_k |
| 91 | + |
| 92 | +@[simp] |
| 93 | +theorem k_mul_i : q.k * q.i = -c₁ • q.j := by |
| 94 | + rw [← i_mul_j, mul_assoc, j_mul_i, mul_neg, i_mul_k, neg_smul] |
| 95 | +#align quaternion_algebra.basis.k_mul_i QuaternionAlgebra.Basis.k_mul_i |
| 96 | + |
| 97 | +@[simp] |
| 98 | +theorem k_mul_j : q.k * q.j = c₂ • q.i := by |
| 99 | + rw [← i_mul_j, mul_assoc, j_mul_j, mul_smul_comm, mul_one] |
| 100 | +#align quaternion_algebra.basis.k_mul_j QuaternionAlgebra.Basis.k_mul_j |
| 101 | + |
| 102 | +@[simp] |
| 103 | +theorem j_mul_k : q.j * q.k = -c₂ • q.i := by |
| 104 | + rw [← i_mul_j, ← mul_assoc, j_mul_i, neg_mul, k_mul_j, neg_smul] |
| 105 | +#align quaternion_algebra.basis.j_mul_k QuaternionAlgebra.Basis.j_mul_k |
| 106 | + |
| 107 | +@[simp] |
| 108 | +theorem k_mul_k : q.k * q.k = -((c₁ * c₂) • (1 : A)) := by |
| 109 | + rw [← i_mul_j, mul_assoc, ← mul_assoc q.j _ _, j_mul_i, ← i_mul_j, ← mul_assoc, mul_neg, ← |
| 110 | + mul_assoc, i_mul_i, smul_mul_assoc, one_mul, neg_mul, smul_mul_assoc, j_mul_j, smul_smul] |
| 111 | +#align quaternion_algebra.basis.k_mul_k QuaternionAlgebra.Basis.k_mul_k |
| 112 | + |
| 113 | +/-- Intermediate result used to define `QuaternionAlgebra.Basis.liftHom`. -/ |
| 114 | +def lift (x : ℍ[R,c₁,c₂]) : A := |
| 115 | + algebraMap R _ x.re + x.imI • q.i + x.imJ • q.j + x.imK • q.k |
| 116 | +#align quaternion_algebra.basis.lift QuaternionAlgebra.Basis.lift |
| 117 | + |
| 118 | +theorem lift_zero : q.lift (0 : ℍ[R,c₁,c₂]) = 0 := by simp [lift] |
| 119 | +#align quaternion_algebra.basis.lift_zero QuaternionAlgebra.Basis.lift_zero |
| 120 | + |
| 121 | +theorem lift_one : q.lift (1 : ℍ[R,c₁,c₂]) = 1 := by simp [lift] |
| 122 | +#align quaternion_algebra.basis.lift_one QuaternionAlgebra.Basis.lift_one |
| 123 | + |
| 124 | +theorem lift_add (x y : ℍ[R,c₁,c₂]) : q.lift (x + y) = q.lift x + q.lift y := by |
| 125 | + simp [lift, add_smul] |
| 126 | + abel |
| 127 | +#align quaternion_algebra.basis.lift_add QuaternionAlgebra.Basis.lift_add |
| 128 | + |
| 129 | +theorem lift_mul (x y : ℍ[R,c₁,c₂]) : q.lift (x * y) = q.lift x * q.lift y := by |
| 130 | + simp only [lift, Algebra.algebraMap_eq_smul_one] |
| 131 | + simp only [add_mul] |
| 132 | + simp_rw [add_mul, mul_add, smul_mul_assoc, mul_smul_comm, one_mul, mul_one, ← Algebra.smul_def, |
| 133 | + smul_add, smul_smul] |
| 134 | + simp only [i_mul_i, j_mul_j, i_mul_j, j_mul_i, i_mul_k, k_mul_i, k_mul_j, j_mul_k, k_mul_k] |
| 135 | + simp only [smul_smul, smul_neg, sub_eq_add_neg, add_smul, ← add_assoc, mul_neg, neg_smul] |
| 136 | + simp only [mul_right_comm _ _ (c₁ * c₂), mul_comm _ (c₁ * c₂)] |
| 137 | + simp only [mul_comm _ c₁, mul_right_comm _ _ c₁] |
| 138 | + simp only [mul_comm _ c₂, mul_right_comm _ _ c₂] |
| 139 | + simp only [← mul_comm c₁ c₂, ← mul_assoc] |
| 140 | + simp [sub_eq_add_neg, add_smul, ← add_assoc] |
| 141 | + abel |
| 142 | +#align quaternion_algebra.basis.lift_mul QuaternionAlgebra.Basis.lift_mul |
| 143 | + |
| 144 | +theorem lift_smul (r : R) (x : ℍ[R,c₁,c₂]) : q.lift (r • x) = r • q.lift x := by |
| 145 | + simp [lift, mul_smul, ← Algebra.smul_def] |
| 146 | +#align quaternion_algebra.basis.lift_smul QuaternionAlgebra.Basis.lift_smul |
| 147 | + |
| 148 | +/-- A `QuaternionAlgebra.Basis` implies an `AlgHom` from the quaternions. -/ |
| 149 | +@[simps!] |
| 150 | +def liftHom : ℍ[R,c₁,c₂] →ₐ[R] A := |
| 151 | + AlgHom.mk' |
| 152 | + { toFun := q.lift |
| 153 | + map_zero' := q.lift_zero |
| 154 | + map_one' := q.lift_one |
| 155 | + map_add' := q.lift_add |
| 156 | + map_mul' := q.lift_mul } q.lift_smul |
| 157 | +#align quaternion_algebra.basis.lift_hom QuaternionAlgebra.Basis.liftHom |
| 158 | + |
| 159 | +/-- Transform a `QuaternionAlgebra.Basis` through an `AlgHom`. -/ |
| 160 | +@[simps i j k] |
| 161 | +def compHom (F : A →ₐ[R] B) : Basis B c₁ c₂ where |
| 162 | + i := F q.i |
| 163 | + i_mul_i := by rw [← F.map_mul, q.i_mul_i, F.map_smul, F.map_one] |
| 164 | + j := F q.j |
| 165 | + j_mul_j := by rw [← F.map_mul, q.j_mul_j, F.map_smul, F.map_one] |
| 166 | + k := F q.k |
| 167 | + i_mul_j := by rw [← F.map_mul, q.i_mul_j] |
| 168 | + j_mul_i := by rw [← F.map_mul, q.j_mul_i, F.map_neg] |
| 169 | +#align quaternion_algebra.basis.comp_hom QuaternionAlgebra.Basis.compHom |
| 170 | + |
| 171 | +end Basis |
| 172 | + |
| 173 | +/-- A quaternionic basis on `A` is equivalent to a map from the quaternion algebra to `A`. -/ |
| 174 | +@[simps] |
| 175 | +def lift : Basis A c₁ c₂ ≃ (ℍ[R,c₁,c₂] →ₐ[R] A) where |
| 176 | + toFun := Basis.liftHom |
| 177 | + invFun := (Basis.self R).compHom |
| 178 | + left_inv q := by ext <;> simp [Basis.lift] |
| 179 | + right_inv F := by |
| 180 | + ext |
| 181 | + dsimp [Basis.lift] |
| 182 | + rw [← F.commutes] |
| 183 | + simp only [← F.commutes, ← F.map_smul, ← F.map_add, mk_add_mk, smul_mk, smul_zero, |
| 184 | + algebraMap_eq] |
| 185 | + congr <;> simp |
| 186 | +#align quaternion_algebra.lift QuaternionAlgebra.lift |
| 187 | + |
| 188 | +end QuaternionAlgebra |
0 commit comments