{-# OPTIONS --cubical --no-import-sorts --safe #-}
module Cubical.Structures.Semigroup where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.Equiv.HalfAdjoint
open import Cubical.Foundations.HLevels
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Univalence
open import Cubical.Foundations.Transport
open import Cubical.Foundations.SIP

open import Cubical.Data.Sigma

open import Cubical.Structures.Axioms
open import Cubical.Structures.Auto

open Iso

private
  variable
     : Level

-- Semigroups as a record, inspired by the Agda standard library:
--
--  https://github.com/agda/agda-stdlib/blob/master/src/Algebra/Bundles.agda#L48
--  https://github.com/agda/agda-stdlib/blob/master/src/Algebra/Structures.agda#L50
--
-- Note that as we are using Path for all equations the IsMagma record
-- would only contain isSet A if we had it.
record IsSemigroup {A : Type } (_·_ : A  A  A) : Type  where

  -- TODO: add no-eta-equality for efficiency? This breaks some proofs later

  constructor issemigroup
  field
    is-set : isSet A
    assoc  : (x y z : A)  x · (y · z)  (x · y) · z

record Semigroup : Type (ℓ-suc ) where

  constructor semigroup

  field
    Carrier     : Type 
    _·_         : Carrier  Carrier  Carrier
    isSemigroup : IsSemigroup _·_

  infixl 7 _·_

  open IsSemigroup isSemigroup public

-- Extractor for the carrier type
⟨_⟩ : Semigroup  Type 
⟨_⟩ = Semigroup.Carrier


record SemigroupEquiv (M N : Semigroup {}) : Type  where

  constructor semigroupequiv

  -- Shorter qualified names
  private
    module M = Semigroup M
    module N = Semigroup N

  field
    e     :  M    N 
    isHom : (x y :  M )  equivFun e (x M.· y)  equivFun e x N.· equivFun e y


-- Develop some theory about Semigroups using various general results
-- that are stated using Σ-types. For this we define Semigroup as a
-- nested Σ-type, prove that it's equivalent to the above record
-- definition and then transport results along this equivalence.
module SemigroupΣTheory {} where

  RawSemigroupStructure : Type   Type 
  RawSemigroupStructure X = X  X  X

  RawSemigroupEquivStr = AutoEquivStr RawSemigroupStructure

  rawSemigroupUnivalentStr : UnivalentStr _ RawSemigroupEquivStr
  rawSemigroupUnivalentStr = autoUnivalentStr RawSemigroupStructure

  SemigroupAxioms : (A : Type )  RawSemigroupStructure A  Type 
  SemigroupAxioms A _·_ = isSet A
                        × ((x y z : A)  x · (y · z)  (x · y) · z)

  SemigroupStructure : Type   Type 
  SemigroupStructure = AxiomsStructure RawSemigroupStructure SemigroupAxioms

  SemigroupΣ : Type (ℓ-suc )
  SemigroupΣ = TypeWithStr  SemigroupStructure

  isPropSemigroupAxioms : (A : Type ) (_·_ : RawSemigroupStructure A)
                         isProp (SemigroupAxioms A _·_)
  isPropSemigroupAxioms _ _ = isPropΣ isPropIsSet λ isSetA  isPropΠ3 λ _ _ _  isSetA _ _

  SemigroupEquivStr : StrEquiv SemigroupStructure 
  SemigroupEquivStr = AxiomsEquivStr RawSemigroupEquivStr SemigroupAxioms

  SemigroupAxiomsIsoIsSemigroup : {A : Type } (_·_ : RawSemigroupStructure A)
                                 Iso (SemigroupAxioms A _·_) (IsSemigroup _·_)
  fun (SemigroupAxiomsIsoIsSemigroup s) (x , y)           = issemigroup x y
  inv (SemigroupAxiomsIsoIsSemigroup s) (issemigroup x y) = (x , y)
  rightInv (SemigroupAxiomsIsoIsSemigroup s) _            = refl
  leftInv (SemigroupAxiomsIsoIsSemigroup s) _             = refl

  SemigroupAxioms≡IsSemigroup : {A : Type } (_·_ : RawSemigroupStructure A)
                               SemigroupAxioms _ _·_  IsSemigroup _·_
  SemigroupAxioms≡IsSemigroup s = isoToPath (SemigroupAxiomsIsoIsSemigroup s)

  Semigroup→SemigroupΣ : Semigroup  SemigroupΣ
  Semigroup→SemigroupΣ (semigroup A _·_ isSemigroup) =
    A , _·_ , SemigroupAxiomsIsoIsSemigroup _ .inv isSemigroup

  SemigroupΣ→Semigroup : SemigroupΣ  Semigroup
  SemigroupΣ→Semigroup (A , _·_ , isSemigroupΣ) =
    semigroup A _·_ (SemigroupAxiomsIsoIsSemigroup _ .fun isSemigroupΣ)

  SemigroupIsoSemigroupΣ : Iso Semigroup SemigroupΣ
  SemigroupIsoSemigroupΣ =
    iso Semigroup→SemigroupΣ SemigroupΣ→Semigroup  _  refl)  _  refl)

  semigroupUnivalentStr : UnivalentStr SemigroupStructure SemigroupEquivStr
  semigroupUnivalentStr = axiomsUnivalentStr _ isPropSemigroupAxioms rawSemigroupUnivalentStr

  SemigroupΣPath : (M N : SemigroupΣ)  (M ≃[ SemigroupEquivStr ] N)  (M  N)
  SemigroupΣPath = SIP semigroupUnivalentStr

  SemigroupEquivΣ : (M N : Semigroup)  Type 
  SemigroupEquivΣ M N = Semigroup→SemigroupΣ M ≃[ SemigroupEquivStr ] Semigroup→SemigroupΣ N

  SemigroupIsoΣPath : {M N : Semigroup}  Iso (SemigroupEquiv M N) (SemigroupEquivΣ M N)
  fun SemigroupIsoΣPath (semigroupequiv e h) = (e , h)
  inv SemigroupIsoΣPath (e , h)            = semigroupequiv e h
  rightInv SemigroupIsoΣPath _             = refl
  leftInv SemigroupIsoΣPath _              = refl

  SemigroupPath : (M N : Semigroup)  (SemigroupEquiv M N)  (M  N)
  SemigroupPath M N =
    SemigroupEquiv M N                                ≃⟨ isoToEquiv SemigroupIsoΣPath 
    SemigroupEquivΣ M N                               ≃⟨ SemigroupΣPath _ _ 
    Semigroup→SemigroupΣ M  Semigroup→SemigroupΣ N ≃⟨ isoToEquiv (invIso (congIso SemigroupIsoSemigroupΣ)) 
    M  N 

-- We now extract the important results from the above module

isPropIsSemigroup : {A : Type } (_·_ : A  A  A)  isProp (IsSemigroup _·_)
isPropIsSemigroup _·_ =
  subst isProp (SemigroupΣTheory.SemigroupAxioms≡IsSemigroup _·_)
        (SemigroupΣTheory.isPropSemigroupAxioms _ _·_)

SemigroupPath : (M N : Semigroup {})  (SemigroupEquiv M N)  (M  N)
SemigroupPath = SemigroupΣTheory.SemigroupPath


-- To rename the fields when using a Semigroup use for example the following:
--
-- open Semigroup M renaming ( Carrier to M ; _·_ to _·M_ )