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

Commit a7be93b

Browse files
committed
feat(data/matrix/hadamard): add the Hadamard product (#8956)
Co-authored-by: l534zhan <84618936+l534zhan@users.noreply.github.com>
1 parent 91824e5 commit a7be93b

File tree

2 files changed

+140
-0
lines changed

2 files changed

+140
-0
lines changed

src/data/matrix/basic.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -236,6 +236,13 @@ variables {n α R}
236236
(diagonal d).map f = diagonal (λ m, f (d m)) :=
237237
by { ext, simp only [diagonal, map_apply], split_ifs; simp [h], }
238238

239+
@[simp] lemma diagonal_conj_transpose [semiring α] [star_ring α] (v : n → α) :
240+
(diagonal v)ᴴ = diagonal (star v) :=
241+
begin
242+
rw [conj_transpose, diagonal_transpose, diagonal_map (star_zero _)],
243+
refl,
244+
end
245+
239246
section one
240247
variables [has_zero α] [has_one α]
241248

src/data/matrix/hadamard.lean

Lines changed: 133 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,133 @@
1+
/-
2+
Copyright (c) 2021 Lu-Ming Zhang. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Lu-Ming Zhang
5+
-/
6+
import data.matrix.notation
7+
import linear_algebra.matrix.trace
8+
9+
/-!
10+
# Hadamard product of matrices
11+
12+
This file defines the Hadamard product `matrix.hadamard`
13+
and contains basic properties about them.
14+
15+
## Main definition
16+
17+
- `matrix.hadamard`: defines the Hadamard product,
18+
which is the pointwise product of two matrices of the same size.
19+
20+
## Notation
21+
22+
* `⊙`: the Hadamard product `matrix.hadamard`;
23+
24+
## References
25+
26+
* <https://en.wikipedia.org/wiki/hadamard_product_(matrices)>
27+
28+
## Tags
29+
30+
hadamard product, hadamard
31+
-/
32+
33+
variables {α β γ m n : Type*}
34+
variables {R : Type*}
35+
36+
namespace matrix
37+
open_locale matrix big_operators
38+
39+
/-- `matrix.hadamard` defines the Hadamard product,
40+
which is the pointwise product of two matrices of the same size.-/
41+
@[simp]
42+
def hadamard [has_mul α] (A : matrix m n α) (B : matrix m n α) : matrix m n α
43+
| i j := A i j * B i j
44+
45+
localized "infix ` ⊙ `:100 := matrix.hadamard" in matrix
46+
47+
section basic_properties
48+
49+
variables (A : matrix m n α) (B : matrix m n α) (C : matrix m n α)
50+
51+
/- commutativity -/
52+
lemma hadamard_comm [comm_semigroup α] : A ⊙ B = B ⊙ A :=
53+
ext $ λ _ _, mul_comm _ _
54+
55+
/- associativity -/
56+
lemma hadamard_assoc [semigroup α] : A ⊙ B ⊙ C = A ⊙ (B ⊙ C) :=
57+
ext $ λ _ _, mul_assoc _ _ _
58+
59+
/- distributivity -/
60+
lemma hadamard_add [distrib α] : A ⊙ (B + C) = A ⊙ B + A ⊙ C :=
61+
ext $ λ _ _, left_distrib _ _ _
62+
63+
lemma add_hadamard [distrib α] : (B + C) ⊙ A = B ⊙ A + C ⊙ A :=
64+
ext $ λ _ _, right_distrib _ _ _
65+
66+
/- scalar multiplication -/
67+
section scalar
68+
69+
@[simp] lemma smul_hadamard [has_mul α] [has_scalar R α] [is_scalar_tower R α α] (k : R) :
70+
(k • A) ⊙ B = k • A ⊙ B :=
71+
ext $ λ _ _, smul_mul_assoc _ _ _
72+
73+
@[simp] lemma hadamard_smul [has_mul α] [has_scalar R α] [smul_comm_class R α α] (k : R):
74+
A ⊙ (k • B) = k • A ⊙ B :=
75+
ext $ λ _ _, mul_smul_comm _ _ _
76+
77+
end scalar
78+
79+
section zero
80+
81+
variables [mul_zero_class α]
82+
83+
@[simp] lemma hadamard_zero : A ⊙ (0 : matrix m n α) = 0 :=
84+
ext $ λ _ _, mul_zero _
85+
86+
@[simp] lemma zero_hadamard : (0 : matrix m n α) ⊙ A = 0 :=
87+
ext $ λ _ _, zero_mul _
88+
89+
end zero
90+
91+
section one
92+
93+
variables [decidable_eq n] [mul_zero_one_class α]
94+
variables (M : matrix n n α)
95+
96+
lemma hadamard_one : M ⊙ (1 : matrix n n α) = diagonal (λ i, M i i) :=
97+
by { ext, by_cases h : i = j; simp [h] }
98+
99+
lemma one_hadamard : (1 : matrix n n α) ⊙ M = diagonal (λ i, M i i) :=
100+
by { ext, by_cases h : i = j; simp [h] }
101+
102+
end one
103+
104+
section diagonal
105+
106+
variables [decidable_eq n] [mul_zero_class α]
107+
108+
lemma diagonal_hadamard_diagonal (v : n → α) (w : n → α) :
109+
diagonal v ⊙ diagonal w = diagonal (v * w) :=
110+
ext $ λ _ _, (apply_ite2 _ _ _ _ _ _).trans (congr_arg _ $ zero_mul 0)
111+
112+
end diagonal
113+
114+
section trace
115+
116+
variables [fintype m] [fintype n]
117+
variables (R) [semiring α] [semiring R] [module R α]
118+
119+
lemma sum_hadamard_eq : ∑ (i : m) (j : n), (A ⊙ B) i j = trace m R α (A ⬝ Bᵀ) :=
120+
rfl
121+
122+
lemma dot_product_vec_mul_hadamard [decidable_eq m] [decidable_eq n] (v : m → α) (w : n → α) :
123+
dot_product (vec_mul v (A ⊙ B)) w = trace m R α (diagonal v ⬝ A ⬝ (B ⬝ diagonal w)ᵀ) :=
124+
begin
125+
rw [←sum_hadamard_eq, finset.sum_comm],
126+
simp [dot_product, vec_mul, finset.sum_mul, mul_assoc],
127+
end
128+
129+
end trace
130+
131+
end basic_properties
132+
133+
end matrix

0 commit comments

Comments
 (0)