├── POPL19.pdf
├── report.pdf
├── slides.pdf
├── deepspecslides.pdf
├── TheoriesAndDataStructures.pdf
├── Makefile
├── Helpers
├── Function2.lagda
├── LeqLemmas.agda
├── Forget.lagda
├── Hardy.lagda
├── EqualityCombinators.lagda
├── DataProperties.lagda
├── FinEquivTypeEquiv.agda
└── Equiv.lagda
├── ltx
├── sectref.sty
├── edcomms.sty
└── theref.sty
├── cheat-sheet.org
├── README.md
├── Structures
├── Experiments
│ ├── FiniteFunctions.agda
│ ├── SetoidEquiv.lagda
│ ├── Proofs.agda
│ ├── Sidequest
│ │ ├── Permutations
│ │ │ ├── FinUtils.lagda
│ │ │ ├── Testing.lagda
│ │ │ ├── BagEquality.lagda
│ │ │ ├── ActionProperties.lagda
│ │ │ ├── PermutationSequences.lagda
│ │ │ ├── Vector.lagda
│ │ │ ├── Basic.lagda
│ │ │ └── Group.lagda
│ │ ├── SummerHole.lagda
│ │ └── Equality.lagda
│ ├── SetoidSetoid.lagda
│ ├── SetoidOfIsos.lagda
│ ├── SubstLemmas.agda
│ ├── RATH.lagda
│ ├── Data
│ │ ├── Product
│ │ │ └── Properties2.agda
│ │ └── Sum
│ │ │ └── Properties2.agda
│ ├── CommMonoid-SET.lagda
│ ├── EquivEquiv.agda
│ ├── ParComp.lagda
│ └── Dependent2.lagda
├── CarrierAlgebra.lagda
├── IndexedUnaryAlgebra.lagda
├── Contractable.lagda
├── Indistinguishable.lagda
├── CommMonoid.lagda
├── DistinguishedSubset.lagda
├── DynamicalSystem.lagda
├── First.lagda
└── Magma.lagda
├── .gitignore
├── tasks.org
├── OurEverything.agda
└── TandDS.tex
/POPL19.pdf:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/JacquesCarette/TheoriesAndDataStructures/HEAD/POPL19.pdf
--------------------------------------------------------------------------------
/report.pdf:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/JacquesCarette/TheoriesAndDataStructures/HEAD/report.pdf
--------------------------------------------------------------------------------
/slides.pdf:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/JacquesCarette/TheoriesAndDataStructures/HEAD/slides.pdf
--------------------------------------------------------------------------------
/deepspecslides.pdf:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/JacquesCarette/TheoriesAndDataStructures/HEAD/deepspecslides.pdf
--------------------------------------------------------------------------------
/TheoriesAndDataStructures.pdf:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/JacquesCarette/TheoriesAndDataStructures/HEAD/TheoriesAndDataStructures.pdf
--------------------------------------------------------------------------------
/Makefile:
--------------------------------------------------------------------------------
1 | all: TandDS.pdf
2 |
3 | TandDS.pdf: TandDS.tex
4 | pdflatex TandDS.tex
5 |
6 | slides: slides.tex
7 | pdflatex slides.tex
8 |
--------------------------------------------------------------------------------
/Helpers/Function2.lagda:
--------------------------------------------------------------------------------
1 | \section{Function2}
2 |
3 | \begin{code}
4 | module Helpers.Function2 where
5 |
6 | -- should be defined in Function ?
7 | infix 4 _$ᵢ _$ₑ_
8 |
9 | _$ᵢ : ∀ {a b} {A : Set a} {B : A → Set b} →
10 | ((x : A) → B x) → {x : A} → B x
11 | _$ᵢ f {x} = f x
12 |
13 | _$ₑ_ : ∀ {a b} {A : Set a} {B : A → Set b} →
14 | ({x : A} → B x) → (x : A) → B x
15 | _$ₑ_ f x = f {x}
16 | \end{code}
17 |
--------------------------------------------------------------------------------
/ltx/sectref.sty:
--------------------------------------------------------------------------------
1 | %% sectref.sty
2 | %
3 | % (C) 2002 by Wolfram Kahl
4 |
5 | \def\LABELS#1{}
6 |
7 | \def\MkCommand#1{\expandafter\def\csname #1\endcsname}
8 |
9 | \def\newsref#1#2{%
10 | \MkCommand{#1label}##1{\label{#1:##1}\LABELS{\textsl{[#1:##1]}}}%
11 | \MkCommand{#1ref}##1{#2\ref{#1:##1}}%
12 | \MkCommand{#1rawref}##1{\ref{#1:##1}}%
13 | \MkCommand{#1pref}##1{\pageref{#1:##1}}%
14 | }
15 | \newsref{subsect}{}
16 | \newsref{sect}{\sectionname~}
17 | \newsref{chapt}{\chaptername~}
18 |
19 | \def\partname{Part}
20 | \newsref{part}{\partname~}
21 |
--------------------------------------------------------------------------------
/cheat-sheet.org:
--------------------------------------------------------------------------------
1 | #+TITLE: org-mode cheat sheet
2 | #+AUTHOR: Jacques Carette
3 | #+EMAIL: carette@mcmaster.ca
4 |
5 | * Preamble
6 | Yes, there are others out there, but this way I can order things in the way
7 | I'm most likely to remember them!
8 |
9 | * org-mode specific
10 | ** Ordered by command
11 | shift-tab - cycle between outline views (whole doc)
12 | tab - fold/unfold single section (if on lhs of title)
13 | C-c C-t - mark task done
14 | f7 - compile latex to pdf
15 |
16 | ** Ordered by task
17 | cycle between outline views: shift-tab
18 | fold/unfold: tab
19 | task done: C-c C-t
20 | latex to pdf: f7
21 |
22 | * emacs specific
23 | ** Ordered by command
24 | C-x C-s: save
25 | C-x C-f: open file
26 |
27 | ** Ordered by task
28 | save: C-x C-s
29 | open: C-x C-f
30 |
--------------------------------------------------------------------------------
/ltx/edcomms.sty:
--------------------------------------------------------------------------------
1 | %
2 | % ``Editor's comments''
3 | %
4 | % Originally created for the RelMiCS Book [Brink-Kahl-Schmidt-1997]
5 | % by Wolfram Kahl
6 | %
7 | % Turned into stand-alone package on 14 Jan. 2000 by Wolfram Kahl
8 | %
9 | % 2007-07-28: Use editor's choices for \edcommsfalse.
10 |
11 | \newif\ifedcomms
12 | \edcommstrue
13 |
14 | \def\eddir#1{\raise1pt\hbox{\footnotesize #1}}
15 |
16 | \long\def\edcomm#1#2{\ifedcomms{\fbox{\bf [ #1:} {\sl #2} \fbox{\bf ]}}\fi}
17 | \def\edmark#1{\ifedcomms{\fbox{\bf [ #1 ]}}\fi}
18 |
19 | \long\def\edchange#1#2#3{\ifedcomms{\fbox{\bf [ #1: \eddir{replace:}} #2
20 | \fbox{\eddir{with:}} #3 \fbox{\bf ]}}\else{#3}\fi}
21 | \long\def\eddelete#1#2{\ifedcomms{\fbox{\bf [ #1: \eddir{delete:}} #2 \fbox{\bf ]}}\fi}
22 | \long\def\edinsert#1#2{\ifedcomms{\fbox{\bf [ #1: \eddir{insert:}} #2
23 | \fbox{\bf ]}}\else{#2}\fi}
24 |
25 |
26 | \def\ectt#1{{\rm\texttt{#1}}}
27 |
28 | \def\unfinished{{\ \fbox{\textbf{ \ \ ???\ \ }}\ }}
29 |
30 |
--------------------------------------------------------------------------------
/README.md:
--------------------------------------------------------------------------------
1 |
TheoriesAndDataStructures
2 |
3 | Showing how some simple mathematical theories naturally give rise to some common data-structures.
4 |
5 | Attempting to answer the following questions:
6 |
7 | - Why do lists pop-up more frequently to the average programmer than, say, their duals: bags?
8 |
9 | - More simply, why do unit and empty types occur so naturally? What about enumerations/sums and records/products?
10 |
11 | - Why is it that dependent sums and products do not pop-up expicitly to the average programmer? They arise naturally all the time as tuples and as classes.
12 |
13 | - How do we get the usual toolbox of functions and helpful combinators for a particular data type? Are they \`\`built into'' the type?
14 |
15 | - Is it that the average programmer works in the category of classical Sets, with functions and propositional equality? Does this result in some \`\`free constructions'' not easily made computable since mathematicians usually work in the category of Setoids but tend to quotient to arrive in \`Sets\` —where quotienting is not computably feasible, in \`Sets\` at-least; and why is that?
16 |
17 | And lots of other stuff.
18 |
--------------------------------------------------------------------------------
/Structures/Experiments/FiniteFunctions.agda:
--------------------------------------------------------------------------------
1 | {-# OPTIONS --without-K #-}
2 |
3 | module FiniteFunctions where
4 |
5 | open import Data.Vec using (tabulate; _∷_)
6 | open import Data.Fin using (Fin; zero; suc)
7 | open import Data.Nat using (ℕ; suc)
8 | open import Relation.Binary.PropositionalEquality
9 | using (_≡_; refl; cong; module ≡-Reasoning)
10 | open import Function using (_∘_)
11 |
12 | ------------------------------------------------------------------------------
13 | -- Important: Extensionality for finite functions
14 |
15 | finext : {n : ℕ} {A : Set} → {f g : Fin n → A} → ((i : Fin n) → f i ≡ g i) →
16 | (tabulate f ≡ tabulate g)
17 | finext {0} _ = refl
18 | finext {suc n} {_} {f} {g} fi≡gi =
19 | begin (tabulate {suc n} f
20 | ≡⟨ refl ⟩
21 | f zero ∷ tabulate {n} (f ∘ suc)
22 | ≡⟨ cong (λ x → x ∷ tabulate {n} (f ∘ suc)) (fi≡gi zero) ⟩
23 | g zero ∷ tabulate {n} (f ∘ suc)
24 | ≡⟨ cong (_∷_ (g zero))
25 | (finext {f = f ∘ suc} {g ∘ suc} (fi≡gi ∘ suc)) ⟩
26 | g zero ∷ tabulate {n} (g ∘ suc)
27 | ≡⟨ refl ⟩
28 | tabulate g ∎)
29 | where open ≡-Reasoning
30 |
31 | ------------------------------------------------------------------------------
32 |
--------------------------------------------------------------------------------
/Helpers/LeqLemmas.agda:
--------------------------------------------------------------------------------
1 | {-# OPTIONS --without-K #-}
2 |
3 | module Helpers.LeqLemmas where
4 |
5 | open import Data.Nat
6 | using (ℕ; suc; _+_; _*_; _<_; _≤_; _≤?_; z≤n; s≤s)
7 | open import Data.Nat.Properties.Simple using (+-comm)
8 | open import Data.Nat.Properties using (n≤m+n; module ≤-Reasoning)
9 | open import Relation.Binary using (Decidable)
10 |
11 | ------------------------------------------------------------------------------
12 | -- Proofs and definitions about ≤ on natural numbers
13 |
14 | _ to ⟨_,_⟩) hiding (Σ-syntax)
26 | Σ∶• : {a b : Level} (A : Set a) (B : A → Set b) → Set (a ⊍ b)
27 | Σ∶• = Data.Product.Σ
28 | infix -666 Σ∶•
29 | syntax Σ∶• A (λ x → B) = Σ x ∶ A • B
30 | \end{spec}
31 |
32 | From private repos of Wolfram Kahl,
33 |
34 | \begin{spec}
35 | import Relation.Binary.Indexed as IS
36 |
37 | record IsDepEquivalence {a b c ℓ} {I : Setoid a b} {A : Setoid.Carrier I → Set c}
38 | (_≈_ : IS.Rel A ℓ)
39 | : Set (a ⊍ b ⊍ c ⊍ ℓ) where
40 |
41 | open Setoid I using () renaming (Carrier to I₀ ; _≈_ to _≈ᵢ_)
42 | field
43 | reflexive : {i j : Setoid.Carrier I} → i ≈ᵢ j → (x : A i) → Σ y ∶ A j • x ≈ y
44 | sym : IS.Symmetric A _≈_
45 | trans : IS.Transitive A _≈_
46 |
47 | -- Derived component saying that |_≈_| is an Indexed equivalence relation:
48 | isEquivalence : IS.IsEquivalence A _≈_
49 | isEquivalence = record
50 | { refl = λ {i} {x} → let x≈x′ = proj₂ (reflexive (Setoid.refl I) x) in trans x≈x′ (sym x≈x′)
51 | ; sym = sym
52 | ; trans = trans
53 | }
54 |
55 | open IS.IsEquivalence isEquivalence public using (refl)
56 |
57 | record DepSetoid {a b : Level} (I : Setoid a b) c ℓ : Set (ℓsuc (c ⊍ a ⊍ b ⊍ ℓ)) where
58 | infix 4 _≈_
59 | open Setoid I using () renaming (Carrier to I₀ ; _≈_ to _≈ᵢ_)
60 | field
61 | Carrier : I₀ → Set c
62 | _≈_ : IS.Rel Carrier ℓ
63 | isDepEquivalence : IsDepEquivalence {I = I} {A = Carrier} _≈_
64 |
65 | open IsDepEquivalence isDepEquivalence public
66 | indexedSetoid : IS.Setoid I₀ c ℓ
67 | indexedSetoid = record { Carrier = Carrier; _≈_ = _≈_; isEquivalence = isEquivalence }
68 |
69 | ΣΣ : {s₁ s₂ c ℓ : Level} (I : Setoid c ℓ) → DepSetoid I s₁ s₂ → Setoid (s₁ ⊍ c) (s₂ ⊍ ℓ)
70 | ΣΣ I S = let open DepSetoid S in record
71 | { Carrier = Σ (Setoid.Carrier I) Carrier
72 | ; _≈_ = λ x y → proj₁ x ≈ᵢ proj₁ y × proj₂ x ≈ proj₂ y
73 | ; isEquivalence = record
74 | { refl = Setoid.refl I , refl
75 | ; sym = λ x≈y → Setoid.sym I (proj₁ x≈y) , sym (proj₂ x≈y)
76 | ; trans = λ x≈y y≈z → Setoid.trans I (proj₁ x≈y) (proj₁ y≈z) , trans (proj₂ x≈y) (proj₂ y≈z)
77 | }
78 | } where open Setoid I using () renaming (_≈_ to _≈ᵢ_)
79 | \end{spec}
80 |
81 | % Quick Folding Instructions:
82 | % C-c C-s :: show/unfold region
83 | % C-c C-h :: hide/fold region
84 | % C-c C-w :: whole file fold
85 | % C-c C-o :: whole file unfold
86 | %
87 | % Local Variables:
88 | % folded-file: t
89 | % eval: (fold-set-marks "%{{{ " "%}}}")
90 | % eval: (fold-whole-buffer)
91 | % fold-internal-margins: 0
92 | % end:
93 |
--------------------------------------------------------------------------------
/Structures/CarrierAlgebra.lagda:
--------------------------------------------------------------------------------
1 | \section{CarrierAlgebras: much ado about the identity}
2 |
3 | Since the forgetful functor is the identity functor, left and right are
4 | trivial.
5 |
6 | %{{{ Imports
7 | \begin{code}
8 | module Structures.CarrierAlgebra where
9 |
10 | open import Level renaming (suc to lsuc; zero to lzero)
11 | open import Function
12 |
13 | open import Categories.Category using (Category; module Category)
14 | open import Categories.Functor using (Functor; Contravariant)
15 | open import Categories.Adjunction using (Adjunction)
16 | open import Categories.Agda using (Sets)
17 | open import Categories.NaturalTransformation using () renaming (id to idT)
18 |
19 | open import Helpers.Function2 using (_$ᵢ; _$ₑ_)
20 | open import Helpers.DataProperties
21 | open import Helpers.EqualityCombinators
22 | \end{code}
23 | %}}}
24 |
25 | %{{{ Carrier and Hom
26 | \subsection{Definition}
27 | \begin{code}
28 | record Carrier {ℓ} : Set (lsuc ℓ) where
29 | constructor car
30 | field
31 | carrier : Set ℓ
32 |
33 | open Carrier
34 |
35 | record Hom {ℓ} (X Y : Carrier {ℓ}) : Set ℓ where
36 | constructor car-hom
37 | field
38 | mor : carrier X → carrier Y
39 |
40 | open Hom
41 | \end{code}
42 | %}}}
43 |
44 | %{{{ Category and forgetful functor U
45 | \subsection{Category and Forgetful Functor}
46 |
47 | \begin{code}
48 | Carriers : (ℓ : Level) → Category _ ℓ ℓ
49 | Carriers ℓ = record
50 | { Obj = Carrier
51 | ; _⇒_ = Hom
52 | ; _≡_ = λ F G → mor F ≐ mor G
53 | ; id = record { mor = id }
54 | ; _∘_ = λ F G → record { mor = mor F ∘ mor G }
55 | ; assoc = ≐-refl
56 | ; identityˡ = ≐-refl
57 | ; identityʳ = ≐-refl
58 | ; equiv = record { IsEquivalence ≐-isEquivalence }
59 | ; ∘-resp-≡ = ∘-resp-≐
60 | }
61 | where open Hom ; open import Relation.Binary using (IsEquivalence)
62 |
63 | Forget : (o : Level) → Functor (Carriers o) (Sets o)
64 | Forget _ = record
65 | { F₀ = carrier
66 | ; F₁ = mor
67 | ; identity = ≡.refl
68 | ; homomorphism = ≡.refl
69 | ; F-resp-≡ = _$ᵢ
70 | }
71 | \end{code}
72 | %}}}
73 |
74 | %{{{ Left and AdjLeft
75 | \subsection{Free Adjunction: Part 1 of a toolkit}
76 |
77 |
78 | \begin{code}
79 | Left : (ℓ : Level) → Functor (Sets ℓ) (Carriers ℓ)
80 | Left ℓ = record
81 | { F₀ = car
82 | ; F₁ = car-hom
83 | ; identity = ≐-refl
84 | ; homomorphism = ≐-refl
85 | ; F-resp-≡ = _$ₑ_
86 | }
87 | \end{code}
88 |
89 | \begin{code}
90 | AdjLeft : (ℓ : Level) → Adjunction (Left ℓ) (Forget ℓ)
91 | AdjLeft ℓ = record { unit = idT ; counit = idT ; zig = ≐-refl ; zag = ≡.refl }
92 | \end{code}
93 | %}}}
94 |
95 | %{{{ Right, diag, AdjRight
96 | \subsection{CoFree Adjunction}
97 |
98 | \begin{code}
99 | Right : (ℓ : Level) → Functor (Sets ℓ) (Carriers ℓ)
100 | Right ℓ = Left ℓ
101 |
102 | AdjRight : (ℓ : Level) → Adjunction (Forget ℓ) (Right ℓ)
103 | AdjRight ℓ = record { unit = idT ; counit = idT ; zig = ≡.refl ; zag = ≐-refl }
104 |
105 | \end{code}
106 |
107 | %}}}
108 |
109 | % Quick Folding Instructions:
110 | % C-c C-s :: show/unfold region
111 | % C-c C-h :: hide/fold region
112 | % C-c C-w :: whole file fold
113 | % C-c C-o :: whole file unfold
114 | %
115 | % Local Variables:
116 | % folded-file: t
117 | % eval: (fold-set-marks "%{{{ " "%}}}")
118 | % eval: (fold-whole-buffer)
119 | % fold-internal-margins: 0
120 | % end:
121 |
--------------------------------------------------------------------------------
/Structures/Experiments/Data/Product/Properties2.agda:
--------------------------------------------------------------------------------
1 | {-# OPTIONS --without-K #-}
2 |
3 | -- Note that this is properly named, but it does depend on our version of
4 | -- Equiv and TypeEquiv for a number of things.
5 |
6 | module Data.Product.Properties2 where
7 |
8 | open import Level
9 | open import Data.Unit using (⊤)
10 | open import Data.Product using (_×_; _,_; proj₁; proj₂)
11 | renaming (map to _×→_)
12 |
13 | import Relation.Binary.PropositionalEquality as P
14 | using (_≡_; refl; cong; cong₂)
15 | open import Function using (id; _∘_)
16 |
17 | open import Equiv using (_≐_)
18 | open import TypeEquiv
19 | using (unite⋆; uniti⋆; unite⋆′; uniti⋆′;
20 | swap⋆; assocl⋆; assocr⋆)
21 |
22 | infixr 12 _×∼_
23 |
24 | ------------------------------------------------------------------------------
25 |
26 | abstract
27 | id×id∼id : {A B : Set} → (id ×→ id) ≐ id {A = A × B}
28 | id×id∼id x = P.refl
29 |
30 | ×∘∼∘× : {A B C D E F : Set} →
31 | {f : A → C} {g : B → D} {h : C → E} {i : D → F} →
32 | (h ∘ f) ×→ (i ∘ g) ≐ (h ×→ i) ∘ (f ×→ g)
33 | ×∘∼∘× x = P.refl
34 |
35 | _×∼_ : {A B C D : Set} → {f₀ g₀ : A → B} {f₁ g₁ : C → D} →
36 | (e₁ : f₀ ≐ g₀) → (e₂ : f₁ ≐ g₁) →
37 | f₀ ×→ f₁ ≐ g₀ ×→ g₁
38 | _×∼_ e₁ e₂ x = P.cong₂ _,_ (e₁ (proj₁ x)) (e₂ (proj₂ x))
39 |
40 | unite⋆-coh : {A B : Set} {f : A → B} →
41 | unite⋆ {zero} ∘ (id ×→ f) ≐ f ∘ unite⋆
42 | unite⋆-coh x = P.refl
43 |
44 | -- and the 'converse', of sorts; g is used here because
45 | -- this is usually applied with g = f⁻¹
46 | uniti⋆-coh : {A B : Set} {g : A → B} →
47 | (id ×→ g) ∘ uniti⋆ {zero} ≐ (uniti⋆ ∘ g)
48 | uniti⋆-coh x = P.refl
49 |
50 | unite⋆′-coh : {A B : Set} {f : A → B} →
51 | unite⋆′ ∘ (f ×→ id) ≐ f ∘ unite⋆′ {_} {zero}
52 | unite⋆′-coh _ = P.refl
53 |
54 | uniti⋆′-coh : {A B : Set} {g : A → B} →
55 | (g ×→ id) ∘ uniti⋆′ {_} {zero} ≐ uniti⋆′ ∘ g
56 | uniti⋆′-coh x = P.refl
57 |
58 | -- wf = well-formed. Need a better name than 'coh' everywhere
59 | assocr⋆-wf : {A B C D E F : Set} →
60 | {f₀ : A → D} {f₁ : B → E} {f₂ : C → F} →
61 | assocr⋆ ∘ ((f₀ ×→ f₁) ×→ f₂) ≐ (f₀ ×→ (f₁ ×→ f₂)) ∘ assocr⋆
62 | assocr⋆-wf _ = P.refl
63 |
64 | assocl⋆-wf : {A B C D E F : Set} →
65 | {f₀ : A → D} {f₁ : B → E} {f₂ : C → F} →
66 | ((f₀ ×→ f₁) ×→ f₂) ∘ assocl⋆ ≐ assocl⋆ ∘ (f₀ ×→ (f₁ ×→ f₂))
67 | assocl⋆-wf _ = P.refl
68 |
69 | triangle×-right : {A B : Set} →
70 | unite⋆′ ×→ id {A = B} ≐ (id {A = A} ×→ unite⋆) ∘ assocr⋆
71 | triangle×-right _ = P.refl
72 |
73 | triangle×-left : {A B : Set} →
74 | uniti⋆′ ×→ id {A = B} ≐ assocl⋆ {A} ∘ (id ×→ uniti⋆)
75 | triangle×-left _ = P.refl
76 |
77 | pentagon×-right : {A B C D : Set} →
78 | assocr⋆ ∘ assocr⋆ {A × B} {C} {D} ≐ id ×→ assocr⋆ ∘ assocr⋆ ∘ assocr⋆ ×→ id
79 | pentagon×-right _ = P.refl
80 |
81 | pentagon×-left : {A B C D : Set} →
82 | assocl⋆ ∘ assocl⋆ {A} {B} {C × D} ≐ assocl⋆ ×→ id ∘ assocl⋆ ∘ id ×→ assocl⋆
83 | pentagon×-left _ = P.refl
84 |
85 | swap⋆-coh : {A B C D : Set} {f : A → C} {g : B → D} →
86 | swap⋆ ∘ (f ×→ g) ≐ (g ×→ f) ∘ swap⋆
87 | swap⋆-coh _ = P.refl
88 |
89 | unite⋆-swap-coh-right : {A : Set} → unite⋆ {zero} {_} {A} ≐ unite⋆′ ∘ swap⋆
90 | unite⋆-swap-coh-right x = P.refl
91 |
92 | unite⋆-swap-coh-left : {A : Set} → uniti⋆ {zero} {A = A} ≐ swap⋆ ∘ uniti⋆′
93 | unite⋆-swap-coh-left x = P.refl
94 |
95 | hexagon×-right : {A B C : Set} →
96 | assocr⋆ {A = B} {C} {A} ∘ swap⋆ ∘ assocr⋆ ≐ (id ×→ swap⋆) ∘ assocr⋆ ∘ (swap⋆ ×→ id)
97 | hexagon×-right _ = P.refl
98 |
99 | hexagon×-left : {A B C : Set} →
100 | assocl⋆ {A = A} {B} {C} ∘ swap⋆ ∘ assocl⋆ ≐ (swap⋆ ×→ id) ∘ assocl⋆ ∘ (id ×→ swap⋆)
101 | hexagon×-left _ = P.refl
102 | ------------------------------------------------------------------------------
103 |
--------------------------------------------------------------------------------
/Structures/Experiments/Sidequest/Permutations/ActionProperties.lagda:
--------------------------------------------------------------------------------
1 | \section{Structures.Sidequest.Permutations.ActionProperties}
2 |
3 | Documenting properties of |_◈_| and |_◇_|; most notably the elimination results,
4 | \begin{spec}
5 | ◈-elimination : ∀ p xs ys → p ◈ xs ≈ₖ ys → xs ≈ₖ p ◇ ys
6 | ◇-elimination : ∀ p xs ys → p ◇ xs ≈ₖ ys → xs ≈ₖ p ◈ ys
7 | \end{spec}
8 |
9 | ( These names are the other way around wrt “inversion”; they ought to be swapped. c.f. Equality.lagda )
10 |
11 | %{{{ Imports
12 | \begin{code}
13 | open import Level using (Level)
14 | open import Relation.Binary using (Setoid)
15 |
16 | open import EqualityCombinators hiding (reflexive)
17 | open import Function using (_$_) renaming (id to Id₀ ; _∘_ to _∘₀_)
18 | open import DataProperties using (_‼_)
19 |
20 | open import Data.Vec
21 | open import Data.Nat hiding (fold ; _*_)
22 | open import Data.Fin hiding (_+_ ; fold ; _≤_)
23 |
24 | open import Structures.Sidequest.Permutations.Basic
25 | open import Structures.Sidequest.Permutations.Vector
26 |
27 | module Structures.Sidequest.Permutations.ActionProperties {ℓ c : Level} (𝒮 : Setoid c ℓ) where
28 |
29 | open import Structures.Sidequest.Equality 𝒮 renaming (_≈_ to _≈ₖ_)
30 |
31 | private
32 | open module ≈₀ = Setoid 𝒮 using (Carrier)
33 | Seq = Vec Carrier
34 | \end{code}
35 |
36 | Subscript 0 for ``underlying'', or `base', equality.
37 | %}}}
38 |
39 | %{{{ ◈-elimination and inversionTheorem
40 | \begin{code}
41 | ◇-cong₂ : {n m : ℕ} {ps : Permutation n m} {xs ys : Seq m}
42 | → xs ≈ₖ ys → ps ◇ xs ≈ₖ ps ◇ ys
43 | ◇-cong₂ []-cong = refl _
44 | ◇-cong₂ {ps = zero ∷ ps} (x≈y ∷-cong xs≈ys) = x≈y ∷-cong ◇-cong₂ xs≈ys
45 | ◇-cong₂ {ps = suc p ∷ ps} eq@(_ ∷-cong xs≈ys)
46 | = lookup-cong₂ {i = p} xs≈ys ∷-cong ◇-cong₂ (remove-cong₂ {i = suc p} eq)
47 |
48 | ◈-elimination : {n m : ℕ} (p : Permutation n m) (xs : Seq n) (ys : Seq m)
49 | → p ◈ xs ≈ₖ ys → xs ≈ₖ p ◇ ys
50 | ◈-elimination p xs _ eq = reflexive (≡.sym (inversionTheorem p xs)) ⟨≈ₖ≈⟩ ◇-cong₂ eq
51 | \end{code}
52 | %}}}
53 |
54 | %{{{ ◇-elimination and inversionTheorem˘
55 | The other form as well:
56 | \edcomm{WK}{There is no point making |ps| implicit!}
57 | \begin{code}
58 | ◈-cong₂ : {n m : ℕ} (ps : Permutation n m) {xs ys : Seq n}
59 | → xs ≈ₖ ys → ps ◈ xs ≈ₖ ps ◈ ys
60 | ◈-cong₂ _ []-cong = refl _
61 | ◈-cong₂ (p ∷ ps) {x ∷ xs} {y ∷ ys} (x≈y ∷-cong xs≈ys) = ≈ₖ-begin
62 | (p ∷ ps) ◈ (x ∷ xs)
63 | ≈ₖ⟨⟩
64 | insert (ps ◈ xs) p x
65 | ≈ₖ⟨ insert-cong₁ {_} {ps ◈ xs} {ps ◈ ys} {p} (◈-cong₂ ps xs≈ys) ⟩
66 | insert (ps ◈ ys) p x
67 | ≈ₖ⟨ insert-cong₃ {_} {ps ◈ ys} {p} {y} x≈y ⟩
68 | insert (ps ◈ ys) p y
69 | ≈ₖ⟨⟩
70 | (p ∷ ps) ◈ (y ∷ ys)
71 | □ₖ
72 |
73 | ◇-elimination : {n m : ℕ} (p : Permutation n m) (xs : Seq m) (ys : Seq n)
74 | → p ◇ xs ≈ₖ ys → xs ≈ₖ p ◈ ys
75 | ◇-elimination p xs ys p◇xs≈ys = ≈ₖ-begin
76 | xs
77 | ≈ₖ≡˘⟨ inversionTheorem˘ p xs ⟩
78 | p ◈ (p ◇ xs)
79 | ≈ₖ⟨ ◈-cong₂ p p◇xs≈ys ⟩
80 | p ◈ ys
81 | □ₖ
82 | \end{code}
83 | %}}}
84 |
85 | %{{{ Id is the unit of the actions
86 | \begin{code}
87 | Id-◈ : {n : ℕ} {xs : Seq n} → Id ◈ xs ≈ₖ xs
88 | Id-◈ {xs = [] } = []-cong
89 | Id-◈ {xs = _ ∷ _} = ≈₀.refl ∷-cong Id-◈
90 |
91 | Id-◇ : {m : ℕ} {xs : Seq m} → Id ◇ xs ≈ₖ xs
92 | Id-◇ {xs = [] } = []-cong
93 | Id-◇ {xs = _ ∷ _} = ≈₀.refl ∷-cong Id-◇
94 | \end{code}
95 | %}}}
96 |
97 |
98 | % Quick Folding Instructions:
99 | % C-c C-s :: show/unfold region
100 | % C-c C-h :: hide/fold region
101 | % C-c C-w :: whole file fold
102 | % C-c C-o :: whole file unfold
103 | %
104 | % Local Variables:
105 | % folded-file: t
106 | % eval: (fold-set-marks "%{{{ " "%}}}")
107 | % eval: (fold-whole-buffer)
108 | % fold-internal-margins: 0
109 | % end:
110 |
--------------------------------------------------------------------------------
/Structures/Experiments/Sidequest/SummerHole.lagda:
--------------------------------------------------------------------------------
1 | \section{Structures.Sidequest.SummerHole}
2 |
3 | The pesky hole from the summer of 2017.
4 |
5 | %{{{ Imports
6 | \begin{code}
7 | open import Structures.CommMonoid using (CommMonoid ; module CommMonoid ; module IsCommutativeMonoid)
8 |
9 | open import Level using (Level)
10 | open import Relation.Binary using (Setoid ; IsEquivalence)
11 | open import Function using (_$_)
12 | open import EqualityCombinators hiding (reflexive)
13 |
14 | -- open import Data.Fin using (Fin ; zero ; suc)
15 | open import Data.Nat hiding (_*_ ; fold)
16 | open import Data.Vec
17 |
18 | module Structures.Sidequest.SummerHole {s₁ s₂} (S : Setoid s₁ s₂) (C : CommMonoid S) where
19 |
20 | open import Structures.Sidequest.Permutations.Basic
21 | open import Structures.Sidequest.Permutations.Vector
22 | open import Structures.Sidequest.Permutations.ActionProperties S
23 | open import Structures.Sidequest.Permutations.Group S
24 | open import Structures.Sidequest.Equality S renaming (_≈_ to _≈ₖ_) hiding (refl ; trans)
25 | open import Structures.Sidequest.Permutations.BagEquality S
26 |
27 | open CommMonoid C
28 | open IsCommutativeMonoid isCommMonoid -- \edcomm{MA}{The field name really oughtn't be abbreviated!}
29 |
30 | private
31 | open module ≈₀ = Setoid S using (Carrier ; _≈_)
32 | Seq = Vec Carrier
33 | \end{code}
34 | %}}}
35 |
36 | \begin{code}
37 | -- fold is a setoid homomorphism
38 |
39 | fold : {n : ℕ} → Seq n → Carrier
40 | fold = foldr (λ _ → Carrier) _*_ e
41 |
42 | fold-cong : {n : ℕ} {xs ys : Seq n} → xs ≈ₖ ys → fold xs ≈ fold ys
43 | fold-cong []-cong = ≈.refl
44 | fold-cong (x≈y ∷-cong xs≈ys) = x≈y ⟨∙⟩ fold-cong xs≈ys
45 | -- commutativity is not used here and so this result is valid for non-commutative monoids as well.
46 |
47 | open import Relation.Binary.SetoidReasoning
48 |
49 | -- commutativity here!
50 | proposition₄ : {n : ℕ} {zs : Seq n} {x y : Carrier}
51 | → fold (x ∷ y ∷ zs) ≈ fold (y ∷ x ∷ zs)
52 | proposition₄ {n} {zs} {x} {y} = begin⟨ S ⟩
53 | fold (x ∷ y ∷ zs)
54 | ≈˘⟨ assoc _ _ _ ⟩
55 | (x * y) * fold zs
56 | ≈⟨ comm _ _ ⟨∙⟩ ≈.refl ⟩
57 | (y * x) * fold zs
58 | ≈⟨ assoc _ _ _ ⟩
59 | fold (y ∷ x ∷ zs)
60 | ∎
61 |
62 | open import Data.Fin hiding (_+_ ; fold ; _≤_)
63 |
64 | proposition₃ : {n : ℕ} {xs : Seq n} {i : Fin (suc n)} {x y : Carrier}
65 | → fold (x ∷ y ∷ xs) ≈ fold (y ∷ insert xs i x)
66 | proposition₃ {.0} {[]} {zero} = proposition₄
67 | proposition₃ {.0} {[]} {suc ()}
68 | proposition₃ {.(suc _)} {x ∷ xs} {zero} = proposition₄
69 | proposition₃ {.(suc _)} {hd ∷ xs} {suc i} {x} {y} = begin⟨ S ⟩
70 | fold (x ∷ y ∷ hd ∷ xs)
71 | ≈⟨ proposition₄ ⟩
72 | fold (y ∷ x ∷ hd ∷ xs)
73 | ≡⟨ ≡.refl ⟩
74 | y * fold (x ∷ hd ∷ xs)
75 | ≈⟨ ≈.refl ⟨∙⟩ proposition₃ ⟩
76 | y * fold (hd ∷ insert xs i x)
77 | ≡⟨ ≡.refl ⟩
78 | fold (y ∷ hd ∷ insert xs i x)
79 | ∎
80 |
81 | proposition₂ : {n : ℕ} {xs : Seq n} {i : Fin (suc n)} {x : Carrier}
82 | → fold (x ∷ xs) ≈ fold (insert xs i x)
83 | proposition₂ {.0} {[]} {zero} = ≈.refl
84 | proposition₂ {.0} {[]} {suc ()}
85 | proposition₂ {.(suc _)} {y ∷ xs} {zero} = ≈.refl
86 | proposition₂ {.(suc _)} {y ∷ xs} {suc i} = proposition₃
87 |
88 | open import Relation.Binary.PropositionalEquality using (inspect; [_])
89 |
90 | proposition₁ : {n : ℕ} {xs : Seq n} {p : Permutation n n} → fold xs ≈ fold (p ◈ xs)
91 | proposition₁ {xs = []} {[]} = ≈.refl
92 | proposition₁ {xs = x ∷ xs} {zero ∷ ps} = ≈.refl ⟨∙⟩ proposition₁
93 | proposition₁ {xs = x ∷ xs} {suc p ∷ ps} with ps ◈ xs | inspect (_◈_ ps) xs
94 | proposition₁ {_} {x ∷ xs} {suc () ∷ ps} | [] | _
95 | proposition₁ {_} {x ∷ xs} {suc p ∷ ps} | x′ ∷ xs′ | [ ps◈xs≈xs′ ] = begin⟨ S ⟩
96 | x * fold xs
97 | ≈⟨ ≈.refl ⟨∙⟩ proposition₁ ⟩
98 | x * fold (ps ◈ xs)
99 | ≡⟨ ≡.cong (λ zs → x * fold zs) ps◈xs≈xs′ ⟩
100 | x * fold (x′ ∷ xs′)
101 | ≡⟨ ≡.refl ⟩
102 | fold (x ∷ x′ ∷ xs′)
103 | ≈⟨ proposition₄ ⟩
104 | fold (x′ ∷ x ∷ xs′)
105 | ≡⟨ ≡.refl ⟩
106 | x′ * fold (x ∷ xs′)
107 | ≈⟨ ≈.refl ⟨∙⟩ proposition₂ ⟩
108 | x′ * fold (insert xs′ p x)
109 | ∎
110 |
111 | -- This is essentially |Multiset.fold-permute|, the pesky-hole from the summer.
112 | proposition₀ : {n : ℕ} {xs ys : Seq n} → xs ≈ₚ ys → fold xs ≈ fold ys
113 | proposition₀ (MkEq p p◈xs≈ys) = ≈.trans proposition₁ (fold-cong p◈xs≈ys)
114 | \end{code}
115 |
116 | % Quick Folding Instructions:
117 | % C-c C-s :: show/unfold region
118 | % C-c C-h :: hide/fold region
119 | % C-c C-w :: whole file fold
120 | % C-c C-o :: whole file unfold
121 | %
122 | % Local Variables:
123 | % folded-file: t
124 | % eval: (fold-set-marks "%{{{ " "%}}}")
125 | % eval: (fold-whole-buffer)
126 | % fold-internal-margins: 0
127 | % end:
128 |
--------------------------------------------------------------------------------
/Helpers/Forget.lagda:
--------------------------------------------------------------------------------
1 | \DeclareUnicodeCharacter{8718}{\ensuremath{ 8718 }}
2 |
3 | \section{Obtaining Forgetful Functors}
4 |
5 | We aim to realise a ``toolkit'' for an data-structure by considering a free
6 | construction and proving it adjoint to a forgetful functor. Since the majority
7 | of our theories are built on the category |Set|, we begin my making a helper
8 | method to produce the forgetful functors from as little information as needed
9 | about the mathematical structure being studied.
10 |
11 | Indeed, it is a common scenario where we have an algebraic structure with a single
12 | carrier set and we are interested in the categories of such structures along
13 | with functions preserving the structure.
14 |
15 | We consider a type of ``algebras'' built upon the category of |Sets|
16 | ---in that, every algebra has a carrier set and every homomorphism is a
17 | essentially a function between carrier sets where the composition of
18 | homomorphisms is essentially the composition of functions and the identity
19 | homomorphism is essentially the identity function.
20 |
21 | Such algebras consistute a category from which we obtain a method to
22 | Forgetful functor builder for single-sorted algebras to Sets.
23 |
24 | %{{{ Imports
25 | begin{ModuleHead}
26 | \begin{code}
27 | module Helpers.Forget where
28 |
29 | open import Level
30 |
31 | open import Helpers.Categorical
32 |
33 | open import Helpers.Function2
34 | open import Function
35 | open import Helpers.EqualityCombinators
36 | \end{code}
37 | end{ModuleHead}
38 | edcomm{MA}{For one reason or another, the module head is not making the imports smaller.}
39 | %}}}
40 |
41 | %{{{ OneSortedAlg
42 |
43 | A |OneSortedAlg| is essentially the details of a forgetful functor from
44 | some category to |Sets|,
45 |
46 | \begin{code}
47 | record OneSortedAlg (ℓ : Level) : Set (suc (suc ℓ)) where
48 | field
49 | Alg : Set (suc ℓ)
50 | Carrier : Alg → Set ℓ
51 | Hom : Alg → Alg → Set ℓ
52 | mor : {A B : Alg} → Hom A B → (Carrier A → Carrier B)
53 | comp : {A B C : Alg} → Hom B C → Hom A B → Hom A C
54 | .comp-is-∘ : {A B C : Alg} {g : Hom B C} {f : Hom A B} → mor (comp g f) ≐ mor g ∘ mor f
55 | Id : {A : Alg} → Hom A A
56 | .Id-is-id : {A : Alg} → mor (Id {A}) ≐ id
57 | \end{code}
58 | %}}}
59 |
60 | %{{{ oneSortedCategory
61 |
62 | The aforementioned claim that algebras and their structure preserving morphisms
63 | form a category can be realised due to the coherency conditions we requested viz
64 | the morphism operation on homomorphisms is functorial.
65 |
66 | \begin{code}
67 | open import Relation.Binary.SetoidReasoning
68 | oneSortedCategory : (ℓ : Level) → OneSortedAlg ℓ → Category (suc ℓ) ℓ ℓ
69 | oneSortedCategory ℓ A = record
70 | { Obj = Alg
71 | ; _⇒_ = Hom
72 | ; _≡_ = λ F G → mor F ≐ mor G
73 | ; id = Id
74 | ; _∘_ = comp
75 | ; assoc = λ {A B C D} {F} {G} {H} → begin⟨ ≐-setoid (Carrier A) (Carrier D) ⟩
76 | mor (comp (comp H G) F) ≈⟨ comp-is-∘ ⟩
77 | mor (comp H G) ∘ mor F ≈⟨ ∘-≐-cong₁ _ comp-is-∘ ⟩
78 | mor H ∘ mor G ∘ mor F ≈˘⟨ ∘-≐-cong₂ (mor H) comp-is-∘ ⟩
79 | mor H ∘ mor (comp G F) ≈˘⟨ comp-is-∘ ⟩
80 | mor (comp H (comp G F)) ∎
81 | ; identityˡ = λ{ {f = f} → comp-is-∘ ⟨≐≐⟩ Id-is-id ∘ mor f }
82 | ; identityʳ = λ{ {f = f} → comp-is-∘ ⟨≐≐⟩ ≡.cong (mor f) ∘ Id-is-id }
83 | ; equiv = record { IsEquivalence ≐-isEquivalence }
84 | ; ∘-resp-≡ = λ f≈h g≈k → comp-is-∘ ⟨≐≐⟩ ∘-resp-≐ f≈h g≈k ⟨≐≐⟩ ≐-sym comp-is-∘
85 | }
86 | where open OneSortedAlg A ; open import Relation.Binary using (IsEquivalence)
87 | \end{code}
88 | %}}}
89 |
90 | %{{{ mkForgetful
91 |
92 | The fact that the algebras are built on the category of sets is captured by the
93 | existence of a forgetful functor.
94 |
95 | \begin{code}
96 | mkForgetful : (ℓ : Level) (A : OneSortedAlg ℓ) → Functor (oneSortedCategory ℓ A) (Sets ℓ)
97 | mkForgetful ℓ A = record
98 | { F₀ = Carrier
99 | ; F₁ = mor
100 | ; identity = Id-is-id $ᵢ
101 | ; homomorphism = comp-is-∘ $ᵢ
102 | ; F-resp-≡ = _$ᵢ
103 | }
104 | where open OneSortedAlg A
105 | \end{code}
106 |
107 | That is, the constituents of a |OneSortedAlgebra| suffice to produce a category
108 | and a so-called presheaf as well.
109 | %}}}
110 |
111 | % Quick Folding Instructions:
112 | % C-c C-s :: show/unfold region
113 | % C-c C-h :: hide/fold region
114 | % C-c C-w :: whole file fold
115 | % C-c C-o :: whole file unfold
116 | %
117 | % Local Variables:
118 | % folded-file: t
119 | % eval: (fold-set-marks "%{{{ " "%}}}")
120 | % eval: (fold-whole-buffer)
121 | % fold-internal-margins: 0
122 | % end:
123 |
124 |
--------------------------------------------------------------------------------
/tasks.org:
--------------------------------------------------------------------------------
1 | #+TODO: TODO | STARTED | DONE
2 |
3 | The TODOs with respect to writing the paper.
4 |
5 | + [X] JC start learning about org mode
6 | + [X] JC Figure out how to expand collapsed entries
7 |
8 | + [X] JC Write introduction pass 1
9 | + [ ] JC Write outline pass 1
10 |
11 | + [X] MA To read:
12 | /From monoids to near-semirings: the essence of MonadPlus and Alternative/,
13 | https://usuarios.fceia.unr.edu.ar/~mauro/pubs/FromMonoidstoNearsemirings.pdf.
14 |
15 | + [ ] MA Write section on Monoids and lists
16 | + [X] MA code for Pointed
17 | + [X] MA code for Pointed + isContr
18 | + [X] MA code for Carrier + all equal
19 | + [ ] MA code for Unary
20 | + [ ] MA code for Involutive
21 | + [ ] MA code for Unary + cyclic-n rule (Involutive is n=2)
22 | + [ ] MA code for Unary + Pointed
23 | + [ ] MA code for Unary + Pointed + unit
24 | + [ ] WK other theories in the Carrier / Pointed / Unary /
25 | UnaryPointed realm?
26 | + [ ] MA code for Magma
27 | + [ ] MA code for semigroup (aka associative magma)
28 | + [ ] MA code for Pointed Magma
29 | + [ ] MA code for left-unital
30 | + [ ] MA code for left-unital semigroup
31 | + [ ] MA code for monoid (unital semigroup)
32 | + [ ] MA code for x*y = x Magma
33 | + [ ] MA code for y*x = x Magma
34 | + [ ] JC figure out which other Magmas should be looked at
35 | + [ ] MA code for Unary Magma
36 | + [ ] MA code for Unary Magma with -(y * x) = (-x) * (-y)
37 | + [ ] MA code for Involutive Magma with "
38 | + [ ] JC explain why commutative monoid (but also Magma) is a problem
39 | + [ ] Spot the patterns in all of the above code
40 | + [ ] MA Survey Haskell, Ocaml, Scala, Racket, Clojure for DS above;
41 | for 'kit' to go with data-structures
42 | + [ ] Write abstract
43 |
44 | Reminders:
45 | + Exploring Magma-based theories :: see
46 | https://en.wikipedia.org/wiki/Magma_(algebra)
47 | where we want to at least explore all the properties that are
48 | affine. These are interesting things said at
49 | https://en.wikipedia.org/wiki/Category_of_magmas which should be
50 | better understood.
51 |
52 |
53 | + unary theories ::
54 |
55 | wikipedia sure doesn't spend much time on these (see
56 | https://en.wikipedia.org/wiki/Algebraic_structure)
57 | but there are some interesting ones, because if the unary operation
58 | is 'f' things like
59 | forall x. f (f x) = x
60 | is *linear*, because x is used exactly once on each side. The
61 | non-linearity of 'f' doesn't count (else associativity wouldn't work
62 | either, as _*_ is used funnily there too). So "iter 17 f x = x" is a
63 | fine axiom here too. [iter is definable in the ground theory]
64 |
65 | This is actually where things started, as 'involution' belongs here.
66 |
67 | And is the first weird one.
68 |
69 | + Pointed unary theories :: E.g., the natural numbers
70 | + Pointer binary theories :: need to figure out which are expressible
71 | + more :: semiring, near-ring, etc. Need a sampling. But quasigroup
72 | (with 3 operations!) would be neat to look at.
73 |
74 |
75 | Also, I think we want to explore
76 | - Free Theories
77 | - Initial Objects
78 | - Cofree Theories (when they exist)
79 |
80 | Then the potential 'future work' is huge. But that can be left for
81 | later. We want to have all the above rock solid first.
82 |
83 | Attempting to answer the following questions:
84 |
85 | + Why do lists pop-up more frequently to the average programmer than,
86 | say, their commutative cousins: bags?
87 |
88 | + More simply, why do unit and empty types occur so naturally? What about enumerations/sums and records/products?
89 |
90 | + Why is it that dependent sums and products do not pop-up expicitly to the average programmer? They arise naturally all the time as tuples and as classes.
91 |
92 | + How do we get the usual toolbox of functions and helpful combinators for a particular data type? Are they ``built into'' the type?
93 |
94 | + Is it that the average programmer works in the category of classical Sets, with functions and propositional equality? Does this result in some ``free constructions'' not easily made computable since mathematicians usually work in the category of Setoids but tend to quotient to arrive in `Sets` ---where quotienting is not computably feasible, in `Sets` at-least; and why is that?
95 |
96 | (load-file "~/org-agda-mode/org-agda-mode.el")
97 | (load-file "~/org-agda-mode/literate.el")
98 |
99 | # Local Variables:
100 | # eval: (progn (org-babel-goto-named-src-block "make-acmart-class") (org-babel-execute-src-block) (outline-hide-sublevels 1))
101 | # eval: (global-set-key (kbd "") (lambda () (interactive) (org-babel-goto-named-src-block "make-readme") (org-babel-execute-src-block) (outline-hide-sublevels 1)))
102 | # eval: (global-set-key (kbd "") (lambda () (interactive) (org-babel-tangle) (async-shell-command (concat "open " (org-latex-export-to-pdf) ))))
103 | # End:
104 |
--------------------------------------------------------------------------------
/Structures/Experiments/CommMonoid-SET.lagda:
--------------------------------------------------------------------------------
1 | %% In this file, only CommMonoid up to CMonoidCat & Forget are defined.
2 | %% The issue is that this has no left adjoint in a constructive setting.
3 | \section{Structures.CommMonoid-SET}
4 |
5 | A SET based variant of |Structure.CommMonoid|.
6 |
7 | %{{{ Imports
8 | \begin{code}
9 | module Structures.CommMonoid-SET where
10 |
11 | open import Level renaming (zero to lzero; suc to lsuc ; _⊔_ to _⊍_) hiding (lift)
12 |
13 | open import Categories.Category using (Category)
14 | open import Categories.Functor using (Functor)
15 |
16 | open import Data.Product using (Σ; proj₁; proj₂; _,_)
17 | open import Function using (id ; _∘_)
18 |
19 | open import Relation.Binary.Sum
20 | import Algebra.FunctionProperties as AFP
21 | open AFP using (Op₂)
22 | \end{code}
23 | %}}}
24 |
25 | %{{{ CommMonoid ; Hom
26 | \subsection{Definitions}
27 |
28 | Some of this is borrowed from the standard library's |Algebra.Structures|
29 | and |Algebra|. But un-nested and made direct.
30 |
31 | Splitting off the properties is useful when defining structures which
32 | are commutative-monoid-like, but differ in other ways. The core
33 | properties can be re-used.
34 | \begin{code}
35 | open import EqualityCombinators
36 |
37 | record IsCommutativeMonoid {ℓ} {A : Set ℓ} (_∙_ : Op₂ A) (ε : A) : Set ℓ where
38 |
39 | open AFP (_≡_ {ℓ} {A = A})
40 | field
41 | left-unit : LeftIdentity ε _∙_
42 | right-unit : RightIdentity ε _∙_
43 | assoc : Associative _∙_
44 | comm : Commutative _∙_
45 | \end{code}
46 |
47 | There are many equivalent ways of defining a |CommMonoid|. But it
48 | boils down to this: Agda's dependent records are \textbf{telescopes}.
49 | Sometimes, one wants to identify a particular initial sub-telescope
50 | that should be shared between two instances. This is hard (impossible?)
51 | to do with holistic records. But if split, via |Σ|, this becomes
52 | easy.
53 |
54 | For our purposes, it is very convenient to split the |Setoid|
55 | part of the definition.
56 |
57 | \begin{code}
58 | record CommMonoid {ℓ} (X : Set ℓ) : Set (lsuc ℓ) where
59 | constructor MkCommMon
60 | field
61 | ε : X
62 | _∙_ : X → X → X
63 | -- \edcomm{MA}{The field name really oughtn't be abbreviated!}
64 | isCommMonoid : IsCommutativeMonoid _∙_ ε
65 |
66 | open IsCommutativeMonoid isCommMonoid public
67 |
68 | _⟨≡⟩_ = ≡.trans
69 |
70 | infix -666 eq-in
71 | eq-in = ≡._≡_
72 | syntax eq-in M x y = x ≡ y ∶ M -- ghost colon
73 |
74 | record Hom {ℓ : Level} {X Y : Set ℓ} (CMX : CommMonoid X) (CMY : CommMonoid Y) : Set ℓ where
75 | constructor MkHom
76 | open CommMonoid CMX using () renaming (ε to ε₁; _∙_ to _∙₁_ )
77 | open CommMonoid CMY using () renaming (ε to ε₂; _∙_ to _∙₂_ )
78 |
79 | field mor : X → Y
80 | field
81 | pres-ε : mor ε₁ ≡ ε₂
82 | pres-∙ : {x y : X} → mor (x ∙₁ y) ≡ mor x ∙₂ mor y
83 |
84 | cong : {a b : X} → a ≡ b → mor a ≡ mor b
85 | cong = ≡.cong mor
86 |
87 | open Hom using (mor)
88 | open Hom renaming (mor to _⟨$⟩_)
89 | \end{code}
90 |
91 | Notice that the last line in the record, |open Π mor public|, lifts the setoid-homomorphism
92 | operation |_⟨$⟩_| and |cong| to work on our monoid homomorphisms directly.
93 | %}}}
94 |
95 | %{{{ CommMonoidCat ; Forget
96 | \subsection{Category and Forgetful Functor}
97 | \begin{code}
98 | CommMonoidCat : (ℓ : Level) → Category (lsuc ℓ) ℓ ℓ
99 | CommMonoidCat ℓ = let open CommMonoid using (eq-in) in record
100 | { Obj = Σ (Set ℓ) CommMonoid
101 | ; _⇒_ = λ{ (X , CMX) (Y , CMY) → Hom CMX CMY}
102 | ; _≡_ = λ { {X , CMX} {Y , CMY} F G → ∀ {x : X} → F ⟨$⟩ x ≡ G ⟨$⟩ x }
103 | ; id = λ { {A , _} → MkHom id ≡.refl ≡.refl }
104 | ; _∘_ = λ { {C = _ , C} F G → let open CommMonoid C in record
105 | { mor = mor F ∘ mor G
106 | ; pres-ε = cong F (pres-ε G) ⟨≡≡⟩ pres-ε F
107 | ; pres-∙ = cong F (pres-∙ G) ⟨≡≡⟩ pres-∙ F
108 | } }
109 | ; assoc = ≡.refl
110 | ; identityˡ = ≡.refl
111 | ; identityʳ = ≡.refl
112 | ; equiv = λ { {_} {B , _} → record
113 | { refl = ≡.refl
114 | ; sym = λ F≈G → ≡.sym F≈G
115 | ; trans = λ F≈G G≈H → F≈G ⟨≡≡⟩ G≈H }
116 | }
117 | ; ∘-resp-≡ = λ { {C = C , _} {f = F} F≈F' G≈G' → cong F G≈G' ⟨≡≡⟩ F≈F' }
118 | }
119 | where open Hom
120 | \end{code}
121 |
122 | \begin{code}
123 | open import Categories.Agda using (Sets)
124 |
125 | Forget : (ℓ o : Level) → Functor (CommMonoidCat ℓ) (Sets ℓ)
126 | Forget ℓ o = record
127 | { F₀ = λ{ (X , CMX) → X }
128 | ; F₁ = λ F → mor F
129 | ; identity = ≡.refl
130 | ; homomorphism = ≡.refl
131 | ; F-resp-≡ = id
132 | }
133 | \end{code}
134 | %}}}
135 |
136 | % Quick Folding Instructions:
137 | % C-c C-s :: show/unfold region
138 | % C-c C-h :: hide/fold region
139 | % C-c C-w :: whole file fold
140 | % C-c C-o :: whole file unfold
141 | %
142 | % Local Variables:
143 | % folded-file: t
144 | % eval: (fold-set-marks "%{{{ " "%}}}")
145 | % eval: (fold-whole-buffer)
146 | % fold-internal-margins: 0
147 | % end:
148 |
--------------------------------------------------------------------------------
/Structures/Experiments/Data/Sum/Properties2.agda:
--------------------------------------------------------------------------------
1 | {-# OPTIONS --without-K #-}
2 |
3 | -- Note that this is properly named, but it does depend on our version of
4 | -- Equiv and TypeEquiv for a number of things.
5 |
6 | module Data.Sum.Properties2 where
7 |
8 | open import Level
9 |
10 | open import DataProperties using (⊥)
11 | open import Data.Sum using (_⊎_; inj₁; inj₂) renaming (map to map⊎)
12 |
13 | import Relation.Binary.PropositionalEquality as P using (_≡_; refl; cong; trans)
14 | open import Function as F using (id; _∘_)
15 |
16 | open import Equiv using (_≐_)
17 | open import TypeEquiv
18 | using (unite₊; uniti₊; unite₊′; uniti₊′;
19 | swap₊; assocl₊; assocr₊)
20 |
21 | infix 8 _⊎→_
22 | infix 12 _⊎∼_
23 |
24 | _⊎→_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} →
25 | (A → C) → (B → D) → (A ⊎ B → C ⊎ D)
26 | _⊎→_ = map⊎
27 |
28 | ------------------------------------------------------------------------------
29 | -- Note that all these lemmas are "simple" in the sense that they
30 | -- are all about ⊎→ (i.e. map⊎) rather than [_,_]
31 |
32 | abstract
33 |
34 | id⊎id∼id : {A B : Set} → (F.id {A = A} ⊎→ F.id {A = B}) ≐ F.id
35 | id⊎id∼id (inj₁ x) = P.refl
36 | id⊎id∼id (inj₂ y) = P.refl
37 |
38 | ⊎∘∼∘⊎ : {A B C D E F : Set} →
39 | {f : A → C} {g : B → D} {h : C → E} {i : D → F} →
40 | (h F.∘ f) ⊎→ (i F.∘ g) ≐ (h ⊎→ i) ∘ (f ⊎→ g)
41 | ⊎∘∼∘⊎ (inj₁ x) = P.refl
42 | ⊎∘∼∘⊎ (inj₂ y) = P.refl
43 |
44 | _⊎∼_ : {A B C D : Set} → {f₀ g₀ : A → B} {f₁ g₁ : C → D} →
45 | (e₁ : f₀ ≐ g₀) → (e₂ : f₁ ≐ g₁) →
46 | f₀ ⊎→ f₁ ≐ g₀ ⊎→ g₁
47 | _⊎∼_ f₀~g₀ _ (inj₁ x) = P.cong inj₁ (f₀~g₀ x)
48 | _⊎∼_ _ f₁~g₁ (inj₂ y) = P.cong inj₂ (f₁~g₁ y)
49 |
50 | unite₊-coh : {A B : Set} {f : A → B} {g : ⊥ → ⊥} →
51 | unite₊ {zero} ∘ (g ⊎→ f) ≐ f ∘ unite₊ {zero}
52 | unite₊-coh (inj₁ ())
53 | unite₊-coh (inj₂ y) = P.refl
54 |
55 | uniti₊-coh : {A B : Set} {f : A → B} {g : ⊥ → ⊥} →
56 | (g ⊎→ f) ∘ uniti₊ {zero} ≐ uniti₊ {zero} ∘ f
57 | uniti₊-coh x = P.refl
58 |
59 | unite₊′-coh : {A B : Set} {f : A → B} {g : ⊥ → ⊥} →
60 | unite₊′ {zero} ∘ (f ⊎→ g) ≐ f ∘ unite₊′ {zero}
61 | unite₊′-coh (inj₁ x) = P.refl
62 | unite₊′-coh (inj₂ ())
63 |
64 | uniti₊′-coh : {A B : Set} {f : A → B} {g : ⊥ {zero} → ⊥ {zero}} →
65 | (f ⊎→ g) ∘ uniti₊′ ≐ (uniti₊′ ∘ f)
66 | uniti₊′-coh x = P.refl
67 |
68 | assocr₊-wf : {A B C D E F : Set} →
69 | {f₀ : A → D} {f₁ : B → E} {f₂ : C → F} →
70 | assocr₊ ∘ ((f₀ ⊎→ f₁) ⊎→ f₂) ≐ (f₀ ⊎→ (f₁ ⊎→ f₂)) ∘ assocr₊
71 | assocr₊-wf (inj₁ (inj₁ x)) = P.refl
72 | assocr₊-wf (inj₁ (inj₂ y)) = P.refl
73 | assocr₊-wf (inj₂ y) = P.refl
74 |
75 | assocl₊-wf : {A B C D E F : Set} →
76 | {f₀ : A → D} {f₁ : B → E} {f₂ : C → F} →
77 | ((f₀ ⊎→ f₁) ⊎→ f₂) ∘ assocl₊ ≐ assocl₊ ∘ (f₀ ⊎→ (f₁ ⊎→ f₂))
78 | assocl₊-wf (inj₁ x) = P.refl
79 | assocl₊-wf (inj₂ (inj₁ x)) = P.refl
80 | assocl₊-wf (inj₂ (inj₂ y)) = P.refl
81 |
82 | triangle⊎-right : {A B : Set} →
83 | unite₊′ {zero} ⊎→ F.id {A = B} ≐ (F.id {A = A} ⊎→ unite₊ {zero}) ∘ assocr₊
84 | triangle⊎-right (inj₁ (inj₁ x)) = P.refl
85 | triangle⊎-right (inj₁ (inj₂ ()))
86 | triangle⊎-right (inj₂ y) = P.refl
87 |
88 | triangle⊎-left : {A B : Set} →
89 | uniti₊′ {zero} ⊎→ F.id {A = B} ≐ assocl₊ ∘ (F.id {A = A} ⊎→ uniti₊ {zero})
90 | triangle⊎-left (inj₁ x) = P.refl
91 | triangle⊎-left (inj₂ y) = P.refl
92 |
93 | pentagon⊎-right : {A B C D : Set} →
94 | assocr₊ ∘ assocr₊ {A = A ⊎ B} {C} {D} ≐ (F.id ⊎→ assocr₊) ∘ assocr₊ ∘ (assocr₊ ⊎→ F.id)
95 | pentagon⊎-right (inj₁ (inj₁ (inj₁ x))) = P.refl
96 | pentagon⊎-right (inj₁ (inj₁ (inj₂ y))) = P.refl
97 | pentagon⊎-right (inj₁ (inj₂ y)) = P.refl
98 | pentagon⊎-right (inj₂ y) = P.refl
99 |
100 | pentagon⊎-left : {A B C D : Set} →
101 | assocl₊ ∘ assocl₊ {A = A} {B} {C ⊎ D} ≐ (assocl₊ ⊎→ F.id) ∘ assocl₊ ∘ (F.id ⊎→ assocl₊)
102 | pentagon⊎-left (inj₁ x) = P.refl
103 | pentagon⊎-left (inj₂ (inj₁ x)) = P.refl
104 | pentagon⊎-left (inj₂ (inj₂ (inj₁ x))) = P.refl
105 | pentagon⊎-left (inj₂ (inj₂ (inj₂ y))) = P.refl
106 |
107 | swap₊-coh : {A B C D : Set} {f : A → C} {g : B → D} →
108 | swap₊ ∘ (f ⊎→ g) ≐ (g ⊎→ f) ∘ swap₊
109 | swap₊-coh (inj₁ x) = P.refl
110 | swap₊-coh (inj₂ y) = P.refl
111 |
112 | unite₊-swap-coh-right : {A : Set} → unite₊ {zero} {zero} {A} ≐ unite₊′ ∘ swap₊
113 | unite₊-swap-coh-right (inj₁ ())
114 | unite₊-swap-coh-right (inj₂ y) = P.refl
115 |
116 | unite₊-swap-coh-left : {A : Set} → uniti₊ {zero} {A = A} ≐ swap₊ ∘ uniti₊′
117 | unite₊-swap-coh-left x = P.refl
118 |
119 | hexagon⊎-right : {A B C : Set} →
120 | assocr₊ {A = B} {C} {A} ∘ swap₊ ∘ assocr₊ ≐ (F.id ⊎→ swap₊) ∘ assocr₊ ∘ (swap₊ ⊎→ F.id)
121 | hexagon⊎-right (inj₁ (inj₁ x)) = P.refl
122 | hexagon⊎-right (inj₁ (inj₂ y)) = P.refl
123 | hexagon⊎-right (inj₂ y) = P.refl
124 |
125 | hexagon⊎-left : {A B C : Set} →
126 | assocl₊ {A = A} {B} {C} ∘ swap₊ ∘ assocl₊ ≐ (swap₊ ⊎→ F.id) ∘ assocl₊ ∘ (F.id ⊎→ swap₊)
127 | hexagon⊎-left (inj₁ x) = P.refl
128 | hexagon⊎-left (inj₂ (inj₁ x)) = P.refl
129 | hexagon⊎-left (inj₂ (inj₂ y)) = P.refl
130 |
131 | ------------------------------------------------------------------------------
132 |
--------------------------------------------------------------------------------
/Helpers/Hardy.lagda:
--------------------------------------------------------------------------------
1 | \section{Items from Hardy's fork of the standard library}
2 |
3 | Material taken from Brad Hardy's variant of the standard library,
4 | https://github.com/bch29/agda-stdlib. In particular, from his
5 | module Algebra.Properties.CommutativeMonoid.
6 |
7 | \begin{code}
8 | module Helpers.Hardy where
9 |
10 | open import Algebra using (CommutativeMonoid)
11 | open import Function using (_∘_)
12 | open import Data.Nat using (ℕ; zero; suc)
13 | open import Data.Fin using (Fin; punchIn; zero; suc)
14 | open import Data.Fin.Permutation as Perm using (Permutation; Permutation′; _⟨$⟩ˡ_; _⟨$⟩ʳ_)
15 | open import Relation.Binary using (module Setoid)
16 |
17 | open import Data.Table as Table
18 | open import Data.Table.Relation.Equality as TE using (_≗_)
19 |
20 | open import Relation.Binary.PropositionalEquality as ≡ using (_≡_)
21 |
22 | module HardyCommutativeMonoid {g₁ g₂} (M : CommutativeMonoid g₁ g₂) where
23 |
24 | open CommutativeMonoid M renaming
25 | (ε to 0#; _∙_ to _+_; ∙-cong to +-cong; identity to +-identity; assoc to +-assoc; comm to +-comm)
26 | open import Algebra.Operations.CommutativeMonoid M using (sumₜ)
27 |
28 | import Relation.Binary.EqReasoning as EqR; open EqR setoid
29 |
30 | module _ {n} where
31 | open Setoid (TE.setoid setoid n) public using () renaming (_≈_ to _≋_)
32 |
33 | -- When summing over a function from a finite set, we can pull out any value and move it to the front.
34 |
35 | sumₜ-punchIn : ∀ {n} t (i : Fin (suc n)) → sumₜ t ≈ lookup t i + sumₜ (rearrange (punchIn i) t)
36 | sumₜ-punchIn f zero = refl
37 | sumₜ-punchIn {zero} t (suc ())
38 | sumₜ-punchIn {suc n} t (suc i) =
39 | let x = head t
40 | y = lookup t (suc i)
41 | z = sumₜ (rearrange (punchIn i) (tail t))
42 | in begin
43 | x + sumₜ (tail t) ≈⟨ +-cong refl (sumₜ-punchIn (tail t) i) ⟩
44 | x + (y + z) ≈⟨ sym (+-assoc _ _ _) ⟩
45 | (x + y) + z ≈⟨ +-cong (+-comm _ _) refl ⟩
46 | (y + x) + z ≈⟨ +-assoc _ _ _ ⟩
47 | y + (x + z) ∎
48 |
49 | -- '_≈_' is a congruence over 'sumTable n'.
50 |
51 | sumₜ-cong : ∀ {n} {t t′ : Table Carrier n} → t ≋ t′ → sumₜ t ≈ sumₜ t′
52 | sumₜ-cong {zero} p = refl
53 | sumₜ-cong {suc n} p = +-cong (p _) (sumₜ-cong (p ∘ suc))
54 |
55 | -- '_≡_' is a congruence over 'sum n'.
56 |
57 | sumₜ-cong≡ : ∀ {n} {t t′ : Table Carrier n} → t ≗ t′ → sumₜ t ≡ sumₜ t′
58 | sumₜ-cong≡ {zero} p = ≡.refl
59 | sumₜ-cong≡ {suc n} p = ≡.cong₂ _+_ (p _) (sumₜ-cong≡ (p ∘ suc))
60 |
61 | -- Any permutation of a table has the same sum as the original.
62 |
63 | sumₜ-permute : ∀ {n} t (π : Permutation′ n) → sumₜ t ≈ sumₜ (rearrange (π ⟨$⟩ʳ_) t)
64 | sumₜ-permute {zero} t π = refl
65 | sumₜ-permute {suc n} t π =
66 | let f = lookup t
67 | 0i = zero
68 | ππ0 = π ⟨$⟩ʳ (π ⟨$⟩ˡ 0i)
69 | in begin
70 | sumₜ t ≡⟨⟩
71 | f 0i + sumₜ (rearrange (punchIn 0i) t) ≈⟨ +-cong refl (sumₜ-permute _ (Perm.remove (π ⟨$⟩ˡ 0i) π)) ⟩
72 | f 0i + sumₜ (rearrange (punchIn 0i ∘ (Perm.remove (π ⟨$⟩ˡ 0i) π ⟨$⟩ʳ_)) t) ≡⟨ ≡.cong₂ _+_ ≡.refl (sumₜ-cong≡ (≡.cong f ∘ ≡.sym ∘ Perm.punchIn-permute′ π 0i)) ⟩
73 | f 0i + sumₜ (rearrange ((π ⟨$⟩ʳ_) ∘ punchIn (π ⟨$⟩ˡ 0i)) t) ≡⟨ ≡.cong₂ _+_ (≡.cong f (≡.sym (Perm.inverseʳ π))) ≡.refl ⟩
74 | f _ + sumₜ (rearrange ((π ⟨$⟩ʳ_) ∘ punchIn (π ⟨$⟩ˡ 0i)) t) ≈⟨ sym (sumₜ-punchIn (rearrange (π ⟨$⟩ʳ_) t) (π ⟨$⟩ˡ 0i)) ⟩
75 | sumₜ (rearrange (π ⟨$⟩ʳ_) t) ∎
76 |
77 | -- A version of 'sumₜ-permute' allowing heterogeneous sum lengths.
78 |
79 | sumₜ-permute′ : ∀ {m n} t (π : Permutation m n) → sumₜ t ≈ sumₜ (rearrange (π ⟨$⟩ʳ_) t)
80 | sumₜ-permute′ t π with Perm.↔⇒≡ π
81 | sumₜ-permute′ t π | ≡.refl = sumₜ-permute t π
82 | \end{code}
83 |
84 | ∑-permute : ∀ {n} f (π : Permutation′ n) → ∑[ i < n ] f i ≈ ∑[ i < n ] f (π ⟨$⟩ʳ i)
85 | ∑-permute = sumₜ-permute ∘ tabulate
86 |
87 | ∑-permute′ : ∀ {m n} f (π : Permutation m n) → ∑[ i < n ] f i ≈ ∑[ i < m ] f (π ⟨$⟩ʳ i)
88 | ∑-permute′ = sumₜ-permute′ ∘ tabulate
89 |
90 | -- The sum over the constantly zero function is zero.
91 | sumₜ-zero : ∀ n → sumₜ (replicate {n} 0#) ≈ 0#
92 |
93 | -- The '∑' operator distributes over addition.
94 | ∑-+-hom : ∀ n (f g : Fin n → Carrier) → ∑[ i < n ] f i + ∑[ i < n ] g i ≈ ∑[ i < n ] (f i + g i)
95 |
96 | -- The '∑' operator commutes with itself.
97 | ∑-comm : ∀ n m (f : Fin n → Fin m → Carrier) → ∑[ i < n ] ∑[ j < m ] f i j ≈ ∑[ j < m ] ∑[ i < n ] f i j
98 |
99 | -- One-point rule: Summing over a pulse gives you the single value picked out by the pulse.
100 | sumₜ-select : ∀ {n i} (t : Table Carrier n) → sumₜ (select 0# i t) ≈ lookup t i
101 |
102 | -- Converting to a table then summing is the same as summing the original list
103 | sumₜ-fromList : ∀ xs → sumₜ (fromList xs) ≡ sumₗ xs
104 |
105 | -- Converting to a list then summing is the same as summing the original table
106 | sumₜ-toList : ∀ {n} (t : Table Carrier n) → sumₜ t ≡ sumₗ (toList t)
107 |
108 | -- If addition is idempotent on a particular value 'x', then summing over any
109 | -- arbitrary number of copies of 'x' gives back 'x'.
110 |
111 | sumₜ-idem-replicate : ∀ n {x} → _+_ IdempotentOn x → sumₜ (Table.replicate {suc n} x) ≈ x
112 |
113 | % Quick Folding Instructions:
114 | % C-c C-s :: show/unfold region
115 | % C-c C-h :: hide/fold region
116 | % C-c C-w :: whole file fold
117 | % C-c C-o :: whole file unfold
118 | %
119 | % Local Variables:
120 | % folded-file: t
121 | % eval: (fold-set-marks "%{{{ " "%}}}")
122 | % eval: (fold-whole-buffer)
123 | % fold-internal-margins: 0
124 | % end:
125 |
--------------------------------------------------------------------------------
/Helpers/EqualityCombinators.lagda:
--------------------------------------------------------------------------------
1 | \section{Equality Combinators}
2 |
3 | Here we export all equality related concepts, including those for propositional and
4 | function extensional equality.
5 |
6 | \begin{code}
7 | module Helpers.EqualityCombinators where
8 |
9 | open import Level
10 | \end{code}
11 |
12 | %{{{ Propositional Equality
13 |
14 | \subsection{Propositional Equality}
15 |
16 | We use one of Agda's features to qualify all propositional equality properties
17 | by ``|≡.|'' for the sake of clarity and to avoid name clashes with similar other
18 | properties.
19 |
20 | \begin{code}
21 | import Relation.Binary.PropositionalEquality
22 | module ≡ = Relation.Binary.PropositionalEquality
23 | open ≡ using (_≡_) public
24 | \end{code}
25 |
26 | We also provide two handy-dandy combinators for common uses of transitivity proofs.
27 |
28 | \begin{code}
29 | infixr 5 _⟨≡≡⟩_ _⟨≡≡˘⟩_
30 |
31 | _⟨≡≡⟩_ = ≡.trans
32 |
33 | _⟨≡≡˘⟩_ : {a : Level} {A : Set a} {x y z : A} → x ≡ y → z ≡ y → x ≡ z
34 | x≈y ⟨≡≡˘⟩ z≈y = x≈y ⟨≡≡⟩ ≡.sym z≈y
35 | \end{code}
36 |
37 | Besides brevity, another reason for this naming is that transitivity
38 | acts as a group operator with inverses provided by for symmetry
39 | and identity is the reflexitivity proof. See trans-reflʳ for example
40 | from the standard library.
41 |
42 | Here's a nifty result: The cong operatrion is functorial in its first argument
43 | via function composition and identity function, but its also functorial in
44 | its second argument a la the previously mentioned group!
45 | The fact “cong p refl ≈ refl” is true by definition, and the second functor law
46 | is as follows:
47 |
48 | \begin{code}
49 | {- Maybe make a PR to agda-std-lib -}
50 | cong-over-trans : ∀ {i j} {A : Set i} {B : Set j} {f : A → B}
51 | → {x y z : A} {p : x ≡ y} (q : y ≡ z)
52 | → ≡.cong f p ⟨≡≡⟩ ≡.cong f q
53 | ≡ ≡.cong f (p ⟨≡≡⟩ q)
54 | cong-over-trans {p = ≡.refl} ≡.refl = ≡.refl
55 | \end{code}
56 |
57 | %}}}
58 |
59 | %{{{ Function Extensionality
60 | \subsection{Function Extensionality}
61 |
62 | We bring into scope pointwise equality, |_≐_|, and provide a proof that it constitutes
63 | an equivalence relation ---where the source and target of the functions being compared
64 | are left implicit.
65 |
66 | \begin{code}
67 | open ≡ using () renaming (_→-setoid_ to ≐-setoid ; _≗_ to _≐_) public
68 |
69 | open import Relation.Binary using (IsEquivalence ; Setoid)
70 |
71 | module _ {a b : Level} {A : Set a} {B : Set b} where
72 |
73 | ≐-isEquivalence : IsEquivalence (_≐_ {A = A} {B = B})
74 | ≐-isEquivalence = record {Setoid (≐-setoid A B) }
75 |
76 | open IsEquivalence ≐-isEquivalence public
77 | renaming (refl to ≐-refl ; sym to ≐-sym ; trans to ≐-trans)
78 |
79 | open import Helpers.Equiv public using (∘-resp-≐) -- To do: port this over here!
80 | renaming (cong∘l to ∘-≐-cong₂ ; cong∘r to ∘-≐-cong₁)
81 |
82 | infixr 5 _⟨≐≐⟩_
83 | _⟨≐≐⟩_ = ≐-trans
84 | \end{code}
85 |
86 | Note that the precedence of this last operator is lower than that of function composition so as to avoid superfluous parenthesis.
87 |
88 | Here is an implicit version of extensional
89 | ---we use it as a transitionary tool since the standard library and the category theory library differ
90 | on their uses of implicit versus explicit variable usage.
91 |
92 | \begin{code}
93 | infixr 5 _≐ᵢ_
94 | _≐ᵢ_ : {a b : Level} {A : Set a} {B : A → Set b}
95 | (f g : (x : A) → B x) → Set (a ⊔ b)
96 | f ≐ᵢ g = ∀{x} → f x ≡ g x
97 | \end{code}
98 | %}}}
99 |
100 | %{{{ Equiv
101 | \subsection{Equiv}
102 |
103 | We form some combinators for HoTT like reasoning.
104 |
105 | \begin{code}
106 | cong₂D : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Set c}
107 | (f : (x : A) → B x → C)
108 | → {x₁ x₂ : A} {y₁ : B x₁} {y₂ : B x₂}
109 | → (x₂≡x₁ : x₂ ≡ x₁) → ≡.subst B x₂≡x₁ y₂ ≡ y₁ → f x₁ y₁ ≡ f x₂ y₂
110 | cong₂D f ≡.refl ≡.refl = ≡.refl
111 |
112 | open import Helpers.Equiv public using (_≃_; id≃; sym≃; trans≃; qinv)
113 |
114 | infix 3 _◻
115 | infixr 2 _≃⟨_⟩_
116 |
117 | _≃⟨_⟩_ : {x y z : Level} (X : Set x) {Y : Set y} {Z : Set z}
118 | → X ≃ Y → Y ≃ Z → X ≃ Z
119 | X ≃⟨ X≃Y ⟩ Y≃Z = trans≃ X≃Y Y≃Z
120 |
121 | _◻ : {x : Level} (X : Set x) → X ≃ X
122 | X ◻ = id≃
123 | \end{code}
124 |
125 | \edcomm{MA}{Consider moving pertinent material here from |Equiv.lagda| at the end.}
126 | %}}}
127 |
128 | %{{{ _⟨≈⁺≈⁺⟩_
129 | \subsection{More Equational Reasoning for |Setoid|}
130 |
131 | A few convenient combinators for equational reasoning in |Setoid|.
132 |
133 | \savecolumns
134 | \begin{code}
135 | module SetoidCombinators {ℓS ℓs : Level} (S : Setoid ℓS ℓs) where
136 | open Setoid S
137 | _⟨≈≈⟩_ = trans
138 | _⟨≈˘≈⟩_ : {a b c : Carrier} → b ≈ a → b ≈ c → a ≈ c
139 | _⟨≈˘≈⟩_ = λ b≈a b≈c → sym b≈a ⟨≈≈⟩ b≈c
140 | _⟨≈≈˘⟩_ : {a b c : Carrier} → a ≈ b → c ≈ b → a ≈ c
141 | _⟨≈≈˘⟩_ = λ a≈b c≈b → a≈b ⟨≈≈⟩ sym c≈b
142 | _⟨≈˘≈˘⟩_ : {a b c : Carrier} → b ≈ a → c ≈ b → a ≈ c
143 | _⟨≈˘≈˘⟩_ = λ b≈a c≈b → b≈a ⟨≈˘≈⟩ sym c≈b
144 | \end{code}
145 | %}}}
146 |
147 | %{{{ inSetoidEquiv ; x ≈⌊ S ⌋ y
148 | \subsection{Localising Equality}
149 | Convenient syntax for when we need to specify which |Setoid|'s equality we are
150 | talking about.
151 |
152 | \begin{code}
153 | infix 4 inSetoidEquiv
154 | inSetoidEquiv : {ℓS ℓs : Level} → (S : Setoid ℓS ℓs) → (x y : Setoid.Carrier S) → Set ℓs
155 | inSetoidEquiv = Setoid._≈_
156 |
157 | syntax inSetoidEquiv S x y = x ≈⌊ S ⌋ y
158 | \end{code}
159 | %}}}
160 |
161 | % Quick Folding Instructions:
162 | % C-c C-s :: show/unfold region
163 | % C-c C-h :: hide/fold region
164 | % C-c C-w :: whole file fold
165 | % C-c C-o :: whole file unfold
166 | %
167 | % Local Variables:
168 | % folded-file: t
169 | % eval: (fold-set-marks "%{{{ " "%}}}")
170 | % eval: (fold-whole-buffer)
171 | % fold-internal-margins: 0
172 | % end:
173 |
--------------------------------------------------------------------------------
/Structures/Experiments/Sidequest/Permutations/PermutationSequences.lagda:
--------------------------------------------------------------------------------
1 | \section{|Structures.Sidequest.Permutations.PermutationSequences}
2 |
3 | The exploration here started from a failed attempt to interpret factorial numbers
4 | as sequences of side-by-side swaps (based on |_𝕩_|),
5 | therefore the original module name ``|CallChanges''.
6 |
7 | The approach implemented via |_𝕫_| based on complete rounds of rotations
8 | should have all the nice properties we can wish for.
9 |
10 | %{{{ Imports
11 | \begin{code}
12 | module Structures.Sidequest.Permutations.PermutationSequences where
13 |
14 | open import Level using (Level)
15 | open import Relation.Binary using (Setoid)
16 | import Relation.Binary.PropositionalEquality as ≡
17 | open ≡ using (_≡_)
18 |
19 | -- open import EqualityCombinators hiding (reflexive)
20 | open import Function using (_$_) renaming (id to Id₀ ; _∘_ to _∘₀_)
21 | open import DataProperties using (_‼_)
22 |
23 | open import Data.Maybe
24 | open import Data.List as List
25 | open import Data.Vec as Vec
26 | open import Data.Nat hiding (fold ; _*_)
27 | open import Data.Nat.Properties.Simple
28 | open import Data.Fin hiding (_+_ ; fold ; _≤_)
29 | open import Data.Product using (_,_)
30 |
31 | open import FinUtils using (suc′)
32 | open import Structures.Sidequest.Permutations.Basic
33 | open import Structures.Sidequest.Permutations.SME
34 | \end{code}
35 |
36 |
37 | \begin{code}
38 | lastPermutation : (n : ℕ) → Permutation n n
39 | lastPermutation zero = []
40 | lastPermutation (suc n) = fromℕ n ∷ lastPermutation n
41 | \end{code}
42 |
43 | \begin{code}
44 | {-# TERMINATING #-} -- The |inject₁| is in the way of the termination checker.
45 | permSME₀ : {n : ℕ} → Permutation n n → SME n n
46 | permSME₀ [] = 𝕀
47 | permSME₀ (_ ∷ []) = 𝕀
48 | permSME₀ {suc (suc n)} (zero ∷ is) = 𝕀 {1} ⊗ permSME₀ is
49 | permSME₀ {suc (suc n)} (suc iₙ ∷ is) = (𝕀 {1} ⊗ permSME₀ (lastPermutation (suc n))) ⨾ 𝕏′ {suc n} zero ⨾ permSME₀ (inject₁ iₙ ∷ is)
50 |
51 | permSME : {m n : ℕ} → Permutation m n → SME m n
52 | permSME p with homogeneity p
53 | ... | ≡.refl = permSME₀ p
54 | \end{code}
55 |
56 | \begin{code}
57 | {-# TERMINATING #-} -- The |inject₁| is in the way of the termination checker.
58 | perm𝕏s : {m n : ℕ} → (Fin m → Fin n) → Permutation (suc m) (suc m) → List (Fin n) → List (Fin n)
59 | perm𝕏s _f (_ ∷ []) = Id₀
60 | perm𝕏s {suc m} f (zero ∷ is) = perm𝕏s (f ∘₀ inject₁) is
61 | perm𝕏s {suc m} f (suc iₘ ∷ is)
62 | = perm𝕏s (f ∘₀ inject₁) (lastPermutation (suc m))
63 | ∘₀ (f (fromℕ m) ∷_)
64 | ∘₀ perm𝕏s f (inject₁ iₘ ∷ is)
65 | \end{code}
66 |
67 | \begin{code}
68 | permSME₁ : {m n : ℕ} → Permutation (suc m) (suc n) → SME (suc m) (suc n)
69 | permSME₁ p with homogeneity p
70 | ... | ≡.refl = List.foldr (λ i r → 𝕏′ i ⨾ r) 𝕀 (perm𝕏s Id₀ p [])
71 | \end{code}
72 |
73 | %{{{ noPermutation ; sucPermutation
74 | \begin{code}
75 | noPermutation : (n : ℕ) → Permutation n n
76 | noPermutation zero = []
77 | noPermutation (suc n) = zero ∷ noPermutation n
78 |
79 | sucPermutation : {n : ℕ} → Permutation n n → Maybe (Permutation n n)
80 | sucPermutation [] = nothing
81 | sucPermutation (i ∷ is) = maybe
82 | (λ is′ → just (i ∷ is′))
83 | (maybe (λ i′ → just (i′ ∷ noPermutation _)) nothing (suc′ i))
84 | (sucPermutation is)
85 | \end{code}
86 | %}}}
87 |
88 | \begin{code}
89 | perm : {n : ℕ} → Permutation (suc n) (suc n) → Vec ℕ (suc n)
90 | perm {n} p = Vec.map toℕ (permSME₁ p ◣ allFin (suc n))
91 | where
92 | open Action (≡.setoid (Fin (suc n)))
93 | \end{code}
94 |
95 | \begin{code}
96 | perms : {n : ℕ} → Permutation (suc n) (suc n) → List (Vec ℕ (suc n))
97 | perms {n} p = List.map (Vec.map toℕ) (execFinList _𝕫_ (perm𝕏s Id₀ p []) (allFin (suc n)))
98 | where
99 | open Action (≡.setoid (Fin (suc n)))
100 | \end{code}
101 | (Using |_𝕪_| instead of |_𝕩_| produces duplicates even earlier.)
102 |
103 | Using |_𝕫_| instead of |_𝕩_| produces no duplicates anymore,
104 | and should be easy to accelerate and to invert.
105 |
106 | Using the original |_◺_|,
107 | the following list of 15 permuted vectors takes seconds to generate via
108 | |C-c C-n|:
109 | \begin{spec}
110 | perms (zero ∷ suc (suc zero) ∷ suc zero ∷ suc zero ∷ zero ∷ [])
111 | \end{spec}
112 | (Using the current |_◺_|, both this and below are instanteneous.)
113 |
114 | Using the original |_◺_| involving |_𝕩_|,
115 | the 24 permuted vectors of the following take almost 400 seconds ---
116 | and contain duplicates!
117 | \begin{spec}
118 | perms (suc zero ∷ zero ∷ zero ∷ zero ∷ zero ∷ [])
119 | \end{spec}
120 |
121 | Example |perm𝕏s| results:
122 | \begin{spec}
123 | List.map toℕ (perm𝕏s Id₀ (suc zero ∷ zero ∷ zero ∷ zero ∷ zero ∷ []) [])
124 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
125 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
126 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
127 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 3 ∷ []
128 |
129 | List.map toℕ (perm𝕏s Id₀ (suc zero ∷ zero ∷ zero ∷ zero ∷ zero ∷ zero ∷ []) [])
130 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
131 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
132 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
133 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 3 ∷
134 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
135 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
136 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
137 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 3 ∷
138 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
139 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
140 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
141 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 3 ∷
142 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
143 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
144 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
145 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 3 ∷
146 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
147 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
148 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 2 ∷
149 | 0 ∷ 1 ∷ 0 ∷ 1 ∷ 0 ∷ 4 ∷ []
150 | \end{spec}
151 | (The |perm𝕏s| calculations have always been fast.)
152 |
153 |
154 | % Quick Folding Instructions:
155 | % C-c C-s :: show/unfold region
156 | % C-c C-h :: hide/fold region
157 | % C-c C-w :: whole file fold
158 | % C-c C-o :: whole file unfold
159 | %
160 | % Local Variables:
161 | % folded-file: t
162 | % eval: (fold-set-marks "%{{{ " "%}}}")
163 | % eval: (fold-whole-buffer)
164 | % fold-internal-margins: 0
165 | % end:
166 |
--------------------------------------------------------------------------------
/OurEverything.agda:
--------------------------------------------------------------------------------
1 | -----------------------------------------------------------------------------------------------
2 | -- The TheoriesAndDataStructures library
3 | --
4 | -- All library modules, along with short descriptions
5 | --
6 | -- Other experimental approaches can be found in the directory Experiments/
7 | --
8 | -- We are using Agda 2.6.0 with standard library version 1.0.1.
9 | -- Items marked with “‼” are not yet up to date and so have been commented out.
10 | -----------------------------------------------------------------------------------------------
11 |
12 | module OurEverything where
13 |
14 | -- Creates forgetful functors from single sorted algebras to |Set|
15 | --
16 | import Helpers.Forget
17 |
18 | import Helpers.Categorical
19 | import Helpers.Hardy
20 |
21 | -- ============================================================================================
22 | -- Variations on Sets
23 | -- ============================================================================================
24 |
25 | -- “The” one-object one-arrow category, along with general free, forgetful, and adjoint constructions.
26 | --
27 | -- Arrows in the one-object one-arrow category correspond to the functions to a singleton set.
28 | -- ‼ open import Structures.OneCat
29 |
30 | -- Category of algebras consisting of an index set and a family of sets on the index set.
31 | -- Along with forgetful functor to Sets with heterogenous equality and free functor, think “Σ”.
32 | --
33 | -- Currently has a hole.
34 | -- ‼ open import Structures.Dependent
35 |
36 | -- Category of heterogenous relations along with a few forgetful and (co)free functors and
37 | -- associated adjunction proofs.
38 | --
39 |
40 | open import Structures.Rel
41 |
42 | -- Category of algebras of a predicate furnished type; c.f., Dependents.
43 | -- Along with a forgetful functor to Sets.
44 | --
45 | -- Many simple-looking holes.
46 | -- ‼ open import Structures.DistinguishedSubset
47 |
48 | -- Category of pairs of sets with a few (co)free constructions.
49 | -- Along with a proof that sum is adjoint to squaring which is adjoint to product.
50 | --
51 | open import Structures.TwoSorted
52 |
53 | -- Category of pointed sets along with adjuntions with Sets and OneCat.
54 | --
55 | open import Structures.Pointed
56 |
57 | -- ============================================================================================
58 | -- Unary Algebras
59 | -- ============================================================================================
60 |
61 | -- Category of algebras consisting of a type endowed with an operator; along with
62 | -- adjunctions with Sets.
63 | -- The free structure corresponds to “performing a method” a number of times.
64 | --
65 | open import Structures.UnaryAlgebra
66 |
67 | -- Category of involutive algebras with a forgetful functor to Sets,
68 | -- adjunctions, and (co)monads.
69 | --
70 | open import Structures.InvolutiveAlgebra
71 |
72 | -- Category of algebras consisting of a carrier with a family of operations on it; i.e., actions.
73 | -- Along with a forgetful functor to Sets and an initial algebra.
74 | --
75 | --
76 | open import Structures.IndexedUnaryAlgebra
77 |
78 | -- ============================================================================================
79 | -- Algebras with binary operators, Boom Hierarchy
80 | -- ============================================================================================
81 |
82 | -- Interface for multisets over a given type, along with a free construction
83 | -- via sequences.
84 | --
85 | -- ‼ open import Structures.Baguette
86 |
87 | -- Category of Magmas along with explicit toolkit of operations for binary trees.
88 | -- Also initial & terminal objects, along with adjunctions with OneCat.
89 | --
90 | open import Structures.Magma
91 |
92 | -- Category of semigroups and an explicit theory of operators for non-empty lists.
93 | -- Along with adjunctions with Sets, Magmas, and OneCat.
94 | -- Also an involved proof of the non-existence of a certain adjunction:
95 | -- There cannot be a (free) associative “extension” of an arbitrary binary operator.
96 | --
97 | -- Contains a postulate.
98 | open import Structures.Semigroup
99 |
100 | -- Category of setoid-based commutative monoids over a given type, with a forgetful functor.
101 | -- Free constructions can be found in baguette.lagda .
102 | --
103 | -- ‼ open import Structures.CommMonoid
104 |
105 | -- Category of monoids with adjunctions with Sets and OneCat.
106 | -- There are holes left intentionally for exposition purposes.
107 | --
108 | open import Structures.Monoid
109 |
110 | -- A theory of bags; intend to be the free multisets.
111 | -- ( A difficult read! )
112 | --
113 | -- ‼ open import Structures.SequencesAsBags
114 |
115 | -- The category of Set-based Abelian groups, forgetful functor to Sets.
116 | -- Work towards free construction; no free functor yet.
117 | --
118 | -- Contains a postulate.
119 | -- ‼ open import Structures.AbelianGroup
120 |
121 | -- ============================================================================================
122 | -- ============================================================================================
123 | -- ============================================================================================
124 | -- Helpers
125 | --
126 | -- The remaining modules are mostly technical ones needed for the structure-theory relationships.
127 | -- ============================================================================================
128 |
129 | -- Re-exports all equality-related concepts
130 | --
131 | import Helpers.EqualityCombinators
132 |
133 | -- Contains properties about sums and products not found in standard library
134 | --
135 | import Helpers.DataProperties
136 |
137 | -- Implicit function application.
138 | --
139 | open import Helpers.Function2
140 |
141 | {-
142 | -- Files that asssit baguette.lagda ;; they are from previous work.
143 | -- Mostly laws about Fin and type isomorphisms for it.
144 | --
145 | open import Helpers.TypeEquiv
146 | open import Helpers.FinEquivPlusTimes
147 | open import Helpers.LeqLemmas
148 | open import Helpers.FinNatLemmas
149 | open import Helpers.FinEquivTypeEquiv
150 | -}
151 |
--------------------------------------------------------------------------------
/Helpers/DataProperties.lagda:
--------------------------------------------------------------------------------
1 | \section{Properties of Sums and Products}
2 |
3 | This module is for those domain-ubiquitous properties that, disappointingly,
4 | we could not locate in the standard library. ---The standard library needs some
5 | sort of ``table of contents \emph{with} subsection'' to make it easier to know
6 | of what is available.
7 |
8 | This module re-exports (some of) the contents of the standard library's |Data.Product| and |Data.Sum|.
9 |
10 | %{{{ Imports
11 | \begin{code}
12 | module Helpers.DataProperties where
13 |
14 | open import Level renaming (suc to lsuc; zero to lzero)
15 | open import Function using (id ; _∘_ ; const)
16 | open import Helpers.EqualityCombinators
17 |
18 | open import Data.Product public using (_×_; proj₁; proj₂; Σ; _,_; swap ; uncurry) renaming (map to _×₁_ ; <_,_> to ⟨_,_⟩)
19 | open import Data.Sum public using (inj₁; inj₂; [_,_]) renaming (map to _⊎₁_)
20 | open import Data.Nat using (ℕ; zero; suc)
21 |
22 | foldn : {ℓ : Level} {P : ℕ → Set ℓ}
23 | → P zero
24 | → (∀ {x} → P x → P (suc x))
25 | →
26 | (n : ℕ) → P n
27 | foldn {P = P} pz ps zero = pz
28 | foldn {P = P} pz ps (suc x) = ps (foldn {P = P} pz ps x)
29 | \end{code}
30 |
31 | \subsection*{Precedence Levels}
32 |
33 | The standard library assigns precedence level of 1 for the infix operator |_⊎_|,
34 | which is rather odd since infix operators ought to have higher precedence that equality
35 | combinators, yet the standard library assigns |_≈⟨_⟩_| a precedence level of 2.
36 | The usage of these two ---e.g. in |CommMonoid.lagda|--- causes an annoying number of
37 | parentheses and so we reassign the level of the infix operator to avoid such a situation.
38 |
39 | \begin{code}
40 | infixr 3 _⊎_
41 | _⊎_ = Data.Sum._⊎_
42 | \end{code}
43 |
44 | %}}}
45 |
46 | %{{{ Generalised Bot and Top
47 | \subsection{Generalised Bot and Top}
48 | To avoid a flurry of |lift|'s, and for the sake of clarity, we define level-polymorphic
49 | empty and unit types.
50 | \begin{code}
51 | open import Level
52 |
53 | data ⊥ {ℓ : Level} : Set ℓ where
54 |
55 | ⊥-elim : {a ℓ : Level} {A : Set a} → ⊥ {ℓ} → A
56 | ⊥-elim ()
57 |
58 | record ⊤ {ℓ : Level} : Set ℓ where
59 | constructor tt
60 | \end{code}
61 | %}}}
62 |
63 | %{{{ sums
64 | \subsection{Sums}
65 |
66 | %{{{ the ⊎ operation on functions is a functorial congruence
67 |
68 | Just as |_⊎_| takes types to types, its ``map'' variant |_⊎₁_| takes
69 | functions to functions and is a functorial congruence:
70 | It preserves identity, distributes over composition, and preserves
71 | extenstional equality.
72 |
73 | \begin{code}
74 | ⊎-id : {a b : Level} {A : Set a} {B : Set b} → id ⊎₁ id ≐ id {A = A ⊎ B}
75 | ⊎-id = [ ≐-refl , ≐-refl ]
76 |
77 | ⊎-∘ : {a b c a' b' c' : Level}
78 | {A : Set a} {A' : Set a'} {B : Set b} {B' : Set b'} {C' : Set c} {C : Set c'}
79 | {f : A → A'} {g : B → B'} {f' : A' → C} {g' : B' → C'}
80 | → (f' ∘ f) ⊎₁ (g' ∘ g) ≐ (f' ⊎₁ g') ∘ (f ⊎₁ g) --- aka “the exchange rule for sums”
81 | ⊎-∘ = [ ≐-refl , ≐-refl ]
82 |
83 | ⊎-cong : {a b c d : Level} {A : Set a} {B : Set b} {C : Set c} {D : Set d} {f f' : A → C} {g g' : B → D}
84 | → f ≐ f' → g ≐ g' → f ⊎₁ g ≐ f' ⊎₁ g'
85 | ⊎-cong f≈f' g≈g' = [ ∘-≐-cong₂ inj₁ f≈f' , ∘-≐-cong₂ inj₂ g≈g' ]
86 | \end{code}
87 | %}}}
88 |
89 | %{{{ ∘-[,] : fusion property for casing
90 |
91 | Composition post-distributes into casing,
92 |
93 | \begin{code}
94 | ∘-[,] : {a b c d : Level} {A : Set a} {B : Set b} {C : Set c} {D : Set d} {f : A → C} {g : B → C} {h : C → D}
95 | → h ∘ [ f , g ] ≐ [ h ∘ f , h ∘ g ] -- aka “fusion”
96 | ∘-[,] = [ ≐-refl , ≐-refl ]
97 | \end{code}
98 |
99 | %}}}
100 |
101 | %{{{ from⊎ : the dual to |diag|
102 |
103 | It is common that a data-type constructor |D : Set → Set| allows us to extract
104 | elements of the underlying type and so we have a natural transfomation |D ⟶ 𝑰|,
105 | where |𝑰| is the identity functor.
106 | These kind of results will occur for our other simple data-strcutres as well.
107 | In particular, this is the case for |D A = 2× A = A ⊎ A|:
108 |
109 | \begin{code}
110 | from⊎ : {ℓ : Level} {A : Set ℓ} → A ⊎ A → A
111 | from⊎ = [ id , id ]
112 |
113 | -- |from⊎| is a natural transformation
114 | --
115 | from⊎-nat : {a b : Level} {A : Set a} {B : Set b}{f : A → B} → f ∘ from⊎ ≐ from⊎ ∘ (f ⊎₁ f)
116 | from⊎-nat = [ ≐-refl , ≐-refl ]
117 |
118 | -- |from⊎| is injective and so is pre-invertible,
119 | --
120 | from⊎-preInverse : {a b : Level} {A : Set a} {B : Set b} → id ≐ from⊎ {A = A ⊎ B} ∘ (inj₁ ⊎₁ inj₂)
121 | from⊎-preInverse = [ ≐-refl , ≐-refl ]
122 | \end{code}
123 |
124 | \edinsert{MA}{A brief mention about co-monads?}
125 |
126 | %}}}
127 |
128 | %}}}
129 |
130 | %{{{ products: diag
131 | \subsection{Products}
132 | Dual to |from⊎|, a natural transformation |2×_ ⟶ 𝑰|, is |diag|, the transformation |𝑰 ⟶ _²|.
133 |
134 | \begin{code}
135 | diag : {ℓ : Level} {A : Set ℓ} (a : A) → A × A
136 | diag a = a , a
137 | \end{code}
138 |
139 | \edinsert{MA}{ A brief mention of Haskell's |const|, which is |diag| curried. Also something about |K| combinator? }
140 |
141 | Z-style notation for sums,
142 | \begin{code}
143 | Σ∶• : {a b : Level} (A : Set a) (B : A → Set b) → Set (a ⊔ b)
144 | Σ∶• = Data.Product.Σ
145 | infix -666 Σ∶•
146 | syntax Σ∶• A (λ x → B) = Σ x ∶ A • B
147 | \end{code}
148 |
149 | %}}}
150 |
151 | %{{{ suc is injective
152 | \begin{code}
153 | open import Data.Nat.Properties
154 | suc-inj : ∀ {i j} → ℕ.suc i ≡ ℕ.suc j → i ≡ j
155 | suc-inj = +-cancelˡ-≡ (ℕ.suc ℕ.zero)
156 | \end{code}
157 | or
158 | \begin{spec}
159 | suc-inj {0} _≡_.refl = _≡_.refl
160 | suc-inj {ℕ.suc i} _≡_.refl = _≡_.refl
161 | \end{spec}
162 |
163 | %}}}
164 |
165 | %{{{ vectors: _‼_
166 |
167 | \begin{code}
168 | open import Data.Vec public using () renaming (lookup to _‼_)
169 | \end{code}
170 | %}}}
171 |
172 | % Quick Folding Instructions:
173 | % C-c C-s :: show/unfold region
174 | % C-c C-h :: hide/fold region
175 | % C-c C-w :: whole file fold
176 | % C-c C-o :: whole file unfold
177 | %
178 | % Local Variables:
179 | % folded-file: t
180 | % eval: (fold-set-marks "%{{{ " "%}}}")
181 | % eval: (fold-whole-buffer)
182 | % fold-internal-margins: 0
183 | % end:
184 |
--------------------------------------------------------------------------------
/Structures/Experiments/EquivEquiv.agda:
--------------------------------------------------------------------------------
1 | {-# OPTIONS --without-K #-}
2 |
3 | module EquivEquiv where
4 |
5 | open import Level using (Level; _⊔_)
6 | open import Data.Product using (_,_; proj₁; proj₂)
7 | open import Data.Sum using (inj₁; inj₂)
8 |
9 | open import Relation.Binary using (Setoid)
10 | import Relation.Binary.PropositionalEquality as P
11 | using (_≡_; refl; sym; trans; cong; module ≡-Reasoning)
12 |
13 | import Relation.Binary.EqReasoning as EqR
14 |
15 | open import Function using (id; _∘_)
16 |
17 | open import Equiv
18 | using (module isqinv; qinv; _≃_; id≃; sym≃; _●_;
19 | _≐_; ≐-refl; ≐-sym; ≐-trans; cong∘r; cong∘l;
20 | _⊎≃_; β₁; β₂; gg; β⊎₁; β⊎₂)
21 |
22 | ------------------------------------------------------------------------------
23 | -- Extensional equivalence of equivalences
24 |
25 | -- We need g to "pin down" the inverse, else we get lots of unsolved
26 | -- metas.
27 |
28 | infix 4 _≋_
29 |
30 | record _≋_ {ℓ ℓ' : Level} {A : Set ℓ} {B : Set ℓ'} (eq₁ eq₂ : A ≃ B) :
31 | Set (ℓ ⊔ ℓ') where
32 | constructor eq
33 | open isqinv
34 | field
35 | f≡ : proj₁ eq₁ ≐ proj₁ eq₂
36 | g≡ : g (proj₂ eq₁) ≐ g (proj₂ eq₂)
37 |
38 | -- the proof could use ∼-Reasoning if we had defined it
39 | g≡′ : g (proj₂ eq₁) ≐ g (proj₂ eq₂)
40 | g≡′ =
41 | ≐-trans (cong∘r g₁ (≐-refl {f = id})) ( -- id ∘ g₁
42 | ≐-trans (cong∘r g₁ (≐-sym (β (proj₂ eq₂))))
43 | (≐-trans (cong∘l g₂ (cong∘r g₁ (≐-sym f≡))) (
44 | (cong∘l g₂ (α (proj₂ eq₁))))))
45 | where
46 | g₁ = g (proj₂ eq₁)
47 | g₂ = g (proj₂ eq₂)
48 | f₁ = proj₁ eq₁
49 | f₂ = proj₁ eq₂
50 |
51 |
52 | -- The equivalence of equivalences is an equivalence relation that
53 | -- respects composition
54 |
55 | id≋ : ∀ {ℓ ℓ'} {A : Set ℓ} {B : Set ℓ'} {x : A ≃ B} → x ≋ x
56 | id≋ = record { f≡ = λ _ → P.refl ; g≡ = λ _ → P.refl }
57 |
58 | sym≋ : ∀ {ℓ ℓ'} {A : Set ℓ} {B : Set ℓ'} {x y : A ≃ B} → x ≋ y → y ≋ x
59 | sym≋ (eq f≡ g≡) = eq (λ a → P.sym (f≡ a)) (λ b → P.sym (g≡ b))
60 |
61 | flip≋ : ∀ {ℓ ℓ'} {A : Set ℓ} {B : Set ℓ'} {x y : A ≃ B} →
62 | x ≋ y → (sym≃ x) ≋ (sym≃ y)
63 | flip≋ (eq f≡ g≡) = eq g≡ f≡
64 |
65 | flip-sym≋ : ∀ {ℓ ℓ'} {A : Set ℓ} {B : Set ℓ'} {x y : A ≃ B}→
66 | x ≋ y → sym≃ y ≋ sym≃ x
67 | flip-sym≋ (eq f≡ g≡) = eq (≐-sym g≡) (≐-sym f≡)
68 |
69 | trans≋ : ∀ {ℓ ℓ'} {A : Set ℓ} {B : Set ℓ'} {x y z : A ≃ B} →
70 | x ≋ y → y ≋ z → x ≋ z
71 | trans≋ (eq f≡ g≡) (eq h≡ i≡) =
72 | eq (λ a → P.trans (f≡ a) (h≡ a)) (λ b → P.trans (g≡ b) (i≡ b))
73 |
74 | _◎_ : {A B C : Set} {f h : B ≃ C} {g i : A ≃ B} → f ≋ h → g ≋ i →
75 | (f ● g) ≋ (h ● i)
76 | _◎_ {f = f} {h} {g} {i}
77 | (eq f≡ g≡) (eq h≡ i≡) = eq fwd bwd
78 | -- eq (λ x → {!!} ) -- P.trans (P.cong f (h≡ x)) (f≡ (i x)))
79 | -- (λ x → {!!} ) -- P.trans (P.cong g⁻¹ (g≡ x)) (i≡ (h⁻¹ x)))
80 | where
81 | open P.≡-Reasoning
82 | fwd : proj₁ (f ● g) ≐ proj₁ (h ● i)
83 | fwd x = begin (
84 | proj₁ (f ● g) x
85 | ≡⟨ β₁ x ⟩
86 | proj₁ f (proj₁ g x)
87 | ≡⟨ f≡ (proj₁ g x) ⟩
88 | proj₁ h (proj₁ g x)
89 | ≡⟨ P.cong (proj₁ h) (h≡ x) ⟩
90 | proj₁ h (proj₁ i x)
91 | ≡⟨ P.sym (β₁ x) ⟩
92 | proj₁ (h ● i) x ∎)
93 | bwd : gg (f ● g) ≐ gg (h ● i)
94 | bwd x =
95 | begin (
96 | gg (f ● g) x
97 | ≡⟨ β₂ x ⟩
98 | gg g (gg f x)
99 | ≡⟨ i≡ (gg f x) ⟩
100 | gg i (gg f x)
101 | ≡⟨ P.cong (gg i) (g≡ x) ⟩
102 | gg i (gg h x)
103 | ≡⟨ P.sym (β₂ x) ⟩
104 | gg (h ● i) x ∎)
105 |
106 | rinv≋ : ∀ {ℓ} {A B : Set ℓ} (x : A ≃ B) →
107 | (x ● (sym≃ x)) ≋ id≃ {A = B}
108 | rinv≋ (_ , qinv _ α _) = eq (≐-trans β₁ α) (≐-trans β₂ α)
109 |
110 | linv≋ : ∀ {ℓ} {A B : Set ℓ} (e : A ≃ B) → ((sym≃ e) ● e) ≋ id≃
111 | linv≋ (_ , qinv _ _ β) = eq (≐-trans β₁ β) (≐-trans β₂ β)
112 |
113 | lid≋ : ∀ {ℓ} {A B : Set ℓ} {f : A ≃ B} → (id≃ ● f) ≋ f
114 | lid≋ = eq β₁ β₂
115 |
116 | rid≋ : ∀ {ℓ} {A B : Set ℓ} {f : A ≃ B} → (f ● id≃) ≋ f
117 | rid≋ = eq β₁ β₂
118 |
119 | --
120 | {-
121 | symsym : ∀ {A B : Set} {f : A ≃ B} → sym≃ (sym≃ f) ≋ f
122 | symsym = eq (λ _ → P.refl) (λ _ → P.refl)
123 | -}
124 |
125 | sym≃● : ∀ {A B C : Set} {g : B ≃ C} {f : A ≃ B} →
126 | sym≃ (g ● f) ≋ sym≃ f ● sym≃ g
127 | sym≃● {g = (g , qinv g⁻¹ _ _)} {(f , qinv f⁻¹ _ _)} =
128 | eq (≐-trans β₂ (≐-sym β₁)) (≐-trans β₁ (≐-sym β₂))
129 |
130 | -- underlying it all, it uses ∘ and ≡, thus associativity is immediate
131 |
132 | ●-assoc : {A B C D : Set} {f : A ≃ B} {g : B ≃ C} {h : C ≃ D} →
133 | ((h ● g) ● f) ≋ (h ● (g ● f))
134 | ●-assoc {f = f} {g} {h} = eq fwd bwd
135 | where
136 | open P.≡-Reasoning
137 | fwd : proj₁ ((h ● g) ● f) ≐ proj₁ (h ● (g ● f))
138 | fwd x = begin (
139 | proj₁ ((h ● g) ● f) x
140 | ≡⟨ β₁ x ⟩
141 | proj₁ (h ● g) (proj₁ f x)
142 | ≡⟨ β₁ (proj₁ f x) ⟩
143 | proj₁ h (proj₁ g (proj₁ f x))
144 | ≡⟨ P.cong (proj₁ h) (P.sym (β₁ x)) ⟩
145 | proj₁ h (proj₁ (g ● f) x)
146 | ≡⟨ P.sym (β₁ x) ⟩
147 | proj₁ (h ● (g ● f)) x ∎)
148 | bwd : gg ((h ● g) ● f) ≐ gg (h ● (g ● f))
149 | bwd x = begin (
150 | gg ((h ● g) ● f) x
151 | ≡⟨ β₂ x ⟩
152 | gg f (gg (h ● g) x)
153 | ≡⟨ P.cong (gg f) (β₂ x) ⟩
154 | gg f (gg g (gg h x))
155 | ≡⟨ P.sym (β₂ (gg h x)) ⟩
156 | gg (g ● f) (gg h x)
157 | ≡⟨ P.sym (β₂ x) ⟩
158 | gg (h ● (g ● f)) x ∎)
159 |
160 | ●-assocl : {A B C D : Set} {f : A ≃ B} {g : B ≃ C} {h : C ≃ D} →
161 | h ● (g ● f) ≋ (h ● g) ● f
162 | ●-assocl {f = f} {g} {h} = sym≋ (●-assoc {f = f} {g} {h})
163 |
164 | -- The setoid of equivalences under ≋
165 |
166 | _S≃_ : ∀ {ℓ ℓ'} (A : Set ℓ) (B : Set ℓ') → Setoid (ℓ ⊔ ℓ') (ℓ ⊔ ℓ')
167 | _S≃_ A B = record
168 | { Carrier = A ≃ B
169 | ; _≈_ = _≋_
170 | ; isEquivalence = record
171 | { refl = id≋
172 | ; sym = sym≋
173 | ; trans = trans≋
174 | }
175 | }
176 |
177 | module ≋-Reasoning where
178 | module _ {a b} {A : Set a} {B : Set b} where
179 | open EqR (A S≃ B) public
180 | hiding (_≡⟨_⟩_; _≡⟨⟩_) renaming (_≈⟨_⟩_ to _≋⟨_⟩_)
181 |
182 | ------------------------------------------------------------------------------
183 |
--------------------------------------------------------------------------------
/Structures/Experiments/ParComp.lagda:
--------------------------------------------------------------------------------
1 | \section{Parallel Composition}
2 |
3 | We need a way to put two |SetoidFamily| ``side by side'' -- a form of
4 | parellel composition. To achieve this requires a certain amount of
5 | infrastructure: parallel composition of relations, and both disjoint
6 | sum and cartesian product of |Setoid|s. So the next couple of sections
7 | proceed with that infrastruction, before diving in to the crux of the
8 | matter.
9 |
10 | %{{{ Imports
11 | \begin{code}
12 | module ParComp where
13 |
14 | open import Level
15 | open import Relation.Binary using (Setoid; REL; Rel)
16 | open import Function using (flip) renaming (id to id₀; _∘_ to _⊚_)
17 | open import Function.Equality using (Π; _⟨$⟩_; cong; id; _⟶_; _∘_)
18 | open import Function.Inverse using () renaming (_InverseOf_ to Inv)
19 | open import Relation.Binary.Product.Pointwise using (_×-setoid_)
20 |
21 | open import Categories.Category using (Category)
22 | open import Categories.Object.Coproduct
23 |
24 | open import DataProperties
25 | open import SetoidEquiv
26 |
27 | open import TypeEquiv using (swap₊; swap⋆)
28 | \end{code}
29 | %}}}
30 |
31 | %{{{ \subsection{Parallel Composition of relations} _∥_ ; [_∥_] ; ∥-sym ; ∥-trans
32 | \subsection{Parallel Composition of relations}
33 |
34 | Parallel composition of heterogeneous relations.
35 |
36 | Note that this is a specialized version of the standard library's
37 | |_⊎-Rel_| (see |Relation.Binary.Sum|); this one gets rid of the
38 | bothersome |₁∼₂| term.
39 |
40 | \begin{code}
41 | data _∥_ {a₁ b₁ c₁ a₂ b₂ c₂ : Level}
42 | {A₁ : Set a₁} {B₁ : Set b₁} {A₂ : Set a₂} {B₂ : Set b₂}
43 | (R₁ : REL A₁ B₁ c₁) (R₂ : REL A₂ B₂ c₂)
44 | : REL (A₁ ⊎ A₂) (B₁ ⊎ B₂) (c₁ ⊔ c₂) where
45 | left : {x : A₁} {y : B₁} (r₁ : R₁ x y) → (R₁ ∥ R₂) (inj₁ x) (inj₁ y)
46 | right : {x : A₂} {y : B₂} (r₂ : R₂ x y) → (R₁ ∥ R₂) (inj₂ x) (inj₂ y)
47 |
48 | elim : {a₁ b₁ a₂ b₂ c₁ c₂ d : Level}
49 | {A₁ : Set a₁} {B₁ : Set b₁} {A₂ : Set a₂} {B₂ : Set b₂}
50 | {R₁ : REL A₁ B₁ c₁} {R₂ : REL A₂ B₂ c₂}
51 | (P : {a : A₁ ⊎ A₂} {b : B₁ ⊎ B₂} (pf : (R₁ ∥ R₂) a b) → Set d)
52 | (F : {a : A₁} {b : B₁} → (f : R₁ a b) → P (left f))
53 | (G : {a : A₂} {b : B₂} → (g : R₂ a b) → P (right g))
54 | {a : A₁ ⊎ A₂} {b : B₁ ⊎ B₂} → (x : (R₁ ∥ R₂) a b) → P x
55 | elim P F G (left r₁) = F r₁
56 | elim P F G (right r₂) = G r₂
57 |
58 | -- If the argument relations are symmetric then so is their parallel composition.
59 | ∥-sym : {a a′ c c′ : Level} {A : Set a} {R₁ : Rel A c}
60 | {A′ : Set a′} {R₂ : Rel A′ c′}
61 | (sym₁ : {x y : A} → R₁ x y → R₁ y x) (sym₂ : {x y : A′} → R₂ x y → R₂ y x)
62 | {x y : A ⊎ A′}
63 | → (R₁ ∥ R₂) x y → (R₁ ∥ R₂) y x
64 | ∥-sym {R₁ = R₁} {R₂ = R₂} sym₁ sym₂ pf =
65 | elim (λ {a b} (x : (R₁ ∥ R₂) a b) → (R₁ ∥ R₂) b a) (left ⊚ sym₁) (right ⊚ sym₂) pf
66 |
67 | ∥-trans : {a a′ ℓ ℓ′ : Level} (A : Setoid a ℓ) (A′ : Setoid a′ ℓ′)
68 | {x y z : Setoid.Carrier A ⊎ Setoid.Carrier A′} →
69 | let R₁ = Setoid._≈_ A in let R₂ = Setoid._≈_ A′ in
70 | (R₁ ∥ R₂) x y → (R₁ ∥ R₂) y z → (R₁ ∥ R₂) x z
71 | ∥-trans A A′ {inj₁ x} (left r₁) (left r₂) = left (Setoid.trans A r₁ r₂)
72 | ∥-trans A A′ {inj₂ y₂} (right r₂) (right r₃) = right (Setoid.trans A′ r₂ r₃)
73 | \end{code}
74 | %}}}
75 |
76 | %{{{ \subsection{Union and product of |Setoid|} |_⊎S_| |_×S_|
77 | \subsection{Union and product of |Setoid|}
78 | \begin{code}
79 | module _ {ℓA₁ ℓa₁ ℓA₂ ℓa₂ : Level} (S₁ : Setoid ℓA₁ ℓa₁) (S₂ : Setoid ℓA₂ ℓa₂) where
80 | infix 3 _⊎S_ _×S_
81 |
82 | open Setoid S₁ renaming (Carrier to s₁; _≈_ to _≈₁_; refl to refl₁; sym to sym₁)
83 | open Setoid S₂ renaming (Carrier to s₂; _≈_ to _≈₂_; refl to refl₂; sym to sym₂)
84 |
85 | _⊎S_ : Setoid (ℓA₁ ⊔ ℓA₂) (ℓa₁ ⊔ ℓa₂)
86 | _⊎S_ = record
87 | { Carrier = s₁ ⊎ s₂
88 | ; _≈_ = _≈₁_ ∥ _≈₂_
89 | ; isEquivalence = record
90 | { refl = λ { {inj₁ x} → left refl₁ ; {inj₂ y} → right refl₂ }
91 | ; sym = ∥-sym sym₁ sym₂
92 | ; trans = ∥-trans S₁ S₂
93 | }
94 | }
95 |
96 | _×S_ : Setoid (ℓA₁ ⊔ ℓA₂) (ℓa₁ ⊔ ℓa₂)
97 | _×S_ = S₁ ×-setoid S₂
98 | \end{code}
99 | %}}}
100 |
101 | %{{{ \subsection{Union of |Setoid| is commutative} |⊎S-comm|
102 | \subsection{Union of |Setoid| is commutative}
103 | \begin{code}
104 | module _ {ℓA₁ ℓa₁ ℓA₂ ℓa₂ : Level} (S₁ : Setoid ℓA₁ ℓa₁) (S₂ : Setoid ℓA₂ ℓa₂) where
105 | ⊎S-comm : (S₁ ⊎S S₂) ≅ (S₂ ⊎S S₁)
106 | ⊎S-comm = record
107 | { to = record { _⟨$⟩_ = swap₊ ; cong = λ { (left r₁) → right r₁ ; (right r₂) → left r₂} }
108 | ; from = record { _⟨$⟩_ = swap₊ ; cong = λ { (left r₁) → right r₁ ; (right r₂) → left r₂} }
109 | ; inverse-of = record
110 | { left-inverse-of = λ { (inj₁ x) → left (refl S₁) ; (inj₂ y) → right (refl S₂)}
111 | ; right-inverse-of = λ { (inj₁ x) → left (refl S₂) ; (inj₂ y) → right (refl S₁)} }
112 | }
113 | where open Setoid
114 | \end{code}
115 | %}}}
116 | %{{{ \subsection{|_⊎S_| is a congruence} |_⊎S₁_|
117 | \subsection{|_⊎S_| is a congruence}
118 | \begin{code}
119 | module _ {ℓA₁ ℓa₁ ℓA₂ ℓa₂ ℓA₃ ℓa₃ ℓA₄ ℓa₄ : Level}
120 | {S₁ : Setoid ℓA₁ ℓa₁} {S₂ : Setoid ℓA₂ ℓa₂} {S₃ : Setoid ℓA₃ ℓa₃} {S₄ : Setoid ℓA₄ ℓa₄} where
121 | _⊎S₁_ : S₁ ≅ S₃ → S₂ ≅ S₄ → (S₁ ⊎S S₂) ≅ (S₃ ⊎S S₄)
122 | 1≅3 ⊎S₁ 2≅4 = record
123 | { to = record
124 | { _⟨$⟩_ = λ { (inj₁ x) → inj₁ (to 1≅3 ⟨$⟩ x) ; (inj₂ y) → inj₂ (to 2≅4 ⟨$⟩ y)}
125 | ; cong = λ { (left r₁) → left (cong (to 1≅3) r₁) ; (right r₂) → right (cong (to 2≅4) r₂)} }
126 | ; from = record
127 | { _⟨$⟩_ = λ { (inj₁ x) → inj₁ (from 1≅3 ⟨$⟩ x) ; (inj₂ y) → inj₂ (from 2≅4 ⟨$⟩ y)}
128 | ; cong = λ { (left r₁) → left (cong (from 1≅3) r₁) ; (right r₂) → right (cong (from 2≅4) r₂)} }
129 | ; inverse-of = record
130 | { left-inverse-of = λ { (inj₁ x) → left (left-inverse-of 1≅3 x)
131 | ; (inj₂ y) → right (left-inverse-of 2≅4 y)}
132 | ; right-inverse-of = λ { (inj₁ x) → left (right-inverse-of 1≅3 x)
133 | ; (inj₂ y) → right (right-inverse-of 2≅4 y)} }
134 | }
135 | where open _≅_
136 | \end{code}
137 | %}}}
138 |
139 | % Quick Folding Instructions:
140 | % C-c C-s :: show/unfold region
141 | % C-c C-h :: hide/fold region
142 | % C-c C-w :: whole file fold
143 | % C-c C-o :: whole file unfold
144 | %
145 | % Local Variables:
146 | % folded-file: t
147 | % eval: (fold-set-marks "%{{{ " "%}}}")
148 | % eval: (fold-whole-buffer)
149 | % fold-internal-margins: 0
150 | % end:
151 |
--------------------------------------------------------------------------------
/Structures/IndexedUnaryAlgebra.lagda:
--------------------------------------------------------------------------------
1 | \section{Indexed Unary Algebras}
2 |
3 | %{{{ Necessary Imports
4 | \begin{code}
5 | module Structures.IndexedUnaryAlgebra where
6 |
7 | open import Level renaming (suc to lsuc; zero to lzero ; _⊔_ to _⊍_)
8 | open import Function hiding (_$_)
9 | open import Data.List
10 |
11 | open import Helpers.Categorical
12 |
13 | open import Helpers.Forget
14 | open import Helpers.Function2
15 | open import Helpers.EqualityCombinators
16 | open import Helpers.DataProperties
17 | \end{code}
18 | %}}}
19 |
20 | %{{{ _UnaryAlg
21 | An \emph{|I| indexed unary algebra} consists of a carrier |Q| and for each
22 | index |i : I| a unary morphism |Opᵢ : Q → Q| on the carrier. In general, an
23 | operation of type |I × Q → Q| is also known as an \emph{|I| action on |Q|},
24 | and pop-up in the study of groups and vector spaces.
25 |
26 | \begin{code}
27 | record UnaryAlg {a} (I : Set a) (ℓ : Level) : Set (lsuc ℓ ⊍ a) where
28 | constructor MkAlg
29 | field
30 | Carrier : Set ℓ
31 | Op : {i : I} → Carrier → Carrier
32 |
33 | -- action form
34 | _·_ : I → Carrier → Carrier
35 | i · c = Op {i} c
36 |
37 | open UnaryAlg
38 | \end{code}
39 | %}}}
40 |
41 | Henceforth we work with a given indexing set,
42 | \begin{code}
43 | module _ {a} (I : Set a) where -- Musa: Most likely ought to name this module.
44 | \end{code}
45 |
46 | %{{{ Hom
47 | Give two unary algebras, over the same indexing set, a morphism between them
48 | is a function of their underlying carriers that respects the actions.
49 |
50 | \begin{code}
51 | record Hom {ℓ} (X Y : UnaryAlg I ℓ) : Set (lsuc ℓ ⊍ a) where
52 | constructor MkHom
53 | infixr 5 mor
54 | field
55 | mor : Carrier X → Carrier Y
56 | preservation : {i : I} → mor ∘ Op X {i} ≐ Op Y {i} ∘ mor
57 |
58 | open Hom using (mor)
59 | open Hom using () renaming (mor to _$_) -- override application to take a |Hom|
60 |
61 | -- arguments can usually be inferred, so implicit variant
62 | preservation : {ℓ : Level} {X Y : UnaryAlg I ℓ} (F : Hom X Y)
63 | → {i : I} {x : Carrier X} → F $ Op X {i} x ≡ Op Y {i} (F $ x)
64 | preservation F = Hom.preservation F _
65 | \end{code}
66 |
67 | Notice that the |preservation| proof looks like a usual homomorphism condition ---after excusing the implicits.
68 | Rendered in action notation, it would take the shape |∀{i x} → mor (i · x) ≡ i · mor x| with the |mor|
69 | ``leap-frogging'' over the action. Admiteddly this form is also common and then |mor| is called an ``equivaraint''
70 | function, yet this sounds like a new unfamilar concept than it really it: Homomorphism.
71 | %}}}
72 |
73 | %{{{ UnaryAlgCat ; Forget
74 |
75 | Unsuprisngly, the indexed unary algebra's form a category.
76 |
77 | \begin{code}
78 | UnaryAlgCat : (ℓ : Level) → Category (lsuc ℓ ⊍ a) (lsuc ℓ ⊍ a) ℓ
79 | UnaryAlgCat ℓ = record
80 | { Obj = UnaryAlg I ℓ
81 | ; _⇒_ = Hom
82 | ; _≡_ = λ F G → mor F ≐ mor G
83 | ; id = λ {A} → MkHom id ≐-refl
84 | ; _∘_ = λ {A} {B} {C} F G → MkHom (mor F ∘ mor G) (λ {i} x → let open ≡.≡-Reasoning {A = Carrier C} in begin
85 | (mor F ∘ mor G ∘ Op A) x
86 | ≡⟨ ≡.cong (mor F) (preservation G) ⟩
87 | (mor F ∘ Op B ∘ mor G) x
88 | ≡⟨ preservation F ⟩
89 | (Op C ∘ mor F ∘ mor G) x
90 | ∎)
91 | ; assoc = ≐-refl
92 | ; identityˡ = ≐-refl
93 | ; identityʳ = ≐-refl
94 | ; equiv = record { IsEquivalence ≐-isEquivalence}
95 | ; ∘-resp-≡ = λ {A} {B} {C} {F} {G} {H} {K} F≈G H≈K x → let open ≡.≡-Reasoning {A = Carrier C} in begin
96 | (mor F ∘ mor H) x
97 | ≡⟨ F≈G _ ⟩
98 | (mor G ∘ mor H) x
99 | ≡⟨ ≡.cong (mor G) (H≈K _) ⟩
100 | (mor G ∘ mor K) x
101 | ∎
102 | }
103 | where
104 | open import Relation.Binary using (IsEquivalence)
105 | \end{code}
106 |
107 | Needless to say, we can ignore the extra structure to arrive at the underlying carrier.
108 |
109 | \begin{code}
110 | Forget : (ℓ : Level) → Functor (UnaryAlgCat ℓ) (Sets ℓ)
111 | Forget ℓ = record
112 | { F₀ = Carrier
113 | ; F₁ = mor
114 | ; identity = ≡.refl
115 | ; homomorphism = ≡.refl
116 | ; F-resp-≡ = λ F≈G {x} → F≈G x
117 | }
118 | \end{code}
119 | %}}}
120 |
121 | %{{{ I* ; the free I-unaryAlgebra
122 |
123 | For each |I| indexed unary algebra |(A, Op)| with an elected element |a₀ : A|,
124 | there is a unique homomorpism |fold : (List I, _∷_) ⟶ (A, Op)| sending |[] ↦ a₀|.
125 |
126 | \begin{code}
127 | module _ (Q : UnaryAlg I a) (q₀ : Carrier Q) where
128 |
129 | open import Data.Unit
130 | I* : UnaryAlg I a
131 | I* = MkAlg (List I) (λ {x} xs → x ∷ xs)
132 |
133 | fold₀ : List I → Carrier Q
134 | fold₀ [] = q₀
135 | fold₀ (x ∷ xs) = Op Q {x} (fold₀ xs)
136 |
137 | fold : Hom I* Q
138 | fold = MkHom fold₀ ≐-refl
139 |
140 | fold-point : fold $ [] ≡ q₀
141 | fold-point = ≡.refl
142 |
143 | fold-unique : (F : Hom I* Q)(point : F $ [] ≡ q₀) → mor F ≐ mor fold
144 | fold-unique F point [] = point
145 | fold-unique F point (x ∷ xs) = let open ≡.≡-Reasoning {A = Carrier Q} in begin
146 | mor F (x ∷ xs)
147 | ≡⟨ preservation F ⟩
148 | Op Q {x} (mor F xs)
149 | ≡⟨ induction-hypothesis ⟩
150 | Op Q {x} (fold₀ xs)
151 | ∎
152 | where induction-hypothesis = ≡.cong (Op Q) (fold-unique F point xs)
153 | \end{code}
154 |
155 | Perhaps it would be better to consider POINTED indexed unary algebras, where this
156 | result may be phrased more concisely.
157 |
158 | %}}}
159 |
160 | WK: A signature with no constant symbols results in there being no closed terms
161 | and so the term algebra is just the empty set of no closed terms quotiented by
162 | the given equations and the resulting algebra has an empty carrier.
163 |
164 | Free : build over generators -- cf Multiset construction in CommMonoid.lagda
165 | Initial : does not require generators
166 |
167 | ToDo ∷ mimic the multiset construction here for generators S “over” IndexedUnaryAlgebras.
168 | WK claims it may have carrier S × List I; then the non-indexed case is simply List ⊤ ≅ ℕ.
169 |
170 | % Quick Folding Instructions:
171 | % C-c C-s :: show/unfold region
172 | % C-c C-h :: hide/fold region
173 | % C-c C-w :: whole file fold
174 | % C-c C-o :: whole file unfold
175 | %
176 | % Local Variables:
177 | % folded-file: t
178 | % eval: (fold-set-marks "%{{{ " "%}}}")
179 | % eval: (fold-whole-buffer)
180 | % fold-internal-margins: 0
181 | % end:
182 |
183 |
--------------------------------------------------------------------------------
/Structures/Contractable.lagda:
--------------------------------------------------------------------------------
1 | \DeclareUnicodeCharacter{7472}{\ensuremath{7472}}
2 |
3 | \section{Contractable Algebras: Working with Pointed Types}
4 |
5 | We consider the theory of \emph{pointed algebras} which consist of a type
6 | along with an elected value of that type along with the only law that makes use
7 | of the point and the carrier: $\forall x \;\bullet\; x = \texttt{thePoint}$.
8 |
9 | At a first glance, this algebra seems useless. However, the aim of this work
10 | is to explore simple elementary algebras and see where they arise in software
11 | engineering. Lest the reader abandon us now, we cut to the chase: This algebra
12 | arises when proxy datatypes are required; i.e., when a \emph{type parameter} is
13 | utilised to convey information, but the actual value is completely irrelevant.
14 |
15 | JC: Examples from Haskell?
16 |
17 | %{{{ Imports
18 | \begin{code}
19 | module Structures.Contractable where
20 |
21 | open import Level renaming (suc to lsuc; zero to lzero)
22 | open import Function using (id ; _∘_)
23 | open import Data.Maybe using (Maybe; just; nothing; maybe; maybe′)
24 |
25 | open import Helpers.Categorical
26 | open import Helpers.Forget
27 | open import Helpers.EqualityCombinators
28 |
29 | open ≡
30 | \end{code}
31 | %}}}
32 |
33 | %{{{ Contractable ; Hom
34 |
35 | \subsection{Definition}
36 |
37 | \begin{code}
38 | record Contractable {a} : Set (lsuc a) where
39 | constructor MkContractable
40 | field
41 | Carrier : Set a
42 | Id : Carrier {- The “Id”entified point -}
43 | collapse : ∀ {x} → x ≡ Id
44 |
45 | open Contractable
46 | \end{code}
47 |
48 | One would expect a ``structure preserving operation'' on such structures to be a function
49 | between the underlying carriers that takes the source's point to the target's point.
50 | This is the case in the brute theory of pointed algebras, whereas contractable algebras
51 | are so refined that the point-preservation property is provided for free.
52 |
53 | \begin{code}
54 | Hom : {ℓ : Level} (X Y : Contractable {ℓ}) → Set ℓ
55 | Hom X Y = Carrier X → Carrier Y
56 |
57 | preservation : {ℓ : Level} {X Y : Contractable {ℓ}}
58 | → {mor : Hom X Y} → mor (Id X) ≡ Id Y
59 | preservation {Y = Y} = collapse Y
60 | \end{code}
61 | %}}}
62 |
63 | %{{{ PointedAlg ; PointedCat ; Forget
64 | \subsection{Category and Forgetful Functors}
65 |
66 | Since there is only one type, or sort, involved in the definition,
67 | we may hazard these structures as ``one sorted algebras'':
68 |
69 | \begin{code}
70 | oneSortedAlg : ∀ {ℓ} → OneSortedAlg ℓ
71 | oneSortedAlg = record
72 | { Alg = Contractable
73 | ; Carrier = Carrier
74 | ; Hom = Hom
75 | ; mor = id
76 | ; comp = λ f g → f ∘ g
77 | ; comp-is-∘ = ≐-refl
78 | ; Id = id
79 | ; Id-is-id = ≐-refl
80 | }
81 | \end{code}
82 |
83 | From which we immediately obtain a category and a forgetful functor.
84 |
85 | \begin{code}
86 | Contractables : (ℓ : Level) → Category (lsuc ℓ) ℓ ℓ
87 | Contractables ℓ = oneSortedCategory ℓ oneSortedAlg
88 |
89 | Forget : (ℓ : Level) → Functor (Contractables ℓ) (Sets ℓ)
90 | Forget ℓ = mkForgetful ℓ oneSortedAlg
91 | \end{code}
92 |
93 | %}}}
94 |
95 | %{{{ Free ; MaybeLeft ; NoRight
96 | \subsection{A Free Construction}
97 |
98 | As discussed earlier, the prime example of pointed algebras are the optional types,
99 | and this claim can be realised as a functor:
100 | \begin{code}
101 | open import Structures.OneCat hiding (initial ; Free ; Forget)
102 |
103 | One𝒞 : {ℓ : Level} → Contractable {ℓ}
104 | One𝒞 = MkContractable One ⋆ (sym uip-One)
105 |
106 | Free : (ℓ : Level) → Functor (Sets ℓ) (Contractables ℓ)
107 | Free ℓ = record
108 | { F₀ = λ _ → One𝒞
109 | ; F₁ = λ _ → id
110 | ; identity = λ _ → refl
111 | ; homomorphism = λ _ → refl
112 | ; F-resp-≡ = λ _ _ → refl
113 | }
114 | \end{code}
115 |
116 | Which is indeed deserving of its name:
117 |
118 | \begin{code}
119 | Left : (ℓ : Level) → Adjunction (Free ℓ) (Forget ℓ)
120 | Left ℓ = record
121 | { unit = record { η = λ _ _ → ⋆ ; commute = λ _ → ≡.refl }
122 | ; counit = record
123 | { η = λ X _ → Id X
124 | ; commute = λ {X} {Y} _ _ → sym (collapse Y)
125 | }
126 | ; zig = λ _ → sym uip-One
127 | ; zag = λ {X} → collapse X
128 | }
129 | \end{code}
130 |
131 | Note that we could have used the isomorphic type Proxy, below, instead of One
132 | but we do not want to increase the amount of new types needlessly.
133 | \begin{code}
134 | data Proxy {ℓ} (A : Set ℓ) : Set ℓ where
135 | MkProxy : Proxy A
136 | \end{code}
137 |
138 | While there is a ``least'' contractable object for any given set, there is,
139 | in-general, no ``largest'' contractable object corresponding to any given set.
140 | That is, there is no co-free functor.
141 |
142 | \begin{code}
143 | open import Helpers.DataProperties
144 |
145 | NoRight : {ℓ : Level} → (CoFree : Functor (Sets ℓ) (Contractables ℓ)) → Adjunction (Forget ℓ) CoFree → ⊥
146 | NoRight {ℓ} (record { F₀ = f }) Adjunct = η (counit Adjunct) ⊥ (Id (f ⊥))
147 | where open Adjunction
148 | open NaturalTransformation
149 | \end{code}
150 | %}}}
151 |
152 | %{{{ Initial object
153 |
154 | A singleton set with its only point determines an initial object in this category.
155 | That is, One is not only the free object on a set but is special that way.
156 |
157 | \begin{code}
158 |
159 | initial : {ℓ : Level} → Initial (Contractables ℓ)
160 | initial = record
161 | { ⊥ = One𝒞
162 | ; ! = λ{ {MkContractable X x p} → (𝑲 x) }
163 | ; !-unique = λ {A} _ _ → sym (collapse A)
164 | }
165 | \end{code}
166 |
167 | %}}}
168 |
169 | %{{{ 0-ary adjoint
170 | \begin{code}
171 | module ZeroAryAdjoint where
172 |
173 | Forget-0 : (ℓ : Level) → Functor (Contractables ℓ) (OneCat ℓ ℓ ℓ)
174 | Forget-0 ℓ = MakeForgetfulFunctor Carrier
175 |
176 | -- OneCat can be, itself, viewed as a pointed set; i.e., an object of Contractables.
177 | Free-0 : (ℓ : Level) → Functor (OneCat ℓ ℓ ℓ) (Contractables ℓ)
178 | Free-0 ℓ = MakeFreeFunctor One𝒞
179 |
180 | Left′ : (ℓ : Level) → Adjunction (Free-0 ℓ) (Forget-0 ℓ)
181 | Left′ ℓ = Make-Free⊢Forget {C = Contractables ℓ} Carrier initial
182 | \end{code}
183 | %}}}
184 |
185 | So much for contractable structures.
186 |
187 | % Quick Folding Instructions:
188 | % C-c C-s :: show/unfold region
189 | % C-c C-h :: hide/fold region
190 | % C-c C-w :: whole file fold
191 | % C-c C-o :: whole file unfold
192 | %
193 | % Local Variables:
194 | % folded-file: t
195 | % eval: (fold-set-marks "%{{{ " "%}}}")
196 | % eval: (fold-whole-buffer)
197 | % fold-internal-margins: 0
198 | % end:
199 |
--------------------------------------------------------------------------------
/Structures/Experiments/Sidequest/Permutations/Vector.lagda:
--------------------------------------------------------------------------------
1 | \section{Structures.Sidequest.Permutations}
2 |
3 | ‼ THIS HAS A POSTULATE ‼
4 |
5 | Documenting the relationship between |Vec| and |Permutation|.
6 |
7 | The following is inspired by copumkin & vmchale's libraries.
8 |
9 | %{{{ Imports
10 | \begin{code}
11 | module Structures.Sidequest.Permutations.Vector where
12 |
13 | open import Level using (Level)
14 | open import Function using (_$_)
15 | open import DataProperties hiding (⟨_,_⟩)
16 | open import EqualityCombinators
17 |
18 | open import Function using (_$_) renaming (id to Id₀ ; _∘_ to _∘₀_)
19 |
20 | open import Data.Vec
21 | open import Data.Vec.Properties using (lookup∘tabulate; lookup-allFin)
22 | open import Data.Nat hiding (fold ; _*_)
23 | open import Data.Fin hiding (_+_ ; fold ; _≤_)
24 |
25 | open import Data.Sum using () renaming (map to _⊎₁_; [_,_] to either)
26 |
27 | open import Structures.Sidequest.Permutations.Basic
28 | \end{code}
29 | %}}}
30 |
31 | Note that the most straightforward implementation of |toVector′| gives us
32 | things "backwards": Elements of |Fin n| of length |m|.
33 | Further, this is completely different than |seeP|, as |toVector′| really gives a
34 | way to see the action on |allFin|.
35 |
36 | \begin{code}
37 | toVector′ : {n m : ℕ} → Permutation n m → Vec (Fin n) m
38 | toVector′ p = p ◈ allFin _
39 | \end{code}
40 |
41 | %{{{ Efforts for deletion
42 |
43 | Attempt to tighten the bound on a Fin; i.e., |Sidequest.idris|.
44 | \begin{code}
45 | tighten : {m : ℕ} → Fin (suc m) → Fin (suc m) ⊎ Fin m
46 | tighten {zero} zero = inj₁ zero
47 | tighten {zero} (suc ())
48 | tighten {suc m} zero = inj₂ zero
49 | tighten {suc m} (suc i) = (suc ⊎₁ suc) (tighten i)
50 |
51 | -- spec : {m : ℕ} {i : Fin (suc m)} (i