Idris2/libs/contrib/System/Random.idr

78 lines
2.1 KiB
Idris
Raw Normal View History

2020-05-18 15:59:07 +03:00
module System.Random
import Data.Fin
import Data.Vect
import Data.List
public export
interface Random a where
randomIO : HasIO io => io a
2020-05-18 15:59:07 +03:00
-- Takes a range (lo, hi), and returns a random value uniformly
-- distributed in the closed interval [lo, hi]. It is unspecified what
-- happens if lo > hi.
randomRIO : HasIO io => (a, a) -> io a
2020-05-18 15:59:07 +03:00
2020-08-24 19:38:29 +03:00
%foreign "scheme:blodwen-random"
prim__randomInt : Int -> PrimIO Int
randomInt : Int -> IO Int
randomInt upperBound = fromPrim (prim__randomInt upperBound)
2020-05-18 15:59:07 +03:00
public export
Random Int where
-- Generate a random value within [-2^31, 2^31-1].
randomIO =
2021-03-04 23:59:56 +03:00
let maxInt = 2147483647 --shiftL 1 31 - 1
minInt = -2147483648 -- negate $ shiftL 1 31
2020-05-18 15:59:07 +03:00
range = maxInt - minInt
2020-08-24 19:38:29 +03:00
in map (+ minInt) $ liftIO $ randomInt range
2020-05-18 15:59:07 +03:00
-- Generate a random value within [lo, hi].
randomRIO (lo, hi) =
let range = hi - lo + 1
2020-08-24 19:38:29 +03:00
in map (+ lo) $ liftIO $ randomInt range
%foreign "scheme:blodwen-random"
prim__randomDouble : PrimIO Double
2020-05-18 15:59:07 +03:00
2020-08-24 19:38:29 +03:00
randomDouble : IO Double
randomDouble = fromPrim prim__randomDouble
2020-05-18 15:59:07 +03:00
public export
Random Double where
-- Generate a random value within [0, 1].
2020-08-24 19:38:29 +03:00
randomIO = liftIO randomDouble
2020-05-18 15:59:07 +03:00
-- Generate a random value within [lo, hi].
2020-08-24 19:38:29 +03:00
randomRIO (lo, hi) = map ((+ lo) . (* (hi - lo))) (liftIO randomDouble)
2020-08-25 14:30:57 +03:00
%foreign "scheme:blodwen-random-seed"
2020-08-24 19:38:29 +03:00
prim__srand : Bits64 -> PrimIO ()
2020-05-18 15:59:07 +03:00
||| Sets the random seed
export
2020-08-24 19:38:29 +03:00
srand : Bits64 -> IO ()
srand n = fromPrim (prim__srand n)
2020-05-18 15:59:07 +03:00
||| Generate a random number in Fin (S `k`)
|||
||| Note that rndFin k takes values 0, 1, ..., k.
public export
rndFin : (n : Nat) -> IO (Fin (S n))
rndFin 0 = pure FZ
rndFin (S k) = do
let intBound = the Int (cast (S k))
randomInt <- randomRIO (0, intBound)
pure $ restrict (S k) (cast randomInt)
||| Select a random element from a vector
public export
rndSelect' : {k : Nat} -> Vect (S k) a -> IO a
rndSelect' xs = pure $ Vect.index !(rndFin k) xs
2020-05-18 15:59:07 +03:00
||| Select a random element from a non-empty list
public export
rndSelect : (elems : List a) -> (0 _ : NonEmpty elems) => IO a
rndSelect (x :: xs) = rndSelect' $ fromList (x :: xs)