haftmann@29815: (* Author: Florian Haftmann, TU Muenchen *) haftmann@22528: haftmann@26265: header {* A HOL random engine *} haftmann@22528: haftmann@22528: theory Random haftmann@29823: imports Code_Index haftmann@22528: begin haftmann@22528: haftmann@29823: notation fcomp (infixl "o>" 60) haftmann@29823: notation scomp (infixl "o\" 60) haftmann@29823: haftmann@29823: haftmann@26265: subsection {* Auxiliary functions *} haftmann@26265: haftmann@29823: definition inc_shift :: "index \ index \ index" where haftmann@26265: "inc_shift v k = (if v = k then 1 else k + 1)" haftmann@26265: haftmann@29823: definition minus_shift :: "index \ index \ index \ index" where haftmann@26265: "minus_shift r k l = (if k < l then r + k - l else k - l)" haftmann@26265: haftmann@29823: fun log :: "index \ index \ index" where haftmann@26265: "log b i = (if b \ 1 \ i < b then 1 else 1 + log b (i div b))" haftmann@26265: haftmann@30495: haftmann@26265: subsection {* Random seeds *} haftmann@26038: haftmann@26038: types seed = "index \ index" haftmann@22528: haftmann@29823: primrec "next" :: "seed \ index \ seed" where haftmann@26265: "next (v, w) = (let haftmann@26265: k = v div 53668; haftmann@26265: v' = minus_shift 2147483563 (40014 * (v mod 53668)) (k * 12211); haftmann@26265: l = w div 52774; haftmann@26265: w' = minus_shift 2147483399 (40692 * (w mod 52774)) (l * 3791); haftmann@26265: z = minus_shift 2147483562 v' (w' + 1) + 1 haftmann@26265: in (z, (v', w')))" haftmann@26265: haftmann@26265: lemma next_not_0: haftmann@26265: "fst (next s) \ 0" haftmann@29823: by (cases s) (auto simp add: minus_shift_def Let_def) haftmann@26265: haftmann@29823: primrec seed_invariant :: "seed \ bool" where haftmann@26265: "seed_invariant (v, w) \ 0 < v \ v < 9438322952 \ 0 < w \ True" haftmann@26265: haftmann@29823: lemma if_same: "(if b then f x else f y) = f (if b then x else y)" haftmann@26265: by (cases b) simp_all haftmann@26265: haftmann@29823: definition split_seed :: "seed \ seed \ seed" where haftmann@26038: "split_seed s = (let haftmann@26038: (v, w) = s; haftmann@26038: (v', w') = snd (next s); haftmann@26265: v'' = inc_shift 2147483562 v; haftmann@26038: s'' = (v'', w'); haftmann@26265: w'' = inc_shift 2147483398 w; haftmann@26038: s''' = (v', w'') haftmann@26038: in (s'', s'''))" haftmann@26038: haftmann@26038: haftmann@26265: subsection {* Base selectors *} haftmann@22528: haftmann@30500: fun iterate :: "index \ ('b \ 'a \ 'b \ 'a) \ 'b \ 'a \ 'b \ 'a" where haftmann@30495: "iterate k f x = (if k = 0 then Pair x else f x o\ iterate (k - 1) f)" haftmann@22528: haftmann@30500: definition range :: "index \ seed \ index \ seed" where haftmann@30495: "range k = iterate (log 2147483561 k) haftmann@30495: (\l. next o\ (\v. Pair (v + l * 2147483561))) 1 haftmann@29823: o\ (\v. Pair (v mod k))" haftmann@26265: haftmann@26265: lemma range: haftmann@30495: "k > 0 \ fst (range k s) < k" haftmann@30495: by (simp add: range_def scomp_apply split_def del: log.simps iterate.simps) haftmann@26038: haftmann@29823: definition select :: "'a list \ seed \ 'a \ seed" where haftmann@29823: "select xs = range (Code_Index.of_nat (length xs)) haftmann@29823: o\ (\k. Pair (nth xs (Code_Index.nat_of k)))" haftmann@29823: haftmann@26265: lemma select: haftmann@26265: assumes "xs \ []" haftmann@26265: shows "fst (select xs s) \ set xs" haftmann@26265: proof - haftmann@29815: from assms have "Code_Index.of_nat (length xs) > 0" by simp haftmann@26265: with range have haftmann@29815: "fst (range (Code_Index.of_nat (length xs)) s) < Code_Index.of_nat (length xs)" by best haftmann@26265: then have haftmann@29815: "Code_Index.nat_of (fst (range (Code_Index.of_nat (length xs)) s)) < length xs" by simp haftmann@26265: then show ?thesis haftmann@29823: by (simp add: scomp_apply split_beta select_def) haftmann@26265: qed haftmann@22528: haftmann@31180: primrec pick :: "(index \ 'a) list \ index \ 'a" where haftmann@31180: "pick (x # xs) i = (if i < fst x then snd x else pick xs (i - fst x))" haftmann@31180: haftmann@31180: lemma pick_member: haftmann@31180: "i < listsum (map fst xs) \ pick xs i \ set (map snd xs)" haftmann@31180: by (induct xs arbitrary: i) simp_all haftmann@31180: haftmann@31180: lemma pick_drop_zero: haftmann@31180: "pick (filter (\(k, _). k > 0) xs) = pick xs" haftmann@31180: by (induct xs) (auto simp add: expand_fun_eq) haftmann@31180: haftmann@31180: definition select_weight :: "(index \ 'a) list \ seed \ 'a \ seed" where haftmann@31180: "select_weight xs = range (listsum (map fst xs)) haftmann@31180: o\ (\k. Pair (pick xs k))" haftmann@31180: haftmann@31180: lemma select_weight_member: haftmann@31180: assumes "0 < listsum (map fst xs)" haftmann@31180: shows "fst (select_weight xs s) \ set (map snd xs)" haftmann@31180: proof - haftmann@31180: from range assms haftmann@31180: have "fst (range (listsum (map fst xs)) s) < listsum (map fst xs)" . haftmann@31180: with pick_member haftmann@31180: have "pick xs (fst (range (listsum (map fst xs)) s)) \ set (map snd xs)" . haftmann@31180: then show ?thesis by (simp add: select_weight_def scomp_def split_def) haftmann@31180: qed haftmann@31180: haftmann@29823: definition select_default :: "index \ 'a \ 'a \ seed \ 'a \ seed" where haftmann@29823: [code del]: "select_default k x y = range k haftmann@29823: o\ (\l. Pair (if l + 1 < k then x else y))" haftmann@26265: haftmann@26265: lemma select_default_zero: haftmann@26265: "fst (select_default 0 x y s) = y" haftmann@29823: by (simp add: scomp_apply split_beta select_default_def) haftmann@26038: haftmann@26265: lemma select_default_code [code]: haftmann@29823: "select_default k x y = (if k = 0 haftmann@29823: then range 1 o\ (\_. Pair y) haftmann@29823: else range k o\ (\l. Pair (if l + 1 < k then x else y)))" haftmann@29823: proof haftmann@29823: fix s haftmann@29823: have "snd (range (Code_Index.of_nat 0) s) = snd (range (Code_Index.of_nat 1) s)" haftmann@29823: by (simp add: range_def scomp_Pair scomp_apply split_beta) haftmann@29823: then show "select_default k x y s = (if k = 0 haftmann@29823: then range 1 o\ (\_. Pair y) haftmann@29823: else range k o\ (\l. Pair (if l + 1 < k then x else y))) s" haftmann@29823: by (cases "k = 0") (simp_all add: select_default_def scomp_apply split_beta) haftmann@26265: qed haftmann@22528: haftmann@26265: haftmann@26265: subsection {* @{text ML} interface *} haftmann@22528: haftmann@22528: ML {* haftmann@26265: structure Random_Engine = haftmann@22528: struct haftmann@22528: haftmann@26038: type seed = int * int; haftmann@22528: haftmann@22528: local haftmann@26038: haftmann@26265: val seed = ref haftmann@26265: (let haftmann@26265: val now = Time.toMilliseconds (Time.now ()); haftmann@26038: val (q, s1) = IntInf.divMod (now, 2147483562); haftmann@26038: val s2 = q mod 2147483398; haftmann@26265: in (s1 + 1, s2 + 1) end); haftmann@26265: haftmann@22528: in haftmann@26038: haftmann@26038: fun run f = haftmann@26038: let haftmann@26265: val (x, seed') = f (! seed); haftmann@26038: val _ = seed := seed' haftmann@26038: in x end; haftmann@26038: haftmann@22528: end; haftmann@22528: haftmann@22528: end; haftmann@22528: *} haftmann@22528: haftmann@31180: hide (open) type seed haftmann@31180: hide (open) const inc_shift minus_shift log "next" seed_invariant split_seed haftmann@31180: iterate range select pick select_weight select_default haftmann@31196: hide (open) fact log_def haftmann@31180: haftmann@29823: no_notation fcomp (infixl "o>" 60) haftmann@29823: no_notation scomp (infixl "o\" 60) haftmann@29823: haftmann@26038: end haftmann@28145: