@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Kyle Miller
5
5
6
6
! This file was ported from Lean 3 source module algebra.quandle
7
- ! leanprover-community/mathlib commit 8631e2d5ea77f6c13054d9151d82b83069680cb1
7
+ ! leanprover-community/mathlib commit 28aa996fc6fb4317f0083c4e6daf79878d81be33
8
8
! Please do not edit these lines, except to modify the commit id
9
9
! if you have ported upstream changes.
10
10
-/
@@ -36,7 +36,7 @@ complements that is analogous to the fundamental group of the
36
36
exterior, and he showed that the quandle associated to an oriented
37
37
knot is invariant up to orientation-reversed mirror image. Racks were
38
38
used by Fenn and Rourke for framed codimension-2 knots and
39
- links. [ FennRourke1992 ]
39
+ links in [ FennRourke1992 ] . Unital shelves are discussed in [ crans2017 ] .
40
40
41
41
The name "rack" came from wordplay by Conway and Wraith for the "wrack
42
42
and ruin" of forgetting everything but the conjugation operation for a
45
45
## Main definitions
46
46
47
47
* `Shelf` is a type with a self-distributive action
48
+ * `UnitalShelf` is a shelf with a left and right unit
48
49
* `Rack` is a shelf whose action for each element is invertible
49
50
* `Quandle` is a rack whose action for an element fixes that element
50
51
* `Quandle.conj` defines a quandle of a group acting on itself by conjugation.
@@ -56,6 +57,10 @@ group.
56
57
* `Rack.EnvelGroup` is left adjoint to `Quandle.Conj` (`toEnvelGroup.map`).
57
58
The universality statements are `toEnvelGroup.univ` and `toEnvelGroup.univ_uniq`.
58
59
60
+ ## Implementation notes
61
+ "Unital racks" are uninteresting (see `Rack.assoc_iff_id`, `UnitalShelf.assoc`), so we do not
62
+ define them.
63
+
59
64
## Notation
60
65
61
66
The following notation is localized in `quandles`:
@@ -98,6 +103,15 @@ class Shelf (α : Type u) where
98
103
self_distrib : ∀ {x y z : α}, act x (act y z) = act (act x y) (act x z)
99
104
#align shelf Shelf
100
105
106
+ /--
107
+ A *unital shelf* is a shelf equipped with an element `1` such that, for all elements `x`,
108
+ we have both `x ◃ 1` and `1 ◃ x` equal `x`.
109
+ -/
110
+ class UnitalShelf (α : Type u) extends Shelf α, One α :=
111
+ (one_act : ∀ a : α, act 1 a = a)
112
+ (act_one : ∀ a : α, act a 1 = a)
113
+ #align unital_shelf UnitalShelf
114
+
101
115
/-- The type of homomorphisms between shelves.
102
116
This is also the notion of rack and quandle homomorphisms.
103
117
-/
@@ -138,6 +152,38 @@ scoped[Quandles] infixr:25 " →◃ " => ShelfHom
138
152
139
153
open Quandles
140
154
155
+ namespace UnitalShelf
156
+ open Shelf
157
+
158
+ variable {S : Type _} [UnitalShelf S]
159
+
160
+ /--
161
+ A monoid is *graphic* if, for all `x` and `y`, the *graphic identity*
162
+ `(x * y) * x = x * y` holds. For a unital shelf, this graphic
163
+ identity holds.
164
+ -/
165
+ lemma act_act_self_eq (x y : S) : (x ◃ y) ◃ x = x ◃ y := by
166
+ have h : (x ◃ y) ◃ x = (x ◃ y) ◃ (x ◃ 1 ) := by rw [act_one]
167
+ rw [h, ←Shelf.self_distrib, act_one]
168
+ #align unital_shelf.act_act_self_eq UnitalShelf.act_act_self_eq
169
+
170
+ lemma act_idem (x : S) : (x ◃ x) = x := by rw [←act_one x, ←Shelf.self_distrib, act_one, act_one]
171
+ #align unital_shelf.act_idem UnitalShelf.act_idem
172
+
173
+ lemma act_self_act_eq (x y : S) : x ◃ (x ◃ y) = x ◃ y := by
174
+ have h : x ◃ (x ◃ y) = (x ◃ 1 ) ◃ (x ◃ y) := by rw [act_one]
175
+ rw [h, ←Shelf.self_distrib, one_act]
176
+ #align unital_shelf.act_self_act_eq UnitalShelf.act_self_act_eq
177
+
178
+ /--
179
+ The associativity of a unital shelf comes for free.
180
+ -/
181
+ lemma assoc (x y z : S) : (x ◃ y) ◃ z = x ◃ y ◃ z := by
182
+ rw [self_distrib, self_distrib, act_act_self_eq, act_self_act_eq]
183
+ #align unital_shelf.assoc UnitalShelf.assoc
184
+
185
+ end UnitalShelf
186
+
141
187
namespace Rack
142
188
143
189
variable {R : Type _} [Rack R]
0 commit comments