Skip to content

Commit 25dba60

Browse files
Add Effect.Foldable and Data.List.Effectful.Foldable implementation (#2300)
* draft initial implementation * added `List` instances * fix module name * non-dependent instance for `vec` * tweak * added Nathan's morphism proofs * add module for `Vec` * tweak `import`s * removed dependency * removed dependency; more permissive `open` * simplify with a local definition of the hom * rename lemma; add a redundant lemma * streamline `import` interface(s) * remove the word instance from any of the comments. Tighten imports. * Update Foldable.agda Fixes the deprecated module import i hope --------- Co-authored-by: Jacques Carette <[email protected]>
1 parent fdb4e57 commit 25dba60

File tree

4 files changed

+212
-0
lines changed

4 files changed

+212
-0
lines changed

CHANGELOG.md

+6
Original file line numberDiff line numberDiff line change
@@ -148,6 +148,12 @@ New modules
148148
```
149149
plus adding `rawX` fields to each of `Relation.Binary.Bundles.X`.
150150

151+
* `Data.List.Effectful.Foldable`: `List` is `Foldable`
152+
153+
* `Data.Vec.Effectful.Foldable`: `Vec` is `Foldable`
154+
155+
* `Effect.Foldable`: implementation of haskell-like `Foldable`
156+
151157
Additions to existing modules
152158
-----------------------------
153159

src/Data/List/Effectful/Foldable.agda

+78
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,78 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- List is Foldable
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Data.List.Effectful.Foldable where
10+
11+
open import Algebra.Bundles using (Monoid)
12+
open import Algebra.Bundles.Raw using (RawMonoid)
13+
open import Algebra.Morphism.Structures using (IsMonoidHomomorphism)
14+
open import Data.List.Base as List using (List; []; _∷_; _++_)
15+
open import Effect.Foldable using (RawFoldableWithDefaults; RawFoldable)
16+
open import Function.Base using (_∘_; id)
17+
open import Level using (Level)
18+
import Relation.Binary.PropositionalEquality.Core as ≡
19+
20+
private
21+
variable
22+
a c ℓ : Level
23+
A : Set a
24+
25+
------------------------------------------------------------------------
26+
-- Root implementation
27+
28+
module _ (M : RawMonoid c ℓ) where
29+
30+
open RawMonoid M
31+
32+
foldMap : (A Carrier) List A Carrier
33+
foldMap f [] = ε
34+
foldMap f (x ∷ xs) = f x ∙ foldMap f xs
35+
36+
------------------------------------------------------------------------
37+
-- Basic implementation using supplied defaults
38+
39+
foldableWithDefaults : RawFoldableWithDefaults (List {a})
40+
foldableWithDefaults = record { foldMap = λ M foldMap M }
41+
42+
------------------------------------------------------------------------
43+
-- Specialised version using optimised implementations
44+
45+
foldable : RawFoldable (List {a})
46+
foldable = record
47+
{ foldMap = λ M foldMap M
48+
; foldr = List.foldr
49+
; foldl = List.foldl
50+
; toList = id
51+
}
52+
53+
------------------------------------------------------------------------
54+
-- Properties
55+
56+
module _ (M : Monoid c ℓ) (f : A Monoid.Carrier M) where
57+
58+
open Monoid M
59+
60+
private
61+
h = foldMap rawMonoid f
62+
63+
[]-homo : h [] ≈ ε
64+
[]-homo = refl
65+
66+
++-homo : xs {ys} h (xs ++ ys) ≈ h xs ∙ h ys
67+
++-homo [] = sym (identityˡ _)
68+
++-homo (x ∷ xs) = trans (∙-congˡ (++-homo xs)) (sym (assoc _ _ _))
69+
70+
foldMap-morphism : IsMonoidHomomorphism (List.++-[]-rawMonoid A) rawMonoid h
71+
foldMap-morphism = record
72+
{ isMagmaHomomorphism = record
73+
{ isRelHomomorphism = record
74+
{ cong = reflexive ∘ ≡.cong h }
75+
; homo = λ xs _ ++-homo xs
76+
}
77+
; ε-homo = []-homo
78+
}

src/Data/Vec/Effectful/Foldable.agda

+51
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,51 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Vec is Foldable
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Data.Vec.Effectful.Foldable where
10+
11+
open import Algebra.Bundles.Raw using (RawMonoid)
12+
open import Data.Nat.Base using (ℕ)
13+
open import Data.Vec.Base using (Vec; []; _∷_; foldr′; foldl′; toList)
14+
open import Effect.Foldable using (RawFoldableWithDefaults; RawFoldable)
15+
open import Function.Base using (id)
16+
open import Level using (Level)
17+
18+
private
19+
variable
20+
a c ℓ : Level
21+
A : Set a
22+
n :
23+
24+
------------------------------------------------------------------------
25+
-- Root implementation
26+
27+
module _ (M : RawMonoid c ℓ) where
28+
29+
open RawMonoid M
30+
31+
foldMap : (A Carrier) Vec A n Carrier
32+
foldMap f [] = ε
33+
foldMap f (x ∷ xs) = f x ∙ foldMap f xs
34+
35+
------------------------------------------------------------------------
36+
-- Basic implementation using supplied defaults
37+
38+
foldableWithDefaults : RawFoldableWithDefaults {a} (λ A Vec A n)
39+
foldableWithDefaults = record { foldMap = λ M foldMap M }
40+
41+
------------------------------------------------------------------------
42+
-- Specialised version using optimised implementations
43+
44+
foldable : RawFoldable {a} (λ A Vec A n)
45+
foldable = record
46+
{ foldMap = λ M foldMap M
47+
; foldr = foldr′
48+
; foldl = foldl′
49+
; toList = toList
50+
}
51+

src/Effect/Foldable.agda

+77
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,77 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Foldable functors
5+
------------------------------------------------------------------------
6+
7+
-- Note that currently the Foldable laws are not included here.
8+
9+
{-# OPTIONS --cubical-compatible --safe #-}
10+
11+
module Effect.Foldable where
12+
13+
open import Algebra.Bundles.Raw using (RawMonoid)
14+
open import Algebra.Bundles using (Monoid)
15+
import Algebra.Construct.Flip.Op as Op
16+
17+
open import Data.List.Base as List using (List; [_]; _++_)
18+
19+
open import Effect.Functor as Fun using (RawFunctor)
20+
21+
open import Function.Base using (id; flip)
22+
open import Function.Endo.Propositional using (∘-id-monoid)
23+
open import Level using (Level; Setω)
24+
open import Relation.Binary.Bundles using (Setoid)
25+
26+
private
27+
variable
28+
f g c ℓ : Level
29+
A : Set f
30+
C : Set c
31+
32+
33+
------------------------------------------------------------------------
34+
-- The type of raw Foldables:
35+
-- all fields can be given non-default implementations
36+
37+
record RawFoldable (F : Set f Set g) : Setω where
38+
field
39+
foldMap : (M : RawMonoid c ℓ) (open RawMonoid M)
40+
(A Carrier) F A Carrier
41+
42+
fold : (M : RawMonoid f ℓ) (open RawMonoid M) F Carrier Carrier
43+
fold M = foldMap M id
44+
45+
field
46+
foldr : (A -> C -> C) -> C -> F A -> C
47+
foldl : (C -> A -> C) -> C -> F A -> C
48+
toList : F A List A
49+
50+
51+
------------------------------------------------------------------------
52+
-- The type of raw Foldables, with default implementations a la haskell
53+
54+
record RawFoldableWithDefaults (F : Set f Set g) : Setω where
55+
field
56+
foldMap : (M : RawMonoid c ℓ) (open RawMonoid M)
57+
(A Carrier) F A Carrier
58+
59+
foldr : (A -> C -> C) -> C -> F A -> C
60+
foldr {C = C} f z t = foldMap M₀ f t z
61+
where M = ∘-id-monoid C ; M₀ = Monoid.rawMonoid M
62+
63+
foldl : (C -> A -> C) -> C -> F A -> C
64+
foldl {C = C} f z t = foldMap M₀ (flip f) t z
65+
where M = Op.monoid (∘-id-monoid C) ; M₀ = Monoid.rawMonoid M
66+
67+
toList : F A List A
68+
toList {A = A} = foldMap (List.++-[]-rawMonoid A) [_]
69+
70+
rawFoldable : RawFoldable F
71+
rawFoldable = record
72+
{ foldMap = foldMap
73+
; foldr = foldr
74+
; foldl = foldl
75+
; toList = toList
76+
}
77+

0 commit comments

Comments
 (0)