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
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
randomIO : HasIO io => io a
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-19 02:04:25 +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.
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
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"
|
2021-10-19 02:04:25 +03:00
|
|
|
"javascript:lambda:(max)=>Math.floor(Math.random() * max)"
|
|
|
|
prim__randomBits32 : Bits32 -> PrimIO Bits32
|
2020-08-24 19:38:29 +03:00
|
|
|
|
2021-10-19 02:04:25 +03:00
|
|
|
randomBits32 : Bits32 -> IO Bits32
|
|
|
|
randomBits32 upperBound = fromPrim (prim__randomBits32 upperBound)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
public export
|
2021-10-19 02:04:25 +03:00
|
|
|
Random Int32 where
|
2020-05-18 15:59:07 +03:00
|
|
|
-- Generate a random value within [-2^31, 2^31-1].
|
2021-10-19 02:04:25 +03:00
|
|
|
randomIO = do
|
|
|
|
let maxInt : Bits32 = 2147483647 -- shiftL 1 31 - 1
|
|
|
|
negMinInt : Bits32 = 2147483648 -- negate $ shiftL 1 31
|
|
|
|
magnitude : Bits32 = maxInt + negMinInt
|
|
|
|
bits32 <- liftIO $ randomBits32 magnitude
|
|
|
|
let int : Integer = cast bits32
|
|
|
|
pure . cast $ int - (cast negMinInt)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
-- Generate a random value within [lo, hi].
|
|
|
|
randomRIO (lo, hi) =
|
2021-10-19 02:04:25 +03:00
|
|
|
let range : Integer = (cast hi) - (cast lo) + 1
|
|
|
|
in pure . cast $ !(liftIO . randomBits32 $ cast range) + cast lo
|
2020-08-24 19:38:29 +03:00
|
|
|
|
|
|
|
%foreign "scheme:blodwen-random"
|
2021-10-19 02:04:25 +03:00
|
|
|
"javascript:lambda:()=>Math.random()"
|
2020-08-24 19:38:29 +03:00
|
|
|
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
|
2021-10-19 02:04:25 +03:00
|
|
|
rndFin : HasIO io => (n : Nat) -> io (Fin (S n))
|
2020-05-18 15:59:07 +03:00
|
|
|
rndFin 0 = pure FZ
|
|
|
|
rndFin (S k) = do
|
2021-10-19 02:04:25 +03:00
|
|
|
let intBound = the Int32 (cast (S k))
|
2020-05-18 15:59:07 +03:00
|
|
|
randomInt <- randomRIO (0, intBound)
|
|
|
|
pure $ restrict (S k) (cast randomInt)
|
|
|
|
|
|
|
|
||| Select a random element from a vector
|
|
|
|
public export
|
2021-10-19 02:04:25 +03:00
|
|
|
rndSelect' : HasIO io => {k : Nat} -> Vect (S k) a -> io a
|
2021-02-16 14:56:43 +03:00
|
|
|
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
|
2021-10-19 02:04:25 +03:00
|
|
|
rndSelect : HasIO io => (elems : List a) -> (0 _ : NonEmpty elems) => io a
|
2021-02-16 14:56:43 +03:00
|
|
|
rndSelect (x :: xs) = rndSelect' $ fromList (x :: xs)
|
2021-10-19 02:04:25 +03:00
|
|
|
|