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

Commit f900513

Browse files
committed
feat(linear_algebra/matrix): slightly generalize smul_left_mul_matrix (#7632)
Two minor changes that make `smul_left_mul_matrix` slightly easier to apply: * the bases `b` and `c` can now be indexed by different types * replace `(i, k)` on the LHS with `ik.1 ik.2` on the RHS (so you don't have to introduce the constructor with `rw ← prod.mk.eta` somewhere deep in your expression)
1 parent b8a6995 commit f900513

File tree

1 file changed

+4
-4
lines changed

1 file changed

+4
-4
lines changed

src/linear_algebra/matrix/to_lin.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -440,7 +440,7 @@ section lmul
440440
variables {R S T : Type*} [comm_ring R] [comm_ring S] [comm_ring T]
441441
variables [algebra R S] [algebra S T] [algebra R T] [is_scalar_tower R S T]
442442
variables {m n : Type*} [fintype m] [decidable_eq m] [fintype n] [decidable_eq n]
443-
variables (b : basis m R S) (c : basis m S T)
443+
variables (b : basis m R S) (c : basis n S T)
444444

445445
open algebra
446446

@@ -492,9 +492,9 @@ lemma left_mul_matrix_injective : function.injective (left_mul_matrix b) :=
492492
... = algebra.lmul R S x' 1 : by rw (linear_map.to_matrix b b).injective h
493493
... = x' : mul_one x'
494494

495-
lemma smul_left_mul_matrix (x) (i j) (k k') :
496-
left_mul_matrix (b.smul c) x (i, k) (j, k') =
497-
left_mul_matrix b (left_mul_matrix c x k k') i j :=
495+
lemma smul_left_mul_matrix (x) (ik jk) :
496+
left_mul_matrix (b.smul c) x ik jk =
497+
left_mul_matrix b (left_mul_matrix c x ik.2 jk.2) ik.1 jk.1 :=
498498
by simp only [left_mul_matrix_apply, linear_map.to_matrix_apply, mul_comm, basis.smul_apply,
499499
basis.smul_repr, finsupp.smul_apply, algebra.lmul_apply, id.smul_eq_mul,
500500
linear_equiv.map_smul, mul_smul_comm]

0 commit comments

Comments
 (0)