src/HOL/Random.thy
author haftmann
Tue, 19 May 2009 16:54:55 +0200
changeset 31205 98370b26c2ce
parent 31203 5c8fb4fd67e0
child 31255 900ebbc35e30
permissions -rw-r--r--
String.literal replaces message_string, code_numeral replaces (code_)index
     1 (* Author: Florian Haftmann, TU Muenchen *)
     2 
     3 header {* A HOL random engine *}
     4 
     5 theory Random
     6 imports Code_Numeral List
     7 begin
     8 
     9 notation fcomp (infixl "o>" 60)
    10 notation scomp (infixl "o\<rightarrow>" 60)
    11 
    12 
    13 subsection {* Auxiliary functions *}
    14 
    15 definition inc_shift :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
    16   "inc_shift v k = (if v = k then 1 else k + 1)"
    17 
    18 definition minus_shift :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
    19   "minus_shift r k l = (if k < l then r + k - l else k - l)"
    20 
    21 fun log :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
    22   "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
    23 
    24 
    25 subsection {* Random seeds *}
    26 
    27 types seed = "code_numeral \<times> code_numeral"
    28 
    29 primrec "next" :: "seed \<Rightarrow> code_numeral \<times> seed" where
    30   "next (v, w) = (let
    31      k =  v div 53668;
    32      v' = minus_shift 2147483563 (40014 * (v mod 53668)) (k * 12211);
    33      l =  w div 52774;
    34      w' = minus_shift 2147483399 (40692 * (w mod 52774)) (l * 3791);
    35      z =  minus_shift 2147483562 v' (w' + 1) + 1
    36    in (z, (v', w')))"
    37 
    38 lemma next_not_0:
    39   "fst (next s) \<noteq> 0"
    40   by (cases s) (auto simp add: minus_shift_def Let_def)
    41 
    42 primrec seed_invariant :: "seed \<Rightarrow> bool" where
    43   "seed_invariant (v, w) \<longleftrightarrow> 0 < v \<and> v < 9438322952 \<and> 0 < w \<and> True"
    44 
    45 definition split_seed :: "seed \<Rightarrow> seed \<times> seed" where
    46   "split_seed s = (let
    47      (v, w) = s;
    48      (v', w') = snd (next s);
    49      v'' = inc_shift 2147483562 v;
    50      s'' = (v'', w');
    51      w'' = inc_shift 2147483398 w;
    52      s''' = (v', w'')
    53    in (s'', s'''))"
    54 
    55 
    56 subsection {* Base selectors *}
    57 
    58 fun iterate :: "code_numeral \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
    59   "iterate k f x = (if k = 0 then Pair x else f x o\<rightarrow> iterate (k - 1) f)"
    60 
    61 definition range :: "code_numeral \<Rightarrow> seed \<Rightarrow> code_numeral \<times> seed" where
    62   "range k = iterate (log 2147483561 k)
    63       (\<lambda>l. next o\<rightarrow> (\<lambda>v. Pair (v + l * 2147483561))) 1
    64     o\<rightarrow> (\<lambda>v. Pair (v mod k))"
    65 
    66 lemma range:
    67   "k > 0 \<Longrightarrow> fst (range k s) < k"
    68   by (simp add: range_def scomp_apply split_def del: log.simps iterate.simps)
    69 
    70 definition select :: "'a list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
    71   "select xs = range (Code_Numeral.of_nat (length xs))
    72     o\<rightarrow> (\<lambda>k. Pair (nth xs (Code_Numeral.nat_of k)))"
    73      
    74 lemma select:
    75   assumes "xs \<noteq> []"
    76   shows "fst (select xs s) \<in> set xs"
    77 proof -
    78   from assms have "Code_Numeral.of_nat (length xs) > 0" by simp
    79   with range have
    80     "fst (range (Code_Numeral.of_nat (length xs)) s) < Code_Numeral.of_nat (length xs)" by best
    81   then have
    82     "Code_Numeral.nat_of (fst (range (Code_Numeral.of_nat (length xs)) s)) < length xs" by simp
    83   then show ?thesis
    84     by (simp add: scomp_apply split_beta select_def)
    85 qed
    86 
    87 primrec pick :: "(code_numeral \<times> 'a) list \<Rightarrow> code_numeral \<Rightarrow> 'a" where
    88   "pick (x # xs) i = (if i < fst x then snd x else pick xs (i - fst x))"
    89 
    90 lemma pick_member:
    91   "i < listsum (map fst xs) \<Longrightarrow> pick xs i \<in> set (map snd xs)"
    92   by (induct xs arbitrary: i) simp_all
    93 
    94 lemma pick_drop_zero:
    95   "pick (filter (\<lambda>(k, _). k > 0) xs) = pick xs"
    96   by (induct xs) (auto simp add: expand_fun_eq)
    97 
    98 lemma pick_same:
    99   "l < length xs \<Longrightarrow> Random.pick (map (Pair 1) xs) (Code_Numeral.of_nat l) = nth xs l"
   100 proof (induct xs arbitrary: l)
   101   case Nil then show ?case by simp
   102 next
   103   case (Cons x xs) then show ?case by (cases l) simp_all
   104 qed
   105 
   106 definition select_weight :: "(code_numeral \<times> 'a) list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
   107   "select_weight xs = range (listsum (map fst xs))
   108    o\<rightarrow> (\<lambda>k. Pair (pick xs k))"
   109 
   110 lemma select_weight_member:
   111   assumes "0 < listsum (map fst xs)"
   112   shows "fst (select_weight xs s) \<in> set (map snd xs)"
   113 proof -
   114   from range assms
   115     have "fst (range (listsum (map fst xs)) s) < listsum (map fst xs)" .
   116   with pick_member
   117     have "pick xs (fst (range (listsum (map fst xs)) s)) \<in> set (map snd xs)" .
   118   then show ?thesis by (simp add: select_weight_def scomp_def split_def) 
   119 qed
   120 
   121 lemma select_weigth_drop_zero:
   122   "Random.select_weight (filter (\<lambda>(k, _). k > 0) xs) = Random.select_weight xs"
   123 proof -
   124   have "listsum (map fst [(k, _)\<leftarrow>xs . 0 < k]) = listsum (map fst xs)"
   125     by (induct xs) auto
   126   then show ?thesis by (simp only: select_weight_def pick_drop_zero)
   127 qed
   128 
   129 lemma select_weigth_select:
   130   assumes "xs \<noteq> []"
   131   shows "Random.select_weight (map (Pair 1) xs) = Random.select xs"
   132 proof -
   133   have less: "\<And>s. fst (Random.range (Code_Numeral.of_nat (length xs)) s) < Code_Numeral.of_nat (length xs)"
   134     using assms by (intro range) simp
   135   moreover have "listsum (map fst (map (Pair 1) xs)) = Code_Numeral.of_nat (length xs)"
   136     by (induct xs) simp_all
   137   ultimately show ?thesis
   138     by (auto simp add: select_weight_def select_def scomp_def split_def
   139       expand_fun_eq pick_same [symmetric])
   140 qed
   141 
   142 definition select_default :: "code_numeral \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
   143   [code del]: "select_default k x y = range k
   144      o\<rightarrow> (\<lambda>l. Pair (if l + 1 < k then x else y))"
   145 
   146 lemma select_default_zero:
   147   "fst (select_default 0 x y s) = y"
   148   by (simp add: scomp_apply split_beta select_default_def)
   149 
   150 lemma select_default_code [code]:
   151   "select_default k x y = (if k = 0
   152     then range 1 o\<rightarrow> (\<lambda>_. Pair y)
   153     else range k o\<rightarrow> (\<lambda>l. Pair (if l + 1 < k then x else y)))"
   154 proof
   155   fix s
   156   have "snd (range (Code_Numeral.of_nat 0) s) = snd (range (Code_Numeral.of_nat 1) s)"
   157     by (simp add: range_def scomp_Pair scomp_apply split_beta)
   158   then show "select_default k x y s = (if k = 0
   159     then range 1 o\<rightarrow> (\<lambda>_. Pair y)
   160     else range k o\<rightarrow> (\<lambda>l. Pair (if l + 1 < k then x else y))) s"
   161     by (cases "k = 0") (simp_all add: select_default_def scomp_apply split_beta)
   162 qed
   163 
   164 
   165 subsection {* @{text ML} interface *}
   166 
   167 ML {*
   168 structure Random_Engine =
   169 struct
   170 
   171 type seed = int * int;
   172 
   173 local
   174 
   175 val seed = ref 
   176   (let
   177     val now = Time.toMilliseconds (Time.now ());
   178     val (q, s1) = IntInf.divMod (now, 2147483562);
   179     val s2 = q mod 2147483398;
   180   in (s1 + 1, s2 + 1) end);
   181 
   182 in
   183 
   184 fun run f =
   185   let
   186     val (x, seed') = f (! seed);
   187     val _ = seed := seed'
   188   in x end;
   189 
   190 end;
   191 
   192 end;
   193 *}
   194 
   195 hide (open) type seed
   196 hide (open) const inc_shift minus_shift log "next" seed_invariant split_seed
   197   iterate range select pick select_weight select_default
   198 
   199 no_notation fcomp (infixl "o>" 60)
   200 no_notation scomp (infixl "o\<rightarrow>" 60)
   201 
   202 end
   203