Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit dffb09a

Browse files
committed
feat(linear_algebra/{clifford,exterior,tensor,free}_algebra): provide canonical images from larger algebras into smaller ones (#5745)
This adds: * `free_algebra.to_tensor` * `tensor_algebra.to_exterior` * `tensor_algebra.to_clifford` Providing the injection in the other direction is more challenging, so is left as future work.
1 parent bea7651 commit dffb09a

File tree

3 files changed

+42
-0
lines changed

3 files changed

+42
-0
lines changed

src/linear_algebra/clifford_algebra.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -165,3 +165,17 @@ alg_equiv.of_alg_hom
165165
(by { ext, simp, })
166166

167167
end clifford_algebra
168+
169+
namespace tensor_algebra
170+
171+
variables {Q}
172+
173+
/-- The canonical image of the `tensor_algebra` in the `clifford_algebra`, which maps
174+
`tensor_algebra.ι R x` to `clifford_algebra.ι Q x`. -/
175+
def to_clifford : tensor_algebra R M →ₐ[R] clifford_algebra Q :=
176+
tensor_algebra.lift R (clifford_algebra.ι Q)
177+
178+
@[simp] lemma to_clifford_ι (m : M) : (tensor_algebra.ι R m).to_clifford = clifford_algebra.ι Q m :=
179+
by simp [to_clifford]
180+
181+
end tensor_algebra

src/linear_algebra/exterior_algebra.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -236,3 +236,17 @@ lemma ι_multi_apply {n : ℕ} (v : fin n → M) :
236236
ι_multi R n v = (list.of_fn $ λ i, ι R (v i)).prod := rfl
237237

238238
end exterior_algebra
239+
240+
namespace tensor_algebra
241+
242+
variables {R M}
243+
244+
/-- The canonical image of the `tensor_algebra` in the `exterior_algebra`, which maps
245+
`tensor_algebra.ι R x` to `exterior_algebra.ι R x`. -/
246+
def to_exterior : tensor_algebra R M →ₐ[R] exterior_algebra R M :=
247+
tensor_algebra.lift R (exterior_algebra.ι R)
248+
249+
@[simp] lemma to_exterior_ι (m : M) : (tensor_algebra.ι R m).to_exterior = exterior_algebra.ι R m :=
250+
by simp [to_exterior]
251+
252+
end tensor_algebra

src/linear_algebra/tensor_algebra.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -141,3 +141,17 @@ lemma ι_left_inverse : function.left_inverse ι_inv (ι R : M → tensor_algebr
141141
λ x, by simp [ι_inv]
142142

143143
end tensor_algebra
144+
145+
namespace free_algebra
146+
147+
variables {R M}
148+
149+
/-- The canonical image of the `free_algebra` in the `tensor_algebra`, which maps
150+
`free_algebra.ι R x` to `tensor_algebra.ι R x`. -/
151+
def to_tensor : free_algebra R M →ₐ[R] tensor_algebra R M :=
152+
free_algebra.lift R (tensor_algebra.ι R)
153+
154+
@[simp] lemma to_tensor_ι (m : M) : (free_algebra.ι R m).to_tensor = tensor_algebra.ι R m :=
155+
by simp [to_tensor]
156+
157+
end free_algebra

0 commit comments

Comments
 (0)