module Algebra.Quasigroup where
Quasigroupsπ
Traditionally, groups are defined as monoids where every element has a (necessarily unique) inverse. However, there is another path to the theory of groups that emphasizes division/subtraction over inverses. This perspective is especially interesting when generalizing groups to non-associative settings; axioms of the form are rather ill-behaved without associativity, as inverses are no longer necessarily unique!
private variable
: Level
β β' : Type β
A B
open Cat.Displayed.Univalence.Thin using (Extensional-Hom) public
Right quasigroupsπ
For the sake of maximum generality, we will build up to the definition of quasigroups in stages, starting with right quasigroups.
Let be a binary operator. is a right quasigroup if is a magma, and there is some binary operator subject to the following axioms:
- For all
- For all
record is-right-quasigroup {β} {A : Type β} (_β_ : A β A β A) : Type β where
no-eta-equality
field
_/_ : A β A β A
: β x y β (x / y) β y β‘ x
/-invl : β x y β (x β y) / y β‘ x
/-invr : is-magma _β_
has-is-magma
open is-magma has-is-magma public
Intuitively, the operation acts like a sort of right-biased division. This is further cemented by noting that the existence of such an operation implies that every the action is an equivalence.
: β x β is-equiv (_β x)
β-equivr .is-eqv y .centre = y / x , /-invl y x
β-equivr x .is-eqv y .paths (l , lx=y) =
β-equivr x (Ξ» l β has-is-set (l β x) y) $
Ξ£-prop-path (_/ x) lx=y β©
y / x β‘Λβ¨ ap (l β x) / x β‘β¨ /-invr l x β©
l β
This in turn implies that is an equivalence.
: β x β is-equiv (_/ x)
/-equivr = inverse-is-equiv (β-equivr x) /-equivr x
As easy corollaries, we get that multiplication and division in are right-cancellative.
opaque: β {x y} (z : A) β x β z β‘ y β z β x β‘ y
β-cancelr = Equiv.injective (_β z , β-equivr z)
β-cancelr z
: β {x y} (z : A) β x / z β‘ y / z β x β‘ y
/-cancelr = Equiv.injective (_/ z , /-equivr z) /-cancelr z
It turns out that being an equivalence is a sufficient condition for a to be a right quasigroup, provided that is a set.
β-equivrβis-right-quasigroup: β {_β_ : A β A β A}
β is-set A
β (β x β is-equiv (_β x))
β is-right-quasigroup _β_
The proof is an exercise in unrolling definitions. If is an equivalence, then serves as a perfectly valid definition of both axioms follow directly from the unit and counit of the equivalence.
{_β_ = _β_} A-set β-eqv =
β-equivrβis-right-quasigroup
β-right-quasiwhere
open is-right-quasigroup
module β-eqv x = Equiv (_β x , β-eqv x)
: is-right-quasigroup _β_
β-right-quasi ._/_ x y = β-eqv.from y x
β-right-quasi ./-invl x y = β-eqv.Ξ΅ y x
β-right-quasi ./-invr x y = β-eqv.Ξ· y x
β-right-quasi .has-is-magma .is-magma.has-is-set = A-set β-right-quasi
We can upgrade this correspondence to an equivalence, as we definitionally get back the same division operation we started with after round-tripping.
is-right-quasigroupββ-equivr: β {_β_ : A β A β A}
β is-right-quasigroup _β_ β (is-set A Γ (β x β is-equiv (_β x)))
{_β_ = _β_} =
is-right-quasigroupββ-equivr
IsoβEquiv $
β¨ has-is-set , β-equivr β© ,(uncurry β-equivrβis-right-quasigroup)
iso (Ξ» _ β prop!)
(Ξ» β-right-quasi β
let open is-right-quasigroup β-right-quasi in
.injective eqv (refl ,β prop!))
Isowhere
open is-right-quasigroup
unquoteDecl eqv = declare-record-iso eqv (quote is-right-quasigroup)
Crucially, this means that the division operation
is actually a property rather
than structure, as both is-equiv
and is-set
are propositions.
is-right-quasigroup-is-prop: β {_β_ : A β A β A}
β is-prop (is-right-quasigroup _β_)
{A = A} =
is-right-quasigroup-is-prop 1 is-right-quasigroupββ-equivr (hlevel 1) Equivβis-hlevel
instance
H-Level-is-right-quasigroup: β {_β_ : A β A β A} {n}
β H-Level (is-right-quasigroup _β_) (suc n)
= prop-instance is-right-quasigroup-is-prop H-Level-is-right-quasigroup
Right quasigroup homomorphismsπ
In the previous section, we proved that the existence of a right division operation was actually a property of a magma, rather than structure. We shall now see that right division is automatically preserved by magma homomorphisms.
We start by defining the corresponding notion of structure for right quasigroups.
A right quasigroup structure on a type is a binary operation such that is a right quasigroup.
record Right-quasigroup-on {β} (A : Type β) : Type β where
no-eta-equality
field
_β_ : A β A β A
: is-right-quasigroup _β_
has-is-right-quasigroup
open is-right-quasigroup has-is-right-quasigroup public
A right quasigroup homomorphism between two right quasigroups is a function such that
record is-right-quasigroup-hom
{β β'} {A : Type β} {B : Type β'}
(f : A β B)
(S : Right-quasigroup-on A) (T : Right-quasigroup-on B)
: Type (β β β')
where
no-eta-equality
private
module A = Right-quasigroup-on S
module B = Right-quasigroup-on T
field
: β x y β f (x A.β y) β‘ f x B.β f y pres-β
Preservation of right division follows almost immediately from right-cancellativity.
: β x y β f (x A./ y) β‘ f x B./ f y
pres-/ =
pres-/ x y .β-cancelr (f y) $
B(x A./ y) B.β f y β‘Λβ¨ pres-β (x A./ y) y β©
f ((x A./ y) A.β y) β‘β¨ ap f (A./-invl x y) β©
f ./-invl (f x) (f y) β©
f x β‘Λβ¨ B(f x B./ f y) B.β f y β
is-right-quasigroup-hom-is-prop: {f : A β B}
β {S : Right-quasigroup-on A} {T : Right-quasigroup-on B}
β is-prop (is-right-quasigroup-hom f S T)
{T = T} =
is-right-quasigroup-hom-is-prop 1 eqv
Isoβis-hlevel! where
open Right-quasigroup-on T
private unquoteDecl eqv = declare-record-iso eqv (quote is-right-quasigroup-hom)
instance
H-Level-is-right-quasigroup-hom: {f : A β B}
β {S : Right-quasigroup-on A} {T : Right-quasigroup-on B}
β β {n} β H-Level (is-right-quasigroup-hom f S T) (suc n)
= prop-instance is-right-quasigroup-hom-is-prop H-Level-is-right-quasigroup-hom
module _ where
open is-right-quasigroup-hom
open Right-quasigroup-on
Right-quasigroup-on-pathp: β {A : I β Type β}
β {S : Right-quasigroup-on (A i0)} {T : Right-quasigroup-on (A i1)}
β PathP (Ξ» i β β (x y : A i) β A i) (S ._β_) (T ._β_)
β PathP (Ξ» i β Right-quasigroup-on (A i)) S T
{S = S} {T = T} p i ._β_ x y =
Right-quasigroup-on-pathp
p i x y{S = S} {T = T} p i .has-is-right-quasigroup =
Right-quasigroup-on-pathp
is-propβpathp(Ξ» i β is-right-quasigroup-is-prop {_β_ = p i})
(S .has-is-right-quasigroup)
(T .has-is-right-quasigroup)
i
Right quasigroups are algebraic structures, so they form a thin structure over
: β β β Thin-structure β Right-quasigroup-on
Right-quasigroup-structure .is-hom f A B .β£_β£ =
Right-quasigroup-structure β
is-right-quasigroup-hom f A B.is-hom f A B .is-tr =
Right-quasigroup-structure β
is-right-quasigroup-hom-is-prop.id-is-hom .pres-β x y =
Right-quasigroup-structure β
refl.β-is-hom f g f-hom g-hom .pres-β x y =
Right-quasigroup-structure β (g-hom .pres-β x y) β f-hom .pres-β (g x) (g y)
ap f .id-hom-unique {A} {S} {T} p q =
Right-quasigroup-structure β (ext (p .pres-β)) Right-quasigroup-on-pathp
This observation lets us neatly organize right quasigroups into a category.
: β β β Precategory (lsuc β) β
Right-quasigroups = Structured-objects (Right-quasigroup-structure β)
Right-quasigroups β
module Right-quasigroups {β} = Cat.Reasoning (Right-quasigroups β)
: (β : Level) β Type (lsuc β)
Right-quasigroup = Right-quasigroups.Ob {β} Right-quasigroup β
module Right-quasigroup {β} (A : Right-quasigroup β) where
open Right-quasigroup-on (A .snd) public
Left quasigroupsπ
We can dualise the definition of right quasigroups to arrive at the notion of a left quasigroup.
Let be a binary operator. is a left quasigroup if is a magma, and there is some binary operator subject to the following axioms:
- For all
- For all
record is-left-quasigroup {β} {A : Type β} (_β_ : A β A β A) : Type β where
no-eta-equality
field
_⧡_ : A β A β A
: β x y β x β (x ⧡ y) β‘ y
⧡-invr : β x y β x ⧡ (x β y) β‘ y
⧡-invl : is-magma _β_
has-is-magma
open is-magma has-is-magma public
Intuitively, we should think of as β divided by β or β over β. This is reinforced by the fact that for every the action is an equivalence.
: β x β is-equiv (x β_)
β-equivl .is-eqv y .centre = x ⧡ y , ⧡-invr x y
β-equivl x .is-eqv y .paths (r , xr=y) =
β-equivl x (Ξ» r β has-is-set (x β r) y) $
Ξ£-prop-path (x ⧡_) xr=y β©
x ⧡ y β‘Λβ¨ ap (x β r) β‘⨠⧡-invl x r β©
x ⧡
r β
: β x β is-equiv (x ⧡_)
⧡-equivl = inverse-is-equiv (β-equivl x) ⧡-equivl x
This implies that that multiplication and division in is left-cancellative.
opaque: β (x : A) {y z} β x β y β‘ x β z β y β‘ z
β-cancell = Equiv.injective (x β_ , β-equivl x)
β-cancell x
: β (x : A) {y z} β x ⧡ y β‘ x ⧡ z β y β‘ z
⧡-cancell = Equiv.injective (x ⧡_ , ⧡-equivl x) ⧡-cancell x
Next, we shall show that left quasigroups are formally dual to right quasigroups.
is-left-quasigroupβop-is-right-quasigroup: β {_β_ : A β A β A}
β is-left-quasigroup _β_ β is-right-quasigroup (Ξ» x y β y β x)
The proof of this fact is rather tedious, so we will omit the details.
is-left-quasigroupβop-is-right-quasigroup: β {_β_ : A β A β A}
β is-left-quasigroup _β_ β is-right-quasigroup (Ξ» x y β y β x)
is-right-quasigroupβop-is-left-quasigroup: β {_β_ : A β A β A}
β is-right-quasigroup _β_ β is-left-quasigroup (Ξ» x y β y β x)
{_β_ = _β_} A-quasi = A-op-quasi
is-left-quasigroupβop-is-right-quasigroup where
open is-left-quasigroup A-quasi
open is-right-quasigroup
: is-right-quasigroup (Ξ» x y β y β x)
A-op-quasi ._/_ x y = y ⧡ x
A-op-quasi ./-invr = flip ⧡-invl
A-op-quasi ./-invl = flip ⧡-invr
A-op-quasi .has-is-magma .is-magma.has-is-set = hlevel 2
A-op-quasi
{_β_ = _β_} A-quasi = A-op-quasi
is-right-quasigroupβop-is-left-quasigroup where
open is-right-quasigroup A-quasi
open is-left-quasigroup
: is-left-quasigroup (Ξ» x y β y β x)
A-op-quasi ._⧡_ x y = y / x
A-op-quasi .⧡-invr = flip /-invl
A-op-quasi .⧡-invl = flip /-invr
A-op-quasi .has-is-magma .is-magma.has-is-set = hlevel 2
A-op-quasi
=
is-left-quasigroupβop-is-right-quasigroup
IsoβEquiv $
is-left-quasigroupβop-is-right-quasigroup ,
iso is-right-quasigroupβop-is-left-quasigroup(Ξ» _ β prop!)
(Ξ» β-left-quasi β
let open is-left-quasigroup β-left-quasi in
.injective eqv (refl ,β prop!))
Isowhere
private unquoteDecl eqv = declare-record-iso eqv (quote is-left-quasigroup)
This in turn means that being a left quasigroup is a property of a binary operation.
is-left-quasigroup-is-prop: β {_β_ : A β A β A}
β is-prop (is-left-quasigroup _β_)
{A = A} =
is-left-quasigroup-is-prop 1 (is-left-quasigroupβop-is-right-quasigroup) (hlevel 1) Equivβis-hlevel
instance
H-Level-is-left-quasigroup: β {_β_ : A β A β A} {n}
β H-Level (is-left-quasigroup _β_) (suc n)
= prop-instance is-left-quasigroup-is-prop H-Level-is-left-quasigroup
β-equivlβis-left-quasigroup: β {_β_ : A β A β A}
β is-set A
β (β x β is-equiv (x β_))
β is-left-quasigroup _β_
=
β-equivlβis-left-quasigroup A-set eqv
is-right-quasigroupβop-is-left-quasigroup $
β-equivrβis-right-quasigroup A-set eqv
is-left-quasigroupββ-equivl: β {_β_ : A β A β A}
β is-left-quasigroup _β_ β (is-set A Γ (β x β is-equiv (x β_)))
=
is-left-quasigroupββ-equivl
is-left-quasigroupβop-is-right-quasigroup βe is-right-quasigroupββ-equivr
Left quasigroup homomorphismsπ
We can continue dualizing to define a notion of homomorphism for left quasigroups, though we shall be much more terse in our development. Following the pattern of right quasigroups, we begin by defining left quasigroup structures.
A left quasigroup structure on a type is a binary operation such that is a left quasigroup.
record Left-quasigroup-on {β} (A : Type β) : Type β where
no-eta-equality
field
_β_ : A β A β A
: is-left-quasigroup _β_
has-is-left-quasigroup
open is-left-quasigroup has-is-left-quasigroup public
A left quasigroup homomorphism between two left quasigroups is a function such that
record is-left-quasigroup-hom
{β β'} {A : Type β} {B : Type β'}
(f : A β B)
(S : Left-quasigroup-on A) (T : Left-quasigroup-on B)
: Type (β β β')
where
no-eta-equality
private
module A = Left-quasigroup-on S
module B = Left-quasigroup-on T
field
: β x y β f (x A.β y) β‘ f x B.β f y pres-β
Dual to right quasigroups, preservation of left division follows from left-cancellativity.
: β x y β f (x A.⧡ y) β‘ f x B.⧡ f y
pres-⧡ =
pres-⧡ x y .β-cancell (f x) $
B.β f (x A.⧡ y) β‘Λβ¨ pres-β x (x A.⧡ y) β©
f x B(x A.β (x A.⧡ y)) β‘β¨ ap f (A.⧡-invr x y) β©
f .⧡-invr (f x) (f y) β©
f y β‘Λβ¨ B.β (f x B.⧡ f y) β f x B
is-left-quasigroup-hom-is-prop: {f : A β B}
β {S : Left-quasigroup-on A} {T : Left-quasigroup-on B}
β is-prop (is-left-quasigroup-hom f S T)
{T = T} =
is-left-quasigroup-hom-is-prop 1 eqv
Isoβis-hlevel! where
open Left-quasigroup-on T
private unquoteDecl eqv = declare-record-iso eqv (quote is-left-quasigroup-hom)
instance
H-Level-is-left-quasigroup-hom: {f : A β B}
β {S : Left-quasigroup-on A} {T : Left-quasigroup-on B}
β β {n} β H-Level (is-left-quasigroup-hom f S T) (suc n)
= prop-instance is-left-quasigroup-hom-is-prop H-Level-is-left-quasigroup-hom
Left quasigroups are algebraic structures, so they form a thin structure over
module _ where
open is-left-quasigroup-hom
open Left-quasigroup-on
Left-quasigroup-on-pathp: β {A : I β Type β}
β {S : Left-quasigroup-on (A i0)} {T : Left-quasigroup-on (A i1)}
β PathP (Ξ» i β β (x y : A i) β A i) (S ._β_) (T ._β_)
β PathP (Ξ» i β Left-quasigroup-on (A i)) S T
{S = S} {T = T} p i ._β_ x y =
Left-quasigroup-on-pathp
p i x y{S = S} {T = T} p i .has-is-left-quasigroup =
Left-quasigroup-on-pathp
is-propβpathp(Ξ» i β is-left-quasigroup-is-prop {_β_ = p i})
(S .has-is-left-quasigroup)
(T .has-is-left-quasigroup)
i
: β β β Thin-structure β Left-quasigroup-on
Left-quasigroup-structure .is-hom f A B .β£_β£ =
Left-quasigroup-structure β
is-left-quasigroup-hom f A B.is-hom f A B .is-tr =
Left-quasigroup-structure β
is-left-quasigroup-hom-is-prop.id-is-hom .pres-β x y =
Left-quasigroup-structure β
refl.β-is-hom f g f-hom g-hom .pres-β x y =
Left-quasigroup-structure β (g-hom .pres-β x y) β f-hom .pres-β (g x) (g y)
ap f .id-hom-unique p q =
Left-quasigroup-structure β (ext (p .pres-β)) Left-quasigroup-on-pathp
This observation lets us assemble left quasigroups into a category.
: β β β Precategory (lsuc β) β
Left-quasigroups = Structured-objects (Left-quasigroup-structure β)
Left-quasigroups β
module Left-quasigroups {β} = Cat.Reasoning (Left-quasigroups β)
: (β : Level) β Type (lsuc β)
Left-quasigroup = Left-quasigroups.Ob {β} Left-quasigroup β
module Left-quasigroup {β} (A : Left-quasigroup β) where
open Left-quasigroup-on (A .snd) public
Quasigroupsπ
A type equipped with a binary operation is a quasigroup if it is a left quasigroup and a right quasigroup.
record is-quasigroup {β} {A : Type β} (_β_ : A β A β A) : Type β where
no-eta-equality
field
: is-left-quasigroup _β_
has-is-left-quasigroup : is-right-quasigroup _β_ has-is-right-quasigroup
open is-left-quasigroup has-is-left-quasigroup public
open is-right-quasigroup has-is-right-quasigroup
hiding (has-is-magma; underlying-set; magma-hlevel; has-is-set) public
Quasigroups obey the latin square property: for every there exists a unique pair such that and
: β x y β is-contr (Ξ£[ l β A ] Ξ£[ r β A ] (l β x β‘ y Γ x β r β‘ y)) β-latin
The proof is an exercise in equivalence yoga: by definition, a quasigroup is both a left and right quasigroup, so multiplication on the left and right is an equivalence, so the types and are both contractible. Moreover, is equivalent to so the latter must also be contractible.
=
β-latin x y 0
Equivβis-hlevel
latin-eqv(Γ-is-hlevel 0 (β-equivr x .is-eqv y) (β-equivl x .is-eqv y))
where
latin-eqv: (Ξ£[ l β A ] Ξ£[ r β A ] (l β x β‘ y Γ x β r β‘ y))
((Ξ£[ l β A ] l β x β‘ y) Γ (Ξ£[ r β A ] x β r β‘ y))
β =
latin-eqv (l β x β‘ y Γ x β r β‘ y) ββ¨ Ξ£-ap idβ (Ξ» l β Ξ£-swapβ) β©
Ξ£[ l β A ] Ξ£[ r β A ] (l β x β‘ y Γ (Ξ£[ r β A ] (x β r) β‘ y)) ββ¨ Ξ£-assoc β©
Ξ£[ l β A ] (Ξ£[ l β A ] l β x β‘ y) Γ (Ξ£[ r β A ] x β r β‘ y) ββ
We also have the following pair of useful identities that relate left and right division.
: β x y β x / (y ⧡ x) β‘ y
/-⧡-cancelr =
/-⧡-cancelr x y (y ⧡ x) β‘Λβ¨ ap (_/ (y ⧡ x)) (⧡-invr y x) β©
x / (y β (y ⧡ x)) / (y ⧡ x) β‘β¨ /-invr y (y ⧡ x) β©
y β
: β x y β (x / y) ⧡ x β‘ y
/-⧡-cancell =
/-⧡-cancell x y (x / y) ⧡ x β‘Λβ¨ ap ((x / y) ⧡_) (/-invl x y) β©
(x / y) ⧡ ((x / y) β y) β‘⨠⧡-invl (x / y) y β©
y β
unquoteDecl H-Level-is-quasigroup = declare-record-hlevel 1 H-Level-is-quasigroup (quote is-quasigroup)
is-quasigroup-is-prop: β {_β_ : A β A β A}
β is-prop (is-quasigroup _β_)
= hlevel 1 is-quasigroup-is-prop
Quasigroup homomorphismsπ
A quasigroup structure on a type is a binary operation such that is a quasigroup.
record Quasigroup-on {β} (A : Type β) : Type β where
no-eta-equality
field
_β_ : A β A β A
: is-quasigroup _β_
has-is-quasigroup
open is-quasigroup has-is-quasigroup public
: Left-quasigroup-on A
left-quasigroup-on .Left-quasigroup-on._β_ = _β_
left-quasigroup-on .Left-quasigroup-on.has-is-left-quasigroup =
left-quasigroup-on
has-is-left-quasigroup
: Right-quasigroup-on A
right-quasigroup-on .Right-quasigroup-on._β_ = _β_
right-quasigroup-on .Right-quasigroup-on.has-is-right-quasigroup =
right-quasigroup-on has-is-right-quasigroup
A quasigroup homomorphism between two quasigroups is a function such that
record is-quasigroup-hom
{β β'} {A : Type β} {B : Type β'}
(f : A β B)
(S : Quasigroup-on A) (T : Quasigroup-on B)
: Type (β β β')
where
no-eta-equality
private
module A = Quasigroup-on S
module B = Quasigroup-on T
field
: β x y β f (x A.β y) β‘ f x B.β f y pres-β
has-is-left-quasigroup-hom: is-left-quasigroup-hom f A.left-quasigroup-on B.left-quasigroup-on
.is-left-quasigroup-hom.pres-β = pres-β
has-is-left-quasigroup-hom
has-is-right-quasigroup-hom: is-right-quasigroup-hom f A.right-quasigroup-on B.right-quasigroup-on
.is-right-quasigroup-hom.pres-β = pres-β
has-is-right-quasigroup-hom
open is-left-quasigroup-hom has-is-left-quasigroup-hom
hiding (pres-β)
public
open is-right-quasigroup-hom has-is-right-quasigroup-hom
hiding (pres-β)
public
is-quasigroup-hom-is-prop: {f : A β B}
β {S : Quasigroup-on A} {T : Quasigroup-on B}
β is-prop (is-quasigroup-hom f S T)
{T = T} =
is-quasigroup-hom-is-prop 1 eqv
Isoβis-hlevel! where
open Quasigroup-on T
private unquoteDecl eqv = declare-record-iso eqv (quote is-quasigroup-hom)
instance
H-Level-is-quasigroup-hom: {f : A β B}
β {S : Quasigroup-on A} {T : Quasigroup-on B}
β β {n} β H-Level (is-quasigroup-hom f S T) (suc n)
= prop-instance is-quasigroup-hom-is-prop H-Level-is-quasigroup-hom
module _ where
open is-quasigroup-hom
open Quasigroup-on
Quasigroup-on-pathp: β {A : I β Type β}
β {S : Quasigroup-on (A i0)} {T : Quasigroup-on (A i1)}
β PathP (Ξ» i β β (x y : A i) β A i) (S ._β_) (T ._β_)
β PathP (Ξ» i β Quasigroup-on (A i)) S T
{S = S} {T = T} p i ._β_ x y =
Quasigroup-on-pathp
p i x y{S = S} {T = T} p i .has-is-quasigroup =
Quasigroup-on-pathp
is-propβpathp(Ξ» i β is-quasigroup-is-prop {_β_ = p i})
(S .has-is-quasigroup)
(T .has-is-quasigroup)
i
Quasigroups and quasigroup homomorphisms form a thin structure.
: β β β Thin-structure β Quasigroup-on Quasigroup-structure
.is-hom f A B .β£_β£ =
Quasigroup-structure β
is-quasigroup-hom f A B.is-hom f A B .is-tr =
Quasigroup-structure β
is-quasigroup-hom-is-prop.id-is-hom .pres-β x y =
Quasigroup-structure β
refl.β-is-hom f g f-hom g-hom .pres-β x y =
Quasigroup-structure β (g-hom .pres-β x y) β f-hom .pres-β (g x) (g y)
ap f .id-hom-unique p q =
Quasigroup-structure β (ext (p .pres-β)) Quasigroup-on-pathp
This gives us a tidy way to construct the category of quasigroups.
: β β β Precategory (lsuc β) β
Quasigroups = Structured-objects (Quasigroup-structure β)
Quasigroups β
module Quasigroups {β} = Cat.Reasoning (Quasigroups β)
: (β : Level) β Type (lsuc β)
Quasigroup = Quasigroups.Ob {β} Quasigroup β
module Quasigroup {β} (A : Quasigroup β) where
open Quasigroup-on (A .snd) public
: Left-quasigroup β
left-quasigroup = A .fst , left-quasigroup-on
left-quasigroup
: Right-quasigroup β
right-quasigroup = A .fst , right-quasigroup-on right-quasigroup
Constructing quasigroupsπ
The interfaces for is-quasigroup
and Quasigroup-on
are deeply nested
and contain some duplication, so we introduce some helper functions for
constructing them.
record make-quasigroup {β} (A : Type β) : Type β where
field
: is-set A
quasigroup-is-set _β_ : A β A β A
_⧡_ : A β A β A
_/_ : A β A β A
: β x y β (x / y) β y β‘ x
/-invl : β x y β (x β y) / y β‘ x
/-invr : β x y β x β (x ⧡ y) β‘ y
⧡-invr : β x y β x ⧡ (x β y) β‘ y ⧡-invl
: is-quasigroup _β_
to-is-quasigroup .is-quasigroup.has-is-left-quasigroup .is-left-quasigroup._⧡_ = _⧡_
to-is-quasigroup .is-quasigroup.has-is-left-quasigroup .is-left-quasigroup.⧡-invr = ⧡-invr
to-is-quasigroup .is-quasigroup.has-is-left-quasigroup .is-left-quasigroup.⧡-invl = ⧡-invl
to-is-quasigroup .is-quasigroup.has-is-left-quasigroup .is-left-quasigroup.has-is-magma .is-magma.has-is-set = quasigroup-is-set
to-is-quasigroup .is-quasigroup.has-is-right-quasigroup .is-right-quasigroup._/_ = _/_
to-is-quasigroup .is-quasigroup.has-is-right-quasigroup .is-right-quasigroup./-invl = /-invl
to-is-quasigroup .is-quasigroup.has-is-right-quasigroup .is-right-quasigroup./-invr = /-invr
to-is-quasigroup .is-quasigroup.has-is-right-quasigroup .is-right-quasigroup.has-is-magma .is-magma.has-is-set = quasigroup-is-set
to-is-quasigroup
: Quasigroup-on A
to-quasigroup-on .Quasigroup-on._β_ = _β_
to-quasigroup-on .Quasigroup-on.has-is-quasigroup = to-is-quasigroup
to-quasigroup-on
: Quasigroup β
to-quasigroup .fst .β£_β£ = A
to-quasigroup .fst .is-tr = quasigroup-is-set
to-quasigroup .snd = to-quasigroup-on
to-quasigroup
open make-quasigroup using (to-is-quasigroup; to-quasigroup-on; to-quasigroup) public