src/HOL/Tools/Mirabelle/mirabelle_util.ML
author desharna
Tue, 18 Jan 2022 17:55:20 +0100
changeset 74986 fc664e4fbf6d
child 76183 8089593a364a
permissions -rw-r--r--
added Mirabelle option -r to randomize the goals before selection
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
74986
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
     1
(*  Title:      HOL/Mirabelle/Tools/mirabelle_util.ML
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
     2
    Author:     Martin Desharnais, MPI-INF Saarbruecken
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
     3
*)
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
     4
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
     5
(* Pseudorandom number generator *)
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
     6
signature PRNG = sig
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
     7
  type state
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
     8
  val initialize : int -> state
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
     9
  val next : state -> int * state
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    10
end
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    11
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    12
(* Pseudorandom algorithms *)
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    13
signature PRNG_ALGORITHMS = sig
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    14
  include PRNG
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    15
  val shuffle : state -> 'a list -> 'a list * state
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    16
end
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    17
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    18
functor PRNG_Algorithms(PRNG : PRNG) : PRNG_ALGORITHMS = struct
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    19
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    20
open PRNG
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    21
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    22
fun shuffle prng_state xs =
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    23
  fold_map (fn x => fn prng_state =>
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    24
    let
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    25
      val (n, prng_state') = next prng_state
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    26
    in ((n, x), prng_state') end) xs prng_state
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    27
  |> apfst (sort (int_ord o apply2 fst))
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    28
  |> apfst (map snd)
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    29
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    30
end
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    31
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    32
(* multiplicative linear congruential generator *)
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    33
structure MLCG_PRNG : PRNG = struct
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    34
  (* The modulus is implicitly 2^64 through using Word64.
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    35
     The multiplier and increment are the same as Newlib and Musl according to Wikipedia.
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    36
     See: https://en.wikipedia.org/wiki/Linear_congruential_generator#Parameters_in_common_use
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    37
   *)
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    38
  val multiplier = Word64.fromInt 6364136223846793005
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    39
  val increment = Word64.fromInt 1
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    40
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    41
  type state = Word64.word
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    42
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    43
  val initialize = Word64.fromInt
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    44
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    45
  fun next s =
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    46
    let
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    47
      open Word64
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    48
      val s' = multiplier * s + increment
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    49
    in
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    50
      (toInt s', s')
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    51
    end
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    52
end
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    53
fc664e4fbf6d added Mirabelle option -r to randomize the goals before selection
desharna
parents:
diff changeset
    54
structure MLCG = PRNG_Algorithms(MLCG_PRNG)