File tree 8 files changed +56
-8
lines changed 8 files changed +56
-8
lines changed Original file line number Diff line number Diff line change
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
Original file line number Diff line number Diff line change @@ -5,8 +5,10 @@ module Haskell.Extra.Delay where
5
5
open import Agda.Builtin.Size public
6
6
7
7
open import Haskell.Prelude
8
- open import Haskell.Prim.Thunk
8
+
9
+ open import Haskell.Data.Maybe
9
10
open import Haskell.Extra.Refinement
11
+ open import Haskell.Prim.Thunk
10
12
11
13
private variable
12
14
x y z : a
Original file line number Diff line number Diff line change @@ -137,6 +137,3 @@ IsJust : Maybe a → Set
137
137
IsJust Nothing = ⊥
138
138
IsJust (Just _) = ⊤
139
139
140
- fromJust : (x : Maybe a) → @0 {IsJust x} → a
141
- fromJust Nothing = error "fromJust Nothing"
142
- fromJust (Just x) = x
Original file line number Diff line number Diff line change @@ -11,7 +11,3 @@ data Maybe {@0 ℓ} (a : Set ℓ) : Set ℓ where
11
11
maybe : ∀ {@0 ℓ₁ ℓ₂} {@0 a : Set ℓ₁} {@0 b : Set ℓ₂} → b → (a → b) → Maybe a → b
12
12
maybe n j Nothing = n
13
13
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
Original file line number Diff line number Diff line change @@ -94,6 +94,7 @@ import FunCon
94
94
import Issue308
95
95
import Issue324
96
96
import Assert
97
+ import Issue377
97
98
98
99
{-# FOREIGN AGDA2HS
99
100
import Issue14
@@ -185,4 +186,5 @@ import FunCon
185
186
import Issue308
186
187
import Issue324
187
188
import Assert
189
+ import Issue377
188
190
#-}
Original file line number Diff line number Diff line change
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 #-}
Original file line number Diff line number Diff line change @@ -89,4 +89,5 @@ import FunCon
89
89
import Issue308
90
90
import Issue324
91
91
import Assert
92
+ import Issue377
92
93
Original file line number Diff line number Diff line change
1
+ module Issue377 where
2
+
3
+ import Data.Maybe (fromMaybe )
4
+
5
+ test :: Integer
6
+ test = fromMaybe 0 (Just 12 )
7
+
You can’t perform that action at this time.
0 commit comments