Skip to content

Commit

Permalink
fix(data/matrix/basic): remove an accidental requirement for a matrix…
Browse files Browse the repository at this point in the history
… to be square (#8372)
  • Loading branch information
l534zhan committed Jul 20, 2021
1 parent fce38f1 commit ed8d597
Showing 1 changed file with 1 addition and 1 deletion.
2 changes: 1 addition & 1 deletion src/data/matrix/basic.lean
Expand Up @@ -353,7 +353,7 @@ instance [has_mul α] [add_comm_monoid α] : has_mul (matrix n n α) := ⟨matri
@[simp] theorem mul_eq_mul [has_mul α] [add_comm_monoid α] (M N : matrix n n α) :
M * N = M ⬝ N := rfl

theorem mul_apply' [has_mul α] [add_comm_monoid α] {M N : matrix n n α} {i k} :
theorem mul_apply' [has_mul α] [add_comm_monoid α] {M : matrix l m α} {N : matrix m n α} {i k} :
(M ⬝ N) i k = dot_product (λ j, M i j) (λ j, N j k) := rfl

@[simp] theorem diagonal_neg [decidable_eq n] [add_group α] (d : n → α) :
Expand Down

0 comments on commit ed8d597

Please sign in to comment.