src/Tools/Spec_Check/random.ML
author wenzelm
Sat, 06 May 2017 21:28:41 +0200
changeset 65750 7f5556f4b584
parent 53164 beb4ee344c22
permissions -rw-r--r--
added option verbose for SQL source;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53164
beb4ee344c22 clarified position of Spec_Check for Isabelle/ML -- it is unrelated to Isabelle/HOL;
wenzelm
parents: 52248
diff changeset
     1
(*  Title:      Tools/Spec_Check/random.ML
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     2
    Author:     Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     3
    Author:     Christopher League
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     4
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     5
Random number engine.
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     6
*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     7
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     8
signature RANDOM =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     9
sig
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    10
  type rand
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    11
  val new : unit -> rand
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    12
  val range : int * int -> rand -> int * rand
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    13
  val split : rand -> rand * rand
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    14
end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    15
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    16
structure Random : RANDOM  =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    17
struct
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    18
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    19
type rand = real
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    20
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    21
val a = 16807.0
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    22
val m = 2147483647.0
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    23
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    24
fun nextrand seed =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    25
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    26
    val t = a * seed
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    27
  in
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    28
    t - m * real(floor(t/m))
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    29
  end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    30
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    31
val new = nextrand o Time.toReal o Time.now
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    32
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    33
fun range (min, max) =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    34
  if min > max then raise Domain (* TODO: raise its own error *)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    35
  else fn r => (min + (floor(real(max-min+1) * r / m)), nextrand r)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    36
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    37
fun split r =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    38
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    39
    val r = r / a
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    40
    val r0 = real(floor r)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    41
    val r1 = r - r0
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    42
  in
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    43
    (nextrand r0, nextrand r1)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    44
  end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    45
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    46
end