--- a/src/HOL/Extraction/Higman.thy Thu Feb 05 11:49:15 2009 +0100
+++ b/src/HOL/Extraction/Higman.thy Thu Feb 05 14:50:43 2009 +0100
@@ -1,5 +1,4 @@
(* Title: HOL/Extraction/Higman.thy
- ID: $Id$
Author: Stefan Berghofer, TU Muenchen
Monika Seisenberger, LMU Muenchen
*)
@@ -7,7 +6,7 @@
header {* Higman's lemma *}
theory Higman
-imports Main "~~/src/HOL/ex/Random"
+imports Main State_Monad Random
begin
text {*
--- a/src/HOL/IsaMakefile Thu Feb 05 11:49:15 2009 +0100
+++ b/src/HOL/IsaMakefile Thu Feb 05 14:50:43 2009 +0100
@@ -311,6 +311,7 @@
HOL-Library: HOL $(LOG)/HOL-Library.gz
+
$(LOG)/HOL-Library.gz: $(OUT)/HOL Library/SetsAndFunctions.thy \
Library/Abstract_Rat.thy \
Library/BigO.thy Library/ContNotDenum.thy Library/Efficient_Nat.thy \
@@ -335,7 +336,8 @@
Library/Mapping.thy Library/Numeral_Type.thy Library/Reflection.thy \
Library/Boolean_Algebra.thy Library/Countable.thy \
Library/RBT.thy Library/Univ_Poly.thy \
- Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML \
+ Library/Random.thy Library/Quickcheck.thy \
+ Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML $(SRC)/HOL/Tools/float_arith.ML \
Library/reify_data.ML Library/reflection.ML
@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
@@ -815,7 +817,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/Coherent.thy ex/Dense_Linear_Order_Ex.thy ex/Eval_Examples.thy \
- ex/Groebner_Examples.thy ex/Random.thy ex/Quickcheck.thy \
+ ex/Groebner_Examples.thy ex/Quickcheck_Generators.thy \
ex/Codegenerator.thy ex/Codegenerator_Pretty.thy \
ex/CodegenSML_Test.thy ex/Formal_Power_Series_Examples.thy \
ex/Commutative_RingEx.thy ex/Efficient_Nat_examples.thy \
--- a/src/HOL/Library/Library.thy Thu Feb 05 11:49:15 2009 +0100
+++ b/src/HOL/Library/Library.thy Thu Feb 05 14:50:43 2009 +0100
@@ -34,9 +34,11 @@
Permutation
Pocklington
Primes
+ Quickcheck
Quicksort
Quotient
Ramsey
+ Random
Reflection
RBT
State_Monad
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Quickcheck.thy Thu Feb 05 14:50:43 2009 +0100
@@ -0,0 +1,85 @@
+(* 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 \<Rightarrow> seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> seed"
+
+text {* Type @{typ "'a itself"} *}
+
+instantiation itself :: ("{type, typerep}") random
+begin
+
+definition
+ "random _ = return (TYPE('a), \<lambda>u. Code_Eval.Const (STR ''TYPE'') TYPEREP('a))"
+
+instance ..
+
+end
+
+
+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
+
+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 \<Colon> bool \<Rightarrow> term list option \<Rightarrow> term list option \<Rightarrow> term list option"}
+ $ result $ @{term "None \<Colon> term list option"} $ (@{term "Some \<Colon> term list \<Rightarrow> term list option "} $ terms);
+ val return = @{term "Pair \<Colon> term list option \<Rightarrow> seed \<Rightarrow> term list option \<times> seed"};
+ fun mk_termtyp ty = HOLogic.mk_prodT (ty, @{typ "unit \<Rightarrow> term"});
+ fun mk_split ty = Sign.mk_const thy
+ (@{const_name split}, [ty, @{typ "unit \<Rightarrow> 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 \<Rightarrow> 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)
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Random.thy Thu Feb 05 14:50:43 2009 +0100
@@ -0,0 +1,182 @@
+(* Author: Florian Haftmann, TU Muenchen
+*)
+
+header {* A HOL random engine *}
+
+theory Random
+imports State_Monad Code_Index
+begin
+
+subsection {* Auxiliary functions *}
+
+definition
+ inc_shift :: "index \<Rightarrow> index \<Rightarrow> index"
+where
+ "inc_shift v k = (if v = k then 1 else k + 1)"
+
+definition
+ minus_shift :: "index \<Rightarrow> index \<Rightarrow> index \<Rightarrow> index"
+where
+ "minus_shift r k l = (if k < l then r + k - l else k - l)"
+
+fun
+ log :: "index \<Rightarrow> index \<Rightarrow> index"
+where
+ "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
+
+subsection {* Random seeds *}
+
+types seed = "index \<times> index"
+
+primrec
+ "next" :: "seed \<Rightarrow> index \<times> seed"
+where
+ "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) \<noteq> 0"
+apply (cases s)
+apply (auto simp add: minus_shift_def Let_def)
+done
+
+primrec
+ seed_invariant :: "seed \<Rightarrow> bool"
+where
+ "seed_invariant (v, w) \<longleftrightarrow> 0 < v \<and> v < 9438322952 \<and> 0 < w \<and> True"
+
+lemma if_same:
+ "(if b then f x else f y) = f (if b then x else y)"
+ by (cases b) simp_all
+
+definition
+ split_seed :: "seed \<Rightarrow> seed \<times> seed"
+where
+ "split_seed s = (let
+ (v, w) = s;
+ (v', w') = snd (next s);
+ v'' = inc_shift 2147483562 v;
+ s'' = (v'', w');
+ w'' = inc_shift 2147483398 w;
+ s''' = (v', w'')
+ in (s'', s'''))"
+
+
+subsection {* Base selectors *}
+
+function
+ range_aux :: "index \<Rightarrow> index \<Rightarrow> seed \<Rightarrow> index \<times> seed"
+where
+ "range_aux k l s = (if k = 0 then (l, s) else
+ let (v, s') = next s
+ in range_aux (k - 1) (v + l * 2147483561) s')"
+by pat_completeness auto
+termination
+ by (relation "measure (nat_of_index o fst)")
+ (auto simp add: index)
+
+definition
+ range :: "index \<Rightarrow> seed \<Rightarrow> index \<times> seed"
+where
+ "range k = (do
+ v \<leftarrow> 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: monad_collapse range_def del: range_aux.simps log.simps)
+qed
+
+definition
+ select :: "'a list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed"
+where
+ "select xs = (do
+ k \<leftarrow> range (index_of_nat (length xs));
+ return (nth xs (nat_of_index k))
+ done)"
+
+lemma select:
+ assumes "xs \<noteq> []"
+ shows "fst (select xs s) \<in> 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: monad_collapse select_def)
+qed
+
+definition
+ select_default :: "index \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> seed \<Rightarrow> 'a \<times> seed"
+where
+ [code del]: "select_default k x y = (do
+ l \<leftarrow> 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: monad_collapse select_default_def)
+
+lemma select_default_code [code]:
+ "select_default k x y = (if k = 0 then do
+ _ \<leftarrow> range 1;
+ return y
+ done else do
+ l \<leftarrow> 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
+ case True then show ?thesis
+ by (simp add: monad_collapse select_default_def range_def)
+qed
+
+
+subsection {* @{text ML} interface *}
+
+ML {*
+structure Random_Engine =
+struct
+
+type seed = int * int;
+
+local
+
+val seed = ref
+ (let
+ val now = Time.toMilliseconds (Time.now ());
+ val (q, s1) = IntInf.divMod (now, 2147483562);
+ val s2 = q mod 2147483398;
+ in (s1 + 1, s2 + 1) end);
+
+in
+
+fun run f =
+ let
+ val (x, seed') = f (! seed);
+ val _ = seed := seed'
+ in x end;
+
+end;
+
+end;
+*}
+
+end
+
--- a/src/HOL/Tools/primrec_package.ML Thu Feb 05 11:49:15 2009 +0100
+++ b/src/HOL/Tools/primrec_package.ML Thu Feb 05 14:50:43 2009 +0100
@@ -247,11 +247,11 @@
val _ = if gen_eq_set (op =) (names1, names2) then ()
else primrec_error ("functions " ^ commas_quote names2 ^
"\nare not mutually recursive");
- val qualify = Binding.qualify
- (space_implode "_" (map (Sign.base_name o #1) defs));
- val spec' = (map o apfst o apfst) qualify spec;
- val simp_atts = map (Attrib.internal o K)
- [Simplifier.simp_add, Code.add_default_eqn_attribute];
+ val prefix = space_implode "_" (map (Sign.base_name o #1) defs);
+ val qualify = Binding.qualify prefix;
+ val spec' = (map o apfst)
+ (fn (b, attrs) => (qualify b, Code.add_default_eqn_attrib :: attrs)) spec;
+ val simp_att = (Attrib.internal o K) Simplifier.simp_add;
in
lthy
|> set_group ? LocalTheory.set_group (serial_string ())
@@ -259,7 +259,7 @@
|-> (fn defs => `(fn ctxt => prove_spec ctxt names1 rec_rewrites defs spec'))
|-> (fn simps => fold_map (LocalTheory.note Thm.theoremK) simps)
|-> (fn simps' => LocalTheory.note Thm.theoremK
- ((qualify (Binding.name "simps"), simp_atts), maps snd simps'))
+ ((qualify (Binding.name "simps"), [simp_att]), maps snd simps'))
|>> snd
end handle PrimrecError (msg, some_eqn) =>
error ("Primrec definition error:\n" ^ msg ^ (case some_eqn
--- a/src/HOL/ex/Quickcheck.thy Thu Feb 05 11:49:15 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 \<Rightarrow> seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> seed"
-
-text {* Type @{typ "'a itself"} *}
-
-instantiation itself :: ("{type, typerep}") random
-begin
-
-definition
- "random _ = return (TYPE('a), \<lambda>u. Code_Eval.Const (STR ''TYPE'') TYPEREP('a))"
-
-instance ..
-
-end
-
-text {* Type @{typ "'a \<Rightarrow> '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 \<Rightarrow> typerep \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> term)
- \<Rightarrow> (seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> seed) \<Rightarrow> (seed \<Rightarrow> seed \<times> seed)
- \<Rightarrow> seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> seed"
-
-code_const random_fun_aux (SML "Random'_Engine.random'_fun")
-
-instantiation "fun" :: ("{eq, term_of}", "{type, random}") random
-begin
-
-definition random_fun :: "index \<Rightarrow> seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> 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 \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
- "collapse f = (do g \<leftarrow> 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 \<Rightarrow> index \<Rightarrow> seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> seed"
- assumes "random' 0 j = (\<lambda>s. undefined)"
- and "\<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 \<Rightarrow> term"});
- fun mk_split thy ty ty' = Sign.mk_const thy
- (@{const_name split}, [ty, @{typ "unit \<Rightarrow> 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 \<Rightarrow> 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\<Colon>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\<Colon>index"})
- else raise TYP
- ("Will not generate random elements for type(s) " ^ quote (hd tycos));
- fun dtyp tyco = ((this_ty, true), random' $ @{term "i\<Colon>index"} $ @{term "j\<Colon>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\<Colon>index"} $ @{term "j\<Colon>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\<Colon>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\<Colon>index"},
- random' $ @{term "i\<Colon>index"} $ @{term "i\<Colon>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, _) \<leftarrow> random n;
- (m, t) \<leftarrow> random n;
- return (if b then (int m, \<lambda>u. Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \<Rightarrow> int)) (t ()))
- else (- int m, \<lambda>u. Code_Eval.App (Code_Eval.Const (STR ''HOL.uminus_class.uminus'') TYPEREP(int \<Rightarrow> int))
- (Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \<Rightarrow> 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 \<Colon> bool \<Rightarrow> term list option \<Rightarrow> term list option \<Rightarrow> term list option"}
- $ result $ @{term "None \<Colon> term list option"} $ (@{term "Some \<Colon> term list \<Rightarrow> term list option "} $ terms);
- val return = @{term "Pair \<Colon> term list option \<Rightarrow> seed \<Rightarrow> term list option \<times> seed"};
- fun mk_termtyp ty = HOLogic.mk_prodT (ty, @{typ "unit \<Rightarrow> term"});
- fun mk_split ty = Sign.mk_const thy
- (@{const_name split}, [ty, @{typ "unit \<Rightarrow> 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 \<Rightarrow> 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 \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> '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 \<Rightarrow> 'a list \<Rightarrow> nat" where
- "occurs a [] = 0"
- | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)"
-
-primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> '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 ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
- quickcheck [generator = code]
- -- {* Also wrong.*}
- oops
-
-lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
- quickcheck [generator = code]
- by (induct xs) auto
-
-primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> '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 \<or> a=b \<longrightarrow> 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 \<Rightarrow> 'a list" where
- "leaves Twig = []"
- | "leaves (Leaf a) = [a]"
- | "leaves (Branch l r) = (leaves l) @ (leaves r)"
-
-primrec plant :: "'a list \<Rightarrow> 'a tree" where
- "plant [] = Twig "
- | "plant (x#xs) = Branch (Leaf x) (plant xs)"
-
-primrec mirror :: "'a tree \<Rightarrow> '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 \<Rightarrow> 'a list" where
- "inOrder (Tip a)= [a]"
- | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)"
-
-primrec root :: "'a ntree \<Rightarrow> '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
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Quickcheck_Generators.thy Thu Feb 05 14:50:43 2009 +0100
@@ -0,0 +1,353 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Experimental counterexample generators *}
+
+theory Quickcheck_Generators
+imports Quickcheck State_Monad
+begin
+
+subsection {* Type @{typ "'a \<Rightarrow> '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 \<Rightarrow> typerep \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> term)
+ \<Rightarrow> (seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> seed) \<Rightarrow> (seed \<Rightarrow> seed \<times> seed)
+ \<Rightarrow> seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> seed"
+
+code_const random_fun_aux (SML "Random'_Engine.random'_fun")
+
+instantiation "fun" :: ("{eq, term_of}", "{type, random}") random
+begin
+
+definition random_fun :: "index \<Rightarrow> seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> 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
+
+
+subsection {* Datatypes *}
+
+definition
+ collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
+ "collapse f = (do g \<leftarrow> 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 \<Rightarrow> index \<Rightarrow> seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> seed"
+ assumes "random' 0 j = (\<lambda>s. undefined)"
+ and "\<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 \<Rightarrow> term"});
+ fun mk_split thy ty ty' = Sign.mk_const thy
+ (@{const_name split}, [ty, @{typ "unit \<Rightarrow> 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 \<Rightarrow> 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\<Colon>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\<Colon>index"})
+ else raise TYP
+ ("Will not generate random elements for type(s) " ^ quote (hd tycos));
+ fun dtyp tyco = ((this_ty, true), random' $ @{term "i\<Colon>index"} $ @{term "j\<Colon>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\<Colon>index"} $ @{term "j\<Colon>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\<Colon>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\<Colon>index"},
+ random' $ @{term "i\<Colon>index"} $ @{term "i\<Colon>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
+*}
+
+
+subsection {* Type @{typ int} *}
+
+instantiation int :: random
+begin
+
+definition
+ "random n = (do
+ (b, _) \<leftarrow> random n;
+ (m, t) \<leftarrow> random n;
+ return (if b then (int m, \<lambda>u. Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \<Rightarrow> int)) (t ()))
+ else (- int m, \<lambda>u. Code_Eval.App (Code_Eval.Const (STR ''HOL.uminus_class.uminus'') TYPEREP(int \<Rightarrow> int))
+ (Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \<Rightarrow> int)) (t ()))))
+ done)"
+
+instance ..
+
+end
+
+
+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 \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> '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 \<Rightarrow> 'a list \<Rightarrow> nat" where
+ "occurs a [] = 0"
+ | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)"
+
+primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> '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 ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
+ quickcheck [generator = code]
+ -- {* Also wrong.*}
+ oops
+
+lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
+ quickcheck [generator = code]
+ by (induct xs) auto
+
+primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> '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 \<or> a=b \<longrightarrow> 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 \<Rightarrow> 'a list" where
+ "leaves Twig = []"
+ | "leaves (Leaf a) = [a]"
+ | "leaves (Branch l r) = (leaves l) @ (leaves r)"
+
+primrec plant :: "'a list \<Rightarrow> 'a tree" where
+ "plant [] = Twig "
+ | "plant (x#xs) = Branch (Leaf x) (plant xs)"
+
+primrec mirror :: "'a tree \<Rightarrow> '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 \<Rightarrow> 'a list" where
+ "inOrder (Tip a)= [a]"
+ | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)"
+
+primrec root :: "'a ntree \<Rightarrow> '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
--- a/src/HOL/ex/ROOT.ML Thu Feb 05 11:49:15 2009 +0100
+++ b/src/HOL/ex/ROOT.ML Thu Feb 05 14:50:43 2009 +0100
@@ -10,7 +10,7 @@
"FuncSet",
"Word",
"Eval_Examples",
- "Quickcheck",
+ "Quickcheck_Generators",
"Term_Of_Syntax",
"Codegenerator",
"Codegenerator_Pretty",
--- a/src/HOL/ex/Random.thy Thu Feb 05 11:49:15 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,183 +0,0 @@
-(* ID: $Id$
- Author: Florian Haftmann, TU Muenchen
-*)
-
-header {* A HOL random engine *}
-
-theory Random
-imports State_Monad Code_Index
-begin
-
-subsection {* Auxiliary functions *}
-
-definition
- inc_shift :: "index \<Rightarrow> index \<Rightarrow> index"
-where
- "inc_shift v k = (if v = k then 1 else k + 1)"
-
-definition
- minus_shift :: "index \<Rightarrow> index \<Rightarrow> index \<Rightarrow> index"
-where
- "minus_shift r k l = (if k < l then r + k - l else k - l)"
-
-fun
- log :: "index \<Rightarrow> index \<Rightarrow> index"
-where
- "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
-
-subsection {* Random seeds *}
-
-types seed = "index \<times> index"
-
-primrec
- "next" :: "seed \<Rightarrow> index \<times> seed"
-where
- "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) \<noteq> 0"
-apply (cases s)
-apply (auto simp add: minus_shift_def Let_def)
-done
-
-primrec
- seed_invariant :: "seed \<Rightarrow> bool"
-where
- "seed_invariant (v, w) \<longleftrightarrow> 0 < v \<and> v < 9438322952 \<and> 0 < w \<and> True"
-
-lemma if_same:
- "(if b then f x else f y) = f (if b then x else y)"
- by (cases b) simp_all
-
-definition
- split_seed :: "seed \<Rightarrow> seed \<times> seed"
-where
- "split_seed s = (let
- (v, w) = s;
- (v', w') = snd (next s);
- v'' = inc_shift 2147483562 v;
- s'' = (v'', w');
- w'' = inc_shift 2147483398 w;
- s''' = (v', w'')
- in (s'', s'''))"
-
-
-subsection {* Base selectors *}
-
-function
- range_aux :: "index \<Rightarrow> index \<Rightarrow> seed \<Rightarrow> index \<times> seed"
-where
- "range_aux k l s = (if k = 0 then (l, s) else
- let (v, s') = next s
- in range_aux (k - 1) (v + l * 2147483561) s')"
-by pat_completeness auto
-termination
- by (relation "measure (nat_of_index o fst)")
- (auto simp add: index)
-
-definition
- range :: "index \<Rightarrow> seed \<Rightarrow> index \<times> seed"
-where
- "range k = (do
- v \<leftarrow> 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: monad_collapse range_def del: range_aux.simps log.simps)
-qed
-
-definition
- select :: "'a list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed"
-where
- "select xs = (do
- k \<leftarrow> range (index_of_nat (length xs));
- return (nth xs (nat_of_index k))
- done)"
-
-lemma select:
- assumes "xs \<noteq> []"
- shows "fst (select xs s) \<in> 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: monad_collapse select_def)
-qed
-
-definition
- select_default :: "index \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> seed \<Rightarrow> 'a \<times> seed"
-where
- [code del]: "select_default k x y = (do
- l \<leftarrow> 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: monad_collapse select_default_def)
-
-lemma select_default_code [code]:
- "select_default k x y = (if k = 0 then do
- _ \<leftarrow> range 1;
- return y
- done else do
- l \<leftarrow> 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
- case True then show ?thesis
- by (simp add: monad_collapse select_default_def range_def)
-qed
-
-
-subsection {* @{text ML} interface *}
-
-ML {*
-structure Random_Engine =
-struct
-
-type seed = int * int;
-
-local
-
-val seed = ref
- (let
- val now = Time.toMilliseconds (Time.now ());
- val (q, s1) = IntInf.divMod (now, 2147483562);
- val s2 = q mod 2147483398;
- in (s1 + 1, s2 + 1) end);
-
-in
-
-fun run f =
- let
- val (x, seed') = f (! seed);
- val _ = seed := seed'
- in x end;
-
-end;
-
-end;
-*}
-
-end
-