| author | wenzelm | 
| Mon, 04 Nov 2024 21:25:26 +0100 | |
| changeset 81344 | 1b9ea66810ff | 
| parent 80175 | 200107cdd3ac | 
| permissions | -rw-r--r-- | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
29806 
diff
changeset
 | 
1  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
| 22528 | 2  | 
|
| 60758 | 3  | 
section \<open>A HOL random engine\<close>  | 
| 22528 | 4  | 
|
5  | 
theory Random  | 
|
| 74101 | 6  | 
imports List Groups_List Code_Numeral  | 
| 22528 | 7  | 
begin  | 
8  | 
||
| 60758 | 9  | 
subsection \<open>Auxiliary functions\<close>  | 
| 26265 | 10  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
11  | 
fun log :: "natural \<Rightarrow> natural \<Rightarrow> natural" where  | 
| 33236 | 12  | 
"log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"  | 
13  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
14  | 
definition inc_shift :: "natural \<Rightarrow> natural \<Rightarrow> natural" where  | 
| 26265 | 15  | 
"inc_shift v k = (if v = k then 1 else k + 1)"  | 
16  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
17  | 
definition minus_shift :: "natural \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> natural" where  | 
| 26265 | 18  | 
"minus_shift r k l = (if k < l then r + k - l else k - l)"  | 
19  | 
||
| 30495 | 20  | 
|
| 60758 | 21  | 
subsection \<open>Random seeds\<close>  | 
| 26038 | 22  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
23  | 
type_synonym seed = "natural \<times> natural"  | 
| 22528 | 24  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
25  | 
primrec "next" :: "seed \<Rightarrow> natural \<times> seed" where  | 
| 26265 | 26  | 
"next (v, w) = (let  | 
27  | 
k = v div 53668;  | 
|
| 33236 | 28  | 
v' = minus_shift 2147483563 ((v mod 53668) * 40014) (k * 12211);  | 
| 26265 | 29  | 
l = w div 52774;  | 
| 33236 | 30  | 
w' = minus_shift 2147483399 ((w mod 52774) * 40692) (l * 3791);  | 
| 26265 | 31  | 
z = minus_shift 2147483562 v' (w' + 1) + 1  | 
32  | 
in (z, (v', w')))"  | 
|
33  | 
||
| 
29823
 
0ab754d13ccd
session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
 
haftmann 
parents: 
29815 
diff
changeset
 | 
34  | 
definition split_seed :: "seed \<Rightarrow> seed \<times> seed" where  | 
| 26038 | 35  | 
"split_seed s = (let  | 
36  | 
(v, w) = s;  | 
|
37  | 
(v', w') = snd (next s);  | 
|
| 26265 | 38  | 
v'' = inc_shift 2147483562 v;  | 
| 33236 | 39  | 
w'' = inc_shift 2147483398 w  | 
40  | 
in ((v'', w'), (v', w'')))"  | 
|
| 26038 | 41  | 
|
42  | 
||
| 60758 | 43  | 
subsection \<open>Base selectors\<close>  | 
| 22528 | 44  | 
|
| 72581 | 45  | 
context  | 
46  | 
includes state_combinator_syntax  | 
|
47  | 
begin  | 
|
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 | 50  | 
"iterate k f x = (if k = 0 then Pair x else f x \<circ>\<rightarrow> iterate (k - 1) f)"  | 
| 22528 | 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 | 53  | 
"range k = iterate (log 2147483561 k)  | 
| 37751 | 54  | 
(\<lambda>l. next \<circ>\<rightarrow> (\<lambda>v. Pair (v + l * 2147483561))) 1  | 
55  | 
\<circ>\<rightarrow> (\<lambda>v. Pair (v mod k))"  | 
|
| 26265 | 56  | 
|
57  | 
lemma range:  | 
|
| 30495 | 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 | 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 | 65  | 
lemma select:  | 
66  | 
assumes "xs \<noteq> []"  | 
|
67  | 
shows "fst (select xs s) \<in> set xs"  | 
|
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 | 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 | 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 | 74  | 
then show ?thesis  | 
| 44921 | 75  | 
by (simp add: split_beta select_def)  | 
| 26265 | 76  | 
qed  | 
| 22528 | 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 | 79  | 
"pick (x # xs) i = (if i < fst x then snd x else pick xs (i - fst x))"  | 
80  | 
||
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 | 84  | 
|
85  | 
lemma pick_drop_zero:  | 
|
86  | 
"pick (filter (\<lambda>(k, _). k > 0) xs) = pick xs"  | 
|
| 
80175
 
200107cdd3ac
Some new simprules – and patches for proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
74101 
diff
changeset
 | 
87  | 
apply (induct xs)  | 
| 
 
200107cdd3ac
Some new simprules – and patches for proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
74101 
diff
changeset
 | 
88  | 
apply (auto simp: fun_eq_iff less_natural.rep_eq split: prod.split)  | 
| 
 
200107cdd3ac
Some new simprules – and patches for proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
74101 
diff
changeset
 | 
89  | 
by (metis diff_zero of_nat_0 of_nat_of_natural)  | 
| 31180 | 90  | 
|
| 
31203
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
91  | 
lemma pick_same:  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
92  | 
"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
 | 
93  | 
proof (induct xs arbitrary: l)  | 
| 
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
94  | 
case Nil then show ?case by simp  | 
| 
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
95  | 
next  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
96  | 
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
 | 
97  | 
qed  | 
| 
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
98  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
99  | 
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
 | 
100  | 
"select_weight xs = range (sum_list (map fst xs))  | 
| 37751 | 101  | 
\<circ>\<rightarrow> (\<lambda>k. Pair (pick xs k))"  | 
| 31180 | 102  | 
|
103  | 
lemma select_weight_member:  | 
|
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
62608 
diff
changeset
 | 
104  | 
assumes "0 < sum_list (map fst xs)"  | 
| 31180 | 105  | 
shows "fst (select_weight xs s) \<in> set (map snd xs)"  | 
106  | 
proof -  | 
|
107  | 
from range assms  | 
|
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
62608 
diff
changeset
 | 
108  | 
have "fst (range (sum_list (map fst xs)) s) < sum_list (map fst xs)" .  | 
| 31180 | 109  | 
with pick_member  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
62608 
diff
changeset
 | 
110  | 
have "pick xs (fst (range (sum_list (map fst xs)) s)) \<in> set (map snd xs)" .  | 
| 31180 | 111  | 
then show ?thesis by (simp add: select_weight_def scomp_def split_def)  | 
112  | 
qed  | 
|
113  | 
||
| 31268 | 114  | 
lemma select_weight_cons_zero:  | 
115  | 
"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
 | 
116  | 
by (simp add: select_weight_def less_natural_def)  | 
| 31268 | 117  | 
|
| 46311 | 118  | 
lemma select_weight_drop_zero:  | 
| 31261 | 119  | 
"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
 | 
120  | 
proof -  | 
| 68386 | 121  | 
have "sum_list (map fst [(k, _)\<leftarrow>xs . 0 < k]) = sum_list (map fst xs)"  | 
| 62608 | 122  | 
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
 | 
123  | 
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
 | 
124  | 
qed  | 
| 
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
125  | 
|
| 46311 | 126  | 
lemma select_weight_select:  | 
| 
31203
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
127  | 
assumes "xs \<noteq> []"  | 
| 31261 | 128  | 
shows "select_weight (map (Pair 1) xs) = select xs"  | 
| 
31203
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
129  | 
proof -  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
130  | 
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
 | 
131  | 
using assms by (intro range) (simp add: less_natural_def)  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
62608 
diff
changeset
 | 
132  | 
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
 | 
133  | 
by (induct xs) simp_all  | 
| 
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
134  | 
ultimately show ?thesis  | 
| 
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
135  | 
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
 | 
136  | 
fun_eq_iff pick_same [symmetric] less_natural_def)  | 
| 
31203
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
137  | 
qed  | 
| 
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31196 
diff
changeset
 | 
138  | 
|
| 72581 | 139  | 
end  | 
140  | 
||
| 26265 | 141  | 
|
| 61799 | 142  | 
subsection \<open>\<open>ML\<close> interface\<close>  | 
| 22528 | 143  | 
|
| 
36538
 
4fe16d49283b
make random engine persistent using code_reflect
 
haftmann 
parents: 
36176 
diff
changeset
 | 
144  | 
code_reflect Random_Engine  | 
| 
 
4fe16d49283b
make random engine persistent using code_reflect
 
haftmann 
parents: 
36176 
diff
changeset
 | 
145  | 
functions range select select_weight  | 
| 
 
4fe16d49283b
make random engine persistent using code_reflect
 
haftmann 
parents: 
36176 
diff
changeset
 | 
146  | 
|
| 60758 | 147  | 
ML \<open>  | 
| 26265 | 148  | 
structure Random_Engine =  | 
| 22528 | 149  | 
struct  | 
150  | 
||
| 
36538
 
4fe16d49283b
make random engine persistent using code_reflect
 
haftmann 
parents: 
36176 
diff
changeset
 | 
151  | 
open Random_Engine;  | 
| 
 
4fe16d49283b
make random engine persistent using code_reflect
 
haftmann 
parents: 
36176 
diff
changeset
 | 
152  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
46311 
diff
changeset
 | 
153  | 
type seed = Code_Numeral.natural * Code_Numeral.natural;  | 
| 22528 | 154  | 
|
155  | 
local  | 
|
| 26038 | 156  | 
|
| 32740 | 157  | 
val seed = Unsynchronized.ref  | 
| 26265 | 158  | 
(let  | 
159  | 
val now = Time.toMilliseconds (Time.now ());  | 
|
| 26038 | 160  | 
val (q, s1) = IntInf.divMod (now, 2147483562);  | 
161  | 
val s2 = q mod 2147483398;  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
162  | 
in apply2 Code_Numeral.natural_of_integer (s1 + 1, s2 + 1) end);  | 
| 26265 | 163  | 
|
| 22528 | 164  | 
in  | 
| 26038 | 165  | 
|
| 
36020
 
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
 
bulwahn 
parents: 
35266 
diff
changeset
 | 
166  | 
fun next_seed () =  | 
| 
 
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
 
bulwahn 
parents: 
35266 
diff
changeset
 | 
167  | 
let  | 
| 
 
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
 
bulwahn 
parents: 
35266 
diff
changeset
 | 
168  | 
    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
 | 
169  | 
val _ = seed := seed'  | 
| 
 
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
 
bulwahn 
parents: 
35266 
diff
changeset
 | 
170  | 
in  | 
| 
 
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
 
bulwahn 
parents: 
35266 
diff
changeset
 | 
171  | 
seed1  | 
| 
 
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
 
bulwahn 
parents: 
35266 
diff
changeset
 | 
172  | 
end  | 
| 
 
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
 
bulwahn 
parents: 
35266 
diff
changeset
 | 
173  | 
|
| 26038 | 174  | 
fun run f =  | 
175  | 
let  | 
|
| 26265 | 176  | 
val (x, seed') = f (! seed);  | 
| 26038 | 177  | 
val _ = seed := seed'  | 
178  | 
in x end;  | 
|
179  | 
||
| 22528 | 180  | 
end;  | 
181  | 
||
182  | 
end;  | 
|
| 60758 | 183  | 
\<close>  | 
| 22528 | 184  | 
|
| 
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
 | 
185  | 
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
 | 
186  | 
hide_const (open) inc_shift minus_shift log "next" split_seed  | 
| 31636 | 187  | 
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
 | 
188  | 
hide_fact (open) range_def  | 
| 31180 | 189  | 
|
| 26038 | 190  | 
end  |