# HG changeset patch # User haftmann # Date 1205347094 -3600 # Node ID 4b63b9e9b10dbf1efb0f20a553b5c64121a039ce # Parent 89e25cc8da7a7bb4811966501cbccfe843a12292 separated Random.thy from Quickcheck.thy diff -r 89e25cc8da7a -r 4b63b9e9b10d src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed Mar 12 19:38:13 2008 +0100 +++ b/src/HOL/IsaMakefile Wed Mar 12 19:38:14 2008 +0100 @@ -674,7 +674,7 @@ ex/Abstract_NAT.thy ex/Antiquote.thy ex/Arith_Examples.thy \ ex/BT.thy ex/BinEx.thy ex/CTL.thy \ ex/Chinese.thy ex/Classical.thy ex/Dense_Linear_Order_Ex.thy \ - ex/Eval_Examples.thy ex/Groebner_Examples.thy ex/Random.thy \ + ex/Eval_Examples.thy ex/Groebner_Examples.thy ex/Random.thy ex/Quickcheck.thy \ ex/Codegenerator.thy ex/Codegenerator_Pretty.thy \ ex/Commutative_RingEx.thy ex/Efficient_Nat_examples.thy ex/Hex_Bin_Examples.thy \ ex/Commutative_Ring_Complete.thy ex/ExecutableContent.thy \ diff -r 89e25cc8da7a -r 4b63b9e9b10d src/HOL/ex/Quickcheck.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Quickcheck.thy Wed Mar 12 19:38:14 2008 +0100 @@ -0,0 +1,383 @@ +(* ID: $Id$ + Author: Florian Haftmann, TU Muenchen +*) + +header {* A simple counterexample generator *} + +theory Quickcheck +imports Random Eval +begin + +subsection {* The @{text random} class *} + +class random = type + + fixes random :: "index \ seed \ 'a \ seed" + +print_classes + +instantiation itself :: (type) random +begin + +definition + "random _ = return TYPE('a)" + +instance .. + +end + +lemma random_aux_if: + fixes random' :: "index \ index \ seed \ 'a \ seed" + assumes "random' 0 j = undefined" + and "\i. random' (Suc_index i) j = rhs2 i" + shows "random' i j s = (if i = 0 then undefined else rhs2 (i - 1) s)" + by (cases i rule: index.exhaust) (insert assms, simp_all add: undefined_fun) + +setup {* +let + exception REC; + fun random_inst tyco thy = + let + val { descr, index, ... } = DatatypePackage.the_datatype thy tyco; + val _ = if length descr > 1 then raise REC else (); + val (raw_vs, _) = DatatypePackage.the_datatype_spec thy tyco; + val vs = (map o apsnd) + (curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort random}) raw_vs; + val ty = Type (tyco, map TFree vs); + val typ_of = DatatypeAux.typ_of_dtyp descr vs; + val SOME (_, _, constrs) = AList.lookup (op =) descr index; + val randomN = NameSpace.base @{const_name random}; + val random_aux_name = randomN ^ "_" ^ Class.type_name tyco ^ "'"; + fun lift_ty ty = StateMonad.liftT ty @{typ seed}; + val ty_aux = @{typ index} --> @{typ index} --> lift_ty ty; + fun random ty = + Const (@{const_name random}, @{typ index} --> lift_ty ty); + val random_aux = Free (random_aux_name, ty_aux); + fun add_cons_arg dty (is_rec, t) = + let + val ty' = typ_of dty; + val rec_call = case try DatatypeAux.dest_DtRec dty + of SOME index' => index = index' + | NONE => false + val random' = if rec_call + then random_aux $ @{term "i\index"} $ @{term "j\index"} + else random ty' $ @{term "j\index"} + val is_rec' = is_rec orelse DatatypeAux.is_rec_type dty; + val t' = StateMonad.mbind ty' ty @{typ seed} random' (Abs ("", ty', t)) + in (is_rec', t') end; + fun mk_cons_t (c, dtys) = + let + val ty' = map typ_of dtys ---> ty; + val t = StateMonad.return ty @{typ seed} (list_comb (Const (c, ty'), + map Bound (length dtys - 1 downto 0))); + val (is_rec, t') = fold_rev add_cons_arg dtys (false, t); + in (is_rec, StateMonad.run ty @{typ seed} t') end; + fun check_empty [] = NONE + | check_empty xs = SOME xs; + fun bundle_cons_ts cons_ts = + let + val ts = map snd cons_ts; + val t = HOLogic.mk_list (lift_ty ty) ts; + val t' = Const (@{const_name select}, HOLogic.listT (lift_ty ty) --> lift_ty (lift_ty ty)) $ t; + val t'' = Const (@{const_name collapse}, lift_ty (lift_ty ty) --> lift_ty ty) $ t'; + in t'' end; + fun bundle_conss (some_rec_t, nonrec_t) = + let + val t = case some_rec_t + of SOME rec_t => Const (@{const_name collapse}, lift_ty (lift_ty ty) --> lift_ty ty) + $ (Const (@{const_name select_default}, + @{typ index} --> lift_ty ty --> lift_ty ty --> lift_ty (lift_ty ty)) + $ @{term "i\index"} $ rec_t $ nonrec_t) + | NONE => nonrec_t + in t end; + val random_rhs = constrs + |> map mk_cons_t + |> List.partition fst + |> apfst (Option.map bundle_cons_ts o check_empty) + |> apsnd bundle_cons_ts + |> bundle_conss; + val random_aux_undef = (HOLogic.mk_Trueprop o HOLogic.mk_eq) + (random_aux $ @{term "0\index"} $ @{term "j\index"}, Const (@{const_name undefined}, lift_ty ty)) + val random_aux_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) + (random_aux $ @{term "Suc_index i"} $ @{term "j\index"}, random_rhs); + val random_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (Const (@{const_name random}, + @{typ index} --> lift_ty ty) $ @{term "i\index"}, + random_aux $ @{term "i\index"} $ @{term "i\index"}); + val del_func = Attrib.internal (fn _ => Thm.declaration_attribute + (fn thm => Context.mapping (Code.del_func thm) I)); + fun add_code simps lthy = + let + val thy = ProofContext.theory_of lthy; + val thm = @{thm random_aux_if} + |> Drule.instantiate' [SOME (Thm.ctyp_of thy ty)] [SOME (Thm.cterm_of thy random_aux)] + |> (fn thm => thm OF simps) + |> singleton (ProofContext.export lthy (ProofContext.init thy)) + in + lthy + |> LocalTheory.theory (PureThy.note Thm.internalK (random_aux_name ^ "_code", thm) + #-> Code.add_func) + end; + in + thy + |> TheoryTarget.instantiation ([tyco], vs, @{sort random}) + |> PrimrecPackage.add_primrec [(random_aux_name, SOME ty_aux, NoSyn)] + [(("", [del_func]), random_aux_undef), (("", [del_func]), random_aux_eq)] + |-> add_code + |> `(fn lthy => Syntax.check_term lthy random_eq) + |-> (fn eq => Specification.definition (NONE, (("", []), eq))) + |> snd + |> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) + |> LocalTheory.exit + |> ProofContext.theory_of + end; + fun add_random_inst [tyco] = (fn thy => random_inst tyco thy handle REC => + (warning ("Will not generated random elements for mutual recursive type " ^ quote tyco); thy)) + | add_random_inst tycos = tap (fn _ => warning + ("Will not generated random elements for mutual recursive type(s) " ^ commas (map quote tycos))); +in DatatypePackage.interpretation add_random_inst end +*} + +instantiation int :: random +begin + +definition + "random n = (do + (b, m) \ random n; + return (if b then int m else - int m) + done)" + +instance .. + +end + +instantiation set :: (random) random +begin + +primrec random_set' :: "index \ index \ seed \ 'a set \ seed" where + "random_set' 0 j = undefined" + | "random_set' (Suc_index i) j = collapse (select_default i + (do x \ random i; xs \ random_set' i j; return (insert x xs) done) + (return {}))" + +lemma random_set'_code [code func]: + "random_set' i j s = (if i = 0 then undefined else collapse (select_default (i - 1) + (do x \ random (i - 1); xs \ random_set' (i - 1) j; return (insert x xs) done) + (return {})) s)" + by (rule random_aux_if random_set'.simps)+ + +definition + "random i = random_set' i i" + +instance .. + +end + +code_reserved SML Quickcheck + + +subsection {* Quickcheck generator *} + +ML {* +structure Quickcheck = +struct + +val eval_ref : (unit -> int -> int * int -> term list option * (int * int)) option ref = ref NONE; + +fun mk_generator_expr prop tys = + let + val bounds = map_index (fn (i, ty) => (i, ty)) tys; + val result = list_comb (prop, map (fn (i, _) => Bound (length tys - i - 1)) bounds); + val terms = map (fn (i, ty) => Const (@{const_name Eval.term_of}, ty --> @{typ term}) $ Bound (length tys - i - 1)) bounds; + val check = @{term "If \ bool \ term list option \ term list option \ term list option"} + $ result $ @{term "None \ term list option"} $ (@{term "Some \ term list \ term list option "} $ HOLogic.mk_list @{typ term} terms); + val return = @{term "Pair \ term list option \ seed \ term list option \ seed"}; + fun mk_bindtyp ty = @{typ seed} --> HOLogic.mk_prodT (ty, @{typ seed}); + fun mk_bindclause (i, ty) t = Const (@{const_name mbind}, mk_bindtyp ty + --> (ty --> mk_bindtyp @{typ "term list option"}) --> mk_bindtyp @{typ "term list option"}) + $ (Const (@{const_name random}, @{typ index} --> mk_bindtyp ty) + $ Bound i) $ Abs ("a", ty, t); + val t = fold_rev mk_bindclause bounds (return $ check); + in Abs ("n", @{typ index}, t) end; + +fun compile_generator_expr thy prop tys = + let + val f = CodePackage.eval_term ("Quickcheck.eval_ref", eval_ref) thy + (mk_generator_expr prop tys) []; + in f #> Random_Engine.run #> (Option.map o map) (Code.postprocess_term thy) end; + +fun VALUE prop tys thy = + let + val t = mk_generator_expr prop tys; + val eq = Logic.mk_equals (Free ("VALUE", fastype_of t), t) + in + thy + |> TheoryTarget.init NONE + |> Specification.definition (NONE, (("", []), eq)) + |> snd + |> LocalTheory.exit + |> ProofContext.theory_of + end; + +end +*} + +(*setup {* Quickcheck.VALUE @{term "\(n::nat) (m::nat) (q::nat). n = m + q + 1"} [@{typ nat}, @{typ nat}, @{typ nat}] *} +export_code VALUE in SML module_name QuickcheckExample file "~~/../../gen_code/quickcheck.ML" +use "~~/../../gen_code/quickcheck.ML" +ML {* Random_Engine.run (QuickcheckExample.range 1) *}*) + +ML {* val f = Quickcheck.compile_generator_expr @{theory} + @{term "\(n::nat) (m::nat) (q::nat). n = m + q + 1"} [@{typ nat}, @{typ nat}, @{typ nat}] *} + +ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 25 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} + +ML {* val f = Quickcheck.compile_generator_expr @{theory} + @{term "\(n::int) (m::int) (q::int). n = m + q + 1"} [@{typ int}, @{typ int}, @{typ int}] *} + +(*definition "FOO = (True, Suc 0)" + +code_module (test) QuickcheckExample + file "~~/../../gen_code/quickcheck'.ML" + contains FOO*) + +ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 25 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 3 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} + +ML {* val f = Quickcheck.compile_generator_expr @{theory} + @{term "\(xs\int list) ys. rev (xs @ ys) = rev xs @ rev ys"} + [@{typ "int list"}, @{typ "int list"}] *} + +ML {* f 15 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 25 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 8 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 8 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 8 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 88 |> (Option.map o map) (Sign.string_of_term @{theory}) *} + +subsection {* Incremental function generator *} + +ML {* +structure Quickcheck = +struct + +open Quickcheck; + +fun random_fun (eq : 'a -> 'a -> bool) + (random : Random_Engine.seed -> 'b * Random_Engine.seed) + (random_split : Random_Engine.seed -> Random_Engine.seed * Random_Engine.seed) + (seed : Random_Engine.seed) = + let + val (seed', seed'') = random_split seed; + val state = ref (seed', []); + fun random_fun' x = + let + val (seed, fun_map) = ! state; + in case AList.lookup (uncurry eq) fun_map x + of SOME y => y + | NONE => let + val (y, seed') = random seed; + val _ = state := (seed', (x, y) :: fun_map); + in y end + end; + in (random_fun', seed'') end; + +end +*} + +axiomatization + random_fun_aux :: "('a \ 'a \ bool) \ (seed \ 'b \ seed) + \ (seed \ seed \ seed) \ seed \ ('a \ 'b) \ seed" + +code_const random_fun_aux (SML "Quickcheck.random'_fun") + +instantiation "fun" :: (term_of, term_of) term_of +begin + +instance .. + +end + +code_const "Eval.term_of :: ('a\term_of \ 'b\term_of) \ _" + (SML "(fn '_ => Const (\"arbitrary\", dummyT))") + +instantiation "fun" :: (eq, "{type, random}") random +begin + +definition + "random n = random_fun_aux (op =) (random n) split_seed" + +instance .. + +end + +ML {* val f = Quickcheck.compile_generator_expr @{theory} + @{term "\f k. int (f k) = k"} [@{typ "int \ nat"}, @{typ int}] *} + +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} + +ML {* val f = Quickcheck.compile_generator_expr @{theory} + @{term "\(A \ nat set) B. card (A \ B) = card A + card B"} [@{typ "nat set"}, @{typ "nat set"}] *} + +ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 3 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 6 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 10 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 10 |> (Option.map o map) (Sign.string_of_term @{theory}) *} + +ML {* val f = Quickcheck.compile_generator_expr @{theory} + @{term "\(s \ string). s \ rev s"} [@{typ string}] *} + +ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 10 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 10 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 10 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 10 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 10 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 8 |> (Option.map o map) (Sign.string_of_term @{theory}) *} +ML {* f 8 |> (Option.map o map) (Sign.string_of_term @{theory}) *} + +end diff -r 89e25cc8da7a -r 4b63b9e9b10d src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Wed Mar 12 19:38:13 2008 +0100 +++ b/src/HOL/ex/ROOT.ML Wed Mar 12 19:38:14 2008 +0100 @@ -14,7 +14,7 @@ "FuncSet", "Word", "Eval_Examples", - "Random" + "Quickcheck" ]; no_document use_thy "Codegenerator"; diff -r 89e25cc8da7a -r 4b63b9e9b10d src/HOL/ex/Random.thy --- a/src/HOL/ex/Random.thy Wed Mar 12 19:38:13 2008 +0100 +++ b/src/HOL/ex/Random.thy Wed Mar 12 19:38:14 2008 +0100 @@ -2,30 +2,75 @@ Author: Florian Haftmann, TU Muenchen *) -header {* A simple term generator *} +header {* A HOL random engine *} theory Random -imports State_Monad Code_Index List Eval +imports State_Monad Code_Index begin -subsection {* The random engine *} +subsection {* Auxiliary functions *} + +definition + inc_shift :: "index \ index \ index" +where + "inc_shift v k = (if v = k then 1 else k + 1)" + +definition + minus_shift :: "index \ index \ index \ index" +where + "minus_shift r k l = (if k < l then r + k - l else k - l)" + +function + log :: "index \ index \ index" +where + "log b i = (if b \ 1 \ i < b then 1 else 1 + log b (i div b))" +by pat_completeness auto +termination + by (relation "measure (nat_of_index o snd)") + (auto simp add: index) + + +subsection {* Random seeds *} types seed = "index \ index" -definition +primrec "next" :: "seed \ index \ seed" where - "next s = (let - (v, w) = s; - k = v div 53668; - v' = 40014 * (v - k * 53668) - k * 12211; - v'' = (if v' < 0 then v' + 2147483563 else v'); - l = w div 52774; - w' = 40692 * (w - l * 52774) - l * 3791; - w'' = (if w' < 0 then w' + 2147483399 else w'); - z = v'' - w''; - z' = (if z < 1 then z + 2147483562 else z) - in (z', (v'', w'')))" + "next (v, w) = (let + k = v div 53668; + v' = minus_shift 2147483563 (40014 * (v mod 53668)) (k * 12211); + l = w div 52774; + w' = minus_shift 2147483399 (40692 * (w mod 52774)) (l * 3791); + z = minus_shift 2147483562 v' (w' + 1) + 1 + in (z, (v', w')))" + +lemma next_not_0: + "fst (next s) \ 0" +apply (cases s) +apply (auto simp add: minus_shift_def Let_def) +done + +primrec + seed_invariant :: "seed \ bool" +where + "seed_invariant (v, w) \ 0 < v \ v < 9438322952 \ 0 < w \ True" + +lemma if_same: + "(if b then f x else f y) = f (if b then x else y)" + by (cases b) simp_all + +(*lemma seed_invariant: + assumes "seed_invariant (index_of_nat v, index_of_nat w)" + and "(index_of_nat z, (index_of_nat v', index_of_nat w')) = next (index_of_nat v, index_of_nat w)" + shows "seed_invariant (index_of_nat v', index_of_nat w')" +using assms +apply (auto simp add: seed_invariant_def) +apply (auto simp add: minus_shift_def Let_def) +apply (simp_all add: if_same cong del: if_cong) +apply safe +unfolding not_less +oops*) definition split_seed :: "seed \ seed \ seed" @@ -33,27 +78,14 @@ "split_seed s = (let (v, w) = s; (v', w') = snd (next s); - v'' = (if v = 2147483562 then 1 else v + 1); + v'' = inc_shift 2147483562 v; s'' = (v'', w'); - w'' = (if w = 2147483398 then 1 else w + 1); + w'' = inc_shift 2147483398 w; s''' = (v', w'') in (s'', s'''))" -text {* Base selectors *} -primrec - pick :: "(index \ 'a) list \ index \ 'a" -where - "pick (x#xs) n = (if n < fst x then snd x else pick xs (n - fst x))" - -function - iLogBase :: "index \ index \ index" -where - "iLogBase b i = (if b \ 1 \ i < b then 1 else 1 + iLogBase b (i div b))" -by pat_completeness auto -termination - by (relation "measure (nat_of_index o snd)") - (auto simp add: index) +subsection {* Base selectors *} function range_aux :: "index \ index \ seed \ index \ seed" @@ -69,122 +101,94 @@ definition range :: "index \ seed \ index \ seed" where - "range k s = (let - b = 2147483561; - n = iLogBase b k; - (v, s') = range_aux n 1 s - in (v mod k, s'))" + "range k = (do + v \ range_aux (log 2147483561 k) 1; + return (v mod k) + done)" + +lemma range: + assumes "k > 0" + shows "fst (range k s) < k" +proof - + obtain v w where range_aux: + "range_aux (log 2147483561 k) 1 s = (v, w)" + by (cases "range_aux (log 2147483561 k) 1 s") + with assms show ?thesis + by (simp add: range_def run_def mbind_def split_def del: range_aux.simps log.simps) +qed definition select :: "'a list \ seed \ 'a \ seed" where - [simp]: "select xs s = (let - (k, s') = range (index_of_nat (length xs)) s - in (nth xs (nat_of_index k), s'))" + "select xs = (do + k \ range (index_of_nat (length xs)); + return (nth xs (nat_of_index k)) + done)" + +lemma select: + assumes "xs \ []" + shows "fst (select xs s) \ set xs" +proof - + from assms have "index_of_nat (length xs) > 0" by simp + with range have + "fst (range (index_of_nat (length xs)) s) < index_of_nat (length xs)" by best + then have + "nat_of_index (fst (range (index_of_nat (length xs)) s)) < length xs" by simp + then show ?thesis + by (auto simp add: select_def run_def mbind_def split_def) +qed definition - select_weight :: "(index \ 'a) list \ seed \ 'a \ seed" + select_default :: "index \ 'a \ 'a \ seed \ 'a \ seed" where - [simp]: "select_weight xs s = (let - (k, s') = range (foldr (op + \ fst) xs 0) s - in (pick xs k, s'))" + [code func del]: "select_default k x y = (do + l \ range k; + return (if l + 1 < k then x else y) + done)" + +lemma select_default_zero: + "fst (select_default 0 x y s) = y" + by (simp add: run_def mbind_def split_def select_default_def) -(*lemma - "select (x#xs) s = select_weight (map (Pair 1) (x#xs)) s" -proof (induct xs) - case Nil show ?case apply (auto simp add: Let_def split_def) +lemma select_default_code [code]: + "select_default k x y = (if k = 0 then do + _ \ range 1; + return y + done else do + l \ range k; + return (if l + 1 < k then x else y) + done)" +proof (cases "k = 0") + case False then show ?thesis by (simp add: select_default_def) next - have map_fst_Pair: "!!xs y. map fst (map (Pair y) xs) = replicate (length xs) y" - proof - - fix xs - fix y - show "map fst (map (Pair y) xs) = replicate (length xs) y" - by (induct xs) simp_all - qed - have pick_nth: "!!xs n. n < length xs \ pick (map (Pair 1) xs) n = nth xs n" - proof - - fix xs - fix n - assume "n < length xs" - then show "pick (map (Pair 1) xs) n = nth xs n" - proof (induct xs arbitrary: n) - case Nil then show ?case by simp - next - case (Cons x xs) show ?case - proof (cases n) - case 0 then show ?thesis by simp - next - case (Suc _) - from Cons have "n < length (x # xs)" by auto - then have "n < Suc (length xs)" by simp - with Suc have "n - 1 < Suc (length xs) - 1" by auto - with Cons have "pick (map (Pair (1\nat)) xs) (n - 1) = xs ! (n - 1)" by auto - with Suc show ?thesis by auto - qed - qed - qed - have sum_length: "!!xs. foldl (op +) 0 (map fst (map (Pair 1) xs)) = length xs" - proof - - have replicate_append: - "!!x xs y. replicate (length (x # xs)) y = replicate (length xs) y @ [y]" - by (simp add: replicate_app_Cons_same) - fix xs - show "foldl (op +) 0 (map fst (map (Pair 1) xs)) = length xs" - unfolding map_fst_Pair proof (induct xs) - case Nil show ?case by simp - next - case (Cons x xs) then show ?case unfolding replicate_append by simp - qed - qed - have pick_nth_random: - "!!x xs s. pick (map (Pair 1) (x#xs)) (fst (random (length (x#xs)) s)) = nth (x#xs) (fst (random (length (x#xs)) s))" - proof - - fix s - fix x - fix xs - have bound: "fst (random (length (x#xs)) s) < length (x#xs)" by (rule random_bound) simp - from pick_nth [OF bound] show - "pick (map (Pair 1) (x#xs)) (fst (random (length (x#xs)) s)) = nth (x#xs) (fst (random (length (x#xs)) s))" . - qed - have pick_nth_random_do: - "!!x xs s. (do n \ random (length (x#xs)); return (pick (map (Pair 1) (x#xs)) n) done) s = - (do n \ random (length (x#xs)); return (nth (x#xs) n) done) s" - unfolding monad_collapse split_def unfolding pick_nth_random .. - case (Cons x xs) then show ?case - unfolding select_weight_def sum_length pick_nth_random_do - by simp -qed*) + case True then show ?thesis + by (simp add: run_def mbind_def split_def select_default_def expand_fun_eq range_def) +qed -text {* The @{text ML} interface *} + +subsection {* @{text ML} interface *} ML {* -structure Quickcheck = +structure Random_Engine = struct type seed = int * int; local -val seed = ref (0, 0); - -fun init () = - let - val now = Time.toNanoseconds (Time.now ()); +val seed = ref + (let + val now = Time.toMilliseconds (Time.now ()); val (q, s1) = IntInf.divMod (now, 2147483562); val s2 = q mod 2147483398; - in seed := (s1 + 1, s2 + 1) end; - + in (s1 + 1, s2 + 1) end); + in -val seed = seed; (* FIXME *) - fun run f = let - val (x, seed') = f (!seed); + val (x, seed') = f (! seed); val _ = seed := seed' - val _ = if fst (! seed) = 0 orelse snd (! seed) = 0 then - (warning "broken random seed"; init ()) - else () in x end; end; @@ -192,325 +196,4 @@ end; *} - -subsection {* The @{text random} class *} - -class random = - fixes random :: "index \ seed \ 'a\{} \ seed" - -instantiation itself :: (type) random -begin - -definition - "random _ = return TYPE('a)" - -instance .. - end - -lemma random_aux_if: - fixes random' :: "index \ index \ seed \ 'a \ seed" - assumes "random' 0 j = undefined" - and "\i. random' (Suc_index i) j = rhs2 i" - shows "random' i j s = (if i = 0 then undefined else rhs2 (i - 1) s)" - by (cases i rule: index.exhaust) (insert assms, simp_all add: undefined_fun) - -setup {* -let - fun random_inst tyco thy = - let - val { descr, index, ... } = DatatypePackage.the_datatype thy tyco; - val (raw_vs, _) = DatatypePackage.the_datatype_spec thy tyco; - val vs = (map o apsnd) - (curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort random}) raw_vs; - val ty = Type (tyco, map TFree vs); - val typ_of = DatatypeAux.typ_of_dtyp descr vs; - val SOME (_, _, constrs) = AList.lookup (op =) descr index; - val randomN = NameSpace.base @{const_name random}; - val random_aux_name = randomN ^ "_" ^ Class.type_name tyco ^ "'"; - fun lift_ty ty = StateMonad.liftT ty @{typ seed}; - val ty_aux = @{typ index} --> @{typ index} --> lift_ty ty; - fun random ty = - Const (@{const_name random}, @{typ index} --> lift_ty ty); - val random_aux = Free (random_aux_name, ty_aux); - fun add_cons_arg dty (is_rec, t) = - let - val ty' = typ_of dty; - val random' = if can DatatypeAux.dest_DtRec dty - then random_aux $ @{term "i\index"} $ @{term "j\index"} - else random ty' $ @{term "j\index"} - val is_rec' = is_rec orelse DatatypeAux.is_rec_type dty; - val t' = StateMonad.mbind ty' ty @{typ seed} random' (Abs ("", ty', t)) - in (is_rec', t') end; - fun mk_cons_t (c, dtys) = - let - val ty' = map typ_of dtys ---> ty; - val t = StateMonad.return ty @{typ seed} (list_comb (Const (c, ty'), - map Bound (length dtys - 1 downto 0))); - val (is_rec, t') = fold_rev add_cons_arg dtys (false, t); - in (is_rec, StateMonad.run ty @{typ seed} t') end; - fun check_empty [] = NONE - | check_empty xs = SOME xs; - fun bundle_cons_ts cons_ts = - let - val ts = map snd cons_ts; - val t = HOLogic.mk_list (lift_ty ty) ts; - val t' = Const (@{const_name select}, HOLogic.listT (lift_ty ty) --> lift_ty (lift_ty ty)) $ t; - val t'' = Const (@{const_name collapse}, lift_ty (lift_ty ty) --> lift_ty ty) $ t'; - in t'' end; - fun bundle_conss (some_rec_t, nonrec_t) = - let - val rec' = case some_rec_t - of SOME rec_t => SOME (HOLogic.mk_prod (@{term "i\index"}, rec_t)) - | NONE => NONE; - val nonrec' = HOLogic.mk_prod (@{term "1\index"}, nonrec_t); - val i_ty = HOLogic.mk_prodT (@{typ index}, lift_ty ty); - val t = HOLogic.mk_list i_ty (the_list rec' @ single nonrec'); - val t' = Const (@{const_name select_weight}, HOLogic.listT i_ty --> lift_ty (lift_ty ty)) $ t; - val t'' = Const (@{const_name collapse}, lift_ty (lift_ty ty) --> lift_ty ty) $ t'; - in t'' end; - val random_rhs = constrs - |> map mk_cons_t - |> List.partition fst - |> apfst (Option.map bundle_cons_ts o check_empty) - |> apsnd bundle_cons_ts - |> bundle_conss; - val random_aux_undef = (HOLogic.mk_Trueprop o HOLogic.mk_eq) - (random_aux $ @{term "0\index"} $ @{term "j\index"}, Const (@{const_name undefined}, lift_ty ty)) - val random_aux_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) - (random_aux $ @{term "Suc_index i"} $ @{term "j\index"}, random_rhs); - val random_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (Const (@{const_name random}, - @{typ index} --> lift_ty ty) $ @{term "i\index"}, - random_aux $ @{term "i\index"} $ @{term "i\index"}); - val del_func = Attrib.internal (fn _ => Thm.declaration_attribute - (fn thm => Context.mapping (Code.del_func thm) I)); - fun add_code simps lthy = - let - val thy = ProofContext.theory_of lthy; - val thm = @{thm random_aux_if} - |> Drule.instantiate' [SOME (Thm.ctyp_of thy ty)] [SOME (Thm.cterm_of thy random_aux)] - |> (fn thm => thm OF simps) - |> singleton (ProofContext.export lthy (ProofContext.init thy)) - in - lthy - |> LocalTheory.theory (PureThy.note Thm.internalK (random_aux_name ^ "_code", thm) - #-> Code.add_func) - end; - in - thy - |> TheoryTarget.instantiation ([tyco], vs, @{sort random}) - |> PrimrecPackage.add_primrec [(random_aux_name, SOME ty_aux, NoSyn)] - [(("", [del_func]), random_aux_undef), (("", [del_func]), random_aux_eq)] - |-> add_code - |> `(fn lthy => Syntax.check_term lthy random_eq) - |-> (fn eq => Specification.definition (NONE, (("", []), eq))) - |> snd - |> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) - |> LocalTheory.exit - |> ProofContext.theory_of - end; - fun add_random_inst [tyco] = (fn thy => case try (random_inst tyco) thy - of SOME thy => thy - | NONE => (warning ("Failed to generate random elements for type" ^ tyco); thy)) - | add_random_inst tycos = tap (fn _ => warning - ("Will not generated random elements for mutual recursive types " ^ commas (map quote tycos))); -in DatatypePackage.interpretation add_random_inst end -*} - -instantiation int :: random -begin - -definition - "random n = (do - (b, m) \ random n; - return (if b then int m else - int m) - done)" - -instance .. - -end - -code_reserved SML Quickcheck - - -subsection {* Quickcheck generator *} - -ML {* -structure Quickcheck = -struct - -open Quickcheck; - -val eval_ref : (unit -> int -> int * int -> term list option * (int * int)) option ref = ref NONE; - -fun mk_generator_expr prop tys = - let - val bounds = map_index (fn (i, ty) => (i, ty)) tys; - val result = list_comb (prop, map (fn (i, _) => Bound (length tys - i - 1)) bounds); - val terms = map (fn (i, ty) => Const (@{const_name Eval.term_of}, ty --> @{typ term}) $ Bound (length tys - i - 1)) bounds; - val check = @{term "If \ bool \ term list option \ term list option \ term list option"} - $ result $ @{term "None \ term list option"} $ (@{term "Some \ term list \ term list option "} $ HOLogic.mk_list @{typ term} terms); - val return = @{term "Pair \ term list option \ seed \ term list option \ seed"}; - fun mk_bindtyp ty = @{typ seed} --> HOLogic.mk_prodT (ty, @{typ seed}); - fun mk_bindclause (i, ty) t = Const (@{const_name mbind}, mk_bindtyp ty - --> (ty --> mk_bindtyp @{typ "term list option"}) --> mk_bindtyp @{typ "term list option"}) - $ (Const (@{const_name random}, @{typ index} --> mk_bindtyp ty) - $ Bound i) $ Abs ("a", ty, t); - val t = fold_rev mk_bindclause bounds (return $ check); - in Abs ("n", @{typ index}, t) end; - -fun compile_generator_expr thy prop tys = - let - val f = CodePackage.eval_term ("Quickcheck.eval_ref", eval_ref) thy - (mk_generator_expr prop tys) []; - in f #> run #> (Option.map o map) (Code.postprocess_term thy) end; - -fun VALUE prop tys thy = - let - val t = mk_generator_expr prop tys; - val eq = Logic.mk_equals (Free ("VALUE", fastype_of t), t) - in - thy - |> TheoryTarget.init NONE - |> Specification.definition (NONE, (("", []), eq)) - |> snd - |> LocalTheory.exit - |> ProofContext.theory_of - end; - -end -*} - -ML {* val f = Quickcheck.compile_generator_expr @{theory} - @{term "\(n::nat) (m::nat) (q::nat). n = m + q + 1"} [@{typ nat}, @{typ nat}, @{typ nat}] *} - -ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 25 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} - -ML {* val f = Quickcheck.compile_generator_expr @{theory} - @{term "\(n::int) (m::int) (q::int). n = m + q + 1"} [@{typ int}, @{typ int}, @{typ int}] *} - -(*setup {* Quickcheck.VALUE @{term "\(n::int) (m::int) (q::int). n = m + q + 1"} [@{typ int}, @{typ int}, @{typ int}] *} - -export_code VALUE in SML module_name QuickcheckExample file "~~/../../gen_code/quickcheck.ML" - -definition "FOO = (True, Suc 0)" - -code_module (test) QuickcheckExample - file "~~/../../gen_code/quickcheck'.ML" - contains FOO*) - -ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 25 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 3 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} - -ML {* val f = Quickcheck.compile_generator_expr @{theory} - @{term "\(xs\nat list) ys. rev (xs @ ys) = rev xs @ rev ys"} - [@{typ "nat list"}, @{typ "nat list"}] *} - -ML {* f 15 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 25 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 1 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 2 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 4 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 5 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 8 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 8 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 8 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 88 |> (Option.map o map) (Sign.string_of_term @{theory}) *} - -subsection {* Incremental function generator *} - -ML {* -structure Quickcheck = -struct - -open Quickcheck; - -fun random_fun (eq : 'a -> 'a -> bool) - (random : seed -> 'b * seed) - (random_split : seed -> seed * seed) - (seed : seed) = - let - val (seed', seed'') = random_split seed; - val state = ref (seed', []); - fun random_fun' x = - let - val (seed, fun_map) = ! state; - in case AList.lookup (uncurry eq) fun_map x - of SOME y => y - | NONE => let - val (y, seed') = random seed; - val _ = state := (seed', (x, y) :: fun_map); - in y end - end; - in (random_fun', seed'') end; - -end -*} - -axiomatization - random_fun_aux :: "('a \ 'a \ bool) \ (seed \ 'b \ seed) - \ (seed \ seed \ seed) \ seed \ ('a \ 'b) \ seed" - -code_const random_fun_aux (SML "Quickcheck.random'_fun") - -instantiation "fun" :: (term_of, term_of) term_of -begin - -instance .. - -end - -code_const "Eval.term_of :: ('a\term_of \ 'b\term_of) \ _" - (SML "(fn '_ => Const (\"arbitrary\", dummyT))") - -instantiation "fun" :: (eq, "{type, random}") random -begin - -definition - "random n = random_fun_aux (op =) (random n) split_seed" - -instance .. - -end - -ML {* val f = Quickcheck.compile_generator_expr @{theory} - @{term "\f k. int (f k) = k"} [@{typ "int \ nat"}, @{typ int}] *} - -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} -ML {* f 20 |> (Option.map o map) (Sign.string_of_term @{theory}) *} - -end