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

Commit 6b556cf

Browse files
feat(combinatorics/adjacency_matrix): defines adjacency matrices of simple graphs (#3672)
defines the adjacency matrix of a simple graph proves lemmas about adjacency matrix that will form the bulk of the proof of the friendship theorem (freek 83) Co-authored-by: Aaron Anderson <awainverse@gmail.com>
1 parent ea9bf31 commit 6b556cf

File tree

1 file changed

+101
-0
lines changed

1 file changed

+101
-0
lines changed

src/combinatorics/adj_matrix.lean

Lines changed: 101 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,101 @@
1+
/-
2+
Copyright (c) 2020 Aaron Anderson. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Author: Aaron Anderson, Jalex Stark.
5+
-/
6+
import linear_algebra.matrix
7+
import data.rel
8+
import combinatorics.simple_graph
9+
10+
/-!
11+
# Adjacency Matrices
12+
13+
This module defines the adjacency matrix of a graph, and provides theorems connecting graph
14+
properties to computational properties of the matrix.
15+
16+
## Main definitions
17+
18+
* `adj_matrix` is the adjacency matrix of a `simple_graph` with coefficients in a given semiring.
19+
20+
-/
21+
22+
open_locale big_operators matrix
23+
open finset matrix simple_graph
24+
25+
universes u v
26+
variables {α : Type u} [fintype α]
27+
variables (R : Type v) [semiring R]
28+
29+
namespace simple_graph
30+
31+
variables (G : simple_graph α) (R) [decidable_rel G.adj]
32+
33+
/-- `adj_matrix G R` is the matrix `A` such that `A i j = (1 : R)` if `i` and `j` are
34+
adjacent in the simple graph `G`, and otherwise `A i j = 0`. -/
35+
def adj_matrix : matrix α α R
36+
| i j := if (G.adj i j) then 1 else 0
37+
38+
variable {R}
39+
40+
@[simp]
41+
lemma adj_matrix_apply (v w : α) : G.adj_matrix R v w = if (G.adj v w) then 1 else 0 := rfl
42+
43+
@[simp]
44+
theorem transpose_adj_matrix : (G.adj_matrix R)ᵀ = G.adj_matrix R :=
45+
by { ext, simp [edge_symm] }
46+
47+
@[simp]
48+
lemma adj_matrix_dot_product (v : α) (vec : α → R) :
49+
dot_product (G.adj_matrix R v) vec = ∑ u in G.neighbor_finset v, vec u :=
50+
by simp [neighbor_finset_eq_filter, dot_product, sum_filter]
51+
52+
@[simp]
53+
lemma dot_product_adj_matrix (v : α) (vec : α → R) :
54+
dot_product vec (G.adj_matrix R v) = ∑ u in G.neighbor_finset v, vec u :=
55+
by simp [neighbor_finset_eq_filter, dot_product, sum_filter, sum_apply]
56+
57+
@[simp]
58+
lemma adj_matrix_mul_vec_apply (v : α) (vec : α → R) :
59+
((G.adj_matrix R).mul_vec vec) v = ∑ u in G.neighbor_finset v, vec u :=
60+
by rw [mul_vec, adj_matrix_dot_product]
61+
62+
@[simp]
63+
lemma adj_matrix_vec_mul_apply (v : α) (vec : α → R) :
64+
((G.adj_matrix R).vec_mul vec) v = ∑ u in G.neighbor_finset v, vec u :=
65+
begin
66+
rw [← dot_product_adj_matrix, vec_mul],
67+
refine congr rfl _, ext,
68+
rw [← transpose_apply (adj_matrix R G) x v, transpose_adj_matrix],
69+
end
70+
71+
@[simp]
72+
lemma adj_matrix_mul_apply (M : matrix α α R) (v w : α) :
73+
(G.adj_matrix R ⬝ M) v w = ∑ u in G.neighbor_finset v, M u w :=
74+
by simp [mul_apply, neighbor_finset_eq_filter, sum_filter]
75+
76+
@[simp]
77+
lemma mul_adj_matrix_apply (M : matrix α α R) (v w : α) :
78+
(M ⬝ G.adj_matrix R) v w = ∑ u in G.neighbor_finset w, M v u :=
79+
by simp [mul_apply, neighbor_finset_eq_filter, sum_filter, edge_symm]
80+
81+
variable (R)
82+
theorem trace_adj_matrix : matrix.trace α R R (G.adj_matrix R) = 0 := by simp
83+
variable {R}
84+
85+
theorem adj_matrix_mul_self_apply_self (i : α) :
86+
((G.adj_matrix R) ⬝ (G.adj_matrix R)) i i = degree G i :=
87+
by simp [degree]
88+
89+
variable {G}
90+
91+
@[simp]
92+
lemma adj_matrix_mul_vec_const_apply {r : R} {v : α} :
93+
(G.adj_matrix R).mul_vec (function.const _ r) v = G.degree v * r :=
94+
by simp [degree]
95+
96+
lemma adj_matrix_mul_vec_const_apply_of_regular {d : ℕ} {r : R} (hd : G.is_regular_of_degree d)
97+
{v : α} :
98+
(G.adj_matrix R).mul_vec (function.const _ r) v = (d * r) :=
99+
by simp [hd v]
100+
101+
end simple_graph

0 commit comments

Comments
 (0)