@@ -16,13 +16,14 @@ open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
16
16
import Agda.Builtin.Unit as Unit0
17
17
open import Data.Unit.Polymorphic.Base
18
18
open import Function.Base using (_∘′_; const; flip)
19
- import IO.Primitive as Prim
19
+ import IO.Primitive.Core as Prim
20
20
21
21
private
22
22
variable
23
- a b : Level
23
+ a b e : Level
24
24
A : Set a
25
25
B : Set b
26
+ E : Set e
26
27
27
28
------------------------------------------------------------------------
28
29
-- The IO monad
@@ -106,18 +107,37 @@ lift′ io = lift (io Prim.>>= λ _ → Prim.pure _)
106
107
ignore : IO A → IO ⊤
107
108
ignore io = io >> pure _
108
109
110
+
111
+ ------------------------------------------------------------------------
109
112
-- Conditional executions
113
+
114
+ -- Only run the action if the boolean is true
110
115
when : Bool → IO {a} ⊤ → IO ⊤
111
116
when true m = m
112
117
when false _ = pure _
113
118
119
+ -- Only run the action if the boolean is false
114
120
unless : Bool → IO {a} ⊤ → IO ⊤
115
121
unless = when ∘′ not
116
122
123
+ -- Run the action if the `Maybe` computation was successful
117
124
whenJust : Maybe A → (A → IO {a} ⊤) → IO ⊤
118
125
whenJust (just a) k = k a
119
126
whenJust nothing _ = pure _
120
127
128
+ -- Run the action if the `E ⊎_` computation was successful
129
+ whenInj₂ : E ⊎ A → (A → IO {a} ⊤) → IO ⊤
130
+ whenInj₂ (inj₂ a) k = k a
131
+ whenInj₂ (inj₁ _) _ = pure _
132
+
133
+
134
+ ------------------------------------------------------------------------
135
+ -- Loops
136
+
137
+ -- Keep running the action forever
138
+ forever : IO {a} ⊤ → IO {a} ⊤
139
+ forever act = seq (♯ act) (♯ forever act)
140
+
121
141
-- Keep running an IO action until we get a value. Convenient when user
122
142
-- input is involved and it may be malformed.
123
143
untilJust : IO (Maybe A) → IO A
@@ -127,7 +147,22 @@ untilJust m = bind (♯ m) λ where
127
147
nothing → ♯ untilJust m
128
148
(just a) → ♯ pure a
129
149
130
- untilRight : {A B : Set a} → (A → IO (A ⊎ B)) → A → IO B
131
- untilRight f x = bind (♯ f x) λ where
132
- (inj₁ x′) → ♯ untilRight f x′
150
+ untilInj₂ : {A B : Set a} → (A → IO (A ⊎ B)) → A → IO B
151
+ untilInj₂ f x = bind (♯ f x) λ where
152
+ (inj₁ x′) → ♯ untilInj₂ f x′
133
153
(inj₂ y) → ♯ pure y
154
+
155
+
156
+
157
+
158
+
159
+
160
+
161
+ ------------------------------------------------------------------------
162
+ -- DEPRECATIONS
163
+
164
+ untilRight = untilInj₂
165
+ {-# WARNING_ON_USAGE untilRight
166
+ "Warning: untilRight was deprecated in v2.1.
167
+ Please use untilInj₂ instead."
168
+ #-}
0 commit comments