Skip to content

Commit dc59724

Browse files
committed
Change Monad polymorphism
1 parent aa62ce6 commit dc59724

File tree

12 files changed

+130
-24
lines changed

12 files changed

+130
-24
lines changed

Class/Applicative/Core.agda

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@ open import Class.Prelude
55
open import Class.Core
66
open import Class.Functor.Core
77

8-
record Applicative (F : Type↑) : Typeω where
8+
record Applicative (F : Type↑ ℓ↑) : Typeω where
99
infixl 4 _<*>_ _⊛_ _<*_ _<⊛_ _*>_ _⊛>_
1010
infix 4 _⊗_
1111

@@ -34,13 +34,13 @@ record Applicative (F : Type↑) : Typeω where
3434

3535
open Applicative ⦃...⦄ public
3636

37-
record Applicative₀ (F : Type↑) : Typeω where
37+
record Applicative₀ (F : Type↑ ℓ↑) : Typeω where
3838
field
3939
overlap ⦃ super ⦄ : Applicative F
4040
ε₀ : F A
4141
open Applicative₀ ⦃...⦄ public
4242

43-
record Alternative (F : Type↑) : Typeω where
43+
record Alternative (F : Type↑ ℓ↑) : Typeω where
4444
infixr 3 _<|>_
4545
field _<|>_ : F A F A F A
4646
open Alternative ⦃...⦄ public

Class/Bifunctor.agda

Lines changed: 7 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -13,7 +13,7 @@ private variable
1313

1414
-- ** indexed/dependent version
1515
record BifunctorI
16-
(F : (A : Type a) (B : A Type b) Type (a ⊔ b)) : Type (lsuc (a ⊔ b)) where
16+
(F : {a b} (A : Type a) (B : A Type b) Type (ℓ↑² a b)) : Typeω where
1717
field
1818
bimap′ : (f : A A′) ( {x} B x C (f x)) F A B F A′ C
1919

@@ -30,11 +30,11 @@ record BifunctorI
3030
open BifunctorI ⦃...⦄ public
3131

3232
instance
33-
Bifunctor-Σ : BifunctorI {a}{b} Σ
33+
Bifunctor-Σ : BifunctorI Σ
3434
Bifunctor-Σ .bimap′ = ×.map
3535

3636
-- ** non-dependent version
37-
record Bifunctor (F : Type a Type b Type (a ⊔ b)) : Type (lsuc (a ⊔ b)) where
37+
record Bifunctor (F : Type↑² ℓ↑²) : Typeω where
3838
field
3939
bimap : {A A′ : Type a} {B B′ : Type b} (A A′) (B B′) F A B F A′ B′
4040

@@ -50,16 +50,15 @@ record Bifunctor (F : Type a → Type b → Type (a ⊔ b)) : Type (lsuc (a ⊔
5050

5151
open Bifunctor ⦃...⦄ public
5252

53-
map₁₂ : {F : Type a Type a Type a} {A B : Type a}
54-
⦃ Bifunctor F ⦄
55-
(A B) F A A F B B
53+
map₁₂ : {F : Type↑² ℓ↑²} ⦃ _ : Bifunctor F ⦄
54+
( {a} {A B : Type a} (A B) F A A F B B)
5655
map₁₂ f = bimap f f
5756
_<$>₁₂_ = map₁₂
5857
infixl 4 _<$>₁₂_
5958

6059
instance
61-
Bifunctor-× : Bifunctor {a}{b} _×_
60+
Bifunctor-× : Bifunctor _×_
6261
Bifunctor-× .bimap f g = ×.map f g
6362

64-
Bifunctor-⊎ : Bifunctor {a}{b} _⊎_
63+
Bifunctor-⊎ : Bifunctor _⊎_
6564
Bifunctor-⊎ .bimap = ⊎.map

Class/Core.agda

Lines changed: 23 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -3,13 +3,32 @@ module Class.Core where
33

44
open import Class.Prelude
55

6-
Type[_↝_] : ℓ ℓ′ Type (lsuc ℓ ⊔ lsuc ℓ′)
6+
-- ** unary type formers
7+
8+
Type[_↝_] : ℓ ℓ′ Type _
79
Type[ ℓ ↝ ℓ′ ] = Type ℓ Type ℓ′
810

9-
Type↑ : Typeω
10-
Type↑ = {ℓ} Type[ ℓ ↝ ℓ ]
11+
Level↑ = Level Level
12+
13+
variable ℓ↑ ℓ↑′ : Level↑
14+
15+
Type↑ : Level↑ Typeω
16+
Type↑ ℓ↑ = {ℓ} Type[ ℓ ↝ ℓ↑ ℓ ]
17+
18+
variable M F : Type↑ ℓ↑
19+
20+
-- ** binary type formers
21+
Type[_∣_↝_] : ℓ ℓ′ ℓ″ Type _
22+
Type[ ℓ ∣ ℓ′ ↝ ℓ″ ] = Type ℓ Type ℓ′ Type ℓ″
1123

12-
module _ (M : Type↑) where
24+
Level↑² = Level Level Level
25+
26+
Type↑² : Level↑² Typeω
27+
Type↑² ℓ↑² = {ℓ ℓ′} Type[ ℓ ∣ ℓ′ ↝ ℓ↑² ℓ ℓ′ ]
28+
29+
variable ℓ↑² ℓ↑²′ : Level Level Level
30+
31+
module _ (M : Type↑ ℓ↑) where
1332
: (A Type ℓ) Type _
1433
_¹ P = {x} M (P x)
1534

@@ -18,6 +37,3 @@ module _ (M : Type↑) where
1837

1938
: (A B C Type ℓ) Type _
2039
_³ _~_~_ = {x y z} M (x ~ y ~ z)
21-
22-
variable
23-
M F : Type↑

Class/Foldable/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -7,6 +7,6 @@ open import Class.Functor
77
open import Class.Semigroup
88
open import Class.Monoid
99

10-
record Foldable (F : Type↑) ⦃ _ : Functor F ⦄ : Typeω where
10+
record Foldable (F : Type↑ ℓ↑) ⦃ _ : Functor F ⦄ : Typeω where
1111
field fold : ⦃ _ : Semigroup A ⦄ ⦃ Monoid A ⦄ F A A
1212
open Foldable ⦃...⦄ public

Class/Functor/Core.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,7 @@ open import Class.Core
66

77
private variable a b c : Level
88

9-
record Functor (F : Type↑) : Typeω where
9+
record Functor (F : Type↑ ℓ↑) : Typeω where
1010
infixl 4 _<$>_ _<$_
1111
infixl 1 _<&>_
1212

@@ -20,7 +20,7 @@ record Functor (F : Type↑) : Typeω where
2020
_<&>_ = flip _<$>_
2121
open Functor ⦃...⦄ public
2222

23-
record FunctorLaws (F : Type↑) ⦃ _ : Functor F ⦄ : Typeω where
23+
record FunctorLaws (F : Type↑ ℓ↑) ⦃ _ : Functor F ⦄ : Typeω where
2424
field
2525
-- preserves identity morphisms
2626
fmap-id : {A : Type a} (x : F A)

Class/Monad/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,7 @@ open import Class.Core
66
open import Class.Functor
77
open import Class.Applicative
88

9-
record Monad (M : Type↑) : Typeω where
9+
record Monad (M : Type↑ ℓ↑) : Typeω where
1010
infixl 1 _>>=_ _>>_ _>=>_
1111
infixr 1 _=<<_ _<=<_
1212

Class/Monad/Instances.agda

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -19,3 +19,15 @@ instance
1919
.return just
2020
._>>=_ Maybe._>>=_
2121
where import Data.Maybe as Maybe
22+
23+
Monad-Sumˡ : Monad (_⊎ A)
24+
Monad-Sumˡ = λ where
25+
.return inj₁
26+
._>>=_ (inj₁ a) f f a
27+
._>>=_ (inj₂ b) f inj₂ b
28+
29+
Monad-Sumʳ : Monad (A ⊎_)
30+
Monad-Sumʳ = λ where
31+
.return inj₂
32+
._>>=_ (inj₁ a) f inj₁ a
33+
._>>=_ (inj₂ b) f f b

Class/Traversable/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,7 @@ open import Class.Core
66
open import Class.Functor.Core
77
open import Class.Monad
88

9-
record Traversable (F : Type↑) ⦃ _ : Functor F ⦄ : Typeω where
9+
record Traversable (F : Type↑ ℓ↑) ⦃ _ : Functor F ⦄ : Typeω where
1010
field sequence : ⦃ Monad M ⦄ F (M A) M (F A)
1111

1212
traverse : ⦃ Monad M ⦄ (A M B) F A M (F B)

Test/Functor.agda

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -30,3 +30,9 @@ _ = map₂′ id (1 , 2 ∷ []) ≡ ((∃ λ n → Vec ℕ n) ∋ (1 , 2 ∷ [])
3030
∋ refl
3131
_ = bimap′ suc (2 ∷_) (0 , []) ≡ ((∃ λ n Vec ℕ n) ∋ (1 , 2 ∷ []))
3232
∋ refl
33+
34+
-- ** cross-level mapping
35+
module _ (X : Type) (Y : Type₁) (Z : Type₂) (g : X Y) (f : Y Z) (xs : List X) where
36+
_ : fmap (f ∘ g) xs
37+
≡ (fmap f ∘ fmap g) xs
38+
_ = fmap-∘ xs

Test/Monad.agda

Lines changed: 62 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,62 @@
1+
{-# OPTIONS --cubical-compatible #-}
2+
module Test.Monad where
3+
4+
open import Data.List using (length)
5+
6+
open import Class.Core
7+
open import Class.Prelude
8+
open import Class.Functor
9+
open import Class.Monad
10+
11+
-- ** maybe monad
12+
13+
pred? : Maybe ℕ
14+
pred? = λ where
15+
0 nothing
16+
(suc n) just n
17+
18+
getPred : Maybe ℕ
19+
getPred = λ n do
20+
x pred? n
21+
return x
22+
23+
_ = getPred 3 ≡ just 2
24+
∋ refl
25+
26+
-- ** reader monad
27+
28+
ReaderT : (M : Type↑ ℓ↑) Type ℓ Type ℓ′ _
29+
ReaderT M A B = A M B
30+
31+
instance
32+
Monad-ReaderT : ⦃ _ : Monad M ⦄ Monad (ReaderT M A)
33+
Monad-ReaderT = λ where
34+
.return λ x _ return x
35+
._>>=_ m k λ a m a >>= λ b k b a
36+
37+
Reader : Type ℓ Type ℓ′ Type _
38+
Reader = ReaderT id
39+
40+
instance
41+
Monad-Id : Monad id
42+
Monad-Id = λ where
43+
.return id
44+
._>>=_ m k k m
45+
{-# INCOHERENT Monad-Id #-}
46+
47+
ask : Reader A A
48+
ask a = a
49+
50+
local : (A B) Reader B C Reader A C
51+
local f r = r ∘ f
52+
53+
runReader : A Reader A B B
54+
runReader = flip _$_
55+
56+
getLength : List (Maybe ℕ)
57+
getLength ys = runReader ys $ local (just 0 ∷_) $ do
58+
xs ask
59+
return (length xs)
60+
61+
_ = getLength (just 1 ∷ nothing ∷ just 2 ∷ []) ≡ 4
62+
∋ refl

Test/Reflection.agda

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,9 @@
1+
module Test.Reflection where
2+
3+
open import Reflection hiding (_>>=_)
4+
open import Data.Nat using (zero)
5+
open import Class.Monad
6+
7+
-- ** cross-level bind
8+
_ : TC Set
9+
_ = getType (quote zero) >>= unquoteTC

standard-library-classes.agda

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -25,11 +25,13 @@ open import Class.HasAdd public
2525
open import Class.HasOrder public
2626
open import Class.Show public
2727
open import Class.ToBool public
28-
open import Class.MonotonePredicate public
28+
open import Class.MonotonePredicate public
2929

3030
-- ** Tests
3131
open import Test.Monoid
3232
open import Test.Functor
33+
open import Test.Monad
3334
open import Test.DecEq
3435
open import Test.Decidable
3536
open import Test.Show
37+
open import Test.Reflection

0 commit comments

Comments
 (0)