open import 1Lab.Type.Sigma

open import Algebra.Magma

open import Cat.Displayed.Univalence.Thin
open import Cat.Displayed.Total
open import Cat.Prelude hiding (_/_)

import Cat.Reasoning
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
  A B : Type β„“

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
    /-invl : βˆ€ x y β†’ (x / y) ⋆ y ≑ x
    /-invr : βˆ€ x y β†’ (x ⋆ y) / y ≑ x
    has-is-magma : 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.

  ⋆-equivr : βˆ€ x β†’ is-equiv (_⋆ x)
  ⋆-equivr x .is-eqv y .centre = y / x , /-invl y x
  ⋆-equivr x .is-eqv y .paths (l , lx=y) =
    Ξ£-prop-path (Ξ» l β†’ has-is-set (l ⋆ x) y) $
      y / x       β‰‘Λ˜βŸ¨ ap (_/ x) lx=y ⟩
      (l ⋆ x) / x β‰‘βŸ¨ /-invr l x ⟩
      l           ∎

This in turn implies that is an equivalence.

  /-equivr : βˆ€ x β†’ is-equiv (_/ x)
  /-equivr x = inverse-is-equiv (⋆-equivr x)

As easy corollaries, we get that multiplication and division in are right-cancellative.

  opaque
    ⋆-cancelr : βˆ€ {x y} (z : A) β†’ x ⋆ z ≑ y ⋆ z β†’ x ≑ y
    ⋆-cancelr z = Equiv.injective (_⋆ z , ⋆-equivr z)

    /-cancelr : βˆ€ {x y} (z : A) β†’ x / z ≑ y / z β†’ x ≑ y
    /-cancelr z = Equiv.injective (_/ z , /-equivr 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.

⋆-equivrβ†’is-right-quasigroup {_⋆_ = _⋆_} A-set ⋆-eqv =
  ⋆-right-quasi
  where
    open is-right-quasigroup

    module ⋆-eqv x = Equiv (_⋆ x , ⋆-eqv x)

    ⋆-right-quasi : 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

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 ⟩ ,
    iso (uncurry ⋆-equivrβ†’is-right-quasigroup)
      (Ξ» _ β†’ prop!)
     (Ξ» ⋆-right-quasi β†’
      let open is-right-quasigroup ⋆-right-quasi in
      Iso.injective eqv (refl ,β‚š prop!))
  where
    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 _⋆_)
is-right-quasigroup-is-prop {A = A} =
  Equivβ†’is-hlevel 1 is-right-quasigroup≃⋆-equivr (hlevel 1)
instance
  H-Level-is-right-quasigroup
    : βˆ€ {_⋆_ : A β†’ A β†’ A} {n}
    β†’ H-Level (is-right-quasigroup _⋆_) (suc n)
  H-Level-is-right-quasigroup = prop-instance is-right-quasigroup-is-prop

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
    has-is-right-quasigroup : 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
    pres-⋆ : βˆ€ x y β†’ f (x A.⋆ y) ≑ f x B.⋆ f y

Preservation of right division follows almost immediately from right-cancellativity.

  pres-/ : βˆ€ x y β†’ f (x A./ y) ≑ f x B./ f y
  pres-/ x y =
    B.⋆-cancelr (f y) $
      f (x A./ y) B.⋆ f y   β‰‘Λ˜βŸ¨ pres-⋆ (x A./ y) y ⟩
      f ((x A./ y) A.⋆ y)   β‰‘βŸ¨ ap f (A./-invl x y) ⟩
      f x                   β‰‘Λ˜βŸ¨ B./-invl (f x) (f y) ⟩
      (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)
is-right-quasigroup-hom-is-prop {T = T} =
  Iso→is-hlevel! 1 eqv
  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)
  H-Level-is-right-quasigroup-hom = prop-instance is-right-quasigroup-hom-is-prop
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
  Right-quasigroup-on-pathp {S = S} {T = T} p i ._⋆_ x y =
    p i x y
  Right-quasigroup-on-pathp {S = S} {T = T} p i .has-is-right-quasigroup =
    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

  Right-quasigroup-structure : βˆ€ β„“ β†’ Thin-structure β„“ Right-quasigroup-on
  Right-quasigroup-structure β„“ .is-hom f A B .∣_∣ =
    is-right-quasigroup-hom f A B
  Right-quasigroup-structure β„“ .is-hom f A B .is-tr =
    is-right-quasigroup-hom-is-prop
  Right-quasigroup-structure β„“ .id-is-hom .pres-⋆ x y =
    refl
  Right-quasigroup-structure β„“ .∘-is-hom f g f-hom g-hom .pres-⋆ x y =
    ap f (g-hom .pres-⋆ x y) βˆ™ f-hom .pres-⋆ (g x) (g y)
  Right-quasigroup-structure β„“ .id-hom-unique {A} {S} {T} p q =
    Right-quasigroup-on-pathp (ext (p .pres-⋆))

This observation lets us neatly organize right quasigroups into a category.

Right-quasigroups : βˆ€ β„“ β†’ Precategory (lsuc β„“) β„“
Right-quasigroups β„“ = Structured-objects (Right-quasigroup-structure β„“)

module Right-quasigroups {β„“} = Cat.Reasoning (Right-quasigroups β„“)

Right-quasigroup : (β„“ : Level) β†’ Type (lsuc β„“)
Right-quasigroup β„“ = Right-quasigroups.Ob {β„“}
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
    ⧡-invr : βˆ€ x y β†’ x ⋆ (x ⧡ y) ≑ y
    ⧡-invl : βˆ€ x y β†’ x ⧡ (x ⋆ y) ≑ y
    has-is-magma : 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.

  ⋆-equivl : βˆ€ x β†’ is-equiv (x ⋆_)
  ⋆-equivl x .is-eqv y .centre = x ⧡ y , ⧡-invr x y
  ⋆-equivl x .is-eqv y .paths (r , xr=y) =
    Ξ£-prop-path (Ξ» r β†’ has-is-set (x ⋆ r) y) $
      x ⧡ y       β‰‘Λ˜βŸ¨ ap (x ⧡_) xr=y ⟩
      x ⧡ (x ⋆ r) β‰‘βŸ¨ ⧡-invl x r ⟩
      r ∎

  ⧡-equivl : βˆ€ x β†’ is-equiv (x ⧡_)
  ⧡-equivl x = inverse-is-equiv (⋆-equivl x)

This implies that that multiplication and division in is left-cancellative.

  opaque
    ⋆-cancell : βˆ€ (x : A) {y z} β†’ x ⋆ y ≑ x ⋆ z β†’ y ≑ z
    ⋆-cancell x = Equiv.injective (x ⋆_ , ⋆-equivl x)

    ⧡-cancell : βˆ€ (x : A) {y z} β†’ x ⧡ y ≑ x ⧡ z β†’ y ≑ z
    ⧡-cancell x = Equiv.injective (x ⧡_ , ⧡-equivl 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)

is-left-quasigroupβ†’op-is-right-quasigroup {_⋆_ = _⋆_} A-quasi = A-op-quasi
  where
    open is-left-quasigroup A-quasi
    open is-right-quasigroup

    A-op-quasi : 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

is-right-quasigroupβ†’op-is-left-quasigroup {_⋆_ = _⋆_} A-quasi = A-op-quasi
  where
    open is-right-quasigroup A-quasi
    open is-left-quasigroup

    A-op-quasi : 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

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
        Iso.injective eqv (refl ,β‚š prop!))
    where
      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 _⋆_)
is-left-quasigroup-is-prop {A = A} =
  Equivβ†’is-hlevel 1 (is-left-quasigroup≃op-is-right-quasigroup) (hlevel 1)
instance
  H-Level-is-left-quasigroup
    : βˆ€ {_⋆_ : A β†’ A β†’ A} {n}
    β†’ H-Level (is-left-quasigroup _⋆_) (suc n)
  H-Level-is-left-quasigroup = prop-instance is-left-quasigroup-is-prop
⋆-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
    has-is-left-quasigroup : 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
    pres-⋆ : βˆ€ x y β†’ f (x A.⋆ y) ≑ f x B.⋆ f y

Dual to right quasigroups, preservation of left division follows from left-cancellativity.

  pres-⧡ : βˆ€ x y β†’ f (x A.⧡ y) ≑ f x B.⧡ f y
  pres-⧡ x y =
    B.⋆-cancell (f x) $
      f x B.⋆ f (x A.⧡ y)     β‰‘Λ˜βŸ¨ pres-⋆ x (x A.⧡ y) ⟩
      f (x A.⋆ (x A.⧡ y))     β‰‘βŸ¨ ap f (A.⧡-invr x y) ⟩
      f y                     β‰‘Λ˜βŸ¨ B.⧡-invr (f x) (f y) ⟩
      f x B.⋆ (f x B.⧡ f y) ∎
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)
is-left-quasigroup-hom-is-prop {T = T} =
  Iso→is-hlevel! 1 eqv
  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)
  H-Level-is-left-quasigroup-hom = prop-instance is-left-quasigroup-hom-is-prop

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
  Left-quasigroup-on-pathp {S = S} {T = T} p i ._⋆_ x y =
    p i x y
  Left-quasigroup-on-pathp {S = S} {T = T} p i .has-is-left-quasigroup =
    is-prop→pathp
      (Ξ» i β†’ is-left-quasigroup-is-prop {_⋆_ = p i})
      (S .has-is-left-quasigroup)
      (T .has-is-left-quasigroup)
      i
  Left-quasigroup-structure : βˆ€ β„“ β†’ Thin-structure β„“ Left-quasigroup-on
  Left-quasigroup-structure β„“ .is-hom f A B .∣_∣ =
    is-left-quasigroup-hom f A B
  Left-quasigroup-structure β„“ .is-hom f A B .is-tr =
    is-left-quasigroup-hom-is-prop
  Left-quasigroup-structure β„“ .id-is-hom .pres-⋆ x y =
    refl
  Left-quasigroup-structure β„“ .∘-is-hom f g f-hom g-hom .pres-⋆ x y =
    ap f (g-hom .pres-⋆ x y) βˆ™ f-hom .pres-⋆ (g x) (g y)
  Left-quasigroup-structure β„“ .id-hom-unique p q =
    Left-quasigroup-on-pathp (ext (p .pres-⋆))

This observation lets us assemble left quasigroups into a category.

Left-quasigroups : βˆ€ β„“ β†’ Precategory (lsuc β„“) β„“
Left-quasigroups β„“ = Structured-objects (Left-quasigroup-structure β„“)

module Left-quasigroups {β„“} = Cat.Reasoning (Left-quasigroups β„“)

Left-quasigroup : (β„“ : Level) β†’ Type (lsuc β„“)
Left-quasigroup β„“ = Left-quasigroups.Ob {β„“}
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
    has-is-left-quasigroup : is-left-quasigroup _⋆_
    has-is-right-quasigroup : 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

  ⋆-latin : βˆ€ x y β†’ is-contr (Ξ£[ l ∈ A ] Ξ£[ r ∈ A ] (l ⋆ x ≑ y Γ— x ⋆ r ≑ y))

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 =
    Equiv→is-hlevel 0
      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 ∈ A ] Ξ£[ r ∈ A ] (l ⋆ x ≑ y Γ— x ⋆ r ≑ y)     β‰ƒβŸ¨ Ξ£-ap id≃ (Ξ» l β†’ Ξ£-swapβ‚‚) ⟩
        Ξ£[ l ∈ A ] (l ⋆ x ≑ y Γ— (Ξ£[ r ∈ A ] (x ⋆ r) ≑ y)) β‰ƒβŸ¨ Ξ£-assoc ⟩
        (Ξ£[ 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.

  /-⧡-cancelr : βˆ€ x y β†’ x / (y ⧡ x) ≑ y
  /-⧡-cancelr x y =
    x / (y ⧡ x)             β‰‘Λ˜βŸ¨ ap (_/ (y ⧡ x)) (⧡-invr y x) ⟩
    (y ⋆ (y ⧡ x)) / (y ⧡ x) β‰‘βŸ¨ /-invr y (y ⧡ x) ⟩
    y                       ∎

  /-⧡-cancell : βˆ€ x y β†’ (x / y) ⧡ x ≑ y
  /-⧡-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 _⋆_)
is-quasigroup-is-prop = hlevel 1

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
    has-is-quasigroup : is-quasigroup _⋆_

  open is-quasigroup has-is-quasigroup public
  left-quasigroup-on : Left-quasigroup-on A
  left-quasigroup-on .Left-quasigroup-on._⋆_ = _⋆_
  left-quasigroup-on .Left-quasigroup-on.has-is-left-quasigroup =
    has-is-left-quasigroup

  right-quasigroup-on : Right-quasigroup-on A
  right-quasigroup-on .Right-quasigroup-on._⋆_ = _⋆_
  right-quasigroup-on .Right-quasigroup-on.has-is-right-quasigroup =
    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
    pres-⋆ : βˆ€ x y β†’ f (x A.⋆ y) ≑ f x B.⋆ f y
  has-is-left-quasigroup-hom
    : is-left-quasigroup-hom f A.left-quasigroup-on B.left-quasigroup-on
  has-is-left-quasigroup-hom .is-left-quasigroup-hom.pres-⋆ = pres-⋆

  has-is-right-quasigroup-hom
    : is-right-quasigroup-hom f A.right-quasigroup-on B.right-quasigroup-on
  has-is-right-quasigroup-hom .is-right-quasigroup-hom.pres-⋆ = pres-⋆

  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)
is-quasigroup-hom-is-prop {T = T} =
  Iso→is-hlevel! 1 eqv
  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)
  H-Level-is-quasigroup-hom = prop-instance is-quasigroup-hom-is-prop
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
  Quasigroup-on-pathp {S = S} {T = T} p i ._⋆_ x y =
    p i x y
  Quasigroup-on-pathp {S = S} {T = T} p i .has-is-quasigroup =
    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.

  Quasigroup-structure : βˆ€ β„“ β†’ Thin-structure β„“ Quasigroup-on
  Quasigroup-structure β„“ .is-hom f A B .∣_∣ =
    is-quasigroup-hom f A B
  Quasigroup-structure β„“ .is-hom f A B .is-tr =
    is-quasigroup-hom-is-prop
  Quasigroup-structure β„“ .id-is-hom .pres-⋆ x y =
    refl
  Quasigroup-structure β„“ .∘-is-hom f g f-hom g-hom .pres-⋆ x y =
    ap f (g-hom .pres-⋆ x y) βˆ™ f-hom .pres-⋆ (g x) (g y)
  Quasigroup-structure β„“ .id-hom-unique p q =
    Quasigroup-on-pathp (ext (p .pres-⋆))

This gives us a tidy way to construct the category of quasigroups.

Quasigroups : βˆ€ β„“ β†’ Precategory (lsuc β„“) β„“
Quasigroups β„“ = Structured-objects (Quasigroup-structure β„“)

module Quasigroups {β„“} = Cat.Reasoning (Quasigroups β„“)

Quasigroup : (β„“ : Level) β†’ Type (lsuc β„“)
Quasigroup β„“ = Quasigroups.Ob {β„“}
module Quasigroup {β„“} (A : Quasigroup β„“) where
  open Quasigroup-on (A .snd) public

  left-quasigroup : Left-quasigroup β„“
  left-quasigroup = A .fst , left-quasigroup-on

  right-quasigroup : Right-quasigroup β„“
  right-quasigroup = A .fst , right-quasigroup-on

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
    quasigroup-is-set : is-set A
    _⋆_ : A β†’ A β†’ A
    _⧡_ : A β†’ A β†’ A
    _/_ : A β†’ A β†’ A

    /-invl : βˆ€ x y β†’ (x / y) ⋆ y ≑ x
    /-invr : βˆ€ x y β†’ (x ⋆ y) / y ≑ x
    ⧡-invr : βˆ€ x y β†’ x ⋆ (x ⧡ y) ≑ y
    ⧡-invl : βˆ€ x y β†’ x ⧡ (x ⋆ y) ≑ y
  to-is-quasigroup : 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-quasigroup-on : Quasigroup-on A
  to-quasigroup-on .Quasigroup-on._⋆_ = _⋆_
  to-quasigroup-on .Quasigroup-on.has-is-quasigroup = to-is-quasigroup

  to-quasigroup : Quasigroup β„“
  to-quasigroup .fst .∣_∣ = A
  to-quasigroup .fst .is-tr = quasigroup-is-set
  to-quasigroup .snd = to-quasigroup-on

open make-quasigroup using (to-is-quasigroup; to-quasigroup-on; to-quasigroup) public