Skip to content

Commit

Permalink
feat(linear_algebra/matrix): define the trace of a square matrix (#1883)
Browse files Browse the repository at this point in the history
* feat(linear_algebra/matrix): define the trace of a square matrix

* Move ring carrier to correct universe

* Add lemma trace_one, and define diag as linear map

* Define diag and trace solely as linear functions

* Diag and trace for module-valued matrices

* Fix cyclic import

* Rename matrix.mul_sum' --> matrix.smul_sum

Co-Authored-By: Johan Commelin <johan@commelin.net>

* Trigger CI

Co-authored-by: Johan Commelin <johan@commelin.net>
Co-authored-by: mergify[bot] <37929162+mergify[bot]@users.noreply.github.com>
  • Loading branch information
3 people committed Jan 26, 2020
1 parent 587b312 commit 497e692
Show file tree
Hide file tree
Showing 3 changed files with 40 additions and 5 deletions.
5 changes: 3 additions & 2 deletions src/data/matrix/basic.lean
Expand Up @@ -8,7 +8,7 @@ Matrices
import algebra.module algebra.pi_instances
import data.fintype

universes u v
universes u v w

def matrix (m n : Type u) [fintype m] [fintype n] (α : Type v) : Type (max u v) :=
m → n → α
Expand Down Expand Up @@ -220,7 +220,8 @@ instance [decidable_eq n] [ring α] : ring (matrix n n α) :=
{ ..matrix.add_comm_group, ..matrix.semiring }

instance [semiring α] : has_scalar α (matrix m n α) := pi.has_scalar
instance [ring α] : module α (matrix m n α) := pi.module _
instance {β : Type w} [ring α] [add_comm_group β] [module α β] :
module α (matrix m n β) := pi.module _

@[simp] lemma smul_val [semiring α] (a : α) (A : matrix m n α) (i : m) (j : n) : (a • A) i j = a * A i j := rfl

Expand Down
36 changes: 35 additions & 1 deletion src/linear_algebra/matrix.lean
Expand Up @@ -36,7 +36,7 @@ noncomputable theory

open set submodule

universes u v
universes u v w
variables {l m n : Type u} [fintype l] [fintype m] [fintype n]

namespace matrix
Expand Down Expand Up @@ -187,6 +187,40 @@ end linear_equiv_matrix

namespace matrix

section trace

variables {R : Type v} {M : Type w} [ring R] [add_comm_group M] [module R M]

/--
The diagonal of a square matrix.
-/
def diag : (matrix n n M) →ₗ[R] n → M := {
to_fun := λ A i, A i i,
add := by { intros, ext, refl, },
smul := by { intros, ext, refl, } }

@[simp] lemma diag_one [decidable_eq n] :
(diag : (matrix n n R) →ₗ[R] n → R) 1 = λ i, 1 := by {
dunfold diag, ext, simp [one_val_eq], }

protected lemma smul_sum {α : Type u} {s : finset α} {f : α → M} {c : R} :
c • s.sum f = s.sum (λx, c • f x) := (s.sum_hom _).symm

/--
The trace of a square matrix.
-/
def trace : (matrix n n M) →ₗ[R] M := {
to_fun := finset.univ.sum ∘ (diag : (matrix n n M) →ₗ[R] n → M),
add := by { intros, apply finset.sum_add_distrib, },
smul := by { intros, simp [matrix.smul_sum], } }

@[simp] lemma trace_one [decidable_eq n] :
(trace : (matrix n n R) →ₗ[R] R) 1 = fintype.card n := by {
have h : @trace n _ R R _ _ _ 1 = finset.univ.sum (@diag n _ R R _ _ _ 1) := rfl,
rw [h, diag_one, finset.sum_const, add_monoid.smul_one], refl, }

end trace

section ring

variables {R : Type v} [comm_ring R]
Expand Down
4 changes: 2 additions & 2 deletions src/ring_theory/algebra.lean
Expand Up @@ -156,8 +156,8 @@ set_option class.instance_max_depth 40
instance matrix_algebra (n : Type u) (R : Type v)
[fintype n] [decidable_eq n] [comm_ring R] : algebra R (matrix n n R) :=
{ to_fun := (λ r, r • 1),
hom := { map_one := by simp,
map_mul := by { intros, simp [mul_smul], },
hom := { map_one := by { ext, simp, },
map_mul := by { intros, ext, simp [mul_assoc], },
map_add := by { intros, simp [add_smul], } },
commutes' := by { intros, simp },
smul_def' := by { intros, simp } }
Expand Down

0 comments on commit 497e692

Please sign in to comment.