Skip to content

Commit af93324

Browse files
committed
feat: port unital shelves to Algebra.Quandle (#2393)
* [`algebra.quandle`@`8631e2d5ea77f6c13054d9151d82b83069680cb1`..`28aa996fc6fb4317f0083c4e6daf79878d81be33`](https://leanprover-community.github.io/mathlib-port-status/file/algebra/quandle?range=8631e2d5ea77f6c13054d9151d82b83069680cb1..28aa996fc6fb4317f0083c4e6daf79878d81be33)
1 parent 0422d4a commit af93324

File tree

1 file changed

+48
-2
lines changed

1 file changed

+48
-2
lines changed

Mathlib/Algebra/Quandle.lean

Lines changed: 48 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Kyle Miller
55
66
! This file was ported from Lean 3 source module algebra.quandle
7-
! leanprover-community/mathlib commit 8631e2d5ea77f6c13054d9151d82b83069680cb1
7+
! leanprover-community/mathlib commit 28aa996fc6fb4317f0083c4e6daf79878d81be33
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -36,7 +36,7 @@ complements that is analogous to the fundamental group of the
3636
exterior, and he showed that the quandle associated to an oriented
3737
knot is invariant up to orientation-reversed mirror image. Racks were
3838
used by Fenn and Rourke for framed codimension-2 knots and
39-
links.[FennRourke1992]
39+
links in [FennRourke1992]. Unital shelves are discussed in [crans2017].
4040
4141
The name "rack" came from wordplay by Conway and Wraith for the "wrack
4242
and ruin" of forgetting everything but the conjugation operation for a
@@ -45,6 +45,7 @@ group.
4545
## Main definitions
4646
4747
* `Shelf` is a type with a self-distributive action
48+
* `UnitalShelf` is a shelf with a left and right unit
4849
* `Rack` is a shelf whose action for each element is invertible
4950
* `Quandle` is a rack whose action for an element fixes that element
5051
* `Quandle.conj` defines a quandle of a group acting on itself by conjugation.
@@ -56,6 +57,10 @@ group.
5657
* `Rack.EnvelGroup` is left adjoint to `Quandle.Conj` (`toEnvelGroup.map`).
5758
The universality statements are `toEnvelGroup.univ` and `toEnvelGroup.univ_uniq`.
5859
60+
## Implementation notes
61+
"Unital racks" are uninteresting (see `Rack.assoc_iff_id`, `UnitalShelf.assoc`), so we do not
62+
define them.
63+
5964
## Notation
6065
6166
The following notation is localized in `quandles`:
@@ -98,6 +103,15 @@ class Shelf (α : Type u) where
98103
self_distrib : ∀ {x y z : α}, act x (act y z) = act (act x y) (act x z)
99104
#align shelf Shelf
100105

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+
101115
/-- The type of homomorphisms between shelves.
102116
This is also the notion of rack and quandle homomorphisms.
103117
-/
@@ -138,6 +152,38 @@ scoped[Quandles] infixr:25 " →◃ " => ShelfHom
138152

139153
open Quandles
140154

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+
141187
namespace Rack
142188

143189
variable {R : Type _} [Rack R]

0 commit comments

Comments
 (0)