Skip to content

Commit 57f4365

Browse files
committed
[ fix #377 ] Add module Haskell.Data.Maybe with fromMaybe and more
1 parent b69790e commit 57f4365

File tree

8 files changed

+56
-8
lines changed

8 files changed

+56
-8
lines changed

lib/Haskell/Data/Maybe.agda

Lines changed: 34 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,34 @@
1+
module Haskell.Data.Maybe where
2+
3+
open import Haskell.Prelude
4+
5+
isJust : Maybe a Bool
6+
isJust Nothing = False
7+
isJust (Just _) = True
8+
9+
isNothing : Maybe a Bool
10+
isNothing Nothing = True
11+
isNothing (Just _) = False
12+
13+
fromJust : (x : Maybe a) @0 {IsJust x} a
14+
fromJust Nothing = error "fromJust Nothing"
15+
fromJust (Just x) = x
16+
17+
fromMaybe : {a : Set} a Maybe a a
18+
fromMaybe d Nothing = d
19+
fromMaybe _ (Just x) = x
20+
21+
listToMaybe : List a Maybe a
22+
listToMaybe [] = Nothing
23+
listToMaybe (x ∷ _) = Just x
24+
25+
maybeToList : Maybe a List a
26+
maybeToList Nothing = []
27+
maybeToList (Just x) = x ∷ []
28+
29+
mapMaybe : (a Maybe b) List a List b
30+
mapMaybe f [] = []
31+
mapMaybe f (x ∷ xs) = maybe id _∷_ (f x) (mapMaybe f xs)
32+
33+
catMaybes : List (Maybe a) List a
34+
catMaybes = mapMaybe id

lib/Haskell/Extra/Delay.agda

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -5,8 +5,10 @@ module Haskell.Extra.Delay where
55
open import Agda.Builtin.Size public
66

77
open import Haskell.Prelude
8-
open import Haskell.Prim.Thunk
8+
9+
open import Haskell.Data.Maybe
910
open import Haskell.Extra.Refinement
11+
open import Haskell.Prim.Thunk
1012

1113
private variable
1214
x y z : a

lib/Haskell/Prelude.agda

Lines changed: 0 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -137,6 +137,3 @@ IsJust : Maybe a → Set
137137
IsJust Nothing =
138138
IsJust (Just _) =
139139

140-
fromJust : (x : Maybe a) @0 {IsJust x} a
141-
fromJust Nothing = error "fromJust Nothing"
142-
fromJust (Just x) = x

lib/Haskell/Prim/Maybe.agda

Lines changed: 0 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,3 @@ data Maybe {@0 ℓ} (a : Set ℓ) : Set ℓ where
1111
maybe : {@0 ℓ₁ ℓ₂} {@0 a : Set ℓ₁} {@0 b : Set ℓ₂} b (a b) Maybe a b
1212
maybe n j Nothing = n
1313
maybe n j (Just x) = j x
14-
15-
fromMaybe : {a : Set} a Maybe a a
16-
fromMaybe d Nothing = d
17-
fromMaybe _ (Just x) = x

test/AllTests.agda

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -94,6 +94,7 @@ import FunCon
9494
import Issue308
9595
import Issue324
9696
import Assert
97+
import Issue377
9798

9899
{-# FOREIGN AGDA2HS
99100
import Issue14
@@ -185,4 +186,5 @@ import FunCon
185186
import Issue308
186187
import Issue324
187188
import Assert
189+
import Issue377
188190
#-}

test/Issue377.agda

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,9 @@
1+
module Issue377 where
2+
3+
open import Haskell.Prelude
4+
open import Haskell.Data.Maybe
5+
6+
test : Integer
7+
test = fromMaybe 0 (Just 12)
8+
9+
{-# COMPILE AGDA2HS test #-}

test/golden/AllTests.hs

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -89,4 +89,5 @@ import FunCon
8989
import Issue308
9090
import Issue324
9191
import Assert
92+
import Issue377
9293

test/golden/Issue377.hs

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,7 @@
1+
module Issue377 where
2+
3+
import Data.Maybe (fromMaybe)
4+
5+
test :: Integer
6+
test = fromMaybe 0 (Just 12)
7+

0 commit comments

Comments
 (0)