Swap matrices #
A swap matrix indexed by i
and j
is the matrix that, when multiplying another matrix
on the left (resp. on the right), swaps the i
-th row with the j
-th row
(resp. the i
-th column with the j
-th column).
Swap matrices are a special case of elementary matrices. For transvections see
Mathlib.LinearAlgebra.Matrix.Transvection
.
Implementation detail #
This is a thin wrapper around (Equiv.swap i j).permMatrix
.
def
Matrix.swap
(R : Type u_1)
{n : Type u_2}
[Zero R]
[One R]
[DecidableEq n]
(i j : n)
:
Matrix n n R
The swap matrix swap R i j
is the identity matrix with the
i
-th and j
-th rows modified such that multiplying by it on the
left (resp. right) corresponds to swapping the i
-th and j
-th row (resp. column).
Equations
- Matrix.swap R i j = Equiv.Perm.permMatrix R (Equiv.swap i j)
Instances For
@[simp]
theorem
Matrix.conjTranspose_swap
{n : Type u_2}
[DecidableEq n]
{R : Type u_3}
[Semiring R]
[StarRing R]
(i j : n)
:
theorem
Matrix.swap_mulVec
{R : Type u_1}
{n : Type u_2}
[Semiring R]
[DecidableEq n]
[Fintype n]
(i j : n)
(a : n → R)
:
theorem
Matrix.vecMul_swap
{R : Type u_1}
{n : Type u_2}
[Semiring R]
[DecidableEq n]
[Fintype n]
(i j : n)
(a : n → R)
:
@[simp]
theorem
Matrix.swap_mulVec_apply
{R : Type u_1}
{n : Type u_2}
[Semiring R]
[DecidableEq n]
[Fintype n]
(i j : n)
(a : n → R)
:
@[simp]
theorem
Matrix.vecMul_swap_apply
{R : Type u_1}
{n : Type u_2}
[Semiring R]
[DecidableEq n]
[Fintype n]
(i j : n)
(a : n → R)
:
def
Matrix.GeneralLinearGroup.swap
(R : Type u_1)
{n : Type u_2}
[CommRing R]
[DecidableEq n]
[Fintype n]
(i j : n)
:
GL n R
Matrix.swap
as an element of GL n R
.
Equations
- Matrix.GeneralLinearGroup.swap R i j = { val := Matrix.swap R i j, inv := Matrix.swap R i j, val_inv := ⋯, inv_val := ⋯ }
Instances For
@[simp]
theorem
Matrix.GeneralLinearGroup.val_swap
(R : Type u_1)
{n : Type u_2}
[CommRing R]
[DecidableEq n]
[Fintype n]
(i j : n)
:
@[simp]
theorem
Matrix.GeneralLinearGroup.val_inv_swap
(R : Type u_1)
{n : Type u_2}
[CommRing R]
[DecidableEq n]
[Fintype n]
(i j : n)
: