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

Commit 944ffff

Browse files
committed
chore(combinatorics/quiver): make quiver a class (#7150)
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent 5e188d2 commit 944ffff

File tree

1 file changed

+127
-104
lines changed

1 file changed

+127
-104
lines changed

src/combinatorics/quiver.lean

Lines changed: 127 additions & 104 deletions
Original file line numberDiff line numberDiff line change
@@ -3,130 +3,153 @@ Copyright (c) 2021 David Wärn. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: David Wärn
55
-/
6-
import tactic
6+
import data.equiv.basic
7+
import order.well_founded
8+
import data.nat.basic
9+
import data.opposite
10+
711
/-!
812
# Quivers
913
10-
This module defines quivers. A quiver `G` on a type `V` of vertices assigns to every
11-
pair `a b : V` of vertices a type `G.arrow a b` of arrows from `a` to `b`. This
14+
This module defines quivers. A quiver on a type `V` of vertices assigns to every
15+
pair `a b : V` of vertices a type `a ⟶ b` of arrows from `a` to `b`. This
1216
is a very permissive notion of directed graph.
17+
18+
## Implementation notes
19+
20+
It would be interesting to try to replace `has_hom` with `quiver` in the definition of a category.
21+
This would be convenient for defining path categories.
22+
23+
Currently `quiver` is defined with `arrow : V → V → Sort v`.
24+
This is different from the category theory setup,
25+
where we insist that morphisms live in some `Type`.
26+
There's some balance here: it's nice to allow `Prop` to ensure there are no multiple arrows,
27+
but it is also results in error-prone universe signatures when constraints require a `Type`.
1328
-/
1429

30+
open opposite
31+
1532
-- We use the same universe order as in category theory.
1633
-- See note [category_theory universes]
1734
universes v u
1835

1936
/-- A quiver `G` on a type `V` of vertices assigns to every pair `a b : V` of vertices
20-
a type `G.arrow a b` of arrows from `a` to `b`. -/
21-
structure quiver (V : Type u) :=
22-
(arrow : V → V → Sort v)
37+
a sort `a ⟶ b` of arrows from `a` to `b`.
38+
39+
For graphs with no repeated edges, one can use `quiver.{0} V`, which ensures
40+
`a ⟶ b : Prop`. For multigraphs, one can use `quiver.{v+1} V`, which ensures
41+
`a ⟶ b : Type v`. -/
42+
class quiver (V : Type u) :=
43+
(hom : V → V → Sort v)
44+
45+
infixr ` ⟶ `:10 := quiver.hom -- type as \h
2346

2447
/-- A wide subquiver `H` of `G` picks out a set `H a b` of arrows from `a` to `b`
25-
for every pair of vertices `a b`. -/
26-
def wide_subquiver {V} (G : quiver V) :=
27-
Π a b : V, set (G.arrow a b)
48+
for every pair of vertices `a b`.
2849
29-
/-- A wide subquiver viewed as a quiver on its own. -/
30-
def wide_subquiver.quiver {V} {G : quiver V} (H : wide_subquiver G) : quiver V :=
31-
⟨λ a b, H a b⟩
50+
NB: this does not work for `Prop`-valued quivers. It requires `G : quiver.{v+1} V`. -/
51+
def wide_subquiver (V) [quiver.{v+1} V] :=
52+
Π a b : V, set (a ⟶ b)
3253

33-
namespace quiver
54+
/-- A type synonym for `V`, when thought of as a quiver having only the arrows from
55+
some `wide_subquiver`. -/
56+
@[nolint unused_arguments has_inhabited_instance]
57+
def wide_subquiver.to_Type (V) [quiver V] (H : wide_subquiver V) : Type u := V
3458

35-
/-- The quiver with no arrows. -/
36-
protected def empty (V) : quiver.{v} V := ⟨λ a b, pempty⟩
59+
instance wide_subquiver_has_coe_to_sort {V} [quiver V] : has_coe_to_sort (wide_subquiver V) :=
60+
{ S := Type u,
61+
coe := λ H, wide_subquiver.to_Type V H, }
3762

38-
instance {V} : inhabited (quiver.{v} V) := ⟨quiver.empty V⟩
39-
instance {V} (G : quiver V) : has_bot (wide_subquiver G) := ⟨λ a b, ∅⟩
40-
instance {V} (G : quiver V) : has_top (wide_subquiver G) := ⟨λ a b, set.univ⟩
41-
instance {V} {G : quiver V} : inhabited (wide_subquiver G) := ⟨⊤⟩
63+
/-- A wide subquiver viewed as a quiver on its own. -/
64+
instance wide_subquiver.quiver {V} [quiver V] (H : wide_subquiver V) : quiver H :=
65+
⟨λ a b, H a b⟩
4266

43-
/-- `G.sum H` takes the disjoint union of the arrows of `G` and `H`. -/
44-
protected def sum {V} (G H : quiver V) : quiver V :=
45-
⟨λ a b, G.arrow a b ⊕ H.arrow a b⟩
67+
namespace quiver
4668

47-
/-- `G.opposite` reverses the direction of all arrows of `G`. -/
48-
protected def opposite {V} (G : quiver V) : quiver V :=
49-
⟨flip G.arrow⟩
69+
/-- A type synonym for a quiver with no arrows. -/
70+
@[nolint has_inhabited_instance]
71+
def empty (V) : Type u := V
5072

51-
/-- `G.symmetrify` adds to `G` the reversal of all arrows of `G`. -/
52-
def symmetrify {V} (G : quiver V) : quiver V :=
53-
G.sum G.opposite
73+
instance empty_quiver (V : Type u) : quiver.{u} (empty V) := ⟨λ a b, pempty⟩
5474

55-
@[simp] lemma empty_arrow {V} (a b : V) : (quiver.empty V).arrow a b = pempty := rfl
75+
@[simp] lemma empty_arrow {V : Type u} (a b : empty V) : (a ⟶ b) = pempty := rfl
5676

57-
@[simp] lemma sum_arrow {V} (G H : quiver V) (a b : V) :
58-
(G.sum H).arrow a b = (G.arrow a b ⊕ H.arrow a b) := rfl
77+
instance {V} [quiver V] : has_bot (wide_subquiver V) := ⟨λ a b, ∅⟩
78+
instance {V} [quiver V] : has_top (wide_subquiver V) := ⟨λ a b, set.univ⟩
79+
instance {V} [quiver V] : inhabited (wide_subquiver V) := ⟨⊤⟩
5980

60-
@[simp] lemma opposite_arrow {V} (G : quiver V) (a b : V) :
61-
G.opposite.arrow a b = G.arrow b a := rfl
81+
/-- `Vᵒᵖ` reverses the direction of all arrows of `V`. -/
82+
instance op_quiver {V} [quiver V] : quiver Vᵒᵖ :=
83+
⟨λ a b, (unop b) ⟶ (unop a)⟩
6284

63-
@[simp] lemma symmetrify_arrow {V} (G : quiver V) (a b : V) :
64-
G.symmetrify.arrow a b = (G.arrow a b ⊕ G.arrow b a) := rfl
85+
/-- A type synonym for the symmetrized quiver (with an arrow both ways for each original arrow).
86+
NB: this does not work for `Prop`-valued quivers. It requires `[quiver.{v+1} V]`. -/
87+
@[nolint has_inhabited_instance]
88+
def symmetrify (V) : Type u := V
6589

66-
@[simp] lemma opposite_opposite {V} (G : quiver V) : G.opposite.opposite = G :=
67-
by { cases G, refl }
90+
instance symmetrify_quiver (V : Type u) [quiver V] : quiver (symmetrify V) :=
91+
⟨λ a b : V, (a ⟶ b) ⊕ (b ⟶ a)⟩
6892

69-
/-- `G.total` is the type of _all_ arrows of `G`. -/
70-
@[ext] protected structure total {V} (G : quiver.{v u} V) : Sort (max (u+1) v) :=
93+
/-- `total V` is the type of _all_ arrows of `V`. -/
94+
@[ext, nolint has_inhabited_instance]
95+
structure total (V : Type u) [quiver.{v} V] : Type (max u v) :=
7196
(source : V)
7297
(target : V)
73-
(arrow : G.arrow source target)
74-
75-
instance {V} [inhabited V] {G : quiver V} [inhabited (G.arrow (default V) (default V))] :
76-
inhabited G.total :=
77-
⟨⟨default V, default V, default _⟩⟩
98+
(arrow : source ⟶ target)
7899

79100
/-- A wide subquiver `H` of `G.symmetrify` determines a wide subquiver of `G`, containing an
80101
an arrow `e` if either `e` or its reversal is in `H`. -/
81-
def wide_subquiver_symmetrify {V} {G : quiver V} :
82-
wide_subquiver G.symmetrify → wide_subquiver G :=
102+
-- Without the explicit universe level in `quiver.{v+1}` Lean comes up with
103+
-- `quiver.{max u_2 u_3 + 1}`. This causes problems elsewhere, so we write `quiver.{v+1}`.
104+
def wide_subquiver_symmetrify {V} [quiver.{v+1} V] :
105+
wide_subquiver (symmetrify V) → wide_subquiver V :=
83106
λ H a b, { e | sum.inl e ∈ H a b ∨ sum.inr e ∈ H b a }
84107

85108
/-- A wide subquiver of `G` can equivalently be viewed as a total set of arrows. -/
86-
def wide_subquiver_equiv_set_total {V} {G : quiver V} :
87-
wide_subquiver G ≃ set G.total :=
109+
def wide_subquiver_equiv_set_total {V} [quiver V] :
110+
wide_subquiver V ≃ set (total V) :=
88111
{ to_fun := λ H, { e | e.arrow ∈ H e.source e.target },
89112
inv_fun := λ S a b, { e | total.mk a b e ∈ S },
90113
left_inv := λ H, rfl,
91114
right_inv := by { intro S, ext, cases x, refl } }
92115

93116
/-- `G.path a b` is the type of paths from `a` to `b` through the arrows of `G`. -/
94-
inductive path {V} (G : quiver.{v u} V) (a : V) : V → Sort (max (u+1) v)
117+
inductive path {V : Type u} [quiver.{v} V] (a : V) : V → Type (max u v)
95118
| nil : path a
96-
| cons : Π {b c : V}, path b → G.arrow b c → path c
119+
| cons : Π {b c : V}, path b → (b ⟶ c) → path c
97120

98121
/-- An arrow viewed as a path of length one. -/
99-
def arrow.to_path {V} {G : quiver V} {a b} (e : G.arrow a b) : G.path a b :=
122+
def hom.to_path {V} [quiver V] {a b : V} (e : a ⟶ b) : path a b :=
100123
path.nil.cons e
101124

102125
namespace path
103126

104-
variables {V : Type*} {G : quiver V}
127+
variables {V : Type u} [quiver V]
105128

106129
/-- The length of a path is the number of arrows it uses. -/
107-
def length {a : V} : Π {b}, G.path a b → ℕ
130+
def length {a : V} : Π {b : V}, path a b → ℕ
108131
| _ path.nil := 0
109132
| _ (path.cons p _) := p.length + 1
110133

111134
@[simp] lemma length_nil {a : V} :
112-
(path.nil : G.path a a).length = 0 := rfl
135+
(path.nil : path a a).length = 0 := rfl
113136

114-
@[simp] lemma length_cons (a b c : V) (p : G.path a b)
115-
(e : G.arrow b c) : (p.cons e).length = p.length + 1 := rfl
137+
@[simp] lemma length_cons (a b c : V) (p : path a b)
138+
(e : b ⟶ c) : (p.cons e).length = p.length + 1 := rfl
116139

117140
/-- Composition of paths. -/
118-
def comp {a b} : Π {c}, G.path a b → G.path b c → G.path a c
141+
def comp {a b : V} : Π {c}, path a b → path b c → path a c
119142
| _ p (path.nil) := p
120143
| _ p (path.cons q e) := (p.comp q).cons e
121144

122-
@[simp] lemma comp_cons {a b c d} (p : G.path a b) (q : G.path b c) (e : G.arrow c d) :
145+
@[simp] lemma comp_cons {a b c d : V} (p : path a b) (q : path b c) (e : c ⟶ d) :
123146
p.comp (q.cons e) = (p.comp q).cons e := rfl
124-
@[simp] lemma comp_nil {a b} (p : G.path a b) : p.comp path.nil = p := rfl
125-
@[simp] lemma nil_comp {a} : ∀ {b} (p : G.path a b), path.nil.comp p = p
147+
@[simp] lemma comp_nil {a b : V} (p : path a b) : p.comp path.nil = p := rfl
148+
@[simp] lemma nil_comp {a : V} : ∀ {b} (p : path a b), path.nil.comp p = p
126149
| a path.nil := rfl
127150
| b (path.cons p e) := by rw [comp_cons, nil_comp]
128-
@[simp] lemma comp_assoc {a b c} : ∀ {d}
129-
(p : G.path a b) (q : G.path b c) (r : G.path c d),
151+
@[simp] lemma comp_assoc {a b c : V} : ∀ {d}
152+
(p : path a b) (q : path b c) (r : path c d),
130153
(p.comp q).comp r = p.comp (q.comp r)
131154
| c p q path.nil := rfl
132155
| d p q (path.cons r e) := by rw [comp_cons, comp_cons, comp_cons, comp_assoc]
@@ -135,33 +158,33 @@ end path
135158

136159
/-- A quiver is an arborescence when there is a unique path from the default vertex
137160
to every other vertex. -/
138-
class arborescence {V} (T : quiver.{v u} V) : Sort (max (u+1) v) :=
161+
class arborescence (V : Type u) [quiver.{v} V] : Type (max u v) :=
139162
(root : V)
140-
(unique_path : Π (b : V), unique (T.path root b))
163+
(unique_path : Π (b : V), unique (path root b))
141164

142165
/-- The root of an arborescence. -/
143-
protected def root {V} (T : quiver V) [arborescence T] : V :=
144-
arborescence.root T
166+
def root (V : Type u) [quiver V] [arborescence V] : V :=
167+
arborescence.root
145168

146-
instance {V} (T : quiver V) [arborescence T] (b : V) : unique (T.path T.root b) :=
169+
instance {V : Type u} [quiver V] [arborescence V] (b : V) : unique (path (root V) b) :=
147170
arborescence.unique_path b
148171

149172
/-- An `L`-labelling of a quiver assigns to every arrow an element of `L`. -/
150-
def labelling {V} (G : quiver V) (L) := Π a b, G.arrow a b → L
173+
def labelling (V : Type u) [quiver V] (L : Sort*) := Π a b : V, (a ⟶ b) → L
151174

152-
instance {V} (G : quiver V) (L) [inhabited L] : inhabited (G.labelling L) :=
175+
instance {V : Type u} [quiver V] (L) [inhabited L] : inhabited (labelling V L) :=
153176
⟨λ a b e, default L⟩
154177

155-
/-- To show that `T : quiver V` is an arborescence with root `r : V`, it suffices to
178+
/-- To show that `[quiver V]` is an arborescence with root `r : V`, it suffices to
156179
- provide a height function `V → ℕ` such that every arrow goes from a
157180
lower vertex to a higher vertex,
158181
- show that every vertex has at most one arrow to it, and
159182
- show that every vertex other than `r` has an arrow to it. -/
160-
noncomputable def arborescence_mk {V} (T : quiver V) (r : V)
183+
noncomputable def arborescence_mk {V : Type u} [quiver V] (r : V)
161184
(height : V → ℕ)
162-
(height_lt : ∀ ⦃a b⦄, T.arrow a b → height a < height b)
163-
(unique_arrow : ∀ ⦃a b c⦄ (e : T.arrow a c) (f : T.arrow b c), a = b ∧ e == f)
164-
(root_or_arrow : ∀ b, b = r ∨ ∃ a, nonempty (T.arrow a b)) : arborescence T :=
185+
(height_lt : ∀ ⦃a b⦄, (a ⟶ b) → height a < height b)
186+
(unique_arrow : ∀ ⦃a b c : V⦄ (e : a ⟶ c) (f : b ⟶ c), a = b ∧ e == f)
187+
(root_or_arrow : ∀ b, b = r ∨ ∃ a, nonempty (a ⟶ b)) : arborescence V :=
165188
{ root := r,
166189
unique_path := λ b, ⟨classical.inhabited_of_nonempty
167190
begin
@@ -174,10 +197,10 @@ noncomputable def arborescence_mk {V} (T : quiver V) (r : V)
174197
exact ⟨p.cons e⟩ }
175198
end,
176199
begin
177-
have height_le : ∀ {a b}, T.path a b → height a ≤ height b,
200+
have height_le : ∀ {a b}, path a b → height a ≤ height b,
178201
{ intros a b p, induction p with b c p e ih, refl,
179202
exact le_of_lt (lt_of_le_of_lt ih (height_lt e)) },
180-
suffices : ∀ p q : T.path r b, p = q,
203+
suffices : ∀ p q : path r b, p = q,
181204
{ intro p, apply this },
182205
intros p q, induction p with a c p e ih; cases q with b _ q f,
183206
{ refl },
@@ -186,85 +209,85 @@ noncomputable def arborescence_mk {V} (T : quiver V) (r : V)
186209
{ rcases unique_arrow e f with ⟨⟨⟩, ⟨⟩⟩, rw ih },
187210
end ⟩ }
188211

189-
/-- `G.rooted_connected r` means that there is a path from `r` to any other vertex. -/
190-
class rooted_connected {V} (G : quiver V) (r : V) : Prop :=
191-
(nonempty_path : ∀ b : V, nonempty (G.path r b))
212+
/-- `rooted_connected r` means that there is a path from `r` to any other vertex. -/
213+
class rooted_connected {V : Type u} [quiver V] (r : V) : Prop :=
214+
(nonempty_path : ∀ b : V, nonempty (path r b))
192215

193216
attribute [instance] rooted_connected.nonempty_path
194217

195218
section geodesic_subtree
196219

197-
variables {V : Type*} (G : quiver.{v+1} V) (r : V) [G.rooted_connected r]
220+
variables {V : Type u} [quiver.{v+1} V] (r : V) [rooted_connected r]
198221

199222
/-- A path from `r` of minimal length. -/
200-
noncomputable def shortest_path (b : V) : G.path r b :=
223+
noncomputable def shortest_path (b : V) : path r b :=
201224
well_founded.min (measure_wf path.length) set.univ set.univ_nonempty
202225

203226
/-- The length of a path is at least the length of the shortest path -/
204-
lemma shortest_path_spec {a : V} (p : G.path r a) :
205-
(G.shortest_path r a).length ≤ p.length :=
227+
lemma shortest_path_spec {a : V} (p : path r a) :
228+
(shortest_path r a).length ≤ p.length :=
206229
not_lt.mp (well_founded.not_lt_min (measure_wf _) set.univ _ trivial)
207230

208231
/-- A subquiver which by construction is an arborescence. -/
209-
def geodesic_subtree : wide_subquiver G :=
210-
λ a b, { e | ∃ p : G.path r a, shortest_path G r b = p.cons e }
232+
def geodesic_subtree : wide_subquiver V :=
233+
λ a b, { e | ∃ p : path r a, shortest_path r b = p.cons e }
211234

212-
noncomputable instance geodesic_arborescence : (G.geodesic_subtree r).quiver.arborescence :=
213-
arborescence_mk _ r (λ a, (G.shortest_path r a).length)
235+
noncomputable instance geodesic_arborescence : arborescence (geodesic_subtree r) :=
236+
arborescence_mk r (λ a, (shortest_path r a).length)
214237
(by { rintros a b ⟨e, p, h⟩,
215238
rw [h, path.length_cons, nat.lt_succ_iff], apply shortest_path_spec })
216239
(by { rintros a b c ⟨e, p, h⟩ ⟨f, q, j⟩, cases h.symm.trans j, split; refl })
217-
(by { intro b, have : ∃ p, G.shortest_path r b = p := ⟨_, rfl⟩,
240+
(by { intro b, have : ∃ p, shortest_path r b = p := ⟨_, rfl⟩,
218241
rcases this with ⟨p, hp⟩, cases p with a _ p e,
219242
{ exact or.inl rfl }, { exact or.inr ⟨a, ⟨⟨e, p, hp⟩⟩⟩ } })
220243

221244
end geodesic_subtree
222245

223-
variables {V : Type*} (G : quiver V)
246+
variables (V : Type u) [quiver.{v+1} V]
224247

225248
/-- A quiver `has_reverse` if we can reverse an arrow `p` from `a` to `b` to get an arrow
226249
`p.reverse` from `b` to `a`.-/
227250
class has_reverse :=
228-
(reverse' : Π {a b}, G.arrow a bG.arrow b a)
251+
(reverse' : Π {a b : V}, (a ⟶ b)(b ⟶ a))
229252

230-
variables {G} [has_reverse G]
253+
instance : has_reverse (symmetrify V) := ⟨λ a b e, e.swap⟩
254+
255+
variables {V} [has_reverse V]
231256

232257
/-- Reverse the direction of an arrow. -/
233-
def arrow.reverse {a b} : G.arrow a bG.arrow b a := has_reverse.reverse'
258+
def reverse {a b : V} : (a ⟶ b)(b ⟶ a) := has_reverse.reverse'
234259

235260
/-- Reverse the direction of a path. -/
236-
def path.reverse {a} : Π {b}, G.path a b → G.path b a
261+
def path.reverse {a : V} : Π {b}, path a b → path b a
237262
| a path.nil := path.nil
238-
| b (path.cons p e) := e.reverse.to_path.comp p.reverse
239-
240-
variable (H : quiver.{v+1} V)
263+
| b (path.cons p e) := (reverse e).to_path.comp p.reverse
241264

242-
instance : has_reverse H.symmetrify := ⟨λ a b e, e.swap⟩
265+
variables (V)
243266

244267
/-- Two vertices are related in the zigzag setoid if there is a
245268
zigzag of arrows from one to the other. -/
246269
def zigzag_setoid : setoid V :=
247-
⟨λ a b, nonempty (H.symmetrify.path a b),
270+
⟨λ a b, nonempty (path (a : symmetrify V) (b : symmetrify V)),
248271
λ a, ⟨path.nil⟩,
249272
λ a b ⟨p⟩, ⟨p.reverse⟩,
250273
λ a b c ⟨p⟩ ⟨q⟩, ⟨p.comp q⟩⟩
251274

252275
/-- The type of weakly connected components of a directed graph. Two vertices are
253276
in the same weakly connected component if there is a zigzag of arrows from one
254277
to the other. -/
255-
def weakly_connected_component : Type* := quotient H.zigzag_setoid
278+
def weakly_connected_component : Type* := quotient (zigzag_setoid V)
256279

257280
namespace weakly_connected_component
258-
variable {H}
281+
variable {V}
259282

260283
/-- The weakly connected component corresponding to a vertex. -/
261-
protected def mk : V → H.weakly_connected_component := quotient.mk'
284+
protected def mk : V → weakly_connected_component V := quotient.mk'
262285

263-
instance : has_coe_t V H.weakly_connected_component := ⟨weakly_connected_component.mk⟩
264-
instance [inhabited V] : inhabited H.weakly_connected_component := ⟨↑(default V)⟩
286+
instance : has_coe_t V (weakly_connected_component V) := ⟨weakly_connected_component.mk⟩
287+
instance [inhabited V] : inhabited (weakly_connected_component V) := ⟨↑(default V)⟩
265288

266289
protected lemma eq (a b : V) :
267-
(a : H.weakly_connected_component) = b ↔ nonempty (H.symmetrify.path a b) :=
290+
(a : weakly_connected_component V) = b ↔ nonempty (path (a : symmetrify V) (b : symmetrify V)) :=
268291
quotient.eq'
269292

270293
end weakly_connected_component

0 commit comments

Comments
 (0)