Source code on Github
{-# OPTIONS --cubical-compatible #-}
module Class.CommutativeMonoid.Core where

open import Class.Prelude
open import Class.Semigroup
open import Class.Monoid

import Algebra as Alg

record CommutativeMonoid c  Carrier : Type (lsuc (c  )) where
  infix  4 _≈_
  field
    _≈_                 : Carrier  Carrier  Type 
     semigroup        : Semigroup Carrier
     monoid           : Monoid Carrier
    isCommutativeMonoid : Alg.IsCommutativeMonoid {c} _≈_ _◇_ ε

module Conversion {c } where
  toBundle :  {Carrier}  CommutativeMonoid c  Carrier  Alg.CommutativeMonoid c 
  toBundle c = record { CommutativeMonoid c }

  fromBundle : (m : Alg.CommutativeMonoid c )
              CommutativeMonoid c  (Alg.CommutativeMonoid.Carrier m)
  fromBundle c = record
    { Alg.CommutativeMonoid c using (_≈_; isCommutativeMonoid)
    ; semigroup = λ where ._◇_  Alg.CommutativeMonoid._∙_ c
    ; monoid    = λ where .ε    Alg.CommutativeMonoid.ε c }