# HG changeset patch # User bulwahn # Date 1292435114 -3600 # Node ID 5391d34b0f4c1107a25b870294234be8f9f37912 # Parent f4d3acf0c4cc0c489d91bb5ac2a724ec464afa50# Parent 5a9543f95cc62734e756a34fcf4ba16a0df96c0f merged diff -r 5a9543f95cc6 -r 5391d34b0f4c src/HOL/Smallcheck.thy --- a/src/HOL/Smallcheck.thy Wed Dec 15 18:20:44 2010 +0100 +++ b/src/HOL/Smallcheck.thy Wed Dec 15 18:45:14 2010 +0100 @@ -117,7 +117,6 @@ where "full_small_fun f d = full_small_fun' f d d" - instance .. end @@ -126,32 +125,46 @@ class check_all = enum + term_of + -fixes check_all :: "('a * (unit \ term) \ term list option) \ term list option" - + fixes check_all :: "('a * (unit \ term) \ term list option) \ term list option" + fixes enum_term_of :: "'a itself \ unit \ term list" + fun check_all_n_lists :: "(('a :: check_all) list * (unit \ term list) \ term list option) \ code_numeral \ term list option" where "check_all_n_lists f n = (if n = 0 then f ([], (%_. [])) else check_all (%(x, xt). check_all_n_lists (%(xs, xst). f ((x # xs), (%_. (xt () # xst ())))) (n - 1)))" -instantiation "fun" :: ("{equal, check_all}", check_all) check_all -begin - -definition mk_map_term :: "'a list \ (unit \ term list) \ (unit \ typerep) \ unit \ term" +definition mk_map_term :: " (unit \ typerep) \ (unit \ typerep) \ (unit \ term list) \ (unit \ term list) \ unit \ term" where - "mk_map_term domm rng T2 = - (%_. let T1 = (Typerep.typerep (TYPE('a))); + "mk_map_term T1 T2 domm rng = + (%_. let T1 = T1 (); T2 = T2 (); update_term = (%g (a, b). Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.Const (STR ''Fun.fun_upd'') (Typerep.Typerep (STR ''fun'') [Typerep.Typerep (STR ''fun'') [T1, T2], - Typerep.Typerep (STR ''fun'') [T1, Typerep.Typerep (STR ''fun'') [T2, Typerep.Typerep (STR ''fun'') [T1, T2]]]])) g) (Code_Evaluation.term_of a)) b) + Typerep.Typerep (STR ''fun'') [T1, + Typerep.Typerep (STR ''fun'') [T2, Typerep.Typerep (STR ''fun'') [T1, T2]]]])) + g) a) b) in - List.foldl update_term (Code_Evaluation.Abs (STR ''x'') T1 (Code_Evaluation.Const (STR ''HOL.undefined'') T2)) (zip domm (rng ())))" + List.foldl update_term (Code_Evaluation.Abs (STR ''x'') T1 (Code_Evaluation.Const (STR ''HOL.undefined'') T2)) (zip (domm ()) (rng ())))" + +instantiation "fun" :: ("{equal, check_all}", check_all) check_all +begin definition - "check_all f = check_all_n_lists (\(ys, yst). f (the o map_of (zip (Enum.enum\'a list) ys), mk_map_term (Enum.enum::'a list) yst (%_. Typerep.typerep (TYPE('b))))) (Code_Numeral.of_nat (length (Enum.enum :: 'a list)))" + "check_all f = + (let + mk_term = mk_map_term (%_. Typerep.typerep (TYPE('a))) (%_. Typerep.typerep (TYPE('b))) (enum_term_of (TYPE('a))); + enum = (Enum.enum :: 'a list) + in check_all_n_lists (\(ys, yst). f (the o map_of (zip enum ys), mk_term yst)) (Code_Numeral.of_nat (length enum)))" +definition enum_term_of_fun :: "('a => 'b) itself => unit => term list" +where + "enum_term_of_fun = (%_ _. let + enum_term_of_a = enum_term_of (TYPE('a)); + mk_term = mk_map_term (%_. Typerep.typerep (TYPE('a))) (%_. Typerep.typerep (TYPE('b))) enum_term_of_a + in map (%ys. mk_term (%_. ys) ()) (Enum.n_lists (length (enum_term_of_a ())) (enum_term_of (TYPE('b)) ())))" + instance .. end @@ -163,6 +176,10 @@ definition "check_all f = f (Code_Evaluation.valtermify ())" +definition enum_term_of_unit :: "unit itself => unit => term list" +where + "enum_term_of_unit = (%_ _. [Code_Evaluation.term_of ()])" + instance .. end @@ -174,6 +191,10 @@ definition "check_all f = (case f (Code_Evaluation.valtermify False) of Some x' \ Some x' | None \ f (Code_Evaluation.valtermify True))" +definition enum_term_of_bool :: "bool itself => unit => term list" +where + "enum_term_of_bool = (%_ _. map Code_Evaluation.term_of (Enum.enum :: bool list))" + instance .. end @@ -185,6 +206,10 @@ definition "check_all f = check_all (%(x, t1). check_all (%(y, t2). f ((x, y), %_. Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.term_of (Pair :: 'a => 'b => ('a * 'b))) (t1 ())) (t2 ()))))" +definition enum_term_of_prod :: "('a * 'b) itself => unit => term list" +where + "enum_term_of_prod = (%_ _. map (%(x, y). Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.term_of (Pair :: 'a => 'b => ('a * 'b))) x) y) (Enum.product (enum_term_of (TYPE('a)) ()) (enum_term_of (TYPE('b)) ())))" + instance .. end @@ -197,6 +222,11 @@ "check_all f = (case check_all (%(a, t). f (Inl a, %_. Code_Evaluation.App (Code_Evaluation.term_of (Inl :: 'a => 'a + 'b)) (t ()))) of Some x' => Some x' | None => check_all (%(b, t). f (Inr b, %_. Code_Evaluation.App (Code_Evaluation.term_of (Inr :: 'b => 'a + 'b)) (t ()))))" +definition enum_term_of_sum :: "('a + 'b) itself => unit => term list" +where + "enum_term_of_sum = (%_ _. map (Code_Evaluation.App (Code_Evaluation.term_of (Inl :: 'a => ('a + 'b)))) (enum_term_of (TYPE('a)) ()) @ + map (Code_Evaluation.App (Code_Evaluation.term_of (Inr :: 'b => ('a + 'b)))) (enum_term_of (TYPE('b)) ()))" + instance .. end @@ -223,6 +253,10 @@ f (Code_Evaluation.valtermify NibbleE) orelse f (Code_Evaluation.valtermify NibbleF)" +definition enum_term_of_nibble :: "nibble itself => unit => term list" +where + "enum_term_of_nibble = (%_ _. map Code_Evaluation.term_of (Enum.enum :: nibble list))" + instance .. end @@ -234,6 +268,10 @@ definition "check_all f = check_all (%(x, t1). check_all (%(y, t2). f (Char x y, %_. Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.term_of Char) (t1 ())) (t2 ()))))" +definition enum_term_of_char :: "char itself => unit => term list" +where + "enum_term_of_char = (%_ _. map Code_Evaluation.term_of (Enum.enum :: char list))" + instance .. end @@ -243,7 +281,13 @@ begin definition - "check_all f = f (Code_Evaluation.valtermify (None :: 'a option)) orelse check_all (%(x, t). f (Some x, %_. Code_Evaluation.App (Code_Evaluation.term_of (Some :: 'a => 'a option)) (t ())))" + "check_all f = f (Code_Evaluation.valtermify (None :: 'a option)) orelse check_all (%(x, t). f (Some x, %_. Code_Evaluation.App + (Code_Evaluation.Const (STR ''Option.option.Some'') + (Typerep.Typerep (STR ''fun'') [Typerep.typerep TYPE('a), Typerep.Typerep (STR ''Option.option'') [Typerep.typerep TYPE('a)]])) (t ())))" + +definition enum_term_of_option :: "'a option itself => unit => term list" +where + "enum_term_of_option = (% _ _. (Code_Evaluation.term_of (None :: 'a option)) # (map (Code_Evaluation.App (Code_Evaluation.term_of (Some :: 'a => 'a option))) (enum_term_of (TYPE('a)) ())))" instance .. @@ -256,6 +300,10 @@ definition "check_all f = f (Code_Evaluation.valtermify Enum.finite_1.a\<^isub>1)" +definition enum_term_of_finite_1 :: "Enum.finite_1 itself => unit => term list" +where + "enum_term_of_finite_1 = (%_ _. [Code_Evaluation.term_of Enum.finite_1.a\<^isub>1])" + instance .. end @@ -266,6 +314,10 @@ definition "check_all f = (case f (Code_Evaluation.valtermify Enum.finite_2.a\<^isub>1) of Some x' \ Some x' | None \ f (Code_Evaluation.valtermify Enum.finite_2.a\<^isub>2))" +definition enum_term_of_finite_2 :: "Enum.finite_2 itself => unit => term list" +where + "enum_term_of_finite_2 = (%_ _. map Code_Evaluation.term_of (Enum.enum :: Enum.finite_2 list))" + instance .. end @@ -276,6 +328,10 @@ definition "check_all f = (case f (Code_Evaluation.valtermify Enum.finite_3.a\<^isub>1) of Some x' \ Some x' | None \ (case f (Code_Evaluation.valtermify Enum.finite_3.a\<^isub>2) of Some x' \ Some x' | None \ f (Code_Evaluation.valtermify Enum.finite_3.a\<^isub>3)))" +definition enum_term_of_finite_3 :: "Enum.finite_3 itself => unit => term list" +where + "enum_term_of_finite_3 = (%_ _. map Code_Evaluation.term_of (Enum.enum :: Enum.finite_3 list))" + instance .. end diff -r 5a9543f95cc6 -r 5391d34b0f4c src/HOL/Tools/smallvalue_generators.ML --- a/src/HOL/Tools/smallvalue_generators.ML Wed Dec 15 18:20:44 2010 +0100 +++ b/src/HOL/Tools/smallvalue_generators.ML Wed Dec 15 18:45:14 2010 +0100 @@ -280,6 +280,52 @@ $ absdummy (T, absdummy (@{typ "unit => term"}, t))) $ Bound i in Abs ("d", @{typ code_numeral}, fold_rev mk_small_closure bounds check) end +(** post-processing of function terms **) + +fun dest_fun_upd (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) = (t0, (t1, t2)) + | dest_fun_upd t = raise TERM ("dest_fun_upd", [t]) + +fun mk_fun_upd T1 T2 (t1, t2) t = + Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2) $ t $ t1 $ t2 + +fun dest_plain_fun t = + case try dest_fun_upd t of + NONE => + (case t of + (Abs (_, _, Const (@{const_name HOL.undefined}, _))) => [] + | _ => raise TERM ("dest_plain_fun", [t])) + | SOME (t0, (t1, t2)) => (t1, t2) :: dest_plain_fun t0 + +fun make_plain_fun T1 T2 tps = + fold_rev (mk_fun_upd T1 T2) tps (absdummy (T1, Const ("_", T2))) + +fun make_set T1 [] = Const (@{const_abbrev Set.empty}, T1 --> @{typ bool}) + | make_set T1 ((_, @{const False}) :: tps) = make_set T1 tps + | make_set T1 ((t1, @{const True}) :: tps) = + Const (@{const_name insert}, T1 --> (T1 --> @{typ bool}) --> T1 --> @{typ bool}) + $ t1 $ (make_set T1 tps) + | make_set T1 ((_, t) :: tps) = raise TERM ("make_set", [t]) + +fun make_map T1 T2 [] = Const (@{const_abbrev Map.empty}, T1 --> T2) + | make_map T1 T2 ((_, Const (@{const_name None}, _)) :: tps) = make_map T1 T2 tps + | make_map T1 T2 ((t1, t2) :: tps) = mk_fun_upd T1 T2 (t1, t2) (make_map T1 T2 tps) + +fun post_process_term t = + case fastype_of t of + Type (@{type_name fun}, [T1, T2]) => + (case try dest_plain_fun t of + SOME tps => + tps + |> map (pairself post_process_term) + |> (case T2 of + @{typ bool} => rev #> make_set T1 + | Type (@{type_name option}, _) => make_map T1 T2 + | _ => make_plain_fun T1 T2) + | NONE => t) + | _ => t + +(** generator compiliation **) + fun compile_generator_expr ctxt t = let val thy = ProofContext.theory_of ctxt @@ -289,7 +335,9 @@ val compile = Code_Runtime.dynamic_value_strict (Counterexample.get, put_counterexample, "Smallvalue_Generators.put_counterexample") thy (SOME target) (fn proc => fn g => g #> (Option.map o map) proc) t' []; - in fn size => rpair NONE (compile size) end; + in + fn size => rpair NONE (compile size |> Option.map (map post_process_term)) + end; (** setup **)