author | wenzelm |
Sat, 30 Nov 2024 13:41:38 +0100 | |
changeset 81512 | c1aa8a61ee65 |
parent 81507 | 08574da77b4a |
permissions | -rw-r--r-- |
41930
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents:
41925
diff
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:
41925
diff
changeset
|
7 |
signature NARROWING_GENERATORS = |
41905 | 8 |
sig |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
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:
42020
diff
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:
43047
diff
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:
43878
diff
changeset
|
13 |
val active : bool Config.T |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
14 |
datatype counterexample = Universal_Counterexample of (term * counterexample) |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
15 |
| Existential_Counterexample of (term * counterexample) list |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
16 |
| Empty_Assignment |
45725
2987b29518aa
the narrowing also indicates if counterexample is potentially spurious
bulwahn
parents:
45685
diff
changeset
|
17 |
val put_counterexample: (unit -> (bool * term list) option) -> Proof.context -> Proof.context |
62979 | 18 |
val put_existential_counterexample : (unit -> counterexample option) -> |
19 |
Proof.context -> Proof.context |
|
20 |
end |
|
41905 | 21 |
|
41930
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents:
41925
diff
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:
42020
diff
changeset
|
25 |
(* configurations *) |
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents:
42020
diff
changeset
|
26 |
|
67149 | 27 |
val allow_existentials = Attrib.setup_config_bool \<^binding>\<open>quickcheck_allow_existentials\<close> (K true) |
28 |
val finite_functions = Attrib.setup_config_bool \<^binding>\<open>quickcheck_finite_functions\<close> (K true) |
|
29 |
val overlord = Attrib.setup_config_bool \<^binding>\<open>quickcheck_narrowing_overlord\<close> (K false) |
|
30 |
val ghc_options = Attrib.setup_config_string \<^binding>\<open>quickcheck_narrowing_ghc_options\<close> (K "") |
|
42023
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents:
42020
diff
changeset
|
31 |
|
62979 | 32 |
|
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:
42616
diff
changeset
|
33 |
(* 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:
42616
diff
changeset
|
34 |
|
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:
42616
diff
changeset
|
35 |
fun mk_partial_term_of (x, T) = |
67149 | 36 |
Const (\<^const_name>\<open>Quickcheck_Narrowing.partial_term_of_class.partial_term_of\<close>, |
37 |
Term.itselfT T --> \<^typ>\<open>narrowing_term\<close> --> \<^typ>\<open>Code_Evaluation.term\<close>) $ Logic.mk_type T $ x |
|
62979 | 38 |
|
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:
42616
diff
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:
42616
diff
changeset
|
40 |
(** 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:
42616
diff
changeset
|
41 |
|
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:
42616
diff
changeset
|
42 |
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:
42616
diff
changeset
|
43 |
let |
67149 | 44 |
val vs = map (fn (v, _) => (v, \<^sort>\<open>typerep\<close>)) raw_vs |
62979 | 45 |
val ty = Type (tyco, map TFree vs) |
46 |
val lhs = |
|
67149 | 47 |
Const (\<^const_name>\<open>partial_term_of\<close>, |
48 |
Term.itselfT ty --> \<^typ>\<open>narrowing_term\<close> --> \<^typ>\<open>Code_Evaluation.term\<close>) $ |
|
49 |
Free ("x", Term.itselfT ty) $ Free ("t", \<^typ>\<open>narrowing_term\<close>) |
|
50 |
val rhs = \<^term>\<open>undefined :: Code_Evaluation.term\<close> |
|
62979 | 51 |
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) |
52 |
fun triv_name_of t = |
|
53 |
(fst o dest_Free o fst o strip_comb o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ |
|
54 |
"_triv" |
|
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:
42616
diff
changeset
|
55 |
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:
42616
diff
changeset
|
56 |
thy |
67149 | 57 |
|> Class.instantiation ([tyco], vs, \<^sort>\<open>partial_term_of\<close>) |
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:
42616
diff
changeset
|
58 |
|> `(fn lthy => Syntax.check_term lthy eq) |
63180 | 59 |
|-> (fn eq => Specification.definition NONE [] [] ((Binding.name (triv_name_of eq), []), eq)) |
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:
42616
diff
changeset
|
60 |
|> snd |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59434
diff
changeset
|
61 |
|> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) |
62979 | 62 |
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:
42616
diff
changeset
|
63 |
|
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:
42616
diff
changeset
|
64 |
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:
42616
diff
changeset
|
65 |
let |
67149 | 66 |
val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\<open>partial_term_of\<close>) |
67 |
andalso Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\<open>typerep\<close> |
|
62979 | 68 |
in if need_inst then add_partial_term_of tyco raw_vs thy else thy 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:
42616
diff
changeset
|
69 |
|
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:
42616
diff
changeset
|
70 |
|
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:
42616
diff
changeset
|
71 |
(** 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:
42616
diff
changeset
|
72 |
|
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:
42616
diff
changeset
|
73 |
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:
42616
diff
changeset
|
74 |
let |
81512
c1aa8a61ee65
tuned: more direct use of Name.context operations;
wenzelm
parents:
81507
diff
changeset
|
75 |
val frees = map (fn a => Free (a, \<^typ>\<open>narrowing_term\<close>)) (Name.invent_global "a" (length tys)) |
62979 | 76 |
val narrowing_term = |
67149 | 77 |
\<^term>\<open>Quickcheck_Narrowing.Narrowing_constructor\<close> $ HOLogic.mk_number \<^typ>\<open>integer\<close> i $ |
78 |
HOLogic.mk_list \<^typ>\<open>narrowing_term\<close> (rev frees) |
|
62979 | 79 |
val rhs = |
67149 | 80 |
fold (fn u => fn t => \<^term>\<open>Code_Evaluation.App\<close> $ t $ u) |
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:
42616
diff
changeset
|
81 |
(map mk_partial_term_of (frees ~~ tys)) |
67149 | 82 |
(\<^term>\<open>Code_Evaluation.Const\<close> $ HOLogic.mk_literal c $ HOLogic.mk_typerep (tys ---> 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:
42616
diff
changeset
|
83 |
val insts = |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59498
diff
changeset
|
84 |
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:
42616
diff
changeset
|
85 |
[Free ("ty", Term.itselfT ty), narrowing_term, rhs] |
62979 | 86 |
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:
42616
diff
changeset
|
87 |
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:
42616
diff
changeset
|
88 |
@{thm partial_term_of_anything} |
60801 | 89 |
|> Thm.instantiate' [SOME cty] insts |
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:
42616
diff
changeset
|
90 |
|> 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:
42616
diff
changeset
|
91 |
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:
42616
diff
changeset
|
92 |
|
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:
42616
diff
changeset
|
93 |
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:
42616
diff
changeset
|
94 |
let |
62979 | 95 |
val algebra = Sign.classes_of thy |
67149 | 96 |
val vs = map (fn (v, sort) => (v, curry (Sorts.inter_sort algebra) \<^sort>\<open>typerep\<close> sort)) raw_vs |
62979 | 97 |
val ty = Type (tyco, map TFree vs) |
98 |
val cs = |
|
99 |
(map o apsnd o apsnd o map o map_atyps) |
|
100 |
(fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs |
|
67149 | 101 |
val const = Axclass.param_of_inst thy (\<^const_name>\<open>partial_term_of\<close>, tyco) |
45344
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents:
45159
diff
changeset
|
102 |
val var_insts = |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59498
diff
changeset
|
103 |
map (SOME o Thm.global_cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) |
67149 | 104 |
[Free ("ty", Term.itselfT ty), \<^term>\<open>Quickcheck_Narrowing.Narrowing_variable p tt\<close>, |
105 |
\<^term>\<open>Code_Evaluation.Free (STR ''_'')\<close> $ HOLogic.mk_typerep ty] |
|
43079
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents:
43047
diff
changeset
|
106 |
val var_eq = |
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents:
43047
diff
changeset
|
107 |
@{thm partial_term_of_anything} |
60801 | 108 |
|> Thm.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:
43047
diff
changeset
|
109 |
|> Thm.varifyT_global |
62979 | 110 |
val eqs = var_eq :: map_index (mk_partial_term_of_eq thy ty) cs |
111 |
in |
|
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:
42616
diff
changeset
|
112 |
thy |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
65905
diff
changeset
|
113 |
|> Code.declare_default_eqns_global (map (rpair true) eqs) |
62979 | 114 |
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:
42616
diff
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:
42616
diff
changeset
|
116 |
fun ensure_partial_term_of_code (tyco, (raw_vs, cs)) thy = |
67149 | 117 |
let val has_inst = Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\<open>partial_term_of\<close> |
62979 | 118 |
in if has_inst then add_partial_term_of_code tyco raw_vs cs thy else thy 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:
42616
diff
changeset
|
119 |
|
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:
42616
diff
changeset
|
120 |
|
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:
42616
diff
changeset
|
121 |
(* 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:
42616
diff
changeset
|
122 |
|
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:
42616
diff
changeset
|
123 |
(** 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:
41953
diff
changeset
|
124 |
|
62979 | 125 |
exception FUNCTION_TYPE |
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
126 |
|
62979 | 127 |
val narrowingN = "narrowing" |
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
128 |
|
67149 | 129 |
fun narrowingT T = \<^typ>\<open>integer\<close> --> Type (\<^type_name>\<open>Quickcheck_Narrowing.narrowing_cons\<close>, [T]) |
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
130 |
|
67149 | 131 |
fun mk_cons c T = Const (\<^const_name>\<open>Quickcheck_Narrowing.cons\<close>, T --> narrowingT T) $ Const (c, T) |
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
132 |
|
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
133 |
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:
41953
diff
changeset
|
134 |
let |
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
135 |
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:
41953
diff
changeset
|
136 |
in |
67149 | 137 |
(U', Const (\<^const_name>\<open>Quickcheck_Narrowing.apply\<close>, |
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
138 |
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:
41953
diff
changeset
|
139 |
end |
59154 | 140 |
|
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
141 |
fun mk_sum (t, u) = |
62979 | 142 |
let val T = fastype_of t |
67149 | 143 |
in Const (\<^const_name>\<open>Quickcheck_Narrowing.sum\<close>, T --> T --> T) $ t $ u end |
62979 | 144 |
|
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
145 |
|
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:
42616
diff
changeset
|
146 |
(** 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:
41953
diff
changeset
|
147 |
|
50046 | 148 |
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:
41953
diff
changeset
|
149 |
let |
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
150 |
fun mk_call T = |
67149 | 151 |
(T, Const (\<^const_name>\<open>Quickcheck_Narrowing.narrowing_class.narrowing\<close>, 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:
41953
diff
changeset
|
152 |
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:
41953
diff
changeset
|
153 |
let |
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
154 |
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:
41953
diff
changeset
|
155 |
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:
41953
diff
changeset
|
156 |
in |
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
157 |
(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:
41953
diff
changeset
|
158 |
end |
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
159 |
fun mk_consexpr simpleT (c, xs) = |
62979 | 160 |
let val Ts = map fst xs |
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
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:
41953
diff
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:
41953
diff
changeset
|
163 |
val rhss = |
58112
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents:
57996
diff
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:
41953
diff
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:
41953
diff
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:
41953
diff
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:
41953
diff
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:
41953
diff
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:
41953
diff
changeset
|
170 |
val eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhss ~~ rhss) |
62979 | 171 |
in eqs end |
59154 | 172 |
|
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
173 |
fun contains_recursive_type_under_function_types xs = |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
174 |
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:
57996
diff
changeset
|
175 |
(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:
43114
diff
changeset
|
176 |
|
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
177 |
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:
41953
diff
changeset
|
178 |
let |
62979 | 179 |
val _ = Old_Datatype_Aux.message config "Creating narrowing generators ..." |
180 |
val narrowingsN = map (prefix (narrowingN ^ "_")) (names @ auxnames) |
|
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
181 |
in |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
182 |
if not (contains_recursive_type_under_function_types descr) then |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
183 |
thy |
67149 | 184 |
|> Class.instantiation (tycos, vs, \<^sort>\<open>narrowing\<close>) |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
185 |
|> Quickcheck_Common.define_functions |
50046 | 186 |
(fn narrowings => mk_equations descr vs narrowings, NONE) |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
187 |
prfx [] narrowingsN (map narrowingT (Ts @ Us)) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59434
diff
changeset
|
188 |
|> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) |
62979 | 189 |
else thy |
190 |
end |
|
191 |
||
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
192 |
|
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
193 |
(* testing framework *) |
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41953
diff
changeset
|
194 |
|
43308
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
bulwahn
parents:
43240
diff
changeset
|
195 |
val target = "Haskell_Quickcheck" |
41930
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents:
41925
diff
changeset
|
196 |
|
62979 | 197 |
|
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:
42616
diff
changeset
|
198 |
(** invocation of Haskell interpreter **) |
41905 | 199 |
|
43702
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43619
diff
changeset
|
200 |
val narrowing_engine = |
63680 | 201 |
File.read \<^file>\<open>~~/src/HOL/Tools/Quickcheck/Narrowing_Engine.hs\<close> |
43702
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43619
diff
changeset
|
202 |
|
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43619
diff
changeset
|
203 |
val pnf_narrowing_engine = |
63680 | 204 |
File.read \<^file>\<open>~~/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs\<close> |
41905 | 205 |
|
206 |
fun exec verbose code = |
|
60956 | 207 |
ML_Context.exec (fn () => |
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
208 |
ML_Compiler0.ML ML_Env.context |
60956 | 209 |
{line = 0, file = "generated code", verbose = verbose, debug = false} code) |
41905 | 210 |
|
43079
4022892a2f28
improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents:
43047
diff
changeset
|
211 |
fun with_overlord_dir name f = |
72511
460d743010bc
clarified signature: overloaded "+" for Path.append;
wenzelm
parents:
72376
diff
changeset
|
212 |
(Path.explode "$ISABELLE_HOME_USER" + Path.basic (name ^ serial_string ())) |
72376 | 213 |
|> Isabelle_System.make_directory |
78706
a4969ab077d2
omit pointless capture/release (see also 26774ccb1c74);
wenzelm
parents:
78705
diff
changeset
|
214 |
|> f |
43379
8c4b383e5143
quickcheck_narrowing returns some timing information
bulwahn
parents:
43329
diff
changeset
|
215 |
|
62979 | 216 |
fun value (contains_existentials, ((genuine_only, (quiet, verbose)), size)) |
75604
39df30349778
more scalable generated files and code export, using Bytes.T;
wenzelm
parents:
74147
diff
changeset
|
217 |
ctxt cookie (code_modules_bytes, _) = |
41905 | 218 |
let |
75604
39df30349778
more scalable generated files and code export, using Bytes.T;
wenzelm
parents:
74147
diff
changeset
|
219 |
val code_modules = (map o apsnd) Bytes.content code_modules_bytes |
45756
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
220 |
val ((is_genuine, counterexample_of), (get, put, put_ml)) = cookie |
58843 | 221 |
fun message s = if quiet then () else writeln s |
222 |
fun verbose_message s = if not quiet andalso verbose then writeln s else () |
|
43585 | 223 |
val current_size = Unsynchronized.ref 0 |
59154 | 224 |
val current_result = Unsynchronized.ref Quickcheck.empty_result |
41930
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents:
41925
diff
changeset
|
225 |
val tmp_prefix = "Quickcheck_Narrowing" |
47843
4da917ed49b7
adding configuration to pass options to the ghc call in quickcheck[narrowing]
bulwahn
parents:
47330
diff
changeset
|
226 |
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:
43047
diff
changeset
|
227 |
val with_tmp_dir = |
59154 | 228 |
if Config.get ctxt overlord then with_overlord_dir else Isabelle_System.with_tmp_dir |
229 |
fun run in_path = |
|
41905 | 230 |
let |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
67149
diff
changeset
|
231 |
fun mk_code_file module = |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
67149
diff
changeset
|
232 |
let |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
67149
diff
changeset
|
233 |
val (paths, base) = split_last module |
69625
94048eac7463
restored quickcheck/narrowing, which got accidentally broken in ef02c5e051e5
haftmann
parents:
69623
diff
changeset
|
234 |
in Path.appends (in_path :: map Path.basic (paths @ [suffix ".hs" base])) end; |
94048eac7463
restored quickcheck/narrowing, which got accidentally broken in ef02c5e051e5
haftmann
parents:
69623
diff
changeset
|
235 |
val generatedN_suffix = suffix ".hs" Code_Target.generatedN; |
94048eac7463
restored quickcheck/narrowing, which got accidentally broken in ef02c5e051e5
haftmann
parents:
69623
diff
changeset
|
236 |
val includes = AList.delete (op =) [generatedN_suffix] code_modules |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
67149
diff
changeset
|
237 |
|> (map o apfst) mk_code_file |
69625
94048eac7463
restored quickcheck/narrowing, which got accidentally broken in ef02c5e051e5
haftmann
parents:
69623
diff
changeset
|
238 |
val code = the (AList.lookup (op =) code_modules [generatedN_suffix]) |
94048eac7463
restored quickcheck/narrowing, which got accidentally broken in ef02c5e051e5
haftmann
parents:
69623
diff
changeset
|
239 |
val code_file = mk_code_file [Code_Target.generatedN] |
69623
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
67149
diff
changeset
|
240 |
val narrowing_engine_file = mk_code_file ["Narrowing_Engine"] |
ef02c5e051e5
explicit model concerning files of generated code
haftmann
parents:
67149
diff
changeset
|
241 |
val main_file = mk_code_file ["Main"] |
62979 | 242 |
val main = |
243 |
"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:
46316
diff
changeset
|
244 |
"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:
46316
diff
changeset
|
245 |
"import System.Environment;\n" ^ |
41933
10f254a4e5b9
adapting Main file generation for Quickcheck_Narrowing
bulwahn
parents:
41932
diff
changeset
|
246 |
"import Narrowing_Engine;\n" ^ |
69625
94048eac7463
restored quickcheck/narrowing, which got accidentally broken in ef02c5e051e5
haftmann
parents:
69623
diff
changeset
|
247 |
"import " ^ Code_Target.generatedN ^ " ;\n\n" ^ |
45685
e2e928af750b
quickcheck narrowing also shows potential counterexamples
bulwahn
parents:
45420
diff
changeset
|
248 |
"main = getArgs >>= \\[potential, size] -> " ^ |
69625
94048eac7463
restored quickcheck/narrowing, which got accidentally broken in ef02c5e051e5
haftmann
parents:
69623
diff
changeset
|
249 |
"Narrowing_Engine.depthCheck (read potential) (read size) (" ^ Code_Target.generatedN ^ |
62979 | 250 |
".value ())\n\n}\n" |
251 |
val _ = |
|
252 |
map (uncurry File.write) |
|
253 |
(includes @ |
|
254 |
[(narrowing_engine_file, |
|
255 |
if contains_existentials then pnf_narrowing_engine else narrowing_engine), |
|
256 |
(code_file, code), (main_file, main)]) |
|
72511
460d743010bc
clarified signature: overloaded "+" for Path.append;
wenzelm
parents:
72376
diff
changeset
|
257 |
val executable = in_path + Path.basic "isabelle_quickcheck_narrowing" |
62979 | 258 |
val cmd = |
259 |
"exec \"$ISABELLE_GHC\" " ^ Code_Haskell.language_params ^ " " ^ ghc_options ^ " " ^ |
|
80910 | 260 |
(implode_space |
72278 | 261 |
(map File.bash_platform_path |
62979 | 262 |
(map fst includes @ [code_file, narrowing_engine_file, main_file]))) ^ |
72278 | 263 |
" -o " ^ File.bash_platform_path executable ^ ";" |
73283 | 264 |
val compilation_time = |
74147
d030b988d470
provide bash_process server for Isabelle/ML and other external programs;
wenzelm
parents:
74142
diff
changeset
|
265 |
Isabelle_System.bash_process (Bash.script cmd) |
73285
0e7a3c055f39
clarified uses of Isabelle_System.bash_process: more checks, fewer messages;
wenzelm
parents:
73284
diff
changeset
|
266 |
|> Process_Result.check |
0e7a3c055f39
clarified uses of Isabelle_System.bash_process: more checks, fewer messages;
wenzelm
parents:
73284
diff
changeset
|
267 |
|> Process_Result.timing_elapsed |> Time.toMilliseconds |
0e7a3c055f39
clarified uses of Isabelle_System.bash_process: more checks, fewer messages;
wenzelm
parents:
73284
diff
changeset
|
268 |
handle ERROR msg => cat_error "Compilation with GHC failed" msg |
73283 | 269 |
val _ = Quickcheck.add_timing ("Haskell compilation", compilation_time) current_result |
45685
e2e928af750b
quickcheck narrowing also shows potential counterexamples
bulwahn
parents:
45420
diff
changeset
|
270 |
fun haskell_string_of_bool v = if v then "True" else "False" |
45756
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
271 |
fun with_size genuine_only k = |
62979 | 272 |
if k > size then (NONE, !current_result) |
43114
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents:
43079
diff
changeset
|
273 |
else |
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents:
43079
diff
changeset
|
274 |
let |
45765
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
bulwahn
parents:
45760
diff
changeset
|
275 |
val _ = verbose_message ("[Quickcheck-narrowing] Test data size: " ^ string_of_int k) |
43585 | 276 |
val _ = current_size := k |
73283 | 277 |
val res = |
74147
d030b988d470
provide bash_process server for Isabelle/ML and other external programs;
wenzelm
parents:
74142
diff
changeset
|
278 |
Isabelle_System.bash_process (Bash.script |
73283 | 279 |
(File.bash_path executable ^ " " ^ haskell_string_of_bool genuine_only ^ " " ^ |
74147
d030b988d470
provide bash_process server for Isabelle/ML and other external programs;
wenzelm
parents:
74142
diff
changeset
|
280 |
string_of_int k)) |
73285
0e7a3c055f39
clarified uses of Isabelle_System.bash_process: more checks, fewer messages;
wenzelm
parents:
73284
diff
changeset
|
281 |
|> Process_Result.check |
0e7a3c055f39
clarified uses of Isabelle_System.bash_process: more checks, fewer messages;
wenzelm
parents:
73284
diff
changeset
|
282 |
val response = Process_Result.out res |
73284 | 283 |
val timing = res |> Process_Result.timing_elapsed |> Time.toMilliseconds; |
73283 | 284 |
val _ = |
285 |
Quickcheck.add_timing |
|
286 |
("execution of size " ^ string_of_int k, timing) current_result |
|
43114
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents:
43079
diff
changeset
|
287 |
in |
73269 | 288 |
if response = "NONE" then with_size genuine_only (k + 1) |
43585 | 289 |
else |
290 |
let |
|
291 |
val output_value = the_default "NONE" |
|
292 |
(try (snd o split_last o filter_out (fn s => s = "") o split_lines) response) |
|
62889 | 293 |
val ml_code = |
294 |
"\nval _ = Context.put_generic_context (SOME (Context.map_proof (" ^ put_ml |
|
62979 | 295 |
^ " (fn () => " ^ output_value ^ ")) (Context.the_generic_context ())))" |
43585 | 296 |
val ctxt' = ctxt |
297 |
|> put (fn () => error ("Bad evaluation for " ^ quote put_ml)) |
|
62979 | 298 |
|> Context.proof_map (exec false ml_code) |
45757
e32dd098f57a
renaming potential flag to genuine_only flag with an inverse semantics
bulwahn
parents:
45756
diff
changeset
|
299 |
val counterexample = get ctxt' () |
45756
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
300 |
in |
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
301 |
if is_genuine counterexample then |
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
302 |
(counterexample, !current_result) |
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
303 |
else |
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
304 |
let |
62979 | 305 |
val cex = Option.map (rpair []) (counterexample_of counterexample) |
306 |
val _ = message (Pretty.string_of (Quickcheck.pretty_counterex ctxt false cex)) |
|
307 |
val _ = message "Quickcheck continues to find a genuine counterexample..." |
|
59434 | 308 |
in with_size true (k + 1) end |
45756
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
309 |
end |
59154 | 310 |
end |
45756
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
311 |
in with_size genuine_only 0 end |
62979 | 312 |
in with_tmp_dir tmp_prefix run end |
41905 | 313 |
|
55757 | 314 |
fun dynamic_value_strict opts cookie ctxt postproc t = |
41905 | 315 |
let |
56920
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
haftmann
parents:
56242
diff
changeset
|
316 |
fun evaluator program _ vs_ty_t deps = |
78705
fde0b195cb7d
clarified signature: avoid association with potentially dangerous Exn.capture;
wenzelm
parents:
75604
diff
changeset
|
317 |
Exn.result (value opts ctxt cookie) |
64957 | 318 |
(Code_Target.compilation_text ctxt target program deps true vs_ty_t) |
62979 | 319 |
in Exn.release (Code_Thingol.dynamic_value ctxt (Exn.map_res o postproc) evaluator t) end |
41905 | 320 |
|
59154 | 321 |
|
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:
42616
diff
changeset
|
322 |
(** counterexample generator **) |
46042 | 323 |
|
59154 | 324 |
datatype counterexample = |
325 |
Universal_Counterexample of (term * counterexample) |
|
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
326 |
| Existential_Counterexample of (term * counterexample) list |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
327 |
| Empty_Assignment |
59154 | 328 |
|
62979 | 329 |
fun map_counterexample _ Empty_Assignment = Empty_Assignment |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
330 |
| map_counterexample f (Universal_Counterexample (t, c)) = |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
331 |
Universal_Counterexample (f t, map_counterexample f c) |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
332 |
| map_counterexample f (Existential_Counterexample cs) = |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
333 |
Existential_Counterexample (map (fn (t, c) => (f t, map_counterexample f c)) cs) |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
334 |
|
59154 | 335 |
structure Data = Proof_Data |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
336 |
( |
59154 | 337 |
type T = |
338 |
(unit -> (bool * term list) option) * |
|
62979 | 339 |
(unit -> counterexample option) |
59154 | 340 |
val empty: T = |
341 |
(fn () => raise Fail "counterexample", |
|
62979 | 342 |
fn () => raise Fail "existential_counterexample") |
343 |
fun init _ = empty |
|
344 |
) |
|
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
345 |
|
59154 | 346 |
val get_counterexample = #1 o Data.get; |
347 |
val get_existential_counterexample = #2 o Data.get; |
|
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
348 |
|
62979 | 349 |
val put_counterexample = Data.map o @{apply 2(1)} o K |
350 |
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:
42020
diff
changeset
|
351 |
|
43240
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents:
43237
diff
changeset
|
352 |
fun finitize_functions (xTs, t) = |
42024
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents:
42023
diff
changeset
|
353 |
let |
43240
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents:
43237
diff
changeset
|
354 |
val (names, boundTs) = split_list xTs |
42024
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents:
42023
diff
changeset
|
355 |
fun mk_eval_ffun dT rT = |
67149 | 356 |
Const (\<^const_name>\<open>Quickcheck_Narrowing.eval_ffun\<close>, |
357 |
Type (\<^type_name>\<open>Quickcheck_Narrowing.ffun\<close>, [dT, rT]) --> dT --> rT) |
|
42024
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents:
42023
diff
changeset
|
358 |
fun mk_eval_cfun dT rT = |
67149 | 359 |
Const (\<^const_name>\<open>Quickcheck_Narrowing.eval_cfun\<close>, |
360 |
Type (\<^type_name>\<open>Quickcheck_Narrowing.cfun\<close>, [rT]) --> dT --> rT) |
|
361 |
fun eval_function (Type (\<^type_name>\<open>fun\<close>, [dT, rT])) = |
|
62979 | 362 |
let |
363 |
val (rt', rT') = eval_function rT |
|
364 |
in |
|
365 |
(case dT of |
|
67149 | 366 |
Type (\<^type_name>\<open>fun\<close>, _) => |
62979 | 367 |
(fn t => absdummy dT (rt' (mk_eval_cfun dT rT' $ incr_boundvars 1 t $ Bound 0)), |
67149 | 368 |
Type (\<^type_name>\<open>Quickcheck_Narrowing.cfun\<close>, [rT'])) |
62979 | 369 |
| _ => |
370 |
(fn t => absdummy dT (rt' (mk_eval_ffun dT rT' $ incr_boundvars 1 t $ Bound 0)), |
|
67149 | 371 |
Type (\<^type_name>\<open>Quickcheck_Narrowing.ffun\<close>, [dT, rT']))) |
62979 | 372 |
end |
67149 | 373 |
| eval_function (T as Type (\<^type_name>\<open>prod\<close>, [fT, sT])) = |
62979 | 374 |
let |
375 |
val (ft', fT') = eval_function fT |
|
376 |
val (st', sT') = eval_function sT |
|
67149 | 377 |
val T' = Type (\<^type_name>\<open>prod\<close>, [fT', sT']) |
378 |
val map_const = Const (\<^const_name>\<open>map_prod\<close>, (fT' --> fT) --> (sT' --> sT) --> T' --> T) |
|
62979 | 379 |
fun apply_dummy T t = absdummy T (t (Bound 0)) |
380 |
in |
|
381 |
(fn t => list_comb (map_const, [apply_dummy fT' ft', apply_dummy sT' st', t]), T') |
|
382 |
end |
|
42024
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents:
42023
diff
changeset
|
383 |
| eval_function T = (I, T) |
43240
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents:
43237
diff
changeset
|
384 |
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:
43237
diff
changeset
|
385 |
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:
42023
diff
changeset
|
386 |
in |
43240
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents:
43237
diff
changeset
|
387 |
(names ~~ boundTs', t') |
42024
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents:
42023
diff
changeset
|
388 |
end |
42023
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents:
42020
diff
changeset
|
389 |
|
67149 | 390 |
fun dest_ffun (Type (\<^type_name>\<open>Quickcheck_Narrowing.ffun\<close>, [dT, rT])) = (dT, rT) |
45039
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
bulwahn
parents:
45002
diff
changeset
|
391 |
|
67149 | 392 |
fun eval_finite_functions (Const (\<^const_name>\<open>Quickcheck_Narrowing.ffun.Constant\<close>, T) $ value) = |
62979 | 393 |
absdummy (fst (dest_ffun (body_type T))) (eval_finite_functions value) |
67149 | 394 |
| eval_finite_functions (Const (\<^const_name>\<open>Quickcheck_Narrowing.ffun.Update\<close>, T) $ a $ b $ f) = |
62979 | 395 |
let |
396 |
val (T1, T2) = dest_ffun (body_type T) |
|
397 |
in |
|
398 |
Quickcheck_Common.mk_fun_upd T1 T2 |
|
399 |
(eval_finite_functions a, eval_finite_functions b) (eval_finite_functions f) |
|
400 |
end |
|
45039
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
bulwahn
parents:
45002
diff
changeset
|
401 |
| eval_finite_functions t = t |
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
bulwahn
parents:
45002
diff
changeset
|
402 |
|
62979 | 403 |
|
43114
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents:
43079
diff
changeset
|
404 |
(** tester **) |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
405 |
|
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
406 |
val rewrs = |
62979 | 407 |
map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) |
408 |
(@{thms all_simps} @ @{thms ex_simps}) @ |
|
409 |
map (HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) |
|
410 |
[@{thm iff_conv_conj_imp}, @{thm not_ex}, @{thm not_all}, |
|
411 |
@{thm meta_eq_to_obj_eq [OF Ex1_def]}] |
|
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
412 |
|
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
413 |
fun make_pnf_term thy t = Pattern.rewrite_term thy rewrs [] t |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
414 |
|
67149 | 415 |
fun strip_quantifiers (Const (\<^const_name>\<open>Ex\<close>, _) $ Abs (x, T, t)) = |
416 |
apfst (cons (\<^const_name>\<open>Ex\<close>, (x, T))) (strip_quantifiers t) |
|
417 |
| strip_quantifiers (Const (\<^const_name>\<open>All\<close>, _) $ Abs (x, T, t)) = |
|
418 |
apfst (cons (\<^const_name>\<open>All\<close>, (x, T))) (strip_quantifiers t) |
|
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
419 |
| strip_quantifiers t = ([], t) |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
420 |
|
62979 | 421 |
fun contains_existentials t = |
67149 | 422 |
exists (fn (Q, _) => Q = \<^const_name>\<open>Ex\<close>) (fst (strip_quantifiers t)) |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
423 |
|
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
424 |
fun mk_property qs t = |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
425 |
let |
67149 | 426 |
fun enclose (\<^const_name>\<open>Ex\<close>, (x, T)) t = |
427 |
Const (\<^const_name>\<open>Quickcheck_Narrowing.exists\<close>, |
|
428 |
(T --> \<^typ>\<open>property\<close>) --> \<^typ>\<open>property\<close>) $ Abs (x, T, t) |
|
429 |
| enclose (\<^const_name>\<open>All\<close>, (x, T)) t = |
|
430 |
Const (\<^const_name>\<open>Quickcheck_Narrowing.all\<close>, |
|
431 |
(T --> \<^typ>\<open>property\<close>) --> \<^typ>\<open>property\<close>) $ Abs (x, T, t) |
|
432 |
in fold_rev enclose qs (\<^term>\<open>Quickcheck_Narrowing.Property\<close> $ t) end |
|
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
433 |
|
67149 | 434 |
fun mk_case_term ctxt p ((\<^const_name>\<open>Ex\<close>, (x, T)) :: qs') (Existential_Counterexample cs) = |
62979 | 435 |
Case_Translation.make_case ctxt Case_Translation.Quiet Name.context (Free (x, T)) (map (fn (t, c) => |
436 |
(t, mk_case_term ctxt (p - 1) qs' c)) cs) |
|
67149 | 437 |
| mk_case_term ctxt p ((\<^const_name>\<open>All\<close>, _) :: qs') (Universal_Counterexample (t, c)) = |
62979 | 438 |
if p = 0 then t else mk_case_term ctxt (p - 1) qs' c |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
439 |
|
62979 | 440 |
val post_process = |
441 |
perhaps (try Quickcheck_Common.post_process_term) o eval_finite_functions |
|
45039
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
bulwahn
parents:
45002
diff
changeset
|
442 |
|
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
443 |
fun mk_terms ctxt qs result = |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
444 |
let |
67149 | 445 |
val ps = filter (fn (_, (\<^const_name>\<open>All\<close>, _)) => true | _ => false) (map_index I qs) |
62979 | 446 |
in |
447 |
map (fn (p, (_, (x, _))) => (x, mk_case_term ctxt p qs result)) ps |
|
448 |
|> map (apsnd post_process) |
|
449 |
end |
|
59154 | 450 |
|
50046 | 451 |
fun test_term ctxt catch_code_errors (t, _) = |
41905 | 452 |
let |
59154 | 453 |
fun dest_result (Quickcheck.Result r) = r |
43585 | 454 |
val opts = |
45765
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
bulwahn
parents:
45760
diff
changeset
|
455 |
((Config.get ctxt Quickcheck.genuine_only, |
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
bulwahn
parents:
45760
diff
changeset
|
456 |
(Config.get ctxt Quickcheck.quiet, Config.get ctxt Quickcheck.verbose)), |
45685
e2e928af750b
quickcheck narrowing also shows potential counterexamples
bulwahn
parents:
45420
diff
changeset
|
457 |
Config.get ctxt Quickcheck.size) |
42361 | 458 |
val thy = Proof_Context.theory_of ctxt |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
459 |
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:
43114
diff
changeset
|
460 |
val pnf_t = make_pnf_term thy t' |
41905 | 461 |
in |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
462 |
if Config.get ctxt allow_existentials andalso contains_existentials pnf_t then |
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
463 |
let |
43240
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents:
43237
diff
changeset
|
464 |
fun wrap f (qs, t) = |
62979 | 465 |
let val (qs1, qs2) = split_list qs |
466 |
in apfst (map2 pair qs1) (f (qs2, t)) end |
|
43240
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents:
43237
diff
changeset
|
467 |
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:
43237
diff
changeset
|
468 |
val (qs, prop_t) = finitize (strip_quantifiers pnf_t) |
46309
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
469 |
val act = if catch_code_errors then try else (fn f => SOME o f) |
59154 | 470 |
val execute = |
471 |
dynamic_value_strict (true, opts) |
|
472 |
((K true, fn _ => error ""), |
|
473 |
(get_existential_counterexample, put_existential_counterexample, |
|
474 |
"Narrowing_Generators.put_existential_counterexample")) |
|
475 |
ctxt (apfst o Option.map o map_counterexample) |
|
476 |
in |
|
62979 | 477 |
(case act execute (mk_property qs prop_t) of |
46309
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
478 |
SOME (counterexample, result) => Quickcheck.Result |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
479 |
{counterexample = Option.map (pair true o mk_terms ctxt qs) counterexample, |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
480 |
evaluation_terms = Option.map (K []) counterexample, |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
481 |
timings = #timings (dest_result result), reports = #reports (dest_result result)} |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
482 |
| NONE => |
59434 | 483 |
(Quickcheck.message ctxt "Conjecture is not executable with Quickcheck-narrowing"; |
62979 | 484 |
Quickcheck.empty_result)) |
43237
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents:
43114
diff
changeset
|
485 |
end |
43240
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents:
43237
diff
changeset
|
486 |
else |
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents:
43237
diff
changeset
|
487 |
let |
45756
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
488 |
val frees = Term.add_frees t [] |
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
489 |
val t' = fold_rev absfree frees t |
46219
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents:
46042
diff
changeset
|
490 |
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:
43237
diff
changeset
|
491 |
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:
43237
diff
changeset
|
492 |
fun ensure_testable t = |
67149 | 493 |
Const (\<^const_name>\<open>Quickcheck_Narrowing.ensure_testable\<close>, |
62979 | 494 |
fastype_of t --> fastype_of t) $ t |
59154 | 495 |
fun is_genuine (SOME (true, _)) = true |
45756
295658b28d3b
quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents:
45728
diff
changeset
|
496 |
| is_genuine _ = false |
62979 | 497 |
val counterexample_of = |
498 |
Option.map (apsnd (curry (op ~~) (map fst frees) o map post_process)) |
|
46309
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
499 |
val act = if catch_code_errors then try else (fn f => SOME o f) |
59154 | 500 |
val execute = |
501 |
dynamic_value_strict (false, opts) |
|
502 |
((is_genuine, counterexample_of), |
|
503 |
(get_counterexample, put_counterexample, |
|
504 |
"Narrowing_Generators.put_counterexample")) |
|
505 |
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:
43237
diff
changeset
|
506 |
in |
62979 | 507 |
(case act execute (ensure_testable (finitize t')) of |
46309
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
508 |
SOME (counterexample, result) => |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
509 |
Quickcheck.Result |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
510 |
{counterexample = counterexample_of counterexample, |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
511 |
evaluation_terms = Option.map (K []) counterexample, |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
512 |
timings = #timings (dest_result result), |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
513 |
reports = #reports (dest_result result)} |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
514 |
| NONE => |
59434 | 515 |
(Quickcheck.message ctxt "Conjecture is not executable with Quickcheck-narrowing"; |
62979 | 516 |
Quickcheck.empty_result)) |
43240
da47097bd589
adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents:
43237
diff
changeset
|
517 |
end |
62979 | 518 |
end |
41905 | 519 |
|
46309
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
520 |
fun test_goals ctxt catch_code_errors insts goals = |
62979 | 521 |
if not (getenv "ISABELLE_GHC" = "") then |
43314
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents:
43308
diff
changeset
|
522 |
let |
62979 | 523 |
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:
45039
diff
changeset
|
524 |
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:
43308
diff
changeset
|
525 |
in |
46309
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
526 |
Quickcheck_Common.collect_results (test_term ctxt catch_code_errors) |
693d8d7bc3cd
catching code generation errors in quickcheck-narrowing
bulwahn
parents:
46219
diff
changeset
|
527 |
(maps (map snd) correct_inst_goals) [] |
43314
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents:
43308
diff
changeset
|
528 |
end |
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents:
43308
diff
changeset
|
529 |
else |
58843 | 530 |
(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:
43308
diff
changeset
|
531 |
("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:
43910
diff
changeset
|
532 |
^ "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:
43910
diff
changeset
|
533 |
^ "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:
43910
diff
changeset
|
534 |
[Quickcheck.empty_result]) |
43114
b9fca691addd
Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents:
43079
diff
changeset
|
535 |
|
62979 | 536 |
|
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:
42616
diff
changeset
|
537 |
(* setup *) |
41905 | 538 |
|
67149 | 539 |
val active = Attrib.setup_config_bool \<^binding>\<open>quickcheck_narrowing_active\<close> (K false) |
43878
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
bulwahn
parents:
43850
diff
changeset
|
540 |
|
58826 | 541 |
val _ = |
542 |
Theory.setup |
|
543 |
(Code.datatype_interpretation ensure_partial_term_of |
|
544 |
#> Code.datatype_interpretation ensure_partial_term_of_code |
|
67149 | 545 |
#> Quickcheck_Common.datatype_interpretation \<^plugin>\<open>quickcheck_narrowing\<close> |
546 |
(\<^sort>\<open>narrowing\<close>, instantiate_narrowing_datatype) |
|
62979 | 547 |
#> Context.theory_map (Quickcheck.add_tester ("narrowing", (active, test_goals)))) |
59154 | 548 |
|
62979 | 549 |
end |