Making Modules with Meta-Programmed Meta-Primitives

Liberating Package Formation from the Backend

Abstract

This article is about implementing a prototype supporting “the next 700 module systems” proposal as an editor extension. In particular, we show how intimately related presentations of a type can be derived automatically from a single generic declaration which we call a PackageFormer.

Think of a language that does not support currying and you need to have a function of 10 arguments that needs to support accepting any number of arguments less than 10, say for partial application. In such languages, one must utilise the builder design pattern, or quickly copy-paste the function 10 times, altering it slightly each time. In general, if such a function definition requires N lines and M forms of the function are needed, then nearly N × M lines of code are written manually.

Our prototype deals with this problem, among others, for functions on types —i.e., type constructors— and reduces this quadratic count to a linear count N + M: One declaration of N lines, then M lines, each being an instantiation of the desired form. These ideas are discussed in the pre-print A Language Feature to Unbundle Data at Will.

Design patterns for theories become library methods! An interesting side-effect of having meta-primitives for packages is that traditional patterns for theories —e.g., homomorphisms, syntax, interpretation functions— can now be codified as general re-usable methods.

The ideas are targeted to the dependently-typed language Agda. However, with little alterations they could easily be made to target other languages, such as Haskell and Coq. The ideas transcend the presentation language, Agda. Ideally, variational definitions would also be in the host language rather than in Lisp but that would require alteration to the host language itself, and Lisp with Emacs hooks is much easier to do; for now. This is an Emacs centric language extension.

The first section below quickly elaborates on our goal, after that is a ‘user manual’ then the remainder of the article serves as literate documentation of the prototype; as well as an opportunity for me to explore emojis. To follow along, it may be useful to look at an Elisp Cheat Sheet.

The user manual, among other things, shows how we can avoid the open ⋯ public ⋯ renaming ⋯ pattern which plagues Agda's standard library —and is much more pronounced in the RATH-Agda, which devotes p27-39 for simple setoid renaming and does much more elsewhere—, as well as showing how to avoid laborious, yet tedious, definitions of homomorphisms in Agda's library. Moreover, the motivating factor of this work is to avoid the pattern of defining a predicate IsX c₀ ⋯ cₙ on constituents cᵢ then packaging the constituents along with a proof of this predicate as a record X.

Tedium is for machines; interesting problems are for people.

Everything here works with Agda version 2.6.0.

Results of tests and examples will be in pink, like this.

1 Aim: Scrap the Repetition

We're going to write a code generator in Lisp that is going to interpret fictitious Agda code —henceforth referred to as “700 code”— into currently legitimate Agda code.

For example, something like the following, henceforth referred to as test:

PackageFormer M-Set : Setwhere
   Scalar  : Set
   Vector  : Set
   _·_     : Scalar Vector Vector
   𝟙       : Scalar
   _×_     : Scalar Scalar Scalar
   leftId  : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   assoc   : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)

Semantics     = M-Set record
Syntax        = M-Set data with tags (Scalar; Vector)
UntypedSyntax = M-Set data with identified carrier (Scalar; Vector)
ScalarSyntax  = M-Set data with identified carrier (Scalar)
Stream        = M-Set parameterised codata with identified carrier (Vector) renaming (_·_₁ to head; _·_₂ to tail)
VectorSyntax  = M-Set data with identified carrier (Vector) and variables (embed) from (Scalar)
NearMonoid    = M-Set record renaming (Scalar to Carrier; Vector to Carrier; _·_ to _⨾_; _×_ to _⨾_)

Will behave as if it were written:

record Semantics : Setwhere
  field
    Scalar  : Set
    Vector  : Set
    _·_     : Scalar Vector Vector
    𝟙       : Scalar
    _×_     : Scalar Scalar Scalar
    leftId  : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
    assoc   : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)

data SyntaxTag : Set where Scalar Vector : SyntaxTag
data Syntax : SyntaxTag Setwhere
    _·_     : Syntax Scalar Syntax Vector Syntax Vector
    𝟙       : Syntax Scalar
    _×_     : Syntax Scalar Syntax Scalar Syntax Scalar

data UntypedSyntax : Setwhere
    _·_     : UntypedSyntax UntypedSyntax UntypedSyntax
    𝟙       : UntypedSyntax
    _×_     : UntypedSyntax UntypedSyntax UntypedSyntax

data ScalarSyntax : Setwhere
    𝟙       : ScalarSyntax
    _×_     : ScalarSyntax ScalarSyntax ScalarSyntax

{- Nutshell: Keep items ending in “Stream Carrier”, then discard that ending,
         then form a subscripted version for each argument.
-}
record Stream (Carrier : Set) : Set₁ where
  coinductive
  field
    head : Carrier
    tail : Stream Carrier

data VectorSyntax (Scalar : Set) : Set₁ where
    embed   : Scalar → VectorSyntax
    _·_     : Scalar → VectorSyntax → VectorSyntax

record NearMonoid : Set₁ where
  field
    Carrier : Set
    _⨾_     : Carrier → Carrier → Carrier
    𝟙       : Carrier
    leftId  : {𝓋 : Carrier}  →  𝟙 ⨾ 𝓋  ≡  𝓋
    assoc   : {a b : Carrier} {𝓋 : Carrier} → (a ⨾ b) ⨾ 𝓋  ≡  a ⨾ (b ⨾ 𝓋)

This is a more than a 200% increase in size; that is, our fictitious code will save us a lot of repetition.

The above is ideal syntax: In Lisp fashion, we are not limiting our vision to what we can currently be done. The actual syntax that is currently supported by this program is surprisingly close to the above, with an occasional - or : inserted. Just look at the following source file and resulting generated code!

package-former.agda
{- This loads the PackageFormer metaprogram; press C-x C-e after the closing “)” below.                 -}
{- (progn (load-file "~/.emacs.d/agda-next-700-module-systems.el") (agda-next-700-module-systems-mode)) -}

module package-former where

open import package-former-generated
open import Level
open import Data.Bool
open import Data.List using (List; _∷_; []; foldr)
import Relation.Binary.PropositionalEquality as ≡; openusing (_≡_)

{- Let's ensure content of User Manual part I actually type checkes -}
{- Feel free to comment this line out. -}
import package-former-user-manual-i

{-
0. There are a number of common use-cases.
1. We can handle all of them & more, since we're extensible.
  - Mention the Lean & Coq, as well as the Agda, repeated fragments.
2. The resulting setup is pragmatic: It is unobtrusive in the
   traditional Agda coding style in that it happens in the background.
3. It fills a particular need; the desire to avoid repetitious code.
-}

{- lisp
(message-box "Hello")
(message-box "World")
-}

{-700
PackageFormer M-Set : Set₁ where
   Scalar  : Set
   Vector  : Set
   _·_     : Scalar → Vector → Vector
   𝟙       : Scalar
   _×_     : Scalar → Scalar → Scalar
   leftId  : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   assoc   : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
-}

{-700
PackageFormer MonoidP : Setwhere

    -- A few declarations
    Carrier : Set
    _⨾_     : Carrier Carrier Carrier
    Id      : Carrier
    assoc   : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)

    -- We have a setoid-like structure; with a default implementation
    _≈_   : Carrier CarrierSet
    _≈_   = _≡_
    ⨾-cong : ∀ {x y xy′} → xx′ →  yy′ → (xy) ≈ (x′ ⨾ y′)
    ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}

    -- For now only one item in a declaration;
    -- namely “Lid” & “Ridcannot be declared in one line.
    Lid : Carrier Carrier
    Lid x = Idx
    Rid : Carrier Carrier
    Rid x = x Id

    -- Agda permits pure, non-pattern-matching, equations betweenfieldsin a record.
    concat : List Carrier Carrier
    concat = foldr _⨾_ Id

    -- More declarations
    leftId  : ∀ {x : Carrier} → (Idx) ≈ x
    rightId : ∀ {x : Carrier} → Rid xx

    -- Since there are no more pure declarations, “fields”, subsequent equations
    -- may use pattern matching.

    Id² : (Id Id) ≈ Id
    Id² = rightId

    concatₚ : List Carrier Carrier
    concatₚ []       = Id
    concatₚ (xxs) = xconcatxs
-}

{-700
-- Variational with empty right hand side.
𝒱-identity =
-}

{-700
MonoidPⁱᵈ = MonoidP identity
-}

{-700
𝒱-no-op = "This is the do-nothing variational"
-}

{-700
-- No variational clauses needed!
MonoidP= MonoidP
-}

{-700
-- Identity of composition ⟴
MonoidPᶜ = MonoidP ⟴
-}

{-700
𝒱-test positional (keyword 3) another = "I have two mandatory arguments and one keyword argument"

Monoid-test = MonoidPtest "positional arg₁" "positional arg₂" :keyword 25
-}

{-   700
𝒱-whoops  = tester 1 2 :keyword 3
-}

{-700
𝒱-record= :kind record :alter-elements (λ es → (--map (map-qualifier (λ _ → "field") it) es))
-}

{-700
M-Set-Record = M-Set record₀
-}

{-lisp
(𝒱 record₁ (discard-equations nil)
 = "Reify a variational as an Agda “record”.
    Elements with equations are construed as
    derivatives of fields  ---the elements
    without any equations--- by default, unless
    DISCARD-EQUATIONS is provided with a non-nil value.
   "
  :kind record
  :alter-elements
    (λ es →
      (thread-last es
      ;; Keep or drop eqns depending ondiscard-equations”
      (--map
        (if discard-equations
            (map-equations (λ _ → nil) it)
            it))
      ;; Unless there's equations, mark elements as fields.
      (--map (map-qualifier
        (λ _ → (unless (element-equations it)
               "field")) it)))))
-}

{-700
Monoid-Record-derived = MonoidP record₁
-}

{-700
Monoid-Record-field = MonoidP record₁ :discard-equations t
-}

{-700
Monoid-Record-derived-again  = MonoidP record
Monoid-Record-derived-again2 = MonoidP record :and-names t
Monoid-Record-field-again    = MonoidP record :discard-equations t
Monoid-Record-no-equationals = MonoidP record :discard-equations t :and-names t
-}

{-700
𝒱-typeclass-attempt = record ⟴ :waist 2
-}

{-700
M-Set-TypeClass = M-Set typeclass-attempt
-}

{-700
𝒱-typeclass= record ⟴ :waist 2 :level dec
MonoidT= MonoidP typeclass₂
-}

{-700
MonoidT₃         = MonoidP record ⟴ :waist 3 :level dec
-- MonoidT₃-again   = MonoidP ⟴ record ⟴ unbundling 3
M-Set-Typeclass₂ = M-Set record ⟴ typeclass₂
-}

{-700
-- Ill-formed in Agda: A defintion is not a parameter!
MonoidP-Typeclass= MonoidP :waist 5
-}

{-700
MonoidT₅ = MonoidP ⟴ unbundling 5 ⟴ record
-}

{-700
-- Intentionally erroenous attempt.
𝒱-primed-attempt = :alter-elements (λ es → (--map (map-name (λ n → (rename-mixfix (λ np → (concat np "′")) n)) it) es))

-- M-Set′-attempt = M-Set recordprimed-attempt
-}

{-700
M-Set′-attempt-raw = M-Set primed-attempt
-}

{-700
𝒱-typeclass height (level 'dec) = record ⟴ :waist height :level level
M-Set2-Typeclass= M-Set typeclass 3 :level 'inc
M-Set0-Typeclass= M-Set typeclass 3
-}

{-700
MR𝕏    = M-Set record ⟴ map (λ e → (map-name (λ n → (rename-mixfix (λ x → (concat x "𝕏")) n)) e))
-}

{-700
MR𝕪    = M-Set-Record rename (λ n → (concat n "𝕪"))
MR-oh  = M-Set-Record rename (λ n → (pcase n ("Scalar" "S") ("𝟙" "ε") (else else)))
-}

{-700
-- MR₁₂   = M-Set-Record decorated "₁" ⟴ decorated "₂" :adjoin-retract nil
the-MR = M-Set-Record co-decorated "the-"
-- MR₃₄   = M-Set-Record subscripted₃ ⟴ subscripted₄ :adjoin-retract nil
MRₜₒ   = M-Set-Record renaming "Scalar to S; Vector to V; · to nice"
NearMonoid = M-Set-Record renaming "Scalar to Carrier; Vector to Carrier; · to ×"
-}

{-700
NearMonoid¹ = M-Set-Record single-sorted "Carrier"
-}

{-   700
ScalarTerm = M-Set data "Scalar"
-}

{-700
M-Set-Sorts = M-Set recordsorts
-}

{-700
MonoidSignature = M-Set-Record generated (λ e → (and (s-contains? "Scalar" (element-type e)) (not (s-contains? "Vector" (element-type e)))))
-}

{-700
MonSig = M-Set-Record signature
-}

{-700
𝒱-empty-module = :kind module :level none :waist 999
Neato = M-Set empty-module
-}

{- A module where the elements are all parameters -}
open Neato using ()

{-700
M-Set-R = M-Set record
M-Set-R₁ = M-Set-R ⟴ open (λ x → (concat x "₁"))
-}

{-700
M-Set-R-SV = M-Set-R opening "Scalar to S; Vector to V"
-}

{-700
Algebra  = M-Set record
Algebra′ = Algebra open-with-decoration "′"
Hom  = Algebra hom
Hom² = Algebra hom ⟴ renaming "map₁ to scalar; pres-𝟙 to unity" :adjoin-retract nil
-}

_ : {Src Tgt : Algebra} → Hom² Src Tgt → Algebra.Scalar Src → Algebra.Scalar Tgt
_ = Hom².scalar

{-700
-- regular expression test --

crazy-name-[]-+-\-^-*-? = M-Set extended-by "_+_ : Scalar; _*_ : Vector; ^ : Set; [_] : Set" :adjoin-retract nilrecord

PackageFormer MagmaP : Setwhere
  Carrier : Set
  op      : Carrier Carrier Carrier

Magma = MagmaPrecord

Pointed   = Magma extended-by "e : let Carrier = Carrier in Carrier"record
Additive+ = Pointed renaming "op to _+_; e to O; Carrier to C"record
Additive× = Additive+ renaming "_+_ to _×_"

crazy-name-test  = Pointed map (λ e → (map-name (λ n → (concat n "/crazy-name-[]-+-\-^-*-?")) e)) ⟴ record
crazy-name-test2 = crazy-name-test map (λ e → (map-name (λ n → (concat n "+2")) e)) ⟴ record
-}

{-700
M-Set-R′ = M-Set-R open-with-decoration "′"
-}
package-former-generated.agda
{- This file is generated ;; do not alter. -}

open import Level
open import Data.Bool
open import Data.List using (List; _∷_; []; foldr)
import Relation.Binary.PropositionalEquality as ≡; open ≡ using (_≡_)
import package-former-user-manual-i
open import Level as Level using (Level)
module package-former-generated where 


{- KindPackageFormerdoes not correspond  to a concrete Agda type.

PackageFormer M-Set : Setwhere
    Scalar      : Set
    Vector      : Set
    _·_     : Scalar Vector Vector
    𝟙       : Scalar
    _×_     : Scalar Scalar Scalar
    leftId      : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
    assoc       : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
 -}





{- Kind “PackageFormer” does not correspond  to a concrete Agda type. 

PackageFormer MonoidP : Set₁ where
    Carrier     : Set
    _⨾_     : Carrier → Carrier → Carrier
    Id      : Carrier
    assoc       : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)
    _≈_     : Carrier → Carrier → Set ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier → Carrier ;   Lid x = Id ⨾ x
    Rid     : Carrier → Carrier ;   Rid x = x ⨾ Id
    concat      : List Carrier → Carrier ;  concat = foldr _⨾_ Id
    leftId      : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    rightId     : ∀ {x : Carrier} → Rid x ≈ x
    Id²     : (Id ⨾ Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier → Carrier ;  concatₚ []       = Id ; concatₚ (x ∷ xs) = x ⨾ concatₚ xs
 -}


{- KindPackageFormerdoes not correspond  to a concrete Agda type.
{- MonoidPⁱᵈ = MonoidP identity -}
PackageFormer MonoidPⁱᵈ : Set₁ where
    Carrier     : Set
    _⨾_     : Carrier → Carrier → Carrier
    Id      : Carrier
    assoc       : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)
    _≈_     : Carrier → Carrier → Set ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier → Carrier ;   Lid x = Id ⨾ x
    Rid     : Carrier → Carrier ;   Rid x = x ⨾ Id
    concat      : List Carrier → Carrier ;  concat = foldr _⨾_ Id
    leftId      : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    rightId     : ∀ {x : Carrier} → Rid x ≈ x
    Id²     : (Id ⨾ Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier → Carrier ;  concatₚ []       = Id ; concatₚ (x ∷ xs) = x ⨾ concatₚ xs
 -}


{- KindPackageFormerdoes not correspond  to a concrete Agda type.
{- MonoidP⁰  = MonoidP -}
PackageFormer MonoidP⁰ : Set₁ where
    Carrier     : Set
    _⨾_     : Carrier → Carrier → Carrier
    Id      : Carrier
    assoc       : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)
    _≈_     : Carrier → Carrier → Set ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier → Carrier ;   Lid x = Id ⨾ x
    Rid     : Carrier → Carrier ;   Rid x = x ⨾ Id
    concat      : List Carrier → Carrier ;  concat = foldr _⨾_ Id
    leftId      : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    rightId     : ∀ {x : Carrier} → Rid x ≈ x
    Id²     : (Id ⨾ Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier → Carrier ;  concatₚ []       = Id ; concatₚ (x ∷ xs) = x ⨾ concatₚ xs
 -}


{- KindPackageFormerdoes not correspond  to a concrete Agda type.
{- MonoidPᶜ = MonoidP ⟴ -}
PackageFormer MonoidPᶜ : Set₁ where
    Carrier     : Set
    _⨾_     : Carrier → Carrier → Carrier
    Id      : Carrier
    assoc       : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)
    _≈_     : Carrier → Carrier → Set ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier → Carrier ;   Lid x = Id ⨾ x
    Rid     : Carrier → Carrier ;   Rid x = x ⨾ Id
    concat      : List Carrier → Carrier ;  concat = foldr _⨾_ Id
    leftId      : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    rightId     : ∀ {x : Carrier} → Rid x ≈ x
    Id²     : (Id ⨾ Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier → Carrier ;  concatₚ []       = Id ; concatₚ (x ∷ xs) = x ⨾ concatₚ xs
 -}


{- KindPackageFormerdoes not correspond  to a concrete Agda type.
{- Monoid-test = MonoidP ⟴ test "positional arg₁" "positional arg₂" :keyword 25 -}
PackageFormer Monoid-test : Set₁ where
    Carrier     : Set
    _⨾_     : Carrier → Carrier → Carrier
    Id      : Carrier
    assoc       : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)
    _≈_     : Carrier → Carrier → Set ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier → Carrier ;   Lid x = Id ⨾ x
    Rid     : Carrier → Carrier ;   Rid x = x ⨾ Id
    concat      : List Carrier → Carrier ;  concat = foldr _⨾_ Id
    leftId      : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    rightId     : ∀ {x : Carrier} → Rid x ≈ x
    Id²     : (Id ⨾ Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier → Carrier ;  concatₚ []       = Id ; concatₚ (x ∷ xs) = x ⨾ concatₚ xs
 -}


{- M-Set-Record = M-Set record₀ -}
record M-Set-Record : Setwhere
    field Scalar        : Set
    field Vector        : Set
    field _·_       : Scalar Vector Vector
    field 𝟙     : Scalar
    field _×_       : Scalar Scalar Scalar
    field leftId        : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
    field assoc     : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)


{- Monoid-Record-derived = MonoidP record₁ -}
record Monoid-Record-derived : Set₁ where
    field Carrier       : Set
    field _⨾_       : Carrier → Carrier → Carrier
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)
    _≈_     : Carrier → Carrier → Set ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier → Carrier ;   Lid x = Id ⨾ x
    Rid     : Carrier → Carrier ;   Rid x = x ⨾ Id
    concat      : List Carrier → Carrier ;  concat = foldr _⨾_ Id
    field leftId        : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid x ≈ x
    Id²     : (Id ⨾ Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier → Carrier ;  concatₚ []       = Id ; concatₚ (x ∷ xs) = x ⨾ concatₚ xs


{- Monoid-Record-field = MonoidP record₁ :discard-equations t -}
record Monoid-Record-field : Setwhere
    field Carrier       : Set
    field _⨾_       : Carrier Carrier Carrier
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)
    field _≈_       : Carrier CarrierSet
    field ⨾-cong        : ∀ {x y xy′} → xx′ →  yy′ → (xy) ≈ (x′ ⨾ y′)
    field Lid       : Carrier Carrier
    field Rid       : Carrier Carrier
    field concat        : List Carrier Carrier
    field leftId        : ∀ {x : Carrier} → (Idx) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid xx
    field Id²       : (Id Id) ≈ Id
    field concatₚ       : List Carrier Carrier


{- Monoid-Record-derived-again  = MonoidP record -}
record Monoid-Record-derived-again : Set₁ where
    field Carrier       : Set
    field _⨾_       : Carrier → Carrier → Carrier
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)
    _≈_     : Carrier → Carrier → Set ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier → Carrier ;   Lid x = Id ⨾ x
    Rid     : Carrier → Carrier ;   Rid x = x ⨾ Id
    concat      : List Carrier → Carrier ;  concat = foldr _⨾_ Id
    field leftId        : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid x ≈ x
    Id²     : (Id ⨾ Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier → Carrier ;  concatₚ []       = Id ; concatₚ (x ∷ xs) = x ⨾ concatₚ xs


{- Monoid-Record-derived-again2 = MonoidP record :and-names t -}
record Monoid-Record-derived-again2 : Setwhere
    field Carrier       : Set
    field _⨾_       : Carrier Carrier Carrier
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)
    _≈_     : Carrier CarrierSet ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y xy′} → xx′ →  yy′ → (xy) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier Carrier ;   Lid x = Idx
    Rid     : Carrier Carrier ;   Rid x = x Id
    concat      : List Carrier Carrier ;  concat = foldr _⨾_ Id
    field leftId        : ∀ {x : Carrier} → (Idx) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid xx
    Id²     : (Id Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier Carrier ;  concatₚ []       = Id ; concatₚ (xxs) = xconcatxs


{- Monoid-Record-field-again    = MonoidP record :discard-equations t -}
record Monoid-Record-field-again : Set₁ where
    field Carrier       : Set
    field _⨾_       : Carrier → Carrier → Carrier
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)
    field _≈_       : Carrier → Carrier → Set
    field ⨾-cong        : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′)
    field Lid       : Carrier → Carrier
    field Rid       : Carrier → Carrier
    field concat        : List Carrier → Carrier
    field leftId        : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid x ≈ x
    field Id²       : (Id ⨾ Id) ≈ Id
    field concatₚ       : List Carrier → Carrier


{- Monoid-Record-no-equationals = MonoidP record :discard-equations t :and-names t -}
record Monoid-Record-no-equationals : Setwhere
    field Carrier       : Set
    field _⨾_       : Carrier Carrier Carrier
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)
    field _≈_       : Carrier CarrierSet
    field Rid       : Carrier Carrier
    field leftId        : ∀ {x : Carrier} → (Idx) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid xx


{- M-Set-TypeClass = M-Set typeclass-attempt -}
record M-Set-TypeClass (Scalar : Set) (Vector : Set) : Set₁ where
    field _·_       : Scalar → Vector → Vector
    field 𝟙     : Scalar
    field _×_       : Scalar → Scalar → Scalar
    field leftId        : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
    field assoc     : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)


{- MonoidT= MonoidP typeclass₂ -}
record MonoidT (Carrier : Set) (_⨾_ : Carrier Carrier Carrier) : Set where
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)
    _≈_     : Carrier CarrierSet ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y xy′} → xx′ →  yy′ → (xy) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier Carrier ;   Lid x = Idx
    Rid     : Carrier Carrier ;   Rid x = x Id
    concat      : List Carrier Carrier ;  concat = foldr _⨾_ Id
    field leftId        : ∀ {x : Carrier} → (Idx) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid xx
    Id²     : (Id Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier Carrier ;  concatₚ []       = Id ; concatₚ (xxs) = xconcatxs


{- MonoidT₃         = MonoidP record ⟴ :waist 3 :level dec -}
record MonoidT₃ (Carrier : Set) (_⨾_ : Carrier → Carrier → Carrier) (Id : Carrier) : Set where
    field assoc     : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)
    _≈_     : Carrier → Carrier → Set ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier → Carrier ;   Lid x = Id ⨾ x
    Rid     : Carrier → Carrier ;   Rid x = x ⨾ Id
    concat      : List Carrier → Carrier ;  concat = foldr _⨾_ Id
    field leftId        : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid x ≈ x
    Id²     : (Id ⨾ Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier → Carrier ;  concatₚ []       = Id ; concatₚ (x ∷ xs) = x ⨾ concatₚ xs


{- M-Set-Typeclass= M-Set recordtypeclass₂ -}
record M-Set-Typeclass (Scalar : Set) (Vector : Set) : Set where
    field _·_       : Scalar Vector Vector
    field 𝟙     : Scalar
    field _×_       : Scalar Scalar Scalar
    field leftId        : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
    field assoc     : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)


{- Kind “PackageFormer” does not correspond  to a concrete Agda type. 
{- MonoidP-Typeclass= MonoidP :waist 5 -}
PackageFormer MonoidP-Typeclass (Carrier : Set) (_⨾_ : Carrier Carrier Carrier) (Id : Carrier) (assoc : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)) (_≈_ : Carrier CarrierSet ; _≈_ = _≡_) : Setwhere
    ⨾-cong      : ∀ {x y xy′} → xx′ →  yy′ → (xy) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier Carrier ;   Lid x = Idx
    Rid     : Carrier Carrier ;   Rid x = x Id
    concat      : List Carrier Carrier ;  concat = foldr _⨾_ Id
    leftId      : ∀ {x : Carrier} → (Idx) ≈ x
    rightId     : ∀ {x : Carrier} → Rid xx
    Id²     : (Id Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier Carrier ;  concatₚ []       = Id ; concatₚ (xxs) = xconcatxs
 -}


{- MonoidT₅ = MonoidP ⟴ unbundling 5 ⟴ record -}
record MonoidT₅ (Carrier : Set) (_⨾_ : Carrier → Carrier → Carrier) (Id : Carrier) (assoc : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)) (_≈_ : Carrier → Carrier → Set) : Set₁ where


{- KindPackageFormerdoes not correspond  to a concrete Agda type.
{- M-Set′-attempt-raw = M-Set primed-attempt -}
PackageFormer M-Set′-attempt-raw : Set₁ where
    Scalar′     : Set
    Vector′     : Set
    _·′_        : Scalar → Vector → Vector
    𝟙′      : Scalar
    _×′_        : Scalar → Scalar → Scalar
    leftId′     : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
    assoc′      : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
 -}


{- MR𝕏    = M-Set recordmap (λ e → (map-name (λ n → (rename-mixfix (λ x → (concat x "𝕏")) n)) e)) -}
record MR𝕏 : Setwhere
    field Scalar𝕏       : Set
    field Vector𝕏       : Set
    field _·𝕏_      : Scalar𝕏 → Vector𝕏 → Vector𝕏
    field 𝟙𝕏        : Scalar𝕏
    field _×𝕏_      : Scalar𝕏 → Scalar𝕏 → Scalar𝕏
    field leftId𝕏       : {𝓋 : Vector𝕏}  →  𝟙𝕏 ·𝕏 𝓋  ≡  𝓋
    field assoc𝕏        : {a b : Scalar𝕏} {𝓋 : Vector𝕏} → (a ×𝕏 b) ·𝕏 𝓋  ≡  a ·𝕏 (b ·𝕏 𝓋)


{- MR𝕪    = M-Set-Record rename (λ n → (concat n "𝕪")) -}
record MR𝕪 : Set₁ where
    field Scalar𝕪       : Set
    field Vector𝕪       : Set
    field _·𝕪_      : Scalar𝕪 → Vector𝕪 → Vector𝕪
    field 𝟙𝕪        : Scalar𝕪
    field _×𝕪_      : Scalar𝕪 → Scalar𝕪 → Scalar𝕪
    field leftId𝕪       : {𝓋 : Vector𝕪}  →  𝟙𝕪 ·𝕪 𝓋  ≡  𝓋
    field assoc𝕪        : {a b : Scalar𝕪} {𝓋 : Vector𝕪} → (a ×𝕪 b) ·𝕪 𝓋  ≡  a ·𝕪 (b ·𝕪 𝓋)
    toM-Set-Record      : let View X = X in View M-Set-Record ; toM-Set-Record = record {Scalar = Scalar𝕪;Vector = Vector𝕪;_·_ = _·𝕪_;𝟙 = 𝟙𝕪;_×_ = _×𝕪_;leftId = leftId𝕪;assoc = assoc𝕪}


{- MR-oh  = M-Set-Record rename (λ n → (pcase n ("Scalar" "S") ("𝟙" "ε") (else else))) -}
record MR-oh : Setwhere
    field S     : Set
    field Vector        : Set
    field _·_       : S Vector Vector
    field ε     : S
    field _×_       : S S S
    field leftId        : {𝓋 : Vector}  →  ε · 𝓋  ≡  𝓋
    field assoc     : {a b : S} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
    toM-Set-Record      : let View X = X in View M-Set-Record ; toM-Set-Record = record {Scalar = S;Vector = Vector;_·_ = _·_;𝟙 = ε;_×_ = _×_;leftId = leftId;assoc = assoc}


{- the-MR = M-Set-Record co-decorated "the-" -}
record the-MR : Set₁ where
    field the-Scalar        : Set
    field the-Vector        : Set
    field _the-·_       : the-Scalar → the-Vector → the-Vector
    field the-𝟙     : the-Scalar
    field _the-×_       : the-Scalar → the-Scalar → the-Scalar
    field the-leftId        : {𝓋 : the-Vector}  →  the-𝟙 the-· 𝓋  ≡  𝓋
    field the-assoc     : {a b : the-Scalar} {𝓋 : the-Vector} → (a the-× b) the-· 𝓋  ≡  a the-· (b the-· 𝓋)
    toM-Set-Record      : let View X = X in View M-Set-Record ; toM-Set-Record = record {Scalar = the-Scalar;Vector = the-Vector;_·_ = _the-·_;𝟙 = the-𝟙;_×_ = _the-×_;leftId = the-leftId;assoc = the-assoc}


{- MRₜₒ   = M-Set-Record renaming "Scalar to S; Vector to V; · to nice" -}
record MRₜₒ : Setwhere
    field S     : Set
    field V     : Set
    field _·_       : S V V
    field 𝟙     : S
    field _×_       : S S S
    field leftId        : {𝓋 : V}  →  𝟙 · 𝓋  ≡  𝓋
    field assoc     : {a b : S} {𝓋 : V} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
    toM-Set-Record      : let View X = X in View M-Set-Record ; toM-Set-Record = record {Scalar = S;Vector = V;_·_ = _·_;𝟙 = 𝟙;_×_ = _×_;leftId = leftId;assoc = assoc}


{- NearMonoid = M-Set-Record renaming "Scalar to Carrier; Vector to Carrier; · to ×" -}
record NearMonoid : Set₁ where
    field Carrier       : Set
    field _·_       : Carrier → Carrier → Carrier
    field 𝟙     : Carrier
    field _×_       : Carrier → Carrier → Carrier
    field leftId        : {𝓋 : Carrier}  →  𝟙 · 𝓋  ≡  𝓋
    field assoc     : {a b : Carrier} {𝓋 : Carrier} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
    toM-Set-Record      : let View X = X in View M-Set-Record ; toM-Set-Record = record {Scalar = Carrier;Vector = Carrier;_·_ = _·_;𝟙 = 𝟙;_×_ = _×_;leftId = leftId;assoc = assoc}


{- NearMonoid¹ = M-Set-Record single-sorted "Carrier" -}
record NearMonoid¹ : Setwhere
    field Carrier       : Set
    field _·_       : Carrier Carrier Carrier
    field 𝟙     : Carrier
    field _×_       : Carrier Carrier Carrier
    field leftId        : {𝓋 : Carrier}  →  𝟙 · 𝓋  ≡  𝓋
    field assoc     : {a b : Carrier} {𝓋 : Carrier} → (a × b) · 𝓋  ≡  a · (b · 𝓋)


{- M-Set-Sorts = M-Set record ⟴ sorts -}
record M-Set-Sorts : Set₁ where
    field Scalar        : Set
    field Vector        : Set


{- MonoidSignature = M-Set-Record generated (λ e → (and (s-contains? "Scalar" (element-type e)) (not (s-contains? "Vector" (element-type e))))) -}
record MonoidSignature : Setwhere
    field Scalar        : Set
    field 𝟙     : Scalar
    field _×_       : Scalar Scalar Scalar


{- MonSig = M-Set-Record signature -}
record MonSig : Set₁ where
    field Scalar        : Set
    field Vector        : Set
    field _·_       : Scalar → Vector → Vector
    field 𝟙     : Scalar
    field _×_       : Scalar → Scalar → Scalar


{- Neato = M-Set empty-module -}
module Neato (Scalar : Set) (Vector : Set) (_·_ : Scalar Vector Vector) (𝟙 : Scalar) (_×_ : Scalar Scalar Scalar) (leftId : {𝓋 : Vector} → 𝟙 · 𝓋 ≡ 𝓋) (assoc : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋 ≡ a · (b · 𝓋)) where


{- M-Set-R = M-Set record -}
record M-Set-R : Set₁ where
    field Scalar        : Set
    field Vector        : Set
    field _·_       : Scalar → Vector → Vector
    field 𝟙     : Scalar
    field _×_       : Scalar → Scalar → Scalar
    field leftId        : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
    field assoc     : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)


{- M-Set-R= M-Set-Ropen (λ x → (concat x "₁")) -}
module M-Set-R (Arg37563 : M-Set-R) where
    Scalar₁     : let open M-Set-R Arg37563 in Set ;    Scalar= M-Set-R.Scalar Arg37563
    Vector₁     : let open M-Set-R Arg37563 in Set ;    Vector= M-Set-R.Vector Arg37563
    _·₁_        : let open M-Set-R Arg37563 in Scalar Vector Vector ;   _·₁_ = M-Set-R._·_ Arg37563
    𝟙₁      : let open M-Set-R Arg37563 in Scalar ; 𝟙₁ = M-Set-R.𝟙 Arg37563
    _×₁_        : let open M-Set-R Arg37563 in Scalar Scalar Scalar ;   _×₁_ = M-Set-R._×_ Arg37563
    leftId₁     : let open M-Set-R Arg37563 in {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 ;   leftId= M-Set-R.leftId Arg37563
    assoc₁      : let open M-Set-R Arg37563 in {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋) ;  assoc= M-Set-R.assoc Arg37563


{- M-Set-R-SV = M-Set-R opening "Scalar to S; Vector to V" -}
module M-Set-R-SV (Arg37564 : M-Set-R) where
    S       : let open M-Set-R Arg37564 in Set ;    S = M-Set-R.Scalar Arg37564
    V       : let open M-Set-R Arg37564 in Set ;    V = M-Set-R.Vector Arg37564
    _       : let open M-Set-R Arg37564 in Scalar → Vector → Vector ;   _ = M-Set-R._·_ Arg37564
    _       : let open M-Set-R Arg37564 in Scalar ; _ = M-Set-R.𝟙 Arg37564
    _       : let open M-Set-R Arg37564 in Scalar → Scalar → Scalar ;   _ = M-Set-R._×_ Arg37564
    _       : let open M-Set-R Arg37564 in {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 ;   _ = M-Set-R.leftId Arg37564
    _       : let open M-Set-R Arg37564 in {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋) ;  _ = M-Set-R.assoc Arg37564


{- Algebra  = M-Set record -}
record Algebra : Setwhere
    field Scalar        : Set
    field Vector        : Set
    field _·_       : Scalar Vector Vector
    field 𝟙     : Scalar
    field _×_       : Scalar Scalar Scalar
    field leftId        : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
    field assoc     : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)


{- Algebra′ = Algebra open-with-decoration "′" -}
module Algebra′ (Arg37565 : Algebra) where
    Scalar′     : let open Algebra Arg37565 in Set ;    Scalar′ = Algebra.Scalar Arg37565
    Vector′     : let open Algebra Arg37565 in Set ;    Vector′ = Algebra.Vector Arg37565
    _·′_        : let open Algebra Arg37565 in Scalar → Vector → Vector ;   _·′_ = Algebra._·_ Arg37565
    𝟙′      : let open Algebra Arg37565 in Scalar ; 𝟙′ = Algebra.𝟙 Arg37565
    _×′_        : let open Algebra Arg37565 in Scalar → Scalar → Scalar ;   _×′_ = Algebra._×_ Arg37565
    leftId′     : let open Algebra Arg37565 in {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 ;   leftId′ = Algebra.leftId Arg37565
    assoc′      : let open Algebra Arg37565 in {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋) ;  assoc′ = Algebra.assoc Arg37565


{- Hom  = Algebra hom -}
record Hom (Src : Algebra) (Tgt : Algebra) : Setwhere
    field map₁      : let open Algebra Src; open Algebra Tgt in Scalar Scalarfield map₂      : let open Algebra Src; open Algebra Tgt in Vector Vectorfield pres-·        : let open Algebra Src; open Algebra Tgt in {x₁ : Scalar} → {x₂ : Vector} →   map₂ (_·_ xx₂)   ≡   _·′_ (mapx₁) (mapx₂)
    field pres-𝟙        : let open Algebra Src; open Algebra Tgt in   map₁ (𝟙 )   ≡   𝟙′
    field pres-×        : let open Algebra Src; open Algebra Tgt in {x₁ : Scalar} → {x₁ : Scalar} →   map₁ (_×_ xx₁)   ≡   _×′_ (mapx₁) (mapx₁)


{- Hom² = Algebra hom ⟴ renaming "map₁ to scalar; pres-𝟙 to unity" :adjoin-retract nil -}
record Hom² (Src : Algebra) (Tgt : Algebra) : Set₁ where
    field scalar        : let open Algebra Src; open Algebra′ Tgt in Scalar → Scalar′
    field map₂      : let open Algebra Src; open Algebra′ Tgt in Vector → Vector′
    field pres-·        : let open Algebra Src; open Algebra′ Tgt in {x₁ : Scalar} → {x₂ : Vector} →   map₂ (_·_ x₁ x₂)   ≡   _·′_ (scalar x₁) (map₂ x₂)
    field unity     : let open Algebra Src; open Algebra′ Tgt in   scalar (𝟙 )   ≡   𝟙′ 
    field pres-×        : let open Algebra Src; open Algebra′ Tgt in {x₁ : Scalar} → {x₁ : Scalar} →   scalar (_×_ x₁ x₁)   ≡   _×′_ (scalar x₁) (scalar x₁)


{- crazy-name-[]-+-\-^-*-? = M-Set extended-by "_+_ : Scalar; _*_ : Vector; ^ : Set; [_] : Set" :adjoin-retract nilrecord -}
record crazy-name-[]-+-\-^-*-? : Setwhere
    field Scalar        : Set
    field Vector        : Set
    field _·_       : Scalar Vector Vector
    field 𝟙     : Scalar
    field _×_       : Scalar Scalar Scalar
    field leftId        : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
    field assoc     : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
    field _+_       : Scalar
    field _*_       : Vector
    field ^     : Set
    field [_]       : Set





{- Kind “PackageFormer” does not correspond  to a concrete Agda type. 

PackageFormer MagmaP : Set₁ where
    Carrier     : Set
    op      : Carrier → Carrier → Carrier
 -}


{- Magma = MagmaPrecord -}
record Magma : Setwhere
    field Carrier       : Set
    field op        : Carrier Carrier Carrier


{- Pointed   = Magma extended-by "e : let Carrier = Carrier in Carrier" ⟴ record -}
record Pointed : Set₁ where
    field Carrier       : Set
    field op        : Carrier → Carrier → Carrier
    field e     : let Carrier = Carrier in Carrier
    toMagma     : let View X = X in View Magma ;    toMagma = record {Carrier = Carrier;op = op}


{- Additive+ = Pointed renaming "op to _+_; e to O; Carrier to C"record -}
record Additive+ : Setwhere
    field C     : Set
    field _+_       : C C C
    field O     : let Carrier = C in C
    toMagma     : let View X = X in View Magma ;    toMagma = record {Carrier = C;op = _+_}
    toPointed       : let View X = X in View Pointed ;  toPointed = record {Carrier = C;op = _+_;e = O}


{- Additive× = Additive+ renaming "_+_ to _×_" -}
record Additive× : Set₁ where
    field C     : Set
    field _×_       : C → C → C
    field O     : let Carrier = C in C
    toMagma     : let View X = X in View Magma ;    toMagma = record {Carrier = C;op = _×_}
    toPointed       : let View X = X in View Pointed ;  toPointed = record {Carrier = C;op = _×_;e = O}
    toAdditive+     : let View X = X in View Additive+ ;    toAdditive+ = record {C = C;_+_ = _×_;O = O}


{- crazy-name-test  = Pointed map (λ e → (map-name (λ n → (concat n "/crazy-name-[]-+-\-^-*-?")) e)) ⟴ record -}
record crazy-name-test : Setwhere
    field Carrier/crazy-name-[]-+--^-*-?        : Set
    field op/crazy-name-[]-+--^-*-?     : Carrier/crazy-name-[]-+--^-*-? → Carrier/crazy-name-[]-+--^-*-? → Carrier/crazy-name-[]-+--^-*-?
    field e/crazy-name-[]-+--^-*-?      : let Carrier = Carrier/crazy-name-[]-+--^-*-? in Carrier/crazy-name-[]-+--^-*-?
    toMagma/crazy-name-[]-+--^-*-?      : let View X = X in View Magma ;    toMagma/crazy-name-[]-+--^-*-? = record {Carrier = Carrier/crazy-name-[]-+--^-*-?;op = op/crazy-name-[]-+--^-*-?}


{- crazy-name-test2 = crazy-name-test map (λ e → (map-name (λ n → (concat n "+2")) e)) ⟴ record -}
record crazy-name-test2 : Set₁ where
    field Carrier/crazy-name-[]-+--^-*-?+2      : Set
    field op/crazy-name-[]-+--^-*-?+2       : Carrier/crazy-name-[]-+--^-*-?+2 → Carrier/crazy-name-[]-+--^-*-?+2 → Carrier/crazy-name-[]-+--^-*-?+2
    field e/crazy-name-[]-+--^-*-?+2        : let Carrier = Carrier/crazy-name-[]-+--^-*-?+2 in Carrier/crazy-name-[]-+--^-*-?+2
    toMagma/crazy-name-[]-+--^-*-?+2        : let View X = X in View Magma ;    toMagma/crazy-name-[]-+--^-*-?+2 = record {Carrier = Carrier/crazy-name-[]-+--^-*-?+2;op = op/crazy-name-[]-+--^-*-?+2}


{- M-Set-R= M-Set-R open-with-decoration "′" -}
module M-Set-R (Arg37691 : M-Set-R) where
    Scalar′     : let open M-Set-R Arg37691 in Set ;    Scalar= M-Set-R.Scalar Arg37691
    Vector′     : let open M-Set-R Arg37691 in Set ;    Vector= M-Set-R.Vector Arg37691
    _·′_        : let open M-Set-R Arg37691 in Scalar Vector Vector ;   _·′_ = M-Set-R._·_ Arg37691
    𝟙′      : let open M-Set-R Arg37691 in Scalar ; 𝟙′ = M-Set-R.𝟙 Arg37691
    _×′_        : let open M-Set-R Arg37691 in Scalar Scalar Scalar ;   _×′_ = M-Set-R._×_ Arg37691
    leftId′     : let open M-Set-R Arg37691 in {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 ;   leftId= M-Set-R.leftId Arg37691
    assoc′      : let open M-Set-R Arg37691 in {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋) ;  assoc= M-Set-R.assoc Arg37691

Above is a sample source file which contains special comments that are picked up by the prototype which then copy-paste-cuts to produce a generated file. The above code is discussed below, in the user manual —in fact, it's generated from the user-manual below.

For this prototype, we have the following constraints:

  1. The type of a PackageFormer is Set ℓ where is the empty string or a parenthesised expression of type Level.
    • In-particular, subscript types are not yet supported.
  2. The where keyword appears on the same line as the PackageFormer key-phrase.
  3. The name of the PackageFormer should not contain PackageFormer as a sub-identifier.
  4. Each element of a PackageFormer spans only one physical line.

There are many useful features outlined in the proposal, such as default implementations, that we hope to include in the future. For now, we just want something that works, is decently documented, and can be useful.

2 Contributions

Current state of affairs:

  1. Expressive & extendable specification language for the library developer.
    • We demonstrate that our meta-primitives permit this below by demonstrating that ubiquitous module combinators can be easily formalised and easily used.
    • E.g., from a theory we can derive its homomorphism type, signature, its termtype, etc; we generate useful constructions inspired from universal algebra.
    • An example of the freedom allotted by the extensible nature of the system is that combinators defined by library developers can, say, utilise auto-generated names when names are irrelevant, use ‘clever’ default names, and allow end-users to supply desirable names on demand.
  2. Unobtrusive and a tremendously simple interface to the end user.
    • Once a library is developed using (the current implementation of) PackageFormers, the end user only needs to reference the resulting generated Agda, without any knowledge of the existence of PackageFormers.
      • Generated modules are necessarily ‘flattened’ for typechecking with Agda.
    • We demonstrate below how end-users can build upon a library by using one line specifications, by showing over over 200 specifications of mathematical structures.
  3. Efficient: Our current implementation processes over 200 specifications in ~3 seconds; yielding typechecked Agda code.
    • It is the typechecking that takes time.
  4. Pragmatic: We demonstrate how common combinators can be defined for library developers, but also how they can be furnished with concrete syntax —inspired by Agda's— for use by end-users.
  5. Minimal: The system is essentially invariant over the underlying type system; with the exception of the meta-primitive :waist which requires a dependent type theory to express ‘unbundling’ component fields as parameters.
  6. Demonstrated expressive power and use-cases.
    • Common boiler-plate idioms in the standard Agda library, and other places, are provided with terse solutions using the PackageFormer system.
      • E.g., automatically generating homomorphism types and wholesale renaming fields using a single function.
    • Over 200 modules are formalised as one-line specifications.
  7. Immediately useable to end-users and library developers.
    • We have provided a large library to experiment with —thanks to the MathScheme group for providing an adaptable source file.
    • In the second part of the user manual, we show how to formulate module combinators using a simple and straightforward subset of Emacs Lisp —a terse introduction is provided.

Intuitively, the nature of the meta-primitives:

  1. A PackageFormer is a context, a signature, an essentially algebraic theory of Freyd, or a generalised algebraic theory of Cartmell. Moreover, it is tagged by some values for practical purposes; e.g., the kind of concrete realisation that is possible in Agda —namely, record, data, module, or the abstract PackageFormer.
  2. A ‘variational’ is a morphism between PackageFormers —taking signatures to signatures.

    Built up from the grammar:

    𝓋 ∷=                   -- empty, identity, variational
       | 𝓋  𝓋            -- composition
       | :kind  𝓀          -- 𝓀 ∈ {record, data, module, PackageFormer}
       | :waist 𝓃          -- n ∈ ℕ, number of initial items to be considered as ‘parameters’
       | :alter-elements 𝒻 -- f : List Elements → List Elements
    

    :alter-elements is a sledgehammer that may result in ill-formed signatures, but we leave it in the system due to its power. Instead, we recommend using the following derivied primitives for signature catenation, map, and filter:

       | extended-by ds  -- Adjoin declarations ds, “name : type”, to a PackageFormer
       | map         f   -- Alter each element by f : Element → Element
       | generated   f   -- Keep the largest well-formed PackageFormer whose elements satisfy predicate f
    

    Why are these sufficient? The first homomorphism theorem of lists —i.e., the fact that lists are free monoids— informs us that all well-behaved functions from a list monoid are determined as folds after maps. Since our signatures are essentially free monoids and the target of our functions are again the same free monoids, the fold is determined and only the map remains. We conjecture this is enough; we are not yet certain.

    Interestingly, we are mostly generic over the underlying type theory.

    Since a PackageFormer corresponds to a signature, then these variationals —at least the well-behaved ones— correspond to signature morphisms. Great difficulty lies in providing semantics for alter-elements; if we demand a well-typedness judgement, “⊢”, from our underlying type theory then we could define (x₀ : τ₀, …, xₙ ∶ τₙ) = Γ :alter-elements f to be well-typed iff f : List Elements → List Elements and Γ,x₀:τ₀,…,xₖ₋₁:τₖ₋₁ ⊢ τₖ Type for k : 0..n.

  3. We have a categorical structure consisting of PackageFormers as objects and those variationals that are signature morphisms.

Research outcomes:

  1. Narrow down the meta-primitives that permit a variety of algorithms for generating universal algebra consructions —the former being the kerneal which has application the latter.
  2. Realise this for Agda, likely using an editor-extension.
  3. Provide a semantics to the existing syntax.
  4. Ensure the resulting semantics is consistent with that of Agda's.
Any questions or feedback are welcome!

3 User Manual I: Simple Use of the System

If the previous section is unclear regarding the aims and uses of this prototype, please consult the pre-print A Language Feature to Unbundle Data at Will or the next 700 module systems proposal.

Herein we demonstrate how to use this system from the perspective of library designers. We use constructs that are discussed in the next section —which are examples of how users may extend the system to produce grouping mechanisms for any desired purpose. The exposition here follows section 2 of the Theory Presentation Combinators [tpc] paper, reiterating many the ideas therein.

The few constructs demonstrated in this section not only create new grouping mechanisms from old ones, but also create maps from the new, child, presentations to the old parent presentations. Maps between grouping mechanisms are sometimes called views. For example, a theory extended by new declarations comes equipped with a map that forgets the new declarations to obtain an instance of the original theory. Such morphisms are tedious to write out, and our system provides them for free. How? You, the user, can implement such features using our 5 meta-primitives —but we have implemented a few for you as examples.

It is important to clarify that the purpose of this work is the development of a core kernel of meta-primitives for modules. This section demonstrates the power and expressivity of the meta-primitives by showcasing a series of ubiquitous combinators which may be defined using the meta-primitives and Lisp. The section afterwards goes into the detail of how to extend the system to build any desired operations on any notion of grouping mechanism.

3.1 Installation

Obtaining the system!

Add the following to the top of your Emacs configuration file, i.e., the ~/.emacs file.

;; Connect to internet repositories for Emacs packages.
(require 'package)
(push '("melpa-stable" . "http://stable.melpa.org/packages/") package-archives)
(package-initialize)
(package-refresh-contents)

;; Obtain & setup installation interface.
(unless (package-installed-p 'use-package)
  (package-install 'use-package))
(require 'use-package)
(setq use-package-always-ensure t)

;; Necessary libraries for producing the prototype
(use-package s)                  ;; “The long lost Emacs string manipulation library”.
(use-package dash)               ;; “A modern list library for Emacs”.
(use-package dash-functional)    ;; Function combinators; e.g., -partial/-cut, -const, -compose, -orfn & -andfn for generalised ∃/∀.
(use-package origami)            ;; Folding away regions of text.
(use-package hydra)              ;; Helpful menus.
(require 'subr-x)                ;; Extra Lisp functions; e.g., when-let.
;; Eye-candy
(use-package spacemacs-common
    :ensure spacemacs-theme
    :config (load-theme 'spacemacs-light t))

;; Next, obtain the Elisp file, load it, and attach it to Agda.
(shell-command (concat "curl "
    "https://raw.githubusercontent.com/alhassy/next-700-module-systems/master/prototype/agda-next-700-module-systems.el"
    ">> ~/.emacs.d/agda-next-700-module-systems.el"))
(load-file "~/.emacs.d/agda-next-700-module-systems.el")

;; Uncomment if you want this extenssion to ALWAYS be active on .agda files.
;; (add-hook 'agda2-mode-hook #'agda-next-700-module-systems-mode)

;; You likely have this in your ~/.emacs file already
(load-file (let ((coding-system-for-read 'utf-8))
                (shell-command-to-string "/usr/local/bin/agda-mode locate")))

When you enable agda-next-700-module-systems-mode,

  1. A menu-bar PackageFormers will be added.
  2. It will allow you to temporarily disable and enable this new feature, as well as providing a help menu. Invoke M-x agda-next-700-module-systems-mode to toggle turning off this feature completely.
  3. The string icon PackageFormer (•̀ᴗ•́)و is displayed in the mode line —near the bottom of Emacs.
  4. You may use C-c c-l as usual, but it will now recognise 700-comments and generate legitimate Agda code from them —more on this later.
    • PackageFormer syntactical items are coloured green, PackageFormer names are coloured yellow, and their instantiation are simply bolded.

If you need any assistance, please contact me!

3.2 Syntax

Superficial glance at the system's syntax.

The prototype works by translating fictitious 700-syntax into legitimate Agda; as follows:

...agda code here...
{-700
       ...700-syntactical items here...
-}
...more agda code...

Since the first section provides an example source fragment with both 700-comments as well as instantiations, we shall only enclose 700-syntax in 700-comments when it is surrounded by other Agda code, and otherwise leave it free standing.

We will provide full source listings at the end of discussions that only display fragments!

700-syntax is defined informally as follows:

⟪700-syntax⟫    ::=  ⟪PackageFormer⟫ | ⟪Instantiation⟫ | ⟪Agda⟫

⟪PackageFormer⟫ ::= PackageFormer ⟪Identifier⟫ : Set (⟪level⟫) where
               ⟪newline-with-indentation⟫ ⟪Element⟫*

⟪Element⟫       ::=  ⟪Identifier⟫ : ⟪Any-Agda-Type⟫

⟪Instantiation⟫ ::= ⟪Identifier⟫ = ⟪Identifier⟫ ⟪VariationalClause⟫

⟪VariationalClause⟫ ::= [⟪Identifier⟫] (:key (value))* (⟴ ⟪VariationalClause⟫)*
  • One derives many presentations of a grouping mechanism by what we call ‘variational clauses’.
    • In a 700-comment, one declares ‘variational’ such as

      𝒱-typeclass height = :kind record :level dec :waist-strings ("field") :waist height

      These are functions whose names begin with 𝒱-, they may have arguments on the left-hand-side, and their right hand side may invoke any of the 5 meta-primitives kind, waist, waist-strings, level, alter-elements with any mixture of arguments and concrete values.

      • To invoke a variational in an instantiation clause, arguments are not positional but instead are passed by name —e.g., :key value.
  • Example uses of the variational clauses could be seen in the package-former.agda listing in the first section above.

3.3 Extension

The simplest situation is where the presentation of one theory is included, verbatim, in another. Concretely, consider Monoid and CommutativeMonoid.

{-700
PackageFormer Monoid : Set₁ where
   Carrier : Set
   _·_     : Carrier → Carrier → Carrier
   assoc   : {x y z : Carrier} → (x · y) · z  ≡  x · (y · z)
   𝕀       : Carrier
   leftId  : {x : Carrier} → 𝕀 · x  ≡ x
   rightId : {x : Carrier} → x · 𝕀  ≡ x
   𝕀-unique : ∀ {e} (lid : ∀ {x} → e · x ≡ x) (rid : ∀ {x} → x · e ≡ x) → e ≡ 𝕀
   𝕀-unique lid rid = ≡.trans (≡.sym leftId) rid

PackageFormer CommutativeMonoid₀ : Set₁ where
   Carrier : Set
   _·_     : Carrier → Carrier → Carrier
   assoc   : {x y z : Carrier} → (x · y) · z  ≡  x · (y · z)
   𝕀       : Carrier
   leftId  : {x : Carrier} →  𝕀 · x  ≡ x
   rightId : {x : Carrier} →  x · 𝕀  ≡ x
   comm    : {x y : Carrier} →  x · y  ≡  y · x
   𝕀-unique : ∀ {e} (lid : ∀ {x} → e · x ≡ x) (rid : ∀ {x} → x · e ≡ x) → e ≡ 𝕀
   𝕀-unique lid rid = ≡.trans (≡.sym leftId) rid
-}

As expected, the only difference is that CommutativeMonoid₀ adds a comm-utative axiom. Thus, given Monoid, it would be more economical to define:

{-700
CommutativeMonoid = Monoid extended-by "comm : {x y : Carrier} →  x · y  ≡  y · x"
-}

Hovering over the left-hand-side gives a tooltip showing the resulting elaboration, which is identical to CommutativeMonoid₀ along with a forgetful operation 😄 The tooltip shows the expanded version of the theory, which is what we want to specify but not what we want to enter manually. To obtain this specification of CommutativeMonoid in the current implementation of Agda, one would likely declare a record with two fields —one being a Monoid and the other being the commutativity constraint— however, this only gives the appearance of the above specification for consumers; those who produce instances of CommutativeMonoid are then forced to know the particular hierarchy and must provide a Monoid value first. It is a happy coincidence that our system alleviates such an issue.

Alternatively, we may reify the new syntactical items as concrete Agda supported record-s as follows.

{-700
MonoidR            = Monoid ⟴ record
CommutativeMonoidR = MonoidR extended-by "comm : {x y : Carrier} →  x · y  ≡  y · x" ⟴ record
-}

neato : CommutativeMonoidR → MonoidR
neato = CommutativeMonoidR.toMonoidR

“Transport” It is important to notice that the derived result 𝕀-unique, while proven in the setting of Monoid, is not only available via the morphism toMonoidR but is also available directly since it is also a member of CommutativeMonoidR.

As a teaser, here's how the extended-by variational is defined —we leave the details of which to the second part of the user manual.

Definition of extended-by
(eval-and-compile
(cl-defun element-retract (parent es &key (new es) name contravariant)
  "Realise a list of elements as an Agda no-op record.

E.g., list “Carrier : Set; e : Carrier”
maps to the following element value.

      toParent : parent
      toParent = record {Carrier = Carrier; e = e}

See also 𝒱-renaming, which may be useful to change ‘toParent’.

NEW is a new updated version of ES, if any.

NAME is the name of the new retract element; by default it's “toParent” or “fromParent”
depending on whether CONTRAVARIANT is true or not.
"

  ;; the name of the newly defined PackageFormer ---which we may access using the special identifier ~$𝑛𝑎𝑚𝑒~.
  (let* ((toParent (or (unless (equal t name) name) (format "%s%s" (if contravariant "from" "to") parent)))
         ;; NAME may be "t", but not the symbol t. This is useful in any 𝒱artional that has an optional adjoin-retract
         ;; argument, which the user may set to be nil, t, or a string to obtain nothing, default name, and given name for the morphism; respectively.
        (τarg     (if contravariant (format "%s → " parent) ""))
        (δvar     (gensym)) ;; unique argument name to avoid accidental shawdowing of any “e in es”.
        (δarg     (if contravariant (format "λ %s → " δvar) ""))
        )
    (car (parse-elements (list
      (format "%s : let View X = X in %sView %s" toParent τarg (if contravariant $𝑛𝑎𝑚𝑒 parent))
      (format "%s = %srecord {%s}" toParent δarg

        (s-join ";"
        (loop for e  in es
              for e′ in new
              unless (or (s-contains-p "let View X = X" (element-type e)) ;; Ignore source view morphisms
                         (element-equations e))                           ;; Ignore “derivied” elements
              collect (if (not contravariant)
                          (format "%s = %s" (element-name e) (element-name e′))
                        (format "%s = %s.%s %s" (element-name e) parent (element-name e′) δvar))))))))))

(𝒱 extended-by ds (adjoin-retract t)
   = "Extend a given presentation by a list of ;-separated declarations.

      The resuling presentation has a “toX” retract method,
      where ‘X’ is the parent presentation. To avoid this,
      set ADJOIN-RETRACT to be nil. To provide a preferred name for
      the morphism, then set ADJOIN-RETRACT to the desired string.
     "
     :alter-elements (λ es → (-concat es (parse-elements (mapcar #'s-trim (s-split ";" ds))) (when adjoin-retract (list (element-retract $𝑝𝑎𝑟𝑒𝑛𝑡 es :name adjoin-retract))))))
)
𝒱-extended-by
  • Simple exercise: Play with this setup to observe that extended-by is an idempotent operation.
  • Advanced exercise: Read the user manual, parts I and II, then produce a variational commuting such that CommutativeMagma can be regained by Magma commuting "·".

    See the flipping variational below.

Anyhow, notice that we may define GroupR —a record-presentation of groups— as an extension of MonoidR using a single extended-by clause where the necessary items are separated by ;.

{-700
GroupR = MonoidR extended-by "_⁻¹ : Carrier → Carrier; left⁻¹ : ∀ {x} → (x ⁻¹) · x ≡ 𝕀; right⁻¹ : ∀ {x} → x · (x ⁻¹) ≡ 𝕀" ⟴ record
-}

A more fine grained approach may be as follows.

{-700
PackageFormer Empty : Set₁ where {- No elements -}
Type  = Empty extended-by "Carrier : Set" :adjoin-retract nilrecord
Magma = Type  extended-by "_·_ : Carrier → Carrier → Carrier"record
CommutativeMagma = Magma extended-by "comm : {x y : Carrier} →  x · y  ≡  y · x"record
-}

3.4 Defining a Concept Only Once

From a library-designer's perspective, our definition of CommutativeMonoid has the commutativity property ‘hard coded’ into it. If we wish to speak of commutative magmas —types with a single commutative operation— we need to hard-code the property once again. If, at a later time, we wish to move from having arguments be implicit to being explicit then we need to track down every hard-coded instance of the property then alter them —having them in-sync becomes an issue.

Instead, the system lets us ‘build upon’ the extended-by combinator: We make an associative list of names and properties, then string-replace the meta-names op, op′, rel with the provided user names.

(𝒱 postulating bop prop (using bop) (adjoin-retract t)
 = "Adjoin a property PROP for a given binary operation BOP.

   PROP may be a string: associative, commutative, idempotent, etc.

   Some properties require another operator or a relation; which may
   be provided via USING.

   ADJOIN-RETRACT is the optional name of the resulting retract morphism.
   Provide nil if you do not want the morphism adjoined.

   With this variational, a definition is only written once.
   "
   extended-by (s-replace "op" bop (s-replace "rel" using (s-replace "op′" using
    (pcase prop
     ("associative"   "assoc : ∀ x y z → op (op x y) z ≡ op x (op y z)")
     ("commutative"   "comm  : ∀ x y   → op x y ≡ op y x")
     ("idempotent"    "idemp : ∀ x     → op x x ≡ x")
     ("involutive"    "inv   : ∀ x     → op (op x) ≡ x") ;; assuming bop is unary
     ("left-unit"     "unitˡ : ∀ x y z → op e x ≡ e")
     ("right-unit"    "unitʳ : ∀ x y z → op x e ≡ e")
     ("distributiveˡ" "distˡ : ∀ x y z → op x (op′ y z) ≡ op′ (op x y) (op x z)")
     ("distributiveʳ" "distʳ : ∀ x y z → op (op′ y z) x ≡ op′ (op y x) (op z x)")
     ("absorptive"    "absorp  : ∀ x y  → op x (op′ x y) ≡ x")
     ("reflexive"     "refl    : ∀ x y  → rel x x")
     ("transitive"    "trans   : ∀ x y z → rel x y → rel y z → rel x z")
     ("antisymmetric" "antisym : ∀ x y → rel x y → rel y x → x ≡ z")
     ("congruence"    "cong    : ∀ x x′ y y′ → rel x x′ → rel y y′ → rel (op x x′) (op y y′)")
     ;; (_ (error "𝒱-postulating does not know the property “%s”" prop))
     )))) :adjoin-retract 'adjoin-retract)
𝒱-postulating

We can extend this database of properties as needed with relative ease. Here is an example use along with its elaboration.

{-700
PackageFormer Magma : Set₁ where
  Carrier : Set
  _·_      : Carrier → Carrier → Carrier

RawRelationalMagma = Magma extended-by "_≈_ : Carrier → Carrier → Set" ⟴ record

RelationalMagma    = RawRelationalMagma postulating "_·_" "congruence" :using "_≈_" ⟴ record
-}
record RawRelationalMagma : Setwhere
    field Carrier       : Set
    field op        : Carrier Carrier Carrier
    toType      : let View X = X in View Type ; toType = record {Carrier = Carrier}
    field _≈_       : Carrier CarrierSet
    toMagma     : let View X = X in View Magma ;    toMagma = record {Carrier = Carrier;op = op}

record RelationalMagma : Setwhere
    field Carrier       : Set
    field op        : Carrier Carrier Carrier
    toType      : let View X = X in View Type ; toType = record {Carrier = Carrier}
    field _≈_       : Carrier CarrierSet
    toMagma     : let View X = X in View Magma ;    toMagma = record {Carrier = Carrier;op = op}
    field cong      : ∀ x xy y′ → _≈_ x x′ → _≈_ y y′ → _≈_ (op x x′) (op y y′)
    toRawRelationalMagma        : let View X = X in View RawRelationalMagma ;   toRawRelationalMagma = record {Carrier = Carrier;op = op;_≈_ = _≈_}

3.5 Renaming

From an end-user perspective, our CommutativeMonoid has one flaw: Such monoids are frequently written additively rather than multiplicatively. Such a change can be rendered conveniently:

{-700
AbealianMonoidR = CommutativeMonoidR renaming "_·_ to _+_"
-}

An Abealian monoid is both a commutative monoid and also, simply, a monoid. The above declaration freely maintains these relationships: The resulting record comes with a new projection toCommutativeMonoidR, and still has the inherited projection toMonoidR.

Since renaming and extension (including postulating) both adjoin retract morphisms, by default, we are lead to wonder how about the result of performing these operations in sequence ‘on the fly’, rather than naming each application. Since P renaming X ⟴ postulating Y comes with a retract toP via the renaming and another, distinctly defined, toP via the postulating, we have that the operations commute if only the first permits the creation of a retract. Here's a concrete example:

{-700
IdempotentMagma  = Magma renaming "_·_ to _⊔_" ⟴ postulating "_⊔_" "idempotent"   :adjoin-retract nil ⟴ record
-}

These both elaborate to the same thing, up to order of constituents.

It is important to realise that the renaming and postulating combinators are user-defined, and could have been defined without adjoining a retract by default; consequently, we would have unconditional commutativity of these combinators. You, as the user, can make these alternative combinators as follows:

{-700

𝒱-renaming′ by = renaming 'by :adjoin-retract nil

𝒱-postulating′ p bop (using) = postulating 'p 'bop :using 'using :adjoin-retract nil

-- Example use: We need the “𝒱-” in the declaration site, but not in use sites, as below.

IdempotentMagma″ = Magma postulating′ "_⊔_" "idempotent" ⟴ renaming′ "_·_ to _⊔_" ⟴ record
-}

Super near stuff!

As expected, simultaneous renaming works too ^_^

{-700
PackageFormer Two : Set₁ where
  Carrier : Set
  𝟘       : Carrier
  𝟙       : Carrier

TwoR = Two record ⟴ renaming′ "𝟘 to 𝟙; 𝟙 to 𝟘"
-}

TwoR is just Two but as an Agda record, so it typechecks.

Finally, renaming is an invertible operation —ignoring the adjoined retracts, Magmaʳʳ is identical to Magma.

{-700
Magmaʳ  = Magma  renaming "_·_  to op"
Magmaʳʳ = Magmaʳ renaming "op   to _·_"
-}

Alternatively, renaming has an optional argument :adjoin-coretract which can be provided with t to use a default name or provided with a string to use a desired name for the inverse part of a projection, fromMagma below.

{-700
Sequential = Magma renaming "op to _⨾_" :adjoin-coretract t
-}
record Sequential : Setwhere
    field Carrier : Set
    field _⨾_     : Carrier Carrier Carrier

    toType : let View X = X in View Type
    toType = record {Carrier = Carrier}

    toMagma : let View X = X in View Magma
    toMagma = record {Carrier = Carrier;op = _⨾_}

    fromMagma : let View X = X in Magma View Sequential
    fromMagma = λ g227742record {Carrier = Magma.Carrier g227742;_⨾_ = Magma.op g227742}

We are using gensym's for λ-arguments to avoid name clashes.

3.6 Union (and intersection)

But even with these features, given GroupR, we would find ourselves writing:

{-700
CommutativeGroupR₀ = GroupR extended-by "comm : {x y : Carrier} →  x · y  ≡  y · x" ⟴ record
-}

This is problematic: We lose the relationship that every commutative group is a commutative monoid. This is not an issue of erroneous hierarchical design: From Monoid, we could orthogonally add a commutativity property or inverse operation; CommutativeGroupR₀ then closes this diamond-loop by adding both features. The simplest way to share structure is to union two presentations:

{-700
CommutativeGroupR = GroupR union CommutativeMonoidR ⟴ record
-}

The resulting record, CommutativeMonoidR, comes with three derived fields ---toMonoidR, toGroupR, toCommutativeMonoidR— that retain the results relationships with its hierarchical construction.

This approach “works” to build a sizeable library, say of the order of 500 concepts, in a fairly economical way [tpc]. The union operation is an instance of a pushout operation, which consists of 5 arguments —three objects and two morphisms— which may be included into the union operation as optional keyword arguments. The more general notion of pushout is required if we were to combine GroupR with AbealianMonoidR, which have non-identical syntactic copies of MonoidR.

Definition of union & intersection -- Ignore for now
(defun find-duplicates (list)
"Return a list that contains each element from LIST that occurs more than once.

Source: https://emacs.stackexchange.com/a/31449/10352"
  (--> list
       (-group-by #'identity it)
       (-filter (lambda (ele) (> (length ele) 2)) it)
       (mapcar #'car it)))
(cl-defmacro alter-elements (elements variational &body rest)
  "Alter ELEMENTS using a given VARIATIONAL along with its arguments, REST.

   The result is a list of elements.

   This is essentially “:alter-elements” but with the ability to work on the elements
   of **any** PackageFormer by using “($𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠-𝑜𝑓 pf)”.

   This method is only well-defined within the RHS of a variational, or instantiation, declaration.
   E.g., use it to alter elements in an “:alter-elements” clause using a predefined variational;
   see 𝒱-union and 𝒱-intersect for sample uses.
  "
  `(funcall (cdr (assoc :alter-elements (,(𝒱- variational) ,@rest))) ,elements))
(𝒱 union pf (renaming₁ "") (renaming₂ "") (adjoin-retract₁ t) (adjoin-retract₂ t)
 = "Union parent PackageFormer with given PF.

    Union the elements of the parent PackageFormer with those of
    the provided PF symbolic name, then adorn the result with two views:
    One to the parent and one to the provided PF.

    If an identifer is shared but has different types, then crash.

    ADJOIN-RETRACTᵢ, for i : 1..2, are the optional names of the resulting morphisms.
    Provide nil if you do not want the morphisms adjoined.
    "
   :alter-elements (λ es →
     (let* ((p (symbol-name 'pf))
            (es₁ (alter-elements es renaming renaming₁ :adjoin-retract nil))
            (es₂ (alter-elements ($𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠-𝑜𝑓 p) renaming renaming₂ :adjoin-retract nil))
            (es′ (-concat es₁ es₂)))

      ;; Ensure no name clashes!
      (loop for n in (find-duplicates (mapcar #'element-name es′))
            for e = (--filter (equal n (element-name it)) es′)
            unless (--all-p (equal (car e) it) e)
            do (-let [debug-on-error nil]
              (error "%s = %s union %s \n\n\t\t ➩ Error: Elements “%s” conflict!\n\n\t\t\t%s"
                     $𝑛𝑎𝑚𝑒 $𝑝𝑎𝑟𝑒𝑛𝑡 p (element-name (car e)) (s-join "\n\t\t\t" (mapcar #'show-element e)))))

   ;; return value
   (-concat
       es′
       (when adjoin-retract₁ (list (element-retract $𝑝𝑎𝑟𝑒𝑛𝑡 es :new es₁ :name adjoin-retract₁)))
       (when adjoin-retract₂ (list (element-retract p     ($𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠-𝑜𝑓 p) :new es₂ :name adjoin-retract₂)))))))
(𝒱 intersect pf (adjoin-retract₁ t) (adjoin-retract₂ t) (renaming₁ "") (renaming₂ "")
 = "Intersect parent PackageFormer with given PF.

  “pf₁ intersect pf₂ :renaming₁ f :renaming₂ g”  ≅  f(pf₁) ∩ g(pf₂)

  This is essentially the pullback:
  { (x, y) ∈ pf₁ × pf₂ ❙ f x = g y }.

  Intersect the elements of the parent PackageFormer with those of
  the provided PF symbolic name, then adorn the result with two views
  from the result PackageFormer to the input PackageFormers.

  ADJOIN-RETRACTᵢ, for i : 1..2, are the optional names of the resulting morphisms.
  Provide nil if you do not want the morphisms adjoined.
  "
 :alter-elements (λ es →
   (let* ((p (symbol-name 'pf))
          (es₁ (alter-elements es renaming renaming₁ :adjoin-retract nil))
          (es₂ (alter-elements ($𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠-𝑜𝑓 p) renaming renaming₂ :adjoin-retract nil))
          (es′      (reverse (intersection
                               (--reject (element-contains "View" it) es₁) es₂
                               :key #'element-name :test #'string-equal)))
          ;; names not mentioned in the intersection.
          (dangling (set-difference (-concat es₁ es₂) es′
                      :key #'element-name :test #'string-equal)))

  ;; drop the dangling names
  (setq es′ (-reject (λ e → (-some (λ d → (element-contains (element-name d) e)) dangling)) es′))

  ;; Get old names so as to adjoin the co-retracts; i.e., the projections with renaming.
  (setq es₁ (alter-elements es′ rename (reify-to-list renaming₁ :inverse t) :adjoin-retract nil))
  (setq es₂ (alter-elements es′ rename (reify-to-list renaming₂ :inverse t) :adjoin-retract nil))

   (-concat
       es′
       (when adjoin-retract₁ (list (element-retract $𝑝𝑎𝑟𝑒𝑛𝑡 es′ :new es₁ :name adjoin-retract₁ :contravariant t)))
       (when adjoin-retract₂ (list (element-retract p     es′ :new es₂ :name adjoin-retract₂ :contravariant t)))))))

The pushout of \(f : X → A\) and \(g : X → B\) is, essentially, the disjoint sum of \(A\) and \(B\) where embedded elements are considered ‘indistinguishable’ when the share the same origin in \(X\) via the paths \(f\) and \(g\). Unfortunately, the resulting ‘indistinguishable’ elements are actually distinguishable: They may be the A-name or the B-name and a choice must be made as to which name is preferred since users actually want to refer to them later on. Hence, to be useful for library construction, the pushout construction actually requires at least another input function that provides canonical names to the supposedly ‘indistinguishable’ elements.

Since a PackageFormer is essentially just a signature —a collection of typed names—, we can make a ‘partial choice of pushout’ to reduce the number of arguments from 6 to 4 by letting the typed-names object \(X\) be ‘inferred’ and encoding the canonical names function into the operations \(f\) and \(g\). The inputs functions \(f, g\) are necessarily signature morphisms —mappings of names that preserve types— and so are simply lists associating names of \(X\) to names of \(A\) and \(B\). If we instead consider \(f′ : X′ ← A\) and \(g′ : X′ ← B\), in the opposite direction, then we may reconstruct a pushout by setting \(X\) to be common image of \(f′, g′\), and set \(f, g\) to be inclusions In-particular, the full identity of \(X′\) is not necessarily relevant for the pushout reconstruction and so it may be omitted. Moreover, the issue of canonical names is resolved: If \(a ∈ A\) is intended to be identified with \(b ∈ B\) such that the resulting element has \(c\) as the chosen canonical name, then we simply require \(f′\, a = c = g′ \, b\).

At first, a pushout construction needs 5 inputs, to be practical it further needs a function for canonical names for a total of 6 inputs. However, a pushout of \(f : X → A\) and \(g : X → B\) is intended to be the ‘smallest object \(P\) that contains a copy of \(A\) and of \(B\) sharing the common substructure \(X\)’, and as such it outputs two functions \(inj₁ : A → P,\, inj₂ : B → P\) that inject the names of \(A\) and \(B\) into \(P\). If we realise \(P\) as a record —a type of models— then the embedding functions are reversed, to obtain projections \(P → A\) and \(P → B\): If we have a model of \(P\), then we can forget some structure and rename via \(f\) and \(g\) to obtain models of \(A\) and \(B\). For the resulting construction to be useful, these names could be automated such as \(toA : P → A\) and \(toB : P → B\) but such a naming scheme does not scale —but we shall use it for default names. As such, we need two more inputs to the pushout construction so the names of the resulting output functions can be used later on. Hence, a practical choice of pushout needs 8 inputs!

Using the above issue to reverse the directions of \(f, g\) via \(f′, g′\), we can infer the shared structure \(X\) and the canonical name function. Likewise, by using \(toChild : P → Child\) default-naming scheme, we may omit the names of the retract functions. If we wish to rename these retracts or simply omit them altogether, we make the optional arguments: Provide :adjoin-retractᵢ "new-function-name" to use a new name, or nil instead of a string to omit the retract. Here are some examples of this construction of mine.

Here we provide all arguments, optional and otherwise.

{-700
TwoBinaryOps = Magma union Magma :renaming₁ "op to _+_" :renaming₂ "op to _×_"  :adjoin-retract₁ "left" :adjoin-retract₂ "right"
-}
record TwoBinaryOps : Setwhere
    field Carrier : Set
    field _+_     : Carrier Carrier Carrier

    toType : let View X = X in View Type
    toType = record {Carrier = Carrier}

    field _×_     : Carrier Carrier Carrier

    left : let View X = X in View Magma
    left = record {Carrier = Carrier;op = _+_}

    right : let View X = X in View Magma
    right = record {Carrier = Carrier;op = _×_}

Remember, this particular user implementation realises X₁ union X₂ :renaming₁ f′ :renaming₂ g′ as the pushout of the inclusions f′ X₁ ∩ g′ X₂ ⟶ Xᵢ where the source is the set-wise intersection of names. Moreover, when either renamingᵢ is omitted, it defaults to the identity function.

The next example is one of the reasons the construction is named ‘union’ instead of ‘pushout’: It's idempotent, if we ignore the addition of the retract.

{-700
MagmaAgain   = Magma union Magma
-}
record MagmaAgain : Setwhere
    field Carrier : Set
    field op      : Carrier Carrier Carrier

    toType : let View X = X in View Type
    toType = record {Carrier = Carrier}

    toMagma : let View X = X in View Magma
    toMagma = record {Carrier = Carrier;op = op}

We may perform disjoint sums —simply distinguish all the names of one of the input objects.

{-700
Magma′    = Magma primed  ⟴ record
SumMagmas = Magma union Magma′ :adjoin-retract₁ nil ⟴ record
-}
record SumMagmas : Setwhere
    field Carrier  : Set
    field op       : Carrier Carrier Carrier

    toType         : let View X = X in View Type
    toType = record {Carrier = Carrier}

    field Carrier′ : Set
    field op′      : Carrier′ → Carrier′ → CarriertoType′ : let View X = X in View Type
    toType= record {Carrier = Carrier′}

    toMagma : let View X = X in View Magma
    toMagma = record {Carrier = Carrier′;op = op′}

    toMagma′ : let View X = X in View MagmatoMagma= record {Carrier= Carrier′;op= op′}

A common scenario is extending a structure, say Magma, into orthogonal directions, such as by making it operation associative or idempotent, then closing the resulting diamond by combining them, to obtain a semilattice. However, the orthogonal extensions may involve different names and so the resulting semilattice presentation can only be formed via pushout; below are three ways to form it.

{-700
Semigroup          = Magma postulating "_·_" "associative" ⟴ record
-- IdempotentMagma = Magma renaming "_·_ to _⊔_" ⟴ postulating "_⊔_" "idempotent"  :adjoin-retract nil ⟴ record

⊔-SemiLattice     = Semigroup union IdempotentMagma :renaming₁ "_·_ to _⊔_" ⟴ record
·-SemiLattice     = Semigroup union IdempotentMagma :renaming₂ "_⊔_ to _·_" ⟴ record
↑-SemiLattice     = Semigroup union IdempotentMagma :renaming₁ "_·_ to _↑_" :renaming₂ "_⊔_ to _↑_" ⟴ record
-}

Let's close with the classic example of forming a ring structure by combining two monoidal structures. This example also serves to further showcasing how using 𝒱-postulating can make for more granular, modular, developments.

{-700
Additive           = Magma renaming "_·_ to _+_" ⟴ postulating "_+_" "commutative"  :adjoin-retract nil ⟴ record
Multiplicative     = Magma renaming "_·_ to _×_" :adjoin-retract nil ⟴ record
AddMult            = Additive union Multiplicative ⟴ record
AlmostNearSemiRing = AddMult ⟴ postulating "_×_" "distributiveˡ" :using "_+_" ⟴ record
-}
record AlmostNearSemiRing : Setwhere
    field Carrier : Set
    field _+_     : Carrier Carrier Carrier

    toType : let View X = X in View Type
    toType = record {Carrier = Carrier}

    toMagma : let View X = X in View Magma
    toMagma = record {Carrier = Carrier;op = _+_}

    field comm       : ∀ x y   → _+_ x y ≡ _+_ y x
    field _×_        : Carrier Carrier Carrier

    toAdditive : let View X = X in View Additive
    toAdditive = record {Carrier = Carrier;_+_ = _+_;comm = comm}

    toMultiplicative : let View X = X in View Multiplicative
    toMultiplicative = record {Carrier = Carrier;_×_ = _×_}

    field distˡ      : ∀ x y z → _×_ x (_+_ y z) ≡ _+_ (_×_ x y) (_×_ x z)

Following the reasoning for pushouts, we implement pullbacks in the same way with the same optional arguments. Here's an example use:

{-700
Just-Carrier    = Additive intersect Multiplicative
Magma-yet-again = Additive intersect Multiplicative :renaming₁ "_+_ to op" :renaming₂ "_×_ to op"
-}

Moreover the absorptive law \(X ∩ (X ∪ Z) = X\) also holds for these operations: Additive intersect AddMult is just Additive, when we ignore all adjoined retracts.

3.7 Duality

Maps between grouping mechanisms are sometimes called views, which are essentially an internalisation of of the variationals in our system. Let's demonstrate an example of how dual concepts are captured concretely in the system.

For example, the dual, or opposite, of a binary operation \(_·_\) is the operation \(_·ᵒᵖ_\) defined by \(x ·ᵒᵖ y ≔ y · x\). Classically in Agda, duality is utilised as follows:

  1. Define a module R _·_ for the desired concepts.
  2. Define a shallow module Rᵒᵖ _·_ that opens R _·ᵒᵖ_ and renames the concepts in R by the dual names.

    The RATH-Agda library performs essentially this approach, for example for obtaining UpperBounds from LowerBounds in the context of a poset.

Unfortunately, this means that any record definitions in R must have its field names be sufficiently generic to play both roles as the original and the dual concept. Admittedly, RATH-Agda's names are well-chosen; e.g., value, boundᵢ, universal to denote a value that is a lower/upper bound of two given elements, satisfying a lub/glb universal property. However, well-chosen names come at an upfront cost: One must take care to provide sufficiently generic names and account for duality at the outset, irrespective of whether one currently cares about the dual or not; otherwise when the dual is later formalised, then the names of the original concept must be refactored throughout a library and its users.

Consider the following heterogeneous algebra.

{-700
PackageFormer LeftUnitalAction : Set₁ where
  Scalar : Set
  Vector : Set
  _·_     : Scalar → Vector → Vector
  𝟙       : Scalar
  leftId  : {x : Vector} → 𝟙 · x ≡ x

-- Let's reify this as a valid Agda record declaration
LeftUnitalActionR  = LeftUnitalAction ⟴ record
-}

Informally, one now ‘defines’ a right unital action by duality, flipping the binary operation and renaming leftId to be rightId. Such informal parlance is in-fact nearly formally, as the following:

{-700
RightUnitalActionR = LeftUnitalActionR flipping "_·_" :renaming "leftId to rightId" ⟴ record
-}

Of-course the resulting representation is semantically identical to the previous one, and so it is furnished with a to⋯ mapping:

forget : RightUnitalActionR LeftUnitalActionR
forget = RightUnitalActionR.toLeftUnitalActionR

Likewise for the RATH-Agda library's example from above, to define semi-lattice structures by duality:

import Data.Product as P

{-700
PackageFormer JoinSemiLattice : Set₁ where
  Carrier : Set
  _⊑_     : Carrier → Carrier → Set
  refl    : ∀ {x} → x ⊑ x
  trans   : ∀ {x y z} → x ⊑ y → y ⊑ z → x ⊑ z
  antisym : ∀ {x y} → x ⊑ y → y ⊑ x → x ≡ y
  _⊔_     : Carrier → Carrier → Carrier
  ⊔-lub   : ∀ {x y z} → x ⊑ z → y ⊑ z → (x ⊔ y) ⊑ z
  ⊔-lub˘  : ∀ {x y z}  → (x ⊔ y) ⊑ z  → x ⊑ z P.× y ⊑ z

JoinSemiLatticeR = JoinSemiLattice record
MeetSemiLatticeR = JoinSemiLatticeR flipping "_⊑_" :renaming "_⊔_ to _⊓_; ⊔-lub to ⊓-glb"
-}

In this example, besides the map from meet semi-lattices to join semi-lattices, the types of the dualised names, such as ⊓-glb, are what one would expect were the definition written out explicitly:

module woah (M : MeetSemiLatticeR) where
  open MeetSemiLatticeR M

  nifty : ∀ {x y z} → zxzyz ⊑ (xy)
  nifty = ⊓-glb

  _ : let _⊒_ = λ x yyx
      in ∀ {x y z} → xyyzxz
  _ = trans
Definition of flipping -- Feel free to ignore until later
(defun flip-type (τ)
  "Given a binary operation's type, as a string, flip the first two types.

E.g., “A → B → C” becomes “B → A → C”.
  "
  (-let [ts (s-split "→" τ)]
   (s-join " → " (list (nth 1 ts) (nth 0 ts) (nth 2 ts)))))
flip-type
(defun flip (name op)
 "If element OP is named NAME, then flip its type; else leave it alone-ish.

If OP mentions NAME, then prefix its type with
“let NAME = λ x y → NAME y x in”, which results in valid Agda
due to its identifier scoping rules.
"
 (cond ((equal name (element-name op))
               (map-type #'flip-type op))
       ((or (element-contains name op) (element-contains (s-replace "_" "" name) op))
               (-let [letin (format "let %s = λ x y → %s y x in " name name)]
                 (thread-last op
                   (map-type (λ τ → (concat letin τ)))
                   (map-equations (λ eqs → (--map (-let [ps (s-split "=" it)] (format "%s = %s %s" (car ps) letin (s-join "=" (cdr ps)))) eqs))))))
       (t op)))
flip
(𝒱 flipping name (renaming "")
 = "Flip a single binary operation, or predicate, NAME.

    Dual constructs usual require some identifiers to be renamed,
    and these may be supplied as a “;”-separated “to”-separated string list, RENAMING.

    There is no support for underscores; mixfix names must be given properly.
  "
    renaming 'renaming :adjoin-retract nil
 ⟴ :alter-elements (λ es →
                      (let ((es′ (--map (flip name it) es)))
                        (-concat es′ (list (flip name (element-retract $𝑝𝑎𝑟𝑒𝑛𝑡 ($𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠-𝑜𝑓 $𝑝𝑎𝑟𝑒𝑛𝑡) :new es′)))))))
𝒱-flipping

3.8 Extracting Little Theories

The extended-by variational allows Agda users to easily employ the tiny theories [little_theories][mathscheme] approach to library design: New structures are built from old ones by augmenting one concept at a time, then one uses mixins such as union, below, to obtain a complex structure. This approach lets us write a program, or proof, in a context that only provides what is necessary for that program-proof and nothing more. In this way, we obtain maximal generality for re-use! This approach can be construed as The Interface Segregation Principle [design-patterns-solid]: /No client should be forced to depend on methods it does not use.

{-700
PackageFormer Empty : Set₁ where {- No elements -}
Type  = Empty extended-by "Carrier : Set"
Magma = Type  extended-by "_·_ : Carrier → Carrier → Carrier"
CommutativeMagma = Magma extended-by "comm : {x y : Carrier} →  x · y  ≡  y · x"
-}

The cool thing here is that CommutativeMagma comes with toMagma, toType, and toEmpty.

However, life is messy and sometimes one may hurriedly create a structure, then later realise that they are being forced to depend on unused methods. Rather than throw an ‘not implemented’ exception or leave them undefined, we may use the keeping variational to extract the smallest well-formed sub-PackageFormer that mentions a given list of identifiers.

For example, suppose we quickly formed Monoid, from earlier, but later wished to utilise other substrata. This is easily achieved with the following declarations.

{-700
Empty″        = Monoid keeping ""
Type″         = Monoid keeping "Carrier"
Magma″        = Monoid keeping "_·_"
Semigroup″    = Monoid keeping "assoc"
PointedMagma″ = Monoid keeping "𝕀; _·_"
-- ↪ Carrier; _·_; 𝕀
-}

Even better, we may go about deriving results —such as theorems or algorithms— in familiar settings, such as Monoid, only to realise that they are more expressive than necessary. Such an observation no longer need to be found by inspection, instead it may be derived mechanically.

{-700
LeftUnitalMagma = Monoid keeping "𝕀-unique" ⟴ record
-}

This expands to the following theory, minimal enough to derive 𝕀-unique.

record LeftUnitalMagma : Set₁ where

   field
     Carrier : Set
     _·_     : Carrier → Carrier → Carrier
     𝕀       : Carrier
     leftId  : {x : Carrier} → 𝕀 · x  ≡ x

   𝕀-unique     : ∀ {e} (lid : ∀ {x} → e · x ≡ x) (rid : ∀ {x} → x · e ≡ x) → e ≡ 𝕀
   𝕀-unique lid rid = ≡.trans (≡.sym leftId) rid

Surprisingly, in some sense, keeping let's us apply the interface segregation principle, or ‘little theories’, after the fact —this is also known as reverse mathematics.

3.9 200+ theories – one line for each

People should enter terse, readable, specifications that expand into useful, typecheckable, code that may be dauntingly larger in textual size.

The following listing of structures was adapted from the source of a MathScheme library tpc,mathscheme, which in turn was inspired by the web lists of Peter Jipsen and John Halleck, and many others from Wikipedia and nlab. Totalling over 200 theories which elaborate into nearly 1500 lines of typechecked Agda, this demonstrates that our systems works; the 750% efficiency savings speak for themselves.

200+ One Line Specfications
{-700
PackageFormer EmptyPF : Set₁ where

Empty   = EmptyPF ⟴ record

-- Don't worry about generating morphisms.

𝒱-extended-by′ ds  =  extended-by ds :adjoin-retract nil ⟴ record
𝒱-union′       pf  =  union pf :adjoin-retract₁ nil :adjoin-retract₂ nil
𝒱-renaming′    pf  =  renaming pf :adjoin-retract nil

-- A simple way to tag names with the operators that belong to.
--
-- This means we may name a property by, say, “commutative”, then later invoke 𝒱-prefixing to obtain, say, “+-commutative”.
-- Hence, names are ‘unqualified’ unless otherwise need to be prefixed due to ambigiuity.
-- EXCEPT is a ;-seperated string.
--
𝒱-prefixing pre (except "") = rename (λ n → (if (or (member n (mapcar #'s-trim (s-split ";" except))) (s-starts-with? pre n)) n (concat pre n))) :adjoin-retract nil
-}

{-700
Carrier                              = Empty extended-by"U : Set"
CarrierS                             = Empty extended-by"S : Set"
MultiCarrier                         = Carrier union CarrierS
PointedCarrier                       = Carrier extended-by"e : U"
Pointed2Carrier                      = Carrier extended-by"e2 : U"
DoublyPointed                        = PointedCarrier union Pointed2Carrier
DoublyPointed𝟘𝟙                      = DoublyPointed renaming"e to 𝟘; e2 to 𝟙"
UnaryOperation                       = Carrier extended-by"prime : U → U"
Magma                                = Carrier extended-by"_*_ : U → U → U"
UnaryRelation                        = Carrier extended-by"R : U → Set"
BinaryRelation                       = Carrier extended-by"R : U → U → Set"
PointedUnarySystem                   = UnaryOperation union PointedCarrier
FixedPoint                           = PointedUnarySystem extended-by"fixes_prime_e : prime e ≡ e"
InvolutiveMagmaSig                   = UnaryOperation union Magma
InvolutivePointedMagmaSig            = PointedUnarySystem union Magma
Involution                           = UnaryOperation extended-by"prime-involutive : ∀ (x : U) → prime (prime x) ≡ x"
UnaryDistributes                     = InvolutiveMagmaSig extended-by"prime-*-distribute : ∀ (x y : U) → prime (x * y) ≡ (prime x * prime y)"
UnaryAntiDistribution                = InvolutiveMagmaSig extended-by"prime-*-antidistribute : ∀ (x y : U) → prime(x * y) ≡ (prime y * prime x)"
AdditiveUnaryAntiDistribution        = UnaryAntiDistribution renaming"_*_ to _+_"
IdempotentUnary                      = UnaryOperation extended-by"prime-idempotent : ∀ (f : U) → prime (prime f) ≡ prime f"
InvolutiveMagma                      = Involution union UnaryAntiDistribution ⟴ :remark "over UnaryOperation"
IrreflexiveRelation                  = BinaryRelation extended-by"R-irreflexive : ∀ (x : U)  →  ¬ (R x x)"
ReflexiveRelation                    = BinaryRelation extended-by"R-reflexive : ∀ (x : U)  →  R x x"
TransitiveRelation                   = BinaryRelation extended-by"R-transitive : ∀ (x y z : U) → R x y → R y z → R x z"
SymmetricRelation                    = BinaryRelation extended-by"R-symmetric : ∀ (x y : U) →  R x y →  R y x"
AntisymmetricRelation                = BinaryRelation extended-by"R-antisymmetric : ∀ (x y : U) → R y x → R x y → x ≡ y"
OrderRelation                        = BinaryRelation renaming"R to _≤_"
ReflexiveOrderRelation               = ReflexiveRelation renaming"R to _≤_"
TransitiveOrderRelation              = TransitiveRelation renaming"R to _≤_"
SymmetricOrderRelation               = SymmetricRelation renaming"R to _≤_"
AntisymmetricOrderRelation           = AntisymmetricRelation renaming"R to _≤_"
Preorder                             = ReflexiveOrderRelation union TransitiveOrderRelation ⟴ :remark "over OrderRelation"
StrictPartialOrder                   = IrreflexiveRelation union TransitiveRelation ⟴ :remark "over BinaryRelation"
EquivalenceRelation                  = ReflexiveRelation union TransitiveRelationunion SymmetricRelation :adjoin-retractnil ⟴ :remark "over BinaryRelation"
PartialOrder                         = Preorder union AntisymmetricOrderRelation  ⟴ :remark "over OrderRelation"
PartiallyOrderedMagmaSig             = Magma union OrderRelation ⟴ :remark "over Carrier"
OrderPreservingMagma                 = PartiallyOrderedMagmaSig extended-by"*-≤-orderPreserving : ∀ (x y u v : U) → x ≤ y → (u * (x * v)) ≤ (u * (y * v))"
PartiallyOrderedMagma                = OrderPreservingMagma union PartialOrder ⟴ :remark "over OrderRelation"
TotalRelation                        = OrderRelation extended-by"≤-total : ∀ (x y : U) → (x ≤ y) ⊎ (y ≤ x)"
TotalPreorder                        = Preorder union TotalRelation ⟴ :remark "over OrderRelation"
TotalOrder                           = TotalPreorder union AntisymmetricOrderRelation ⟴ :remark "over Preorder"
OrderedMagma                         = PartiallyOrderedMagma union TotalRelation ⟴ :remark "over OrderRelation"
LeftCanonicalOrder                   = PartiallyOrderedMagmaSig extended-by"≤-*-leftCanonicalOrder : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a * c)"
RightCanonicalOrder                  = PartiallyOrderedMagmaSig extended-by"≤-*-rightCanonicalOrder : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c * a)"
LeftCanonicallyOrderedMagma          = OrderedMagma union LeftCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig"
RightCanonicallyOrderedMagma         = OrderedMagma union RightCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig"
CanonicallyOrderedMagma              = LeftCanonicallyOrderedMagma union RightCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig"
Chain                                = TotalOrder
AdditiveMagma                        = Magma renaming"_*_ to _+_"
LeftDivisionMagma                    = Magma renaming"_*_ to _╲_"
RightDivisionMagma                   = Magma renaming"_*_ to _╱_"
LeftOperation                        = MultiCarrier extended-by"_⟫_ : U → S → S"
RightOperation                       = MultiCarrier extended-by"_⟪_ : S → U → S"
IdempotentMagma                      = Magma extended-by"*-idempotent : ∀ (x : U) → (x * x) ≡ x"
IdempotentAdditiveMagma              = IdempotentMagma renaming"_*_ to _+_"
SelectiveMagma                       = Magma extended-by"*-selective : ∀ (x y : U) → (x * y ≡ x) ⊎ (x * y ≡ y)"
SelectiveAdditiveMagma               = SelectiveMagma renaming"_*_ to _+_"
PointedMagma                         = Magma union PointedCarrier ⟴ :remark "over Carrier"
Pointed𝟘Magma                        = PointedMagma renaming"e to 𝟘"
AdditivePointed1Magma                = PointedMagma renaming"_*_ to _+_; e to 𝟙"
LeftPointAction                      = PointedMagma extended-by "pointactLeft  :  U → U; pointactLeft x = e * x"
RightPointAction                     = PointedMagma extended-by "pointactRight  :  U → U; pointactRight x = x * e"
CommutativeMagma                     = Magma extended-by"*-commutative  :  ∀ (x y : U) →  (x * y) ≡ (y * x)"
CommutativeAdditiveMagma             = CommutativeMagma renaming"_*_ to _+_"
PointedCommutativeMagma              = PointedMagma union CommutativeMagma ⟴ :remark "over Magma"
AntiAbsorbent                        = Magma extended-by"*-anti-self-absorbent  : ∀ (x y : U) → (x * (x * y)) ≡ y"
SteinerMagma                         = CommutativeMagma union AntiAbsorbent ⟴ :remark "over Magma"
Squag                                = SteinerMagma union IdempotentMagma ⟴ :remark "over Magma"
PointedSteinerMagma                  = PointedMagma union SteinerMagma ⟴ :remark "over Magma"
UnipotentPointedMagma                = PointedMagma extended-by"unipotent  : ∀ (x : U) →  (x * x) ≡ e"
Sloop                                = PointedSteinerMagma union UnipotentPointedMagma ⟴ :remark "over PointedMagma"
LeftDistributiveMagma                = Magma extended-by"*-leftDistributive : ∀ (x y z : U) → (x * (y * z)) ≡ ((x * y) * (x * z))"
RightDistributiveMagma               = Magma extended-by"*-rightDistributive : ∀ (x y z : U) → ((y * z) * x) ≡ ((y * x) * (z * x))"
LeftCancellativeMagma                = Magma extended-by"*-leftCancellative  :  ∀ (x y z : U) → z * x ≡ z * y → x ≡ y"
RightCancellativeMagma               = Magma extended-by"*-rightCancellative  : ∀ (x y z : U) →  x * z ≡ y * z → x ≡ y"
CancellativeMagma                    = LeftCancellativeMagma union RightCancellativeMagma ⟴ :remark "over Magma"
LeftUnital                           = PointedMagma extended-by"*-leftIdentity : ∀ (x : U) →  e * x ≡ x"
RightUnital                          = PointedMagma extended-by"*-rightIdentity : ∀ (x : U) →  x * e ≡ x"
Unital                               = LeftUnital union RightUnital ⟴ :remark "over PointedMagma"
LeftBiMagma                          = LeftDivisionMagma union Magma :adjoin-retract"toMagma*" ⟴ :remark "over Carrier"
RightBiMagma                         = LeftBiMagma renaming"_╲_ to _╱_"
LeftCancellative                     = LeftBiMagma extended-by"*-╲-leftCancel : ∀ (x y : U) →  x * (x ╲ y) ≡ y"
LeftCancellativeOp                   = LeftBiMagma extended-by"╲-╲-leftCancel : ∀ (x y : U) →  x ╲ (x * y) ≡ y"
LeftQuasiGroup                       = LeftCancellative union LeftCancellativeOp ⟴ :remark "over LeftBiMagma"
RightCancellative                    = RightBiMagma extended-by"╱-*-rightCancel : ∀ (x y : U) →  (y ╱ x) * x ≡ y"
RightCancellativeOp                  = RightBiMagma extended-by"*-╱-rightCancel : ∀ (x y : U) → (y * x) ╱ x ≡ y"
RightQuasiGroup                      = RightCancellative union RightCancellativeOp ⟴ :remark "over RightBiMagma"
QuasiGroup                           = LeftQuasiGroup union RightQuasiGroup
MedialMagma                          = Magma extended-by"mediate : ∀ (w x y z : U) →  (x * y) * (z * w) ≡ (x * z) * (y * w)"
MedialQuasiGroup                     = QuasiGroup union MedialMagma ⟴ :remark "over Magma"
MoufangLaw                           = Magma extended-by"*-moufangLaw : ∀ (e x y z : U)  →  (y * e ≡ y)  →  ((x * y) * z) * x  ≡  x * (y * ((e * z) * x))"
MoufangQuasigroup                    = QuasiGroup union MoufangLaw ⟴ :remark "over Magma"
LeftLoop                             = RightUnital union LeftQuasiGroup ⟴ :remark "over Magma"
Loop                                 = Unital union QuasiGroup ⟴ :remark "over Magma"
-- Should conservatively extend Loop with leftInv and rightInv.
MoufangAlgebra                       = Magma extended-by"*-MoufangIdentity : ∀ (x y z : U) → (z * x) * (y * z) ≡ (z * (x * y)) * z"
MoufangLoop                          = Loop union MoufangAlgebra ⟴ :remark "over Magma"
LeftShelfSig                         = Magma renaming"_*_ to _|>_"
LeftShelf                            = LeftDistributiveMagma renaming"_*_ to _|>_"
RightShelfSig                        = Magma renaming"_*_ to _<|_ "
RightShelf                           = RightDistributiveMagma renaming"_*_ to _<|_ "
RackSig                              = LeftShelfSig union RightShelfSig ⟴ :remark "over Carrier"
Shelf                                = LeftShelf union RightShelf ⟴ :remark "over RackSig"
LeftBinaryInverse                    = RackSig extended-by"|>-<|-absorption : ∀ (x y : U) →  (x |> y) <| x ≡ y"
RightBinaryInverse                   = RackSig extended-by"|>-<|-co-absorption : ∀ (x y : U) →  x |> (y <| x) ≡ y"
Rack                                 = RightShelf union LeftShelfunion LeftBinaryInverse union′ ⟴ union RightBinaryInverse ⟴ :remark "over RackSig"
LeftIdempotence                      = IdempotentMagma renaming"_*_ to _|>_"
RightIdempotence                     = IdempotentMagma renaming"_*_ to _<|_"
LeftSpindle                          = LeftShelf union LeftIdempotence ⟴ :remark "over LeftShelfSig"
RightSpindle                         = RightShelf union RightIdempotence ⟴ :remark "over RightShelfSig"
Quandle                              = Rack union LeftSpindleunion RightSpindle ⟴ :adjoin-retractnil ⟴ :remark "over Shelf"
RightSelfInverse                     = LeftShelfSig extended-by"rightSelfInverse_|> : ∀ (x y : U) →  (x |> y) |> y ≡ x"
Kei                                  = LeftSpindle union RightSelfInverse ⟴ :remark "over LeftShelfSig"
-- A “Shellis a BiUnital LeftBiMagma where the additive 𝟘 is a zero for _*_.
Semigroup                            = Magma extended-by"*-associative : ∀ (x y z : U) →  (x * y) * z  ≡  x * (y * z)"
AdditiveSemigroup                    = Semigroup renaming"_*_ to _+_"
CommutativeSemigroup                 = Semigroup union CommutativeMagma ⟴ :remark "over Magma"
AdditiveCommutativeSemigroup         = CommutativeSemigroup renaming"_*_ to _+_"
LeftCancellativeSemigroup            = Semigroup union LeftCancellativeMagma ⟴ :remark "over Magma"
RightCancellativeSemigroup           = Semigroup union RightCancellativeMagma ⟴ :remark "over Magma"
CancellativeSemigroup                = Semigroup union CancellativeMagma ⟴ :remark "over Magma"
CancellativeCommutativeSemigroup     = CommutativeSemigroup union CancellativeSemigroup ⟴ :remark "over Semigroup"
InvolutiveSemigroup                  = Semigroup union InvolutiveMagma ⟴ :remark "over PointedMagma"
PartiallyOrderedSemigroup            = PartiallyOrderedMagma union Semigroup ⟴ :remark "over PartiallyOrderedMagmaSig"
OrderedSemigroup                     = PartiallyOrderedSemigroup union TotalRelation ⟴ :remark "over OrderRelation"
CommutativePartiallyOrderedSemigroup = CommutativeSemigroup union PartiallyOrderedSemigroup ⟴ :remark "over Semigroup"
CommutativeOrderedSemigroup          = CommutativeSemigroup union OrderedSemigroup ⟴ :remark "over Semigroup"
Band                                 = Semigroup union IdempotentMagma ⟴ :remark "over Magma"
CommutativeBand                      = Band union CommutativeMagma ⟴ :remark "over Magma"
MiddleAbsorption                     = Magma extended-by"*-middleAbsorb : ∀ (x y z : U) →  x * (y * z)  ≡  x * z"
MiddleCommute                        = Magma extended-by"*-middleCommute : ∀ (x y z : U) → (x * y) * (z * x)  ≡  (x * z) * (y * x)"
RectangularBand                      = Band union MiddleAbsorption ⟴ :remark "over Magma"
NormalBand                           = Band union MiddleCommute ⟴ :remark "over Magma"
RightMonoid                          = RightUnital union Semigroup ⟴ :remark "over Magma"
LeftMonoid                           = LeftUnital union Semigroup ⟴ :remark "over Magma"
Monoid                               = Unital union Semigroup ⟴ :remark "over Magma"
AdditiveMonoid                       = Monoid renaming"_*_ to _+_; e to 𝟘"
DoubleMonoid                         = Monoid union AdditiveMonoid ⟴ :remark "over Carrier"
Monoid1                              = Monoid renaming"e to 𝟙"
CommutativeMonoid                    = Monoid union CommutativeSemigroup ⟴ :remark "over Semigroup"
SelectiveMonoid                      = Monoid union SelectiveMagma ⟴ :remark "over Magma"
CancellativeMonoid                   = Monoid union CancellativeMagma ⟴ :remark "over Magma"
CancellativeCommutativeMonoid        = CancellativeMonoid union CommutativeMonoid ⟴ :remark "over Monoid"
LeftZero                             = PointedMagma extended-by"*-leftZero : ∀ (x : U) → e * x ≡ e"
RightZero                            = PointedMagma extended-by"*-rightZero : ∀ (x : U) → x * e ≡ e"
Zero                                 = LeftZero union RightZero ⟴ :remark "over PointedMagma"
Left0                                = LeftZero renaming"e to 𝟘"
Right0                               = RightZero renaming"e to 𝟘"
ComplementSig                        = UnaryOperation renaming"prime to compl"
CommutativeMonoid1                   = CommutativeMonoid renaming"e to 𝟙"
AdditiveCommutativeMonoid            = CommutativeMonoid renaming"_*_ to _+_; e to 𝟘"
---------------------- MA: At this point, there are no tooltips. ----------------------
PartiallyOrderedMonoid               = PartiallyOrderedMagma union Monoid ⟴ :remark "over PartiallyOrderedMagmaSig"
OrderedMonoid                        = PartiallyOrderedMonoid union TotalRelation ⟴ :remark "over OrderRelation"
CommutativePartiallyOrderedMonoid    = CommutativeMonoid union PartiallyOrderedMonoid ⟴ :remark "over Monoid"
CommutativeOrderedMonoid             = CommutativeMonoid union OrderedMonoid ⟴ :remark "over Monoid"
-- For these canonically preordered structures, transitivity and reflexivity follow from the definition ofby *, but they are included for clarity.
LeftCanonicallyPreorderedMonoid      = Monoid union LeftCanonicalOrderunion Preorder ⟴ :adjoin-retractnil ⟴ :remark "over PartiallyOrderedMagmaSig"
RightCanonicallyPreorderedMonoid     = Monoid union RightCanonicalOrderunion Preorder ⟴ :adjoin-retractnil ⟴ :remark "over PartiallyOrderedMagmaSig"
CanonicallyPreorderedMonoid          = LeftCanonicallyPreorderedMonoid union RightCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig"
LeftCanonicallyOrderedMonoid         = PartiallyOrderedMonoid union LeftCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig"
RightCanonicallyOrderedMonoid        = PartiallyOrderedMonoid union LeftCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig"
CanonicallyOrderedMonoid             = LeftCanonicallyOrderedMonoid union RightCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig"
AdditiveCanonicallyOrderedMonoid     = CanonicallyOrderedMonoid renaming"_*_ to _+_; e to 𝟘"
HemiGroup                            = CanonicallyOrderedMonoid union CancellativeMagma ⟴ :remark "over Magma"
AdditiveHemiGroup                    = HemiGroup renaming"_*_ to _+_; e to 𝟘"
BooleanGroup                         = Monoid union UnipotentPointedMagma ⟴ :remark "over PointedMagma"
InverseSig                           = InvolutivePointedMagmaSig renaming"prime to inv"
LeftInverse                          = InverseSig extended-by"*-leftInverse : ∀ (x : U) →  x * (inv x) ≡ e"
RightInverse                         = InverseSig extended-by"*-rightInverse : ∀ (x : U) → (inv x) * x ≡ e"
Inverse                              = LeftInverse union RightInverse ⟴ :remark "over InverseSig"
PseudoInverseSig                     = InvolutiveMagmaSig renaming"prime to inv"
PseudoInverse                        = PseudoInverseSig extended-by"*-quasiLeftInverse : ∀ (x : U) →  x * ((inv x) * x) ≡  x"
PseudoInvolution                     = PseudoInverseSig extended-by"*-quasiRightInverse : ∀ (x : U) → (inv x * x) * inv x ≡ inv x"
RegularSemigroup                     = Semigroup union PseudoInverse ⟴ :remark "over Magma"
InverseSemigroup                     = PseudoInverse union PseudoInvolution ⟴ :remark "over PseudoInverseSig"
Group                                = Monoid union Inverse ⟴ :remark "over InverseSig"
Group1                               = Group renaming"e to 𝟙"
AdditiveGroup                        = Group renaming"_*_ to _+_; e to 𝟘; inv to neg"
AbelianGroup                         = Group1 union CommutativeMonoid1 ⟴ :remark "over Monoid1"
AbelianAdditiveGroup                 = AdditiveGroup union CommutativeAdditiveMagma ⟴ :remark "over AdditiveMagma"
PartiallyOrderedGroup                = PartiallyOrderedMagma union Group ⟴ :remark "over PartiallyOrderedMagmaSig"
OrderedGroup                         = PartiallyOrderedGroup union TotalRelation ⟴ :remark "over OrderRelation"
AbelianPartiallyOrderedGroup         = PartiallyOrderedMagma union AbelianGroup ⟴ :remark "over PartiallyOrderedMagmaSig"
AbelianOrderedGroup                  = AbelianPartiallyOrderedGroup union TotalRelation ⟴ :remark "over OrderRelation"
RingoidSig                           = Magma union AdditiveMagma :adjoin-retract"*-isMagma" :adjoin-retract"+-isMagma" ⟴ :remark "over Carrier"
Pointed0Sig                          = PointedCarrier renaming"e to 𝟘"
Pointed1Sig                          = PointedCarrier renaming"e to 𝟙"
Ringoid0Sig                          = RingoidSig union Pointed0Sig ⟴ :remark "over Carrier"
Ringoid1Sig                          = RingoidSig union Pointed1Sig ⟴ :remark "over Carrier"
Ringoid01Sig                         = Ringoid0Sig union Ringoid1Sig ⟴ :remark "over RingoidSig"
LeftRingoid                          = RingoidSig extended-by"*-+-leftDistributivity : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)"
RightRingoid                         = RingoidSig extended-by"*-+-rightDistributivity : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)"
Ringoid                              = LeftRingoid  union RightRingoid ⟴ :remark "over RingoidSig"
NonassociativeNondistributiveRing    = AbelianAdditiveGroup union Magma ⟴ :remark "over Carrier"
NonassociativeRing                   =  NonassociativeNondistributiveRing union Ringoid ⟴ :remark "over RingoidSig"
PrimeRingoidSig                      = RingoidSig  union UnaryOperation ⟴ :remark "over Carrier"
AnddeMorgan                          = PrimeRingoidSig extended-by"prime-*-+-deMorgan : ∀ (x y z : U) → prime (x * y) ≡ (prime x) + (prime y)"
OrdeMorgan                           = PrimeRingoidSig extended-by"prime-+-*-deMorgan : ∀ (x y z : U) → prime (x + y) ≡ (prime x) * (prime y)"
DualdeMorgan                         = OrdeMorgan union AnddeMorgan ⟴ :remark "over PrimeRingoidSig"
LeftPreSemiring                      = LeftRingoid  union Semigroupunion AdditiveCommutativeSemigroup  ⟴ :adjoin-retractnil ⟴ :remark "over RingoidSig"
RightPreSemiring                     = RightRingoid union Semigroupunion AdditiveCommutativeSemigroup ⟴ :adjoin-retractnil ⟴ :remark "over RingoidSig"
PreSemiring                          = LeftPreSemiring union RightRingoid ⟴ :remark "over RingoidSig"
NearSemiring                         = AdditiveSemigroup  union Semigroupunion RightRingoid ⟴ :adjoin-retractnil ⟴ :remark "over RingoidSig"
NearSemifield = NearSemiring union Group ⟴ :remark "over Semigroup"
Semifield = NearSemifield union LeftRingoid ⟴ :remark "over RingoidSig"
NearRing = AdditiveGroup union Semigroupunion RightRingoid ⟴ :remark "over RingoidSig"
Rng = AbelianAdditiveGroup union Semigroup  Ringoid ⟴ :remark "over RingoidSig"
Semiring = AdditiveCommutativeMonoid union Monoid1union Ringoidunion Left0union Right0 ⟴ :remark "over Ringoid0Sig"

SemiRng       = AdditiveCommutativeMonoid union Semigroupunion Ringoid ⟴ :remark "over RingoidSig"
LeftPreDioid  = LeftPreSemiring union AdditiveCanonicallyOrderedMonoid ⟴ :remark "over AdditiveMonoid"
RightPreDioid = RightPreSemiring union AdditiveCanonicallyOrderedMonoid ⟴ :remark "over AdditiveMonoid"
PreDioid      = LeftPreDioid union RightRingoid ⟴ :remark "over RingoidSig"
LeftDioid     = LeftPreDioid union Monoidunion Left0union Right0 ⟴ :remark "over DoubleMonoid"
RightDioid    = RightPreDioid union Monoidunion Left0union Right0 ⟴ :remark "over DoubleMonoid"
Dioid         = LeftDioid union RightRingoid ⟴ :remark "over RingoidSig"
~1500 Lines of Typechecked Agda
{- This file is generated ;; do not alter. -}

import Relation.Binary.PropositionalEquality as ≡; open ≡ using (_≡_)
open import Relation.Nullary using (¬_)
open import Data.Sum using (_⊎_)
open import Data.Product
open import Relation.Binary.PropositionalEquality using () renaming (_≡_ to _↔_) -- just to get by.
open import Level as Level using (Level)
module math-scheme-generated where 


{- KindPackageFormerdoes not correspond  to a concrete Agda type.

PackageFormer EmptyPF : Setwhere
 -}


{- Empty   = EmptyPF ⟴ record -}
record Empty : Set₁ where


{- Carrier                              = Empty extended-by"U : Set" -}
record Carrier : Setwhere
    field U     : Set


{- CarrierS                             = Empty extended-by′ "S : Set" -}
record CarrierS : Set₁ where
    field S     : Set


{- MultiCarrier                         = Carrier union CarrierS -}
record MultiCarrier : Setwhere
    field U     : Set
    field S     : Set


{- PointedCarrier                       = Carrier extended-by′ "e : U" -}
record PointedCarrier : Set₁ where
    field U     : Set
    field e     : U


{- Pointed2Carrier                      = Carrier extended-by"e2 : U" -}
record Pointed2Carrier : Setwhere
    field U     : Set
    field e2        : U


{- DoublyPointed                        = PointedCarrier union′ Pointed2Carrier -}
record DoublyPointed : Set₁ where
    field U     : Set
    field e     : U
    field e2        : U


{- DoublyPointed𝟘𝟙                      = DoublyPointed renaming"e to 𝟘; e2 to 𝟙" -}
record DoublyPointed𝟘𝟙 : Setwhere
    field U     : Set
    field 𝟘     : U
    field 𝟙     : U


{- UnaryOperation                       = Carrier extended-by′ "prime : U → U" -}
record UnaryOperation : Set₁ where
    field U     : Set
    field prime     : U → U


{- Magma                                = Carrier extended-by"_*_ : U → U → U" -}
record Magma : Setwhere
    field U     : Set
    field _*_       : U U U


{- UnaryRelation                        = Carrier extended-by′ "R : U → Set" -}
record UnaryRelation : Set₁ where
    field U     : Set
    field R     : U → Set


{- BinaryRelation                       = Carrier extended-by"R : U → U → Set" -}
record BinaryRelation : Setwhere
    field U     : Set
    field R     : U USet


{- PointedUnarySystem                   = UnaryOperation union′ PointedCarrier -}
record PointedUnarySystem : Set₁ where
    field U     : Set
    field prime     : U → U
    field e     : U


{- FixedPoint                           = PointedUnarySystem extended-by"fixes_prime_e : prime e ≡ e" -}
record FixedPoint : Setwhere
    field U     : Set
    field prime     : U U
    field e     : U
    field fixes_prime_e     : prime ee


{- InvolutiveMagmaSig                   = UnaryOperation union′ Magma -}
record InvolutiveMagmaSig : Set₁ where
    field U     : Set
    field prime     : U → U
    field _*_       : U → U → U


{- InvolutivePointedMagmaSig            = PointedUnarySystem union Magma -}
record InvolutivePointedMagmaSig : Setwhere
    field U     : Set
    field prime     : U U
    field e     : U
    field _*_       : U U U


{- Involution                           = UnaryOperation extended-by′ "prime-involutive : ∀ (x : U) → prime (prime x) ≡ x" -}
record Involution : Set₁ where
    field U     : Set
    field prime     : U → U
    field prime-involutive      : ∀ (x : U) → prime (prime x) ≡ x


{- UnaryDistributes                     = InvolutiveMagmaSig extended-by"prime-*-distribute : ∀ (x y : U) → prime (x * y) ≡ (prime x * prime y)" -}
record UnaryDistributes : Setwhere
    field U     : Set
    field prime     : U U
    field _*_       : U U U
    field prime-*-distribute        : ∀ (x y : U) → prime (x * y) ≡ (prime x * prime y)


{- UnaryAntiDistribution                = InvolutiveMagmaSig extended-by′ "prime-*-antidistribute : ∀ (x y : U) → prime(x * y) ≡ (prime y * prime x)" -}
record UnaryAntiDistribution : Set₁ where
    field U     : Set
    field prime     : U → U
    field _*_       : U → U → U
    field prime-*-antidistribute        : ∀ (x y : U) → prime(x * y) ≡ (prime y * prime x)


{- AdditiveUnaryAntiDistribution        = UnaryAntiDistribution renaming"_*_ to _+_" -}
record AdditiveUnaryAntiDistribution : Setwhere
    field U     : Set
    field prime     : U U
    field _+_       : U U U
    field prime-+-antidistribute        : ∀ (x y : U) → prime(x + y) ≡ (prime y + prime x)


{- IdempotentUnary                      = UnaryOperation extended-by′ "prime-idempotent : ∀ (f : U) → prime (prime f) ≡ prime f" -}
record IdempotentUnary : Set₁ where
    field U     : Set
    field prime     : U → U
    field prime-idempotent      : ∀ (f : U) → prime (prime f) ≡ prime f


{- InvolutiveMagma                      = Involution union UnaryAntiDistribution ⟴ :remark "over UnaryOperation" -}
record InvolutiveMagma : Setwhere
    field U     : Set
    field prime     : U U
    field prime-involutive      : ∀ (x : U) → prime (prime x) ≡ x
    field _*_       : U U U
    field prime-*-antidistribute        : ∀ (x y : U) → prime(x * y) ≡ (prime y * prime x)


{- IrreflexiveRelation                  = BinaryRelation extended-by′ "R-irreflexive : ∀ (x : U)  →  ¬ (R x x)" -}
record IrreflexiveRelation : Set₁ where
    field U     : Set
    field R     : U → U → Set
    field R-irreflexive     : ∀ (x : U) → ¬ (R x x)


{- ReflexiveRelation                    = BinaryRelation extended-by"R-reflexive : ∀ (x : U)  →  R x x" -}
record ReflexiveRelation : Setwhere
    field U     : Set
    field R     : U USet
    field R-reflexive       : ∀ (x : U) → R x x


{- TransitiveRelation                   = BinaryRelation extended-by′ "R-transitive : ∀ (x y z : U) → R x y → R y z → R x z" -}
record TransitiveRelation : Set₁ where
    field U     : Set
    field R     : U → U → Set
    field R-transitive      : ∀ (x y z : U) → R x y → R y z → R x z


{- SymmetricRelation                    = BinaryRelation extended-by"R-symmetric : ∀ (x y : U) →  R x y →  R y x" -}
record SymmetricRelation : Setwhere
    field U     : Set
    field R     : U USet
    field R-symmetric       : ∀ (x y : U) → R x y R y x


{- AntisymmetricRelation                = BinaryRelation extended-by′ "R-antisymmetric : ∀ (x y : U) → R y x → R x y → x ≡ y" -}
record AntisymmetricRelation : Set₁ where
    field U     : Set
    field R     : U → U → Set
    field R-antisymmetric       : ∀ (x y : U) → R y x → R x y → x ≡ y


{- OrderRelation                        = BinaryRelation renaming"R to _≤_" -}
record OrderRelation : Setwhere
    field U     : Set
    field _≤_       : U USet


{- ReflexiveOrderRelation               = ReflexiveRelation renaming′ "R to _≤_" -}
record ReflexiveOrderRelation : Set₁ where
    field U     : Set
    field _≤_       : U → U → Set
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x


{- TransitiveOrderRelation              = TransitiveRelation renaming"R to _≤_" -}
record TransitiveOrderRelation : Setwhere
    field U     : Set
    field _≤_       : U USet
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z


{- SymmetricOrderRelation               = SymmetricRelation renaming′ "R to _≤_" -}
record SymmetricOrderRelation : Set₁ where
    field U     : Set
    field _≤_       : U → U → Set
    field ≤-symmetric       : ∀ (x y : U) → _≤_ x y → _≤_ y x


{- AntisymmetricOrderRelation           = AntisymmetricRelation renaming"R to _≤_" -}
record AntisymmetricOrderRelation : Setwhere
    field U     : Set
    field _≤_       : U USet
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy


{- Preorder                             = ReflexiveOrderRelation union′ TransitiveOrderRelation ⟴ :remark "over OrderRelation" -}
record Preorder : Set₁ where
    field U     : Set
    field _≤_       : U → U → Set
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z


{- StrictPartialOrder                   = IrreflexiveRelation union TransitiveRelation ⟴ :remark "over BinaryRelation" -}
record StrictPartialOrder : Setwhere
    field U     : Set
    field R     : U USet
    field R-irreflexive     : ∀ (x : U) → ¬ (R x x)
    field R-transitive      : ∀ (x y z : U) → R x y R y z R x z


{- EquivalenceRelation                  = ReflexiveRelation union′ TransitiveRelation ⟴ union′ SymmetricRelation :adjoin-retract₁ nil ⟴ :remark "over BinaryRelation" -}
record EquivalenceRelation : Set₁ where
    field U     : Set
    field R     : U → U → Set
    field R-reflexive       : ∀ (x : U) → R x x
    field R-transitive      : ∀ (x y z : U) → R x y → R y z → R x z
    field R-symmetric       : ∀ (x y : U) → R x y → R y x


{- PartialOrder                         = Preorder union AntisymmetricOrderRelation  ⟴ :remark "over OrderRelation" -}
record PartialOrder : Setwhere
    field U     : Set
    field _≤_       : U USet
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy


{- PartiallyOrderedMagmaSig             = Magma union′ OrderRelation ⟴ :remark "over Carrier" -}
record PartiallyOrderedMagmaSig : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set


{- OrderPreservingMagma                 = PartiallyOrderedMagmaSig extended-by"*-≤-orderPreserving : ∀ (x y u v : U) → x ≤ y → (u * (x * v)) ≤ (u * (y * v))" -}
record OrderPreservingMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y u v : U) → xy → (u * (x * v)) ≤ (u * (y * v))


{- PartiallyOrderedMagma                = OrderPreservingMagma union′ PartialOrder ⟴ :remark "over OrderRelation" -}
record PartiallyOrderedMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y


{- TotalRelation                        = OrderRelation extended-by"≤-total : ∀ (x y : U) → (x ≤ y) ⊎ (y ≤ x)" -}
record TotalRelation : Setwhere
    field U     : Set
    field _≤_       : U USet
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)


{- TotalPreorder                        = Preorder union′ TotalRelation ⟴ :remark "over OrderRelation" -}
record TotalPreorder : Set₁ where
    field U     : Set
    field _≤_       : U → U → Set
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-total       : ∀ (x y : U) → (x ≤ y) ⊎ (y ≤ x)


{- TotalOrder                           = TotalPreorder union AntisymmetricOrderRelation ⟴ :remark "over Preorder" -}
record TotalOrder : Setwhere
    field U     : Set
    field _≤_       : U USet
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy


{- OrderedMagma                         = PartiallyOrderedMagma union′ TotalRelation ⟴ :remark "over OrderRelation" -}
record OrderedMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field ≤-total       : ∀ (x y : U) → (x ≤ y) ⊎ (y ≤ x)


{- LeftCanonicalOrder                   = PartiallyOrderedMagmaSig extended-by"≤-*-leftCanonicalOrder : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a * c)" -}
record LeftCanonicalOrder : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field ≤-*-leftCanonicalOrder        : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] ba * c)


{- RightCanonicalOrder                  = PartiallyOrderedMagmaSig extended-by′ "≤-*-rightCanonicalOrder : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c * a)" -}
record RightCanonicalOrder : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set
    field ≤-*-rightCanonicalOrder       : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c * a)


{- LeftCanonicallyOrderedMagma          = OrderedMagma union LeftCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record LeftCanonicallyOrderedMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)
    field ≤-*-leftCanonicalOrder        : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] ba * c)


{- RightCanonicallyOrderedMagma         = OrderedMagma union′ RightCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record RightCanonicallyOrderedMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field ≤-total       : ∀ (x y : U) → (x ≤ y) ⊎ (y ≤ x)
    field ≤-*-rightCanonicalOrder       : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c * a)


{- CanonicallyOrderedMagma              = LeftCanonicallyOrderedMagma union RightCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record CanonicallyOrderedMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)
    field ≤-*-leftCanonicalOrder        : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] ba * c)
    field ≤-*-rightCanonicalOrder       : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] bc * a)


{- Chain                                = TotalOrder -}
record Chain : Set₁ where
    field U     : Set
    field _≤_       : U → U → Set
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-total       : ∀ (x y : U) → (x ≤ y) ⊎ (y ≤ x)
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y


{- AdditiveMagma                        = Magma renaming"_*_ to _+_" -}
record AdditiveMagma : Setwhere
    field U     : Set
    field _+_       : U U U


{- LeftDivisionMagma                    = Magma renaming′ "_*_ to _╲_" -}
record LeftDivisionMagma : Set₁ where
    field U     : Set
    field _╲_       : U → U → U


{- RightDivisionMagma                   = Magma renaming"_*_ to _╱_" -}
record RightDivisionMagma : Setwhere
    field U     : Set
    field _╱_       : U U U


{- LeftOperation                        = MultiCarrier extended-by′ "_⟫_ : U → S → S" -}
record LeftOperation : Set₁ where
    field U     : Set
    field S     : Set
    field _⟫_       : U → S → S


{- RightOperation                       = MultiCarrier extended-by"_⟪_ : S → U → S" -}
record RightOperation : Setwhere
    field U     : Set
    field S     : Set
    field _⟪_       : S U S


{- IdempotentMagma                      = Magma extended-by′ "*-idempotent : ∀ (x : U) → (x * x) ≡ x" -}
record IdempotentMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-idempotent      : ∀ (x : U) → (x * x) ≡ x


{- IdempotentAdditiveMagma              = IdempotentMagma renaming"_*_ to _+_" -}
record IdempotentAdditiveMagma : Setwhere
    field U     : Set
    field _+_       : U U U
    field +-idempotent      : ∀ (x : U) → (x + x) ≡ x


{- SelectiveMagma                       = Magma extended-by′ "*-selective : ∀ (x y : U) → (x * y ≡ x) ⊎ (x * y ≡ y)" -}
record SelectiveMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-selective       : ∀ (x y : U) → (x * y ≡ x) ⊎ (x * y ≡ y)


{- SelectiveAdditiveMagma               = SelectiveMagma renaming"_*_ to _+_" -}
record SelectiveAdditiveMagma : Setwhere
    field U     : Set
    field _+_       : U U U
    field +-selective       : ∀ (x y : U) → (x + yx) ⊎ (x + yy)


{- PointedMagma                         = Magma union′ PointedCarrier ⟴ :remark "over Carrier" -}
record PointedMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U


{- Pointed𝟘Magma                        = PointedMagma renaming"e to 𝟘" -}
record Pointed𝟘Magma : Setwhere
    field U     : Set
    field _*_       : U U U
    field 𝟘     : U


{- AdditivePointed1Magma                = PointedMagma renaming′ "_*_ to _+_; e to 𝟙" -}
record AdditivePointed1Magma : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field 𝟙     : U


{- LeftPointAction                      = PointedMagma extended-by "pointactLeft  :  U → U; pointactLeft x = e * x" -}
record LeftPointAction : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    pointactLeft        : U U ;   pointactLeft x = e * x
    toPointedMagma      : let View X = X in View PointedMagma ; toPointedMagma = record {U = U;_*_ = _*_;e = e}


{- RightPointAction                     = PointedMagma extended-by "pointactRight  :  U → U; pointactRight x = x * e" -}
record RightPointAction : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    pointactRight       : U → U ;   pointactRight x = x * e
    toPointedMagma      : let View X = X in View PointedMagma ; toPointedMagma = record {U = U;_*_ = _*_;e = e}


{- CommutativeMagma                     = Magma extended-by"*-commutative  :  ∀ (x y : U) →  (x * y) ≡ (y * x)" -}
record CommutativeMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)


{- CommutativeAdditiveMagma             = CommutativeMagma renaming′ "_*_ to _+_" -}
record CommutativeAdditiveMagma : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)


{- PointedCommutativeMagma              = PointedMagma union CommutativeMagma ⟴ :remark "over Magma" -}
record PointedCommutativeMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)


{- AntiAbsorbent                        = Magma extended-by′ "*-anti-self-absorbent  : ∀ (x y : U) → (x * (x * y)) ≡ y" -}
record AntiAbsorbent : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-anti-self-absorbent     : ∀ (x y : U) → (x * (x * y)) ≡ y


{- SteinerMagma                         = CommutativeMagma union AntiAbsorbent ⟴ :remark "over Magma" -}
record SteinerMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field *-anti-self-absorbent     : ∀ (x y : U) → (x * (x * y)) ≡ y


{- Squag                                = SteinerMagma union′ IdempotentMagma ⟴ :remark "over Magma" -}
record Squag : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field *-anti-self-absorbent     : ∀ (x y : U) → (x * (x * y)) ≡ y
    field *-idempotent      : ∀ (x : U) → (x * x) ≡ x


{- PointedSteinerMagma                  = PointedMagma union SteinerMagma ⟴ :remark "over Magma" -}
record PointedSteinerMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field *-anti-self-absorbent     : ∀ (x y : U) → (x * (x * y)) ≡ y


{- UnipotentPointedMagma                = PointedMagma extended-by′ "unipotent  : ∀ (x : U) →  (x * x) ≡ e" -}
record UnipotentPointedMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field unipotent     : ∀ (x : U) → (x * x) ≡ e


{- Sloop                                = PointedSteinerMagma union UnipotentPointedMagma ⟴ :remark "over PointedMagma" -}
record Sloop : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field *-anti-self-absorbent     : ∀ (x y : U) → (x * (x * y)) ≡ y
    field unipotent     : ∀ (x : U) → (x * x) ≡ e


{- LeftDistributiveMagma                = Magma extended-by′ "*-leftDistributive : ∀ (x y z : U) → (x * (y * z)) ≡ ((x * y) * (x * z))" -}
record LeftDistributiveMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-leftDistributive        : ∀ (x y z : U) → (x * (y * z)) ≡ ((x * y) * (x * z))


{- RightDistributiveMagma               = Magma extended-by"*-rightDistributive : ∀ (x y z : U) → ((y * z) * x) ≡ ((y * x) * (z * x))" -}
record RightDistributiveMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-rightDistributive       : ∀ (x y z : U) → ((y * z) * x) ≡ ((y * x) * (z * x))


{- LeftCancellativeMagma                = Magma extended-by′ "*-leftCancellative  :  ∀ (x y z : U) → z * x ≡ z * y → x ≡ y" -}
record LeftCancellativeMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-leftCancellative        : ∀ (x y z : U) → z * x ≡ z * y → x ≡ y


{- RightCancellativeMagma               = Magma extended-by"*-rightCancellative  : ∀ (x y z : U) →  x * z ≡ y * z → x ≡ y" -}
record RightCancellativeMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-rightCancellative       : ∀ (x y z : U) → x * zy * zxy


{- CancellativeMagma                    = LeftCancellativeMagma union′ RightCancellativeMagma ⟴ :remark "over Magma" -}
record CancellativeMagma : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-leftCancellative        : ∀ (x y z : U) → z * x ≡ z * y → x ≡ y
    field *-rightCancellative       : ∀ (x y z : U) → x * z ≡ y * z → x ≡ y


{- LeftUnital                           = PointedMagma extended-by"*-leftIdentity : ∀ (x : U) →  e * x ≡ x" -}
record LeftUnital : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx


{- RightUnital                          = PointedMagma extended-by′ "*-rightIdentity : ∀ (x : U) →  x * e ≡ x" -}
record RightUnital : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x


{- Unital                               = LeftUnital union RightUnital ⟴ :remark "over PointedMagma" -}
record Unital : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex


{- LeftBiMagma                          = LeftDivisionMagma union′ Magma :adjoin-retract₂ "toMagma*" ⟴ :remark "over Carrier" -}
record LeftBiMagma : Set₁ where
    field U     : Set
    field _╲_       : U → U → U
    field _*_       : U → U → U


{- RightBiMagma                         = LeftBiMagma renaming"_╲_ to _╱_" -}
record RightBiMagma : Setwhere
    field U     : Set
    field _╱_       : U U U
    field _*_       : U U U


{- LeftCancellative                     = LeftBiMagma extended-by′ "*-╲-leftCancel : ∀ (x y : U) →  x * (x ╲ y) ≡ y" -}
record LeftCancellative : Set₁ where
    field U     : Set
    field _╲_       : U → U → U
    field _*_       : U → U → U
    field *-╲-leftCancel        : ∀ (x y : U) → x * (x ╲ y) ≡ y


{- LeftCancellativeOp                   = LeftBiMagma extended-by"╲-╲-leftCancel : ∀ (x y : U) →  x ╲ (x * y) ≡ y" -}
record LeftCancellativeOp : Setwhere
    field U     : Set
    field _╲_       : U U U
    field _*_       : U U U
    field ╲-╲-leftCancel        : ∀ (x y : U) → x ╲ (x * y) ≡ y


{- LeftQuasiGroup                       = LeftCancellative union′ LeftCancellativeOp ⟴ :remark "over LeftBiMagma" -}
record LeftQuasiGroup : Set₁ where
    field U     : Set
    field _╲_       : U → U → U
    field _*_       : U → U → U
    field *-╲-leftCancel        : ∀ (x y : U) → x * (x ╲ y) ≡ y
    field ╲-╲-leftCancel        : ∀ (x y : U) → x ╲ (x * y) ≡ y


{- RightCancellative                    = RightBiMagma extended-by"╱-*-rightCancel : ∀ (x y : U) →  (y ╱ x) * x ≡ y" -}
record RightCancellative : Setwhere
    field U     : Set
    field _╱_       : U U U
    field _*_       : U U U
    field ╱-*-rightCancel       : ∀ (x y : U) → (yx) * xy


{- RightCancellativeOp                  = RightBiMagma extended-by′ "*-╱-rightCancel : ∀ (x y : U) → (y * x) ╱ x ≡ y" -}
record RightCancellativeOp : Set₁ where
    field U     : Set
    field _╱_       : U → U → U
    field _*_       : U → U → U
    field *-╱-rightCancel       : ∀ (x y : U) → (y * x) ╱ x ≡ y


{- RightQuasiGroup                      = RightCancellative union RightCancellativeOp ⟴ :remark "over RightBiMagma" -}
record RightQuasiGroup : Setwhere
    field U     : Set
    field _╱_       : U U U
    field _*_       : U U U
    field ╱-*-rightCancel       : ∀ (x y : U) → (yx) * xy
    field *-╱-rightCancel       : ∀ (x y : U) → (y * x) ╱ xy


{- QuasiGroup                           = LeftQuasiGroup union′ RightQuasiGroup -}
record QuasiGroup : Set₁ where
    field U     : Set
    field _╲_       : U → U → U
    field _*_       : U → U → U
    field *-╲-leftCancel        : ∀ (x y : U) → x * (x ╲ y) ≡ y
    field ╲-╲-leftCancel        : ∀ (x y : U) → x ╲ (x * y) ≡ y
    field _╱_       : U → U → U
    field ╱-*-rightCancel       : ∀ (x y : U) → (y ╱ x) * x ≡ y
    field *-╱-rightCancel       : ∀ (x y : U) → (y * x) ╱ x ≡ y


{- MedialMagma                          = Magma extended-by"mediate : ∀ (w x y z : U) →  (x * y) * (z * w) ≡ (x * z) * (y * w)" -}
record MedialMagma : Setwhere
    field U     : Set
    field _*_       : U U U
    field mediate       : ∀ (w x y z : U) → (x * y) * (z * w) ≡ (x * z) * (y * w)


{- MedialQuasiGroup                     = QuasiGroup union′ MedialMagma ⟴ :remark "over Magma" -}
record MedialQuasiGroup : Set₁ where
    field U     : Set
    field _╲_       : U → U → U
    field _*_       : U → U → U
    field *-╲-leftCancel        : ∀ (x y : U) → x * (x ╲ y) ≡ y
    field ╲-╲-leftCancel        : ∀ (x y : U) → x ╲ (x * y) ≡ y
    field _╱_       : U → U → U
    field ╱-*-rightCancel       : ∀ (x y : U) → (y ╱ x) * x ≡ y
    field *-╱-rightCancel       : ∀ (x y : U) → (y * x) ╱ x ≡ y
    field mediate       : ∀ (w x y z : U) → (x * y) * (z * w) ≡ (x * z) * (y * w)


{- MoufangLaw                           = Magma extended-by"*-moufangLaw : ∀ (e x y z : U)  →  (y * e ≡ y)  →  ((x * y) * z) * x  ≡  x * (y * ((e * z) * x))" -}
record MoufangLaw : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-moufangLaw      : ∀ (e x y z : U) → (y * ey) → ((x * y) * z) * xx * (y * ((e * z) * x))


{- MoufangQuasigroup                    = QuasiGroup union′ MoufangLaw ⟴ :remark "over Magma" -}
record MoufangQuasigroup : Set₁ where
    field U     : Set
    field _╲_       : U → U → U
    field _*_       : U → U → U
    field *-╲-leftCancel        : ∀ (x y : U) → x * (x ╲ y) ≡ y
    field ╲-╲-leftCancel        : ∀ (x y : U) → x ╲ (x * y) ≡ y
    field _╱_       : U → U → U
    field ╱-*-rightCancel       : ∀ (x y : U) → (y ╱ x) * x ≡ y
    field *-╱-rightCancel       : ∀ (x y : U) → (y * x) ╱ x ≡ y
    field *-moufangLaw      : ∀ (e x y z : U) → (y * e ≡ y) → ((x * y) * z) * x ≡ x * (y * ((e * z) * x))


{- LeftLoop                             = RightUnital union LeftQuasiGroup ⟴ :remark "over Magma" -}
record LeftLoop : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field _╲_       : U U U
    field *-╲-leftCancel        : ∀ (x y : U) → x * (xy) ≡ y
    field ╲-╲-leftCancel        : ∀ (x y : U) → x ╲ (x * y) ≡ y


{- Loop                                 = Unital union′ QuasiGroup ⟴ :remark "over Magma" -}
record Loop : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field _╲_       : U → U → U
    field *-╲-leftCancel        : ∀ (x y : U) → x * (x ╲ y) ≡ y
    field ╲-╲-leftCancel        : ∀ (x y : U) → x ╲ (x * y) ≡ y
    field _╱_       : U → U → U
    field ╱-*-rightCancel       : ∀ (x y : U) → (y ╱ x) * x ≡ y
    field *-╱-rightCancel       : ∀ (x y : U) → (y * x) ╱ x ≡ y


{- MoufangAlgebra                       = Magma extended-by"*-MoufangIdentity : ∀ (x y z : U) → (z * x) * (y * z) ≡ (z * (x * y)) * z" -}
record MoufangAlgebra : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-MoufangIdentity     : ∀ (x y z : U) → (z * x) * (y * z) ≡ (z * (x * y)) * z


{- MoufangLoop                          = Loop union′ MoufangAlgebra ⟴ :remark "over Magma" -}
record MoufangLoop : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field _╲_       : U → U → U
    field *-╲-leftCancel        : ∀ (x y : U) → x * (x ╲ y) ≡ y
    field ╲-╲-leftCancel        : ∀ (x y : U) → x ╲ (x * y) ≡ y
    field _╱_       : U → U → U
    field ╱-*-rightCancel       : ∀ (x y : U) → (y ╱ x) * x ≡ y
    field *-╱-rightCancel       : ∀ (x y : U) → (y * x) ╱ x ≡ y
    field *-MoufangIdentity     : ∀ (x y z : U) → (z * x) * (y * z) ≡ (z * (x * y)) * z


{- LeftShelfSig                         = Magma renaming"_*_ to _|>_" -}
record LeftShelfSig : Setwhere
    field U     : Set
    field _|>_      : U U U


{- LeftShelf                            = LeftDistributiveMagma renaming′ "_*_ to _|>_" -}
record LeftShelf : Set₁ where
    field U     : Set
    field _|>_      : U → U → U
    field |>-leftDistributive       : ∀ (x y z : U) → (x |> (y |> z)) ≡ ((x |> y) |> (x |> z))


{- RightShelfSig                        = Magma renaming"_*_ to _<|_ " -}
record RightShelfSig : Setwhere
    field U     : Set
    field _<|_      : U U U


{- RightShelf                           = RightDistributiveMagma renaming′ "_*_ to _<|_ " -}
record RightShelf : Set₁ where
    field U     : Set
    field _<|_      : U → U → U
    field <|-rightDistributive      : ∀ (x y z : U) → ((y <| z) <| x) ≡ ((y <| x) <| (z <| x))


{- RackSig                              = LeftShelfSig union RightShelfSig ⟴ :remark "over Carrier" -}
record RackSig : Setwhere
    field U     : Set
    field _|>_      : U U U
    field _<|_      : U U U


{- Shelf                                = LeftShelf union′ RightShelf ⟴ :remark "over RackSig" -}
record Shelf : Set₁ where
    field U     : Set
    field _|>_      : U → U → U
    field |>-leftDistributive       : ∀ (x y z : U) → (x |> (y |> z)) ≡ ((x |> y) |> (x |> z))
    field _<|_      : U → U → U
    field <|-rightDistributive      : ∀ (x y z : U) → ((y <| z) <| x) ≡ ((y <| x) <| (z <| x))


{- LeftBinaryInverse                    = RackSig extended-by"|>-<|-absorption : ∀ (x y : U) →  (x |> y) <| x ≡ y" -}
record LeftBinaryInverse : Setwhere
    field U     : Set
    field _|>_      : U U U
    field _<|_      : U U U
    field |>-<|-absorption      : ∀ (x y : U) → (x |> y) <| xy


{- RightBinaryInverse                   = RackSig extended-by′ "|>-<|-co-absorption : ∀ (x y : U) →  x |> (y <| x) ≡ y" -}
record RightBinaryInverse : Set₁ where
    field U     : Set
    field _|>_      : U → U → U
    field _<|_      : U → U → U
    field |>-<|-co-absorption       : ∀ (x y : U) → x |> (y <| x) ≡ y


{- Rack                                 = RightShelf union LeftShelfunion LeftBinaryInverse union′ ⟴ union RightBinaryInverse ⟴ :remark "over RackSig" -}
record Rack : Setwhere
    field U     : Set
    field _<|_      : U U U
    field <|-rightDistributive      : ∀ (x y z : U) → ((y <| z) <| x) ≡ ((y <| x) <| (z <| x))
    field _|>_      : U U U
    field |>-leftDistributive       : ∀ (x y z : U) → (x |> (y |> z)) ≡ ((x |> y) |> (x |> z))
    field |>-<|-absorption      : ∀ (x y : U) → (x |> y) <| xy
    field |>-<|-co-absorption       : ∀ (x y : U) → x |> (y <| x) ≡ y


{- LeftIdempotence                      = IdempotentMagma renaming′ "_*_ to _|>_" -}
record LeftIdempotence : Set₁ where
    field U     : Set
    field _|>_      : U → U → U
    field |>-idempotent     : ∀ (x : U) → (x |> x) ≡ x


{- RightIdempotence                     = IdempotentMagma renaming"_*_ to _<|_" -}
record RightIdempotence : Setwhere
    field U     : Set
    field _<|_      : U U U
    field <|-idempotent     : ∀ (x : U) → (x <| x) ≡ x


{- LeftSpindle                          = LeftShelf union′ LeftIdempotence ⟴ :remark "over LeftShelfSig" -}
record LeftSpindle : Set₁ where
    field U     : Set
    field _|>_      : U → U → U
    field |>-leftDistributive       : ∀ (x y z : U) → (x |> (y |> z)) ≡ ((x |> y) |> (x |> z))
    field |>-idempotent     : ∀ (x : U) → (x |> x) ≡ x


{- RightSpindle                         = RightShelf union RightIdempotence ⟴ :remark "over RightShelfSig" -}
record RightSpindle : Setwhere
    field U     : Set
    field _<|_      : U U U
    field <|-rightDistributive      : ∀ (x y z : U) → ((y <| z) <| x) ≡ ((y <| x) <| (z <| x))
    field <|-idempotent     : ∀ (x : U) → (x <| x) ≡ x


{- Quandle                              = Rack union′ LeftSpindle ⟴ union′ RightSpindle ⟴ :adjoin-retract₁ nil ⟴ :remark "over Shelf" -}
record Quandle : Set₁ where
    field U     : Set
    field _<|_      : U → U → U
    field <|-rightDistributive      : ∀ (x y z : U) → ((y <| z) <| x) ≡ ((y <| x) <| (z <| x))
    field _|>_      : U → U → U
    field |>-leftDistributive       : ∀ (x y z : U) → (x |> (y |> z)) ≡ ((x |> y) |> (x |> z))
    field |>-<|-absorption      : ∀ (x y : U) → (x |> y) <| x ≡ y
    field |>-<|-co-absorption       : ∀ (x y : U) → x |> (y <| x) ≡ y
    field |>-idempotent     : ∀ (x : U) → (x |> x) ≡ x
    field <|-idempotent     : ∀ (x : U) → (x <| x) ≡ x


{- RightSelfInverse                     = LeftShelfSig extended-by"rightSelfInverse_|> : ∀ (x y : U) →  (x |> y) |> y ≡ x" -}
record RightSelfInverse : Setwhere
    field U     : Set
    field _|>_      : U U U
    field rightSelfInverse_|>       : ∀ (x y : U) → (x |> y) |> yx


{- Kei                                  = LeftSpindle union′ RightSelfInverse ⟴ :remark "over LeftShelfSig" -}
record Kei : Set₁ where
    field U     : Set
    field _|>_      : U → U → U
    field |>-leftDistributive       : ∀ (x y z : U) → (x |> (y |> z)) ≡ ((x |> y) |> (x |> z))
    field |>-idempotent     : ∀ (x : U) → (x |> x) ≡ x
    field rightSelfInverse_|>       : ∀ (x y : U) → (x |> y) |> y ≡ x


{- Semigroup                            = Magma extended-by"*-associative : ∀ (x y z : U) →  (x * y) * z  ≡  x * (y * z)" -}
record Semigroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)


{- AdditiveSemigroup                    = Semigroup renaming′ "_*_ to _+_" -}
record AdditiveSemigroup : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)


{- CommutativeSemigroup                 = Semigroup union CommutativeMagma ⟴ :remark "over Magma" -}
record CommutativeSemigroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)


{- AdditiveCommutativeSemigroup         = CommutativeSemigroup renaming′ "_*_ to _+_" -}
record AdditiveCommutativeSemigroup : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)


{- LeftCancellativeSemigroup            = Semigroup union LeftCancellativeMagma ⟴ :remark "over Magma" -}
record LeftCancellativeSemigroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-leftCancellative        : ∀ (x y z : U) → z * xz * yxy


{- RightCancellativeSemigroup           = Semigroup union′ RightCancellativeMagma ⟴ :remark "over Magma" -}
record RightCancellativeSemigroup : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-rightCancellative       : ∀ (x y z : U) → x * z ≡ y * z → x ≡ y


{- CancellativeSemigroup                = Semigroup union CancellativeMagma ⟴ :remark "over Magma" -}
record CancellativeSemigroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-leftCancellative        : ∀ (x y z : U) → z * xz * yxy
    field *-rightCancellative       : ∀ (x y z : U) → x * zy * zxy


{- CancellativeCommutativeSemigroup     = CommutativeSemigroup union′ CancellativeSemigroup ⟴ :remark "over Semigroup" -}
record CancellativeCommutativeSemigroup : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field *-leftCancellative        : ∀ (x y z : U) → z * x ≡ z * y → x ≡ y
    field *-rightCancellative       : ∀ (x y z : U) → x * z ≡ y * z → x ≡ y


{- InvolutiveSemigroup                  = Semigroup union InvolutiveMagma ⟴ :remark "over PointedMagma" -}
record InvolutiveSemigroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field prime     : U U
    field prime-involutive      : ∀ (x : U) → prime (prime x) ≡ x
    field prime-*-antidistribute        : ∀ (x y : U) → prime(x * y) ≡ (prime y * prime x)


{- PartiallyOrderedSemigroup            = PartiallyOrderedMagma union′ Semigroup ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record PartiallyOrderedSemigroup : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)


{- OrderedSemigroup                     = PartiallyOrderedSemigroup union TotalRelation ⟴ :remark "over OrderRelation" -}
record OrderedSemigroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)


{- CommutativePartiallyOrderedSemigroup = CommutativeSemigroup union′ PartiallyOrderedSemigroup ⟴ :remark "over Semigroup" -}
record CommutativePartiallyOrderedSemigroup : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y


{- CommutativeOrderedSemigroup          = CommutativeSemigroup union OrderedSemigroup ⟴ :remark "over Semigroup" -}
record CommutativeOrderedSemigroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)


{- Band                                 = Semigroup union′ IdempotentMagma ⟴ :remark "over Magma" -}
record Band : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-idempotent      : ∀ (x : U) → (x * x) ≡ x


{- CommutativeBand                      = Band union CommutativeMagma ⟴ :remark "over Magma" -}
record CommutativeBand : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-idempotent      : ∀ (x : U) → (x * x) ≡ x
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)


{- MiddleAbsorption                     = Magma extended-by′ "*-middleAbsorb : ∀ (x y z : U) →  x * (y * z)  ≡  x * z" -}
record MiddleAbsorption : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-middleAbsorb        : ∀ (x y z : U) → x * (y * z) ≡ x * z


{- MiddleCommute                        = Magma extended-by"*-middleCommute : ∀ (x y z : U) → (x * y) * (z * x)  ≡  (x * z) * (y * x)" -}
record MiddleCommute : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-middleCommute       : ∀ (x y z : U) → (x * y) * (z * x) ≡ (x * z) * (y * x)


{- RectangularBand                      = Band union′ MiddleAbsorption ⟴ :remark "over Magma" -}
record RectangularBand : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-idempotent      : ∀ (x : U) → (x * x) ≡ x
    field *-middleAbsorb        : ∀ (x y z : U) → x * (y * z) ≡ x * z


{- NormalBand                           = Band union MiddleCommute ⟴ :remark "over Magma" -}
record NormalBand : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-idempotent      : ∀ (x : U) → (x * x) ≡ x
    field *-middleCommute       : ∀ (x y z : U) → (x * y) * (z * x) ≡ (x * z) * (y * x)


{- RightMonoid                          = RightUnital union′ Semigroup ⟴ :remark "over Magma" -}
record RightMonoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)


{- LeftMonoid                           = LeftUnital union Semigroup ⟴ :remark "over Magma" -}
record LeftMonoid : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)


{- Monoid                               = Unital union′ Semigroup ⟴ :remark "over Magma" -}
record Monoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)


{- AdditiveMonoid                       = Monoid renaming"_*_ to _+_; e to 𝟘" -}
record AdditiveMonoid : Setwhere
    field U     : Set
    field _+_       : U U U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)


{- DoubleMonoid                         = Monoid union′ AdditiveMonoid ⟴ :remark "over Carrier" -}
record DoubleMonoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field _+_       : U → U → U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + x ≡ x
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)


{- Monoid1                              = Monoid renaming"e to 𝟙" -}
record Monoid1 : Setwhere
    field U     : Set
    field _*_       : U U U
    field 𝟙     : U
    field *-leftIdentity        : ∀ (x : U) → 𝟙 * xx
    field *-rightIdentity       : ∀ (x : U) → x * 𝟙 ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)


{- CommutativeMonoid                    = Monoid union′ CommutativeSemigroup ⟴ :remark "over Semigroup" -}
record CommutativeMonoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)


{- SelectiveMonoid                      = Monoid union SelectiveMagma ⟴ :remark "over Magma" -}
record SelectiveMonoid : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-selective       : ∀ (x y : U) → (x * yx) ⊎ (x * yy)


{- CancellativeMonoid                   = Monoid union′ CancellativeMagma ⟴ :remark "over Magma" -}
record CancellativeMonoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-leftCancellative        : ∀ (x y z : U) → z * x ≡ z * y → x ≡ y
    field *-rightCancellative       : ∀ (x y z : U) → x * z ≡ y * z → x ≡ y


{- CancellativeCommutativeMonoid        = CancellativeMonoid union CommutativeMonoid ⟴ :remark "over Monoid" -}
record CancellativeCommutativeMonoid : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-leftCancellative        : ∀ (x y z : U) → z * xz * yxy
    field *-rightCancellative       : ∀ (x y z : U) → x * zy * zxy
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)


{- LeftZero                             = PointedMagma extended-by′ "*-leftZero : ∀ (x : U) → e * x ≡ e" -}
record LeftZero : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftZero        : ∀ (x : U) → e * x ≡ e


{- RightZero                            = PointedMagma extended-by"*-rightZero : ∀ (x : U) → x * e ≡ e" -}
record RightZero : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-rightZero       : ∀ (x : U) → x * ee


{- Zero                                 = LeftZero union′ RightZero ⟴ :remark "over PointedMagma" -}
record Zero : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftZero        : ∀ (x : U) → e * x ≡ e
    field *-rightZero       : ∀ (x : U) → x * e ≡ e


{- Left0                                = LeftZero renaming"e to 𝟘" -}
record Left0 : Setwhere
    field U     : Set
    field _*_       : U U U
    field 𝟘     : U
    field *-leftZero        : ∀ (x : U) → 𝟘 * x ≡ 𝟘


{- Right0                               = RightZero renaming′ "e to 𝟘" -}
record Right0 : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field 𝟘     : U
    field *-rightZero       : ∀ (x : U) → x * 𝟘 ≡ 𝟘


{- ComplementSig                        = UnaryOperation renaming"prime to compl" -}
record ComplementSig : Setwhere
    field U     : Set
    field compl     : U U


{- CommutativeMonoid1                   = CommutativeMonoid renaming′ "e to 𝟙" -}
record CommutativeMonoid1 : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field 𝟙     : U
    field *-leftIdentity        : ∀ (x : U) → 𝟙 * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * 𝟙 ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)


{- AdditiveCommutativeMonoid            = CommutativeMonoid renaming"_*_ to _+_; e to 𝟘" -}
record AdditiveCommutativeMonoid : Setwhere
    field U     : Set
    field _+_       : U U U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)


{- PartiallyOrderedMonoid               = PartiallyOrderedMagma union′ Monoid ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record PartiallyOrderedMonoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)


{- OrderedMonoid                        = PartiallyOrderedMonoid union TotalRelation ⟴ :remark "over OrderRelation" -}
record OrderedMonoid : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)


{- CommutativePartiallyOrderedMonoid    = CommutativeMonoid union′ PartiallyOrderedMonoid ⟴ :remark "over Monoid" -}
record CommutativePartiallyOrderedMonoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y


{- CommutativeOrderedMonoid             = CommutativeMonoid union OrderedMonoid ⟴ :remark "over Monoid" -}
record CommutativeOrderedMonoid : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)


{- LeftCanonicallyPreorderedMonoid      = Monoid union′ LeftCanonicalOrder ⟴ union′ Preorder ⟴ :adjoin-retract₁ nil ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record LeftCanonicallyPreorderedMonoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field _≤_       : U → U → Set
    field ≤-*-leftCanonicalOrder        : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a * c)
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z


{- RightCanonicallyPreorderedMonoid     = Monoid union RightCanonicalOrderunion Preorder ⟴ :adjoin-retractnil ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record RightCanonicallyPreorderedMonoid : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field _≤_       : U USet
    field ≤-*-rightCanonicalOrder       : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] bc * a)
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z


{- CanonicallyPreorderedMonoid          = LeftCanonicallyPreorderedMonoid union′ RightCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record CanonicallyPreorderedMonoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field _≤_       : U → U → Set
    field ≤-*-leftCanonicalOrder        : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a * c)
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-*-rightCanonicalOrder       : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c * a)


{- LeftCanonicallyOrderedMonoid         = PartiallyOrderedMonoid union LeftCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record LeftCanonicallyOrderedMonoid : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field ≤-*-leftCanonicalOrder        : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] ba * c)


{- RightCanonicallyOrderedMonoid        = PartiallyOrderedMonoid union′ LeftCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record RightCanonicallyOrderedMonoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field ≤-*-leftCanonicalOrder        : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a * c)


{- CanonicallyOrderedMonoid             = LeftCanonicallyOrderedMonoid union RightCanonicalOrder ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record CanonicallyOrderedMonoid : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field ≤-*-leftCanonicalOrder        : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] ba * c)
    field ≤-*-rightCanonicalOrder       : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] bc * a)


{- AdditiveCanonicallyOrderedMonoid     = CanonicallyOrderedMonoid renaming′ "_*_ to _+_; e to 𝟘" -}
record AdditiveCanonicallyOrderedMonoid : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field _≤_       : U → U → Set
    field +-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U + (x + v)) ≤ (U + (y + v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + x ≡ x
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field ≤-+-leftCanonicalOrder        : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a + c)
    field ≤-+-rightCanonicalOrder       : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c + a)


{- HemiGroup                            = CanonicallyOrderedMonoid union CancellativeMagma ⟴ :remark "over Magma" -}
record HemiGroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field ≤-*-leftCanonicalOrder        : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] ba * c)
    field ≤-*-rightCanonicalOrder       : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] bc * a)
    field *-leftCancellative        : ∀ (x y z : U) → z * xz * yxy
    field *-rightCancellative       : ∀ (x y z : U) → x * zy * zxy


{- AdditiveHemiGroup                    = HemiGroup renaming′ "_*_ to _+_; e to 𝟘" -}
record AdditiveHemiGroup : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field _≤_       : U → U → Set
    field +-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U + (x + v)) ≤ (U + (y + v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + x ≡ x
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field ≤-+-leftCanonicalOrder        : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a + c)
    field ≤-+-rightCanonicalOrder       : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c + a)
    field +-leftCancellative        : ∀ (x y z : U) → z + x ≡ z + y → x ≡ y
    field +-rightCancellative       : ∀ (x y z : U) → x + z ≡ y + z → x ≡ y


{- BooleanGroup                         = Monoid union UnipotentPointedMagma ⟴ :remark "over PointedMagma" -}
record BooleanGroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field unipotent     : ∀ (x : U) → (x * x) ≡ e


{- InverseSig                           = InvolutivePointedMagmaSig renaming′ "prime to inv" -}
record InverseSig : Set₁ where
    field U     : Set
    field inv       : U → U
    field e     : U
    field _*_       : U → U → U


{- LeftInverse                          = InverseSig extended-by"*-leftInverse : ∀ (x : U) →  x * (inv x) ≡ e" -}
record LeftInverse : Setwhere
    field U     : Set
    field inv       : U U
    field e     : U
    field _*_       : U U U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ e


{- RightInverse                         = InverseSig extended-by′ "*-rightInverse : ∀ (x : U) → (inv x) * x ≡ e" -}
record RightInverse : Set₁ where
    field U     : Set
    field inv       : U → U
    field e     : U
    field _*_       : U → U → U
    field *-rightInverse        : ∀ (x : U) → (inv x) * x ≡ e


{- Inverse                              = LeftInverse union RightInverse ⟴ :remark "over InverseSig" -}
record Inverse : Setwhere
    field U     : Set
    field inv       : U U
    field e     : U
    field _*_       : U U U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ e
    field *-rightInverse        : ∀ (x : U) → (inv x) * xe


{- PseudoInverseSig                     = InvolutiveMagmaSig renaming′ "prime to inv" -}
record PseudoInverseSig : Set₁ where
    field U     : Set
    field inv       : U → U
    field _*_       : U → U → U


{- PseudoInverse                        = PseudoInverseSig extended-by"*-quasiLeftInverse : ∀ (x : U) →  x * ((inv x) * x) ≡  x" -}
record PseudoInverse : Setwhere
    field U     : Set
    field inv       : U U
    field _*_       : U U U
    field *-quasiLeftInverse        : ∀ (x : U) → x * ((inv x) * x) ≡ x


{- PseudoInvolution                     = PseudoInverseSig extended-by′ "*-quasiRightInverse : ∀ (x : U) → (inv x * x) * inv x ≡ inv x" -}
record PseudoInvolution : Set₁ where
    field U     : Set
    field inv       : U → U
    field _*_       : U → U → U
    field *-quasiRightInverse       : ∀ (x : U) → (inv x * x) * inv x ≡ inv x


{- RegularSemigroup                     = Semigroup union PseudoInverse ⟴ :remark "over Magma" -}
record RegularSemigroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field inv       : U U
    field *-quasiLeftInverse        : ∀ (x : U) → x * ((inv x) * x) ≡ x


{- InverseSemigroup                     = PseudoInverse union′ PseudoInvolution ⟴ :remark "over PseudoInverseSig" -}
record InverseSemigroup : Set₁ where
    field U     : Set
    field inv       : U → U
    field _*_       : U → U → U
    field *-quasiLeftInverse        : ∀ (x : U) → x * ((inv x) * x) ≡ x
    field *-quasiRightInverse       : ∀ (x : U) → (inv x * x) * inv x ≡ inv x


{- Group                                = Monoid union Inverse ⟴ :remark "over InverseSig" -}
record Group : Setwhere
    field U     : Set
    field _*_       : U U U
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field inv       : U U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ e
    field *-rightInverse        : ∀ (x : U) → (inv x) * xe


{- Group1                               = Group renaming′ "e to 𝟙" -}
record Group1 : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field 𝟙     : U
    field *-leftIdentity        : ∀ (x : U) → 𝟙 * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * 𝟙 ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field inv       : U → U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ 𝟙
    field *-rightInverse        : ∀ (x : U) → (inv x) * x ≡ 𝟙


{- AdditiveGroup                        = Group renaming"_*_ to _+_; e to 𝟘; inv to neg" -}
record AdditiveGroup : Setwhere
    field U     : Set
    field _+_       : U U U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field neg       : U U
    field +-leftInverse     : ∀ (x : U) → x + (neg x) ≡ 𝟘
    field +-rightInverse        : ∀ (x : U) → (neg x) + x ≡ 𝟘


{- AbelianGroup                         = Group1 union′ CommutativeMonoid1 ⟴ :remark "over Monoid1" -}
record AbelianGroup : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field 𝟙     : U
    field *-leftIdentity        : ∀ (x : U) → 𝟙 * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * 𝟙 ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field inv       : U → U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ 𝟙
    field *-rightInverse        : ∀ (x : U) → (inv x) * x ≡ 𝟙
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)


{- AbelianAdditiveGroup                 = AdditiveGroup union CommutativeAdditiveMagma ⟴ :remark "over AdditiveMagma" -}
record AbelianAdditiveGroup : Setwhere
    field U     : Set
    field _+_       : U U U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field neg       : U U
    field +-leftInverse     : ∀ (x : U) → x + (neg x) ≡ 𝟘
    field +-rightInverse        : ∀ (x : U) → (neg x) + x ≡ 𝟘
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)


{- PartiallyOrderedGroup                = PartiallyOrderedMagma union′ Group ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record PartiallyOrderedGroup : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field inv       : U → U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ e
    field *-rightInverse        : ∀ (x : U) → (inv x) * x ≡ e


{- OrderedGroup                         = PartiallyOrderedGroup union TotalRelation ⟴ :remark "over OrderRelation" -}
record OrderedGroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field inv       : U U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ e
    field *-rightInverse        : ∀ (x : U) → (inv x) * xe
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)


{- AbelianPartiallyOrderedGroup         = PartiallyOrderedMagma union′ AbelianGroup ⟴ :remark "over PartiallyOrderedMagmaSig" -}
record AbelianPartiallyOrderedGroup : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _≤_       : U → U → Set
    field *-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field 𝟙     : U
    field *-leftIdentity        : ∀ (x : U) → 𝟙 * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * 𝟙 ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field inv       : U → U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ 𝟙
    field *-rightInverse        : ∀ (x : U) → (inv x) * x ≡ 𝟙
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)


{- AbelianOrderedGroup                  = AbelianPartiallyOrderedGroup union TotalRelation ⟴ :remark "over OrderRelation" -}
record AbelianOrderedGroup : Setwhere
    field U     : Set
    field _*_       : U U U
    field _≤_       : U USet
    field *-≤-orderPreserving       : ∀ (x y U v : U) → xy (U * (x * v)) ≤ (U * (y * v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field 𝟙     : U
    field *-leftIdentity        : ∀ (x : U) → 𝟙 * xx
    field *-rightIdentity       : ∀ (x : U) → x * 𝟙 ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field inv       : U U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ 𝟙
    field *-rightInverse        : ∀ (x : U) → (inv x) * x ≡ 𝟙
    field *-commutative     : ∀ (x y : U) → (x * y) ≡ (y * x)
    field ≤-total       : ∀ (x y : U) → (xy) ⊎ (yx)


{- RingoidSig                           = Magma union′ AdditiveMagma :adjoin-retract₁ "*-isMagma" :adjoin-retract₂ "+-isMagma" ⟴ :remark "over Carrier" -}
record RingoidSig : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U


{- Pointed0Sig                          = PointedCarrier renaming"e to 𝟘" -}
record Pointed0Sig : Setwhere
    field U     : Set
    field 𝟘     : U


{- Pointed1Sig                          = PointedCarrier renaming′ "e to 𝟙" -}
record Pointed1Sig : Set₁ where
    field U     : Set
    field 𝟙     : U


{- Ringoid0Sig                          = RingoidSig union Pointed0Sig ⟴ :remark "over Carrier" -}
record Ringoid0Sig : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field 𝟘     : U


{- Ringoid1Sig                          = RingoidSig union′ Pointed1Sig ⟴ :remark "over Carrier" -}
record Ringoid1Sig : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U
    field 𝟙     : U


{- Ringoid01Sig                         = Ringoid0Sig union Ringoid1Sig ⟴ :remark "over RingoidSig" -}
record Ringoid01Sig : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field 𝟘     : U
    field 𝟙     : U


{- LeftRingoid                          = RingoidSig extended-by′ "*-+-leftDistributivity : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)" -}
record LeftRingoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)


{- RightRingoid                         = RingoidSig extended-by"*-+-rightDistributivity : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)" -}
record RightRingoid : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)


{- Ringoid                              = LeftRingoid  union′ RightRingoid ⟴ :remark "over RingoidSig" -}
record Ringoid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)


{- NonassociativeNondistributiveRing    = AbelianAdditiveGroup union Magma ⟴ :remark "over Carrier" -}
record NonassociativeNondistributiveRing : Setwhere
    field U     : Set
    field _+_       : U U U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field neg       : U U
    field +-leftInverse     : ∀ (x : U) → x + (neg x) ≡ 𝟘
    field +-rightInverse        : ∀ (x : U) → (neg x) + x ≡ 𝟘
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _*_       : U U U


{- NonassociativeRing                   =  NonassociativeNondistributiveRing union′ Ringoid ⟴ :remark "over RingoidSig" -}
record NonassociativeRing : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + x ≡ x
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field neg       : U → U
    field +-leftInverse     : ∀ (x : U) → x + (neg x) ≡ 𝟘
    field +-rightInverse        : ∀ (x : U) → (neg x) + x ≡ 𝟘
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _*_       : U → U → U
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)


{- PrimeRingoidSig                      = RingoidSig  union UnaryOperation ⟴ :remark "over Carrier" -}
record PrimeRingoidSig : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field prime     : U U


{- AnddeMorgan                          = PrimeRingoidSig extended-by′ "prime-*-+-deMorgan : ∀ (x y z : U) → prime (x * y) ≡ (prime x) + (prime y)" -}
record AnddeMorgan : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U
    field prime     : U → U
    field prime-*-+-deMorgan        : ∀ (x y z : U) → prime (x * y) ≡ (prime x) + (prime y)


{- OrdeMorgan                           = PrimeRingoidSig extended-by"prime-+-*-deMorgan : ∀ (x y z : U) → prime (x + y) ≡ (prime x) * (prime y)" -}
record OrdeMorgan : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field prime     : U U
    field prime-+-*-deMorgan        : ∀ (x y z : U) → prime (x + y) ≡ (prime x) * (prime y)


{- DualdeMorgan                         = OrdeMorgan union′ AnddeMorgan ⟴ :remark "over PrimeRingoidSig" -}
record DualdeMorgan : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U
    field prime     : U → U
    field prime-+-*-deMorgan        : ∀ (x y z : U) → prime (x + y) ≡ (prime x) * (prime y)
    field prime-*-+-deMorgan        : ∀ (x y z : U) → prime (x * y) ≡ (prime x) + (prime y)


{- LeftPreSemiring                      = LeftRingoid  union Semigroupunion AdditiveCommutativeSemigroup  ⟴ :adjoin-retractnil ⟴ :remark "over RingoidSig" -}
record LeftPreSemiring : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)


{- RightPreSemiring                     = RightRingoid union′  Semigroup ⟴ union′ AdditiveCommutativeSemigroup ⟴ :adjoin-retract₁ nil ⟴ :remark "over RingoidSig" -}
record RightPreSemiring : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)


{- PreSemiring                          = LeftPreSemiring union RightRingoid ⟴ :remark "over RingoidSig" -}
record PreSemiring : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)


{- NearSemiring                         = AdditiveSemigroup  union′ Semigroup ⟴ union′ RightRingoid ⟴ :adjoin-retract₁ nil ⟴ :remark "over RingoidSig" -}
record NearSemiring : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field _*_       : U → U → U
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)


{- NearSemifield = NearSemiring union Group ⟴ :remark "over Semigroup" -}
record NearSemifield : Setwhere
    field U     : Set
    field _+_       : U U U
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field inv       : U U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ e
    field *-rightInverse        : ∀ (x : U) → (inv x) * xe


{- Semifield = NearSemifield union′ LeftRingoid ⟴ :remark "over RingoidSig" -}
record Semifield : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field _*_       : U → U → U
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field inv       : U → U
    field *-leftInverse     : ∀ (x : U) → x * (inv x) ≡ e
    field *-rightInverse        : ∀ (x : U) → (inv x) * x ≡ e
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)


{- NearRing = AdditiveGroup union Semigroupunion RightRingoid ⟴ :remark "over RingoidSig" -}
record NearRing : Setwhere
    field U     : Set
    field _+_       : U U U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field neg       : U U
    field +-leftInverse     : ∀ (x : U) → x + (neg x) ≡ 𝟘
    field +-rightInverse        : ∀ (x : U) → (neg x) + x ≡ 𝟘
    field _*_       : U U U
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)


{- Rng = AbelianAdditiveGroup union′ Semigroup  Ringoid ⟴ :remark "over RingoidSig" -}
record Rng : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + x ≡ x
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field neg       : U → U
    field +-leftInverse     : ∀ (x : U) → x + (neg x) ≡ 𝟘
    field +-rightInverse        : ∀ (x : U) → (neg x) + x ≡ 𝟘
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _*_       : U → U → U
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)


{- Semiring = AdditiveCommutativeMonoid union Monoid1union Ringoidunion Left0union Right0 ⟴ :remark "over Ringoid0Sig" -}
record Semiring : Setwhere
    field U     : Set
    field _+_       : U U U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _*_       : U U U
    field 𝟙     : U
    field *-leftIdentity        : ∀ (x : U) → 𝟙 * xx
    field *-rightIdentity       : ∀ (x : U) → x * 𝟙 ≡ x
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)
    field *-leftZero        : ∀ (x : U) → 𝟘 * x ≡ 𝟘
    field *-rightZero       : ∀ (x : U) → x * 𝟘 ≡ 𝟘


{- SemiRng       = AdditiveCommutativeMonoid union′ Semigroup ⟴ union′ Ringoid ⟴ :remark "over RingoidSig" -}
record SemiRng : Set₁ where
    field U     : Set
    field _+_       : U → U → U
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + x ≡ x
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _*_       : U → U → U
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)


{- LeftPreDioid  = LeftPreSemiring union AdditiveCanonicallyOrderedMonoid ⟴ :remark "over AdditiveMonoid" -}
record LeftPreDioid : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _≤_       : U USet
    field +-≤-orderPreserving       : ∀ (x y U v : U) → xy (U + (x + v)) ≤ (U + (y + v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field ≤-+-leftCanonicalOrder        : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] ba + c)
    field ≤-+-rightCanonicalOrder       : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] bc + a)


{- RightPreDioid = RightPreSemiring union′ AdditiveCanonicallyOrderedMonoid ⟴ :remark "over AdditiveMonoid" -}
record RightPreDioid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _≤_       : U → U → Set
    field +-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U + (x + v)) ≤ (U + (y + v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + x ≡ x
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field ≤-+-leftCanonicalOrder        : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a + c)
    field ≤-+-rightCanonicalOrder       : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c + a)


{- PreDioid      = LeftPreDioid union RightRingoid ⟴ :remark "over RingoidSig" -}
record PreDioid : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _≤_       : U USet
    field +-≤-orderPreserving       : ∀ (x y U v : U) → xy (U + (x + v)) ≤ (U + (y + v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field ≤-+-leftCanonicalOrder        : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] ba + c)
    field ≤-+-rightCanonicalOrder       : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] bc + a)
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)


{- LeftDioid     = LeftPreDioid union′ Monoid ⟴ union′ Left0 ⟴ union′ Right0 ⟴ :remark "over DoubleMonoid" -}
record LeftDioid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _≤_       : U → U → Set
    field +-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U + (x + v)) ≤ (U + (y + v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + x ≡ x
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field ≤-+-leftCanonicalOrder        : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a + c)
    field ≤-+-rightCanonicalOrder       : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c + a)
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-leftZero        : ∀ (x : U) → 𝟘 * x ≡ 𝟘
    field *-rightZero       : ∀ (x : U) → x * 𝟘 ≡ 𝟘


{- RightDioid    = RightPreDioid union Monoidunion Left0union Right0 ⟴ :remark "over DoubleMonoid" -}
record RightDioid : Setwhere
    field U     : Set
    field _*_       : U U U
    field _+_       : U U U
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)
    field *-associative     : ∀ (x y z : U) → (x * y) * zx * (y * z)
    field +-associative     : ∀ (x y z : U) → (x + y) + zx + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _≤_       : U USet
    field +-≤-orderPreserving       : ∀ (x y U v : U) → xy (U + (x + v)) ≤ (U + (y + v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x yxy
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + xx
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field ≤-+-leftCanonicalOrder        : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] ba + c)
    field ≤-+-rightCanonicalOrder       : ∀ (a b : U) → (ab) ↔ (Σ[ c U ] bc + a)
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * xx
    field *-rightIdentity       : ∀ (x : U) → x * ex
    field *-leftZero        : ∀ (x : U) → 𝟘 * x ≡ 𝟘
    field *-rightZero       : ∀ (x : U) → x * 𝟘 ≡ 𝟘


{- Dioid         = LeftDioid union′ RightRingoid ⟴ :remark "over RingoidSig" -}
record Dioid : Set₁ where
    field U     : Set
    field _*_       : U → U → U
    field _+_       : U → U → U
    field *-+-leftDistributivity        : ∀ (x y z : U) → x * (y + z) ≡ (x * y) + (x * z)
    field *-associative     : ∀ (x y z : U) → (x * y) * z ≡ x * (y * z)
    field +-associative     : ∀ (x y z : U) → (x + y) + z ≡ x + (y + z)
    field +-commutative     : ∀ (x y : U) → (x + y) ≡ (y + x)
    field _≤_       : U → U → Set
    field +-≤-orderPreserving       : ∀ (x y U v : U) → x ≤ y → (U + (x + v)) ≤ (U + (y + v))
    field ≤-reflexive       : ∀ (x : U) → _≤_ x x
    field ≤-transitive      : ∀ (x y z : U) → _≤_ x y → _≤_ y z → _≤_ x z
    field ≤-antisymmetric       : ∀ (x y : U) → _≤_ y x → _≤_ x y → x ≡ y
    field 𝟘     : U
    field +-leftIdentity        : ∀ (x : U) → 𝟘 + x ≡ x
    field +-rightIdentity       : ∀ (x : U) → x + 𝟘 ≡ x
    field ≤-+-leftCanonicalOrder        : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ a + c)
    field ≤-+-rightCanonicalOrder       : ∀ (a b : U) → (a ≤ b) ↔ (Σ[ c ∈ U ] b ≡ c + a)
    field e     : U
    field *-leftIdentity        : ∀ (x : U) → e * x ≡ x
    field *-rightIdentity       : ∀ (x : U) → x * e ≡ x
    field *-leftZero        : ∀ (x : U) → 𝟘 * x ≡ 𝟘
    field *-rightZero       : ∀ (x : U) → x * 𝟘 ≡ 𝟘
    field *-+-rightDistributivity       : ∀ (x y z : U) → (y + z) * x ≡ (y * x) + (z * x)

If anything, this elaboration demonstrates our tool as useful engineering result. The main novelty being the ability for library users to extend the collection of operations on packages, modules, and then have it immediately applicable to Agda, an executable programming language.

Since the resulting expanded code is typechecked by Agda, we encountered a number of places where non-trivial assumptions accidentally got-by the MathScheme team; for example, in a number of places, an arbitrary binary operation occurred multiple times leading to ambiguous terms, since no associativity was declared. Even if there was an implicit associativity criterion, one would then expect multiple copies of such structures, one for each parenthesisation. Moreover, there were also certain semantic concerns about the design hierarchy that we think are out-of-place, but we chose to leave them as is —e.g., one would think that a “partially ordered magma” would consist of a set, an order relation, and a binary operation that is monotonic in both arguments wrt to the given relation; however, PartiallyOrderedMagma instead comes with a single monotonicity axiom which is only equivalent to the two monotonicity claims in the setting of a monoidal operation. Nonetheless, we are grateful for the source file provided by the MathScheme team.

  • Unlike other systems, ours does not come with a static set of module operators —it grows dynamically, possibly by you, the user.

We implore the readers to build upon our code of theories above by, for example, define the notion of homomorphism for every single one of the theories. Besides being tiresome, such a manual process is also error-prone. Instead, one can automatically derive this concept! ( The construction of such a variational is delegated to the more involved second part of the user manual. )

Likewise, for other concepts from universal algebra —which is useful to computer science in the setting of specifications.

3.10 Summary of Sample Variationals Provided With The System

Below are some variationals that can be used immediately by a user —no Lisp knowledge required. If you simply mention one of them in a comment, then load the file with C-c C-l, then you may hover over them to see their full documentation. You may also look at the second part of the user manual below to see accompanying examples.

Name Description
record Reify a PackageFormer as a valid Agda record.
extended-by ⊙ Extend a PackageFormer by a string-“;”-list of declaration.
keeping ★ Largest well-formed PackageFormer consisting of a given list of elements.
union ⊙ Union two PackageFormers into a new one, maintaining relationships.
flipping ⊙ Dualise a binary operation or predicate.
unbundling Consider the first N elements, which may have definitions, as parameters.
data Reify a PackageFormer as a valid Agda algebraic data type.
open Reify a given PackageFormer as a parameterised Agda “module” declaration.
opening ★ Open a record as a module exposing only the given names.
open-with-decoration Open a record, exposing all elements, with a given decoration.
sorts Keep only the types declared in a grouping mechanism.
signature Keep only the elements that target a sort, drop all else.
rename ⊙ Apply a Name → Name function to the elements of a PackageFormer.
renaming ⊙ ★ Rename elements using a list of “to”-separated pairs.
decorated ⊙ Append all element names by a given string.
codecorated ⊙ Prepend all element names by a given string.
primed ⊙ Prime all element names.
subscriptedᵢ ⊙ Append all element names by subscript i : 0..9.
hom Formulate the notion of homomorphism of parent PackageFormer algebras.
Argument is a string of “;”-separated items consisting of “to”-separated pairs.
  • keeping has input being “;”-separated string of proper names only.
The resulting PackageFormer is furnished with an element toP : P; to P = record {oldᵢ = newᵢ} where P is the name of the parent, or source, PackageFormer. This element provides a way to view the resulting child presentation as an instance of the original presentation.
  • If we dislike the toP naming, we can simply invoke renaming.
  • Sometimes these ‘maps’ are invertible; we leave it as an exercise to the interested reader to formulate toP⁻¹ in the variationals, when possible.
  • Invoke a variational with :adjoin-retract nil if you do not want this map to be added; or use t in-place of nil if you insist on it being added.

Below are the five meta-primitives from which all variationals are borne, followed by a few others that are useful for extending the system by making your own grouping mechanisms and operations on them. Using these requires a small amount of Lisp, which is tersely reviewed in the second part of the user manual, below.

Name Description
:waist Consider the first N elements as, possibly ill-formed, parameters.
:kind Valid Agda grouping mechanisms: record, data, module.
:level The Agda level of a PackageFormer.
:alter-elements Apply a List Element → List Element function over a PackageFormer.
Compose two variational clauses in left-to-right sequence.
map ⊙ Map a Element → Element function over a PackageFormer.
generated Keep the sub-PackageFormer whose elements satisfy a given predicate.

4 User Manual II: Extending the System

Herein we demonstrate how with a little bit of Lisp, one may create any desired form of grouping mechanism as well as operation between groupings.

We begin with a brief review of Lisp basics such as list operations then immediately dive into creating an Agda module that contains our new user-defined constructs, within 700-comments.

4.1 Lisp Interface

Summary of useful functions for programming variationals.

  • A PackageFormer is nothing more than a list of elements.
  • An element is a record of 4 items: A qualifier, a name, a type, and a list of equational clauses for that name.

Below are functions that will be used in subsequent subsections in order to form user-defined variationals. Consult these tables as necessary and look at an Elisp Cheat Sheet, if need be.

  • Elements

    Let ℰ denote one of qualifier, name, type, equations in the following list of useful functions.

    (make-element 𝓆 𝓃 𝓉 ℯ) An element value is formed.
    (element-ℰ e) Project component ℰ from element e.
    (map-ℰ f e) Return a copy of e with component ℰ updated by function f.
    (element-replace old new e) Replace all string occurances of old by new in element e.
  • Lists

    (list x₀ … xₙ) Form a list of elements xᵢ.
    (car xs) Obtain first element of list xs.
    (cdr xs) Obtain all but first element of xs.
    (cons x xs) Form a new list with car x and cdr xs.
    (mapcar (λ it → ⋯it⋯) xs) Map the given function on xs.
    (--map (⋯it⋯) xs) Map the implicit function (λ it → ⋯it⋯) on xs.
    (-cons* x₀ … xₙ xs) (cons x₀ (cons x₁ (⋯ (cons xₙ xs)))).
       
  • Strings

    (concat s₀ … sₙ) Concatenate strings sᵢ.
    (s-replace old new s) Replace all string occurrences of old by new in string e.
    (rename-mixfix f op) Rename string op according to function f by ‘leaping over’ underscores.
      E.g., f, op = (λ x → (concat x "′")), _⊕_ ⇒ _⊕′_.

You can always see the documentation of an item with C-h o!

4.2 User Manual Header

A literate programming approach to a user manual.

In order for our manual's examples to be up-to-date, we will take a literate approach to producing them. Namely, the Agda code here is ‘tangled’ from this prose into an Agda file which can then be checked by an Agda process. Whence, this file is the de-facto source.

Let's start-off with a usual Agda header:

{- This loads the PackageFormer metaprogram; press C-x C-e after the closing “)” below.                 -}
{- (progn (load-file "~/.emacs.d/agda-next-700-module-systems.el") (agda-next-700-module-systems-mode)) -}

module package-former where

open import package-former-generated
open import Level
open import Data.Bool
open import Data.List using (List; _∷_; []; foldr)
import Relation.Binary.PropositionalEquality as ≡; openusing (_≡_)

{- Let's ensure content of User Manual part I actually type checkes -}
{- Feel free to comment this line out. -}
import package-former-user-manual-i

To make the resulting Agda file somewhat self-contained, in case anyone wishes to read that or load it into Agda and play with it, let's add a blurb.

{-
0. There are a number of common use-cases.
1. We can handle all of them & more, since we're extensible.
  - Mention the Lean & Coq, as well as the Agda, repeated fragments.
2. The resulting setup is pragmatic: It is unobtrusive in the
   traditional Agda coding style in that it happens in the background.
3. It fills a particular need; the desire to avoid repetitious code.
-}

Before getting to the meat of things, it is important to note that comments begun with {- and followed by a space are treated as usual Agda comments, whereas those without a following space such as {-700 and {-lisp are picked-up by our meta-program.

For example, having no space between “{-” and “lisp” would cause the following block to be executed as a Lisp form.

{- lisp
(message-box "Hello")
(message-box "World")
-}

Alternatively, here is the PackageFormer for M-Sets from the introduction. It is a useful example since it is multi-sorted.

{-700
PackageFormer M-Set : Set₁ where
   Scalar  : Set
   Vector  : Set
   _·_     : Scalar → Vector → Vector
   𝟙       : Scalar
   _×_     : Scalar → Scalar → Scalar
   leftId  : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   assoc   : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
-}

Let us also introduce a slightly more syntactically-involved example: A PackageFormer with equations. The equations, depending on our perspective, —i.e., the variational invoked— may be thought of as:

  • Derived elements; e.g., in a record, they are a definitional extension and for an ADT, they are methods defined on the constructors.
  • Coherence constraints; e.g., in a record, we may interpret an equation 𝓁 = 𝓇 as an additional axiom ∀ {⋯} → ℓ ≡ 𝓇 —e.g., when a user may supply an efficient definition of 𝓁 but is constrained to have a particular behaviour 𝓇.
  • Rewrite rules; e.g., in an ADT, an equation may simply act as an alias and is to be used in rewriting the remainder of the ADT declaration.
  • Ignored components; e.g., in a record, we may ignore the equations altogether and lift the associated names into being fields —e.g., _≈_ would usually be lifted into a field and its stringent implementation via _≡_ is used as a motivating or simplifying factor.
{-700
PackageFormer MonoidP : Set₁ where

    -- A few declarations
    Carrier : Set
    _⨾_     : Carrier → Carrier → Carrier
    Id      : Carrier
    assoc   : ∀ {x y z} → (x ⨾ y) ⨾ z ≡ x ⨾ (y ⨾ z)

    -- We have a setoid-like structure; with a default implementation
    _≈_   : Carrier → Carrier → Set
    _≈_   = _≡_
    ⨾-cong : ∀ {x y x′ y′} → x ≈ x′ →  y ≈ y′ → (x ⨾ y) ≈ (x′ ⨾ y′)
    ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}

    -- For now only one item in a declaration;
    -- namely “Lid” & “Rid” cannot be declared in one line.
    Lid : Carrier → Carrier
    Lid x = Id ⨾ x
    Rid : Carrier → Carrier
    Rid x = x ⨾ Id

    -- Agda permits pure, non-pattern-matching, equations between “fields” in a record.
    concat : List Carrier → Carrier
    concat = foldr _⨾_ Id

    -- More declarations
    leftId  : ∀ {x : Carrier} → (Id ⨾ x) ≈ x
    rightId : ∀ {x : Carrier} → Rid x ≈ x

    -- Since there are no more pure declarations, “fields”, subsequent equations
    -- may use pattern matching.

    Id² : (Id ⨾ Id) ≈ Id
    Id² = rightId

    concatₚ : List Carrier → Carrier
    concatₚ []       = Id
    concatₚ (x ∷ xs) = x ⨾ concatₚ xs
-}

Notice that there is no particular segregation of declarations and equations. Simply: A declaration may optionally have an associated equation; however once an equation uses pattern matching then all subsequent declarations must also have equations —this is a constraint of the current Agda implementation—; as such, the equation for ⨾-cong uses Agda's pattern-matching-λ.

4.3 Empty Variationals

A summary of the key features of 𝒱ariationals

The simplest user definable variational is the empty one:

{-700
-- Variational with empty right hand side.
𝒱-identity =
-}

The prefix 𝒱- signals to the Elisp meta-program that this particular equation is intended to be a variational and should be loaded into Emacs as such. Indeed, you may view the documentation and elaborated Lisp of this definition using C-h o RET 𝒱-identity.

Remember: The prefix 𝒱- only occurs at the definition site, the call site omits it. Why? We have augmented the Emacs system with a new functional definition, and the 𝒱- serves as a namespace delimiter.

Indeed, we may now perform the following invocation —within 700-comments.

{-700
MonoidPⁱᵈ = MonoidP identity
-}

Loading the meta-program using Agda's usual C-c C-l lets us hover over MonoidPⁱᵈ to see its elaboration is precisely that of MonoidP.

Moreover, to be useful, all variationals have tooltips showing their user-defined documentation. If we hover over identity, we are informed that it is undocumented. User documentation is optional and may appear immediately following the =, as follows.

{-700
𝒱-no-op = "This is the do-nothing variational"
-}

We may substitute equals-for-equals:

{-700
-- No variational clauses needed!
MonoidP⁰  = MonoidP
-}

Doing nothing is meaningful with respect to a composition operation. Momentarily, “⟴” is forwards composition: We ‘thread’ the Pf through the variationals vᵢ in order.

{-700
-- Identity of composition ⟴
MonoidPᶜ = MonoidP ⟴
-}
Operationally: Pf ⟴ v ≈ Pf v ⟴ ≈ Pf v

We may also augment a variational with positional and (optional) keyword arguments that have default values. The keyword arguments along with their default value, if any, are enclosed in parenthesis.

{-700
𝒱-test positional (keyword 3) another = "I have two mandatory arguments and one keyword argument"

Monoid-test = MonoidP ⟴ test "positional arg₁" "positional arg₂" :keyword 25
-}

We are not doing anything with the arguments here; we shall return to this in later subsections.

4.4 ⇨ Errors

Even though this is a prototype, we wish it to be useful to ourselves and to others —especially those who take a quick glance, think they got it, and try things out only to not have them work immediately. As such, we have implemented a cute little error-reporting system.

If you try to load, C-c C-l, but your 700-syntax is wrong, you get an immediate error explaining why ♥‿♥

For example, suppose we accidentally wrote tester instead of test, which we defined at the end of the previous section, as in the following. ( The space before 700 is so that this crashing block is not in effect. )

{-   700
𝒱-whoops  = tester 1 2 :keyword 3
-}

When we try to load our Agda file the Agda process is interrupted and we are warned:

700: Did you mistype a variational’s name: “tester” is not defined.

    ⇨   whoops = tester 1 2 :keyword 3
    ⇨   Use the PackageFormer menu to see which variationals are defined.

The 700 system informs us of our fault in “quotes”, suggests a solution, and points to the offending declaration hierarchy. Neato (•̀ᴗ•́)و

The “quotes” help, in this case, when there are multiple variationals being invoked in a clause.

Of-course, we do not attempt to cover all possible errors —e.g., wrong number of arguments or division by zero— instead relying on Emacs Lisp's native error mechanism.

4.5 Records and Meta-Primitives :kind & :alter-elements

Let's begin with the simplest thing: Realising these fictitious ‘PackageFormers’ as records.

An Agda ‘record’ is just a PackageFormer where the qualifier PackageFormer has been replaced with record and each element is qualified by Agda keyword field. We may declare this particular configuration using the meta-primitives :kind and :alter-elements, as follows.

{-700
𝒱-record₀ = :kind record :alter-elements (λ es → (--map (map-qualifier (λ _ → "field") it) es))
-}

Huh? The :kind part was already explained, the :alter-elements is the powerhouse of our system. It takes a function with argument being the list of PackageFormer elements, es, then we perform a functorial list map where each element is implicitly referred to as it. Then the map function is to alter the qualifier of an element by replacing it with the string "field". In Agda syntax this corresponds to: λ es → map (λ it → (map-qualifier (λ _ → "field") it)) es.
Notice that the Agda form and Lisp form are only one outer parenthesis off from each other —Lisp is easy!

The :key value syntax is inspired from Lisp

Unsurprisingly, we have elected to name this grouping mechanism configuration as 𝒱-record. Let's try it out.

{-700
M-Set-Record = M-Set record₀
-}

The system picks this up, looks up M-Set which was defined in the first section earlier, looks up the variational record, then runs that configuration to generate:

{- M-Set-Record = M-Set record₀ -}
record M-Set-Record : Set₁ where
   field Scalar     : Set
   field Vector     : Set
   field _·_        : Scalar → Vector → Vector
   field 𝟙      : Scalar
   field _×_        : Scalar → Scalar → Scalar
   field leftId     : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   field assoc      : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)

Nothing too remarkable; the keyword field has been inserted and the rewrite PackageFormer ↦ record has been performed. The above is the exact generated result of the system —the comment indicates the source of this generated code.

4.6 Equation Accommodating Record Variational

Since record formation is a variational that is likely to be used often, it is sensible to document it —which in turn is attached to all occurences of the variational name via tooltips. Moreover, let's strengthen it to accomodate PackageFormers with equations.

{-lisp
(𝒱 record₁ (discard-equations nil)
 = "Reify a variational as an Agda “record”.
    Elements with equations are construed as
    derivatives of fields  ---the elements
    without any equations--- by default, unless
    DISCARD-EQUATIONS is provided with a non-nil value.
   "
  :kind record
  :alter-elements
    (λ es →
      (thread-last es
      ;; Keep or drop eqns depending on “discard-equations”
      (--map
        (if discard-equations
            (map-equations (λ _ → nil) it)
            it))
      ;; Unless there's equations, mark elements as fields.
      (--map (map-qualifier
        (λ _ → (unless (element-equations it)
               "field")) it)))))
-}

Unlike 𝒱-identity from a previous section, we have decided to split this definition into multiple lines by enclosing it in {-lisp ⋯ -}~. Such blocks may contain arbitrary Lisp to be executed and so all contents must be Lisp forms —notice the 𝒱-⋯ from 700-blocks has been exchanged for a parenthesised (𝒱 ⋯) within lisp-blocks.

Let's try this out.

First, using only the default value —which doesn't discard equations.

{-700
Monoid-Record-derived = MonoidP record₁
-}
record Monoid-Record-derived : Setwhere
    field Carrier       : Set
    field _⨾_       : Carrier Carrier Carrier
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)
    _≈_     : Carrier CarrierSet ; _≈_   = _≡_
    ⨾-cong      : ∀ {x y xy′} → xx′ →  yy′ → (xy) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier Carrier ;   Lid x = Idx
    Rid     : Carrier Carrier ;   Rid x = x Id
    concat      : List Carrier Carrier ;  concat = foldr _⨾_ Id
    field leftId        : ∀ {x : Carrier} → (Idx) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid xx
    Id²     : (Id Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier Carrier ;  concatₚ []       = Id ; concatₚ (xxs) = xconcatxs

Second, discarding equations and lifting all elements into field-s.

{-700
Monoid-Record-field = MonoidP record₁ :discard-equations t
-}
record Monoid-Record-cons : Setwhere
    field Carrier       : Set
    field _⨾_       : Carrier Carrier Carrier
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)
    field _≈_       : Carrier CarrierSet
    field ⨾-cong        : ∀ {x y xy′} → xx′ →  yy′ → (xy) ≈ (x′ ⨾ y′)
    field Lid       : Carrier Carrier
    field Rid       : Carrier Carrier
    field concat        : List Carrier Carrier
    field leftId        : ∀ {x : Carrier} → (Idx) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid xx
    field Id²       : (Id Id) ≈ Id
    field concatₚ       : List Carrier Carrier

Let's also codify the converse operation of marking a grouping mechanism abstract to avoid elaboration.

(𝒱 PackageFormer = "Mark a grouping mechanism as abstract, so that it is NOT elaborated into concrete Agda." :kind PackageFormer)

4.7 A Coherent Equation Accommodating Record Variational

Yet another option to handling equations is to drop the names that have equations associated with them. To tackle such a scenario requires the remaining elements to be well-defined and so requires “the largest sub-PackageFormer”.

Coherent relationships are just graphs in disguise, so let's abstract away the details and solve a graph-theoretic problem. In {-lisp ⋯ -} blocks we may have arbitrary Emacs Lisp code and so include the following —which has a large number of shortcomings, but the aim is a simple demonstration of Lisp code for the Agda user, not to be robust Lisp code. The name, graph-map, may not be ideal but it seems good enough, for now.

;; p ≈ symptom; f ≈ medicine; adj ≈ neighbouring dependency
;;
(cl-defun graph-map (p f adj xs &optional keep-only-marked)
  "Map the nodes XS satisfying P by F along adjacency ADJ.
<<docs('graph-map)>>"
  (let* (;; Using -map instead of -filter since nodes may become
         ;; sickly later on, position matters.
         (sickly (-map p xs))
         ;; Obtain the items that are currently ‘sickly’.
         (get-sickly (lambda ()
                       (--filter it (--zip-with (when it other) sickly xs))))
         ;; infected x  ≡ x has a sickly neighbour
         (infected (λ x → (--any (funcall adj x it) (funcall get-sickly)))))

     ;; Propogate sickness.
     (loop for _ in xs
           do (loop for x in xs
                    for i from 0
                    do (when (funcall infected x) (setf (nth i sickly) t))))

     ;; Apply medication to sickly elements only.
     (--filter it (--map (if (-contains-p (funcall get-sickly) it)
                (funcall f it)
                (unless keep-only-marked it))
            xs))))

Here's how this works —the following is what the incantation above <<docs('graph-map)>> refers to, and the reader may ignore all <<…>> as they are a backend ‘literate programming’ utility.

  • F is performed on nodes satisfying P, all neighbours are then considered to satisfy P and the process repeats recursively.
  • E.g., nodes exhibiting symptoms P are given medicine F, and their sickness spreads to their neighbours who in turn become ill thereby requiring medication, and the process continues.
  • ADJ is a binary relation denoting adjacency.
    • (adj x y) ≈ “x depends on, or is a neighbour, of y.”
  • For example, a graph of 10 nodes, with an edge between multiples; where nodes 3, 4, 5 are initally ill.

       (graph-map (λ x → (-contains-p '(3 4 5) x))
                  (λ x → (format "medicated-%s" x))
                  (λ x y → (zerop (mod x y)))
                  '(1 2 3 4 5 6 7 8 9 10))(1 2 medicated-3 medicated-4 medicated-5 medicated-6 7
           medicated-8 medicated-9 medicated-10)
    

Testing this graph-theoretic solution for our setting shows it to be a reasonable fit.

;; Example: Dropping the implementations of the first 2 items.
(setq i -1)
(graph-map (λ _ → (incf i) (< i 2))
           (λ x → (map-equations (λ _ → nil) x))
           ;; x depends on y  ≡  x mentions y in its type or equations.
           (λ x y → (s-contains? (s-replace "_" " " (element-name y)) (s-join " " (cons (element-type x) (element-equations x)))))
           (parse-elements '("A : Set" "_≈_ : A → A → Set" "_≈_ = _≡_" "easy : ∀ {x} → x ≈ x" "easy = refl"
                             "another : ∀ {x} → Set" "another = easy" "by : Set₁" "by = Set"))))
⇒
  A       : Set
  eq      : A → A → Set    ;; implementation dropped
  easy    : ∀ {x} → x ≈ x  ;; ditto, since it depends on ≈'s implementation
  another : ∀ {x} → Set    ;; ditto, since it depends on easy's implementation
  by      : Set₁
  by      = Set₁

Let's introduce a dedicated form for element values:

  • Mark elements in a given list, and recursively mark all those that depend on them. Return the list of elements with the marked ones being altered.
  • MARK and ALTER are expressions mentioning IT, a value of ELEMENTS, and so are implicit functional expressions.
  • Only the MARKED elements are kept.
(cl-defmacro --graph-map (mark alter elements &optional (keep-only-marked t))
  "Recursively ALTER and MARK elements and their dependents.
<<docs('--graph-map)>>"
  `(graph-map (λ it → ,mark)
              (λ it → ,alter)
              ;; x depends on y  ≡  x mentions y, with all or no undescores,
              ;;                    in its type or equations.
              (λ x y →
                 (or (s-contains? (s-replace "_" " " (element-name x))
                                  (s-join " " (cons (element-type y)
                                                    (element-equations y))))
                     (s-contains? (element-name x)
                                  (s-join " " (cons (element-type y)
                                                    (element-equations y))))))
              ,elements ,keep-only-marked))

Now the previous example may be invoked as:

(setq i -1)
(--graph-map (progn (incf i) (< i 3))
             (map-equations (λ _ → nil) it)
             (parse-elements '("A : Set" "_≈_ : A → A → Set" "_≈_ = _≡_"
                               "easy : ∀ {x} → x ≈ x" "easy = refl"
                               "another : ∀ {x} → Set" "another = easy" "by : Set₁" "by = Set"))))

With these pieces in hand, let's form

(𝒱 record (discard-equations nil) (and-names nil)
 = "Reify a variational as an Agda “record”.

    By default, elements with equations are construed as
    derivatives of fields  ---the elements
    without any equations.

    ⇨ DISCARD-EQUATIONS is nil by default.
      If provided with a non-nil value, equations are dropped indiscriminately.

    ⇨ AND-NAMES is nil by default and only takes
      effect when DISCARD-EQUATIONS is active.
      If provided with a non-nil value, names with
      equations are dropped altogether; but some may be kept
      if they are needed for some fields to be well-defined.
   "
  :kind record
  :alter-elements
    (λ es →
      (thread-last es

      (funcall (λ es′ → (if (not discard-equations) es′
               (--map (map-equations (-const nil) (map-qualifier (-const (when (element-equations it) 'eqns)) it)) es′))))

      (funcall (λ es′ → (if (not and-names) es′
        (--graph-map (not (equal 'eqns (element-qualifier it))) it es′))))

      ;; Unless there's equations, mark elements as fields.
      (--map (map-qualifier
        (λ _ → (unless (element-equations it)
               "field")) it)))))

We can obtain the previous variationals rcordᵢ as well as new presentations.

{-700
Monoid-Record-derived-again  = MonoidP record
Monoid-Record-derived-again2 = MonoidP record :and-names t
Monoid-Record-field-again    = MonoidP record :discard-equations t
Monoid-Record-no-equationals = MonoidP record :discard-equations t :and-names t
-}

The last form yields:

record Monoid-Record-no-equationals : Setwhere
    field Carrier       : Set
    field _⨾_       : Carrier Carrier Carrier
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)

4.8 Typeclasses —Parameterised Records— and Meta-Primitives :waist & :level

We mentioned the “waist” before, but what is it exactly? I propose that the difference between ‘field’ and ‘parameter’ is an illusion —as is that of ‘input’ and ‘output’ when one considers relations rather than deterministic functions.

For example, let's alter the previous variation declaration to lift the waist up 2 positions.

{-700
𝒱-typeclass-attempt = record ⟴ :waist 2
-}

Notice we have avoided repeating the definition of the record variational from earlier by making use of composition. More on it later, but it suffices to say that above we could replace record ⟴ with the exact text of 𝒱-record = ⋯ right-hand-side and all would continue work.

Trying this out, below, one notices that the first two elements of the PackageFormer have been lifted into being parameters, while the rest have been construed as fields.

{-700
M-Set-TypeClass = M-Set typeclass-attempt
-}
record M-Set-TypeClass (Scalar : Set) (Vector : Set) : Setwhere
   field _·_        : Scalar Vector Vector
   field 𝟙      : Scalar
   field _×_        : Scalar Scalar Scalar
   field leftId     : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   field assoc      : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)

While this typechecks according to Agda standards, it is not ideal to human standards since the level of the resulting package is larger than necessary. The meta-primitive :level allows us to inc-rement or dec-crement the current level of a PackageFormer, so we may instead define:

{-700
𝒱-typeclass₂ = record ⟴ :waist 2 :level dec
MonoidT₂      = MonoidP typeclass₂
-}
record MonoidT (Carrier : Set) (_⨾_ : Carrier Carrier Carrier) : Set where
    field Id        : Carrier
    field assoc     : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)
    field leftId        : ∀ {x : Carrier} → Idxx
    field rightId       : ∀ {x : Carrier} → x Idx

For fun, here are a few more to play with:

{-700
MonoidT₃         = MonoidP record ⟴ :waist 3 :level dec
-- MonoidT₃-again   = MonoidP ⟴ record ⟴ unbundling 3
M-Set-Typeclass₂ = M-Set record ⟴ typeclass₂
-}

In particular, the last example suggest that our composition is idempotent, but this is clearly not the case. Indeed, here's a pretty alternative to the meta-primitive :waist that is not ⟴-idempotent but is in-fact a homomorphism: unbundling n ⟴ unbundling m ≈ unbundling (n + m).

(𝒱 unbundling n
 = "Make the first N elements as parameters to the PackageFormer.

    Any elements in above the waist line have their equations dropped.
    As such, unbundling is not invertible.
   "
   :waist n
   :alter-elements (λ es →
     (-let [i 0]
       (--graph-map (progn (incf i) (<= i n))
                    (map-equations (-const nil) it)
                    es))))

( The graph-map operation was defined in the previous section. )

Incidentally, this solves the problem of lifting the waist to include elements with equations.

{-700
-- Ill-formed in Agda: A defintion is not a parameter!
MonoidP-Typeclass₅ = MonoidP :waist 5
-}
{- Kind “PackageFormer” does not correspond  to a concrete Agda type.
{- MonoidP-Typeclass= MonoidP :waist 5 -}
PackageFormer MonoidP-Typeclass (Carrier : Set) (_⨾_ : Carrier Carrier Carrier) (Id : Carrier) (assoc : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)) (_≈_ : Carrier CarrierSet ; _≈_ = _≡_) : Setwhere
    ⨾-cong      : ∀ {x y xy′} → xx′ →  yy′ → (xy) ≈ (x′ ⨾ y′) ;  ⨾-cong = λ{ ≡.refl ≡.refl → ≡.refl}
    Lid     : Carrier Carrier ;   Lid x = Idx
    Rid     : Carrier Carrier ;   Rid x = x Id
    concat      : List Carrier Carrier ;  concat = foldr _⨾_ Id
    leftId      : ∀ {x : Carrier} → (Idx) ≈ x
    rightId     : ∀ {x : Carrier} → Rid xx
    Id²     : (Id Id) ≈ Id ;  Id² = rightId
    concatₚ     : List Carrier Carrier ;  concatₚ []       = Id ; concatₚ (xxs) = xconcatxs -}
{-700
MonoidT₅ = MonoidP ⟴ unbundling 5 ⟴ record
-}
record MonoidT (Carrier : Set) (_⨾_ : Carrier Carrier Carrier) (Id : Carrier) (assoc : ∀ {x y z} → (xy) ⨾ zx ⨾ (yz)) (_≈_ : Carrier CarrierSet) : Setwhere
    field ⨾-cong        : ∀ {x y xy′} → xx′ →  yy′ → (xy) ≈ (x′ ⨾ y′)
    field Lid       : Carrier Carrier
    field Rid       : Carrier Carrier
    field concat        : List Carrier Carrier
    field leftId        : ∀ {x : Carrier} → (Idx) ≈ x
    field rightId       : ∀ {x : Carrier} → Rid xx
    field Id²       : (Id Id) ≈ Id
    field concatₚ       : List Carrier Carrier

:smile:

4.9 Primed Decoration and rename-mixfix

When we have two occurrences of a structure, we may want one of them to be decorated say with a prime so as to disambiguate them easily rather than have to qualify all of their components.

{-700
-- Intentionally erroenous attempt.
𝒱-primed-attempt = :alter-elements (λ es → (--map (map-name (λ n → (rename-mixfix (λ np → (concat np "′")) n)) it) es))

-- M-Set′-attempt = M-Set record ⟴ primed-attempt
-}
record M-Set′-attempt : Setwhere
   field Scalar′        : Set
   field Vector′        : Set
   field _·′_       : Scalar Vector Vector
   field 𝟙′     : Scalar
   field _×′_       : Scalar Scalar Scalar
   field leftId′        : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   field assoc′     : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
{- This is crashing Agda code; as expected. Working improvements below. -}

There are number of issues to address.

  1. The system comes with a Lisp methods map-name and map-type that yield the name and type part, respectively, of a PackageFormer element.
  2. The :key value pairs have legitimate Lisp for the value positions.

    The basics of list processing, such as maps/filters/folds, with Lisp suffice for a rich inventory of possible configurations. Moreover, the functional nature of such higher-order functions ought to be familiar to any Agda coder worth their salt.

    Here's a terse tutorial rendered as an Elisp Cheat Sheet.

  3. One would expect catenating a prime to the mixfix name _×_ would yield _×_′ but above it yielded _×′_. Indeed, the former would yield confusing expressions of the form 1 × 2 ′ whereas the latter permits 1 ×′ 2. It is with this pragmatic usage that rename-mixfix performs a rewrite to a name by jumping over the Agda mixfix marker, _, if it occurs at the start or end of a name.

    As an additional example, the name _≈_∶_, under the above scheme, would have rewritten to _≈_∶′_ thereby allowing terms such as x ≈ y ∶ A → f x ≈ f y ∶′ B —a elegant way to express that, say, f is a setoid homomorphism. If the prime scheme were instead a prepend, we would have obtained the name _′≈_∶_.

Notice that we have fields such as 𝟙′ : Scalar whose type is a free variable: Scalar no longer refers to any field. As such, the above code is ill-typed. The solution then is to propagate any changes a name has down to its siblings. We will return to this later in the form of a map variational.

4.10 First-class PackageFormers

The previous example caused an Agda typechecking error, however if we do not invoke the record variational then the result is a non-Agda syntactical item which can only be the subject of further alteration via PackageFormer combinators.

{-700
M-Set′-attempt-raw = M-Set primed-attempt
-}
{- Kind “PackageFormer” does not correspond  to a concrete Agda type.
PackageFormer M-Set′-attempt-raw : Set₁ where
   Scalar′      : Set
   Vector′      : Set
   _·′_     : Scalar → Vector → Vector
   𝟙′       : Scalar
   _×′_     : Scalar → Scalar → Scalar
   leftId′      : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   assoc′       : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋) -}

This is interesting: We have not generated a concrete —legitimate Agda construct— but instead yielded a new abstract grouping mechanism which may be instantiated later on, whenever desired.

Likewise, we have already declared M-Set-Record = M-Set record and now we may apply our priming variational.

{- M-Set-Record′ = M-Set-Record primed-attempt -}
record M-Set-Record′ : Set₁ where
   field Scalar′        : Set
   field Vector′        : Set
   field _·′_       : Scalar → Vector → Vector
   field 𝟙′     : Scalar
   field _×′_       : Scalar → Scalar → Scalar
   field leftId′        : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   field assoc′     : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
We may apply variationals even to concrete Agda packaging constructs!

Of-course, we may simply want to obtain M-Set-Record′ without having first to define M-Set-Record, and so we may use the variational composition operator .

{- M-Set-Record′ = M-Set record ⟴ primed-attempt -}
record M-Set-Record′ : Set₁ where
   field Scalar′        : Set
   field Vector′        : Set
   field _·′_       : Scalar → Vector → Vector
   field 𝟙′     : Scalar
   field _×′_       : Scalar → Scalar → Scalar
   field leftId′        : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   field assoc′     : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)

Since the record and primed configurations are ‘disjoint’, they commute with respect to composition. The reader may want to confirm the following identifications:

      M-Set-Record′ M-Set recordprimed-attempt M-Set primed-attemptrecord M-SetrecordM-Set-Record primed-attempt

It is important to remember that these primed perspectives do not typecheck in Agda due to the free-variable issue mentioned earlier. We are only demonstrating composition, , in this section; in a later section we fix-up primed.

4.11 Variationals with Arguments ---map, rename, [co]decorated, subscriptedᵢ, renaming

Thus far our variationals have been nullary, let's consider otherwise.

For example, let's add arguments to the typeclass variational from earlier.

{-700
𝒱-typeclass height (level 'dec) = record ⟴ :waist height :level level
M-Set2-Typeclass₃ = M-Set typeclass 3 :level 'inc
M-Set0-Typeclass₃ = M-Set typeclass 3
-}
record M-Set2-Typeclass (Scalar : Set) (Vector : Set) (_·_ : Scalar Vector Vector) : Setwhere
   field 𝟙      : Scalar
   field _×_        : Scalar Scalar Scalar
   field leftId     : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   field assoc      : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)

record M-Set0-Typeclass (Scalar : Set) (Vector : Set) (_·_ : Scalar Vector Vector) : Set where
   field 𝟙      : Scalar
   field _×_        : Scalar Scalar Scalar
   field leftId     : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   field assoc      : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)
  • Argument come before the = in a variational's definition and the may be used as if they were constants on the right-hand side.
    1. Above we introduced the named arguments height and level.
    2. The first is positional, and the second is a keyword argument with default value being a decrement value.
    3. We then passed the argument level to the meta-primitive :level.
  • Invocation of variationals has positional arguments first then named arguments afterwards. One supplies a named argument in the form :argument-name the-value —this is Lisp-inspired syntax.

    Consequently, order is irrelevant for named arguments.

    • Supplying :key value pairs where the key is not a named argument of the variational yields a error message indicating the allowable keys.

Let's code up the priming operation as a reusable pattern —call it map.

(eval-and-compile
(𝒱 map elements (support-mixfix-names nil) (adjoin-retract nil) (adjoin-coretract nil)
   = "Apply function ELEMENTS that acts on PackageFormer elements,
      then propogate all new name changes to subsequent elements.

      There is minimal support for mixfix names, but it may be
      ignored by setting SUPPORT-MIXFIX-NAMES to be nil.

      When ADJOIN-RETRACT is non-nil, we adjoin a “record {oldᵢ = nameᵢ}”
      view morphism; i.e., record translation.

      Clauses “f = f” are considered to occur only in views, record translations,
      and so only the RHS occurance is updated to a new name.
      C.f. the definition of element-retract.
      "
     :alter-elements (lambda (es)

    (let* ((es′    (mapcar elements es))
           (names  (mapcar #'element-name es))
           (names′ (mapcar #'element-name es′)))

      ;; Replace all occurances of old names with corresponding new ones.
      (loop for old in names
            for new in names′
            do (setq es′ (--map (element-replace old new it :support-mixfix-names support-mixfix-names) es′)))

     ;; return value
     (-concat es′ (when adjoin-retract (list (element-retract $𝑝𝑎𝑟𝑒𝑛𝑡 es :new es′ :name adjoin-retract)))
                  (when adjoin-coretract (list (element-retract $𝑝𝑎𝑟𝑒𝑛𝑡 es′ :new es :name adjoin-coretract :contravariant t)))))))
)

This is a prototype; ideally the variational definition would be rendered in Agda code. Rather than using functional combinators such as unzip, map, zip-with, for diversity, we used imperative constructs.

Important observations include:

  1. The Lisp code lives in a {-lisp ⋯ -} block.
  2. 700-comments have single-line 𝒱-name args = rhs, whereas lisp-blocks have multi-line (𝒱 name args = rhs) —the dash after the 𝒱 is gone and outer-most parenthesis are added.
  3. To provide minimal accommodation for mixfix names, we simply remove the Agda argument indicator ‘_’ when performing rewrites.

    E.g., Agda let's you declare a name such as _⊕_ and use it without mentioning the underscore as in x ⊕ y and so the rename _⊕_ ↦ _⊕′_ would have no effect since _⊕_ does not occur as a substring in x ⊕ y, whence the need to ignore the underscores.

Trying it out:

{-700
MR𝕏    = M-Set record ⟴ map (λ e → (map-name (λ n → (rename-mixfix (λ x → (concat x "𝕏")) n)) e))
-}
record MR𝕏 : Setwhere
   field Scalar𝕏        : Set
   field Vector𝕏        : Set
   field _·𝕏_       : Scalar𝕏 → Vector𝕏 → Vector𝕏
   field 𝟙𝕏     : Scalar𝕏
   field _×𝕏_       : Scalar𝕏 → Scalar𝕏 → Scalar𝕏
   field leftId𝕏        : {𝓋 : Vector𝕏}  →  𝟙𝕏 ·𝕏 𝓋  ≡  𝓋
   field assoc𝕏     : {a b : Scalar𝕏} {𝓋 : Vector𝕏} → (a ×𝕏 b) ·𝕏 𝓋  ≡  a ·𝕏 (b ·𝕏 𝓋)

Now for some useful corollaries.

(𝒱 rename f (support-mixfix-names t) (adjoin-retract t)
  =  "Rename elements using a string-to-string function F acting on names.

      There is minimal support for mixfix names, which may be ignored
      by setting SUPPORT-MIXFIX-NAMES to be nil.

      When ADJOIN-RETRACT is non-nil, we adjoin a “record {oldᵢ = nameᵢ}”
      view morphism; i.e., record translation.
     "
     map (λ e → (map-name (λ n → (rename-mixfix f n (not 'support-mixfix-names))) e))
         :support-mixfix-names 'support-mixfix-names
         :adjoin-retract 'adjoin-retract)

Let's try this out.

{-700
MR𝕪    = M-Set-Record rename (λ n → (concat n "𝕪"))
MR-oh  = M-Set-Record rename (λ n → (pcase n ("Scalar" "S") ("𝟙" "ε") (else else)))
-}
record MR𝕪 : Setwhere
   field Scalar𝕪        : Set
   field Vector𝕪        : Set
   field _·𝕪_       : Scalar𝕪 → Vector𝕪 → Vector𝕪
   field 𝟙𝕪     : Scalar𝕪
   field _×𝕪_       : Scalar𝕪 → Scalar𝕪 → Scalar𝕪
   field leftId𝕪        : {𝓋 : Vector𝕪}  →  𝟙𝕪 ·𝕪 𝓋  ≡  𝓋
   field assoc𝕪     : {a b : Scalar𝕪} {𝓋 : Vector𝕪} → (a ×𝕪 b) ·𝕪 𝓋  ≡  a ·𝕪 (b ·𝕪 𝓋)


{- MR-oh  = M-Set record ⟴ rename (λ n → (pcase n ("Scalar" "S") ("𝟙" "ε") (else else))) -}
record MR-oh : Set₁ where
   field S      : Set
   field Vector     : Set
   field _·_        : S → Vector → Vector
   field ε      : S
   field _×_        : S → S → S
   field leftId     : {𝓋 : Vector}  →  ε · 𝓋  ≡  𝓋
   field assoc      : {a b : S} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)

We would expect it to be common to prefix and suffix symbols, so let's make variationals for these patterns.

(𝒱 decorated by
  = "Rename all elements by suffixing string BY to them."
     rename (λ name → (concat name by)))

(𝒱 co-decorated by
  = "Rename all elements by prefixing string BY to them."
     rename (λ name → (concat by name)))

(𝒱 primed
  = "All elements are renamed with a postfix prime."
    decorated "′")

Likewise, for the casing approach, let's make a “to list”. For now, such lists are necessarily enclosed in double-quotes.

  • Given a string of “;”-separated items consisting of “to”-separated pairs, interpret it as a Lisp function where “to”-pairs denote mapping clauses.
  • E.g., “x₀ to y₀; …; xₙ to yₙ” becomes the function sending value xᵢ to yᵢ, and behaves as the identity function otherwise unless OTHERWISE is provided, in which case it acts as a fallback.
  • Concretely:

          (reify-to-list "1 to x; 2 to y; p to q")(λ arg → (pcase arg ("1" "x") ("2" "y") ("p" "q") (otherwise otherwise)))
    
;; Neato: (reify-to-list "x₀; ⋯; xₙ" nil) ⇒ (λ x ↦ If ∃ i • x ≈ xᵢ then "" else nil)
;; KEY is a function applied to the input argument /before/ casing on LHS ↦ RHS names.
;; INVERSE means to interpret clauses “x to y” as mappings “y ↦ x”.
   (cl-defun reify-to-list (str &key (otherwise 'otherwise) (key #'identity) inverse)
   "Transform “to list” STR with default OTHERWISE into a Lisp function.
<<docs('reify-to-list)>>"
   (let (clauses)
     (thread-last str
       (s-split ";")
       (--map (s-split " to " it))
       (--map (list (s-trim (car it)) (s-trim (or (cadr it) "")))) ;; accomodate empty str.
       (funcall (λ cs → (if inverse (--map (-rotate 1 it) cs) cs)))
       (-cons* 'pcase `(,key arg))
       (setq clauses))
     `(lambda (arg) ,(append clauses `((otherwise ,otherwise))))))

;; (reify-to-list "a to b; c to d" :inverse t) ;; neato!


(𝒱 renaming by (adjoin-retract t) (adjoin-coretract nil)
  = "Rename elements using BY, a “;”-separated string of “to”-separated pairs.

      There is minimal support for mixfix names, which may be ignored
      by setting SUPPORT-MIXFIX-NAMES to be nil.

      When ADJOIN-RETRACT is non-nil, we adjoin a “record {oldᵢ = nameᵢ}”
      view morphism; i.e., record translation.
      Likewise for ADJOIN-CORETRACT results in the inverse morphism,
      “record {nameᵢ = oldᵢ}”.
"
  map (λ e → (map-name (λ n → (funcall (reify-to-list by) n)) e))
         :adjoin-retract adjoin-retract
         :adjoin-coretract adjoin-coretract
         )
𝒱-renaming

It is common in Agda to provide “to”-lists, so we've provide a variant that supports those instead of forcing users to produce functions explicitly.

We may also prefer writing subscriptedᵢ rather than decorated "ᵢ". With a bit of Lisp meta-programming, we can generate these variationals.

(defun to-subscript (n)
  "Associate digit N with its subscript.

If N ∈ 0..9, then yield ₙ, else N."

  (or (nth n '("₀" "₁" "₂" "₃" "₄" "₅" "₆" "₇" "₈" "₉")) n))

;; Let's make a family of variationals.
(loop for i from 0 to 9
      for ᵢ    = (to-subscript i)
      for docs = (format "Subscript all elementes by suffixing them with %s." i)
      do (eval `(𝒱 ,(intern (format "subscripted%s")) = ,docs decorated ,ᵢ)))

Here are some example uses.

{-700
-- MR₁₂   = M-Set-Record decorated "₁" ⟴ decorated "₂" :adjoin-retract nil
the-MR = M-Set-Record co-decorated "the-"
-- MR₃₄   = M-Set-Record subscripted₃ ⟴ subscripted₄ :adjoin-retract nil
MRₜₒ   = M-Set-Record renaming "Scalar to S; Vector to V; · to nice"
NearMonoid = M-Set-Record renaming "Scalar to Carrier; Vector to Carrier; · to ×"
-}
{- MR₁₂   = M-Set record ⟴ decorated "₁" ⟴ decorated "₂" -}
record MR₁₂ : Set₁ where
   field Scalar₁₂       : Set
   field Vector₁₂       : Set
   field _·₁₂_      : Scalar₁₂ → Vector₁₂ → Vector₁₂
   field 𝟙₁₂        : Scalar₁₂
   field _×₁₂_      : Scalar₁₂ → Scalar₁₂ → Scalar₁₂
   field leftId₁₂       : {𝓋 : Vector₁₂}  →  𝟙₁₂ ·₁₂ 𝓋  ≡  𝓋
   field assoc₁₂        : {a b : Scalar₁₂} {𝓋 : Vector₁₂} → (a ×₁₂ b) ·₁₂ 𝓋  ≡  a ·₁₂ (b ·₁₂ 𝓋)

{- the-MR = M-Set recordco-decorated "the-" -}
record the-MR : Setwhere
   field the-Scalar     : Set
   field the-Vector     : Set
   field _the-·_        : the-Scalarthe-Vectorthe-Vector
   field the-𝟙      : the-Scalar
   field _the-×_        : the-Scalarthe-Scalarthe-Scalar
   field the-leftId     : {𝓋 : the-Vector}  →  the-𝟙 the-· 𝓋  ≡  𝓋
   field the-assoc      : {a b : the-Scalar} {𝓋 : the-Vector} → (a the-× b) the-· 𝓋  ≡  a the-· (b the-· 𝓋)

{- MR₃₄   = M-Set record ⟴ subscripted₃ ⟴ subscripted₄ -}
record MR₃₄ : Set₁ where
   field Scalar₃₄       : Set
   field Vector₃₄       : Set
   field _·₃₄_      : Scalar₃₄ → Vector₃₄ → Vector₃₄
   field 𝟙₃₄        : Scalar₃₄
   field _×₃₄_      : Scalar₃₄ → Scalar₃₄ → Scalar₃₄
   field leftId₃₄       : {𝓋 : Vector₃₄}  →  𝟙₃₄ ·₃₄ 𝓋  ≡  𝓋
   field assoc₃₄        : {a b : Scalar₃₄} {𝓋 : Vector₃₄} → (a ×₃₄ b) ·₃₄ 𝓋  ≡  a ·₃₄ (b ·₃₄ 𝓋)

{- MRₜₒ = M-Set recordrenaming "Scalar to S; Vector to V; · to nice" -}
record MRₜₒ : Setwhere
   field S      : Set
   field V      : Set
   field _nice_     : S V V
   field 𝟙      : S
   field _×_        : S S S
   field leftId     : {𝓋 : V}  →  𝟙 nice 𝓋  ≡  𝓋
   field assoc      : {a b : S} {𝓋 : V} → (a × b) nice 𝓋  ≡  a nice (b nice 𝓋)

{- NearMonoid = M-Set record ⟴ renaming "Scalar to Carrier; Vector to Carrier; · to ×" -}
record NearMonoid : Set₁ where
   field Carrier        : Set
   field _×_        : Carrier → Carrier → Carrier
   field 𝟙      : Carrier
   field leftId     : {𝓋 : Carrier}  →  𝟙 × 𝓋  ≡  𝓋
   field assoc      : {a b : Carrier} {𝓋 : Carrier} → (a × b) × 𝓋  ≡  a × (b × 𝓋)

Some observations are in order:

  1. Example M₁₂ demonstrates that composition, ⟴, is sequential from left to right. That is, “⟴” is just forwards composition: We thread the given PackageFormer through the variationals vᵢ in order. Operationally:

    Pf v₀ ⟴ ⋯ ⟴ vₙ ≈ ((Pf v₀) v₁) ⋯) vₙ
    Pf ⟴ v ≈ Pf v ⟴ ≈ Pf v

    Note: In the concrete syntax, parenthesisation is not permitted.

  2. Notice that the NearMonoid example demonstrates multiplicity of PackageFormer elements is irrelevant. That is, elements are algebraically a free monoid with the axiom xs ⊕ ys ⊕ xs ≈ xs ⊕ ys.
  3. Notice that we wanted Agda-style renaming via to-lists, so we simply code that up! This is so cool: We can just extend the system with whatever pattern we prefer! No more bending to the will of language designers! More power to the user!

    For example, we can codify the previous NearMonoid scheme into a top-level pattern.

(defun is-sort (element)
  "Check whether the target of ELEMENT’s type is ‘Set’."
  (s-contains? "Set" (target (element-type element))))
  ;; Method ‘target’ is defined in the next subsection, on ADTs.

(𝒱 single-sorted with-sort
  = "Replace all nullary sorts with the provided WITH-SORT string
     as the name of the new single sort, the universe of discourse."
    map (λ e → (if (is-sort e) (map-name (λ _ → with-sort) e) e)))

Then the previous PackageFormer can be obtained with: Note that the following differs from NearMonoid since it has two binary operations: Our new variational one alters the number and name of sorts, not other elements.

{-700
NearMonoid¹ = M-Set-Record single-sorted "Carrier"
-}
record NearMonoid¹ : Setwhere
   field Carrier        : Set
   field _·_        : Carrier Carrier Carrier
   field 𝟙      : Carrier
   field _×_        : Carrier Carrier Carrier
   field leftId     : {𝓋 : Carrier}  →  𝟙 · 𝓋  ≡  𝓋
   field assoc      : {a b : Carrier} {𝓋 : Carrier} → (a × b) · 𝓋  ≡  a · (b · 𝓋)

Exercise: Write a variational remove-sorts that strips out all sorts from a PackageFormer. If elements depend on sorts, as they normally do, then one must remove them as well; ignore this for now, and we shall return to subgenerated PackageFormers later on.

4.12 Forming Syntax and the Special $𝑛𝑎𝑚𝑒 Variable

Records provide a semantics, what if we want the syntax?

Since data declarations consist of constructors, whose target type necessarily begins with the name of the data-type being defined, let's only keep those fields and drop the rest. To do so, we use the helper function target which takes a declaration name : type0 → ⋯ → typeN and yields typeN.

(defun target (thing)
  "Return final type mentioned in THING, a string declaration.

Given a type-name ‘[name :] τ₀ → ⋯ → τₙ’, yield ‘τₙ’;
the ‘name’ porition is irrelevant."
  (car (-take-last 1 (s-split "→" thing))))

With this in hand, a data presentation requires a designated carrier which is used to keep only those elements that target it. Finally, as data constructor must target the type being defined, we alter the filtered elements by changing every instance of the carrier name with the name of the newly defined PackageFormer —which we may access using the special identifier $𝑛𝑎𝑚𝑒. In a lisp block, we formalise this algorithm as follows.

(𝒱 data carrier
  = "Reify as an Agda “data” type.

     Only elements targeting CARRIER are kept.
    "
    :kind  data
    :level dec
    :alter-elements (lambda (es)
      (thread-last es
        (--filter (s-contains? carrier (target (element-type it))))
        (--map (map-type (λ τ → (s-replace carrier $𝑛𝑎𝑚𝑒 τ)) it)))))

For example:

{-   700
ScalarTerm = M-Set data "Scalar"
-}
data ScalarTerm : Set where
   𝟙        : ScalarTerm
   _×_      : ScalarTerm ScalarTerm ScalarTerm

Again: The meta-primitive :alter-elements is instructed to map over those elements e that contain the carrier in their target type by replacing the given carrier with the newly-minted $𝑛𝑎𝑚𝑒 of the grouping mechanism being constructed. Those that do not contain the given carrier in their target type are filtered out.

Notice that $𝑛𝑎𝑚𝑒 is a special variable that refers to the newly defined PackageFormer's name.

  • It is written using \Mi with Agda input; e.g., \Min gives 𝑛.
  • The ‘$’ is intended to further mark the special nature of this variable.

4.13 Subpackages with generated, sorts, signature

A common grouping operation is to zoom-in to the minimal well-formed package that contains only certain specified elements. For example, in our M-Set grouping, we may want to keep only 𝟙 but to be well-defined we are forced to also keep the elements on which it depends —namely, Scalar.

In particular, the following naive approach only works if the elements are independent of one another —which is rarely the case for Agda users.

;; cute, but too brutish.
(𝒱 generated by = :alter-elements (lambda (es) (-filter by es)))

The coherent scheme is straightforward to implement. For clarity, rather than efficiency, the algorithm below forms a list yeses of the elements that should be kept then traverses the elements list, adding all elements needed to ensure that list is coherent. Moreover, for generality, we consider a predicate rather than an explicit listing of items to be retained.

(𝒱 generated by
  = "Keep the largest well-formed PackageFormer whose elements satisfy BY.

     BY is a predicate on elements.
    "
    :alter-elements (λ es → (--graph-map (funcall by it) it es)))

Here's an immediate application: Obtaining the types declared in a grouping mechanism.

(𝒱 sorts
 = "Obtaining the types declared in a grouping mechanism.

   For now, only base types; i.e., items targeting “Set”.
   "
   generated (λ e → (s-contains? "Set" (target (element-type e)))))
{-700
M-Set-Sorts = M-Set record ⟴ sorts
-}
record M-Set-Sorts : Setwhere
   field Scalar     : Set
   field Vector     : Set

We can even obtain a sub-signature wholesale:

{-700
MonoidSignature = M-Set-Record generated (λ e → (and (s-contains? "Scalar" (element-type e)) (not (s-contains? "Vector" (element-type e)))))
-}
record MonoidSignature : Setwhere
   field Scalar     : Set
   field 𝟙      : Scalar
   field _×_        : Scalar Scalar Scalar

This pattern of having a lawless grouping seems sufficiently desirable that we may codify it.

(defun targets-a-sort (element)
  "Check whether the given ELEMENT targets a sort.

The sorts considered refer to those of the *current* PacakgeFormer."
  (--any (s-contains? it (target (element-type element)))
         (-map #'element-name (-filter #'is-sort $𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠))))

(𝒱 signature
  = "Keep only the elements that target a sort, drop all else."
    generated (λ e → (targets-a-sort e)))

Here's an example.

{-700
MonSig = M-Set-Record signature
-}
record MonSig : Setwhere
   field Scalar     : Set
   field Vector     : Set
   field _·_        : Scalar Vector Vector
   field 𝟙      : Scalar
   field _×_        : Scalar Scalar Scalar

Neato! Those were some nifty applications!

For practicality, let's also introduce a more concrete syntax analogous to that of renaming:

(𝒱 keeping those
  = "Keep THOSE elements, a “;”-separated string of proper names,
    along with the elements that ensure THOSE is well-defined.
 "
    generated (reify-to-list those :otherwise nil :key #'element-name))

4.14 Shallow Renaming with Agda's open ⋯ public ⋯ renaming ⋯

The previous approach to renaming altered field names literally which is not desirable when one only wants to refer to field names of multiple instances of the same record —e.g., when forming homomorphisms.

A common pattern in Agda is then to open the record and perform the desired shallow renames. This pattern is so common that the standard library is littered with instances of it. We can codify the pattern as a method rather than as a manual technique.

Let's go from zero to one-hundred —again: There's a Lisp Cheat Sheet that should have been consulted at one point.

Zero: A module where the elements are all parameters.

{-700
𝒱-empty-module = :kind module :level none :waist 999
Neato = M-Set empty-module
-}

{- A module where the elements are all parameters -}
open Neato using ()
module Neato (Scalar : Set) (Vector : Set) (_·_ : Scalar Vector Vector) (𝟙 : Scalar) (_×_ : Scalar Scalar Scalar) (leftId : {𝓋 : Vector} → 𝟙 · 𝓋 ≡ 𝓋) (assoc : ∀ {a b 𝓋} → (a × b) · 𝓋 ≡ a · (b · 𝓋)) where

One-hundred: A one-parameter module where elements may be renamed.

(𝒱 open with (avoid-mixfix-renaming nil)
  =
    "Reify a given PackageFormer as a *parameterised* Agda “module” declaration.

     WITH is a renaming, string to string, function that is applied to the parent record that will
     be opened and reexported as a module.

     AVOID-MIXFIX-RENAMING is optional; by default renaming “jumps over” underscores,
     but providing a non-nil value for this argument leaves underscores alone.
     It is a matter of having, say, default “_⊕ₙ_” versus “_⊕_ₙ”.

     The resulting module has a parameter, whose name is irrelevant but is
     of the form “Arg𝒹𝒹𝒹𝒹” for some digits 𝒹 in order to minimise clash with
     any user-defined names.

     Besides the addition of a new parameter, all element qualifiers are discarded.
    "
    :kind module
    :level none
    :waist 1
    :alter-elements  (lambda (fs)
      (let ((kind "{! !}") ((format "Ar%s" (gensym))))
        (cons (make-element :name:type $𝑝𝑎𝑟𝑒𝑛𝑡)
          (--map (let ((name (if avoid-mixfix-renaming (with (element-name it)) (rename-mixfix with (element-name it)))))
            (make-element :name name
                          :type (format "let open %s %s in %s" $𝑝𝑎𝑟𝑒𝑛𝑡 ℛ (element-type it))
                          :equations (list (format "%s = %s.%s %s" name $𝑝𝑎𝑟𝑒𝑛𝑡 (element-name it))))) fs)))))

Notice that we do not need any open ⋯ public since all elements are top-level. We are not making using of Agda's renaming facility. An example may clarify this observation.

{-700
M-Set-R = M-Set record
M-Set-R₁ = M-Set-R ⟴ open (λ x → (concat x "₁"))
-}
module M-Set-R (Arg6926 : M-Set-R) where
   Scalar₁      : let open M-Set-R Arg6926 in Set ; Scalar= M-Set-R.Scalar Arg6926
   Vector₁      : let open M-Set-R Arg6926 in Set ; Vector= M-Set-R.Vector Arg6926
   _·₁_     : let open M-Set-R Arg6926 in Scalar Vector Vector ;    _·₁_ = M-Set-R._·_ Arg6926
   𝟙₁       : let open M-Set-R Arg6926 in Scalar ;  𝟙₁ = M-Set-R.𝟙 Arg6926
   _×₁_     : let open M-Set-R Arg6926 in Scalar Scalar Scalar ;    _×₁_ = M-Set-R._×_ Arg6926
   leftId₁      : let open M-Set-R Arg6926 in {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 ;    leftId= M-Set-R.leftId Arg6926
   assoc₁       : let open M-Set-R Arg6926 in ∀ {a b 𝓋} → (a × b) · 𝓋  ≡  a · (b · 𝓋) ; assoc= M-Set-R.assoc Arg6926

In-case you've skipped over the above source documentation for open, it's time to read it.

Notice that a module opening depends on a record, whence the first declaration of M-Set-R.

These kind of open-renamings are so common that the tedium is actually acceptable by most users —it shouldn't be and now it no longer has to be that way.

It is common in Agda to provide “to”-lists, so let's provide a variant that supports those instead of forcing users to produce functions explicitly.

(𝒱 opening with
  = "Open a record as a module exposing only the names mentioned in WITH.

    WITH is a string of “;”-separated items consisting of “to”-separated pairs.
    "
    open (λ x → (funcall (reify-to-list with :otherwise "_") x)) :avoid-mixfix-renaming t)

    ;; Alternatively, we could have used ‘trash’ names,
    ;; something like (format "%s" (gensym)), instead of "_".
{-700
M-Set-R-SV = M-Set-R opening "Scalar to S; Vector to V"
-}

This opens the M-Set-R record exposing only S and V —the rest are ignored using Agda's _ mechanism.

module M-Set-R-SV (Arg6933 : M-Set-R) where
   S        : let open M-Set-R Arg6933 in Set ; S = M-Set-R.Scalar Arg6933
   V        : let open M-Set-R Arg6933 in Set ; V = M-Set-R.Vector Arg6933
   _        : let open M-Set-R Arg6933 in Scalar Vector Vector ;    _ = M-Set-R._·_ Arg6933
   _        : let open M-Set-R Arg6933 in Scalar ;  _ = M-Set-R.𝟙 Arg6933
   _        : let open M-Set-R Arg6933 in Scalar Scalar Scalar ;    _ = M-Set-R._×_ Arg6933
   _        : let open M-Set-R Arg6933 in {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 ;    _ = M-Set-R.leftId Arg6933
   _        : let open M-Set-R Arg6933 in ∀ {a b 𝓋} → (a × b) · 𝓋  ≡  a · (b · 𝓋) ; _ = M-Set-R.assoc Arg6933

After simplifying the let-expressions, this module definition is equivalent to the following —the types of which may be seen with Agda's C-c C-o call.

module M-Set-R-SV (Arg : M-Set-R) where
  S : (Arg : M-Set-R) → Set ; S = M-Set-R.Scalar Arg
  V : (Arg : M-Set-R) → Set ; V = M-Set-R.Vector Arg

Let's provide an even more common feature: Opening records with a decoration. For example, when we have two algebraic structures, we might want the first to be subscripted with ₁ and the second with ₂ —this is different than subscriptedᵢ from above, which produces a new record rather than opening it with renames.

(𝒱 open-with-decoration ddd
  = "Open a record, exposing all elements, with decoration DDD.

    DDD is a string.
   "
   open (λ x → (concat x ddd)))

Here's an example.

{-700
M-Set-R′ = M-Set-R open-with-decoration "′"
-}
module M-Set-R (Arg6938 : M-Set-R) where
   Scalar′      : let open M-Set-R Arg6938 in Set ; Scalar= M-Set-R.Scalar Arg6938
   Vector′      : let open M-Set-R Arg6938 in Set ; Vector= M-Set-R.Vector Arg6938
   _·′_     : let open M-Set-R Arg6938 in Scalar Vector Vector ;    _·′_ = M-Set-R._·_ Arg6938
   𝟙′       : let open M-Set-R Arg6938 in Scalar ;  𝟙′ = M-Set-R.𝟙 Arg6938
   _×′_     : let open M-Set-R Arg6938 in Scalar Scalar Scalar ;    _×′_ = M-Set-R._×_ Arg6938
   leftId′      : let open M-Set-R Arg6938 in {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 ;    leftId= M-Set-R.leftId Arg6938
   assoc′       : let open M-Set-R Arg6938 in ∀ {a b 𝓋} → (a × b) · 𝓋  ≡  a · (b · 𝓋) ; assoc= M-Set-R.assoc Arg6938

Neato petito :smile:

It is important to observe that ‘openings’ are lossy: They lose the types of the declarations and so cannot be used further to construct new pacaking mechanisms. They are a terminal construction.

In the next section, we make use of such openings to actually produce homomorphism constructions.

4.15 Automatically deriving homomorphism definitions ♥‿♥

The definition of “structure preservation” is, nearly always, mechanical to formulate and that's just what we shall do to avoid having to write it out by hand ever again —which the current approach in the Agda standard library.

The idea is not too complicated:

  1. Suppose you have an operation _·_ : Scalar → Vector → Vector.
  2. Suppose you have a numbering of the sorts; e.g., sort₁ = Scalar, sort₂ = Vector.
  3. Form functions mapᵢ : sortᵢ → sortᵢ′
  4. Include implicit arguments in the type: {x₁ : Scalar} → {x₂ : Vector} → Vector.
  5. The target type Vector = sort₂ means we need to apply map₂ to the expression formed from the operation's name along with the arguments.
    • The left hand side is thus map₂ (_·_ x₁ x₂).
  6. For the right hand side, we use the target-space's name, say _·′_, along with mapᵢ applied to xᵢ for each i mentioned in the type.
  7. The result: pres-· : {x₁ : Scalar} → {x₂ : Vector} → map₂ (_·_ x₁ x₂) ≡ _·′_ (map₁ x₁) (map₂ x₂).

First, we need a helper that forms the preservation formulae. For example:

(show-element (homify (make-element :name "_·_" :type "Scalar → Vector → Vector")
                      '( ("Scalar" . 4) ("Vector" . 1))))
⇒
  pres-· : {x₄ : Scalar} → {x₁ : Vector}
         → map₁ (_·_ x₄ x₁)   ≡   _·′_ (map₄ x₄) (map₁ x₁)

With this as a specification, in a lisp block:

(defun homify (element sort)
  "Given a typed name, produce the associating “preservation” formula.

E.g.,
  _·_    : Scalar → Vector → Vector
  pres-· : {x₁ : Scalar} → {x₂ : Vector} → map₂ (x₁ · x₂) = map₁ x₁ ·′ map₂ x₂


Type τ gets variable xᵢ provided (i, τ) ∈ SORT;
likewise we think of mapᵢ : τ → τ′.
Notice that the target name is primed, “·′”

ELEMENT is the typed-name and SORT is the alist of numbered sorts."
  (letf* ((sorts     (mapcar #'car sort))
          (index     (λ it → (to-subscript (cdr (assoc it sort)))))

          (tn→       (s-split " → " (element-type element)))
          (arg-count (1- (length tn→)))

          (all-indicies  (mapcar index
                                 (--filter (member (s-trim it) sorts) tn→)))
          (indicies  (-drop-last 1 all-indicies))
          (tgt-idx   (car (-take-last 1 all-indicies)))

          (op        (element-name element))
          (args      (--map (concat "x" it) indicies))
          (lhs       (format "map%s (%s %s)" tgt-idx op (s-join " " args)))

          (op′       (rename-mixfix (lambda (n) (concat n "′")) op))
          (map-args  (--map (format "(map%s x%s)" it it) indicies))
          (rhs       (format "%s %s" op′ (s-join " " map-args)))

          (target    (format "  %s   ≡   %s" lhs rhs)))

    ;; Change the target type.
    (setq tn→ (--map (when (assoc it sort)
                       (format "{x%s : %s}" (funcall index it) it)) tn→))
    (setf (nth arg-count tn→) target)

    ;; Stick it all together, with an updated name.
    (make-element
     :name (format "pres-%s" (s-replace "_" "" (element-name element)))
     :type (s-join " → " tn→))))

Then, we form the variational as follows —also in a lisp block.

(𝒱 hom
  = "Formulate the notion of homomorphism of $𝑝𝑎𝑟𝑒𝑛𝑡 algebras.

     ➩ $𝑝𝑎𝑟𝑒𝑛𝑡 must be an existing record type used in the resulting formulation.
    "
    record:waist 2
    :alter-elements (lambda (es)

      (let (maps eqns sorts (𝒮𝓇𝒸 "Src") (𝒯ℊ𝓉 "Tgt"))

        ;; Construct the mapᵢ : sortᵢ → sortᵢ′; keeping track of (sort . i) pairs.
        (loop for e in es
              for i from 1
         do
           (when (is-sort e)
             (push (cons (element-name e) i) sorts)
             (push (make-element
                      :qualifier "field"
                      :name (format "map%s" (to-subscript i))
                      :type (format "%s → %s′" (element-name e) (element-name e)))
                   maps))

            (when (and (targets-a-sort e) (not (is-sort e)))
              (push (homify e sorts) eqns)))

      ;; Ensure we have a source and target space as elements.
      (-cons*
       (make-element :qualifier "field" :name 𝒮𝓇𝒸 :type $𝑝𝑎𝑟𝑒𝑛𝑡)
       (make-element :qualifier "field" :name 𝒯ℊ𝓉 :type $𝑝𝑎𝑟𝑒𝑛𝑡)
       (--map
        (map-type (λ τ → (format "let open %s %s; open %s′ %s in %s"
                                 $𝑝𝑎𝑟𝑒𝑛𝑡 𝒮𝓇𝒸 $𝑝𝑎𝑟𝑒𝑛𝑡 𝒯ℊ𝓉 τ))
                  (map-qualifier (λ _ → "field") it))
        (reverse (-concat eqns maps)))))))

Here are two examples. Note that the latter allows us to rename the mapᵢ as we wish —which may be preferable to extending the variational to accommodate for new names.

{-700
Algebra  = M-Set record
Algebra′ = Algebra open-with-decoration "′"
Hom  = Algebra hom
Hom² = Algebra hom ⟴ renaming "map₁ to scalar; pres-𝟙 to unity" :adjoin-retract nil
-}

_ : {Src Tgt : Algebra} → Hom² Src Tgt → Algebra.Scalar Src → Algebra.Scalar Tgt
_ = Hom².scalar
{- Hom  = Algebra hom -}
record Hom (Src : Algebra) (Tgt : Algebra) : Set₁ where
   field map₁       : let open Algebra Src; open Algebra′ Tgt in Scalar → Scalar′
   field map₂       : let open Algebra Src; open Algebra′ Tgt in Vector → Vector′
   field pres-·     : let open Algebra Src; open Algebra′ Tgt in {x₁ : Scalar} → {x₂ : Vector} →   map₂ (_·_ x₁ x₂)   ≡   _·′_ (map₁ x₁) (map₂ x₂)
   field pres-𝟙     : let open Algebra Src; open Algebra′ Tgt in   map₁ (𝟙 )   ≡   𝟙′
   field pres-×     : let open Algebra Src; open Algebra′ Tgt in {x₁ : Scalar} → {x₁ : Scalar} →   map₁ (_×_ x₁ x₁)   ≡   _×′_ (map₁ x₁) (map₁ x₁)


{- Hom² = Algebra homrenaming "map₁ to scalar; pres-𝟙 to unity" -}
record Hom² (Src : Algebra) (Tgt : Algebra) : Setwhere
   field scalar     : let open Algebra Src; open Algebra Tgt in Scalar Scalarfield map₂       : let open Algebra Src; open Algebra Tgt in Vector Vectorfield pres-·     : let open Algebra Src; open Algebra Tgt in {x₁ : Scalar} → {x₂ : Vector} →   map₂ (_·_ xx₂)   ≡   _·′_ (scalar x₁) (mapx₂)
   field unity      : let open Algebra Src; open Algebra Tgt in   scalar (𝟙 )   ≡   𝟙′
   field pres-×     : let open Algebra Src; open Algebra Tgt in {x₁ : Scalar} → {x₁ : Scalar} →   scalar (_×_ xx₁)   ≡   _×′_ (scalar x₁) (scalar x₁)

This is so cool (•̀ᴗ•́)و

We leave it to the reader to derive other constructs from a theory presentation. Examples can be found in these Metaprogramming Agda slides: Homomorphism equality, application to carrier elements, isomorphisms, isomorphisms where only one direction needs to preserve the structure and an automatically derivable proof that the other direction is also structure preserving, endomorphism and automorphism types, kernels, product & sum & other categorical types.

Challenge
Design a scheme to produce simple Cartesian products from a given theory.
  1. The only variable to this problem is an arbitrary record, say it is M.

    For this exercise to be tractable, assume M consists of declarations of sort symbols, function symbols, and nullary (non-implication) equations which may have implicit arguments.

  2. Ensure you understood the definition of the homomorphism scheme above.
  3. Mimic the homomorphism scheme to produce a typed Prod where Prod A₀ A₁ consists of a M value, say P, and two homomorphisms Hom P Aᵢ.
  4. Write a Lisp code that produces a function MakeProduct : (A₀ A₁ : M) → Prod A₀ A₁.
    • The projection morphisms are straightforward.
    • Every n-ary function f could be defined by fₚ = zipₙ f₀ f₁.
    • Every equation e could be defined by eₚ = cong₂ _,_ e₀ e₁.
  5. If you have actually attempted this, then go on to include the remaining artefacts to make the construction an actual categorical product.

5 Strings and Things

Since our prototype is intended to be as minimally obtrusive as possible, we will need to extract our special 700-syntactical items between delimited tokens and process them.

The following subsections introduce:

get-children
Obtaining intended items from a hierarchical listing.
sub-string-delimited-here
Finding the shortest substring between a prefix 𝑳 and a postfix 𝑹 by using the ‘metavariable’ $here; e.g., “𝑳 $here 𝑹”.
buffer-substring-delimited-whole-buffer
Yield all portions of the buffer enclosed in the given delimiters, as a list of strings.
rename-mixfix
Renamaing Agda mixifix names where the rename operation ignores the outermost (Agda) argument position markers ‘_’.
extract-imports
Obtain all ‘import’ clauses so that they can be ported over to the generated file.

First we include an Emacs Lisp library header which declares important meta-data —with the aim of having the resulting software be distributed with Emacs's largest package repository, MELPA. Installing this package, from within Emacs, also automatically installs its dependencies, recursively.

;;; agda-next-700-module-systems.el --- Making Modules with Meta-Programmed Meta-Primitives, for Agda

;; Author: Musa Al-hassy <alhassy@gmail.com>
;; Version: 1.0
;; Package-Requires: ((s "1.12.0") (dash "2.16.0") (origami "1.0")  (emacs "24.4"))
;; Keywords: agda, modules, packages, theories, languages, convenience, maint, tools
;; URL: https://alhassy.github.io/next-700-module-systems

;; Copyright (c) 2019 Musa Al-hassy

;; This program is free software; you can redistribute it and/or modify
;; it under the terms of the GNU General Public License as published by
;; the Free Software Foundation, either version 3 of the License, or
;; (at your option) any later version.

;; This program is distributed in the hope that it will be useful,
;; but WITHOUT ANY WARRANTY; without even the implied warranty of
;; MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
;; GNU General Public License for more details.

;; You should have received a copy of the GNU General Public License
;; along with this program.  If not, see <http://www.gnu.org/licenses/>.

;;; Commentary:

;; This program is intended to reduce the burden in selecting
;; the form of records and other kinds of packages in Agda.
;; For example, the decision of whether a record element should be
;; declared as a field or as a parameter no longer needs to be performed
;; prematurely but rather may be selected when necessary.

;; We cannot use lexical binding since the package provides a macro called 𝒱
;; that acts as an interface for a DSL that generates Lisp functions
;; which requires EVAL to bind staged formal arguments to their runtime values.

;; This el file has been tangled from a literate, org-mode, file.
;; See the documentation on:
;; https://alhassy.github.io/next-700-module-systems/prototype/package-former.html

;;; Code:

First some useful libraries:

;; String and list manipulation libraries
;; https://github.com/magnars/dash.el
;; https://github.com/magnars/s.el
(require 's)               ;; “The long lost Emacs string manipulation library”
(require 'dash)            ;; “A modern list library for Emacs”
(require 'dash-functional) ;; Function library; ‘-const’, ‘-compose’, ‘-orfn’, ‘-not’, ‘-partial’, etc.
(require 'origami)         ;; Folding away regions of text
(require 'subr-x)          ;; Extra Lisp functions; e.g., ‘when-let’.
(require 'ert)             ;; Testing framework; ‘should’ for assertions
(require 'cl-lib)          ;; New Common Lisp library; ‘cl-???’ forms.

5.1 Finding Children in the Wild

Being a prototype, we are talking a mostly string-based approach to working with hierarchical phrases. For example, consider the following todo list,

+ item 0
+ item 1
  - subitem 1.1
    * subsubitem 1.1.1
  - subitem 1.2
+ item 2
  - subitem 2.2
+ item 3
  • item 0
  • item 1
    • subitem 1.1
      • subsubitem 1.1.1
    • subitem 1.2
  • item 2
    • subitem 2.2
  • item 3

We would think that item 1 has two ‘children’, and, moreover, one grand-child. Whereas item 2 has a single child and item 3 is barren.

Here's my intuitive algorithm: We obtain the indentation of the first child, then all subsequent lines with at least that much indentation have the same ancestor.

get-indentation Implementation
;; (pf--declare-type pf--get-indentation : string integer)
(defun pf--get-indentation (string)
  "How many spaces are there at the front of STRING?

Property: The resulting number is ≤ length STRING."
  (if string (length (s-shared-start string (s-repeat (length string) " "))) 0))
get-children Implementation

Go into ‘the-wild’ seeking out the first occurence of ‘parent’, who once found, ought to have a minimal indentation for its children.

“Minimal” in that if there are items with a greater indentation, then they are children of children and should be kept.

The first input argument is of type ‘string’, the second argument may be of type ‘string’ or ‘list’ of strings —if it's a string, we split along new lines—, the optional ‘then’ is a function acting on children strings.

Result is the parent followed by its children, as a list of lines, where each child has been altered using the optional THEN function. Moreover, we also return the rest of the unconsidered portion of THE-WILD:

Result list: (unconsidered-prefix-of-the-wild (cons parent-line children-lines) unconsidered-remaining-lines)

The first element is the porition that does not contain an occurence of PARENT. The second is the parent and its children, if possible. The third is the remainder of THE-WILD.

Implementation: Look at the indentation of the first child, then use that as a lower bound to find the indentation of the remaining children.

(cl-defun pf--get-children (parent the-wild &key (then #'identity))
  "Returns indented text under a PARENT line.
<<docs('get-children)>>"
  (let ((lines (if (stringp the-wild) (s-lines the-wild) the-wild))
        (indentation -1)
        unconsidered
        prefix
        lines&more
        parent-line)

    ;; Ensure: lines ≈ (cons (not-here-prefix) (cons parent-here more-lines) )
    (setq lines (--split-with (not (s-contains? parent it)) lines))

    ;; Discard prefix, for now.
    (setq prefix (car lines))
    (setq lines (cadr lines))

    ;; Discard parent, but remember its contextual line.
    (setq parent-line (car lines))
    (setq lines (cdr lines))

    ;; How far is the first child indented? At least 1 space; otherwise no children.
    (setq indentation (max 1 (pf--get-indentation (car lines))))

    ;; Keep only the children that have at least this level of indentation.
    (setq lines&more
          (--split-with (<= indentation (pf--get-indentation it)) lines))
    (setq lines (car lines&more))
    (setq unconsidered (cadr lines&more))

    ;; Alter the children according to the given function.
    (setq lines (mapcar then lines))

    ;; Yield the parent line along with the children lines;
    ;; and the unconsumed wild's prefix and suffix.
    `(,prefix ,(cons parent-line lines) ,unconsidered)))

Let's try this out on our example hierarchy, eh, from earlier.

(cadr (pf--get-children "+ item 1" eh))
+ item 1 - subitem 1.1 * subsubitem 1.1.1 - subitem 1.2

(pf–load-package-former (cadr (pf–get-children "PackageFormer" "PackageFormer Test : Set₁ where PackageFormer Another : Set₁ where ")))

Excellent! Let's looks at the other parents.

(pf--get-children "+ item 2" eh)
+ item 0 + item 1 - subitem 1.1 * subsubitem 1.1.1 - subitem 1.2
+ item 2 - subitem 2.2      
+ item 3        

Notice that we found the parent + item 2 and its only child - subitem 1.2, and we kept the prefix of eh that did not contain the parent as well as the remaining unconsidered portion of eh. —Moreover, it looks like we obtained the transpose of the example hierarchy 😛

Finally, the barren parents.

(pf--get-children "+ item 0" eh)
+ item 0            
+ item 1 - subitem 1.1 * subsubitem 1.1.1 - subitem 1.2 + item 2 - subitem 2.2 + item 3

As well as,

(pf--get-children "+ item 3" eh)
+ item 0 + item 1 - subitem 1.1 * subsubitem 1.1.1 - subitem 1.2 + item 2 - subitem 2.2
+ item 3            

Everything before it is considered the prefix. Yay :smile:

Before we move on, let's try altering a child clause; e.g., I'd like * subitem 1.1.1 to be renamed to * subitem that is super deep.

(cadr (pf--get-children "+ item 1" eh
 :then (lambda (x) (s-replace "1.1.1" "that is super deep" x))))
+ item 1 - subitem 1.1 * subsubitem that is super deep - subitem 1.2

Nice :grin:

Now the moment of truth, let's try this out on our example.

(cadr (pf--get-children "PackageFormer" test))
| PackageFormer M-Set : Setwhere | Scalar  : Set | Vector  : Set | _·_     : Scalar Vector Vector | 𝟙       : Scalar | _×_     : Scalar Scalar Scalar | leftId  : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 | assoc   : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋) |

Also, does the list variant work:

(cadr (pf--get-children "PackageFormer" (s-lines test)))
| PackageFormer M-Set : Setwhere | Scalar  : Set | Vector  : Set | _·_     : Scalar Vector Vector | 𝟙       : Scalar | _×_     : Scalar Scalar Scalar | leftId  : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 | assoc   : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋) |

Test-driven development doesn't seem bad 😲

These pretty-coloured tests and results may be nice for exposition, however for maintenance it is ideal to include unit tests that can be checked without human intervention. M-x ert <RET> t <RET>, after executing the following block, will report which tests pass and tries to explain why tests fail.

Tests
(ert-deftest get-ind ()
  (loop for s in '(nil "" "x" "  x" "  x ")
    do (should (<= (pf--get-indentation s) (length s))))
  )

(ert-deftest get-child ()
  (-let [eh
"+ item 1
  - subitem 1.1
    * subsubitem 1.1.1
  - subitem 1.2
+ item 2
  - subitem 2.2
+ item 3"]

    ;; Consider each line above as a parent, with ‘eh’ as the wild.
    (loop for parent in (s-split "\n" eh) do
      (let* ((cs (pf--get-children parent eh))
         (children (cdadr cs)))

      ;; Result is a list of lists: Each is either nil or a cons.
      (loop for r in cs do (should (listp r)))

      ;; The parent line contains the parent.
      (should (equal parent (caadr cs)))

      ;; The children all have the same indentation.
      (loop for c in children for d in children
            do (should (equal (pf--get-indentation c) (pf--get-indentation d))))

      ;; Extensionality: Orginal input can be regained from resulting parts.
      (should (equal eh (s-trim (s-join "\n" (--map (s-join "\n" it) cs)))))))))

5.2 Substrings Delimited by Tokens

How do we find a string delimited by two tokens?

Before we can get to the real stuff, we need to produce a few low-level —string manipulation— utilities, so that we can work with higher-level abstract datatypes.

  • substring-delimited: Given prefix and suffix, this operation takes a string of the form ⋯‘prefix’⟪needle⟫‘suffix’⋯ and yields needle.
  • substring-delimited-here: Given "⟪prefix⟫ $here ⟪suffix⟫" this operation takes a string of the form ⋯‘prefix’⟪needle⟫‘suffix’⋯ and yields needle.
substring-delimited Implementation

We convert all adjacent whitespace characters to a single space in the input STRING and trim any surrounding whitespace from the resulting output needle string.

  • NOTE: Delimiters PREFIX and SUFFIX may be empty.
(cl-defun pf--substring-delimited (prefix suffix string)
  "Assuming “STRING ≈ ⋯PREFIX⟪needle⟫SUFFIX⋯”, return the first such needle.
<<docs('substring-delimited)>>"

  (unless (stringp string)
    (error "PF--SUBSTRING-DELIMITED: Argument STRING must be a string"))
  (let* ((new (s-collapse-whitespace string)))
    (when (not (s-blank? prefix))
      (setq new (car (-take-last
                      (if (equal prefix suffix) 2 1) (s-split prefix new)))))
    (when (not (s-blank? suffix))
      (setq new (car (s-split suffix new))))
    (s-trim new)))
(ert-deftest subst-delimit ()
  (-let [str "𝟘 𝟙 𝟚 𝟛 𝟜 𝟝 𝟜 𝟞"] ;; Intentionally repeated ‘𝟜’.
    ;; Pattern for loop: (prefix postfix expected-needle :comment))
    (loop for it in `( ( "" "" ,str            :Identity)
               ( "𝟘" "𝟞" "𝟙 𝟚 𝟛 𝟜 𝟝 𝟜"  :Boundaries)
               ( "" "𝟞" "𝟘 𝟙 𝟚 𝟛 𝟜 𝟝 𝟜" :NoLeft)
               ( "𝟘" "" "𝟙 𝟚 𝟛 𝟜 𝟝 𝟜 𝟞" :NoRight)
               ( "𝟠" ""  ,str          :BogusL)
               ( "" "∞"  ,str          :BogusR)
               ( "𝟠" "∞" ,str          :BogusLR)
             )
      do (should (equal (third it)
                        (pf--substring-delimited (first it) (second it) str))))

    (should (equal "𝟛" (pf--substring-delimited "𝟚" "𝟜" str)))

    ;; Identical boundaries.
    (should (equal "𝟙" (pf--substring-delimited "𝑳" "𝑳" "𝑳 𝟙 𝑳")))
    (should (equal ""  (pf--substring-delimited "𝑳" "𝑳" "𝑳 𝑳")))
    (should (equal ""  (pf--substring-delimited "𝑳" "𝑳" "𝑳𝑳")))

    ;; Multiple occurances of prefix or postfix
    (should (equal "y"  (pf--substring-delimited "𝑳" "𝑹" "𝑳 x 𝑳 y 𝑹")))
    (should (equal "x"  (pf--substring-delimited "𝑳" "𝑹" "𝑳 x 𝑹 y 𝑹")))
pf--substring-delimited-here Implementation
  • That is, assuming “CONTEXT ≈ PREFIX $here SUFFIX” and “STRING ≈ ⋯PREFIX ⟪needle⟫ SUFFIX⋯”, return the first such needle.

    That is, we place template CONTEXT “on top of” provide STRING, then return whatever falls under position ‘$here’.

  • NOTE: PREFIX and SUFFIX cannot be empty strings!
  • We convert all adjacent whitespace characters to a single space in the input ‘string’ and trim any surrounding whitespace from the resulting output needle string.
(cl-defun pf--substring-delimited-here (context string) "\
Assuming “CONTEXT ≈ PREFIX $here SUFFIX” yield the value of needle ‘$here’.
<<docs('substring-delimited-here)>>"

  (-let [pre-post (s-split "$here" context)]
    (pf--substring-delimited (s-trim (car pre-post))
                             (s-trim (cadr pre-post))
                             string)))
(ert-deftest subst-delimit-here ()
  (-let [str "𝟘 𝟙 𝟚 𝟛 𝟜 𝟝 𝟜 𝟞"] ;; Intentionally repeated ‘𝟜’.
    ;; Pattern for loop: (prefix postfix expected-needle :comment)
    (loop for it in `( ( "$here" ,str              :Identity)
               ( "𝟘 $here 𝟞" "𝟙 𝟚 𝟛 𝟜 𝟝 𝟜"  :Boundaries)
               ( "$here 𝟞" "𝟘 𝟙 𝟚 𝟛 𝟜 𝟝 𝟜"  :NoLeft)
               ( "𝟘 $here"  "𝟙 𝟚 𝟛 𝟜 𝟝 𝟜 𝟞" :NoRight)
               ( "𝟠 $here"   ,str          :BogusL)
               ( "$here ∞"   ,str          :BogusR)
               ( "𝟠 $here ∞" ,str          :BogusLR)
             )
      do (should (equal (second it)
                        (pf--substring-delimited-here (first it) str))))

    ;; Longest substring
    (should (equal "𝟛" (pf--substring-delimited-here "𝟚 $here 𝟜" str)))

    ;; Identical boundaries.
    (should (equal "𝟙" (pf--substring-delimited-here "𝟘 $here 𝟘" "𝟘 𝟙 𝟘")))
    (should (equal ""  (pf--substring-delimited-here "𝟘 $here 𝟘" "𝟘 𝟘")))
    (should (equal ""  (pf--substring-delimited-here "𝟘 $here 𝟘" "𝟘𝟘")))

    ;; Multiple occurances of prefix or postfix
    (should (equal "y"  (pf--substring-delimited-here "𝑳 $here 𝑹" "𝑳 x 𝑳 y 𝑹")))
    (should (equal "x"  (pf--substring-delimited-here "𝑳 $here 𝑹" "𝑳 x 𝑹 y 𝑹")))

    ;; Space irrelevance for keyword ‘$here’:
    (should (equal "𝟙" (pf--substring-delimited-here "𝑳 $here 𝑹" "𝑳 𝟙 𝑹")))
    (should (equal "𝟙" (pf--substring-delimited-here "𝑳 $here𝑹" "𝑳 𝟙 𝑹")))
    (should (equal "𝟙" (pf--substring-delimited-here "𝑳$here 𝑹" "𝑳 𝟙 𝑹")))
    (should (equal "𝟙" (pf--substring-delimited-here "𝑳$here𝑹" "𝑳 𝟙 𝑹")))
    (should (equal "𝟙" (pf--substring-delimited-here "𝑳      $here  𝑹" "𝑳 𝟙 𝑹")))
    ))

Suppose a user provides us with an awkwardly spaced PackageFormer header, our string manipulation setup is robust enough to get at the constituents:

(-let [header "PackageFormer  Semigroup   (  v : Variation) : Set (  ℓexpr)   where"]
  ;; Three kinds of invocations; the last is my preferred choice ♥‿♥
  `( ,(pf--substring-delimited "PackageFormer " "(" header)
     ,(pf--substring-delimited-here "PackageFormer $here (" header)))
Semigroup Semigroup

The aim is to eventually have an interface that interacts with an buffer containing Agda code. To that end, we propose that our fictitious syntax be directly embedded via special comments, {-700 ⋯ -}, henceforth referred to as “700-comments”.

  • (pf--buffer-substring-delimited starting-regexp ending-regexp) yields the next portion of the buffer as a string, relative to the current position of the cursor, that is contained in the ‘parenthesis’ starting-regexp and ending-regexp.
  • (pf--buffer-substring-delimited-whole-buffer starting-regexp ending-regexp) yields all portions of the buffer, contained in the ‘parenthesis’ starting-regexp and ending-regexp, as a list of strings.
    • Cursor position is saved.
    • This function let's us obtain the contents of all 700-comments.

Where we have the following value:

(defvar pf-folding nil
  "Should 700 and Lisp blocks be folded away when 𝑪-𝑐 𝑪-𝑙.")
pf--buffer-substring-delimited Implementation
  • Get the current buffer's next available substring that is delimited between the regexp tokens START up to END, exclusively.
  • If no tokens are found, an error is thrown.
  • MORE is a function that is called on the found instance: It is a function of the start and end positions of the occurance.
  • REGEXP indicates whether we are using regular expression strings, or literals. It is nil by default.
;; pf--declare-type has no support for optionals yet
(cl-defun pf--buffer-substring-delimited
    (start end &optional more &key (regexp t))
  "Return next delimited substring in the current buffer.
<<docs('buffer-substring-delimited)>>"
  (let (start-pos end-pos sp ep content)
    (if regexp (re-search-forward start) (search-forward start))
    (setq start-pos (point))
    (backward-word)
    (setq sp (point))

    (if regexp (re-search-forward end) (search-forward end))
    (setq ep (point))
    (backward-word)
    (setq end-pos (point))

    (setq content  (buffer-substring-no-properties start-pos end-pos))

    (when more (funcall more sp ep))
    (when pf-folding (origami-close-node-recursively (current-buffer) (point)))

    content))
pf--buffer-substring-delimited-whole-buffer Implementation
  • Return a list of all substrings in the current buffer that are delimited by regexp tokens START and END, exclusively.
  • MORE is a function that is called on the found instance: It is a function of the start and end positions of the occurance.
;; pf--declare-type has no support for optionals yet
(cl-defun pf--buffer-substring-delimited-whole-buffer (start end &optional more)
  "Return all delimited substrings in the current buffer.
<<docs('buffer-substring-delimited-whole-buffer)>>"
  ;; Colour 700 keywords red “'error”
  (highlight-phrase start 'error)
  (highlight-phrase end 'error)
  (save-excursion
    (let ((l nil) (continue t))
     (goto-char (point-min))

     (while continue
       (condition-case nil
     ;; attemptClause
     (setq l (cons (pf--buffer-substring-delimited start end more) l))
     ;; recoveryBody
     (error (setq continue nil))))

     ;; We've collected items as we saw them, so ‘l’ is in reverse.
    (reverse l))))

Here are some possible invocations, the last one being our use case.

;; Get text delimited by quotes
(pf--buffer-substring-delimited "^\"" "^\"")

;; Get text delimited by usual Agda comments
(pf--buffer-substring-delimited "^\{-" "^-\}")

;; Execute the following in an Agda buffer to see this function in action.
(setq it (pf--buffer-substring-delimited-whole-buffer "^\{-700" "^-\}"))

5.3 Agda Mixfix Renaming and Imports

Renamaing Agda mixifix names where the rename operation ignores the outermost position markers ‘_’.

  • Given an Agda mixfix operator OP, apply a function on strings F on the inner-most delimiting tokens of the operator, in-particular ignoring outer argument markers ‘_’.
  • For example, if you wish to decorate an operator with a prime or a subscript, we cannot simply catenate else we obtain “_⊕_₁” rather than “_⊕₁_”.
  • Here are some sample results, assuming “f ≈ (λ it → (format “₀%s¹” it))”:
    • ₀⊕¹
    • _[_⊗_] ↦ _₀[_⊗_]¹
    • helo ↦ ₀helo¹
    • he-lo ↦ ₀he-lo¹
  • AVOID-MIXFIX-RENAMING is optional; by default renaming “jumps over” underscores, but providing a non-nil value for this argument leaves underscores alone.

    It is a matter of having, say, default “_⊕ₙ_” versus “_⊕”.

(defun rename-mixfix (f op &optional avoid-mixfix-renaming)
  "Rename an operation by “leaping over” Agda positional markers.
<<docs('rename-mixfix)>>"
  (let* ((parts (s-split "_" op))
         (front (s-blank? (first parts)))
         (rear (s-blank? (car (last parts)))))

    (if avoid-mixfix-renaming
        (funcall f op)
      (--> (concat (when front "_") "$here" (when rear "_"))
           (pf--substring-delimited-here it op)
           (funcall f it)
           (concat (when front "_") it (when rear "_"))))))

We also want to prefix the generated file with the imports of the current file.

extract-imports Implementation
(defvar pf-generated-suffix "-generated"
  "The suffix applied to a file's name to produce it's generated counterpart.")

;; Sometimes we may want the full name due to files being in a nested
;; directory hierarchy: (file-name-sans-extension buffer-file-name)
(defun pf--generated-file-name ()
  "Name of the generated file."
  (concat (file-name-sans-extension (buffer-name)) pf-generated-suffix))

(cl-defun pf--extract-imports ()
  "Return substring of buffer whose lines mention “import”.

Throw away any that mention the substring “⟪FileName⟫-generated”."
  (thread-last (buffer-substring-no-properties (point-min) (point-max))
    (s-split "\n")
    (--filter (s-contains? "import " it))
    (--remove (s-contains? (pf--generated-file-name) it))
    (s-join "\n")))

So much string meddling, hopefully no more 🙈 :hearnoevil: :speaknoevil:

5.4 “λ” for the Agda User

Let's make this prototype more accessible to the Agda user by providing an abbreviation of Lisp functions (lambda (x₀ … xₙ) body) into the Agda-like syntax (λ x₀ ⋯ xₙ → body).

(defmacro λ (&rest body)
  "Implementing Agda style, interactive, lambdas; ideally for inline use:

“λ α β … ω → BODY”  becomes an interactive function with arguments α, …, ω.

The args list may be empty, in which case the separator “→” may be omitted
entirely, if desired.

A BODY must always be supplied, even if the literal nil."
  (let* ((parts (-split-on '→ body)) args rest)

    (if (<= 2 (length parts))
        (progn (setq args (car parts)) (setq rest (cadr parts)))
         ;; Otherwise, only one part was found ---no arguments were provided.
         (setq args nil) (setq rest (car parts)))

   `(lambda ,args (interactive) ,@rest)))

6 The package-former Datatype

For this prototype's constraints, a PackageFormer will generally declared as

PackageFormer Name : Set level where

The body, , of such a declaration consists of a number of name-type declarations and, equations ---not yet supported— so let's form a type to work with these components rather than meddle with strings all the time.

  • ‘docstring’: Relevant documentation about this structure; e.g., what is the instance declaration that generated this type, if any.
  • ‘kind’: PackageFormer, record, data, module, function, etc.
  • ‘name’: The name of the grouping mechanism schema.
  • ‘level’: The universe level that the instantiations will inhabit. The universe level of the PackageFormer.
  • Finally, the children fields are the typed-names that constitute the body of the grouping mechanism. As long as consistent indentation is selected, it does not matter how much. As such, we keep track of these indentation numerics ourselves in case we need to tweak them.
  • The first ‘waist’-many elements are considered parameters.

TODO: Eventually need to support variations?

(defstruct pf--package-former
  "Record of components that form a PackageFormer.
<<docs('package-former)>>"
  docstring
  kind
  name
  level

  waist ;; Delimits elements into parameters and fields.

  ;; children
  indentation ;; useful for when new elements are added.
  elements
)

We will keep track of all such declarations in a global list, and provide a minimal user-interface to it.

(defvar pf--package-formers nil
  "The list of PackageFormer schema declarations in the current Agda buffer.")

(defun $𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠-𝑜𝑓 (pf)
 "Return elements of a given PackageFormer's name, PF.

 This is provided to users; it is one of the few utilities that
 makes use of implementation specfic details.
"

;; (message-box (format "%s" (nth 1 (backtrace-frame 30))))
;; -let [here (progn (goto-char (point-min)) (search-forward pf) (thing-at-point 'line 'no-properties))]

 ;; Ensure the given PackageFormer is defined.
 (pf--ensure (assoc pf pf--package-formers)
             (format "Undefined PackageFormer “%s”" pf)
             (list pf)
             "Ensure you spelled the name correctly."
             "Use the PackageFormer menu to see which PackageFormers are defined.")

 ;; Return its elements.
 (pf--package-former-elements (cdr (assoc pf pf--package-formers))))

Let's also declare a global variable for keeping track of whether 700-syntactical should be coloured or not. Below we use a particular set of default colours.

(defvar pf-highlighting t
  "Should PackageFormer syntactical items be coloured?

➩ Yellow for PackageFormer content.
➩ Red for delimiters 700 and Lisp.
➩ Green for names of variationals.")

6.1 Locally Opening a PackageFormer

It will get rather redundant to write (pf--package-former-X p) to project the constituents of a PackageFormer p. As such, let's introduce a useful macro to “open p” locally.

;; An anaphoric macro ^_^
(defmacro pf--open-pf (p &rest body)
  "Open a package-former P so no qualifiers are required in form BODY."
  `(let*
    ((docstring             (pf--package-former-docstring ,p))
     (kind                  (pf--package-former-kind ,p))
     (name                  (pf--package-former-name ,p))
     (level                 (pf--package-former-level ,p))
     (waist                 (pf--package-former-waist ,p))
     (indentation           (pf--package-former-indentation ,p))
     (elements              (pf--package-former-elements ,p))
     (parameters            (-take waist elements))
     (fields                (-drop waist elements)))
    ,@body))

( Lisp convention would advise this function to be named with-pf, but I'm using the prefix open, as it is closer to the object language, Agda. )

It is crucial to realise that we have just established a convention that partitions the elements of a PackageFormer:

  • parameters are the elements before the waist line.
  • fields are the elements after the waist line.

6.2 Elements

A PackageFormer is a list of elements.

(defstruct element
  qualifier ;; E.g., “private, field”
  name      ;; The lhs of an equation and a typed-name
  type      ;; The type of a typed-name
  equations ;; List of definitional clauses: “same-name-as-above args = term”
)

Let's produce the associated useful map functions.

(loop for place in '(qualifier name type equations)
      do
      (-let [loc (intern (format "element-%s" place))]
        (eval `(defun ,(intern (format "map-%s" place)) (f e)
           ,(format "Alter the ‘%s’ field of an ‘element’ value." place)
           (-let [e′ (copy-element e)]
             (setf (,loc e′) (funcall f (,loc e′)))
             e′)))))

Now a method for replacing textual strings within an element structure.

(defun element-contains (needle e)
  "Check whether string NEEDLE occurs anywhere in element E."
    (--any it
           (-concat (loop for place in '(element-qualifier element-name element-type)
                         collect (eval `(when (,place e) (string-match-p (format "\\b%s\\b" (regexp-quote needle)) (,place e)))))
                   (loop for eq in (element-equations e)
                         collect (string-match-p (format "\\b%s\\b" (regexp-quote needle))
                                 (s-join "=" (cdr (s-split "=" eq)))))))) ;; Ensure we do not consider the LHS of an equation.
(cl-defun element-replace (old new e &key (support-mixfix-names nil) (avoid-altering-names nil))
  "Replace every occurance of word OLD by string NEW in element E.

By default, not AVOID-ALTERING-NAMES, we also make the “OLD ↦ NEW” rewrite if OLD occurs as a subexpression
in the name of the given element E. Example use: We have a name “*-comm” and make the rewrite “_*_ ↦ _+_”,
then we have the desirable name “+-comm”.

We account for “OLD = [qualifier.]OLD” translations, as in function ELEMENT-RETRACT,
by transforming them into “OLD = [qualifier.]NEW”.

There is little support for mixfix names, where terms such as “x + y ↦ x * y”
when the replacement of “_+_” by “_*_” is invoked."
  (let* ((e′    (copy-element e))
         (old′  (s-replace "_" " " old))
         (old′=  (format "%s%s%s" (gensym) old "="))
         (old′-=  (format "\\b%s =\\b" old′))
         (word (if (or (= (length old′) 1) (string-match-p "[[:alnum:]].*[[:alnum]]" old′)) "\\b" "")) ;; regexp word marker ;; Do we have a word, or something strange like _+_ ?
         (temp  (format "%s%s" new (gensym)))
         (new′  (s-replace "_" " " new))
         (offend  (format "%s = \\(.+\\)?%s" (regexp-quote temp) (regexp-quote temp))) ;; “l = [qualifier.]l”
         (correct (format "%s = \\1%s" old new)))     ;; “\\1” refers to the matched qualifier.

    ;; Also account for “l = l” translations; c.f., element-retract.
    ;; E.g., with “old, new ≔ y, x” we have
    ;; “let x = y in f x  ⟿  let x = temp in f x  ⟿  let x = x in f x”
    ;; Without the ‘temp’ switch, we would have had: “let x = y in f x  ⟿ let x = x in f x ⟿ let y = x in f x”,
    ;; which may be fine in a let-clause, but ruins a record having “x” as a field.
    (loop for (this . that) in (-concat `((,(regexp-quote old)  . ,temp)  (,offend . ,correct) (,(regexp-quote temp) . ,new))
                                         (when (s-contains-p "_" old) (list (cons (regexp-quote old′) new′))))
          do

          (loop for place in '(element-qualifier element-type)
                do (eval `(when (,place e′)
                            (setf (,place e′)
                                  (replace-regexp-in-string (concat word ,this word)
                                                            (concat that "") (,place e′) t)))))

          ;; Replacements in the equations as well; but not the LHS!
          (setq e′ (map-equations (λ eqs → (--map (replace-regexp-in-string (concat word this word) that it t) eqs)) e′))
   )

    ;; Replacement in the name portion.
    ;; (message-box (format "%s ↦ %s in %s" this that (element-name e′)))
    (unless avoid-altering-names
      (setq old′ (regexp-quote (s-replace "_" "" old)))
      (setq new′ (s-replace "_" "" new))
      (setq e′ (cond
                ((not (string-match-p "[[:alnum:]]" old′))
                   (map-name (λ n → (replace-regexp-in-string old′ new′ n)) e′))
                ;; Else we have old′ is a letter[s] or number[s], so we make the replacements provided it's NOT next to an alphanumeric character, such as _,/,-,⋯
                ((string-match-p (format "%s\\([^[:alnum:]]\\)" old′) (element-name e′))
                   (map-name (λ n → (replace-regexp-in-string (format "%s\\([^[:alnum:]]\\)" old′) (format  "%s\\1" new′) n)) e′))
                ((string-match-p (format "\\([^[:alnum:]]\\)%s" old′) (element-name e′))
                   (map-name (λ n → (replace-regexp-in-string (format "\\([^[:alnum:]]\\)%s" old′) (format  "\\1%s" new′) n)) e′))
                (t e′))))

    ;; return value
    e′))

Given a typed name, return the name.

(defun parse-name (element)
  "Given an string representation of an ELEMENT, yield the ‘name’ component.

The shape of the input may be “qualifier lhs ~ rhs” where ‘~’ is either ‘:’
or ‘=’.  The qualifier is a ‘special’ word: field, private."
  (let ((lhs (s-split " " (car (s-split " = " (car (s-split " : " element)))))))
    (if (and (< 1 (length lhs)) (pf--special (nth 0 lhs)))
        (cadr lhs)
      (car lhs))))

Parsing a list of interleaved elements in concrete syntax.

  • Given a list of PackageFormer ELEMENTS, as strings, parse them into the ‘element’ datatype. Declarations and equations may be interspersed, as along as equations of names follow their declarations.
  • The order is preserved in-case there are declarations that make use of definitions.
  • Types must always be supplied —in general, type inference is undecidable in DTLs.
(defun parse-elements (elements)
  "Parse string representation of elements into the ‘element’ record type.
<<docs('parse-elements)>>"
  (-let [es (mapcar #'list elements)]
    ;; Maintain a list of related items.
    (loop for i from 0
          for e in es
          for name = (parse-name (car e))
          do (loop for j from 0 to (1- i)
                   do
                   ;; If the name of ‘e’ occurs in the prefix,
                   ;; then move ‘e’ to the location in the prefix,
                   ;; and zero-out the current location.
                   (when (equal name (parse-name (or (car (nth j es)) "")))
                     ;; Use an empty string in-case the location is nil.
                     (setf (nth j es) (append (nth j es) e))
                     (setf (nth i es) nil))))

    ;; Drop the nils.
    (setq es (--reject (not it) es))

    ;; We now have a list of related items,
    ;; with the car of each being a qualified typed-name
    ;; and the cdr of each being a list of equational clauses
    ;; associated with that name.
    (loop for e in es
          for τ    = (s-split " : " (car e))
          for nom  = (parse-name (car τ))
          ;; In case there is no qualifier; regexp-quote to avoid regexp ops in a name.
          for qual = (-let [pts (s-split (regexp-quote nom) (car τ))] (when (= 2 (length pts)) (car pts)))
          for qual = (car (s-split nom (car τ)))
          for _    = (pf--ensure (cdr τ)
                                 (format "Type not supplied for %s!" nom)
                                 (s-join "\n\t\t" (-take 5 elements))
                                 "⋯ Add a type!")
          for ty   = (s-join " : " (cdr τ))
          collect
          (make-element :qualifier (unless (s-blank? qual) qual)
                        :name nom
                        :type ty
                        :equations (cdr e)))))

Notice that the equations do not necessarily immediately follow their associated declarations.

6.3 Well-formed checks —Error reporting

What is the above “fancy cons” we mentioned? Here it is.

  • Ensure CONDITION is true and defined, otherwise emit MESSAGE and indicate the offending CONTEXT. If there are any SUGGESTIONS to the user, then we show those too.
  • If CONDITION is defined and non-nil, whence true, we return it.
(defmacro  -ensure (condition message context &rest suggestions)
  "Ensure provided CONDITION is true, otherwise report an error.
<<docs('ensure)>>"
  `(let* ((ლ\(ಠ益ಠ\)ლ
           (format "700: %s\n\n\t⇨\t%s%s%s" ,message ,context
                   (if (quote ,suggestions) "\n" "")
                   (s-join "\n" (--map (format "\t⇨\t%s" it)
                                       (quote ,suggestions)))))
          ;; Try to evaluate the condition.
          (res (condition-case nil ,condition (error ლ\(ಠ益ಠ\)ლ))))

     ;; If we've made it here, then the condition is defined.
     ;; It remains to check that it's true.
     (or res (error ლ\(ಠ益ಠ\)ლ))))
  • This operation checks that the VALUE of KEY is well-formed according to 700-specifications —which are stated explicitly within this method— and if it is well-formed we return the VALUE interpreted along with the KEY.
  • When the value is not well-formed, we use the provided CONTEXT in an error message. No error is reported if VALUE is an ARGument, ARGS, of a variational begin declared.
(defun -wf (key value &optional context args)
  "Report an error unless provided key-value are well-formed.
<<docs('wf)>>"
  (let* ((case
            (pcase key
              (:kind `(,(-contains? '(record data module PackageFormer) value)
                       This kind “ ,value ” is not support by Agda!
                       Valid kinds: record⨾ data⨾ module⨾ PackageFormer!))
              (:waist `(,(numberp value)
                        The waist should be a number⨾ which “ ,value ” is not!))
              (:level `(,(-contains? '(inc dec none) value)
                        The “level” must be “inc” or “dec” or “none”⨾
                        which “ ,value ” is not!))))
        (condition (car case))
        (message   (mapconcat #'prin1-to-string (cdr case) " ")))

    ;; TODO: Acount for alter-elements well-formedness?
    ;; (:alter-elements (functionp value)
    ;; (format "Componenet alter-elements should be a function;
    ;; which “%s” is not." value))

    (when case
      ( -ensure (or condition (-contains? args value)) message context))

    ;; Return the key-value as a pair for further processing.
    ;; :kind and :level values are symbols and so cannot be evaluated furthur.
    (cons key
          (if (or (-contains? args value) (-contains? '(:kind :level) key))
              value
            (eval value)))))

6.4 Package Former Parsing and Pretty Printing

With this in hand, let's produce a robust parser.

  • The input LINES must be a list of lines forming a full PackageFormer declaration; e.g., obtained by calling ‘pf–get-children’.
  • It is parsed and a ‘package-former’ value is returned.
  • Whitespace is stripped off of items.
  • Docstrings are ignored.
(defun pf--load-package-former (lines)
  "Load a string representation of a ‘package-former’ into our global list.
<<docs('load-package-former)>>"
  (when (not lines)
      (error "PF--LOAD-PACKAGE-FORMER: Error: Input must be non-empty list"))

  (let* (pf
         (header (or (car lines) ""))
         (name (pf--substring-delimited-here "PackageFormer $here :" header))
         (level (pf--substring-delimited-here "Set $here where" header)))

    (when pf-highlighting
      (mapc (λ it → (highlight-phrase (s-trim it) 'hi-yellow)) (cdr lines)))

    (setq pf
          (make-pf--package-former
           :kind                     "PackageFormer"
           :name                     name
           ;; level’ may be “”, that's okay.
           ;; It may be a subscript or implicitly zero, so no space after ‘Set’.
           :level                    level
           :waist                    0
           ;; TODO: Currently no parameter support for arbitrary PackageFormers.
           :indentation              (max 4 (pf--get-indentation (cadr lines))) ;; 4 spaces is the minimum
           :elements  (parse-elements (--remove (s-starts-with? "-- " it)
                                                (--map (s-trim it)
                                                       (cdr lines))))))

      (push (cons name pf) pf--package-formers)

      ;; return value
      pf))

Let's try this out.

;; (equal 4 (pf--package-former-indentation (pf--load-package-former (cadr (pf--get-children "PackageFormer" "PackageFormer Empty : Set where")))))

(pf--load-package-former (cadr (pf--get-children "PackageFormer" test)))
#s(package-former nil PackageFormer M-Set ₁ 0 nil 3 (#s(element nil Scalar Set nil) #s(element nil Vector Set nil) #s(element nil · Scalar → Vector → Vector nil) #s(element nil 𝟙 Scalar nil) #s(element nil × Scalar → Scalar → Scalar nil) #s(element nil leftId {𝓋 : Vector} → 𝟙 · 𝓋 ≡ 𝓋 nil) #s(element nil assoc {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋 ≡ a · (b · 𝓋) nil)))

Conversely, let's have a pretty printer.

  • Special elements F, for whatever reason are exceptional, and so are maked as singleton lists and their indentation is lessened. That is, these denote sibling fields rather than more children.
  • Special elements include: field, private.
  • See ‘show-package-former’ for their use and how their printed.
(defun pf--special (f)
  "Test whether an element F is special or not.
<<docs('special)>>"
  (--any? (s-contains? it f) '("field" "private" "open" "top-level" "sibling")))

A “sibling” item has no indentation.

(cl-defun show-element (e &optional omit-qualifier)
  "Render an ‘element’ value E in the form:

   qualifier name : type ; equational-clause₀ ; ⋯ ; equational-clauseₙ

Optional OMIT-QUALIFIER is useful for when
elements are in a parameter position."
  (s-join " ;\t"
          (cons
           (format "%s%s\t\t: %s"
                   (-let [it (element-qualifier e)]
                     (if (or (not it) omit-qualifier) "" (format "%s " it)))
                   (element-name e)
                   (element-type e))
           (element-equations e))))
(cl-defun pf--show-package-former (pf)
  "Pretty print a package-former PF record value."
  (pf--open-pf pf
    (s-join "\n"
      (-cons*

       ;; The documentation string
       (when docstring (format "{- %s -}" docstring))

       ;; The schema declaration
       (s-collapse-whitespace
        (s-join " "
                (list kind
                      name
                      (s-join " "
                              (--map (concat "("
                                             (show-element it :omit-qualifier)
                                             ")")
                                     parameters))
                      (unless (equal level 'none) (concat ": Set" level))
                      "where")))

       ;; The elements of a PackageFormer
       (thread-last fields
         (--map (format "%s%s"
                        (s-repeat indentation " ")
                        (show-element it))))))))

Let's test it out by introducing a whole new local variable and trying to include the field Agda keyword.

(let* ((raw (cadr (pf--get-children "PackageFormer" test)))
       (pf (pf--load-package-former raw))
       (waist 2))
  (setf (pf--package-former-waist pf) waist)

  ;; mark all items as "fields"
  (--map (setf (element-qualifier it) "field") (pf--package-former-elements pf))

  ;; inject new items at the waist line
  (loop for new in (parse-elements '("private n : ℕ" "n = 3" "𝒮ℯ𝓉 : Set₁" "𝒮ℯ𝓉 = Set"))
        do (push new
        (cdr (nthcdr (1- waist) (pf--package-former-elements pf)))))

   (pf--show-package-former pf))
PackageFormer M-Set (Scalar : Set) (Vector : Set) : Setwhere
   𝒮ℯ𝓉      : Set₁ ;    𝒮ℯ𝓉 = Set
   private  n       :  ;   n = 3
   field _·_        : Scalar Vector Vector
   field 𝟙      : Scalar
   field _×_        : Scalar Scalar Scalar
   field leftId     : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋
   field assoc      : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)

Notice that the \(𝒮ℯ𝓉 = Set\) alias is a top-level item, with respect to the current PackageFormer.

We can phrase an approximation of the opinion that parsing and showing should be inverses.

(-let [pf (cadr (pf--get-children "PackageFormer" test))]
  (should (equal (s-collapse-whitespace (concat "\n" (s-join "\n" pf)))
     (s-collapse-whitespace (pf--show-package-former (pf--load-package-former pf))))))
t

( In Lisp, t denotes “true”! )

Unit Tests
(ert-deftest pf-parse ()

  ;; Error on empty list of lines.
   (should-error (pf--load-package-former nil))

   ;; No crash on empty line.
   (should (pf--load-package-former (list "")))

   ;; No crash on PackageFormer with no elements.
   (should (pf--load-package-former (list "PackageFormer PF : Set ℓ where")))

   ;; Levels
   (should (equal "ℓ" (pf--package-former-level (pf--load-package-former (list "PackageFormer PF : Set ℓ where")))))
   ;;
   (should (equal "" (pf--package-former-level (pf--load-package-former (list "PackageFormer PF : Set  where")))))
   ;;
   (should (equal "₃" (pf--package-former-level (pf--load-package-former (list "PackageFormer PF : Set₃ where")))))
   ;;
   (should (equal "(Level.suc ℓ)" (pf--package-former-level (pf--load-package-former (list "PackageFormer PF : Set (Level.suc ℓ) where")))))

   ;; Full parsing.
   (-let [pf (pf--load-package-former (cadr (pf--get-children "PackageFormer" test)))]
     (equal (format "%s" pf)
            "#s(pf--package-former nil PackageFormer M-Set ₁ 0 nil 3 (Scalar  : Set Vector  : Set _·_     : Scalar → Vector → Vector 𝟙       : Scalar _×_     : Scalar → Scalar → Scalar leftId  : {𝓋 : Vector}  →  𝟙 · 𝓋  ≡  𝓋 assoc   : {a b : Scalar} {𝓋 : Vector} → (a × b) · 𝓋  ≡  a · (b · 𝓋)))"))

  (-let [pf (cadr (pf--get-children "PackageFormer" test))]
    (should (equal (s-concat "\n" (s-join "\n" pf))
                   (pf--show-package-former (pf--load-package-former pf)))))

)

7 Variational Language

A variational has the syntactic form specfied by

        𝓋   ::=  identifier (identifier)* = 𝓋𝒸
        𝓋𝒸  ::= [identifier] (:key value)* (⟴ 𝓋𝒸)*

        identifier ≈ key ≈ value ≈ string of text

With the intention that l a = r is a list of key-value pairs determined from the right-hand side where the arguments a are to be considered place-holders. Whenever one mentions free variables, or terms, one actually speaks of functions. Hence, let's reify these as functions.

Let's keep track of the user defined variationals, so that we can display them to the user via menus, as well as highlight them, and attach tooltips to them.

(defvar pf--variationals nil
  "Association list of Agda-user defined variational operators.")

Let's be somewhat agonstic regardiing the variational composition operator; users are welcome to use any other symbol.

(defvar pf-variational-composition-operator "⟴"
  "The operator that composes varitionals.")

7.1 Well-formed checks —Error reporting

What is the above “fancy cons” we mentioned? Here it is.

  • Ensure CONDITION is true and defined, otherwise emit MESSAGE and indicate the offending CONTEXT. If there are any SUGGESTIONS to the user, then we show those too.
  • If CONDITION is defined and non-nil, whence true, we return it.
(defmacro pf--ensure (condition message context &rest suggestions)
  "Ensure provided CONDITION is true, otherwise report an error.
<<docs('ensure)>>"
  `(let* ((ლ\(ಠ益ಠ\)ლ
           (format "700: %s\n\n\t⇨\t%s%s%s" ,message ,context
                   (if (quote ,suggestions) "\n" "")
                   (s-join "\n" (--map (format "\t⇨\t%s" it)
                                       (quote ,suggestions)))))
          ;; Try to evaluate the condition.
          (res (condition-case nil ,condition (error ლ\(ಠ益ಠ\)ლ))))

     ;; If we've made it here, then the condition is defined.
     ;; It remains to check that it's true.
     (or res (error ლ\(ಠ益ಠ\)ლ))))
  • This operation checks that the VALUE of KEY is well-formed according to 700-specifications —which are stated explicitly within this method— and if it is well-formed we return the VALUE interpreted along with the KEY.
  • When the value is not well-formed, we use the provided CONTEXT in an error message. No error is reported if VALUE is an ARGument, ARGS, of a variational begin declared.
(cl-defun pf--wf (key value &optional context args)
  "Report an error unless provided key-value are well-formed.
<<docs('wf)>>"
  (let* ((case
            (pcase key
              (:kind `(,(-contains? '(record data module PackageFormer) value)
                       This kind “ ,value ” is not support by Agda!
                       Valid kinds: record⨾ data⨾ module⨾ PackageFormer!))
              (:waist `(,(numberp value)
                        The waist should be a number⨾ which “ ,value ” is not!))
              (:level `(,(-contains? '(inc dec none) value)
                        The “level” must be “inc” or “dec” or “none”⨾
                        which “ ,value ” is not!))))
        (condition (car case))
        (message   (mapconcat #'prin1-to-string (cdr case) " ")))

    ;; TODO: Acount for alter-elements well-formedness?
    ;; (:alter-elements (functionp value)
    ;; (format "Componenet alter-elements should be a function;
    ;; which “%s” is not." value))

    (when case
      (pf--ensure (or condition (-contains? args value)) message context))

    ;; Return the key-value as a pair for further processing.
    ;; :kind and :level values are symbols and so cannot be evaluated furthur.
    (cons key
          (if (or (-contains? args value) (-contains? '(:kind :level) key))
              value
            (eval value)))))

7.2 𝒱𝒸, 𝒱-, and 𝒱

Variationals are refieid as functions with the prefix 𝒱-; doing otherwise may override existing utilities in Emacs —for example, ‘record’ is a useful name for a variational but is a super important utility function in Emacs. The apporach we take is to allow users to write ‘record’ but the implementation will refer to it with a prefix.

Parsing variational clauses is then straightforward:

(when nil *parent-context* nil
  "For error report; what is the current parent context of a child item.

   Should be set whenver a parent invokes a child.
   Since we have no grandchildren, we only need one level.
")
  • If there is a ‘label’, then yield ‘(label :key value ⋯)’ since ‘label’ is assumed to exist as a variational having the given keys as arguments. The result should be a list of pairs.

    BODY-LIST consists of elements of this shape.

  • If there is no label, the parse the list of pairs.
  • For example,

    (cl-defun 𝒱-test (&key height kind) (list (format "%s & %s" height kind)))

    (𝒱𝒸 '(test :height 3 :kind 'data)) ≈ (test :height 3 :kind data) ≈ (“3 & data”)

    (𝒱𝒸 '( :height 3 :kind data)) ≈ ((:height . 3) (:kind . data))

  • Newer items c₀ ⟴ ⋯ ⟴ cₙ should be at the front of the list; access should then be using ‘assoc’.
  • CONTEXT is the parent context that contains an invocation of this method.
  • ARGS is the list of names that are bound, and so are variational args.
(defun 𝒱𝒸 (body-list &optional context args)
  "Parse a single 𝒱ariational 𝒸lause, “[label] (:key :value)*”, as a list.
<<docs('𝒱𝒸)>>"
  (let (res)
    (loop for clause in (-split-on '⟴ body-list)
          do (setq res (-concat
                        ;; Symbols starting with “:” are keywords.
                        (if (not (keywordp (car clause)))
                            ;; Function invocation case
                            ;; We turn everything into a string so that we may
                            ;; prepend the function name with a 𝒱-
                            ;; then turn that into Lisp with the first eval
                            ;; then invoke the resulting function call
                            ;; with the second eval.
                            (progn
                              ;; The variational being called is defined.
                              (pf--ensure (fboundp (𝒱- (car clause)))
                                          (format "%s %s “%s” %s"
                                                  "Did you mistype a"
                                                  "variational's name:"
                                                  (car clause)
                                                  "is not defined.")
                                          context
                                          "Use the PackageFormer menu
                                                 to see which variationals
                                                 are defined.")
                              (eval `( ,(𝒱- (car clause)) ,@(--map (if (or (keywordp it) (ignore-errors (macrop (car it)))) it `(quote ,it)) (cdr clause)))))
                          ;; List of key-value pairs
                          `,(loop for key   in clause by #'cddr
                                  for value in (cdr clause) by #'cddr
                                  collect (pf--wf key value context args)))
                        ;; “pf--wf” is just a fancy “cons”.
                        ;; Newer items c₀ ⟴ ⋯ ⟴ cₙ should be at the
                        ;; front of the list;
                        ;; access should then be using assoc.
                        res)))
    res))

;; (defmacro nope (&rest b) `(f ,@(--map (if (keywordp it) it `(quote ,it)) b)))
;; (equal '(f 'a :b 'c) (macroexpand '(nope a :b c)))
;;
;; (macrop (car '(lambda a))) ;; (ignore-errors (car '(a b)))

Where we have used the following helper to prefix Lisp code with “𝒱-”.

(defun 𝒱- (name)
  "Prefix the Lisp data NAME with a “𝒱-” then yield that as a Lisp datum."
  (should (symbolp name))
  (thread-last name
    (format "𝒱-%s")
    read-from-string
    car))

We now handle variational declarations by introducing a DSL; i.e., a macro that expects the syntax outlined earlier.

  • The grammar:

    𝓋 ::= [docstring] identifier ([“(”]identifier[“)”])* = 𝓋𝒸

    𝓋𝒸 ::= [identifier] (:key value)* (⟴ 𝓋𝒸)*

  • The result is a function NAME prefixed by 𝒱- whose BODY is an alist obtained from the aforementioned key-value pairs.
  • E.g., (𝒱 tes positional (keyword 3) = :kind data) This defines a variational with one positional and one keyword argument having 3 as default.
  • The resulting generated function has its code embeded as a docstring viewable with “𝑪-𝒉 𝒐” —catented after any provided user documentation.
(defmacro 𝒱 (name &rest body)
  "Reify as Lisp a variational declaration using the variational grammar.
<<docs('𝒱)>>"
  ;; Main code follows.
  (let* ((context (mapconcat (λ x → (prin1-to-string x t)) (cons name body) " "))
         (args-body (-split-on '= body))
         args pargs kargs argnames docs body res actual-code)
    (pcase (length args-body)
      (2 (setq args (car args-body)
               body (cadr args-body)))
      (_ (setq body (car args-body))))

    ;; Realise the arguments as either 𝒫ositinal or 𝒦ey arguments.
    (loop for a in args
          do (if (consp a) (push a kargs) (push a pargs)))

    ;; The arguments are in reverse now, which doesn't matter for keywords
    ;; yet is crucial for positional arguments. So let's fix that.
    (setq pargs (reverse pargs))

    ;; Keep track of only the argument names, omitting any default values.
    (setq argnames (append pargs (mapcar #'car kargs)))

    ;; Set any documentation string and reify the body's variational clauses.
    (when (stringp (car body)) (setq docs (car body) body (cdr body)))
    (setq res (𝒱𝒸 body context argnames))

    ;; I want to be able to actually, visually, see the resulting
    ;; generated definition of a function.
    ;; Hence, I embed its source code as a string in the code.
    ;;
    ;; I'm using strings so that they appear in the docstring via C-h o.
    ;;
    (setq actual-code
    `(cl-defun ,(𝒱- name) (,@pargs &key ,@kargs)

       ;; Stage the formal names *now*, then evaluate their values at run time.
       ;; Traverse the list of pairs and change the nested formal names with the
       ;; given values. Praise the Lord!
      (let* ((give-goal (quote ,res)) (give-goal₀ give-goal))
        (when (quote ,argnames)

          "Stage the formal names *now*, then evaluate their values at run time."
          (loop for arg in (quote ,argnames)
                do (setq give-goal (cl-subst (eval arg) arg give-goal)))

          ;; TODO, maybe.
          ;; "Check that substituted values are well-typed"
          ;; (--map (pf--wf (car it) (or (cdr it)
          ;;                        ;; Mention which argument is not supplied.
          ;;                         (format "No Value for :%s Provided!"
          ;;                       (cdr (assoc (car it) (reverse give-goal₀))))
          )

         give-goal)))

    ;; Now set the code as a documentation string in it, after the fact.
    (setq docs (format "Arguments:\t%s %s\n%s" pargs (reverse kargs)
                       (if (not docs) "Undocumented user-defined variational."
                         ;; Keep paragraph structure, but ignore whitespace.
                         (thread-last docs
                           (s-split "\n\n")
                           (mapcar #'s-collapse-whitespace)
                           (mapcar #'s-trim)
                           (s-join "\n\n")
                           (s-word-wrap 70)
                           (format "\n%s\n\n\n")))))
              ;; When the user provides documentation, they may not want to see
             ;; the raw and expansions, so we pad extra whitespace before them.

    (put (𝒱- name) 'function-documentation
         (format "%s\n⟪User Definition⟫\n\n%s\n\n⟪Lisp Elaboration⟫\n\n%s"
                 docs context (pp-to-string actual-code)))
    ;; Register this new item in our list of variationals.
    (push (cons name docs) pf--variationals)
    ;; Return value:
    actual-code))

Consequently, any item declared with 𝒱 now has a docstring containing its user-facing definition as well as its Lisp realisation ^_^ —simply press “C-h o ENTER” on its name.

Here's an example.

(𝒱 test positional (keyword 3) another = "I have two mandatory arguments and one keyword argument")

(𝒱-test "positional arg₁" "positional arg₂" :keyword 25) ;; ⇒ nil ^_^

7.3 Loading Variationals: Super Simple Conversion From String to Lisp

  • A VARIATION-STRING line is something like:

    𝒱-name x₀ … xₙ = ([label₀] :key₀ val₁ ⋯ :keyₘ valₘ ⟴)*

  • The result is a list of 3-tuples (name (x₀ ⋯ xₙ) ((key₀ val₀) ⋯ (keyₘ valₘ))), containing the clause's name, argument list, and key-value pairs.
  • For now, the RHS must be an expression of the form “:key₀ value₀ ⋯ :keyₙ valueₙ”
    • where the valueᵢ are legitmate Lisp expressions
    • and the LHS is an atomic name, possibly with argument names.
(cl-defun pf--load-variational (variation-string)
  "Obtain lines of the buffer that start with “𝒱-” as a Lisp alist.
<<docs('load-variational)>>"
  (thread-last variation-string
    (s-replace "𝒱-" "𝒱 ")
    (format "(%s)")
    read-from-string
    car
    eval))
Unit Tests
(ert-deftest variationals-𝒱𝒸 ()

  (should (equal (𝒱- 'nice)
                 '𝒱-nice))

  (should (equal (𝒱𝒸 '(:height 3 :kind 'data))
                '((:height . 3) (:kind . data))))


  ;; Error along with “noice”.
  (should-error (𝒱𝒸 '(:height 3 :kind datda) 'noice nil))

  ;; nice error.
  (should-error (𝒱𝒸 '(:level 3)))

  ;;
  (cl-defun 𝒱-test (&key height kind) `( (first . ,height) (second . ,kind)))
  ;;
  (should (equal (𝒱𝒸 '(test :height 3 :kind 'three ⟴ :kind 'module))
                 '((:kind . module) (first . 3) (second . three))))
  ;;
  ;; NOTE: 𝒱-tests′ :kind is optional
  (should (equal (𝒱𝒸 '(test :height 3 ⟴ :kind 'module))
                 '((:kind . module) (first . 3) (second))))
  ;;
  (should (equal (𝒱𝒸 '(:height 3 ⟴ :kind 'module))
                 '((:kind . module) (:height . 3))))

  ;; Recursively place 3 (new) wherever 'it (old) occurs.
  ;; This' a standard Lisp utility.
  (should (equal
           (subst 3 'it '(1 2 it 4 (5 it) 7 (+ 8 it)))
           '(1 2 3 4 (5 3) 7 (+ 8 3))))
)
(ert-deftest variationals-𝒱 ()

  ;; Nullary
  (should (𝒱 test₀  = :kind 'record :waist 3))
  (should (equal (𝒱-test₀)
               '((:kind . record) (:waist . 3))))

  ;; Unary
  (should (𝒱 test₁ heightish = :kind 'record :waist heightish))
  (should (equal (𝒱-test₁ :heightish 6)
                 '((:kind . record) (:waist . 6))))

  ;; Invoking the previously defined variational
  (should (𝒱 test₂  = :kind 'data ⟴ test₁ :heightish 2))
  (should (equal (𝒱-test₂)
                 '((:kind . record) (:waist . 2) (:kind . data))))

  ;; See a nice error message ^_^
  (should-error (𝒱 test₃ = :kind recordd))
)
(ert-deftest variationals-loading ()

  (should (pf--load-variational "𝒱-tc this height = :level this :waist height"))

  ;; NEATO! (Has desired error)
  ;; (-let [*parent-context* "woadh"]
  ;;   (𝒱-tc :height 'no :this 'inc))
  ;;
  ;; Does not pass: I've commented out the type checking in 𝒱 above, for now.

  (should (𝒱-tc :height 9 :this 'inc))

  (should (equal (𝒱𝒸 '(:a 'b ⟴ tc :height 1))
                 '((:level) (:waist . 1) (:a . b))))

;
)

8 Loading an Agda Buffer

Before we can parse an Agda buffer, we need to be able to parse an instantiation declaration; for which we would later generate code.

8.1 Loading an Instance —The Core Utility

An instance declaration is of the form:

PF′ = PF variationalargs₁ ⟴  ⋯ ⟴ variationalₙ (argsₙ)

This gives rise to a simple nice structure:

(defstruct pf-instance-declaration
  "Record of components for an PackageFormer instance declaration:
   ⟪name⟫ = ⟪package-former⟫ (⟴ ⟪variation⟫ [⟪args⟫])*
  "

  docstring      ;; What the declaration looks like, useful for reference.
  name           ;; Left-hand side
  package-former ;; Parent grouping mechanism
  alterations    ;; List of variationals along with their arguments.
)

Loading an instantiation into our list is now trivial.

  • If the current LINE string is an instance declaration, then produce a new PackageFormer from it. Else, do nothing.
  • Whitespace is automatically collopased from LINE.
  • Nil elements are discarded; e.g., due to a filter.
  • Duplicates are discarded; e.g., due to a rename.
  • Variational clauses may mention
    • $𝑛𝑎𝑚𝑒: The name of the PackageFormer currently being declared; i.e., the LHS name.
    • $𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠: Many variationals will act on individal elements, but may check a property relative to all elements and this name allows us to avoid having variationals that simply accomodate for binary functions that operate on an individual element while also needing to refer to all elements. For example, a variational that keeps an element if it's related to another element somehow.
  • SHOW-IT: For testing, and presentation, this optional value indicates whether the resulting PackageFormer should be pretty-printed or not.
(defvar pf-consider-newly-named-declarations-only nil
  "Are we updating only new named declarations? If so, don't waste time doing work.")

(defun pf--load-instance-declaration (line &optional show-it)
  "Reify concrete instance declarations as ‘package-former’ values.
<<docs('load-instance-declaration)>>"
  (letf* (
     (pieces (s-split " " (s-collapse-whitespace line)))
     ($𝑛𝑎𝑚𝑒      (nth 0 pieces))
     ($𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠    nil)
     (eqSymb     (nth 1 pieces))
     ($𝑝𝑎𝑟𝑒𝑛𝑡     (nth 2 pieces))
     (variations (nthcdr 3 pieces))
     (alterations nil)
     (self (copy-pf--package-former (cdr (assoc $𝑝𝑎𝑟𝑒𝑛𝑡 pf--package-formers))))
     (;; If componenet ‘c’ is in the ‘alterations’ list of the
      ;; instance declaration, then evalaute any given ‘more’ code,
      ;; get the value for ‘c’ and turn it
      ;; into a string, if ‘str’ is true, then set the new PackageFormer's ‘c’
      ;; componenet to be this value.
      ;; Well-formedness checks happen at the 𝒱 and 𝒱𝒸 stages, see below.
      (lambda (c &optional str more)
        (when-let ((it (cdr (assoc (intern (format ":%s" c)) alterations))))
          (eval `(progn ,@more))
          (when str (setq it (format "%s" it)))
          (eval `(setf (,(car (read-from-string
                               (format "pf--package-former-%s" c))) self)
                       it))))))

    ;; Ensure instance declaration is well-formed.
    (pf--ensure (and (not (s-blank? (s-trim $𝑛𝑎𝑚𝑒))) (equal "=" eqSymb) $𝑝𝑎𝑟𝑒𝑛𝑡)
                (concat "An instance declaration is of the form "
                        "“new-name = parent-package-former "
                        "variational-clauses”.")
                line)

    (catch 'early-exit

    ;; Are we updating only new named declarations? If so, don't waste time doing work.
    (when pf-consider-newly-named-declarations-only
      (when (assoc $𝑛𝑎𝑚𝑒 pf--package-formers) (throw 'early-exit nil)))

    ;; Let's not overwrite existing PackageFormers.
    (pf--ensure (or (not (assoc $𝑛𝑎𝑚𝑒 pf--package-formers)) (equal $𝑛𝑎𝑚𝑒 "_"))
                (format "PackageFormer “%s” is already defined; use a new name."
                        $𝑛𝑎𝑚𝑒)
                line)
                ;; "Also, “_” is the name of anonymous PackageFormers."

    ;; Ensure the PackageFormer to be instantiated is defined.
    (pf--ensure self
                (format "Parent “%s” not defined." $𝑝𝑎𝑟𝑒𝑛𝑡)
                line
                "Use the PackageFormer Emacs menu to see which PackageFormers are defined."
                "Perhaps you did not enclose the parent in 700-comments?")

    ;; Update the new PackageFormer with a docstring of its instantiation
    ;; as well as its name.
    (setf (pf--package-former-docstring self) line)
    (setf (pf--package-former-name self) $𝑛𝑎𝑚𝑒)
    (setq $𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠 ; Copy so that user does not inadvertently
                  ; alter shared memory locations!
          (loop for e in (pf--package-former-elements self)
                collect (copy-element e)))

    ;; Parse the “𝓋₀ ⟴ ⋯ ⟴ 𝓋ₙ” portion of an instance declaration.
     (thread-last  variations
       (s-join " ")     ;; Stick the rest back together.
       (format "'(%s)") ;; Construe as a lisp list
       read-from-string
       cadar
       (setq variations))
     ;;
     (setq alterations (𝒱𝒸 variations line))

     ;; c.f. ⁉ above
     ;; Now that the aterations have been parsed, let's attach
     ;; the new components of the PackageFormer being made.

      ;; :kind ≈ The vocabulary that replaces “PackageFormer”.
      (funcall ⁉ 'kind 'string-please)

      ;; :waist ≈ The division between parameters and remaining elements.
      (funcall ⁉ 'waist)

      ;; :level ≈ Either 'inc or 'dec, for increment or decrementing the level.
      (funcall ⁉ 'level nil ;; 'string-please
         '((let* ((lvl (pf--package-former-level self))
                  (toLevel (lambda (n)
                             (s-join "" (-concat (-repeat n "Level.suc (")
                                                 (list "Level.zero")
                                                 (-repeat n ")")))))
                 (subs `("" "₁" "₂" "₃" "₄" "₅" "₆" "₇" "₈" "₉"
                         ,(funcall toLevel 10)))
                 (here (-elem-index (s-trim lvl) subs)))
             (setq it
                   (if here

                       (pcase it
                         ('inc (nth (1+ here) subs))
                         ('dec (nth (1- here) subs)))

                     (pcase it
                       ('inc (format "Level.suc (%s)" lvl))
                       ('dec (s-join "suc"
                                     (cdr (s-split "suc" lvl :omit-nulls)))))))

             (unless it (setq it 'none)))))

      ;; :alter-elements ≈ Access the typed name constituents list.
      ;; Perform *all* element alterations,
      ;; in the left-to-right ⟴ order; if any at all.
      (loop for ae in (reverse (mapcar #'cdr (--filter
                                              (equal ':alter-elements (car it))
                                              alterations)))
            do
            (setq $𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠
                  (cl-remove-duplicates (--filter it (funcall ae $𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠))
                                     :test #'equal :from-end t)))
              ;; Filter in only the non-nil constituents & those not
              ;; starting with “--” & remove duplicates.
              ;; We do this each time, rather than at the end, since
              ;; variationals may loop over all possible elements and we do not
              ;;  want to consider intermediary nils or duplicates.
      (setf (pf--package-former-elements self) $𝑒𝑙𝑒𝑚𝑒𝑛𝑡𝑠)

      ;; We've just formed a new PackageFormer, which can be modified,
      ;; specialised, later on.
      (add-to-list 'pf--package-formers (cons $𝑛𝑎𝑚𝑒 self))
      (if show-it (pf--show-package-former self) $𝑛𝑎𝑚𝑒))))

8.2 pf--load-700-comments and lisp blocks

We already two global lists —for our loaded PackageFormers and varitionals declarared— we need two additional globals: One for the contents of 700-comments, primarily for basic efficiency, and the other to port whatever is within 700-comments but is not 700-syntax. The latter is for when we want definitions or generalised variables to be accessible in both 700-comments and the surrounding script. Which is achieved since the matter is exported to the generated file in the Elisp stage, then imported at the Agda level.

However, since content to be ported may occur between PackageFormer definitions and instance declarations, we need to remember the position of items that are ported. The simplest thing to do is simply place items for porting into the package-formers list —then take care to check if an item is for porting or not when printing PackageFormers to a generated file.

(defvar pf--annotations nil
  "The contents of the PackageFormer annotations.

If this variable does not change, we short-circut all processing.")
  • Parse comments of the form “{-700 ⋯ -}” and add all PackageFormer declarations to the ‘package-formers’ list and load all instantations to the list as well.
  • We also execute any valid Lisp code in “{-lisp -}” comments; which may contain an arbitrary number of Lisp forms —a ‘progn’ is auto provided.

    Lisp is executed before any pf–annotations are; which is preferable due to Lisp's dynamic scope.

(cl-defun pf--load-pf--annotations ()
  "Parse and load {-700⋯-} syntax.
<<docs('load-pf--annotations)>>"
  (interactive)

  ;; First, let's run all the lisp. We enclose each in a progn in-case the user
  ;; has multiple forms in a single lisp-block.
  (loop for (lispstart . lispend) in '(("^\{-lisp" . "^-\}"))
        do (loop for lispstr in
                 (pf--buffer-substring-delimited-whole-buffer lispstart lispend)
                 do
                 (eval (car (read-from-string (format "(progn %s)" lispstr))))))

  ;; For now, ‘item’ is a PackageFormer,
  ;; instantiation declaration, or other Agda code.
  (let (item lines 700-cmnts)

    ;; Catenate all pf--annotations into a single string.
    (setq 700-cmnts (s-join "\n"
                            (pf--buffer-substring-delimited-whole-buffer
                             "^\{-700"
                             "^-\}")))

    (if (equal pf--annotations 700-cmnts)

        (message "pf--annotations Unchanged.")

      ;; Update global.
      (setq pf--annotations 700-cmnts)

      ;; View comments as a sequence of lines, ignore empty lines
      ;; ---which are not in our grammar.
      (setq lines (--remove (s-blank? (s-collapse-whitespace it))
                            (s-lines pf--annotations)))

      ;; Traverse the pf--annotations:
      ;; ➩ Skip comments; lines starting with “-- ”.
      ;; ➩ If we see a “𝒱-lhs = rhs” equation, then load it as a variational.
      ;; ➩ If we view a “lhs = rhs” equation, then load it as an
      ;;   instance delcaration.
      ;; ➩ If we view a PackageFormer declaration, then load it into our
      ;;   package-formers list.
      (while lines
        (setq item (car lines))

        (if (not (s-blank? (s-shared-start "-- " (s-trim item))))
            (setq lines (cdr lines))

          (if (not (s-blank? (s-shared-start "𝒱-" item)))
              (progn (pf--load-variational item) (setq lines (cdr lines)))

            (if (s-contains? " = " item)
                (progn (pf--load-instance-declaration item)
                       (setq lines (cdr lines)))

              ;; Else we have a PackageFormer declaration
              ;; and other possiblly-non-700 items.
              (setq item (pf--get-children "PackageFormer" lines))

              ;; port non-700 items to generated file
              (push (cons 'porting (s-join "\n" (car item)))
                    pf--package-formers)

              ;; acknowledge PackageFormer declaration, if any
              (when (cadr item) (pf--load-package-former (cadr item)))

              ;; Update lines to be the unconsidered porition
              ;; of the wild comments.
              (setq lines (caddr item))))))

      (message "Finished parsing pf--annotations."))))

9 Emacs Interface

9.1 Tooltips

It gets rather tedious to jump to the generated files to see the elaborations of 700-syntactical items. As such, let's tie existing occurrences of a PackageFormer's name to its elaboration.

We only add tooltips to PHRASE as a standalone word, not as a subword.

The PHRASE is taken literally; no regexp operators are recognised.

Useful info on tooltips:

The second resource above shows how to alter the phrase's font to indicate that it has a tooltip. It is not desirable for us, since we want to add onto Agda's colouring.

;; Nearly instantaneous display of tooltips.
(setq tooltip-delay 0)

;; Give user 30 seconds before tooltip automatically disappears.
(setq tooltip-hide-delay 30)

(defun pf--tooltipify (phrase notification)
  "Add a tooltip to every instance of PHRASE to show NOTIFICATION.
<<docs('tooltipify)>>"
  (should (stringp phrase))
  (should (stringp notification))
  (save-excursion  ;; Return cursour to current-point afterwards.
    (goto-char 1)
    ;; The \b are for empty-string at the start or end of a word.
    (while (search-forward-regexp (format "\\b%s\\b" (regexp-quote phrase)) (point-max) t)
      (put-text-property (match-beginning 0)
                         (match-end 0)
                         'help-echo
                         (s-trim notification)))))

New textual occurrences of a name obtain tooltips when C-c C-l is invoked.

9.2 Advising our Beloved C-c C-l

Let's give the current buffer access to the location of the generated file.

(defun pf--insert-generated-import (name-of-generated-file)
  "Insert an import pointing to the generated file.

In the current file, find the top-most module declaration
then insert an import to NAME-OF-GENERATED-FILE."
  (interactive)
  (save-excursion
    (goto-char (point-min))
    (condition-case _
        ;; attemptClause:
        (re-search-forward (concat "open import " name-of-generated-file))
      ;; recoveryBody:
      (error ;; (message-box (format "%s" the-err))
       (re-search-forward "\\(module.*\\)")
       (replace-match (concat "\\1\nopen import " name-of-generated-file))))))

The aim is to process test enclosed in {-700 ⋯ -} comments, produce legitimate Agda from that, and ensure the generated Agda is accessible to the current buffer automatically.

(defvar pf-waiting-for-agda-threshhold 14 "\
How long we should wait for Agda before giving up on colouring and tooltips.

Default is to wait 4 × 0.5 milliseconds.
Why? An inital ‘agda2-load’ of a ~300 line file may take some time.")

(defvar pf-multiple-files nil
  "Whether the list of PackageFormers should be reset with each “C- c C-l”, or
   should be left persistant as in when we are loading over multiple files.")

(defun pf--reify-package-formers (orig-fun &rest args)
  "Parse package-former syntax and produce Agda when possible.

ORIG-FUN is intended to be the Agda loading process with arguments ARGS."
  (interactive)
  (let* (printed-pfs
        (parent-dir (s-join "/"
                            (-drop-last 1 (s-split "/" buffer-file-truename))))
        (generatedmodule  (pf--generated-file-name))
        (newfile (concat parent-dir "/" generatedmodule ".agda"))
        (parent-imports (pf--extract-imports)))

    ;; Load variationals, PackageFormers, instantiations, and porting list.
    ;; Setting the following to nil each time is not ideal.
    (setq   pf--variationals
             (-take-last ♯standard-variationals pf--variationals)
             ;; take last n items, those being exported into the .el.
             pf--package-formers       (when (or pf-consider-newly-named-declarations-only pf-multiple-files) pf--package-formers)
             pf--annotations           nil)

    (pf--load-pf--annotations)

    (with-temp-buffer
      (goto-char (point-min))

      ;; Copy/paste imports from parent file.
      (insert (s-join "\n" `(
         "{- This file is generated ;; do not alter. -}\n"
         ,parent-imports
         "open import Level as Level using (Level)"
         ,(format "module %s where " generatedmodule))))

     ;; Print the package-formers
      (setq printed-pfs
            (--map (cond
               ((equal 'porting (car it)) (format "%s" (cdr it)))
               ((equal "_" (car it))      "") ;; Anonymous PackageFormers
               (t (format
                (if (equal "PackageFormer" (pf--package-former-kind (cdr it)))
                    (concat "{- Kind “PackageFormer” does not correspond "
                            " to a concrete Agda type. \n%s\n -}")
                       "%s") (pf--show-package-former (cdr it)))))
             (reverse pf--package-formers)))
      ;;
      (insert (s-join "\n\n\n" printed-pfs))
      ;; (setq package-formers nil) ;; So no accidental

      ;; Replace tabs with spaces
      (untabify (point-min) (point-max))

      (write-region (goto-char (point-min)) (goto-char (point-max)) newfile))

    (pf--insert-generated-import generatedmodule)

    ;; Need to revert buffer to discard old colours.
    ;; (save-buffer) (revert-buffer t t t)

    ;; call agda2-load
    (apply orig-fun args)

  ;; Agda attaches “jump to definition” tooltips; we add to those.
  ;; For some reason we need a slight delay between when Agda is done checking
  ;; and when we can add on our tooltips.
  ;; Attach tooltips only for existing occurrences; update happens with C-c C-l.
  ;; Wait until Agda is finished highlighting, then do ours (งಠ_ಠ)ง
  (-let [counter 0] ; agda2-in-progress
    (while agda2-highlight-in-progress
      (when (> counter pf-waiting-for-agda-threshhold)
        (error (concat "PackageFormer ∷ "
                       "Items generated, but not coloured; "
                       "Agda seems busy...")))
        (incf counter)
        (sleep-for 0.5))) ;; In case Agda errors on a term, no more waiting.
    (loop for (name . pf) in pf--package-formers
          do (unless (equal 'porting name)
               (pf--tooltipify name (pf--show-package-former pf))))
    ;; Special anonymous names.
    (pf--tooltipify "_" "“_” is the name of anonymous PackageFormers, which cannot be elaborated.")

    ;; Let's also add tooltips for the variationals & colour them.
    (loop for (v . docs) in pf--variationals
          do (pf--tooltipify (format "%s" v) docs)
          ;; For beauty, let's colour variational names green.
          ;; Only colour occurances that have a space before or after.
          (when pf-highlighting
            (highlight-phrase (format "[- \\| ]%s " v) 'hi-green)))

    (message "PackageFormer ∷ All the best coding! (•̀ᴗ•́)و")))

Personal note: Using (write-file "Generated.agda") means we make a file then the temporary buffer visits the Agda file, which loads the Agda process therein, which is undesirable since it could leave Agda working on the buffer even after it has been killed!

  • This would necessiate calling (agda2-restart) afterwards.
  • Instead we write the whole region, without visiting the resuting file.

Let's quickly add a menu bar that allows users to enable or disable using PackageFormer's; along with a brief help menu.

Names with a double dash are intended to be private, not to be touched by users.

(defvar pf--menu-bar (make-sparse-keymap "PackageFormer"))
Enabling the feature
(define-key pf--menu-bar [pf-enable-package-formers]
  '(menu-item "Enable PackageFormer Generation" pf-enable-package-formers))

(defun pf-enable-package-formers ()
 "Add a menubar, and make Agda's 𝑪-𝑐 𝑪-𝒍 consider package-former syntax."
 (interactive)
 (define-key global-map [menu-bar pf--menu] (cons "PackageFormer" pf--menu-bar))
 (advice-add 'agda2-load :around #'pf--reify-package-formers)
 (message-box (concat "C-c C-l now reifies PackageFormer annotations"
                      "into legitimate Agda.")))
Disabling the feature
(define-key pf--menu-bar [pf-disable-package-formers]
  '(menu-item "Disable PackageFormer Generation" pf-disable-package-formers))

(defun pf-disable-package-formers ()
  "Remove menubar, and make Agda's 𝑪-𝑐 𝑪-𝒍 act as normal."
 (interactive)
 (define-key global-map [menu-bar pf--menu] nil)
 (advice-remove 'agda2-load #'pf--reify-package-formers)
 (setq global-mode-string (remove "PackageFormer (•̀ᴗ•́)و " global-mode-string))
  (message-box "C-c C-l now behaves as it always has."))
About menu
(define-key pf--menu-bar [pf-package-formers-about]
  '(menu-item "About PackageFormers" pf-package-formers-about))

(defun pf-package-formers-about ()
  "Show help about the PackageFormer system."
 (interactive)
 (switch-to-buffer "*PackageFormer-About*") (insert "\
This is an editor extension prototyping “the next 700 module systems”
proposed research.

An informal documentation, with examples, page can be found at
https://alhassy.github.io/next-700-module-systems/prototype/package-former.html

The technical matter can be found at
https://alhassy.github.io/next-700-module-systems/

If you experience anything “going wrong” or have any ideas for improvement,
please contact Musa Al-hassy at alhassy@gmail.com; thank-you ♥‿♥"))
Bare Bones Agda
(define-key pf--menu-bar [pf--bare-bones]
  '(menu-item "Copy file with PackageFormer annotations stripped away"
              pf--bare-bones))

(defun pf-bare-bones ()
  "Duplicate the current buffer with PackageFormer annotations stripped away."
 (interactive)
 (let* ((src (file-name-sans-extension (buffer-name)))
        (src-agda (format "%s.agda" src))
        (bare-agda (format "%s-bare.agda" src)))
   (with-temp-buffer
     (insert-file-contents src-agda)
     (goto-char (point-min))
       (re-search-forward (format "module %s" src))
       (replace-match (format "module %s-bare" src))
     (loop for pre in '("^\{-lisp" "^\{-700")
      do
      (goto-char (point-min))
      (pf--buffer-substring-delimited-whole-buffer pre "^-\}"
           (lambda (sp ep)
             (save-excursion
             (goto-char (- sp 2))
             (push-mark ep)
             (setq mark-active t)
             (delete-region (- sp 2) ep)))))
     (write-file bare-agda))
     (message "%s_Bare.agda has been written." src)))
Menu of Defined PackageFormer Contents
(define-key pf--menu-bar [pf-show-variationals]
  '(menu-item "Show all registered variationals" pf-show-variationals))

(defun pf-show-variationals ()
  "Show all user declared 𝒱ariationals in another buffer."
 (interactive)
 (occur "𝒱[ \\|-]"))

(define-key pf--menu-bar [pf-show-pfs]
  '(menu-item "Show all concrete PackageFormers" pf-show-pfs))

(defun pf-show-pfs ()
  "Show all user declared PackageFormer's in another buffer."
 (interactive)
 (occur "PackageFormer .* where"))
Folding Away PackageFormer Annotations
(define-key pf--menu-bar [pf-fold-annotations]
  '(menu-item "Toggle folding away “700” and “lisp” blocks"
              pf-fold-annotations))

(defun pf-fold-annotations ()
  "Fold all items enclosed in Agda comments “{- ⋯ -}”."
  (interactive)
  (setq pf-folding (not pf-folding))
  (if pf-folding
      (message "Use “C-c f t” to toggle folding blocks")
    (origami-open-all-nodes (current-buffer))
    (message "Blocks “700” and “lisp” have been unfolded.")))

;; Basic origami support for Agda.
(push (cons 'agda2-mode (origami-markers-parser "{-" "-}"))
      origami-parser-alist)

;; Along with a hydra for super quick navigation
;; and easily folding, unfolding blocks!
(defhydra folding-with-origami-mode (global-map "C-c f")
  ("h" origami-close-node-recursively "Hide")
  ("o" origami-open-node-recursively  "Open")
  ("t" origami-toggle-all-nodes  "Toggle buffer")
  ("n" origami-next-fold "Next")
  ("p" origami-previous-fold "Previous"))

Let's pack these together into a minor mode.

Minor mode
;;;###autoload
(define-minor-mode agda-next-700-module-systems-mode "\
An editor extension prototyping “the next 700 module systems” proposed research.

An informal documentation, with examples, page can be found at
https://alhassy.github.io/next-700-module-systems-proposal/PackageFormer.html

The technical matter can be found at
https://alhassy.github.io/next-700-module-systems-proposal/

If you experience anything “going wrong” or have any ideas for improvement,
please contact Musa Al-hassy at alhassy@gmail.com; thank-you ♥‿♥"
  ;; Icon to display indicating the mode is enabled.
  :lighter " PackageFormer (•̀ᴗ•́)و"
  :require 'foo

  ;; Toggle the menu bar
  (define-key global-map [menu-bar pf--menu]
    (and agda-next-700-module-systems-mode (cons "PackageFormer" pf--menu-bar)))

  (letf (( (symbol-function 'message-box) #'message))
    (if agda-next-700-module-systems-mode
        ;; Initilisation
        (pf-enable-package-formers)

      ;; Closing
      (pf-disable-package-formers))))

The file ought to also contain one or more autoload magic comments, as explained in Packaging Basics.

10 Future & Related Work

Well, that was a lot of Lisp I had to learn (งಠ_ಠ)ง

Hopefully the resulting prototype will be useful to others; drop me a line if you're interested in this effort or have any feedback or pointers!

★ ★ ★

10.1 Features Wish-list

We check them off, rather than deleting them, to produce a ‘tada’ list of features ;-)

  • [X] For now, arguments must be enclosed in parenthesis.
    • [X] This is likely to be dropped in a future edition of the system.
  • [X] Colouring for 700-syntactical items.
  • [X] Tooltips
  • [ ] Some menu items are a bit buggy —this is low-priority.
  • [ ] MA: WK: hiding ↦ dropping
  • [ ] Use pf–buffer-substring-delimited-whole-buffer to parse multiple pf–annotations!
  • [ ] Add a level component to an instance structure; reduce it if Carrier is a parameter and otherwise leave it alone. Instead, let the level of a PackageFormer denote the level of the typeclass instantiation, then with this in mind we increase the level component of an instance structure only for those variations that keep the carrier as a field. When we move to multi-sorted, as in Graphs, this issue will need to be revisited.
  • [X] Currently can only perform simple variational clauses; need to support complex clauses.
  • [X] Support for equations in PackageFormers: Easy, nothing new needed! Thank-you alter-elements!
  • [X] Examples from the user-manual should be exported into an Agda file which is then easily checked; or the resulting generated code should be generated when export happens.
  • [ ] We can only perform “LHS = PF 𝓋𝒸”, do we also want stand-alone “PF 𝓋𝒸”?
    • Maybe not? Maybe it suffices to use, say, “_ = PF 𝓋𝒸”?
  • [X] Need to implement a front-end system to extend variational clauses.
  • [ ] The elements of a PackageFormer should be an alist, '(name . type).
    • Maybe not, the front-end lets users treat them as strings, easily.
    • Maybe such a view is useful for the user? It is provided via get-name and get-type.
  • [X] Refactor load-instantions function to make use of an alter field rather than overloading alter-fields in reify-instances.
  • [X] Refactor instantiate to make use of an instance structure, rather than 13 arguments.
    • In my defence, 13 arguments was because I was trying to find out what were the key ingredients necessary to form different variational presentations. It seems I only needed 5 ^_^
  • [ ] Why does ‘pf–buffer-substring-delimited-whole-buffer’ return a list of strings? Why not join its result to simply return a list?
  • [X] Generated.agda needs to inherit all open/import declarations from parent.
  • [ ] Give an example PackageFormer with some definitions or derived constructs, then hoist-up your waist so that the non-defined items are module parameters.
  • [X] PF's should account for equations. For simplicity they become components of a record & module, but derived operations on datatypes.
  • [ ] Demonstrate how generative modules can be emulated.
  • [ ] The global variables package-formers & instance-declaration should be buffer specific?
  • [X] Assign to a local var, check equality against global pf–annotations, if identical, no more processing since everything already generated.
  • [ ] Make use of docstring so that when a user enters a key sequence or selects from a menu, we can show them a listing of all ‘major’ components of a program: An org-mode file is displayed with an enumeration of the items, each being a link to the source, and only their docstring is shown. This is a nice ‘overview’ of the program source.
  • [X] Currently it looks like we are a minor mode, but this is not true. We only have a menu and the icon (•̀ᴗ•́)و is displayed when our feature is supported.
    • Now a minor mode.
  • [ ] Generate all instances of a PackageFormer schema.
  • MR = M-Set primed crashes!
  • [ ] Look into Specware.
    • Structures are created using colimits.
    • Are hierarchical considerations leaked by the resulting structures?
  • [X]

    PackageFormer.ext ↦ package-former.ext

    ⇨ kebab-case is conventional of Lisp

  • [X]

    Arguments, in tooltip renditions, are in reverse order; c.f. 𝒱-open.

    This is because 𝒱 lists them in the wrong order!

  • [X] colouring should be whole words only; see tooltipify
  • [X]

    Improved concept of PackageFormer constituents

    Element ≈ [qualifier] [lhs] [:] [type] [list-of definitional clauses]

    Eqn ≈ [same-lhs-as-before] [args] ≈ [term]

  • [X] extensions of structures can be formed;
  • [X] Form sum of two PackageFormers!
  • [X]

    Form library of common or useful variations or combinators, using the meta-primitives.

    renaming
    Given a function f, apply it to all names.
    map-elements f
    Apply function f to each element.
    prefix-elements
    Given a list of new elements, prepend them to the current list of elements.
    postfix-elements
    Given a list of new elements, postpend them to the current list of elements.
    drop-elements
    Given a predicate, drop the fields that satisfy it.
    • Moreover, drop additional fields if they depend on dropped fields? (maybe).
    opening f
    Module renaming according to f.
    manifest
    Mention one form; briefly mention others.

    Note that package formation has been liberated from the backend and brought to the user via our 5 meta-primitives: preamble, kind, waist, waist-strings, level, alter-elements.

  • [X] Allow “:level” to be ‘inc, dec, omit’ where the last option omits the “: Set ℓ” declaration altogether.
    • Useful for module setups.
  • [X] Make a “raw” variation that allows meta-primitives to be invoked whenever.
    • Maybe called 𝓌𝒾𝓉𝒽 ? ➱ Look at load variations, if there are any 𝓌𝒾𝓉𝒽 clauses, then add them to the variation's pairs. assoc ensures only latest items are picked up and so these would overwrite whatever the variation wanted; neato. Perhaps check that 𝓌𝒾𝓉𝒽 clauses only allow meta-primitives?
      • This is accomplished by a combination of a meta-primitive followed by “⟴”.
      • Unless needed otherwise, this is a non-issue.