1(* Author: Florian Haftmann, TU Muenchen *)
2
3section \<open>A HOL random engine\<close>
4
5theory Random
6imports List Groups_List
7begin
8
9notation fcomp (infixl "\<circ>>" 60)
10notation scomp (infixl "\<circ>\<rightarrow>" 60)
11
12
13subsection \<open>Auxiliary functions\<close>
14
15fun log :: "natural \<Rightarrow> natural \<Rightarrow> natural" where
16  "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
17
18definition inc_shift :: "natural \<Rightarrow> natural \<Rightarrow> natural" where
19  "inc_shift v k = (if v = k then 1 else k + 1)"
20
21definition minus_shift :: "natural \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> natural" where
22  "minus_shift r k l = (if k < l then r + k - l else k - l)"
23
24
25subsection \<open>Random seeds\<close>
26
27type_synonym seed = "natural \<times> natural"
28
29primrec "next" :: "seed \<Rightarrow> natural \<times> seed" where
30  "next (v, w) = (let
31     k =  v div 53668;
32     v' = minus_shift 2147483563 ((v mod 53668) * 40014) (k * 12211);
33     l =  w div 52774;
34     w' = minus_shift 2147483399 ((w mod 52774) * 40692) (l * 3791);
35     z =  minus_shift 2147483562 v' (w' + 1) + 1
36   in (z, (v', w')))"
37
38definition split_seed :: "seed \<Rightarrow> seed \<times> seed" where
39  "split_seed s = (let
40     (v, w) = s;
41     (v', w') = snd (next s);
42     v'' = inc_shift 2147483562 v;
43     w'' = inc_shift 2147483398 w
44   in ((v'', w'), (v', w'')))"
45
46
47subsection \<open>Base selectors\<close>
48
49fun iterate :: "natural \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
50  "iterate k f x = (if k = 0 then Pair x else f x \<circ>\<rightarrow> iterate (k - 1) f)"
51
52definition range :: "natural \<Rightarrow> seed \<Rightarrow> natural \<times> seed" where
53  "range k = iterate (log 2147483561 k)
54      (\<lambda>l. next \<circ>\<rightarrow> (\<lambda>v. Pair (v + l * 2147483561))) 1
55    \<circ>\<rightarrow> (\<lambda>v. Pair (v mod k))"
56
57lemma range:
58  "k > 0 \<Longrightarrow> fst (range k s) < k"
59  by (simp add: range_def split_def less_natural_def del: log.simps iterate.simps)
60
61definition select :: "'a list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
62  "select xs = range (natural_of_nat (length xs))
63    \<circ>\<rightarrow> (\<lambda>k. Pair (nth xs (nat_of_natural k)))"
64     
65lemma select:
66  assumes "xs \<noteq> []"
67  shows "fst (select xs s) \<in> set xs"
68proof -
69  from assms have "natural_of_nat (length xs) > 0" by (simp add: less_natural_def)
70  with range have
71    "fst (range (natural_of_nat (length xs)) s) < natural_of_nat (length xs)" by best
72  then have
73    "nat_of_natural (fst (range (natural_of_nat (length xs)) s)) < length xs" by (simp add: less_natural_def)
74  then show ?thesis
75    by (simp add: split_beta select_def)
76qed
77
78primrec pick :: "(natural \<times> 'a) list \<Rightarrow> natural \<Rightarrow> 'a" where
79  "pick (x # xs) i = (if i < fst x then snd x else pick xs (i - fst x))"
80
81lemma pick_member:
82  "i < sum_list (map fst xs) \<Longrightarrow> pick xs i \<in> set (map snd xs)"
83  by (induct xs arbitrary: i) (simp_all add: less_natural_def)
84
85lemma pick_drop_zero:
86  "pick (filter (\<lambda>(k, _). k > 0) xs) = pick xs"
87  by (induct xs) (auto simp add: fun_eq_iff less_natural_def minus_natural_def)
88
89lemma pick_same:
90  "l < length xs \<Longrightarrow> Random.pick (map (Pair 1) xs) (natural_of_nat l) = nth xs l"
91proof (induct xs arbitrary: l)
92  case Nil then show ?case by simp
93next
94  case (Cons x xs) then show ?case by (cases l) (simp_all add: less_natural_def)
95qed
96
97definition select_weight :: "(natural \<times> 'a) list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
98  "select_weight xs = range (sum_list (map fst xs))
99   \<circ>\<rightarrow> (\<lambda>k. Pair (pick xs k))"
100
101lemma select_weight_member:
102  assumes "0 < sum_list (map fst xs)"
103  shows "fst (select_weight xs s) \<in> set (map snd xs)"
104proof -
105  from range assms
106    have "fst (range (sum_list (map fst xs)) s) < sum_list (map fst xs)" .
107  with pick_member
108    have "pick xs (fst (range (sum_list (map fst xs)) s)) \<in> set (map snd xs)" .
109  then show ?thesis by (simp add: select_weight_def scomp_def split_def) 
110qed
111
112lemma select_weight_cons_zero:
113  "select_weight ((0, x) # xs) = select_weight xs"
114  by (simp add: select_weight_def less_natural_def)
115
116lemma select_weight_drop_zero:
117  "select_weight (filter (\<lambda>(k, _). k > 0) xs) = select_weight xs"
118proof -
119  have "sum_list (map fst [(k, _)\<leftarrow>xs . 0 < k]) = sum_list (map fst xs)"
120    by (induct xs) (auto simp add: less_natural_def natural_eq_iff)
121  then show ?thesis by (simp only: select_weight_def pick_drop_zero)
122qed
123
124lemma select_weight_select:
125  assumes "xs \<noteq> []"
126  shows "select_weight (map (Pair 1) xs) = select xs"
127proof -
128  have less: "\<And>s. fst (range (natural_of_nat (length xs)) s) < natural_of_nat (length xs)"
129    using assms by (intro range) (simp add: less_natural_def)
130  moreover have "sum_list (map fst (map (Pair 1) xs)) = natural_of_nat (length xs)"
131    by (induct xs) simp_all
132  ultimately show ?thesis
133    by (auto simp add: select_weight_def select_def scomp_def split_def
134      fun_eq_iff pick_same [symmetric] less_natural_def)
135qed
136
137
138subsection \<open>\<open>ML\<close> interface\<close>
139
140code_reflect Random_Engine
141  functions range select select_weight
142
143ML \<open>
144structure Random_Engine =
145struct
146
147open Random_Engine;
148
149type seed = Code_Numeral.natural * Code_Numeral.natural;
150
151local
152
153val seed = Unsynchronized.ref 
154  (let
155    val now = Time.toMilliseconds (Time.now ());
156    val (q, s1) = IntInf.divMod (now, 2147483562);
157    val s2 = q mod 2147483398;
158  in apply2 Code_Numeral.natural_of_integer (s1 + 1, s2 + 1) end);
159
160in
161
162fun next_seed () =
163  let
164    val (seed1, seed') = @{code split_seed} (! seed)
165    val _ = seed := seed'
166  in
167    seed1
168  end
169
170fun run f =
171  let
172    val (x, seed') = f (! seed);
173    val _ = seed := seed'
174  in x end;
175
176end;
177
178end;
179\<close>
180
181hide_type (open) seed
182hide_const (open) inc_shift minus_shift log "next" split_seed
183  iterate range select pick select_weight
184hide_fact (open) range_def
185
186no_notation fcomp (infixl "\<circ>>" 60)
187no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
188
189end
190