Lemmas for concrete matrices matrix (fin m) (fin n) α #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file contains alternative definitions of common operators on matrices that expand
definitionally to the expected expression when evaluated on !![] notation.
This allows "proof by reflection", where we prove A = !![A 0 0, A 0 1; A 1 0, A 1 1] by defining
matrix.eta_expand A to be equal to the RHS definitionally, and then prove that
A = eta_expand A.
The definitions in this file should normally not be used directly; the intent is for the
corresponding *_eq lemmas to be used in a place where they are definitionally unfolded.
Main definitionss #
∀ with better defeq for ∀ x : matrix (fin m) (fin n) α, P x.
Equations
- matrix.forall P = fin_vec.forall (λ (r : fin n → α), matrix.forall (λ (A : matrix (fin m) (fin n) α), P (⇑matrix.of (matrix.vec_cons r A))))
- matrix.forall P = P (⇑matrix.of matrix.vec_empty)
∃ with better defeq for ∃ x : matrix (fin m) (fin n) α, P x.
Equations
- matrix.exists P = fin_vec.exists (λ (r : fin n → α), matrix.exists (λ (A : matrix (fin m) (fin n) α), P (⇑matrix.of (matrix.vec_cons r A))))
- matrix.exists P = P (⇑matrix.of matrix.vec_empty)
matrix.tranpose with better defeq for fin
Equations
- A.transposeᵣ = ⇑matrix.of (matrix.vec_cons (fin_vec.map (λ (v : fin (n + 1) → α), v 0) A) (A.submatrix id fin.succ).transposeᵣ)
- A.transposeᵣ = ⇑matrix.of matrix.vec_empty
matrix.dot_product with better defeq for fin
Equations
- matrix.dot_productᵣ a b = fin_vec.sum (fin_vec.seq (fin_vec.map has_mul.mul a) b)
This can be used to prove
example (a b c d : α) [has_mul α] [add_comm_monoid α] :
  dot_product ![a, b] ![c, d] = a * c + b * d :=
(dot_productᵣ_eq _ _).symm
matrix.mul with better defeq for fin
Equations
- A.mulᵣ B = ⇑matrix.of (fin_vec.map (λ (v₁ : fin m → α), fin_vec.map (λ (v₂ : fin m → α), matrix.dot_productᵣ v₁ v₂) B.transpose) A)
This can be used to prove
example [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) :
  !![a₁₁, a₁₂;
     a₂₁, a₂₂] ⬝ !![b₁₁, b₁₂;
                    b₂₁, b₂₂] =
  !![a₁₁*b₁₁ + a₁₂*b₂₁, a₁₁*b₁₂ + a₁₂*b₂₂;
     a₂₁*b₁₁ + a₂₂*b₂₁, a₂₁*b₁₂ + a₂₂*b₂₂] :=
(mulᵣ_eq _ _).symm
This can be used to prove
example [non_unital_non_assoc_semiring α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁ b₂ : α) :
  !![a₁₁, a₁₂;
     a₂₁, a₂₂].mul_vec ![b₁, b₂] = ![a₁₁*b₁ + a₁₂*b₂, a₂₁*b₁ + a₂₂*b₂] :=
(mul_vecᵣ_eq _ _).symm
matrix.vec_mul with better defeq for fin
Equations
- matrix.vec_mulᵣ v A = fin_vec.map (λ (a : fin l → α), matrix.dot_productᵣ v a) A.transpose
This can be used to prove
example [non_unital_non_assoc_semiring α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁ b₂ : α) :
  vec_mul ![b₁, b₂] !![a₁₁, a₁₂;
                       a₂₁, a₂₂] = ![b₁*a₁₁ + b₂*a₂₁, b₁*a₁₂ + b₂*a₂₂] :=
(vec_mulᵣ_eq _ _).symm
Expand A to !![A 0 0, ...; ..., A m n]
Equations
- A.eta_expand = ⇑matrix.of (fin_vec.eta_expand (λ (i : fin m), fin_vec.eta_expand (λ (j : fin n), A i j)))