Skip to content

Commit 87b26d9

Browse files
authored
[ refactor ] Clean up the dependency graph (#1435)
* [ refactor ] Clean up the dependency graph
1 parent 948b8a6 commit 87b26d9

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

57 files changed

+1753
-1098
lines changed

.gitignore

+1
Original file line numberDiff line numberDiff line change
@@ -5,6 +5,7 @@
55
*.hi
66
*.lagda.el
77
*.o
8+
*.svg
89
*.tix
910
*.vim
1011
*~

CHANGELOG.md

+44-1
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,10 @@ The library has been tested using Agda 2.6.1 and 2.6.1.1.
66
Highlights
77
----------
88

9+
* Drastically reorganised the module hierarchy in the dependency graph of
10+
the `IO` module so that we may compile a program as simple as hello world
11+
without pulling upwards of 130 modules.
12+
913
* First verified implementation of a sorting algorithm (available from `Data.List.Sort`).
1014

1115
Bug-fixes
@@ -14,6 +18,16 @@ Bug-fixes
1418
Non-backwards compatible changes
1519
--------------------------------
1620

21+
* `Data.List.Relation.Binary.Lex.Core` has been thinned to minimise its
22+
dependencies. The more complex properties (`transitive`, `respects₂`,
23+
`[]<[]-⇔`, `∷<∷-⇔`, and `decidable`) have been moved to
24+
`Data.List.Relation.Binary.Lex`.
25+
26+
* `Data.String.Base` has been thinned to minimise its dependencies. The more
27+
complex functions (`parensIfSpace`, `wordsBy`, `words`, `linesBy`, `lines`,
28+
`rectangle`, `rectangleˡ`, `rectangleʳ`, `rectangleᶜ`) have been moved to
29+
`Data.String`.
30+
1731
Deprecated modules
1832
------------------
1933

@@ -116,6 +130,30 @@ New modules
116130
Data.Nat.PseudoRandom.LCG
117131
```
118132

133+
* Broke up `Data.List.Relation.Binary.Pointwise` and introduced:
134+
```
135+
Data.List.Relation.Binary.Pointwise.Base
136+
Data.List.Relation.Binary.Pointwise.Properties
137+
```
138+
139+
* Broke up `Codata.Musical.Colist` into a multitude of modules:
140+
```
141+
Codata.Musical.Colist.Base
142+
Codata.Musical.Colist.Properties
143+
Codata.Musical.Colist.Bisimilarity
144+
Codata.Musical.Colist.Relation.Unary.All
145+
Codata.Musical.Colist.Relation.Unary.All.Properties
146+
Codata.Musical.Colist.Relation.Unary.Any
147+
Codata.Musical.Colist.Relation.Unary.Any.Properties
148+
```
149+
150+
* Broke up `IO` into a many smaller modules:
151+
```
152+
IO.Base
153+
IO.Finite
154+
IO.Infinite
155+
```
156+
119157
* Instantiate a homogeneously indexed bundle at a particular index
120158
```
121159
Relation.Binary.Indexed.Homogeneous.Construct.At
@@ -130,6 +168,11 @@ New modules
130168
Other minor additions
131169
---------------------
132170

171+
* Added new function in `Data.Char.Base`:
172+
```agda
173+
_≈ᵇ_ : (c d : Char) → Bool
174+
```
175+
133176
* Added new proofs in `Algebra.Morphism.GroupMonomorphism`:
134177
```agda
135178
⁻¹-distrib-∙ : ((x ◦ y) ⁻¹₂ ≈₂ (x ⁻¹₂) ◦ (y ⁻¹₂)) → ((x ∙ y) ⁻¹₁ ≈₁ (x ⁻¹₁) ∙ (y ⁻¹₁))
@@ -370,7 +413,7 @@ Other minor additions
370413

371414
* Added new proofs to `Data.List.Relation.Binary.Pointwise`:
372415
```agda
373-
foldr⁺ : (R w x → R y z → R (w • y) (x ◦ z)) →
416+
foldr⁺ : (R w x → R y z → R (w • y) (x ◦ z)) →
374417
R e f → Pointwise R xs ys → R (foldr _•_ e xs) (foldr _◦_ f ys)
375418
lookup⁻ : length xs ≡ length ys →
376419
(toℕ i ≡ toℕ j → R (lookup xs i) (lookup ys j)) →

GNUmakefile

+1-2
Original file line numberDiff line numberDiff line change
@@ -23,8 +23,7 @@ Everything.agda:
2323
# command `cabal install` is needed by cabal-install <= 2.4.*. I did
2424
# not found any problem running both commands with different versions
2525
# of cabal-install. See Issue #1001.
26-
cabal clean && cabal build && cabal install
27-
cabal exec -- GenerateEverything
26+
cabal run GenerateEverything
2827

2928
.PHONY: listings
3029
listings: Everything.agda

GenerateEverything.hs

+12
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,13 @@ unsafeModules :: [FilePath]
2626
unsafeModules = map modToFile
2727
[ "Codata.Musical.Cofin"
2828
, "Codata.Musical.Colist"
29+
, "Codata.Musical.Colist.Base"
30+
, "Codata.Musical.Colist.Properties"
31+
, "Codata.Musical.Colist.Bisimilarity"
32+
, "Codata.Musical.Colist.Relation.Unary.All"
33+
, "Codata.Musical.Colist.Relation.Unary.All.Properties"
34+
, "Codata.Musical.Colist.Relation.Unary.Any"
35+
, "Codata.Musical.Colist.Relation.Unary.Any.Properties"
2936
, "Codata.Musical.Colist.Infinite-merge"
3037
, "Codata.Musical.Conat"
3138
, "Codata.Musical.Costring"
@@ -39,7 +46,12 @@ unsafeModules = map modToFile
3946
, "Foreign.Haskell.Maybe"
4047
, "Foreign.Haskell.Pair"
4148
, "IO"
49+
, "IO.Base"
50+
, "IO.Infinite"
51+
, "IO.Finite"
4252
, "IO.Primitive"
53+
, "IO.Primitive.Infinite"
54+
, "IO.Primitive.Finite"
4355
, "Relation.Binary.PropositionalEquality.TrustMe"
4456
, "System.Environment"
4557
, "System.Environment.Primitive"

README/Debug/Trace.agda

+2-1
Original file line numberDiff line numberDiff line change
@@ -59,10 +59,11 @@ div m n = just (go m m) where
5959
open import Level using (0ℓ)
6060
open import IO
6161

62+
main : Main
6263
main =
6364
let r = trace "Call to div" (div 4 2)
6465
j = λ n trace "Forcing the result wrapped in just." (putStrLn (show n)) in
65-
run {a = 0ℓ} (maybe′ j (return _) r)
66+
run (maybe′ j (return _) r)
6667

6768
-- We get the following trace where we can see that checking that the
6869
-- maybe-solution is just-headed does not force the natural number. Once forced,

README/IO.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ module README.IO where
1111
open import Level
1212
open import Data.Nat.Base
1313
open import Data.Nat.Show using (show)
14-
open import Data.String.Base using (String; _++_; lines)
14+
open import Data.String using (String; _++_; lines)
1515
open import Data.Unit.Polymorphic
1616
open import IO
1717

graph.sh

+32
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,32 @@
1+
#!/bin/sh
2+
3+
### You can call this script like so to generate a dependency graph
4+
### of the `Data.List.Base` module:
5+
### ./graph.sh src/Data/List/Base.agda
6+
7+
### Allow users to pick the agda executable they want by prefixing
8+
### the call with `AGDA=agda-X.Y.Z` and default to agda in case
9+
### nothing was picked
10+
AGDA=${AGDA:-"agda"}
11+
12+
### Grab the directory and name of the target agda file
13+
DIR=$(dirname $1)
14+
BASE=$(basename $1 ".agda")
15+
FILE=_build/${DIR}/${BASE}
16+
17+
### Prepare the directory for the dot & tmp files
18+
mkdir -p _build/$DIR
19+
20+
### Generate the dot file for the target agda file
21+
${AGDA} -i. -isrc/ --dependency-graph=${FILE}.dot $1
22+
23+
### Trim the graph to remove transitive dependencies. Without that the
24+
### graphs get too big too quickly and are impossible to render
25+
tred ${FILE}.dot > ${FILE}2.dot
26+
mv ${FILE}2.dot ${FILE}.dot
27+
28+
### Generate an svg representation of the graph
29+
dot -Tsvg ${FILE}.dot > ${FILE}.svg
30+
31+
### Add a symlink to it in the base directory
32+
ln -is ${FILE}.svg ${BASE}.svg

src/Codata/Colist.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -10,16 +10,16 @@ module Codata.Colist where
1010

1111
open import Level using (Level)
1212
open import Size
13-
open import Data.Unit
13+
open import Data.Unit.Base
1414
open import Data.Nat.Base
1515
open import Data.Product using (_×_ ; _,_)
1616
open import Data.These.Base using (These; this; that; these)
1717
open import Data.Maybe.Base using (Maybe; nothing; just)
1818
open import Data.List.Base using (List; []; _∷_)
1919
open import Data.List.NonEmpty as List⁺ using (List⁺; _∷_)
2020
open import Data.Vec.Base as Vec using (Vec; []; _∷_)
21-
open import Data.Vec.Bounded as Vec≤ using (Vec≤)
22-
open import Function
21+
open import Data.Vec.Bounded.Base as Vec≤ using (Vec≤)
22+
open import Function.Base using (_$′_; _∘′_; id; _∘_)
2323

2424
open import Codata.Thunk using (Thunk; force)
2525
open import Codata.Conat as Conat using (Conat ; zero ; suc)

src/Codata/Cowriter.agda

+3-3
Original file line numberDiff line numberDiff line change
@@ -17,15 +17,15 @@ open import Codata.Thunk using (Thunk; force)
1717
open import Codata.Conat
1818
open import Codata.Delay using (Delay; later; now)
1919
open import Codata.Stream as Stream using (Stream; _∷_)
20-
open import Data.Unit
20+
open import Data.Unit.Base
2121
open import Data.List.Base using (List; []; _∷_)
2222
open import Data.List.NonEmpty using (List⁺; _∷_)
2323
open import Data.Nat.Base as Nat using (ℕ; zero; suc)
2424
open import Data.Product as Prod using (_×_; _,_)
2525
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂)
2626
open import Data.Vec.Base using (Vec; []; _∷_)
27-
open import Data.Vec.Bounded as Vec≤ using (Vec≤; _,_)
28-
open import Function
27+
open import Data.Vec.Bounded.Base as Vec≤ using (Vec≤; _,_)
28+
open import Function.Base using (_$_; _∘′_; id)
2929

3030
private
3131
variable

src/Codata/Delay.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -19,7 +19,7 @@ open import Data.Maybe.Base hiding (map ; fromMaybe ; zipWith ; alignWith ; zip
1919
open import Data.Product as P hiding (map ; zip)
2020
open import Data.Sum.Base as S hiding (map)
2121
open import Data.These.Base as T using (These; this; that; these)
22-
open import Function
22+
open import Function.Base using (id)
2323

2424
------------------------------------------------------------------------
2525
-- Definition

0 commit comments

Comments
 (0)