| author | wenzelm | 
| Mon, 23 Mar 2015 17:01:47 +0100 | |
| changeset 59785 | 4e6ab5831cc0 | 
| parent 59621 | 291934bac95e | 
| child 60801 | 7664e0916eec | 
| 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 | 
| 48270 | 12 | val ghc_options : string Config.T (* FIXME prefer settings, i.e. getenv (!?) *) | 
| 43891 
4690f76f327a
making active configuration public in narrowing-based quickcheck
 bulwahn parents: 
43878diff
changeset | 13 | val active : bool Config.T | 
| 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 | end; | 
| 20 | ||
| 41930 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 bulwahn parents: 
41925diff
changeset | 21 | structure Narrowing_Generators : NARROWING_GENERATORS = | 
| 41905 | 22 | struct | 
| 23 | ||
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42020diff
changeset | 24 | (* configurations *) | 
| 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42020diff
changeset | 25 | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 26 | 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 | 27 | 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 | 28 | val overlord = Attrib.setup_config_bool @{binding quickcheck_narrowing_overlord} (K false)
 | 
| 47843 
4da917ed49b7
adding configuration to pass options to the ghc call in quickcheck[narrowing]
 bulwahn parents: 
47330diff
changeset | 29 | val ghc_options = Attrib.setup_config_string @{binding quickcheck_narrowing_ghc_options} (K "")
 | 
| 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})
 | 
| 56242 | 36 | $ Logic.mk_type T $ x | 
| 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 | 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 | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59434diff
changeset | 57 | |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) | 
| 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 | 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 | 
| 48272 | 62 |     val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco @{sort partial_term_of})
 | 
| 63 |       andalso Sorts.has_instance (Sign.classes_of thy) tyco @{sort typerep};
 | |
| 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 | 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))
 | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50046diff
changeset | 72 |     val narrowing_term = @{term Quickcheck_Narrowing.Narrowing_constructor} $ HOLogic.mk_number @{typ integer} i
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50046diff
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 = | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59498diff
changeset | 78 | map (SOME o Thm.global_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]
 | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59498diff
changeset | 80 | val cty = Thm.global_ctyp_of thy ty; | 
| 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 | 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; | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51143diff
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 = | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59498diff
changeset | 97 | map (SOME o Thm.global_cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) | 
| 46758 
4106258260b3
choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
 bulwahn parents: 
46587diff
changeset | 98 |         [Free ("ty", Term.itselfT ty), @{term "Quickcheck_Narrowing.Narrowing_variable 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}
 | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59498diff
changeset | 102 | |> Drule.instantiate' [SOME (Thm.global_ctyp_of thy ty)] var_insts | 
| 43079 
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 | 
| 48272 | 113 |     val has_inst = Sorts.has_instance (Sign.classes_of thy) tyco @{sort partial_term_of};
 | 
| 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 | 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 = | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50046diff
changeset | 126 |   @{typ integer} --> Type (@{type_name Quickcheck_Narrowing.narrowing_cons}, [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 | 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_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 | 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_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 | 131 | 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 | 132 | 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 | 133 | 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 | 134 |     (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 | 135 | 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 | 136 | end | 
| 59154 | 137 | |
| 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 | 138 | 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 | 139 | 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 | 140 | 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 | 141 | 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 | 142 |     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 | 143 | 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 | 144 | |
| 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 | 145 | (** 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 | 146 | |
| 50046 | 147 | fun mk_equations descr vs narrowings = | 
| 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 | 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 | 149 | fun mk_call T = | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42184diff
changeset | 150 |       (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 | 151 | 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 | 152 | 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 | 153 | 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 | 154 | 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 | 155 | 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 | 156 | (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 | 157 | 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 | 158 | 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 | 159 | 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 | 160 | 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 | 161 | 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 | 162 | 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 | 163 | val rhss = | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
57996diff
changeset | 164 | Old_Datatype_Aux.interpret_construction descr vs | 
| 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 | 165 |         { 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 | 166 | |> (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 | 167 | |> 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 | 168 | |> 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 | 169 | 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 | 170 | 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 | 171 | 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 | 172 | 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 | 173 | end | 
| 59154 | 174 | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 175 | fun contains_recursive_type_under_function_types xs = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 176 | exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT => | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
57996diff
changeset | 177 | (case Old_Datatype_Aux.strip_dtyp dT of (_ :: _, Old_Datatype_Aux.DtRec _) => true | _ => false)))) xs | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 178 | |
| 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 | 179 | 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 | 180 | let | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
57996diff
changeset | 181 | val _ = Old_Datatype_Aux.message config "Creating narrowing generators ..."; | 
| 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 | 182 | 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 | 183 | in | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 184 | if not (contains_recursive_type_under_function_types descr) then | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 185 | thy | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 186 |       |> Class.instantiation (tycos, vs, @{sort narrowing})
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 187 | |> Quickcheck_Common.define_functions | 
| 50046 | 188 | (fn narrowings => mk_equations descr vs narrowings, NONE) | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 189 | prfx [] narrowingsN (map narrowingT (Ts @ Us)) | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59434diff
changeset | 190 | |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 191 | else | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 192 | 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 | 193 | 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 | 194 | |
| 
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 | (* 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 | 196 | |
| 43308 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 bulwahn parents: 
43240diff
changeset | 197 | val target = "Haskell_Quickcheck" | 
| 41930 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 bulwahn parents: 
41925diff
changeset | 198 | |
| 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 | 199 | (** invocation of Haskell interpreter **) | 
| 41905 | 200 | |
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 201 | val narrowing_engine = | 
| 56135 | 202 |   File.read @{path "~~/src/HOL/Tools/Quickcheck/Narrowing_Engine.hs"}
 | 
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 203 | |
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43619diff
changeset | 204 | val pnf_narrowing_engine = | 
| 56135 | 205 |   File.read @{path "~~/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs"}
 | 
| 41905 | 206 | |
| 207 | fun exec verbose code = | |
| 208 | ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") verbose code) | |
| 209 | ||
| 43079 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 210 | 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 | 211 | let | 
| 43602 | 212 | val path = | 
| 213 | 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 | 214 | 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 | 215 | in Exn.release (Exn.capture f path) end; | 
| 43379 
8c4b383e5143
quickcheck_narrowing returns some timing information
 bulwahn parents: 
43329diff
changeset | 216 | |
| 
8c4b383e5143
quickcheck_narrowing returns some timing information
 bulwahn parents: 
43329diff
changeset | 217 | fun elapsed_time description e = | 
| 
8c4b383e5143
quickcheck_narrowing returns some timing information
 bulwahn parents: 
43329diff
changeset | 218 |   let val ({elapsed, ...}, result) = Timing.timing e ()
 | 
| 
8c4b383e5143
quickcheck_narrowing returns some timing information
 bulwahn parents: 
43329diff
changeset | 219 | in (result, (description, Time.toMilliseconds elapsed)) end | 
| 59154 | 220 | |
| 50046 | 221 | fun value (contains_existentials, ((genuine_only, (quiet, verbose)), size)) ctxt cookie (code_modules, _) = | 
| 41905 | 222 | let | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 223 | val ((is_genuine, counterexample_of), (get, put, put_ml)) = cookie | 
| 58843 | 224 | fun message s = if quiet then () else writeln s | 
| 225 | fun verbose_message s = if not quiet andalso verbose then writeln s else () | |
| 43585 | 226 | val current_size = Unsynchronized.ref 0 | 
| 59154 | 227 | val current_result = Unsynchronized.ref Quickcheck.empty_result | 
| 41930 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 bulwahn parents: 
41925diff
changeset | 228 | val tmp_prefix = "Quickcheck_Narrowing" | 
| 47843 
4da917ed49b7
adding configuration to pass options to the ghc call in quickcheck[narrowing]
 bulwahn parents: 
47330diff
changeset | 229 | val ghc_options = Config.get ctxt ghc_options | 
| 43079 
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
 bulwahn parents: 
43047diff
changeset | 230 | val with_tmp_dir = | 
| 59154 | 231 | if Config.get ctxt overlord then with_overlord_dir else Isabelle_System.with_tmp_dir | 
| 232 | fun run in_path = | |
| 41905 | 233 | let | 
| 48568 | 234 | fun mk_code_file name = Path.append in_path (Path.basic (name ^ ".hs")) | 
| 235 | val generatedN = Code_Target.generatedN | |
| 236 | val includes = AList.delete (op =) generatedN code_modules |> (map o apfst) mk_code_file; | |
| 237 | val code = the (AList.lookup (op =) code_modules generatedN) | |
| 238 | val code_file = mk_code_file generatedN | |
| 239 | val narrowing_engine_file = mk_code_file "Narrowing_Engine" | |
| 240 | val main_file = mk_code_file "Main" | |
| 41905 | 241 |         val main = "module Main where {\n\n" ^
 | 
| 46335 
0fd9ab902b5a
using fully qualified module names in Haskell source, which seems to be required by GHC 7.0.4
 bulwahn parents: 
46316diff
changeset | 242 | "import System.IO;\n" ^ | 
| 
0fd9ab902b5a
using fully qualified module names in Haskell source, which seems to be required by GHC 7.0.4
 bulwahn parents: 
46316diff
changeset | 243 | "import System.Environment;\n" ^ | 
| 41933 
10f254a4e5b9
adapting Main file generation for Quickcheck_Narrowing
 bulwahn parents: 
41932diff
changeset | 244 | "import Narrowing_Engine;\n" ^ | 
| 48568 | 245 | "import " ^ generatedN ^ " ;\n\n" ^ | 
| 45685 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45420diff
changeset | 246 | "main = getArgs >>= \\[potential, size] -> " ^ | 
| 48568 | 247 |           "Narrowing_Engine.depthCheck (read potential) (read size) (" ^ generatedN ^ ".value ())\n\n" ^
 | 
| 41905 | 248 | "}\n" | 
| 48568 | 249 | val _ = map (uncurry File.write) (includes @ | 
| 250 | [(narrowing_engine_file, if contains_existentials then pnf_narrowing_engine else narrowing_engine), | |
| 55676 | 251 | (code_file, code), (main_file, main)]) | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 252 | 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 | 253 | val cmd = "exec \"$ISABELLE_GHC\" " ^ Code_Haskell.language_params ^ " " ^ | 
| 47843 
4da917ed49b7
adding configuration to pass options to the ghc call in quickcheck[narrowing]
 bulwahn parents: 
47330diff
changeset | 254 | ghc_options ^ " " ^ | 
| 48568 | 255 | (space_implode " " (map File.shell_path (map fst includes @ [code_file, narrowing_engine_file, main_file]))) ^ | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 256 | " -o " ^ executable ^ ";" | 
| 50046 | 257 | val (_, compilation_time) = | 
| 48568 | 258 | elapsed_time "Haskell compilation" (fn () => Isabelle_System.bash cmd) | 
| 43585 | 259 | val _ = Quickcheck.add_timing compilation_time current_result | 
| 45685 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45420diff
changeset | 260 | 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 | 261 | 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 | 262 | 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 | 263 | if k > size then | 
| 43585 | 264 | (NONE, !current_result) | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 265 | else | 
| 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 266 | let | 
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 267 |               val _ = verbose_message ("[Quickcheck-narrowing] Test data size: " ^ string_of_int k)
 | 
| 43585 | 268 | val _ = current_size := k | 
| 269 |               val ((response, _), timing) = elapsed_time ("execution of size " ^ string_of_int k)
 | |
| 45685 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45420diff
changeset | 270 | (fn () => Isabelle_System.bash_output | 
| 45760 
3b5a735897c3
inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
 bulwahn parents: 
45757diff
changeset | 271 | (executable ^ " " ^ haskell_string_of_bool genuine_only ^ " " ^ string_of_int k)) | 
| 43585 | 272 | val _ = Quickcheck.add_timing timing current_result | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 273 | in | 
| 43585 | 274 | if response = "NONE\n" then | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 275 | with_size genuine_only (k + 1) | 
| 43585 | 276 | else | 
| 277 | let | |
| 278 | val output_value = the_default "NONE" | |
| 279 | (try (snd o split_last o filter_out (fn s => s = "") o split_lines) response) | |
| 280 | |> translate_string (fn s => if s = "\\" then "\\\\" else s) | |
| 281 |                   val ml_code = "\nval _ = Context.set_thread_data (SOME (Context.map_proof (" ^ put_ml
 | |
| 282 | ^ " (fn () => " ^ output_value ^ ")) (ML_Context.the_generic_context ())))"; | |
| 283 | val ctxt' = ctxt | |
| 284 |                     |> put (fn () => error ("Bad evaluation for " ^ quote put_ml))
 | |
| 59154 | 285 | |> Context.proof_map (exec false ml_code); | 
| 45757 
e32dd098f57a
renaming potential flag to genuine_only flag with an inverse semantics
 bulwahn parents: 
45756diff
changeset | 286 | val counterexample = get ctxt' () | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 287 | in | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 288 | if is_genuine counterexample then | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 289 | (counterexample, !current_result) | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 290 | else | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 291 | let | 
| 59434 | 292 | val cex = Option.map (rpair []) (counterexample_of counterexample); | 
| 293 | val _ = message (Pretty.string_of (Quickcheck.pretty_counterex ctxt false cex)); | |
| 294 | val _ = message "Quickcheck continues to find a genuine counterexample..."; | |
| 295 | in with_size true (k + 1) end | |
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 296 | end | 
| 59154 | 297 | end | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 298 | in with_size genuine_only 0 end | 
| 43585 | 299 | in | 
| 45418 | 300 | with_tmp_dir tmp_prefix run | 
| 43585 | 301 | end; | 
| 41905 | 302 | |
| 55757 | 303 | fun dynamic_value_strict opts cookie ctxt postproc t = | 
| 41905 | 304 | let | 
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56242diff
changeset | 305 | fun evaluator program _ vs_ty_t deps = | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 306 | Exn.interruptible_capture (value opts ctxt cookie) | 
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56242diff
changeset | 307 | (Code_Target.evaluator ctxt target program deps true vs_ty_t); | 
| 55757 | 308 | in Exn.release (Code_Thingol.dynamic_value ctxt (Exn.map_result o postproc) evaluator t) end; | 
| 41905 | 309 | |
| 59154 | 310 | |
| 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 | 311 | (** counterexample generator **) | 
| 46042 | 312 | |
| 59154 | 313 | datatype counterexample = | 
| 314 | Universal_Counterexample of (term * counterexample) | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 315 | | Existential_Counterexample of (term * counterexample) list | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 316 | | Empty_Assignment | 
| 59154 | 317 | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 318 | fun map_counterexample f Empty_Assignment = Empty_Assignment | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 319 | | map_counterexample f (Universal_Counterexample (t, c)) = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 320 | Universal_Counterexample (f t, map_counterexample f c) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 321 | | map_counterexample f (Existential_Counterexample cs) = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 322 | 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 | 323 | |
| 59154 | 324 | structure Data = Proof_Data | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 325 | ( | 
| 59154 | 326 | type T = | 
| 327 | (unit -> (bool * term list) option) * | |
| 328 | (unit -> counterexample option); | |
| 329 | val empty: T = | |
| 330 | (fn () => raise Fail "counterexample", | |
| 331 | fn () => raise Fail "existential_counterexample"); | |
| 332 | fun init _ = empty; | |
| 333 | ); | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 334 | |
| 59154 | 335 | val get_counterexample = #1 o Data.get; | 
| 336 | val get_existential_counterexample = #2 o Data.get; | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 337 | |
| 59154 | 338 | val put_counterexample = Data.map o @{apply 2(1)} o K;
 | 
| 339 | val put_existential_counterexample = Data.map o @{apply 2(2)} o K;
 | |
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42020diff
changeset | 340 | |
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 341 | 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 | 342 | let | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 343 | 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 | 344 | fun mk_eval_ffun dT rT = | 
| 59154 | 345 |       Const (@{const_name "Quickcheck_Narrowing.eval_ffun"},
 | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 346 |         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 | 347 | fun mk_eval_cfun dT rT = | 
| 59154 | 348 |       Const (@{const_name "Quickcheck_Narrowing.eval_cfun"},
 | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 349 |         Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT]) --> dT --> rT)
 | 
| 50046 | 350 |     fun eval_function (Type (@{type_name fun}, [dT, rT])) =
 | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 351 | let | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 352 | 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 | 353 | in | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 354 | case dT of | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 355 |           Type (@{type_name fun}, _) =>
 | 
| 44241 | 356 | (fn t => absdummy dT (rt' (mk_eval_cfun dT rT' $ incr_boundvars 1 t $ Bound 0)), | 
| 357 |               Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT']))
 | |
| 358 | | _ => | |
| 359 | (fn t => absdummy dT (rt' (mk_eval_ffun dT rT' $ incr_boundvars 1 t $ Bound 0)), | |
| 360 |               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 | 361 | end | 
| 47330 
8fe04753a210
rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
 bulwahn parents: 
46758diff
changeset | 362 |       | eval_function (T as Type (@{type_name prod}, [fT, sT])) =
 | 
| 
8fe04753a210
rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
 bulwahn parents: 
46758diff
changeset | 363 | let | 
| 
8fe04753a210
rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
 bulwahn parents: 
46758diff
changeset | 364 | val (ft', fT') = eval_function fT | 
| 
8fe04753a210
rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
 bulwahn parents: 
46758diff
changeset | 365 | val (st', sT') = eval_function sT | 
| 
8fe04753a210
rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
 bulwahn parents: 
46758diff
changeset | 366 |           val T' = Type (@{type_name prod}, [fT', sT'])
 | 
| 55932 | 367 |           val map_const = Const (@{const_name map_prod}, (fT' --> fT) --> (sT' --> sT) --> T' --> T)
 | 
| 47330 
8fe04753a210
rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
 bulwahn parents: 
46758diff
changeset | 368 | fun apply_dummy T t = absdummy T (t (Bound 0)) | 
| 
8fe04753a210
rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
 bulwahn parents: 
46758diff
changeset | 369 | in | 
| 
8fe04753a210
rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
 bulwahn parents: 
46758diff
changeset | 370 | (fn t => list_comb (map_const, [apply_dummy fT' ft', apply_dummy sT' st', t]), T') | 
| 
8fe04753a210
rudimentary handling of products in finitize_functions in Quickcheck-Narrowing
 bulwahn parents: 
46758diff
changeset | 371 | end | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 372 | | eval_function T = (I, T) | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 373 | 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 | 374 | 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 | 375 | in | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 376 | (names ~~ boundTs', t') | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 377 | 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 | 378 | |
| 45039 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 379 | 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 | 380 | |
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 381 | 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 | 382 | 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 | 383 |   | 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 | 384 | let | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 385 | val (T1, T2) = dest_ffun (body_type T) | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 386 | in | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 387 | Quickcheck_Common.mk_fun_upd T1 T2 | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 388 | (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 | 389 | end | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 390 | | eval_finite_functions t = t | 
| 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 391 | |
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 392 | (** tester **) | 
| 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 | val rewrs = | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 395 | 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 | 396 |       (@{thms all_simps} @ @{thms ex_simps})
 | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 397 | @ map (HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) | 
| 46316 | 398 |         [@{thm iff_conv_conj_imp}, @{thm not_ex}, @{thm not_all},
 | 
| 399 |          @{thm meta_eq_to_obj_eq [OF Ex1_def]}]
 | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 400 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 401 | 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 | 402 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 403 | fun strip_quantifiers (Const (@{const_name Ex}, _) $ Abs (x, T, t)) =
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 404 |     apfst (cons (@{const_name Ex}, (x, T))) (strip_quantifiers t)
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 405 |   | strip_quantifiers (Const (@{const_name All}, _) $ Abs (x, T, t)) =
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 406 |     apfst (cons (@{const_name All}, (x, T))) (strip_quantifiers t)
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 407 | | strip_quantifiers t = ([], t) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 408 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 409 | 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 | 410 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 411 | fun mk_property qs t = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 412 | let | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 413 |     fun enclose (@{const_name Ex}, (x, T)) t =
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 414 |         Const (@{const_name Quickcheck_Narrowing.exists}, (T --> @{typ property}) --> @{typ property})
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 415 | $ Abs (x, T, t) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 416 |       | enclose (@{const_name All}, (x, T)) t =
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 417 |         Const (@{const_name Quickcheck_Narrowing.all}, (T --> @{typ property}) --> @{typ property})
 | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 418 | $ Abs (x, T, t) | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 419 | 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 | 420 |     fold_rev enclose qs (@{term Quickcheck_Narrowing.Property} $ t)
 | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 421 | end | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 422 | |
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 423 | fun mk_case_term ctxt p ((@{const_name Ex}, (x, T)) :: qs') (Existential_Counterexample cs) =
 | 
| 51673 
4dfa00e264d8
separate data used for case translation from the datatype package
 traytel parents: 
51672diff
changeset | 424 | Case_Translation.make_case ctxt Case_Translation.Quiet Name.context (Free (x, T)) (map (fn (t, c) => | 
| 43317 | 425 | (t, mk_case_term ctxt (p - 1) qs' c)) cs) | 
| 50046 | 426 |   | mk_case_term ctxt p ((@{const_name All}, _) :: qs') (Universal_Counterexample (t, c)) =
 | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 427 | 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 | 428 | |
| 45039 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 429 | 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 | 430 | |
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 431 | fun mk_terms ctxt qs result = | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 432 | let | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 433 | val | 
| 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 434 |       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 | 435 | in | 
| 50046 | 436 | map (fn (p, (_, (x, _))) => (x, mk_case_term ctxt p qs result)) ps | 
| 45039 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
45002diff
changeset | 437 | |> map (apsnd post_process) | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 438 | end | 
| 59154 | 439 | |
| 50046 | 440 | fun test_term ctxt catch_code_errors (t, _) = | 
| 41905 | 441 | let | 
| 59154 | 442 | fun dest_result (Quickcheck.Result r) = r | 
| 43585 | 443 | val opts = | 
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 444 | ((Config.get ctxt Quickcheck.genuine_only, | 
| 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45760diff
changeset | 445 | (Config.get ctxt Quickcheck.quiet, Config.get ctxt Quickcheck.verbose)), | 
| 45685 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45420diff
changeset | 446 | Config.get ctxt Quickcheck.size) | 
| 42361 | 447 | val thy = Proof_Context.theory_of ctxt | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 448 | 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 | 449 | val pnf_t = make_pnf_term thy t' | 
| 41905 | 450 | in | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 451 | 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 | 452 | let | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 453 | fun wrap f (qs, t) = | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 454 | 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 | 455 | 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 | 456 | 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 | 457 | val (qs, prop_t) = finitize (strip_quantifiers pnf_t) | 
| 46309 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 458 | val act = if catch_code_errors then try else (fn f => SOME o f) | 
| 59154 | 459 | val execute = | 
| 460 | dynamic_value_strict (true, opts) | |
| 461 | ((K true, fn _ => error ""), | |
| 462 | (get_existential_counterexample, put_existential_counterexample, | |
| 463 | "Narrowing_Generators.put_existential_counterexample")) | |
| 464 | ctxt (apfst o Option.map o map_counterexample) | |
| 465 | in | |
| 466 | case act execute (mk_property qs prop_t) of | |
| 46309 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 467 | SOME (counterexample, result) => Quickcheck.Result | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 468 |             {counterexample = Option.map (pair true o mk_terms ctxt qs) counterexample,
 | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 469 | evaluation_terms = Option.map (K []) counterexample, | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 470 | timings = #timings (dest_result result), reports = #reports (dest_result result)} | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 471 | | NONE => | 
| 59434 | 472 | (Quickcheck.message ctxt "Conjecture is not executable with Quickcheck-narrowing"; | 
| 46309 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 473 | Quickcheck.empty_result) | 
| 43237 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 bulwahn parents: 
43114diff
changeset | 474 | end | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 475 | else | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 476 | let | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 477 | val frees = Term.add_frees t [] | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 478 | val t' = fold_rev absfree frees t | 
| 46219 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
46042diff
changeset | 479 | fun wrap f t = uncurry (fold_rev Term.abs) (f (strip_abs t)) | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 480 | 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 | 481 | fun ensure_testable t = | 
| 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 482 |           Const (@{const_name Quickcheck_Narrowing.ensure_testable}, fastype_of t --> fastype_of t) $ t
 | 
| 59154 | 483 | fun is_genuine (SOME (true, _)) = true | 
| 45756 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 484 | | is_genuine _ = false | 
| 
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
 bulwahn parents: 
45728diff
changeset | 485 | val counterexample_of = Option.map (apsnd (curry (op ~~) (map fst frees) o map post_process)) | 
| 46309 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 486 | val act = if catch_code_errors then try else (fn f => SOME o f) | 
| 59154 | 487 | val execute = | 
| 488 | dynamic_value_strict (false, opts) | |
| 489 | ((is_genuine, counterexample_of), | |
| 490 | (get_counterexample, put_counterexample, | |
| 491 | "Narrowing_Generators.put_counterexample")) | |
| 492 | ctxt (apfst o Option.map o apsnd o map) | |
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 493 | in | 
| 59154 | 494 | case act execute (ensure_testable (finitize t')) of | 
| 46309 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 495 | SOME (counterexample, result) => | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 496 | Quickcheck.Result | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 497 |              {counterexample = counterexample_of counterexample,
 | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 498 | evaluation_terms = Option.map (K []) counterexample, | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 499 | timings = #timings (dest_result result), | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 500 | reports = #reports (dest_result result)} | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 501 | | NONE => | 
| 59434 | 502 | (Quickcheck.message ctxt "Conjecture is not executable with Quickcheck-narrowing"; | 
| 46309 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 503 | Quickcheck.empty_result) | 
| 43240 
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
 bulwahn parents: 
43237diff
changeset | 504 | end | 
| 41905 | 505 | end; | 
| 506 | ||
| 46309 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 507 | fun test_goals ctxt catch_code_errors insts goals = | 
| 43314 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 508 | if (not (getenv "ISABELLE_GHC" = "")) then | 
| 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 509 | let | 
| 59434 | 510 | 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 | 511 | 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 | 512 | in | 
| 46309 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 513 | Quickcheck_Common.collect_results (test_term ctxt catch_code_errors) | 
| 
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
 bulwahn parents: 
46219diff
changeset | 514 | (maps (map snd) correct_inst_goals) [] | 
| 43314 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 515 | end | 
| 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 516 | else | 
| 58843 | 517 | (if Config.get ctxt Quickcheck.quiet then () else writeln | 
| 43314 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 bulwahn parents: 
43308diff
changeset | 518 |       ("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 | 519 | ^ "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 | 520 | ^ "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 | 521 | [Quickcheck.empty_result]) | 
| 43114 
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
 bulwahn parents: 
43079diff
changeset | 522 | |
| 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 | 523 | (* setup *) | 
| 41905 | 524 | |
| 46587 
d3bcc356cc60
preliminarily switching quickcheck-narrowing off by default (probably it should only be invoked if concrete testing does not work)
 bulwahn parents: 
46335diff
changeset | 525 | val active = Attrib.setup_config_bool @{binding quickcheck_narrowing_active} (K false);
 | 
| 43878 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 bulwahn parents: 
43850diff
changeset | 526 | |
| 58826 | 527 | val _ = | 
| 528 | Theory.setup | |
| 529 | (Code.datatype_interpretation ensure_partial_term_of | |
| 530 | #> Code.datatype_interpretation ensure_partial_term_of_code | |
| 531 |     #> Quickcheck_Common.datatype_interpretation @{plugin quickcheck_narrowing}
 | |
| 532 |       (@{sort narrowing}, instantiate_narrowing_datatype)
 | |
| 533 |     #> Context.theory_map (Quickcheck.add_tester ("narrowing", (active, test_goals))));
 | |
| 59154 | 534 | |
| 44272 | 535 | end; |