Idris2/tests/linear/system_concurrency_session/Main.idr
G. Allais bcf8598f99
[ new ] System.Concurrency.(Linear/Session) (#3294)
* [ refactor ] moving Data.OpenUnion to base

* [ new ] System.Concurrency.(Linear/Session)

* [ test ] for the new feature

Fixing other tests impacted by the refactoring

* [ cleanup ] move definitions around, touch up docs

* [ fix ] re-export linear notations for Control.Linear.LIO
2024-06-05 13:53:30 +01:00

29 lines
1003 B
Idris

import Control.Linear.LIO
import System.Concurrency.Session
main : IO ()
main = LIO.run $ do
let nm1 : String := "natrando"
let nm2 : String := "computer"
putStrLn "main: Forking two threads \{nm1} and \{nm2}"
res <- fork (Send Nat $ Send Nat $ Recv Nat End)
(\ ch => do let m : Nat := 100
putStrLn "\{nm1}: picked the natural \{show m}"
ch <- send ch m
let n : Nat := 50
putStrLn "\{nm1}: picked the natural \{show n}"
ch <- send ch n
(s # ch) <- recv ch
end ch
pure (m, n))
(\ ch => do (m # ch) <- recv ch
(n # ch) <- recv ch
putStrLn "\{nm2}: summing \{show m} and \{show n}"
let s = m + n
ch <- send ch s
end ch
pure s)
let mn = fst res
let mplusn = snd res
putStrLn {io = L IO} "main: Threads have finished and returned \{show mn} and \{show mplusn}"