# HG changeset patch # User haftmann # Date 1242398356 -7200 # Node ID ced8171602837304b205a20e12857efe0bb42841 # Parent 27afaaa6547a8dd31c4425b03b87991e366882d1 experimental addition of quickcheck diff -r 27afaaa6547a -r ced817160283 src/HOL/Complex_Main.thy --- a/src/HOL/Complex_Main.thy Fri May 15 16:39:15 2009 +0200 +++ b/src/HOL/Complex_Main.thy Fri May 15 16:39:16 2009 +0200 @@ -9,6 +9,7 @@ Ln Taylor Integration + "Library/Quickcheck" begin end diff -r 27afaaa6547a -r ced817160283 src/HOL/Library/Quickcheck.thy --- a/src/HOL/Library/Quickcheck.thy Fri May 15 16:39:15 2009 +0200 +++ b/src/HOL/Library/Quickcheck.thy Fri May 15 16:39:16 2009 +0200 @@ -3,43 +3,22 @@ header {* A simple counterexample generator *} theory Quickcheck -imports Random Code_Eval Map +imports Main Real Random begin +notation fcomp (infixl "o>" 60) +notation scomp (infixl "o\" 60) + + subsection {* The @{text random} class *} class random = typerep + - fixes random :: "index \ seed \ ('a \ (unit \ term)) \ seed" - -text {* Type @{typ "'a itself"} *} - -instantiation itself :: ("{type, typerep}") random -begin - -definition - "random _ = Pair (TYPE('a), \u. Code_Eval.Const (STR ''TYPE'') TYPEREP('a))" - -instance .. - -end + fixes random :: "index \ Random.seed \ ('a \ (unit \ term)) \ Random.seed" subsection {* Quickcheck generator *} ML {* -structure StateMonad = -struct - -fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); -fun liftT' sT = sT --> sT; - -fun return T sT x = Const (@{const_name Pair}, T --> liftT T sT) $ x; - -fun scomp T1 T2 sT f g = Const (@{const_name scomp}, - liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; - -end; - structure Quickcheck = struct @@ -58,24 +37,26 @@ val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) 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 "} $ terms); - val return = @{term "Pair \ term list option \ seed \ term list option \ seed"}; + val return = @{term "Pair \ term list option \ Random.seed \ term list option \ Random.seed"}; + fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); fun mk_termtyp ty = HOLogic.mk_prodT (ty, @{typ "unit \ term"}); + fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp}, + liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; fun mk_split ty = Sign.mk_const thy - (@{const_name split}, [ty, @{typ "unit \ term"}, StateMonad.liftT @{typ "term list option"} @{typ seed}]); + (@{const_name split}, [ty, @{typ "unit \ term"}, liftT @{typ "term list option"} @{typ Random.seed}]); fun mk_scomp_split ty t t' = - StateMonad.scomp (mk_termtyp ty) @{typ "term list option"} @{typ seed} t (*FIXME*) + mk_scomp (mk_termtyp ty) @{typ "term list option"} @{typ Random.seed} t (mk_split ty $ Abs ("", ty, Abs ("", @{typ "unit \ term"}, t'))); fun mk_bindclause (_, _, i, ty) = mk_scomp_split ty - (Sign.mk_const thy (@{const_name random}, [ty]) $ Bound i) - val t = fold_rev mk_bindclause bounds (return $ check); - in Abs ("n", @{typ index}, t) end; + (Sign.mk_const thy (@{const_name random}, [ty]) $ Bound i); + in Abs ("n", @{typ index}, fold_rev mk_bindclause bounds (return $ check)) end; fun compile_generator_expr thy t = let val tys = (map snd o fst o strip_abs) t; val t' = mk_generator_expr thy t tys; val f = Code_ML.eval (SOME target) ("Quickcheck.eval_ref", eval_ref) - (fn proc => fn g => fn s => g (s + 1) #>> (Option.map o map) proc) thy t' []; + (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' []; in f #> Random_Engine.run end; end @@ -87,7 +68,35 @@ *} -subsection {* Type @{typ "'a \ 'b"} *} +subsection {* Fundamental types*} + +definition (in term_syntax) + "termify_bool b = (if b then termify True else termify False)" + +instantiation bool :: random +begin + +definition + "random i = Random.range i o\ (\k. Pair (termify_bool (k div 2 = 0)))" + +instance .. + +end + +definition (in term_syntax) + "termify_itself TYPE('a\typerep) = termify TYPE('a)" + +instantiation itself :: (typerep) random +begin + +definition random_itself :: "index \ Random.seed \ ('a itself \ (unit \ term)) \ Random.seed" where + "random_itself _ = Pair (termify_itself TYPE('a))" + +instance .. + +end + +text {* Type @{typ "'a \ 'b"} *} ML {* structure Random_Engine = @@ -123,10 +132,9 @@ end *} -axiomatization - random_fun_aux :: "typerep \ typerep \ ('a \ 'a \ bool) \ ('a \ term) - \ (seed \ ('b \ (unit \ term)) \ seed) \ (seed \ seed \ seed) - \ seed \ (('a \ 'b) \ (unit \ term)) \ seed" +axiomatization random_fun_aux :: "typerep \ typerep \ ('a \ 'a \ bool) \ ('a \ term) + \ (Random.seed \ ('b \ (unit \ term)) \ Random.seed) \ (Random.seed \ Random.seed \ Random.seed) + \ Random.seed \ (('a \ 'b) \ (unit \ term)) \ Random.seed" code_const random_fun_aux (Quickcheck "Random'_Engine.random'_fun") -- {* With enough criminal energy this can be abused to derive @{prop False}; @@ -136,8 +144,8 @@ instantiation "fun" :: ("{eq, term_of}", "{type, random}") random begin -definition random_fun :: "index \ seed \ (('a \ 'b) \ (unit \ term)) \ seed" where - "random n = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of (random n) split_seed" +definition random_fun :: "index \ Random.seed \ (('a \ 'b) \ (unit \ term)) \ Random.seed" where + "random n = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of (random n) Random.split_seed" instance .. @@ -145,4 +153,80 @@ code_reserved Quickcheck Random_Engine + +subsection {* Numeric types *} + +function (in term_syntax) termify_numeral :: "index \ int \ (unit \ term)" where + "termify_numeral k = (if k = 0 then termify Int.Pls + else (if k mod 2 = 0 then termify Int.Bit0 else termify Int.Bit1) <\> termify_numeral (k div 2))" + by pat_completeness auto + +declare (in term_syntax) termify_numeral.psimps [simp del] + +termination termify_numeral by (relation "measure Code_Index.nat_of") + (simp_all add: index) + +definition (in term_syntax) termify_int_number :: "index \ int \ (unit \ term)" where + "termify_int_number k = termify number_of <\> termify_numeral k" + +definition (in term_syntax) termify_nat_number :: "index \ nat \ (unit \ term)" where + "termify_nat_number k = (nat \ number_of, snd (termify (number_of :: int \ nat))) <\> termify_numeral k" + +declare termify_nat_number_def [simplified snd_conv, code] + +instantiation nat :: random +begin + +definition random_nat :: "index \ Random.seed \ (nat \ (unit \ term)) \ Random.seed" where + "random_nat i = Random.range (i + 1) o\ (\k. Pair (termify_nat_number k))" + +instance .. + end + +definition (in term_syntax) term_uminus :: "int \ (unit \ term) \ int \ (unit \ term)" where + [code inline]: "term_uminus k = termify uminus <\> k" + +instantiation int :: random +begin + +definition + "random i = Random.range (2 * i + 1) o\ (\k. Pair (if k \ i + then let j = k - i in termify_int_number j + else let j = i - k in term_uminus (termify_int_number j)))" + +instance .. + +end + +definition (in term_syntax) term_fract :: "int \ (unit \ term) \ int \ (unit \ term) \ rat \ (unit \ term)" where + [code inline]: "term_fract k l = termify Fract <\> k <\> l" + +instantiation rat :: random +begin + +definition + "random i = random i o\ (\num. Random.range (i + 1) o\ (\denom. Pair (term_fract num (termify_int_number denom))))" + +instance .. + +end + +definition (in term_syntax) term_ratreal :: "rat \ (unit \ term) \ real \ (unit \ term)" where + [code inline]: "term_ratreal k = termify Ratreal <\> k" + +instantiation real :: random +begin + +definition + "random i = random i o\ (\r. Pair (term_ratreal r))" + +instance .. + +end + + +no_notation fcomp (infixl "o>" 60) +no_notation scomp (infixl "o\" 60) + +end