Skip to content

Commit c234c72

Browse files
gallaisandreasabel
authored andcommitted
[ new ] IO buffering, and loops (#2367)
* [ new ] IO conditionals & loops * [ new ] stdin, stdout, stderr, buffering, etc. * [ fix ] and test for the new IO primitives * [ fix ] compilation * [ fix ] more import fixing * [ done (?) ] with compilation fixes * [ test ] add test to runner * [ doc ] explain the loop * [ compat ] add deprecated stub * [ fix ] my silly mistake * [ doc ] list re-exports in CHANGELOG
1 parent 1442ee7 commit c234c72

22 files changed

+295
-46
lines changed

CHANGELOG.md

+34-1
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,7 @@ Non-backwards compatible changes
2626
parametrized by _raw_ bundles, and as such take a proof of transitivity.
2727
* The definitions in `Algebra.Module.Morphism.Construct.Identity` are now
2828
parametrized by _raw_ bundles, and as such take a proof of reflexivity.
29+
* The module `IO.Primitive` was moved to `IO.Primitive.Core`.
2930

3031
Other major improvements
3132
------------------------
@@ -66,6 +67,11 @@ Deprecated names
6667
*-pres-∣ ↦ Data.Nat.Divisibility.*-pres-∣
6768
```
6869

70+
* In `IO.Base`:
71+
```agda
72+
untilRight ↦ untilInj₂
73+
```
74+
6975
New modules
7076
-----------
7177

@@ -137,6 +143,12 @@ New modules
137143
Data.Container.Indexed.Relation.Binary.Equality.Setoid
138144
```
139145

146+
* New IO primitives to handle buffering
147+
```agda
148+
IO.Primitive.Handle
149+
IO.Handle
150+
```
151+
140152
* `System.Random` bindings:
141153
```agda
142154
System.Random.Primitive
@@ -418,13 +430,34 @@ Additions to existing modules
418430
between : String → String → String → String
419431
```
420432

433+
* Re-exported new types and functions in `IO`:
434+
```agda
435+
BufferMode : Set
436+
noBuffering : BufferMode
437+
lineBuffering : BufferMode
438+
blockBuffering : Maybe ℕ → BufferMode
439+
Handle : Set
440+
stdin : Handle
441+
stdout : Handle
442+
stderr : Handle
443+
hSetBuffering : Handle → BufferMode → IO ⊤
444+
hGetBuffering : Handle → IO BufferMode
445+
hFlush : Handle → IO ⊤
446+
```
447+
448+
* Added new functions in `IO.Base`:
449+
```agda
450+
whenInj₂ : E ⊎ A → (A → IO ⊤) → IO ⊤
451+
forever : IO ⊤ → IO ⊤
452+
```
453+
421454
* In `Data.Word.Base`:
422455
```agda
423456
_≤_ : Rel Word64 zero
424457
```
425458

426459
* Added new definition in `Relation.Binary.Construct.Closure.Transitive`
427-
```
460+
```agda
428461
transitive⁻ : Transitive _∼_ → TransClosure _∼_ ⇒ _∼_
429462
```
430463

GenerateEverything.hs

+3
Original file line numberDiff line numberDiff line change
@@ -56,11 +56,14 @@ unsafeModules = map modToFile
5656
, "IO"
5757
, "IO.Base"
5858
, "IO.Categorical"
59+
, "IO.Handle"
5960
, "IO.Infinite"
6061
, "IO.Instances"
6162
, "IO.Effectful"
6263
, "IO.Finite"
6364
, "IO.Primitive"
65+
, "IO.Primitive.Core"
66+
, "IO.Primitive.Handle"
6467
, "IO.Primitive.Infinite"
6568
, "IO.Primitive.Finite"
6669
, "Relation.Binary.PropositionalEquality.TrustMe"

doc/README/IO.agda

+20-4
Original file line numberDiff line numberDiff line change
@@ -13,6 +13,7 @@ open import Data.Nat.Base
1313
open import Data.Nat.Show using (show)
1414
open import Data.String.Base using (String; _++_; lines)
1515
open import Data.Unit.Polymorphic
16+
open import Function.Base using (_$_)
1617
open import IO
1718

1819
------------------------------------------------------------------------
@@ -86,10 +87,21 @@ cat fp = do
8687
let ls = lines content
8788
List.mapM′ putStrLn ls
8889

89-
open import Codata.Musical.Notation
90-
open import Codata.Musical.Colist
91-
open import Data.Bool
92-
open import Data.Unit.Polymorphic.Base
90+
------------------------------------------------------------------------
91+
-- TOP-LEVEL LOOP
92+
93+
-- If you simply want to repeat the same action over and over again, you
94+
-- can use `forever` e.g. the following defines a REPL that echos whatever
95+
-- the user types
96+
97+
echo : IO ⊤
98+
echo = do
99+
hSetBuffering stdout noBuffering
100+
forever $ do
101+
putStr "echo< "
102+
str getLine
103+
putStrLn ("echo> " ++ str)
104+
93105

94106
------------------------------------------------------------------------
95107
-- GUARDEDNESS
@@ -101,6 +113,10 @@ open import Data.Unit.Polymorphic.Base
101113
-- In this case you cannot use the convenient combinators that make `do`-notations
102114
-- and have to revert back to the underlying coinductive constructors.
103115

116+
open import Codata.Musical.Notation
117+
open import Codata.Musical.Colist using (Colist; []; _∷_)
118+
open import Data.Bool
119+
open import Data.Unit.Polymorphic.Base
104120

105121
-- Whether a colist is finite is semi decidable: just let the user wait until
106122
-- you reach the end!

src/Foreign/Haskell/Coerce.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -35,7 +35,7 @@ open import Level using (Level; _⊔_)
3535
open import Agda.Builtin.Nat
3636
open import Agda.Builtin.Int
3737

38-
import IO.Primitive as STD
38+
import IO.Primitive.Core as STD
3939
import Data.List.Base as STD
4040
import Data.List.NonEmpty.Base as STD
4141
import Data.Maybe.Base as STD

src/IO.agda

+2-1
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ open import Data.Unit.Polymorphic.Base
1414
open import Data.String.Base using (String)
1515
import Data.Unit.Base as Unit0
1616
open import Function.Base using (_∘_; flip)
17-
import IO.Primitive as Prim
17+
import IO.Primitive.Core as Prim
1818
open import Level
1919

2020
private
@@ -27,6 +27,7 @@ private
2727
-- Re-exporting the basic type and functions
2828

2929
open import IO.Base public
30+
open import IO.Handle public
3031

3132
------------------------------------------------------------------------
3233
-- Utilities

src/IO/Base.agda

+40-5
Original file line numberDiff line numberDiff line change
@@ -16,13 +16,14 @@ open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
1616
import Agda.Builtin.Unit as Unit0
1717
open import Data.Unit.Polymorphic.Base
1818
open import Function.Base using (_∘′_; const; flip)
19-
import IO.Primitive as Prim
19+
import IO.Primitive.Core as Prim
2020

2121
private
2222
variable
23-
a b : Level
23+
a b e : Level
2424
A : Set a
2525
B : Set b
26+
E : Set e
2627

2728
------------------------------------------------------------------------
2829
-- The IO monad
@@ -106,18 +107,37 @@ lift′ io = lift (io Prim.>>= λ _ → Prim.pure _)
106107
ignore : IO A IO ⊤
107108
ignore io = io >> pure _
108109

110+
111+
------------------------------------------------------------------------
109112
-- Conditional executions
113+
114+
-- Only run the action if the boolean is true
110115
when : Bool IO {a} ⊤ IO ⊤
111116
when true m = m
112117
when false _ = pure _
113118

119+
-- Only run the action if the boolean is false
114120
unless : Bool IO {a} ⊤ IO ⊤
115121
unless = when ∘′ not
116122

123+
-- Run the action if the `Maybe` computation was successful
117124
whenJust : Maybe A (A IO {a} ⊤) IO ⊤
118125
whenJust (just a) k = k a
119126
whenJust nothing _ = pure _
120127

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+
121141
-- Keep running an IO action until we get a value. Convenient when user
122142
-- input is involved and it may be malformed.
123143
untilJust : IO (Maybe A) IO A
@@ -127,7 +147,22 @@ untilJust m = bind (♯ m) λ where
127147
nothing ♯ untilJust m
128148
(just a) ♯ pure a
129149

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′
133153
(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+
#-}

src/IO/Finite.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,7 @@ open import Data.Unit.Polymorphic.Base
1212
open import Agda.Builtin.String
1313
import Data.Unit.Base as Unit0
1414
open import IO.Base
15-
import IO.Primitive as Prim
15+
import IO.Primitive.Core as Prim
1616
import IO.Primitive.Finite as Prim
1717
open import Level
1818

src/IO/Handle.agda

+41
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,41 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- IO handles: simple bindings to Haskell types and functions
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --guardedness #-}
8+
9+
module IO.Handle where
10+
11+
open import Level using (Level)
12+
open import Data.Unit.Polymorphic.Base using (⊤)
13+
open import IO.Base using (IO; lift; lift′)
14+
15+
private variable a : Level
16+
17+
------------------------------------------------------------------------
18+
-- Re-exporting types and constructors
19+
20+
open import IO.Primitive.Handle as Prim public
21+
using ( BufferMode
22+
; noBuffering
23+
; lineBuffering
24+
; blockBuffering
25+
; Handle
26+
; stdin
27+
; stdout
28+
; stderr
29+
)
30+
31+
------------------------------------------------------------------------
32+
-- Wrapping functions
33+
34+
hSetBuffering : Handle BufferMode IO {a} ⊤
35+
hSetBuffering hdl bm = lift′ (Prim.hSetBuffering hdl bm)
36+
37+
hGetBuffering : Handle IO BufferMode
38+
hGetBuffering hdl = lift (Prim.hGetBuffering hdl)
39+
40+
hFlush : Handle IO {a} ⊤
41+
hFlush hdl = lift′ (Prim.hFlush hdl)

src/IO/Infinite.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -13,7 +13,7 @@ open import Agda.Builtin.String
1313
open import Data.Unit.Polymorphic.Base
1414
import Data.Unit.Base as Unit0
1515
open import IO.Base
16-
import IO.Primitive as Prim
16+
import IO.Primitive.Core as Prim
1717
import IO.Primitive.Infinite as Prim
1818
open import Level
1919

src/IO/Primitive.agda

+5-28
Original file line numberDiff line numberDiff line change
@@ -1,38 +1,15 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Primitive IO: simple bindings to Haskell types and functions
4+
-- This module is DEPRECATED. Please use IO.Primitive.Core instead
55
------------------------------------------------------------------------
66

7-
-- NOTE: the contents of this module should be accessed via `IO`.
8-
97
{-# OPTIONS --cubical-compatible #-}
108

119
module IO.Primitive where
1210

13-
open import Level using (Level)
14-
private
15-
variable
16-
a : Level
17-
A B : Set a
18-
19-
------------------------------------------------------------------------
20-
-- The IO monad
21-
22-
open import Agda.Builtin.IO public
23-
using (IO)
24-
25-
infixl 1 _>>=_
26-
27-
postulate
28-
pure : A IO A
29-
_>>=_ : IO A (A IO B) IO B
30-
31-
{-# COMPILE GHC pure = \_ _ -> return #-}
32-
{-# COMPILE GHC _>>=_ = \_ _ _ _ -> (>>=) #-}
33-
{-# COMPILE UHC pure = \_ _ x -> UHC.Agda.Builtins.primReturn x #-}
34-
{-# COMPILE UHC _>>=_ = \_ _ _ _ x y -> UHC.Agda.Builtins.primBind x y #-}
11+
open import IO.Primitive.Core public
3512

36-
-- Haskell-style alternative syntax
37-
return : A IO A
38-
return = pure
13+
{-# WARNING_ON_IMPORT
14+
"IO.Primitive was deprecated in v2.1. Use IO.Primitive.Core instead."
15+
#-}

src/IO/Primitive/Core.agda

+38
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,38 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Primitive IO: simple bindings to Haskell types and functions
5+
------------------------------------------------------------------------
6+
7+
-- NOTE: the contents of this module should be accessed via `IO`.
8+
9+
{-# OPTIONS --cubical-compatible #-}
10+
11+
module IO.Primitive.Core where
12+
13+
open import Level using (Level)
14+
private
15+
variable
16+
a : Level
17+
A B : Set a
18+
19+
------------------------------------------------------------------------
20+
-- The IO monad
21+
22+
open import Agda.Builtin.IO public
23+
using (IO)
24+
25+
infixl 1 _>>=_
26+
27+
postulate
28+
pure : A IO A
29+
_>>=_ : IO A (A IO B) IO B
30+
31+
{-# COMPILE GHC pure = \_ _ -> return #-}
32+
{-# COMPILE GHC _>>=_ = \_ _ _ _ -> (>>=) #-}
33+
{-# COMPILE UHC pure = \_ _ x -> UHC.Agda.Builtins.primReturn x #-}
34+
{-# COMPILE UHC _>>=_ = \_ _ _ _ x y -> UHC.Agda.Builtins.primBind x y #-}
35+
36+
-- Haskell-style alternative syntax
37+
return : A IO A
38+
return = pure

0 commit comments

Comments
 (0)