diff -r 07f53494cf20 -r ab8c54355f2e src/HOL/ex/Quickcheck.thy --- a/src/HOL/ex/Quickcheck.thy Fri Feb 06 13:43:19 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,413 +0,0 @@ -(* Author: Florian Haftmann, TU Muenchen *) - -header {* A simple counterexample generator *} - -theory Quickcheck -imports Random Code_Eval Map -begin - -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 _ = return (TYPE('a), \u. Code_Eval.Const (STR ''TYPE'') TYPEREP('a))" - -instance .. - -end - -text {* Type @{typ "'a \ 'b"} *} - -ML {* -structure Random_Engine = -struct - -open Random_Engine; - -fun random_fun (T1 : typ) (T2 : typ) (eq : 'a -> 'a -> bool) (term_of : 'a -> term) - (random : Random_Engine.seed -> ('b * (unit -> term)) * 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', [], Const (@{const_name undefined}, T1 --> T2)); - val fun_upd = Const (@{const_name fun_upd}, - (T1 --> T2) --> T1 --> T2 --> T1 --> T2); - fun random_fun' x = - let - val (seed, fun_map, f_t) = ! state; - in case AList.lookup (uncurry eq) fun_map x - of SOME y => y - | NONE => let - val t1 = term_of x; - val ((y, t2), seed') = random seed; - val fun_map' = (x, y) :: fun_map; - val f_t' = fun_upd $ f_t $ t1 $ t2 (); - val _ = state := (seed', fun_map', f_t'); - in y end - end; - fun term_fun' () = #3 (! state); - in ((random_fun', term_fun'), seed'') end; - -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" - -code_const random_fun_aux (SML "Random'_Engine.random'_fun") - -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" - -instance .. - -end - -code_reserved SML Random_Engine - -text {* Datatypes *} - -definition - collapse :: "('a \ ('a \ 'b \ 'a) \ 'a) \ 'a \ 'b \ 'a" where - "collapse f = (do g \ f; g done)" - -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 return}, 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; -*} - -lemma random'_if: - fixes random' :: "index \ index \ seed \ ('a \ (unit \ term)) \ seed" - assumes "random' 0 j = (\s. 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) - -setup {* -let - exception REC of string; - exception TYP of string; - fun mk_collapse thy ty = Sign.mk_const thy - (@{const_name collapse}, [@{typ seed}, ty]); - fun term_ty ty = HOLogic.mk_prodT (ty, @{typ "unit \ term"}); - fun mk_split thy ty ty' = Sign.mk_const thy - (@{const_name split}, [ty, @{typ "unit \ term"}, StateMonad.liftT (term_ty ty') @{typ seed}]); - fun mk_scomp_split thy ty ty' t t' = - StateMonad.scomp (term_ty ty) (term_ty ty') @{typ seed} t - (mk_split thy ty ty' $ Abs ("", ty, Abs ("", @{typ "unit \ term"}, t'))) - fun mk_cons thy this_ty (c, args) = - let - val tys = map (fst o fst) args; - val c_ty = tys ---> this_ty; - val c = Const (c, tys ---> this_ty); - val t_indices = map (curry ( op * ) 2) (length tys - 1 downto 0); - val c_indices = map (curry ( op + ) 1) t_indices; - val c_t = list_comb (c, map Bound c_indices); - val t_t = Abs ("", @{typ unit}, Eval.mk_term Free Typerep.typerep - (list_comb (c, map (fn k => Bound (k + 1)) t_indices)) - |> map_aterms (fn t as Bound _ => t $ @{term "()"} | t => t)); - val return = StateMonad.return (term_ty this_ty) @{typ seed} - (HOLogic.mk_prod (c_t, t_t)); - val t = fold_rev (fn ((ty, _), random) => - mk_scomp_split thy ty this_ty random) - args return; - val is_rec = exists (snd o fst) args; - in (is_rec, t) end; - fun mk_conss thy ty [] = NONE - | mk_conss thy ty [(_, t)] = SOME t - | mk_conss thy ty ts = SOME (mk_collapse thy (term_ty ty) $ - (Sign.mk_const thy (@{const_name select}, [StateMonad.liftT (term_ty ty) @{typ seed}]) $ - HOLogic.mk_list (StateMonad.liftT (term_ty ty) @{typ seed}) (map snd ts))); - fun mk_clauses thy ty (tyco, (ts_rec, ts_atom)) = - let - val SOME t_atom = mk_conss thy ty ts_atom; - in case mk_conss thy ty ts_rec - of SOME t_rec => mk_collapse thy (term_ty ty) $ - (Sign.mk_const thy (@{const_name select_default}, [StateMonad.liftT (term_ty ty) @{typ seed}]) $ - @{term "i\index"} $ t_rec $ t_atom) - | NONE => t_atom - end; - fun mk_random_eqs thy vs tycos = - let - val this_ty = Type (hd tycos, map TFree vs); - val this_ty' = StateMonad.liftT (term_ty this_ty) @{typ seed}; - val random_name = NameSpace.base @{const_name random}; - val random'_name = random_name ^ "_" ^ Class.type_name (hd tycos) ^ "'"; - fun random ty = Sign.mk_const thy (@{const_name random}, [ty]); - val random' = Free (random'_name, - @{typ index} --> @{typ index} --> this_ty'); - fun atom ty = if Sign.of_sort thy (ty, @{sort random}) - then ((ty, false), random ty $ @{term "j\index"}) - else raise TYP - ("Will not generate random elements for type(s) " ^ quote (hd tycos)); - fun dtyp tyco = ((this_ty, true), random' $ @{term "i\index"} $ @{term "j\index"}); - fun rtyp tyco tys = raise REC - ("Will not generate random elements for mutual recursive type " ^ quote (hd tycos)); - val rhss = DatatypePackage.construction_interpretation thy - { atom = atom, dtyp = dtyp, rtyp = rtyp } vs tycos - |> (map o apsnd o map) (mk_cons thy this_ty) - |> (map o apsnd) (List.partition fst) - |> map (mk_clauses thy this_ty) - val eqss = map ((apsnd o map) (HOLogic.mk_Trueprop o HOLogic.mk_eq) o (fn rhs => ((this_ty, random'), [ - (random' $ @{term "0\index"} $ @{term "j\index"}, Abs ("s", @{typ seed}, - Const (@{const_name undefined}, HOLogic.mk_prodT (term_ty this_ty, @{typ seed})))), - (random' $ @{term "Suc_index i"} $ @{term "j\index"}, rhs) - ]))) rhss; - in eqss end; - fun random_inst [tyco] thy = - let - 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 ((this_ty, random'), eqs') = singleton (mk_random_eqs thy vs) tyco; - val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) - (Sign.mk_const thy (@{const_name random}, [this_ty]) $ @{term "i\index"}, - random' $ @{term "i\index"} $ @{term "i\index"}) - val del_func = Attrib.internal (fn _ => Thm.declaration_attribute - (fn thm => Context.mapping (Code.del_eqn thm) I)); - fun add_code simps lthy = - let - val thy = ProofContext.theory_of lthy; - val thm = @{thm random'_if} - |> Drule.instantiate' [SOME (Thm.ctyp_of thy this_ty)] [SOME (Thm.cterm_of thy random')] - |> (fn thm => thm OF simps) - |> singleton (ProofContext.export lthy (ProofContext.init thy)); - val c = (fst o dest_Const o fst o strip_comb o fst - o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm; - in - lthy - |> LocalTheory.theory (Code.del_eqns c - #> PureThy.add_thm ((Binding.name (fst (dest_Free random') ^ "_code"), thm), [Thm.kind_internal]) - #-> Code.add_eqn) - end; - in - thy - |> TheoryTarget.instantiation ([tyco], vs, @{sort random}) - |> PrimrecPackage.add_primrec - [(Binding.name (fst (dest_Free random')), SOME (snd (dest_Free random')), NoSyn)] - (map (fn eq => ((Binding.empty, [del_func]), eq)) eqs') - |-> add_code - |> `(fn lthy => Syntax.check_term lthy eq) - |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq))) - |> snd - |> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) - |> LocalTheory.exit_global - end - | random_inst tycos thy = raise REC - ("Will not generate random elements for mutual recursive type(s) " ^ commas (map quote tycos)); - fun add_random_inst tycos thy = random_inst tycos thy - handle REC msg => (warning msg; thy) - | TYP msg => (warning msg; thy) -in DatatypePackage.interpretation add_random_inst end -*} - -text {* Type @{typ int} *} - -instantiation int :: random -begin - -definition - "random n = (do - (b, _) \ random n; - (m, t) \ random n; - return (if b then (int m, \u. Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \ int)) (t ())) - else (- int m, \u. Code_Eval.App (Code_Eval.Const (STR ''HOL.uminus_class.uminus'') TYPEREP(int \ int)) - (Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \ int)) (t ())))) - done)" - -instance .. - -end - - -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 thy prop tys = - let - val bound_max = length tys - 1; - val bounds = map_index (fn (i, ty) => - (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) tys; - val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds); - 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"}; - fun mk_termtyp ty = HOLogic.mk_prodT (ty, @{typ "unit \ term"}); - fun mk_split ty = Sign.mk_const thy - (@{const_name split}, [ty, @{typ "unit \ term"}, StateMonad.liftT @{typ "term list option"} @{typ seed}]); - fun mk_scomp_split ty t t' = - StateMonad.scomp (mk_termtyp ty) @{typ "term list option"} @{typ seed} t (*FIXME*) - (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; - -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_term ("Quickcheck.eval_ref", eval_ref) thy t' []; - in f #> Random_Engine.run #> (Option.map o map) (Code.postprocess_term thy) end; - -end -*} - -setup {* - Quickcheck.add_generator ("code", Quickcheck.compile_generator_expr o ProofContext.theory_of) -*} - -subsection {* Examples *} - -theorem "map g (map f xs) = map (g o f) xs" - quickcheck [generator = code] - by (induct xs) simp_all - -theorem "map g (map f xs) = map (f o g) xs" - quickcheck [generator = code] - oops - -theorem "rev (xs @ ys) = rev ys @ rev xs" - quickcheck [generator = code] - by simp - -theorem "rev (xs @ ys) = rev xs @ rev ys" - quickcheck [generator = code] - oops - -theorem "rev (rev xs) = xs" - quickcheck [generator = code] - by simp - -theorem "rev xs = xs" - quickcheck [generator = code] - oops - -primrec app :: "('a \ 'a) list \ 'a \ 'a" where - "app [] x = x" - | "app (f # fs) x = app fs (f x)" - -lemma "app (fs @ gs) x = app gs (app fs x)" - quickcheck [generator = code] - by (induct fs arbitrary: x) simp_all - -lemma "app (fs @ gs) x = app fs (app gs x)" - quickcheck [generator = code] - oops - -primrec occurs :: "'a \ 'a list \ nat" where - "occurs a [] = 0" - | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)" - -primrec del1 :: "'a \ 'a list \ 'a list" where - "del1 a [] = []" - | "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))" - -lemma "Suc (occurs a (del1 a xs)) = occurs a xs" - -- {* Wrong. Precondition needed.*} - quickcheck [generator = code] - oops - -lemma "xs ~= [] \ Suc (occurs a (del1 a xs)) = occurs a xs" - quickcheck [generator = code] - -- {* Also wrong.*} - oops - -lemma "0 < occurs a xs \ Suc (occurs a (del1 a xs)) = occurs a xs" - quickcheck [generator = code] - by (induct xs) auto - -primrec replace :: "'a \ 'a \ 'a list \ 'a list" where - "replace a b [] = []" - | "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs)) - else (x#(replace a b xs)))" - -lemma "occurs a xs = occurs b (replace a b xs)" - quickcheck [generator = code] - -- {* Wrong. Precondition needed.*} - oops - -lemma "occurs b xs = 0 \ a=b \ occurs a xs = occurs b (replace a b xs)" - quickcheck [generator = code] - by (induct xs) simp_all - - -subsection {* Trees *} - -datatype 'a tree = Twig | Leaf 'a | Branch "'a tree" "'a tree" - -primrec leaves :: "'a tree \ 'a list" where - "leaves Twig = []" - | "leaves (Leaf a) = [a]" - | "leaves (Branch l r) = (leaves l) @ (leaves r)" - -primrec plant :: "'a list \ 'a tree" where - "plant [] = Twig " - | "plant (x#xs) = Branch (Leaf x) (plant xs)" - -primrec mirror :: "'a tree \ 'a tree" where - "mirror (Twig) = Twig " - | "mirror (Leaf a) = Leaf a " - | "mirror (Branch l r) = Branch (mirror r) (mirror l)" - -theorem "plant (rev (leaves xt)) = mirror xt" - quickcheck [generator = code] - --{* Wrong! *} - oops - -theorem "plant (leaves xt @ leaves yt) = Branch xt yt" - quickcheck [generator = code] - --{* Wrong! *} - oops - -datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree" - -primrec inOrder :: "'a ntree \ 'a list" where - "inOrder (Tip a)= [a]" - | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)" - -primrec root :: "'a ntree \ 'a" where - "root (Tip a) = a" - | "root (Node f x y) = f" - -theorem "hd (inOrder xt) = root xt" - quickcheck [generator = code] - --{* Wrong! *} - oops - -lemma "int (f k) = k" - quickcheck [generator = code] - oops - -end