mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-27 13:40:15 +03:00
Make Racket backend channels behave same as Chez.
Add tests for both that demonstrate equivalent behavior.
This commit is contained in:
parent
c2dd824c58
commit
795eeb23c5
@ -307,13 +307,13 @@
|
||||
;; Channels
|
||||
|
||||
(define (blodwen-make-channel ty)
|
||||
(make-channel))
|
||||
(make-async-channel 1))
|
||||
|
||||
(define (blodwen-channel-get ty chan)
|
||||
(channel-get chan))
|
||||
(async-channel-get chan))
|
||||
|
||||
(define (blodwen-channel-put ty chan val)
|
||||
(channel-put chan val))
|
||||
(async-channel-put chan val))
|
||||
|
||||
;; Mutex
|
||||
|
||||
|
@ -300,7 +300,7 @@ chezTests = MkTestPool "Chez backend" [] (Just Chez)
|
||||
, "semaphores002"
|
||||
, "perf001"
|
||||
, "reg001"
|
||||
, "channels001", "channels002", "channels003", "channels004", "channels005"
|
||||
, "channels001", "channels002", "channels003", "channels004", "channels005", "channels006"
|
||||
]
|
||||
|
||||
refcTests : IO TestPool
|
||||
@ -312,6 +312,7 @@ racketTests = MkTestPool "Racket backend" [] (Just Racket)
|
||||
, "semaphores001", "semaphores002"
|
||||
, "futures001"
|
||||
, "mutex001", "mutex002", "mutex003", "mutex004", "mutex005"
|
||||
, "channels006"
|
||||
, "conditions001" , "conditions002" , "conditions003" , "conditions004"
|
||||
, "conditions005"
|
||||
-- , "conditions006"
|
||||
|
39
tests/chez/channels006/Main.idr
Normal file
39
tests/chez/channels006/Main.idr
Normal file
@ -0,0 +1,39 @@
|
||||
module Main
|
||||
|
||||
import System
|
||||
import System.Concurrency
|
||||
|
||||
producer : Channel Nat -> IO ()
|
||||
producer ch =
|
||||
do send 1
|
||||
send 2
|
||||
send 3
|
||||
send 4
|
||||
where
|
||||
send : Nat -> IO ()
|
||||
send i =
|
||||
do putStrLn $ "> " ++ show i
|
||||
channelPut ch i
|
||||
|
||||
consumer : Channel Nat -> IO ()
|
||||
consumer ch =
|
||||
do recv
|
||||
recv
|
||||
recv
|
||||
recv
|
||||
where
|
||||
recv : IO ()
|
||||
recv =
|
||||
do usleep 100000
|
||||
v <- channelGet ch
|
||||
putStrLn $ "< " ++ show v
|
||||
|
||||
main : IO ()
|
||||
main =
|
||||
do ch <- makeChannel
|
||||
p <- fork $ producer ch
|
||||
c <- fork $ consumer ch
|
||||
threadWait c
|
||||
threadWait p
|
||||
putStrLn "done"
|
||||
|
9
tests/chez/channels006/expected
Normal file
9
tests/chez/channels006/expected
Normal file
@ -0,0 +1,9 @@
|
||||
> 1
|
||||
> 2
|
||||
< 1
|
||||
> 3
|
||||
< 2
|
||||
> 4
|
||||
< 3
|
||||
< 4
|
||||
done
|
3
tests/chez/channels006/run
Normal file
3
tests/chez/channels006/run
Normal file
@ -0,0 +1,3 @@
|
||||
rm -rf build
|
||||
|
||||
$1 --no-banner --no-color --console-width 0 --cg chez Main.idr --exec main
|
39
tests/racket/channels006/Main.idr
Normal file
39
tests/racket/channels006/Main.idr
Normal file
@ -0,0 +1,39 @@
|
||||
module Main
|
||||
|
||||
import System
|
||||
import System.Concurrency
|
||||
|
||||
producer : Channel Nat -> IO ()
|
||||
producer ch =
|
||||
do send 1
|
||||
send 2
|
||||
send 3
|
||||
send 4
|
||||
where
|
||||
send : Nat -> IO ()
|
||||
send i =
|
||||
do putStrLn $ "> " ++ show i
|
||||
channelPut ch i
|
||||
|
||||
consumer : Channel Nat -> IO ()
|
||||
consumer ch =
|
||||
do recv
|
||||
recv
|
||||
recv
|
||||
recv
|
||||
where
|
||||
recv : IO ()
|
||||
recv =
|
||||
do usleep 100000
|
||||
v <- channelGet ch
|
||||
putStrLn $ "< " ++ show v
|
||||
|
||||
main : IO ()
|
||||
main =
|
||||
do ch <- makeChannel
|
||||
p <- fork $ producer ch
|
||||
c <- fork $ consumer ch
|
||||
threadWait c
|
||||
threadWait p
|
||||
putStrLn "done"
|
||||
|
9
tests/racket/channels006/expected
Normal file
9
tests/racket/channels006/expected
Normal file
@ -0,0 +1,9 @@
|
||||
> 1
|
||||
> 2
|
||||
< 1
|
||||
> 3
|
||||
< 2
|
||||
> 4
|
||||
< 3
|
||||
< 4
|
||||
done
|
3
tests/racket/channels006/run
Normal file
3
tests/racket/channels006/run
Normal file
@ -0,0 +1,3 @@
|
||||
rm -rf build
|
||||
|
||||
$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main
|
Loading…
Reference in New Issue
Block a user