Skip to content

Commit 8d4b5a7

Browse files
authored
Merge branch 'agda:master' into decFinSubset
2 parents 5d0f79f + 2ce5eb3 commit 8d4b5a7

31 files changed

+1138
-94
lines changed

.github/workflows/ci-ubuntu.yml

Lines changed: 6 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -48,8 +48,8 @@ on:
4848
########################################################################
4949

5050
env:
51-
GHC_VERSION: 9.10.1
52-
CABAL_VERSION: 3.12.1.0
51+
GHC_VERSION: 9.12.2
52+
CABAL_VERSION: 3.16.0.0
5353
CABAL_V1_INSTALL: cabal v1-install --ghc-options='-O1 +RTS -M6G -RTS'
5454
CABAL_INSTALL: cabal install --overwrite-policy=always --ghc-options='-O1 +RTS -M6G -RTS'
5555
AGDA: agda -Werror +RTS -M5G -H3.5G -A128M -RTS -i. -isrc -idoc
@@ -76,11 +76,12 @@ jobs:
7676
if [[ '${{ github.ref }}' == 'refs/heads/experimental' \
7777
|| '${{ github.base_ref }}' == 'experimental' ]]; then
7878
# Pick Agda version for experimental
79-
echo "AGDA_COMMIT=ef912c68fd329ad3046d156e3c1a70a7fec19ba1" >> "${GITHUB_ENV}";
79+
echo "AGDA_COMMIT=3d04bacca842729f9c0869b9287256321b5f450f" >> "${GITHUB_ENV}";
80+
# Andreas, 2025-10-07: 3d04bacca842729f9c0869b9287256321b5f450f is tags/v2.8.0
8081
echo "AGDA_HTML_DIR=html/experimental" >> "${GITHUB_ENV}"
8182
else
8283
# Pick Agda version for master
83-
echo "AGDA_COMMIT=tags/v2.8.0-rc3" >> "${GITHUB_ENV}";
84+
echo "AGDA_COMMIT=tags/v2.8.0" >> "${GITHUB_ENV}";
8485
echo "AGDA_HTML_DIR=html/master" >> "${GITHUB_ENV}"
8586
fi
8687
@@ -148,7 +149,7 @@ jobs:
148149

149150
# By default github actions do not pull the repo
150151
- name: Checkout stdlib
151-
uses: actions/checkout@v4
152+
uses: actions/checkout@v5
152153

153154
- name: Test stdlib
154155
run: |

CHANGELOG.md

Lines changed: 70 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -11,6 +11,8 @@ Bug-fixes
1111

1212
* Fix a type error in `README.Data.Fin.Relation.Unary.Top` within the definition of `>-weakInduction`.
1313

14+
* Fix a typo in `Algebra.Morphism.Construct.DirectProduct`.
15+
1416
Non-backwards compatible changes
1517
--------------------------------
1618

@@ -47,18 +49,59 @@ Deprecated names
4749
New modules
4850
-----------
4951

52+
* `Algebra.Properties.BooleanRing`.
53+
54+
* `Algebra.Properties.BooleanSemiring`.
55+
56+
* `Algebra.Properties.CommutativeRing`.
57+
58+
* `Algebra.Properties.Semiring`.
59+
5060
* `Data.List.Relation.Binary.Permutation.Algorithmic{.Properties}` for the Choudhury and Fiore definition of permutation, and its equivalence with `Declarative` below.
5161

5262
* `Data.List.Relation.Binary.Permutation.Declarative{.Properties}` for the least congruence on `List` making `_++_` commutative, and its equivalence with the `Setoid` definition.
5363

64+
* `Effect.Monad.Random` and `Effect.Monad.Random.Instances` for an mtl-style randomness monad constraint.
65+
5466
Additions to existing modules
5567
-----------------------------
5668

69+
* In `Algebra.Bundles`:
70+
```agda
71+
record BooleanSemiring _ _ : Set _
72+
record BooleanRing _ _ : Set _
73+
```
74+
75+
* In `Algebra.Consequences.Propositional`:
76+
```agda
77+
binomial-expansion : Associative _∙_ → _◦_ DistributesOver _∙_ →
78+
∀ w x y z → ((w ∙ x) ◦ (y ∙ z)) ≡ ((((w ◦ y) ∙ (w ◦ z)) ∙ (x ◦ y)) ∙ (x ◦ z))
79+
```
80+
81+
* In `Algebra.Consequences.Setoid`:
82+
```agda
83+
binomial-expansion : Congruent₂ _∙_ → Associative _∙_ → _◦_ DistributesOver _∙_ →
84+
∀ w x y z → ((w ∙ x) ◦ (y ∙ z)) ≈ ((((w ◦ y) ∙ (w ◦ z)) ∙ (x ◦ y)) ∙ (x ◦ z))
85+
```
86+
87+
* In `Algebra.Lattice.Properties.BooleanAlgebra.XorRing`:
88+
```agda
89+
⊕-∧-isBooleanRing : IsBooleanRing _⊕_ _∧_ id ⊥ ⊤
90+
⊕-∧-booleanRing : BooleanRing _ _
91+
```
92+
5793
* In `Algebra.Properties.RingWithoutOne`:
5894
```agda
5995
[-x][-y]≈xy : ∀ x y → - x * - y ≈ x * y
6096
```
6197

98+
* In `Algebra.Structures`:
99+
```agda
100+
record IsBooleanSemiring + * 0# 1# : Set _
101+
record IsBooleanRing + * - 0# 1# : Set _
102+
```
103+
NB. the latter is based on `IsCommutativeRing`, with the former on `IsSemiring`.
104+
62105
* In `Data.Fin.Permutation.Components`:
63106
```agda
64107
transpose[i,i,j]≡j : (i j : Fin n) → transpose i i j ≡ j
@@ -91,6 +134,27 @@ Additions to existing modules
91134

92135
* In `Data.Vec.Properties`:
93136
```agda
137+
updateAt-take : (xs : Vec A (m + n)) (i : Fin m) (f : A → A) →
138+
updateAt (take m xs) i f ≡ take m (updateAt xs (inject≤ i (m≤m+n m n)) f)
139+
140+
truncate-zipWith : (f : A → B → C) (m≤n : m ≤ n) (xs : Vec A n) (ys : Vec B n) →
141+
truncate m≤n (zipWith f xs ys) ≡ zipWith f (truncate m≤n xs) (truncate m≤n ys)
142+
143+
truncate-zipWith-truncate : (f : A → B → C) (m≤n : m ≤ n) (n≤o : n ≤ o) (xs : Vec A o) (ys : Vec B n) →
144+
truncate m≤n (zipWith f (truncate n≤o xs) ys) ≡
145+
zipWith f (truncate (≤-trans m≤n n≤o) xs) (truncate m≤n ys)
146+
147+
truncate-updateAt : (m≤n : m ≤ n) (xs : Vec A n) (i : Fin m) (f : A → A) →
148+
updateAt (truncate m≤n xs) i f ≡
149+
truncate m≤n (updateAt xs (inject≤ i m≤n) f)
150+
151+
updateAt-truncate : (xs : Vec A (m + n)) (i : Fin m) (f : A → A) →
152+
updateAt (truncate (m≤m+n m n) xs) i f ≡
153+
truncate (m≤m+n m n) (updateAt xs (inject≤ i (m≤m+n m n)) f)
154+
155+
map-truncate : (f : A → B) (m≤n : m ≤ n) (xs : Vec A n) →
156+
map f (truncate m≤n xs) ≡ truncate m≤n (map f xs)
157+
94158
padRight-lookup : (m≤n : m ≤ n) (a : A) (xs : Vec A m) (i : Fin m) → lookup (padRight m≤n a xs) (inject≤ i m≤n) ≡ lookup xs i
95159
96160
padRight-map : (f : A → B) (m≤n : m ≤ n) (a : A) (xs : Vec A m) → map f (padRight m≤n a xs) ≡ padRight m≤n (f a) (map f xs)
@@ -115,3 +179,9 @@ Additions to existing modules
115179
¬¬-η : A → ¬ ¬ A
116180
contradiction′ : ¬ A → A → Whatever
117181
```
182+
183+
* In `System.Random`:
184+
```agda
185+
randomIO : IO Bool
186+
randomRIO : RandomRIO {A = Bool} _≤_
187+
```

GenerateEverything.hs

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -59,6 +59,8 @@ unsafeModules = map modToFile
5959
, "Debug.Trace"
6060
, "Effect.Monad.IO"
6161
, "Effect.Monad.IO.Instances"
62+
, "Effect.Monad.Random"
63+
, "Effect.Monad.Random.Instances"
6264
, "Foreign.Haskell"
6365
, "Foreign.Haskell.Coerce"
6466
, "Foreign.Haskell.Either"

src/Algebra/Bundles.agda

Lines changed: 80 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -921,6 +921,41 @@ record Quasiring c ℓ : Set (suc (c ⊔ ℓ)) where
921921
; rawMonoid to *-rawMonoid
922922
)
923923

924+
record BooleanSemiring c ℓ : Set (suc (c ⊔ ℓ)) where
925+
infixl 7 _*_
926+
infixl 6 _+_
927+
infix 4 _≈_
928+
field
929+
Carrier : Set c
930+
_≈_ : Rel Carrier ℓ
931+
_+_ : Op₂ Carrier
932+
_*_ : Op₂ Carrier
933+
0# : Carrier
934+
1# : Carrier
935+
isBooleanSemiring : IsBooleanSemiring _≈_ _+_ _*_ 0# 1#
936+
937+
open IsBooleanSemiring isBooleanSemiring public
938+
939+
semiring : Semiring _ _
940+
semiring = record { isSemiring = isSemiring }
941+
942+
open Semiring semiring public
943+
using ( _≉_; +-rawMagma; +-magma; +-unitalMagma; +-commutativeMagma
944+
; +-semigroup; +-commutativeSemigroup
945+
; *-rawMagma; *-magma; *-semigroup
946+
; +-rawMonoid; +-monoid; +-commutativeMonoid
947+
; *-rawMonoid; *-monoid
948+
; rawNearSemiring ; rawSemiring; nearSemiring
949+
; semiringWithoutOne; semiringWithoutAnnihilatingZero
950+
)
951+
952+
*-idempotentMonoid : IdempotentMonoid c ℓ
953+
*-idempotentMonoid = record { isIdempotentMonoid = *-isIdempotentMonoid }
954+
955+
open IdempotentMonoid *-idempotentMonoid public
956+
using () renaming (band to *-band)
957+
958+
924959
------------------------------------------------------------------------
925960
-- Bundles with 2 binary operations, 1 unary operation & 1 element
926961
------------------------------------------------------------------------
@@ -1106,7 +1141,10 @@ record CommutativeRing c ℓ : Set (suc (c ⊔ ℓ)) where
11061141
ring : Ring _ _
11071142
ring = record { isRing = isRing }
11081143

1109-
open Ring ring public using (_≉_; rawRing; +-invertibleMagma; +-invertibleUnitalMagma; +-group; +-abelianGroup)
1144+
open Ring ring public
1145+
using (_≉_; rawRing
1146+
; +-invertibleMagma; +-invertibleUnitalMagma
1147+
; +-group; +-abelianGroup)
11101148

11111149
commutativeSemiring : CommutativeSemiring _ _
11121150
commutativeSemiring =
@@ -1124,6 +1162,47 @@ record CommutativeRing c ℓ : Set (suc (c ⊔ ℓ)) where
11241162
; commutativeSemiringWithoutOne
11251163
)
11261164

1165+
1166+
record BooleanRing c ℓ : Set (suc (c ⊔ ℓ)) where
1167+
infix 8 -_
1168+
infixl 7 _*_
1169+
infixl 6 _+_
1170+
infix 4 _≈_
1171+
field
1172+
Carrier : Set c
1173+
_≈_ : Rel Carrier ℓ
1174+
_+_ : Op₂ Carrier
1175+
_*_ : Op₂ Carrier
1176+
-_ : Op₁ Carrier
1177+
0# : Carrier
1178+
1# : Carrier
1179+
isBooleanRing : IsBooleanRing _≈_ _+_ _*_ -_ 0# 1#
1180+
1181+
open IsBooleanRing isBooleanRing public
1182+
using (isCommutativeRing; *-idem)
1183+
1184+
open IsCommutativeRing isCommutativeRing public
1185+
1186+
commutativeRing : CommutativeRing _ _
1187+
commutativeRing = record { isCommutativeRing = isCommutativeRing }
1188+
1189+
open CommutativeRing commutativeRing public
1190+
using
1191+
(_≉_; rawRing
1192+
; +-invertibleMagma; +-invertibleUnitalMagma
1193+
; +-group; +-abelianGroup
1194+
; +-rawMagma; +-magma; +-unitalMagma; +-commutativeMagma
1195+
; +-semigroup; +-commutativeSemigroup
1196+
; *-rawMagma; *-magma; *-commutativeMagma; *-semigroup; *-commutativeSemigroup
1197+
; +-rawMonoid; +-monoid; +-commutativeMonoid
1198+
; *-rawMonoid; *-monoid; *-commutativeMonoid
1199+
; nearSemiring; semiringWithoutOne
1200+
; semiringWithoutAnnihilatingZero; semiring
1201+
; commutativeSemiringWithoutOne; commutativeSemiring
1202+
; ring
1203+
)
1204+
1205+
11271206
------------------------------------------------------------------------
11281207
-- Bundles with 3 binary operations
11291208
------------------------------------------------------------------------

src/Algebra/Consequences/Propositional.agda

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -42,6 +42,7 @@ open Base public
4242
; subst∧comm⇒sym
4343
; wlog
4444
; sel⇒idem
45+
; binomial-expansion
4546
-- plus all the deprecated versions of the above
4647
; comm+assoc⇒middleFour
4748
; identity+middleFour⇒assoc
@@ -101,6 +102,15 @@ module _ {_∙_ _◦_ : Op₂ A} (∙-comm : Commutative _∙_) where
101102
comm⇒sym[distribˡ] : x Symmetric (λ y z (x ◦ (y ∙ z)) ≡ ((x ◦ y) ∙ (x ◦ z)))
102103
comm⇒sym[distribˡ] = Base.comm⇒sym[distribˡ] (cong₂ _◦_) ∙-comm
103104

105+
module _ {_∙_ _◦_ : Op₂ A}
106+
(∙-assoc : Associative _∙_)
107+
(distrib : _◦_ DistributesOver _∙_)
108+
where
109+
110+
binomial-expansion : w x y z
111+
((w ∙ x) ◦ (y ∙ z)) ≡ ((((w ◦ y) ∙ (w ◦ z)) ∙ (x ◦ y)) ∙ (x ◦ z))
112+
binomial-expansion = Base.binomial-expansion {_∙_} {_◦_} (cong₂ _) ∙-assoc distrib
113+
104114
------------------------------------------------------------------------
105115
-- Selectivity
106116

src/Algebra/Consequences/Setoid.agda

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -292,6 +292,21 @@ module _ {_∙_ _◦_ : Op₂ A}
292292
(x ◦ (x ∙ z)) ∙ (y ◦ (x ∙ z)) ≈⟨ ◦-distribʳ-∙ _ _ _ ⟨
293293
(x ∙ y) ◦ (x ∙ z) ∎
294294

295+
module _ {_∙_ _◦_ : Op₂ A}
296+
(∙-cong : Congruent₂ _∙_)
297+
(∙-assoc : Associative _∙_)
298+
(distrib@(distribˡ , distribʳ) : _◦_ DistributesOver _∙_)
299+
where
300+
301+
binomial-expansion : w x y z
302+
((w ∙ x) ◦ (y ∙ z)) ≈ ((((w ◦ y) ∙ (w ◦ z)) ∙ (x ◦ y)) ∙ (x ◦ z))
303+
binomial-expansion w x y z = begin
304+
(w ∙ x) ◦ (y ∙ z) ≈⟨ distribʳ _ _ _ ⟩
305+
(w ◦ (y ∙ z)) ∙ (x ◦ (y ∙ z)) ≈⟨ ∙-cong (distribˡ _ _ _) (distribˡ _ _ _) ⟩
306+
((w ◦ y) ∙ (w ◦ z)) ∙ ((x ◦ y) ∙ (x ◦ z)) ≈⟨ ∙-assoc _ _ _ ⟨
307+
(((w ◦ y) ∙ (w ◦ z)) ∙ (x ◦ y)) ∙ (x ◦ z) ∎
308+
309+
295310
------------------------------------------------------------------------
296311
-- Ring-like structures
297312

src/Algebra/Lattice/Properties/BooleanAlgebra.agda

Lines changed: 17 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -6,30 +6,31 @@
66

77
{-# OPTIONS --cubical-compatible --safe #-}
88

9-
open import Algebra.Lattice.Bundles
9+
open import Algebra.Lattice.Bundles using (BooleanAlgebra)
1010

1111
module Algebra.Lattice.Properties.BooleanAlgebra
12-
{b₁ b₂} (B : BooleanAlgebra b₁ b₂)
12+
{c ℓ} (booleanAlgebra : BooleanAlgebra c ℓ)
1313
where
1414

15-
open BooleanAlgebra B
15+
open import Algebra.Bundles
16+
using (CommutativeSemiring; CommutativeRing; BooleanRing)
17+
open import Algebra.Core using (Op₂)
18+
open import Data.Product.Base using (_,_)
19+
open import Function.Base using (id; _$_; _⟨_⟩_)
20+
open import Function.Bundles using (_⇔_; module Equivalence)
1621

17-
import Algebra.Lattice.Properties.DistributiveLattice as DistribLatticeProperties
18-
open import Algebra.Core using (Op₁; Op₂)
22+
open BooleanAlgebra booleanAlgebra
1923
open import Algebra.Structures _≈_
2024
open import Algebra.Definitions _≈_
2125
open import Algebra.Consequences.Setoid setoid
22-
open import Algebra.Bundles using (CommutativeSemiring; CommutativeRing)
2326
open import Algebra.Lattice.Structures _≈_
2427
open import Relation.Binary.Reasoning.Setoid setoid
25-
open import Function.Base using (id; _$_; _⟨_⟩_)
26-
open import Function.Bundles using (_⇔_; module Equivalence)
27-
open import Data.Product.Base using (_,_)
28+
2829

2930
------------------------------------------------------------------------
3031
-- Export properties from distributive lattices
3132

32-
open DistribLatticeProperties distributiveLattice public
33+
open import Algebra.Lattice.Properties.DistributiveLattice distributiveLattice public
3334

3435
------------------------------------------------------------------------
3536
-- The dual construction is also a boolean algebra
@@ -529,6 +530,12 @@ module XorRing
529530
{ isCommutativeRing = ⊕-∧-isCommutativeRing
530531
}
531532

533+
⊕-∧-isBooleanRing : IsBooleanRing _⊕_ _∧_ id ⊥ ⊤
534+
⊕-∧-isBooleanRing = record
535+
{ isCommutativeRing = ⊕-∧-isCommutativeRing ; *-idem = ∧-idem }
536+
537+
⊕-∧-booleanRing : BooleanRing _ _
538+
⊕-∧-booleanRing = record { isBooleanRing = ⊕-∧-isBooleanRing }
532539

533540
infixl 6 _⊕_
534541

src/Algebra/Morphism/Construct/DirectProduct.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -133,7 +133,7 @@ module Monoid-Export {M : RawMonoid a ℓ₁} {N : RawMonoid b ℓ₂} where
133133
module N = RawMonoid N
134134

135135
module _ {refl : Reflexive M._≈_} where
136-
proj₁ = Proj₁.isMonoidHomomorphism M M refl
136+
proj₁ = Proj₁.isMonoidHomomorphism M N refl
137137

138138
module _ {refl : Reflexive N._≈_} where
139139
proj₂ = Proj₂.isMonoidHomomorphism M N refl

0 commit comments

Comments
 (0)