src/HOL/Library/DAList.thy
changeset 51126 df86080de4cb
parent 49834 b27bbb021df1
child 51143 0a2371e7ced3
equal deleted inserted replaced
51125:f90874d3a246 51126:df86080de4cb
   115   ('key, 'value) alist \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
   115   ('key, 'value) alist \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
   116   [code_unfold]: "valterm_update k v a = Code_Evaluation.valtermify update {\<cdot>} k {\<cdot>} v {\<cdot>}a"
   116   [code_unfold]: "valterm_update k v a = Code_Evaluation.valtermify update {\<cdot>} k {\<cdot>} v {\<cdot>}a"
   117 
   117 
   118 fun (in term_syntax) random_aux_alist 
   118 fun (in term_syntax) random_aux_alist 
   119 where
   119 where
   120   "random_aux_alist i j = (if i = 0 then Pair valterm_empty else Quickcheck.collapse (Random.select_weight [(i, Quickcheck.random j \<circ>\<rightarrow> (%k. Quickcheck.random j \<circ>\<rightarrow> (%v. random_aux_alist (i - 1) j \<circ>\<rightarrow> (%a. Pair (valterm_update k v a))))), (1, Pair valterm_empty)]))"
   120   "random_aux_alist i j = (if i = 0 then Pair valterm_empty else Quickcheck_Random.collapse (Random.select_weight [(i, Quickcheck_Random.random j \<circ>\<rightarrow> (%k. Quickcheck_Random.random j \<circ>\<rightarrow> (%v. random_aux_alist (i - 1) j \<circ>\<rightarrow> (%a. Pair (valterm_update k v a))))), (1, Pair valterm_empty)]))"
   121 
   121 
   122 instantiation alist :: (random, random) random
   122 instantiation alist :: (random, random) random
   123 begin
   123 begin
   124 
   124 
   125 definition random_alist
   125 definition random_alist