Skip to content

Commit 831125c

Browse files
Merge in v2.1.1 and update to Agda 2.7.0
2 parents 4972983 + 039e1f4 commit 831125c

File tree

9 files changed

+55
-33
lines changed

9 files changed

+55
-33
lines changed

.github/workflows/ci-ubuntu.yml

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -76,11 +76,11 @@ jobs:
7676
if [[ '${{ github.ref }}' == 'refs/heads/experimental' \
7777
|| '${{ github.base_ref }}' == 'experimental' ]]; then
7878
# Pick Agda version for experimental
79-
echo "AGDA_COMMIT=4d36cb37f8bfb765339b808b13356d760aa6f0ec" >> "${GITHUB_ENV}";
79+
echo "AGDA_COMMIT=18cc53941e924b144c0f1f3953280ef726009f7e" >> "${GITHUB_ENV}";
8080
echo "AGDA_HTML_DIR=html/experimental" >> "${GITHUB_ENV}"
8181
else
8282
# Pick Agda version for master
83-
echo "AGDA_COMMIT=tags/v2.6.4.3" >> "${GITHUB_ENV}";
83+
echo "AGDA_COMMIT=tags/v2.7.0" >> "${GITHUB_ENV}";
8484
echo "AGDA_HTML_DIR=html/master" >> "${GITHUB_ENV}"
8585
fi
8686

CHANGELOG.md

Lines changed: 2 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1,10 +1,9 @@
11
Version 2.2-dev
22
===============
33

4-
The library has been tested using Agda 2.6.4.3.
4+
The library has been tested using Agda 2.7.0.
55

6-
Highlights
7-
----------
6+
* Fixed `Reflection.AST.Definition` to take into account Agda now exposes a `Quantity` argument on the reflection `constructor` constructor.
87

98
Bug-fixes
109
---------

CHANGELOG/v2.1.1.md

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,8 @@
1+
Version 2.1.1
2+
=============
3+
4+
The library has been tested using Agda 2.7.0.
5+
6+
* Fixed `Reflection.AST.Definition` to take into account Agda now exposes a `Quantity` argument on the reflection `constructor` constructor.
7+
8+
* In `Reflection.AST.Show` added new function `showQuantity`

CITATION.cff

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,6 @@ message: "If you use this software, please cite it as below."
33
authors:
44
- name: "The Agda Community"
55
title: "Agda Standard Library"
6-
version: 2.1
7-
date-released: 2024-07-17
6+
version: 2.1.1
7+
date-released: 2024-09-03
88
url: "https://github.com/agda/agda-stdlib"

doc/README.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,7 @@
33
module README where
44

55
------------------------------------------------------------------------
6-
-- The Agda standard library, version 2.0
6+
-- The Agda standard library, version 2.1.1
77
--
88
-- Authors: Nils Anders Danielsson, Matthew Daggitt, Guillaume Allais
99
-- with contributions from Andreas Abel, Stevan Andjelkovic,
@@ -19,7 +19,7 @@ module README where
1919
-- and other anonymous contributors.
2020
------------------------------------------------------------------------
2121

22-
-- This version of the library has been tested using Agda 2.6.4.
22+
-- This version of the library has been tested using Agda 2.7.0
2323

2424
-- The library comes with a .agda-lib file, for use with the library
2525
-- management system.

doc/installation-guide.md

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -3,19 +3,19 @@ Installation instructions
33

44
Note: the full story on installing Agda libraries can be found at [readthedocs](http://agda.readthedocs.io/en/latest/tools/package-system.html).
55

6-
Use version v2.1 of the standard library with Agda 2.6.4.X. You can find the correct version of the library to use for different Agda versions on the [Agda Uncyclo](https://wiki.portal.chalmers.se/agda/Libraries/StandardLibrary).
6+
Use version v2.1.1 of the standard library with Agda 2.7.0.
77

88
1. Navigate to a suitable directory `$HERE` (replace appropriately) where
99
you would like to install the library.
1010

11-
2. Download the tarball of v2.1 of the standard library. This can either be
11+
2. Download the tarball of v2.1.1 of the standard library. This can either be
1212
done manually by visiting the Github repository for the library, or via the
1313
command line as follows:
1414
```
15-
wget -O agda-stdlib.tar https://github.com/agda/agda-stdlib/archive/v2.1.tar.gz
15+
wget -O agda-stdlib.tar https://github.com/agda/agda-stdlib/archive/v2.1.1.tar.gz
1616
```
1717
Note that you can replace `wget` with other popular tools such as `curl` and that
18-
you can replace `2.1` with any other version of the library you desire.
18+
you can replace `2.1.1` with any other version of the library you desire.
1919

2020
3. Extract the standard library from the tarball. Again this can either be
2121
done manually or via the command line as follows:
@@ -26,7 +26,7 @@ Use version v2.1 of the standard library with Agda 2.6.4.X. You can find the cor
2626
4. [ OPTIONAL ] If using [cabal](https://www.haskell.org/cabal/) then run
2727
the commands to install via cabal:
2828
```
29-
cd agda-stdlib-2.1
29+
cd agda-stdlib-2.1.1
3030
cabal install
3131
```
3232

@@ -40,7 +40,7 @@ Use version v2.1 of the standard library with Agda 2.6.4.X. You can find the cor
4040
6. Register the standard library with Agda's package system by adding
4141
the following line to `$HOME/.agda/libraries`:
4242
```
43-
$HERE/agda-stdlib-2.1/standard-library.agda-lib
43+
$HERE/agda-stdlib-2.1.1/standard-library.agda-lib
4444
```
4545

4646
Now, the standard library is ready to be used either:

src/Reflection/AST/Definition.agda

Lines changed: 25 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -15,9 +15,10 @@ open import Relation.Nullary.Decidable.Core using (map′; _×-dec_;
1515
open import Relation.Binary.Definitions using (DecidableEquality)
1616
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong; cong₂)
1717

18-
import Reflection.AST.Argument as Arg
19-
import Reflection.AST.Name as Name
20-
import Reflection.AST.Term as Term
18+
import Reflection.AST.Argument as Arg
19+
import Reflection.AST.Argument.Quantity as Quantity
20+
import Reflection.AST.Name as Name
21+
import Reflection.AST.Term as Term
2122

2223
------------------------------------------------------------------------
2324
-- Re-exporting type publically
@@ -56,8 +57,14 @@ record′-injective₂ refl = refl
5657
record′-injective : {c c′ fs fs′} record′ c fs ≡ record′ c′ fs′ c ≡ c′ × fs ≡ fs′
5758
record′-injective = < record′-injective₁ , record′-injective₂ >
5859

59-
constructor′-injective : {c c′} constructor′ c ≡ constructor′ c′ c ≡ c′
60-
constructor′-injective refl = refl
60+
constructor′-injective₁ : {c c′ q q′} constructor′ c q ≡ constructor′ c′ q′ c ≡ c′
61+
constructor′-injective₁ refl = refl
62+
63+
constructor′-injective₂ : {c c′ q q′} constructor′ c q ≡ constructor′ c′ q′ q ≡ q′
64+
constructor′-injective₂ refl = refl
65+
66+
constructor′-injective : {c c′ q q′} constructor′ c q ≡ constructor′ c′ q′ c ≡ c′ × q ≡ q′
67+
constructor′-injective = < constructor′-injective₁ , constructor′-injective₂ >
6168

6269
infix 4 _≟_
6370

@@ -70,39 +77,40 @@ data-type pars cs ≟ data-type pars′ cs′ =
7077
record′ c fs ≟ record′ c′ fs′ =
7178
map′ (uncurry (cong₂ record′)) record′-injective
7279
(c Name.≟ c′ ×-dec List.≡-dec (Arg.≡-dec Name._≟_) fs fs′)
73-
constructor′ d ≟ constructor′ d′ =
74-
map′ (cong constructor′) constructor′-injective (d Name.≟ d′)
80+
constructor′ d q ≟ constructor′ d′ q′ =
81+
map′ (uncurry (cong₂ constructor′)) constructor′-injective
82+
(d Name.≟ d′ ×-dec q Quantity.≟ q′)
7583
axiom ≟ axiom = yes refl
7684
primitive′ ≟ primitive′ = yes refl
7785

7886
-- antidiagonal
7987
function cs ≟ data-type pars cs₁ = no (λ ())
8088
function cs ≟ record′ c fs = no (λ ())
81-
function cs ≟ constructor′ d = no (λ ())
89+
function cs ≟ constructor′ d q = no (λ ())
8290
function cs ≟ axiom = no (λ ())
8391
function cs ≟ primitive′ = no (λ ())
8492
data-type pars cs ≟ function cs₁ = no (λ ())
8593
data-type pars cs ≟ record′ c fs = no (λ ())
86-
data-type pars cs ≟ constructor′ d = no (λ ())
94+
data-type pars cs ≟ constructor′ d q = no (λ ())
8795
data-type pars cs ≟ axiom = no (λ ())
8896
data-type pars cs ≟ primitive′ = no (λ ())
8997
record′ c fs ≟ function cs = no (λ ())
9098
record′ c fs ≟ data-type pars cs = no (λ ())
91-
record′ c fs ≟ constructor′ d = no (λ ())
99+
record′ c fs ≟ constructor′ d q = no (λ ())
92100
record′ c fs ≟ axiom = no (λ ())
93101
record′ c fs ≟ primitive′ = no (λ ())
94-
constructor′ d ≟ function cs = no (λ ())
95-
constructor′ d ≟ data-type pars cs = no (λ ())
96-
constructor′ d ≟ record′ c fs = no (λ ())
97-
constructor′ d ≟ axiom = no (λ ())
98-
constructor′ d ≟ primitive′ = no (λ ())
102+
constructor′ d q ≟ function cs = no (λ ())
103+
constructor′ d q ≟ data-type pars cs = no (λ ())
104+
constructor′ d q ≟ record′ c fs = no (λ ())
105+
constructor′ d q ≟ axiom = no (λ ())
106+
constructor′ d q ≟ primitive′ = no (λ ())
99107
axiom ≟ function cs = no (λ ())
100108
axiom ≟ data-type pars cs = no (λ ())
101109
axiom ≟ record′ c fs = no (λ ())
102-
axiom ≟ constructor′ d = no (λ ())
110+
axiom ≟ constructor′ d q = no (λ ())
103111
axiom ≟ primitive′ = no (λ ())
104112
primitive′ ≟ function cs = no (λ ())
105113
primitive′ ≟ data-type pars cs = no (λ ())
106114
primitive′ ≟ record′ c fs = no (λ ())
107-
primitive′ ≟ constructor′ d = no (λ ())
115+
primitive′ ≟ constructor′ d q = no (λ ())
108116
primitive′ ≟ axiom = no (λ ())

src/Reflection/AST/Instances.agda

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -17,6 +17,7 @@ import Reflection.AST.Abstraction as Abstraction
1717
import Reflection.AST.Argument as Argument
1818
import Reflection.AST.Argument.Visibility as Visibility
1919
import Reflection.AST.Argument.Relevance as Relevance
20+
import Reflection.AST.Argument.Quantity as Quantity
2021
import Reflection.AST.Argument.Information as Information
2122
import Reflection.AST.Pattern as Pattern
2223
import Reflection.AST.Term as Term
@@ -37,6 +38,7 @@ instance
3738
Meta-≡-isDecEquivalence = isDecEquivalence Meta._≟_
3839
Visibility-≡-isDecEquivalence = isDecEquivalence Visibility._≟_
3940
Relevance-≡-isDecEquivalence = isDecEquivalence Relevance._≟_
41+
Quantity-≡-isDecEquivalence = isDecEquivalence Quantity._≟_
4042
ArgInfo-≡-isDecEquivalence = isDecEquivalence Information._≟_
4143
Pattern-≡-isDecEquivalence = isDecEquivalence Pattern._≟_
4244
Clause-≡-isDecEquivalence = isDecEquivalence Term._≟-Clause_

src/Reflection/AST/Show.agda

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -24,6 +24,7 @@ open import Relation.Nullary.Decidable.Core using (yes; no)
2424

2525
open import Reflection.AST.Abstraction hiding (map)
2626
open import Reflection.AST.Argument hiding (map)
27+
open import Reflection.AST.Argument.Quantity
2728
open import Reflection.AST.Argument.Relevance
2829
open import Reflection.AST.Argument.Visibility
2930
open import Reflection.AST.Argument.Modality
@@ -58,6 +59,10 @@ showVisibility visible = "visible"
5859
showVisibility hidden = "hidden"
5960
showVisibility instance′ = "instance"
6061

62+
showQuantity : Quantity String
63+
showQuantity quantity-0 = "quantity-0"
64+
showQuantity quantity-ω = "quantity-ω"
65+
6166
showLiteral : Literal String
6267
showLiteral (nat x) = ℕ.show x
6368
showLiteral (word64 x) = ℕ.show (Word64.toℕ x)
@@ -144,6 +149,6 @@ showDefinition (data-type pars cs) =
144149
"datatype" <+> ℕ.show pars <+> braces (intersperse ", " (map showName cs))
145150
showDefinition (record′ c fs) =
146151
"record" <+> showName c <+> braces (intersperse ", " (map (showName ∘′ unArg) fs))
147-
showDefinition (constructor′ d) = "constructor" <+> showName d
152+
showDefinition (constructor′ d q) = "constructor" <+> showName d <+> showQuantity q
148153
showDefinition axiom = "axiom"
149154
showDefinition primitive′ = "primitive"

0 commit comments

Comments
 (0)