bulwahn@36020: haftmann@29815: (* Author: Florian Haftmann, TU Muenchen *) haftmann@22528: haftmann@26265: header {* A HOL random engine *} haftmann@22528: haftmann@22528: theory Random haftmann@31205: imports Code_Numeral List haftmann@22528: begin haftmann@22528: haftmann@37751: notation fcomp (infixl "\>" 60) haftmann@37751: notation scomp (infixl "\\" 60) haftmann@29823: haftmann@29823: haftmann@26265: subsection {* Auxiliary functions *} haftmann@26265: haftmann@51143: fun log :: "natural \ natural \ natural" where haftmann@33236: "log b i = (if b \ 1 \ i < b then 1 else 1 + log b (i div b))" haftmann@33236: haftmann@51143: definition inc_shift :: "natural \ natural \ natural" where haftmann@26265: "inc_shift v k = (if v = k then 1 else k + 1)" haftmann@26265: haftmann@51143: definition minus_shift :: "natural \ natural \ natural \ natural" where haftmann@26265: "minus_shift r k l = (if k < l then r + k - l else k - l)" haftmann@26265: haftmann@30495: haftmann@26265: subsection {* Random seeds *} haftmann@26038: haftmann@51143: type_synonym seed = "natural \ natural" haftmann@22528: haftmann@51143: primrec "next" :: "seed \ natural \ seed" where haftmann@26265: "next (v, w) = (let haftmann@26265: k = v div 53668; haftmann@33236: v' = minus_shift 2147483563 ((v mod 53668) * 40014) (k * 12211); haftmann@26265: l = w div 52774; haftmann@33236: w' = minus_shift 2147483399 ((w mod 52774) * 40692) (l * 3791); haftmann@26265: z = minus_shift 2147483562 v' (w' + 1) + 1 haftmann@26265: in (z, (v', w')))" 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@33236: w'' = inc_shift 2147483398 w haftmann@33236: in ((v'', w'), (v', w'')))" haftmann@26038: haftmann@26038: haftmann@26265: subsection {* Base selectors *} haftmann@22528: haftmann@51143: fun iterate :: "natural \ ('b \ 'a \ 'b \ 'a) \ 'b \ 'a \ 'b \ 'a" where haftmann@37751: "iterate k f x = (if k = 0 then Pair x else f x \\ iterate (k - 1) f)" haftmann@22528: haftmann@51143: definition range :: "natural \ seed \ natural \ seed" where haftmann@30495: "range k = iterate (log 2147483561 k) haftmann@37751: (\l. next \\ (\v. Pair (v + l * 2147483561))) 1 haftmann@37751: \\ (\v. Pair (v mod k))" haftmann@26265: haftmann@26265: lemma range: haftmann@30495: "k > 0 \ fst (range k s) < k" haftmann@51143: by (simp add: range_def split_def less_natural_def del: log.simps iterate.simps) haftmann@26038: haftmann@29823: definition select :: "'a list \ seed \ 'a \ seed" where haftmann@51143: "select xs = range (natural_of_nat (length xs)) haftmann@51143: \\ (\k. Pair (nth xs (nat_of_natural k)))" haftmann@29823: haftmann@26265: lemma select: haftmann@26265: assumes "xs \ []" haftmann@26265: shows "fst (select xs s) \ set xs" haftmann@26265: proof - haftmann@51143: from assms have "natural_of_nat (length xs) > 0" by (simp add: less_natural_def) haftmann@26265: with range have haftmann@51143: "fst (range (natural_of_nat (length xs)) s) < natural_of_nat (length xs)" by best haftmann@26265: then have haftmann@51143: "nat_of_natural (fst (range (natural_of_nat (length xs)) s)) < length xs" by (simp add: less_natural_def) haftmann@26265: then show ?thesis huffman@44921: by (simp add: split_beta select_def) haftmann@26265: qed haftmann@22528: haftmann@51143: primrec pick :: "(natural \ 'a) list \ natural \ '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@51143: by (induct xs arbitrary: i) (simp_all add: less_natural_def) haftmann@31180: haftmann@31180: lemma pick_drop_zero: haftmann@31180: "pick (filter (\(k, _). k > 0) xs) = pick xs" haftmann@51143: by (induct xs) (auto simp add: fun_eq_iff less_natural_def minus_natural_def) haftmann@31180: haftmann@31203: lemma pick_same: haftmann@51143: "l < length xs \ Random.pick (map (Pair 1) xs) (natural_of_nat l) = nth xs l" haftmann@31203: proof (induct xs arbitrary: l) haftmann@31203: case Nil then show ?case by simp haftmann@31203: next haftmann@51143: case (Cons x xs) then show ?case by (cases l) (simp_all add: less_natural_def) haftmann@31203: qed haftmann@31203: haftmann@51143: definition select_weight :: "(natural \ 'a) list \ seed \ 'a \ seed" where haftmann@31180: "select_weight xs = range (listsum (map fst xs)) haftmann@37751: \\ (\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@31268: lemma select_weight_cons_zero: haftmann@31268: "select_weight ((0, x) # xs) = select_weight xs" haftmann@51143: by (simp add: select_weight_def less_natural_def) haftmann@31268: bulwahn@46311: lemma select_weight_drop_zero: haftmann@31261: "select_weight (filter (\(k, _). k > 0) xs) = select_weight xs" haftmann@31203: proof - haftmann@31203: have "listsum (map fst [(k, _)\xs . 0 < k]) = listsum (map fst xs)" haftmann@51143: by (induct xs) (auto simp add: less_natural_def, simp add: plus_natural_def) haftmann@31203: then show ?thesis by (simp only: select_weight_def pick_drop_zero) haftmann@31203: qed haftmann@31203: bulwahn@46311: lemma select_weight_select: haftmann@31203: assumes "xs \ []" haftmann@31261: shows "select_weight (map (Pair 1) xs) = select xs" haftmann@31203: proof - haftmann@51143: have less: "\s. fst (range (natural_of_nat (length xs)) s) < natural_of_nat (length xs)" haftmann@51143: using assms by (intro range) (simp add: less_natural_def) haftmann@51143: moreover have "listsum (map fst (map (Pair 1) xs)) = natural_of_nat (length xs)" haftmann@31203: by (induct xs) simp_all haftmann@31203: ultimately show ?thesis haftmann@31203: by (auto simp add: select_weight_def select_def scomp_def split_def haftmann@51143: fun_eq_iff pick_same [symmetric] less_natural_def) haftmann@31203: qed haftmann@31203: haftmann@26265: haftmann@26265: subsection {* @{text ML} interface *} haftmann@22528: haftmann@36538: code_reflect Random_Engine haftmann@36538: functions range select select_weight haftmann@36538: haftmann@22528: ML {* haftmann@26265: structure Random_Engine = haftmann@22528: struct haftmann@22528: haftmann@36538: open Random_Engine; haftmann@36538: haftmann@51143: type seed = Code_Numeral.natural * Code_Numeral.natural; haftmann@22528: haftmann@22528: local haftmann@26038: wenzelm@32740: val seed = Unsynchronized.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@51143: in pairself Code_Numeral.natural_of_integer (s1 + 1, s2 + 1) end); haftmann@26265: haftmann@22528: in haftmann@26038: bulwahn@36020: fun next_seed () = bulwahn@36020: let bulwahn@36020: val (seed1, seed') = @{code split_seed} (! seed) bulwahn@36020: val _ = seed := seed' bulwahn@36020: in bulwahn@36020: seed1 bulwahn@36020: end bulwahn@36020: 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: wenzelm@36176: hide_type (open) seed wenzelm@36176: hide_const (open) inc_shift minus_shift log "next" split_seed haftmann@31636: iterate range select pick select_weight wenzelm@36176: hide_fact (open) range_def haftmann@31180: haftmann@37751: no_notation fcomp (infixl "\>" 60) haftmann@37751: no_notation scomp (infixl "\\" 60) haftmann@29823: haftmann@26038: end haftmann@51143: