src/HOL/Random.thy
author wenzelm
Thu, 08 Aug 2019 12:11:40 +0200
changeset 70490 c42a0a0a9a8d
parent 68386 98cf1c823c48
child 72581 de581f98a3a1
permissions -rw-r--r--
prefer named lemmas -- more compact proofterms;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29815
9e94b7078fa5 mandatory prefix for index conversion operations
haftmann
parents: 29806
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
     2
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59058
diff changeset
     3
section \<open>A HOL random engine\<close>
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
     4
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
     5
theory Random
58101
e7ebe5554281 separated listsum material
haftmann
parents: 57242
diff changeset
     6
imports List Groups_List
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
     7
begin
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
     8
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 36538
diff changeset
     9
notation fcomp (infixl "\<circ>>" 60)
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 36538
diff changeset
    10
notation scomp (infixl "\<circ>\<rightarrow>" 60)
29823
0ab754d13ccd session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
haftmann
parents: 29815
diff changeset
    11
0ab754d13ccd session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
haftmann
parents: 29815
diff changeset
    12
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59058
diff changeset
    13
subsection \<open>Auxiliary functions\<close>
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    14
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    15
fun log :: "natural \<Rightarrow> natural \<Rightarrow> natural" where
33236
haftmann
parents: 32740
diff changeset
    16
  "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
haftmann
parents: 32740
diff changeset
    17
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    18
definition inc_shift :: "natural \<Rightarrow> natural \<Rightarrow> natural" where
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    19
  "inc_shift v k = (if v = k then 1 else k + 1)"
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    20
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    21
definition minus_shift :: "natural \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> natural" where
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    22
  "minus_shift r k l = (if k < l then r + k - l else k - l)"
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    23
30495
haftmann
parents: 29823
diff changeset
    24
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59058
diff changeset
    25
subsection \<open>Random seeds\<close>
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
    26
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    27
type_synonym seed = "natural \<times> natural"
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
    28
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    29
primrec "next" :: "seed \<Rightarrow> natural \<times> seed" where
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    30
  "next (v, w) = (let
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    31
     k =  v div 53668;
33236
haftmann
parents: 32740
diff changeset
    32
     v' = minus_shift 2147483563 ((v mod 53668) * 40014) (k * 12211);
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    33
     l =  w div 52774;
33236
haftmann
parents: 32740
diff changeset
    34
     w' = minus_shift 2147483399 ((w mod 52774) * 40692) (l * 3791);
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    35
     z =  minus_shift 2147483562 v' (w' + 1) + 1
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    36
   in (z, (v', w')))"
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    37
29823
0ab754d13ccd session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
haftmann
parents: 29815
diff changeset
    38
definition split_seed :: "seed \<Rightarrow> seed \<times> seed" where
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
    39
  "split_seed s = (let
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
    40
     (v, w) = s;
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
    41
     (v', w') = snd (next s);
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    42
     v'' = inc_shift 2147483562 v;
33236
haftmann
parents: 32740
diff changeset
    43
     w'' = inc_shift 2147483398 w
haftmann
parents: 32740
diff changeset
    44
   in ((v'', w'), (v', w'')))"
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
    45
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
    46
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59058
diff changeset
    47
subsection \<open>Base selectors\<close>
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
    48
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    49
fun iterate :: "natural \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 36538
diff changeset
    50
  "iterate k f x = (if k = 0 then Pair x else f x \<circ>\<rightarrow> iterate (k - 1) f)"
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
    51
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    52
definition range :: "natural \<Rightarrow> seed \<Rightarrow> natural \<times> seed" where
30495
haftmann
parents: 29823
diff changeset
    53
  "range k = iterate (log 2147483561 k)
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 36538
diff changeset
    54
      (\<lambda>l. next \<circ>\<rightarrow> (\<lambda>v. Pair (v + l * 2147483561))) 1
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 36538
diff changeset
    55
    \<circ>\<rightarrow> (\<lambda>v. Pair (v mod k))"
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    56
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    57
lemma range:
30495
haftmann
parents: 29823
diff changeset
    58
  "k > 0 \<Longrightarrow> fst (range k s) < k"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    59
  by (simp add: range_def split_def less_natural_def del: log.simps iterate.simps)
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
    60
29823
0ab754d13ccd session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
haftmann
parents: 29815
diff changeset
    61
definition select :: "'a list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    62
  "select xs = range (natural_of_nat (length xs))
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    63
    \<circ>\<rightarrow> (\<lambda>k. Pair (nth xs (nat_of_natural k)))"
29823
0ab754d13ccd session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
haftmann
parents: 29815
diff changeset
    64
     
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    65
lemma select:
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    66
  assumes "xs \<noteq> []"
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    67
  shows "fst (select xs s) \<in> set xs"
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    68
proof -
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    69
  from assms have "natural_of_nat (length xs) > 0" by (simp add: less_natural_def)
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    70
  with range have
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    71
    "fst (range (natural_of_nat (length xs)) s) < natural_of_nat (length xs)" by best
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    72
  then have
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    73
    "nat_of_natural (fst (range (natural_of_nat (length xs)) s)) < length xs" by (simp add: less_natural_def)
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    74
  then show ?thesis
44921
58eef4843641 tuned proofs
huffman
parents: 42163
diff changeset
    75
    by (simp add: split_beta select_def)
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
    76
qed
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
    77
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    78
primrec pick :: "(natural \<times> 'a) list \<Rightarrow> natural \<Rightarrow> 'a" where
31180
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
    79
  "pick (x # xs) i = (if i < fst x then snd x else pick xs (i - fst x))"
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
    80
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
    81
lemma pick_member:
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 62608
diff changeset
    82
  "i < sum_list (map fst xs) \<Longrightarrow> pick xs i \<in> set (map snd xs)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    83
  by (induct xs arbitrary: i) (simp_all add: less_natural_def)
31180
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
    84
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
    85
lemma pick_drop_zero:
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
    86
  "pick (filter (\<lambda>(k, _). k > 0) xs) = pick xs"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    87
  by (induct xs) (auto simp add: fun_eq_iff less_natural_def minus_natural_def)
31180
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
    88
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
    89
lemma pick_same:
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    90
  "l < length xs \<Longrightarrow> Random.pick (map (Pair 1) xs) (natural_of_nat l) = nth xs l"
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
    91
proof (induct xs arbitrary: l)
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
    92
  case Nil then show ?case by simp
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
    93
next
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    94
  case (Cons x xs) then show ?case by (cases l) (simp_all add: less_natural_def)
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
    95
qed
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
    96
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
    97
definition select_weight :: "(natural \<times> 'a) list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 62608
diff changeset
    98
  "select_weight xs = range (sum_list (map fst xs))
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 36538
diff changeset
    99
   \<circ>\<rightarrow> (\<lambda>k. Pair (pick xs k))"
31180
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   100
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   101
lemma select_weight_member:
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 62608
diff changeset
   102
  assumes "0 < sum_list (map fst xs)"
31180
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   103
  shows "fst (select_weight xs s) \<in> set (map snd xs)"
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   104
proof -
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   105
  from range assms
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 62608
diff changeset
   106
    have "fst (range (sum_list (map fst xs)) s) < sum_list (map fst xs)" .
31180
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   107
  with pick_member
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 62608
diff changeset
   108
    have "pick xs (fst (range (sum_list (map fst xs)) s)) \<in> set (map snd xs)" .
31180
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   109
  then show ?thesis by (simp add: select_weight_def scomp_def split_def) 
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   110
qed
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   111
31268
3ced22320ceb added lemma select_weight_cons_zero
haftmann
parents: 31261
diff changeset
   112
lemma select_weight_cons_zero:
3ced22320ceb added lemma select_weight_cons_zero
haftmann
parents: 31261
diff changeset
   113
  "select_weight ((0, x) # xs) = select_weight xs"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
   114
  by (simp add: select_weight_def less_natural_def)
31268
3ced22320ceb added lemma select_weight_cons_zero
haftmann
parents: 31261
diff changeset
   115
46311
56fae81902ce random instance for sets
bulwahn
parents: 44921
diff changeset
   116
lemma select_weight_drop_zero:
31261
900ebbc35e30 dropped superfluos prefixes
haftmann
parents: 31205
diff changeset
   117
  "select_weight (filter (\<lambda>(k, _). k > 0) xs) = select_weight xs"
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   118
proof -
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
   119
  have "sum_list (map fst [(k, _)\<leftarrow>xs . 0 < k]) = sum_list (map fst xs)"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 61799
diff changeset
   120
    by (induct xs) (auto simp add: less_natural_def natural_eq_iff)
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   121
  then show ?thesis by (simp only: select_weight_def pick_drop_zero)
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   122
qed
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   123
46311
56fae81902ce random instance for sets
bulwahn
parents: 44921
diff changeset
   124
lemma select_weight_select:
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   125
  assumes "xs \<noteq> []"
31261
900ebbc35e30 dropped superfluos prefixes
haftmann
parents: 31205
diff changeset
   126
  shows "select_weight (map (Pair 1) xs) = select xs"
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   127
proof -
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
   128
  have less: "\<And>s. fst (range (natural_of_nat (length xs)) s) < natural_of_nat (length xs)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
   129
    using assms by (intro range) (simp add: less_natural_def)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 62608
diff changeset
   130
  moreover have "sum_list (map fst (map (Pair 1) xs)) = natural_of_nat (length xs)"
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   131
    by (induct xs) simp_all
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   132
  ultimately show ?thesis
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   133
    by (auto simp add: select_weight_def select_def scomp_def split_def
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
   134
      fun_eq_iff pick_same [symmetric] less_natural_def)
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   135
qed
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31196
diff changeset
   136
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
   137
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 60758
diff changeset
   138
subsection \<open>\<open>ML\<close> interface\<close>
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   139
36538
4fe16d49283b make random engine persistent using code_reflect
haftmann
parents: 36176
diff changeset
   140
code_reflect Random_Engine
4fe16d49283b make random engine persistent using code_reflect
haftmann
parents: 36176
diff changeset
   141
  functions range select select_weight
4fe16d49283b make random engine persistent using code_reflect
haftmann
parents: 36176
diff changeset
   142
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59058
diff changeset
   143
ML \<open>
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
   144
structure Random_Engine =
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   145
struct
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   146
36538
4fe16d49283b make random engine persistent using code_reflect
haftmann
parents: 36176
diff changeset
   147
open Random_Engine;
4fe16d49283b make random engine persistent using code_reflect
haftmann
parents: 36176
diff changeset
   148
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 46311
diff changeset
   149
type seed = Code_Numeral.natural * Code_Numeral.natural;
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   150
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   151
local
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   152
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31636
diff changeset
   153
val seed = Unsynchronized.ref 
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
   154
  (let
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
   155
    val now = Time.toMilliseconds (Time.now ());
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   156
    val (q, s1) = IntInf.divMod (now, 2147483562);
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   157
    val s2 = q mod 2147483398;
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58889
diff changeset
   158
  in apply2 Code_Numeral.natural_of_integer (s1 + 1, s2 + 1) end);
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
   159
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   160
in
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   161
36020
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 35266
diff changeset
   162
fun next_seed () =
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 35266
diff changeset
   163
  let
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 35266
diff changeset
   164
    val (seed1, seed') = @{code split_seed} (! seed)
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 35266
diff changeset
   165
    val _ = seed := seed'
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 35266
diff changeset
   166
  in
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 35266
diff changeset
   167
    seed1
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 35266
diff changeset
   168
  end
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 35266
diff changeset
   169
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   170
fun run f =
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   171
  let
26265
4b63b9e9b10d separated Random.thy from Quickcheck.thy
haftmann
parents: 26261
diff changeset
   172
    val (x, seed') = f (! seed);
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   173
    val _ = seed := seed'
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   174
  in x end;
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   175
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   176
end;
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   177
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   178
end;
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59058
diff changeset
   179
\<close>
22528
8501c4a62a3c cleaned up HOL/ex/Code*.thy
haftmann
parents:
diff changeset
   180
36176
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 36020
diff changeset
   181
hide_type (open) seed
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 36020
diff changeset
   182
hide_const (open) inc_shift minus_shift log "next" split_seed
31636
138625ae4067 where there is nothing, nothing can be hidden
haftmann
parents: 31633
diff changeset
   183
  iterate range select pick select_weight
36176
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 36020
diff changeset
   184
hide_fact (open) range_def
31180
dae7be64d614 hide names in theory Random
haftmann
parents: 30500
diff changeset
   185
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 36538
diff changeset
   186
no_notation fcomp (infixl "\<circ>>" 60)
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 36538
diff changeset
   187
no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
29823
0ab754d13ccd session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
haftmann
parents: 29815
diff changeset
   188
26038
55a02586776d towards quickcheck
haftmann
parents: 24994
diff changeset
   189
end