| author | berghofe | 
| Tue, 10 Jan 2012 23:59:37 +0100 | |
| changeset 46181 | 49c3e0ef9d70 | 
| parent 46042 | ab32a87ba01a | 
| child 46219 | 426ed18eba43 | 
| permissions | -rw-r--r-- | 
| 41930 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 bulwahn parents: 
41925diff
changeset | 1 | (* Title: HOL/Tools/Quickcheck/narrowing_generators.ML | 
| 41905 | 2 | Author: Lukas Bulwahn, TU Muenchen | 
| 3 | ||
| 41938 | 4 | Narrowing-based counterexample generation. | 
| 41905 | 5 | *) | 
| 6 | ||
| 41930 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 bulwahn parents: 
41925diff
changeset | 7 | signature NARROWING_GENERATORS = | 
| 41905 | 8 | sig | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 9 | val allow_existentials : bool Config.T | 
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42020diff
changeset | 10 | val finite_functions : bool Config.T | 
| 43079 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 11 | val overlord : bool Config.T | 
| 43891 
4690f76f327a
making active configuration public in narrowing-based quickcheck
 bulwahn parents: 
43878diff
changeset | 12 | val active : bool Config.T | 
| 45418 | 13 | val test_term: Proof.context -> term * term list -> Quickcheck.result | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 14 | datatype counterexample = Universal_Counterexample of (term * counterexample) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 15 | | Existential_Counterexample of (term * counterexample) list | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 16 | | Empty_Assignment | 
| 45725 
2987b29518aa
the narrowing also indicates if counterexample is potentially spurious
 bulwahn parents: 
45685diff
changeset | 17 | val put_counterexample: (unit -> (bool * term list) option) -> Proof.context -> Proof.context | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 18 | val put_existential_counterexample : (unit -> counterexample option) -> Proof.context -> Proof.context | 
| 41905 | 19 | val setup: theory -> theory | 
| 20 | end; | |
| 21 | ||
| 41930 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 bulwahn parents: 
41925diff
changeset | 22 | structure Narrowing_Generators : NARROWING_GENERATORS = | 
| 41905 | 23 | struct | 
| 24 | ||
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42020diff
changeset | 25 | (* configurations *) | 
| 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42020diff
changeset | 26 | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 27 | val allow_existentials = Attrib.setup_config_bool @{binding quickcheck_allow_existentials} (K true)
 | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42361diff
changeset | 28 | val finite_functions = Attrib.setup_config_bool @{binding quickcheck_finite_functions} (K true)
 | 
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 29 | val overlord = Attrib.setup_config_bool @{binding quickcheck_narrowing_overlord} (K false)
 | 
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42020diff
changeset | 30 | |
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 31 | (* partial_term_of instances *) | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 32 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 33 | fun mk_partial_term_of (x, T) = | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 34 |   Const (@{const_name Quickcheck_Narrowing.partial_term_of_class.partial_term_of},
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 35 |     Term.itselfT T --> @{typ narrowing_term} --> @{typ Code_Evaluation.term})
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 36 |       $ Const ("TYPE", Term.itselfT T) $ x
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 37 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 38 | (** formal definition **) | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 39 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 40 | fun add_partial_term_of tyco raw_vs thy = | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 41 | let | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 42 |     val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 43 | val ty = Type (tyco, map TFree vs); | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 44 |     val lhs = Const (@{const_name partial_term_of},
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 45 |         Term.itselfT ty --> @{typ narrowing_term} --> @{typ Code_Evaluation.term})
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 46 |       $ Free ("x", Term.itselfT ty) $ Free ("t", @{typ narrowing_term});
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 47 |     val rhs = @{term "undefined :: Code_Evaluation.term"};
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 48 | val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 49 | fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 50 | o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv"; | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 51 | in | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 52 | thy | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 53 |     |> Class.instantiation ([tyco], vs, @{sort partial_term_of})
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 54 | |> `(fn lthy => Syntax.check_term lthy eq) | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 55 | |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq))) | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 56 | |> snd | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 57 | |> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 58 | end; | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 59 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 60 | fun ensure_partial_term_of (tyco, (raw_vs, _)) thy = | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 61 | let | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 62 |     val need_inst = not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort partial_term_of})
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 63 |       andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 64 | in if need_inst then add_partial_term_of tyco raw_vs thy else thy end; | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 65 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 66 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 67 | (** code equations for datatypes **) | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 68 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 69 | fun mk_partial_term_of_eq thy ty (i, (c, (_, tys))) = | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 70 | let | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43317diff
changeset | 71 |     val frees = map Free (Name.invent_names Name.context "a" (map (K @{typ narrowing_term}) tys))
 | 
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 72 |     val narrowing_term = @{term "Quickcheck_Narrowing.Ctr"} $ HOLogic.mk_number @{typ code_int} i
 | 
| 43079 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 73 |       $ (HOLogic.mk_list @{typ narrowing_term} (rev frees))
 | 
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 74 |     val rhs = fold (fn u => fn t => @{term "Code_Evaluation.App"} $ t $ u)
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 75 | (map mk_partial_term_of (frees ~~ tys)) | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 76 |         (@{term "Code_Evaluation.Const"} $ HOLogic.mk_literal c $ HOLogic.mk_typerep (tys ---> ty))
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 77 | val insts = | 
| 45344 
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
 wenzelm parents: 
45159diff
changeset | 78 | map (SOME o Thm.cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) | 
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 79 |         [Free ("ty", Term.itselfT ty), narrowing_term, rhs]
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 80 | val cty = Thm.ctyp_of thy ty; | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 81 | in | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 82 |     @{thm partial_term_of_anything}
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 83 | |> Drule.instantiate' [SOME cty] insts | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 84 | |> Thm.varifyT_global | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 85 | end | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 86 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 87 | fun add_partial_term_of_code tyco raw_vs raw_cs thy = | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 88 | let | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 89 | val algebra = Sign.classes_of thy; | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 90 | val vs = map (fn (v, sort) => | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 91 |       (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 92 | val ty = Type (tyco, map TFree vs); | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 93 | val cs = (map o apsnd o apsnd o map o map_atyps) | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 94 | (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs; | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 95 |     val const = AxClass.param_of_inst thy (@{const_name partial_term_of}, tyco);
 | 
| 45344 
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
 wenzelm parents: 
45159diff
changeset | 96 | val var_insts = | 
| 
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
 wenzelm parents: 
45159diff
changeset | 97 | map (SOME o Thm.cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) | 
| 43079 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 98 |         [Free ("ty", Term.itselfT ty), @{term "Quickcheck_Narrowing.Var p tt"},
 | 
| 45344 
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
 wenzelm parents: 
45159diff
changeset | 99 |           @{term "Code_Evaluation.Free (STR ''_'')"} $ HOLogic.mk_typerep ty];
 | 
| 43079 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 100 | val var_eq = | 
| 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 101 |       @{thm partial_term_of_anything}
 | 
| 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 102 | |> Drule.instantiate' [SOME (Thm.ctyp_of thy ty)] var_insts | 
| 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 103 | |> Thm.varifyT_global | 
| 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 104 | val eqs = var_eq :: map_index (mk_partial_term_of_eq thy ty) cs; | 
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 105 | in | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 106 | thy | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 107 | |> Code.del_eqns const | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 108 | |> fold Code.add_eqn eqs | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 109 | end; | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 110 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 111 | fun ensure_partial_term_of_code (tyco, (raw_vs, cs)) thy = | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 112 | let | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 113 |     val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort partial_term_of};
 | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 114 | in if has_inst then add_partial_term_of_code tyco raw_vs cs thy else thy end; | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 115 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 116 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 117 | (* narrowing generators *) | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 118 | |
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 119 | (** narrowing specific names and types **) | 
| 41963 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 120 | |
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 121 | exception FUNCTION_TYPE; | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 122 | |
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 123 | val narrowingN = "narrowing"; | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 124 | |
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 125 | fun narrowingT T = | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 126 |   @{typ Quickcheck_Narrowing.code_int} --> Type (@{type_name Quickcheck_Narrowing.cons}, [T])
 | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 127 | |
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 128 | fun mk_empty T = Const (@{const_name Quickcheck_Narrowing.empty}, narrowingT T)
 | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 129 | |
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 130 | fun mk_cons c T = Const (@{const_name Quickcheck_Narrowing.cons}, T --> narrowingT T) $ Const (c, T)
 | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 131 | |
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 132 | fun mk_apply (T, t) (U, u) = | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 133 | let | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 134 | val (_, U') = dest_funT U | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 135 | in | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 136 |     (U', Const (@{const_name Quickcheck_Narrowing.apply},
 | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 137 | narrowingT U --> narrowingT T --> narrowingT U') $ u $ t) | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 138 | end | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 139 | |
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 140 | fun mk_sum (t, u) = | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 141 | let | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 142 | val T = fastype_of t | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 143 | in | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 144 |     Const (@{const_name Quickcheck_Narrowing.sum}, T --> T --> T) $ t $ u
 | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 145 | end | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 146 | |
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 147 | (** deriving narrowing instances **) | 
| 41963 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 148 | |
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 149 | fun mk_equations descr vs tycos narrowings (Ts, Us) = | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 150 | let | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 151 | fun mk_call T = | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42184diff
changeset | 152 |       (T, Const (@{const_name "Quickcheck_Narrowing.narrowing_class.narrowing"}, narrowingT T))
 | 
| 41963 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 153 | fun mk_aux_call fTs (k, _) (tyco, Ts) = | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 154 | let | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 155 | val T = Type (tyco, Ts) | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 156 | val _ = if not (null fTs) then raise FUNCTION_TYPE else () | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 157 | in | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 158 | (T, nth narrowings k) | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 159 | end | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 160 | fun mk_consexpr simpleT (c, xs) = | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 161 | let | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 162 | val Ts = map fst xs | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 163 | in snd (fold mk_apply xs (Ts ---> simpleT, mk_cons c (Ts ---> simpleT))) end | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 164 | fun mk_rhs exprs = foldr1 mk_sum exprs | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 165 | val rhss = | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 166 | Datatype_Aux.interpret_construction descr vs | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 167 |         { atyp = mk_call, dtyp = mk_aux_call }
 | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 168 | |> (map o apfst) Type | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 169 | |> map (fn (T, cs) => map (mk_consexpr T) cs) | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 170 | |> map mk_rhs | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 171 | val lhss = narrowings | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 172 | val eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhss ~~ rhss) | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 173 | in | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 174 | eqs | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 175 | end | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 176 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 177 | fun contains_recursive_type_under_function_types xs = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 178 | exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT => | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 179 | (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype.DtRec _) => true | _ => false)))) xs | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 180 | |
| 41963 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 181 | fun instantiate_narrowing_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 182 | let | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 183 | val _ = Datatype_Aux.message config "Creating narrowing generators ..."; | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 184 | val narrowingsN = map (prefix (narrowingN ^ "_")) (names @ auxnames); | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 185 | in | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 186 | if not (contains_recursive_type_under_function_types descr) then | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 187 | thy | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 188 |       |> Class.instantiation (tycos, vs, @{sort narrowing})
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 189 | |> Quickcheck_Common.define_functions | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 190 | (fn narrowings => mk_equations descr vs tycos narrowings (Ts, Us), NONE) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 191 | prfx [] narrowingsN (map narrowingT (Ts @ Us)) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 192 | |> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 193 | else | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 194 | thy | 
| 41963 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 195 | end; | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 196 | |
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 197 | (* testing framework *) | 
| 
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
 bulwahn parents: 
41953diff
changeset | 198 | |
| 43308 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 bulwahn parents: 
43240diff
changeset | 199 | val target = "Haskell_Quickcheck" | 
| 41930 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 bulwahn parents: 
41925diff
changeset | 200 | |
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 201 | (** invocation of Haskell interpreter **) | 
| 41905 | 202 | |
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 203 | val narrowing_engine = | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 204 | Context.>>> (Context.map_theory_result | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 205 | (Thy_Load.use_file (Path.explode "Tools/Quickcheck/Narrowing_Engine.hs"))) | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 206 | |
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 207 | val pnf_narrowing_engine = | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 208 | Context.>>> (Context.map_theory_result | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 209 | (Thy_Load.use_file (Path.explode "Tools/Quickcheck/PNF_Narrowing_Engine.hs"))) | 
| 41905 | 210 | |
| 211 | fun exec verbose code = | |
| 212 | ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") verbose code) | |
| 213 | ||
| 43079 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 214 | fun with_overlord_dir name f = | 
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 215 | let | 
| 43602 | 216 | val path = | 
| 217 | Path.append (Path.explode "$ISABELLE_HOME_USER") (Path.basic (name ^ serial_string ())) | |
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 218 | val _ = Isabelle_System.mkdirs path; | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 219 | in Exn.release (Exn.capture f path) end; | 
| 43379 
8c4b383e5143
quickcheck_narrowing returns some timing information
 bulwahn parents: 
43329diff
changeset | 220 | |
| 
8c4b383e5143
quickcheck_narrowing returns some timing information
 bulwahn parents: 
43329diff
changeset | 221 | fun elapsed_time description e = | 
| 
8c4b383e5143
quickcheck_narrowing returns some timing information
 bulwahn parents: 
43329diff
changeset | 222 |   let val ({elapsed, ...}, result) = Timing.timing e ()
 | 
| 
8c4b383e5143
quickcheck_narrowing returns some timing information
 bulwahn parents: 
43329diff
changeset | 223 | in (result, (description, Time.toMilliseconds elapsed)) end | 
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 224 | |
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 225 | fun value (contains_existentials, ((genuine_only, (quiet, verbose)), size)) ctxt cookie (code, value_name) = | 
| 41905 | 226 | let | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 227 | val ((is_genuine, counterexample_of), (get, put, put_ml)) = cookie | 
| 43308 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 bulwahn parents: 
43240diff
changeset | 228 | fun message s = if quiet then () else Output.urgent_message s | 
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 229 | fun verbose_message s = if not quiet andalso verbose then Output.urgent_message s else () | 
| 43585 | 230 | val current_size = Unsynchronized.ref 0 | 
| 231 | val current_result = Unsynchronized.ref Quickcheck.empty_result | |
| 232 | fun excipit () = | |
| 233 | "Quickcheck ran out of time while testing at size " ^ string_of_int (!current_size) | |
| 41930 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 bulwahn parents: 
41925diff
changeset | 234 | val tmp_prefix = "Quickcheck_Narrowing" | 
| 43079 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 235 | val with_tmp_dir = | 
| 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 236 | if Config.get ctxt overlord then with_overlord_dir else Isabelle_System.with_tmp_dir | 
| 41905 | 237 | fun run in_path = | 
| 238 | let | |
| 44751 
f523923d8182
avoid "Code" as structure name (cf. 3bc39cfe27fe)
 bulwahn parents: 
44321diff
changeset | 239 | val code_file = Path.append in_path (Path.basic "Generated_Code.hs") | 
| 41930 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 bulwahn parents: 
41925diff
changeset | 240 | val narrowing_engine_file = Path.append in_path (Path.basic "Narrowing_Engine.hs") | 
| 41905 | 241 | val main_file = Path.append in_path (Path.basic "Main.hs") | 
| 242 |         val main = "module Main where {\n\n" ^
 | |
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 243 | "import System;\n" ^ | 
| 41933 
10f254a4e5b9
adapting Main file generation for Quickcheck_Narrowing
 bulwahn parents: 
41932diff
changeset | 244 | "import Narrowing_Engine;\n" ^ | 
| 44751 
f523923d8182
avoid "Code" as structure name (cf. 3bc39cfe27fe)
 bulwahn parents: 
44321diff
changeset | 245 | "import Generated_Code;\n\n" ^ | 
| 45685 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45420diff
changeset | 246 | "main = getArgs >>= \\[potential, size] -> " ^ | 
| 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45420diff
changeset | 247 | "Narrowing_Engine.depthCheck (read potential) (read size) (Generated_Code.value ())\n\n" ^ | 
| 41905 | 248 | "}\n" | 
| 44751 
f523923d8182
avoid "Code" as structure name (cf. 3bc39cfe27fe)
 bulwahn parents: 
44321diff
changeset | 249 |         val code' = prefix "module Generated_Code where {\n\ndata Typerep = Typerep String [Typerep];\n"
 | 
| 
f523923d8182
avoid "Code" as structure name (cf. 3bc39cfe27fe)
 bulwahn parents: 
44321diff
changeset | 250 |           (unprefix "module Generated_Code where {" code)
 | 
| 41909 
383bbdad1650
replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
 bulwahn parents: 
41908diff
changeset | 251 | val _ = File.write code_file code' | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 252 | val _ = File.write narrowing_engine_file | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 253 | (if contains_existentials then pnf_narrowing_engine else narrowing_engine) | 
| 41905 | 254 | val _ = File.write main_file main | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 255 | val executable = File.shell_path (Path.append in_path (Path.basic "isabelle_quickcheck_narrowing")) | 
| 44926 
de3ed037c9a5
create central list for language extensions used by the haskell code generator
 noschinl parents: 
44852diff
changeset | 256 | val cmd = "exec \"$ISABELLE_GHC\" " ^ Code_Haskell.language_params ^ " " ^ | 
| 41946 | 257 | (space_implode " " (map File.shell_path [code_file, narrowing_engine_file, main_file])) ^ | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 258 | " -o " ^ executable ^ ";" | 
| 43850 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43702diff
changeset | 259 | val (result, compilation_time) = | 
| 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43702diff
changeset | 260 | elapsed_time "Haskell compilation" (fn () => Isabelle_System.bash cmd) | 
| 43585 | 261 | val _ = Quickcheck.add_timing compilation_time current_result | 
| 45685 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45420diff
changeset | 262 | fun haskell_string_of_bool v = if v then "True" else "False" | 
| 43850 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43702diff
changeset | 263 | val _ = if Isabelle_System.bash cmd <> 0 then error "Compilation with GHC failed" else () | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 264 | fun with_size genuine_only k = | 
| 43308 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 bulwahn parents: 
43240diff
changeset | 265 | if k > size then | 
| 43585 | 266 | (NONE, !current_result) | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 267 | else | 
| 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 268 | let | 
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 269 |               val _ = verbose_message ("[Quickcheck-narrowing] Test data size: " ^ string_of_int k)
 | 
| 43585 | 270 | val _ = current_size := k | 
| 271 |               val ((response, _), timing) = elapsed_time ("execution of size " ^ string_of_int k)
 | |
| 45685 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45420diff
changeset | 272 | (fn () => Isabelle_System.bash_output | 
| 45760 
3b5a735897c3
inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
 bulwahn parents: 
45757diff
changeset | 273 | (executable ^ " " ^ haskell_string_of_bool genuine_only ^ " " ^ string_of_int k)) | 
| 43585 | 274 | val _ = Quickcheck.add_timing timing current_result | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 275 | in | 
| 43585 | 276 | if response = "NONE\n" then | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 277 | with_size genuine_only (k + 1) | 
| 43585 | 278 | else | 
| 279 | let | |
| 280 | val output_value = the_default "NONE" | |
| 281 | (try (snd o split_last o filter_out (fn s => s = "") o split_lines) response) | |
| 282 | |> translate_string (fn s => if s = "\\" then "\\\\" else s) | |
| 283 |                   val ml_code = "\nval _ = Context.set_thread_data (SOME (Context.map_proof (" ^ put_ml
 | |
| 284 | ^ " (fn () => " ^ output_value ^ ")) (ML_Context.the_generic_context ())))"; | |
| 285 | val ctxt' = ctxt | |
| 286 |                     |> put (fn () => error ("Bad evaluation for " ^ quote put_ml))
 | |
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 287 | |> Context.proof_map (exec false ml_code); | 
| 45757 
e32dd098f57a
renaming potential flag to genuine_only flag with an inverse semantics
 bulwahn parents: 
45756diff
changeset | 288 | val counterexample = get ctxt' () | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 289 | in | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 290 | if is_genuine counterexample then | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 291 | (counterexample, !current_result) | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 292 | else | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 293 | let | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 294 | val cex = Option.map (rpair []) (counterexample_of counterexample) | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 295 | in | 
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 296 | (message (Pretty.string_of (Quickcheck.pretty_counterex ctxt false cex)); | 
| 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 297 | message "Quickcheck continues to find a genuine counterexample..."; | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 298 | with_size true (k + 1)) | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 299 | end | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 300 | end | 
| 43585 | 301 | end | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 302 | in with_size genuine_only 0 end | 
| 43585 | 303 | in | 
| 45418 | 304 | with_tmp_dir tmp_prefix run | 
| 43585 | 305 | end; | 
| 41905 | 306 | |
| 43308 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 bulwahn parents: 
43240diff
changeset | 307 | fun dynamic_value_strict opts cookie thy postproc t = | 
| 41905 | 308 | let | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 309 | val ctxt = Proof_Context.init_global thy | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 310 | fun evaluator naming program ((_, vs_ty), t) deps = | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 311 | Exn.interruptible_capture (value opts ctxt cookie) | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 312 | (Code_Target.evaluator thy target naming program deps (vs_ty, t)); | 
| 41905 | 313 | in Exn.release (Code_Thingol.dynamic_value thy (Exn.map_result o postproc) evaluator t) end; | 
| 314 | ||
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 315 | (** counterexample generator **) | 
| 46042 | 316 | |
| 317 | (* FIXME just one data slot (record) per program unit *) | |
| 41932 
e8f113ce8a94
adapting Quickcheck_Narrowing and example file to new names
 bulwahn parents: 
41930diff
changeset | 318 | structure Counterexample = Proof_Data | 
| 41905 | 319 | ( | 
| 45725 
2987b29518aa
the narrowing also indicates if counterexample is potentially spurious
 bulwahn parents: 
45685diff
changeset | 320 | type T = unit -> (bool * term list) option | 
| 41936 
9792a882da9c
renaming tester from lazy_exhaustive to narrowing
 bulwahn parents: 
41933diff
changeset | 321 | fun init _ () = error "Counterexample" | 
| 41905 | 322 | ) | 
| 323 | ||
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 324 | datatype counterexample = Universal_Counterexample of (term * counterexample) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 325 | | Existential_Counterexample of (term * counterexample) list | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 326 | | Empty_Assignment | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 327 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 328 | fun map_counterexample f Empty_Assignment = Empty_Assignment | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 329 | | map_counterexample f (Universal_Counterexample (t, c)) = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 330 | Universal_Counterexample (f t, map_counterexample f c) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 331 | | map_counterexample f (Existential_Counterexample cs) = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 332 | Existential_Counterexample (map (fn (t, c) => (f t, map_counterexample f c)) cs) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 333 | |
| 46042 | 334 | (* FIXME just one data slot (record) per program unit *) | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 335 | structure Existential_Counterexample = Proof_Data | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 336 | ( | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 337 | type T = unit -> counterexample option | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 338 | fun init _ () = error "Counterexample" | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 339 | ) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 340 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 341 | val put_existential_counterexample = Existential_Counterexample.put | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 342 | |
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 343 | val put_counterexample = Counterexample.put | 
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42020diff
changeset | 344 | |
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 345 | fun finitize_functions (xTs, t) = | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 346 | let | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 347 | val (names, boundTs) = split_list xTs | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 348 | fun mk_eval_ffun dT rT = | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 349 |       Const (@{const_name "Quickcheck_Narrowing.eval_ffun"}, 
 | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 350 |         Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT]) --> dT --> rT)
 | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 351 | fun mk_eval_cfun dT rT = | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 352 |       Const (@{const_name "Quickcheck_Narrowing.eval_cfun"}, 
 | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 353 |         Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT]) --> dT --> rT)
 | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 354 |     fun eval_function (T as Type (@{type_name fun}, [dT, rT])) =
 | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 355 | let | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 356 | val (rt', rT') = eval_function rT | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 357 | in | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 358 | case dT of | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 359 |           Type (@{type_name fun}, _) =>
 | 
| 44241 | 360 | (fn t => absdummy dT (rt' (mk_eval_cfun dT rT' $ incr_boundvars 1 t $ Bound 0)), | 
| 361 |               Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT']))
 | |
| 362 | | _ => | |
| 363 | (fn t => absdummy dT (rt' (mk_eval_ffun dT rT' $ incr_boundvars 1 t $ Bound 0)), | |
| 364 |               Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT']))
 | |
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 365 | end | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 366 | | eval_function T = (I, T) | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 367 | val (tt, boundTs') = split_list (map eval_function boundTs) | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 368 | val t' = subst_bounds (map2 (fn f => fn x => f x) (rev tt) (map_index (Bound o fst) boundTs), t) | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 369 | in | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 370 | (names ~~ boundTs', t') | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 371 | end | 
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42020diff
changeset | 372 | |
| 45039 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 373 | fun dest_ffun (Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT])) = (dT, rT)
 | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 374 | |
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 375 | fun eval_finite_functions (Const (@{const_name "Quickcheck_Narrowing.ffun.Constant"}, T) $ value) =
 | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 376 | absdummy (fst (dest_ffun (body_type T))) (eval_finite_functions value) | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 377 |   | eval_finite_functions (Const (@{const_name "Quickcheck_Narrowing.ffun.Update"}, T) $ a $ b $ f) =
 | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 378 | let | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 379 | val (T1, T2) = dest_ffun (body_type T) | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 380 | in | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 381 | Quickcheck_Common.mk_fun_upd T1 T2 | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 382 | (eval_finite_functions a, eval_finite_functions b) (eval_finite_functions f) | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 383 | end | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 384 | | eval_finite_functions t = t | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 385 | |
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 386 | (** tester **) | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 387 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 388 | val rewrs = | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 389 | map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 390 |       (@{thms all_simps} @ @{thms ex_simps})
 | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 391 | @ map (HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 392 |         [@{thm iff_conv_conj_imp}, @{thm not_ex}, @{thm not_all}]
 | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 393 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 394 | fun make_pnf_term thy t = Pattern.rewrite_term thy rewrs [] t | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 395 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 396 | fun strip_quantifiers (Const (@{const_name Ex}, _) $ Abs (x, T, t)) =
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 397 |     apfst (cons (@{const_name Ex}, (x, T))) (strip_quantifiers t)
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 398 |   | strip_quantifiers (Const (@{const_name All}, _) $ Abs (x, T, t)) =
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 399 |     apfst (cons (@{const_name All}, (x, T))) (strip_quantifiers t)
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 400 | | strip_quantifiers t = ([], t) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 401 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 402 | fun contains_existentials t = exists (fn (Q, _) => Q = @{const_name Ex}) (fst (strip_quantifiers t))
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 403 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 404 | fun mk_property qs t = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 405 | let | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 406 |     fun enclose (@{const_name Ex}, (x, T)) t =
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 407 |         Const (@{const_name Quickcheck_Narrowing.exists}, (T --> @{typ property}) --> @{typ property})
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 408 | $ Abs (x, T, t) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 409 |       | enclose (@{const_name All}, (x, T)) t =
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 410 |         Const (@{const_name Quickcheck_Narrowing.all}, (T --> @{typ property}) --> @{typ property})
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 411 | $ Abs (x, T, t) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 412 | in | 
| 45002 
df36896aae0f
removing superfluous definition in the quickcheck narrowing invocation as the code generator now generates valid Haskell code with necessary type annotations without a separate definition
 bulwahn parents: 
44926diff
changeset | 413 |     fold_rev enclose qs (@{term Quickcheck_Narrowing.Property} $ t)
 | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 414 | end | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 415 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 416 | fun mk_case_term ctxt p ((@{const_name Ex}, (x, T)) :: qs') (Existential_Counterexample cs) =
 | 
| 45891 
d73605c829cc
clarified module dependencies: Datatype_Data, Datatype_Case, Rep_Datatype;
 wenzelm parents: 
45765diff
changeset | 417 | Datatype_Case.make_case ctxt Datatype_Case.Quiet [] (Free (x, T)) (map (fn (t, c) => | 
| 43317 | 418 | (t, mk_case_term ctxt (p - 1) qs' c)) cs) | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 419 |   | mk_case_term ctxt p ((@{const_name All}, (x, T)) :: qs') (Universal_Counterexample (t, c)) =
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 420 | if p = 0 then t else mk_case_term ctxt (p - 1) qs' c | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 421 | |
| 45039 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 422 | val post_process = (perhaps (try Quickcheck_Common.post_process_term)) o eval_finite_functions | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 423 | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 424 | fun mk_terms ctxt qs result = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 425 | let | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 426 | val | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 427 |       ps = filter (fn (_, (@{const_name All}, _)) => true | _ => false) (map_index I qs)
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 428 | in | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 429 | map (fn (p, (_, (x, T))) => (x, mk_case_term ctxt p qs result)) ps | 
| 45039 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 430 | |> map (apsnd post_process) | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 431 | end | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 432 | |
| 45418 | 433 | fun test_term ctxt (t, eval_terms) = | 
| 41905 | 434 | let | 
| 43585 | 435 | fun dest_result (Quickcheck.Result r) = r | 
| 436 | val opts = | |
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 437 | ((Config.get ctxt Quickcheck.genuine_only, | 
| 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 438 | (Config.get ctxt Quickcheck.quiet, Config.get ctxt Quickcheck.verbose)), | 
| 45685 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45420diff
changeset | 439 | Config.get ctxt Quickcheck.size) | 
| 42361 | 440 | val thy = Proof_Context.theory_of ctxt | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 441 | val t' = fold_rev (fn (x, T) => fn t => HOLogic.mk_all (x, T, t)) (Term.add_frees t []) t | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 442 | val pnf_t = make_pnf_term thy t' | 
| 41905 | 443 | in | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 444 | if Config.get ctxt allow_existentials andalso contains_existentials pnf_t then | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 445 | let | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 446 | fun wrap f (qs, t) = | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 447 | let val (qs1, qs2) = split_list qs in | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 448 | apfst (map2 pair qs1) (f (qs2, t)) end | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 449 | val finitize = if Config.get ctxt finite_functions then wrap finitize_functions else I | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 450 | val (qs, prop_t) = finitize (strip_quantifiers pnf_t) | 
| 43585 | 451 | val (counterexample, result) = dynamic_value_strict (true, opts) | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 452 | ((K true, fn _ => error ""), (Existential_Counterexample.get, Existential_Counterexample.put, | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 453 | "Narrowing_Generators.put_existential_counterexample")) | 
| 45002 
df36896aae0f
removing superfluous definition in the quickcheck narrowing invocation as the code generator now generates valid Haskell code with necessary type annotations without a separate definition
 bulwahn parents: 
44926diff
changeset | 454 | thy (apfst o Option.map o map_counterexample) (mk_property qs prop_t) | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 455 | in | 
| 43585 | 456 | Quickcheck.Result | 
| 45728 
123e3a9a3bb3
compilations return genuine flag to quickcheck framework
 bulwahn parents: 
45725diff
changeset | 457 |          {counterexample = Option.map (pair true o mk_terms ctxt qs) counterexample,
 | 
| 43585 | 458 | evaluation_terms = Option.map (K []) counterexample, | 
| 459 | timings = #timings (dest_result result), reports = #reports (dest_result result)} | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 460 | end | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 461 | else | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 462 | let | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 463 | val frees = Term.add_frees t [] | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 464 | val t' = fold_rev absfree frees t | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 465 | fun wrap f t = list_abs (f (strip_abs t)) | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 466 | val finitize = if Config.get ctxt finite_functions then wrap finitize_functions else I | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 467 | fun ensure_testable t = | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 468 |           Const (@{const_name Quickcheck_Narrowing.ensure_testable}, fastype_of t --> fastype_of t) $ t
 | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 469 | fun is_genuine (SOME (true, _)) = true | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 470 | | is_genuine _ = false | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 471 | val counterexample_of = Option.map (apsnd (curry (op ~~) (map fst frees) o map post_process)) | 
| 43585 | 472 | val (counterexample, result) = dynamic_value_strict (false, opts) | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 473 | ((is_genuine, counterexample_of), | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 474 | (Counterexample.get, Counterexample.put, "Narrowing_Generators.put_counterexample")) | 
| 45725 
2987b29518aa
the narrowing also indicates if counterexample is potentially spurious
 bulwahn parents: 
45685diff
changeset | 475 | thy (apfst o Option.map o apsnd o map) (ensure_testable (finitize t')) | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 476 | in | 
| 43585 | 477 | Quickcheck.Result | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 478 |          {counterexample = counterexample_of counterexample,
 | 
| 43585 | 479 | evaluation_terms = Option.map (K []) counterexample, | 
| 480 | timings = #timings (dest_result result), reports = #reports (dest_result result)} | |
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 481 | end | 
| 41905 | 482 | end; | 
| 483 | ||
| 45419 
10ba32c347b0
quickcheck fails with code generator errors only if one tester is invoked
 bulwahn parents: 
45418diff
changeset | 484 | fun test_goals ctxt _ insts goals = | 
| 43314 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 485 | if (not (getenv "ISABELLE_GHC" = "")) then | 
| 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 486 | let | 
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 487 |       val _ = Quickcheck.message ctxt ("Testing conjecture with Quickcheck-narrowing...")
 | 
| 45159 
3f1d1ce024cb
moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
 bulwahn parents: 
45039diff
changeset | 488 | val correct_inst_goals = Quickcheck_Common.instantiate_goals ctxt insts goals | 
| 43314 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 489 | in | 
| 45418 | 490 | Quickcheck_Common.collect_results (test_term ctxt) (maps (map snd) correct_inst_goals) [] | 
| 43314 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 491 | end | 
| 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 492 | else | 
| 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 493 | (if Config.get ctxt Quickcheck.quiet then () else Output.urgent_message | 
| 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 494 |       ("Environment variable ISABELLE_GHC is not set. To use narrowing-based quickcheck, please set "
 | 
| 43911 
a1da544e2652
more information for the user how to deactivate quickcheck_narrowing if he does not want to use it
 bulwahn parents: 
43910diff
changeset | 495 | ^ "this variable to your GHC Haskell compiler in your settings file. " | 
| 
a1da544e2652
more information for the user how to deactivate quickcheck_narrowing if he does not want to use it
 bulwahn parents: 
43910diff
changeset | 496 | ^ "To deactivate narrowing-based quickcheck, set quickcheck_narrowing_active to false."); | 
| 
a1da544e2652
more information for the user how to deactivate quickcheck_narrowing if he does not want to use it
 bulwahn parents: 
43910diff
changeset | 497 | [Quickcheck.empty_result]) | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 498 | |
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 499 | (* setup *) | 
| 41905 | 500 | |
| 44272 | 501 | val active = Attrib.setup_config_bool @{binding quickcheck_narrowing_active} (K true);
 | 
| 43878 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 bulwahn parents: 
43850diff
changeset | 502 | |
| 41905 | 503 | val setup = | 
| 43047 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 504 | Code.datatype_interpretation ensure_partial_term_of | 
| 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 bulwahn parents: 
42616diff
changeset | 505 | #> Code.datatype_interpretation ensure_partial_term_of_code | 
| 45923 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45891diff
changeset | 506 |   #> Quickcheck_Common.datatype_interpretation (@{sort narrowing}, instantiate_narrowing_datatype)
 | 
| 43878 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 bulwahn parents: 
43850diff
changeset | 507 |   #> Context.theory_map (Quickcheck.add_tester ("narrowing", (active, test_goals)))
 | 
| 41905 | 508 | |
| 44272 | 509 | end; |