author | wenzelm |
Mon, 26 Nov 2012 17:13:44 +0100 | |
changeset 50231 | 81a067b188b8 |
parent 49441 | 0ae4216a1783 |
child 51092 | 5e6398b48030 |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: HOL/Mutabelle/mutabelle_extra.ML |
34965 | 2 |
Author: Stefan Berghofer, Jasmin Blanchette, Lukas Bulwahn, TU Muenchen |
3 |
||
41408 | 4 |
Invokation of Counterexample generators. |
34965 | 5 |
*) |
41408 | 6 |
|
34965 | 7 |
signature MUTABELLE_EXTRA = |
8 |
sig |
|
9 |
||
10 |
val take_random : int -> 'a list -> 'a list |
|
11 |
||
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
12 |
datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error | Solved | Unsolved |
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
13 |
type timings = (string * int) list |
34965 | 14 |
|
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
15 |
type mtd = string * (theory -> term -> outcome * timings) |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35092
diff
changeset
|
16 |
|
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
17 |
type mutant_subentry = term * (string * (outcome * timings)) list |
34965 | 18 |
type detailed_entry = string * bool * term * mutant_subentry list |
19 |
||
20 |
type subentry = string * int * int * int * int * int * int |
|
21 |
type entry = string * bool * subentry list |
|
22 |
type report = entry list |
|
23 |
||
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
24 |
val quickcheck_mtd : (Proof.context -> Proof.context) -> string -> mtd |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
25 |
|
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
26 |
val solve_direct_mtd : mtd |
46641 | 27 |
val try0_mtd : mtd |
40974 | 28 |
(* |
40971
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
29 |
val sledgehammer_mtd : mtd |
40974 | 30 |
*) |
41190 | 31 |
val nitpick_mtd : mtd |
32 |
||
34965 | 33 |
val refute_mtd : mtd |
34 |
||
35 |
val freezeT : term -> term |
|
36 |
val thms_of : bool -> theory -> thm list |
|
37 |
||
38 |
val string_for_report : report -> string |
|
39 |
val write_report : string -> report -> unit |
|
40 |
val mutate_theorems_and_write_report : |
|
46454
d72ab6bf6e6d
making num_mutations a configuration that can be changed with the mutabelle bash command
bulwahn
parents:
46453
diff
changeset
|
41 |
theory -> int * int -> mtd list -> thm list -> string -> unit |
34965 | 42 |
|
43 |
val random_seed : real Unsynchronized.ref |
|
44 |
end; |
|
45 |
||
46 |
structure MutabelleExtra : MUTABELLE_EXTRA = |
|
47 |
struct |
|
48 |
||
49 |
(* Own seed; can't rely on the Isabelle one to stay the same *) |
|
50 |
val random_seed = Unsynchronized.ref 1.0; |
|
51 |
||
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
52 |
(* Another Random engine *) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
53 |
|
34965 | 54 |
exception RANDOM; |
55 |
||
56 |
fun rmod x y = x - y * Real.realFloor (x / y); |
|
57 |
||
58 |
local |
|
59 |
val a = 16807.0; |
|
60 |
val m = 2147483647.0; |
|
61 |
in |
|
62 |
||
63 |
fun random () = CRITICAL (fn () => |
|
64 |
let val r = rmod (a * ! random_seed) m |
|
65 |
in (random_seed := r; r) end); |
|
66 |
||
67 |
end; |
|
68 |
||
69 |
fun random_range l h = |
|
70 |
if h < l orelse l < 0 then raise RANDOM |
|
71 |
else l + Real.floor (rmod (random ()) (real (h - l + 1))); |
|
72 |
||
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
73 |
fun take_random 0 _ = [] |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
74 |
| take_random _ [] = [] |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
75 |
| take_random n xs = |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
76 |
let val j = random_range 0 (length xs - 1) in |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
77 |
Library.nth xs j :: take_random (n - 1) (nth_drop j xs) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
78 |
end |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
79 |
|
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
80 |
(* possible outcomes *) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
81 |
|
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
82 |
datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error | Solved | Unsolved |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
83 |
|
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
84 |
fun string_of_outcome GenuineCex = "GenuineCex" |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
85 |
| string_of_outcome PotentialCex = "PotentialCex" |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
86 |
| string_of_outcome NoCex = "NoCex" |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
87 |
| string_of_outcome Donno = "Donno" |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
88 |
| string_of_outcome Timeout = "Timeout" |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
89 |
| string_of_outcome Error = "Error" |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
90 |
| string_of_outcome Solved = "Solved" |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
91 |
| string_of_outcome Unsolved = "Unsolved" |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
92 |
|
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
93 |
type timings = (string * int) list |
34965 | 94 |
|
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
95 |
type mtd = string * (theory -> term -> outcome * timings) |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35092
diff
changeset
|
96 |
|
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
97 |
type mutant_subentry = term * (string * (outcome * timings)) list |
34965 | 98 |
type detailed_entry = string * bool * term * mutant_subentry list |
99 |
||
100 |
type subentry = string * int * int * int * int * int * int |
|
101 |
type entry = string * bool * subentry list |
|
102 |
type report = entry list |
|
103 |
||
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
104 |
(* possible invocations *) |
34965 | 105 |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
106 |
(** quickcheck **) |
34965 | 107 |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
108 |
fun invoke_quickcheck change_options quickcheck_generator thy t = |
46452
e4f1cda51df6
increase timeout to 30 seconds; changing mutabelle script
bulwahn
parents:
46434
diff
changeset
|
109 |
TimeLimit.timeLimit (seconds 30.0) |
34965 | 110 |
(fn _ => |
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
111 |
let |
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:
44845
diff
changeset
|
112 |
val ctxt' = change_options (Proof_Context.init_global thy) |
46376
110ba6344446
mutabelle must handle the case where quickcheck returns multiple results
bulwahn
parents:
46326
diff
changeset
|
113 |
val (result :: _) = case Quickcheck.active_testers ctxt' of |
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:
44845
diff
changeset
|
114 |
[] => error "No active testers for quickcheck" |
45428
aa35c9454a95
quickcheck invocations in mutabelle must not catch codegenerator errors internally
bulwahn
parents:
45397
diff
changeset
|
115 |
| [tester] => tester ctxt' false [] [(t, [])] |
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:
44845
diff
changeset
|
116 |
| _ => error "Multiple active testers for quickcheck" |
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
117 |
in |
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
118 |
case Quickcheck.counterexample_of result of |
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
119 |
NONE => (NoCex, Quickcheck.timings_of result) |
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
120 |
| SOME _ => (GenuineCex, Quickcheck.timings_of result) |
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
121 |
end) () |
40931 | 122 |
handle TimeLimit.TimeOut => |
43018 | 123 |
(Timeout, [("timelimit", Real.floor (!Try.auto_time_limit))]) |
34965 | 124 |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
125 |
fun quickcheck_mtd change_options quickcheck_generator = |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
126 |
("quickcheck_" ^ quickcheck_generator, invoke_quickcheck change_options quickcheck_generator) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
127 |
|
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
128 |
(** solve direct **) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
129 |
|
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
130 |
fun invoke_solve_direct thy t = |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
131 |
let |
42361 | 132 |
val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (Proof_Context.init_global thy) |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
133 |
in |
43030 | 134 |
case Solve_Direct.solve_direct state of |
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
135 |
(true, _) => (Solved, []) |
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
136 |
| (false, _) => (Unsolved, []) |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
137 |
end |
34965 | 138 |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
139 |
val solve_direct_mtd = ("solve_direct", invoke_solve_direct) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
140 |
|
46641 | 141 |
(** try0 **) |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
142 |
|
46641 | 143 |
fun invoke_try0 thy t = |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
144 |
let |
42361 | 145 |
val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (Proof_Context.init_global thy) |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
146 |
in |
46641 | 147 |
case Try0.try0 (SOME (seconds 5.0)) ([], [], [], []) state of |
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
148 |
true => (Solved, []) |
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
149 |
| false => (Unsolved, []) |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
150 |
end |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
151 |
|
46641 | 152 |
val try0_mtd = ("try0", invoke_try0) |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
153 |
|
40971
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
154 |
(** sledgehammer **) |
40974 | 155 |
(* |
40971
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
156 |
fun invoke_sledgehammer thy t = |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
157 |
if can (Goal.prove_global thy (Term.add_free_names t []) [] t) |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
158 |
(fn {context, ...} => Sledgehammer_Tactics.sledgehammer_with_metis_tac context 1) then |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
159 |
(Solved, ([], NONE)) |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
160 |
else |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
161 |
(Unsolved, ([], NONE)) |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
162 |
|
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
163 |
val sledgehammer_mtd = ("sledgehammer", invoke_sledgehammer) |
40974 | 164 |
*) |
45397 | 165 |
|
34965 | 166 |
fun invoke_refute thy t = |
167 |
let |
|
45397 | 168 |
val params = [] |
169 |
val res = Refute.refute_term (Proof_Context.init_global thy) params [] t |
|
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39555
diff
changeset
|
170 |
val _ = Output.urgent_message ("Refute: " ^ res) |
34965 | 171 |
in |
45397 | 172 |
(rpair []) (case res of |
34965 | 173 |
"genuine" => GenuineCex |
174 |
| "likely_genuine" => GenuineCex |
|
175 |
| "potential" => PotentialCex |
|
176 |
| "none" => NoCex |
|
177 |
| "unknown" => Donno |
|
45397 | 178 |
| _ => Error) |
34965 | 179 |
end |
45397 | 180 |
handle Refute.REFUTE (loc, details) => |
34965 | 181 |
(error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^ |
182 |
".")) |
|
183 |
val refute_mtd = ("refute", invoke_refute) |
|
184 |
||
41190 | 185 |
(** nitpick **) |
34965 | 186 |
|
187 |
fun invoke_nitpick thy t = |
|
188 |
let |
|
42361 | 189 |
val ctxt = Proof_Context.init_global thy |
34965 | 190 |
val state = Proof.init ctxt |
41190 | 191 |
val (res, _) = Nitpick.pick_nits_in_term state |
47715
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents:
47108
diff
changeset
|
192 |
(Nitpick_Isar.default_params thy []) Nitpick.Normal 1 1 1 [] [] [] t |
41190 | 193 |
val _ = Output.urgent_message ("Nitpick: " ^ res) |
34965 | 194 |
in |
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
195 |
(rpair []) (case res of |
41190 | 196 |
"genuine" => GenuineCex |
197 |
| "likely_genuine" => GenuineCex |
|
198 |
| "potential" => PotentialCex |
|
199 |
| "none" => NoCex |
|
200 |
| "unknown" => Donno |
|
201 |
| _ => Error) |
|
34965 | 202 |
end |
41190 | 203 |
|
34965 | 204 |
val nitpick_mtd = ("nitpick", invoke_nitpick) |
205 |
||
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
206 |
(* filtering forbidden theorems and mutants *) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
207 |
|
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38857
diff
changeset
|
208 |
val comms = [@{const_name HOL.eq}, @{const_name HOL.disj}, @{const_name HOL.conj}] |
34965 | 209 |
|
210 |
val forbidden = |
|
211 |
[(* (@{const_name "power"}, "'a"), *) |
|
35325
4123977b469d
adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
bulwahn
parents:
35324
diff
changeset
|
212 |
(*(@{const_name induct_equal}, "'a"), |
4123977b469d
adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
bulwahn
parents:
35324
diff
changeset
|
213 |
(@{const_name induct_implies}, "'a"), |
4123977b469d
adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
bulwahn
parents:
35324
diff
changeset
|
214 |
(@{const_name induct_conj}, "'a"),*) |
34965 | 215 |
(@{const_name "undefined"}, "'a"), |
216 |
(@{const_name "default"}, "'a"), |
|
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
217 |
(@{const_name "dummy_pattern"}, "'a::{}"), |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
218 |
(@{const_name "HOL.simp_implies"}, "prop => prop => prop"), |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
219 |
(@{const_name "bot_fun_inst.bot_fun"}, "'a"), |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
220 |
(@{const_name "top_fun_inst.top_fun"}, "'a"), |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
221 |
(@{const_name "Pure.term"}, "'a"), |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
222 |
(@{const_name "top_class.top"}, "'a"), |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
223 |
(@{const_name "Quotient.Quot_True"}, "'a")(*, |
34965 | 224 |
(@{const_name "uminus"}, "'a"), |
225 |
(@{const_name "Nat.size"}, "'a"), |
|
35092
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
34974
diff
changeset
|
226 |
(@{const_name "Groups.abs"}, "'a") *)] |
34965 | 227 |
|
228 |
val forbidden_thms = |
|
229 |
["finite_intvl_succ_class", |
|
230 |
"nibble"] |
|
231 |
||
232 |
val forbidden_consts = |
|
46433
b67bb064de1e
mutabelle ignores theorems with internal constants
bulwahn
parents:
46376
diff
changeset
|
233 |
[@{const_name nibble_pair_of_char}, @{const_name "TYPE"}, |
b67bb064de1e
mutabelle ignores theorems with internal constants
bulwahn
parents:
46376
diff
changeset
|
234 |
@{const_name Datatype.dsum}, @{const_name Datatype.usum}] |
34965 | 235 |
|
236 |
fun is_forbidden_theorem (s, th) = |
|
237 |
let val consts = Term.add_const_names (prop_of th) [] in |
|
46711
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents:
46641
diff
changeset
|
238 |
exists (member (op =) (Long_Name.explode s)) forbidden_thms orelse |
36692
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents:
36610
diff
changeset
|
239 |
exists (member (op =) forbidden_consts) consts orelse |
46711
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents:
46641
diff
changeset
|
240 |
length (Long_Name.explode s) <> 2 orelse |
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents:
46641
diff
changeset
|
241 |
String.isPrefix "type_definition" (List.last (Long_Name.explode s)) orelse |
34965 | 242 |
String.isSuffix "_def" s orelse |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
243 |
String.isSuffix "_raw" s orelse |
46711
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents:
46641
diff
changeset
|
244 |
String.isPrefix "term_of" (List.last (Long_Name.explode s)) |
34965 | 245 |
end |
246 |
||
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
247 |
val forbidden_mutant_constnames = |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
248 |
["HOL.induct_equal", |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
249 |
"HOL.induct_implies", |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
250 |
"HOL.induct_conj", |
46314
f6532c597300
adding some more forbidden constant names for the mutated conjecture generation
bulwahn
parents:
46310
diff
changeset
|
251 |
"HOL.induct_forall", |
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
252 |
@{const_name undefined}, |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
253 |
@{const_name default}, |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
254 |
@{const_name dummy_pattern}, |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
255 |
@{const_name "HOL.simp_implies"}, |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
256 |
@{const_name "bot_fun_inst.bot_fun"}, |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
257 |
@{const_name "top_fun_inst.top_fun"}, |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
258 |
@{const_name "Pure.term"}, |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
259 |
@{const_name "top_class.top"}, |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
260 |
(*@{const_name "HOL.equal"},*) |
40971
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
261 |
@{const_name "Quotient.Quot_True"}, |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
262 |
@{const_name "equal_fun_inst.equal_fun"}, |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
263 |
@{const_name "equal_bool_inst.equal_bool"}, |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
264 |
@{const_name "ord_fun_inst.less_eq_fun"}, |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
265 |
@{const_name "ord_fun_inst.less_fun"}, |
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
266 |
@{const_name Meson.skolem}, |
43111 | 267 |
@{const_name ATP.fequal}, |
46326 | 268 |
@{const_name ATP.fEx}, |
42435
c3d880b13aa9
adding two further code-generator internal constants to the blacklist of Mutabelle
bulwahn
parents:
42361
diff
changeset
|
269 |
@{const_name transfer_morphism}, |
c3d880b13aa9
adding two further code-generator internal constants to the blacklist of Mutabelle
bulwahn
parents:
42361
diff
changeset
|
270 |
@{const_name enum_prod_inst.enum_all_prod}, |
46314
f6532c597300
adding some more forbidden constant names for the mutated conjecture generation
bulwahn
parents:
46310
diff
changeset
|
271 |
@{const_name enum_prod_inst.enum_ex_prod}, |
46315
40522961d4b1
adding another internal constant to mutabelle's blacklust
bulwahn
parents:
46314
diff
changeset
|
272 |
@{const_name Quickcheck.catch_match}, |
46434
6d2af424d0f8
adding some forbidden constant names for mutabelle
bulwahn
parents:
46433
diff
changeset
|
273 |
@{const_name Quickcheck_Exhaustive.unknown}, |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46711
diff
changeset
|
274 |
@{const_name Num.Bit0}, @{const_name Num.Bit1} |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
275 |
(*@{const_name "==>"}, @{const_name "=="}*)] |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
276 |
|
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
277 |
val forbidden_mutant_consts = |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
278 |
[ |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
279 |
(@{const_name "Groups.zero_class.zero"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
280 |
(@{const_name "Groups.one_class.one"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
281 |
(@{const_name "Groups.plus_class.plus"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
282 |
(@{const_name "Groups.minus_class.minus"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
283 |
(@{const_name "Groups.times_class.times"}, @{typ "prop => prop => prop"}), |
44064
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
huffman
parents:
43883
diff
changeset
|
284 |
(@{const_name "Fields.inverse_class.divide"}, @{typ "prop => prop => prop"}), |
44845 | 285 |
(@{const_name "Lattices.inf_class.inf"}, @{typ "prop => prop => prop"}), |
286 |
(@{const_name "Lattices.sup_class.sup"}, @{typ "prop => prop => prop"}), |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
287 |
(@{const_name "Orderings.bot_class.bot"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
288 |
(@{const_name "Orderings.ord_class.min"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
289 |
(@{const_name "Orderings.ord_class.max"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
290 |
(@{const_name "Divides.div_class.mod"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
291 |
(@{const_name "Divides.div_class.div"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
292 |
(@{const_name "GCD.gcd_class.gcd"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
293 |
(@{const_name "GCD.gcd_class.lcm"}, @{typ "prop => prop => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
294 |
(@{const_name "Orderings.bot_class.bot"}, @{typ "bool => prop"}), |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
295 |
(@{const_name "Groups.one_class.one"}, @{typ "bool => prop"}), |
46326 | 296 |
(@{const_name "Groups.zero_class.zero"},@{typ "bool => prop"}), |
297 |
(@{const_name "equal_class.equal"},@{typ "bool => bool => bool"})] |
|
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
298 |
|
34965 | 299 |
fun is_forbidden_mutant t = |
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
300 |
let |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
301 |
val const_names = Term.add_const_names t [] |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
302 |
val consts = Term.add_consts t [] |
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
303 |
in |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
304 |
exists (String.isPrefix "Nitpick") const_names orelse |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
305 |
exists (String.isSubstring "_sumC") const_names orelse |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
306 |
exists (member (op =) forbidden_mutant_constnames) const_names orelse |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
307 |
exists (member (op =) forbidden_mutant_consts) consts |
34965 | 308 |
end |
309 |
||
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
310 |
(* executable via quickcheck *) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
311 |
|
40248 | 312 |
fun is_executable_term thy t = |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
313 |
let |
42361 | 314 |
val ctxt = Proof_Context.init_global thy |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
315 |
in |
46452
e4f1cda51df6
increase timeout to 30 seconds; changing mutabelle script
bulwahn
parents:
46434
diff
changeset
|
316 |
can (TimeLimit.timeLimit (seconds 30.0) |
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:
44845
diff
changeset
|
317 |
(Quickcheck.test_terms |
45165
f4896c792316
adding testing of quickcheck narrowing with finite types to mutabelle script; modified is_executable in mutabelle_extra
bulwahn
parents:
45159
diff
changeset
|
318 |
((Context.proof_map (Quickcheck.set_active_testers ["exhaustive"]) #> |
f4896c792316
adding testing of quickcheck narrowing with finite types to mutabelle script; modified is_executable in mutabelle_extra
bulwahn
parents:
45159
diff
changeset
|
319 |
Config.put Quickcheck.finite_types true #> |
41106 | 320 |
Config.put Quickcheck.finite_type_size 1 #> |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
321 |
Config.put Quickcheck.size 1 #> Config.put Quickcheck.iterations 1) ctxt) |
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:
44845
diff
changeset
|
322 |
(false, false) [])) (map (rpair [] o Object_Logic.atomize_term thy) |
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:
44845
diff
changeset
|
323 |
(fst (Variable.import_terms true [t] ctxt))) |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
324 |
end |
40248 | 325 |
|
34965 | 326 |
fun is_executable_thm thy th = is_executable_term thy (prop_of th) |
327 |
||
328 |
val freezeT = |
|
329 |
map_types (map_type_tvar (fn ((a, i), S) => |
|
330 |
TFree (if i = 0 then a else a ^ "_" ^ string_of_int i, S))) |
|
331 |
||
332 |
fun thms_of all thy = |
|
333 |
filter |
|
334 |
(fn th => (all orelse Context.theory_name (theory_of_thm th) = Context.theory_name thy) |
|
335 |
(* andalso is_executable_thm thy th *)) |
|
336 |
(map snd (filter_out is_forbidden_theorem (Mutabelle.all_unconcealed_thms_of thy))) |
|
337 |
||
41409 | 338 |
fun count x = (length oo filter o equal) x |
34965 | 339 |
|
49441
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
bulwahn
parents:
47715
diff
changeset
|
340 |
fun elapsed_time description e = |
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
bulwahn
parents:
47715
diff
changeset
|
341 |
let val ({elapsed, ...}, result) = Timing.timing e () |
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
bulwahn
parents:
47715
diff
changeset
|
342 |
in (result, (description, Time.toMilliseconds elapsed)) end |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
343 |
(* |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
344 |
fun unsafe_invoke_mtd thy (mtd_name, invoke_mtd) t = |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
345 |
let |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
346 |
val _ = Output.urgent_message ("Invoking " ^ mtd_name) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
347 |
val ((res, (timing, reports)), time) = cpu_time "total time" (fn () => invoke_mtd thy t |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
348 |
handle ERROR s => (tracing s; (Error, ([], NONE)))) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
349 |
val _ = Output.urgent_message (" Done") |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
350 |
in (res, (time :: timing, reports)) end |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
351 |
*) |
34965 | 352 |
fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t = |
353 |
let |
|
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39555
diff
changeset
|
354 |
val _ = Output.urgent_message ("Invoking " ^ mtd_name) |
49441
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
bulwahn
parents:
47715
diff
changeset
|
355 |
val (res, timing) = elapsed_time "total time" |
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
bulwahn
parents:
47715
diff
changeset
|
356 |
(fn () => case try (invoke_mtd thy) t of |
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
bulwahn
parents:
47715
diff
changeset
|
357 |
SOME (res, timing) => res |
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39555
diff
changeset
|
358 |
| NONE => (Output.urgent_message ("**** PROBLEMS WITH " ^ Syntax.string_of_term_global thy t); |
49441
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
bulwahn
parents:
47715
diff
changeset
|
359 |
Error)) |
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39555
diff
changeset
|
360 |
val _ = Output.urgent_message (" Done") |
49441
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
bulwahn
parents:
47715
diff
changeset
|
361 |
in (res, [timing]) end |
34965 | 362 |
|
363 |
(* theory -> term list -> mtd -> subentry *) |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
364 |
|
34965 | 365 |
fun test_mutants_using_one_method thy mutants (mtd_name, invoke_mtd) = |
366 |
let |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
367 |
val res = map (fst o safe_invoke_mtd thy (mtd_name, invoke_mtd)) mutants |
34965 | 368 |
in |
369 |
(mtd_name, count GenuineCex res, count PotentialCex res, count NoCex res, |
|
370 |
count Donno res, count Timeout res, count Error res) |
|
371 |
end |
|
372 |
||
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
373 |
(* creating entries *) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
374 |
|
34965 | 375 |
fun create_entry thy thm exec mutants mtds = |
36743
ce2297415b54
prefer Thm.get_name_hint, which is closer to a user-space idea of "theorem name";
wenzelm
parents:
36692
diff
changeset
|
376 |
(Thm.get_name_hint thm, exec, map (test_mutants_using_one_method thy mutants) mtds) |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
377 |
|
34965 | 378 |
fun create_detailed_entry thy thm exec mutants mtds = |
379 |
let |
|
380 |
fun create_mutant_subentry mutant = (mutant, |
|
381 |
map (fn (mtd_name, invoke_mtd) => |
|
382 |
(mtd_name, safe_invoke_mtd thy (mtd_name, invoke_mtd) mutant)) mtds) |
|
383 |
in |
|
36743
ce2297415b54
prefer Thm.get_name_hint, which is closer to a user-space idea of "theorem name";
wenzelm
parents:
36692
diff
changeset
|
384 |
(Thm.get_name_hint thm, exec, prop_of thm, map create_mutant_subentry mutants) |
34965 | 385 |
end |
386 |
||
387 |
(* (theory -> thm -> bool -> term list -> mtd list -> 'a) -> theory -> mtd list -> thm -> 'a *) |
|
46454
d72ab6bf6e6d
making num_mutations a configuration that can be changed with the mutabelle bash command
bulwahn
parents:
46453
diff
changeset
|
388 |
fun mutate_theorem create_entry thy (num_mutations, max_mutants) mtds thm = |
34965 | 389 |
let |
390 |
val exec = is_executable_thm thy thm |
|
43277 | 391 |
val _ = tracing (if exec then "EXEC" else "NOEXEC") |
34965 | 392 |
val mutants = |
393 |
(if num_mutations = 0 then |
|
394 |
[Thm.prop_of thm] |
|
395 |
else |
|
396 |
Mutabelle.mutate_mix (Thm.prop_of thm) thy comms forbidden |
|
397 |
num_mutations) |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
398 |
|> tap (fn muts => tracing ("mutants: " ^ string_of_int (length muts))) |
34965 | 399 |
|> filter_out is_forbidden_mutant |
400 |
val mutants = |
|
401 |
if exec then |
|
402 |
let |
|
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39555
diff
changeset
|
403 |
val _ = Output.urgent_message ("BEFORE PARTITION OF " ^ |
41491 | 404 |
string_of_int (length mutants) ^ " MUTANTS") |
34965 | 405 |
val (execs, noexecs) = List.partition (is_executable_term thy) (take_random (20 * max_mutants) mutants) |
41491 | 406 |
val _ = tracing ("AFTER PARTITION (" ^ string_of_int (length execs) ^ |
407 |
" vs " ^ string_of_int (length noexecs) ^ ")") |
|
34965 | 408 |
in |
409 |
execs @ take_random (Int.max (0, max_mutants - length execs)) noexecs |
|
410 |
end |
|
411 |
else |
|
412 |
mutants |
|
413 |
val mutants = mutants |
|
414 |
|> map Mutabelle.freeze |> map freezeT |
|
415 |
(* |> filter (not o is_forbidden_mutant) *) |
|
42429 | 416 |
|> map_filter (try (Sign.cert_term thy)) |
417 |
|> filter (is_some o try (Thm.cterm_of thy)) |
|
418 |
|> filter (is_some o try (Syntax.check_term (Proof_Context.init_global thy))) |
|
40971
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
bulwahn
parents:
40932
diff
changeset
|
419 |
|> take_random max_mutants |
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39555
diff
changeset
|
420 |
val _ = map (fn t => Output.urgent_message ("MUTANT: " ^ Syntax.string_of_term_global thy t)) mutants |
34965 | 421 |
in |
422 |
create_entry thy thm exec mutants mtds |
|
423 |
end |
|
424 |
||
425 |
(* theory -> mtd list -> thm list -> report *) |
|
46453
9e83b7c24b05
making max_mutants an option that can be changed in the Mutabelle-script
bulwahn
parents:
46452
diff
changeset
|
426 |
val mutate_theorems = map oooo mutate_theorem |
34965 | 427 |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35092
diff
changeset
|
428 |
fun string_of_mutant_subentry thy thm_name (t, results) = |
34965 | 429 |
"mutant: " ^ Syntax.string_of_term_global thy t ^ "\n" ^ |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35092
diff
changeset
|
430 |
space_implode "; " |
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35092
diff
changeset
|
431 |
(map (fn (mtd_name, (outcome, timing)) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^ |
34965 | 432 |
"\n" |
433 |
||
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
434 |
(* string -> string *) |
39555
ccb223a4d49c
added XML.content_of convenience -- cover XML.body, which is the general situation;
wenzelm
parents:
39324
diff
changeset
|
435 |
val unyxml = XML.content_of o YXML.parse_body |
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
436 |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35092
diff
changeset
|
437 |
fun string_of_mutant_subentry' thy thm_name (t, results) = |
35380
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents:
35325
diff
changeset
|
438 |
let |
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
439 |
(* fun string_of_report (Quickcheck.Report {iterations = i, raised_match_errors = e, |
35380
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents:
35325
diff
changeset
|
440 |
satisfied_assms = s, positive_concl_tests = p}) = |
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents:
35325
diff
changeset
|
441 |
"errors: " ^ string_of_int e ^ "; conclusion tests: " ^ string_of_int p |
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents:
35325
diff
changeset
|
442 |
fun string_of_reports NONE = "" |
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents:
35325
diff
changeset
|
443 |
| string_of_reports (SOME reports) = |
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents:
35325
diff
changeset
|
444 |
cat_lines (map (fn (size, [report]) => |
42089
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
445 |
"size " ^ string_of_int size ^ ": " ^ string_of_report report) (rev reports))*) |
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
bulwahn
parents:
42032
diff
changeset
|
446 |
fun string_of_mtd_result (mtd_name, (outcome, timing)) = |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
447 |
mtd_name ^ ": " ^ string_of_outcome outcome |
49441
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
bulwahn
parents:
47715
diff
changeset
|
448 |
^ "(" ^ space_implode "; " (map (fn (s, t) => (s ^ ": " ^ string_of_int t)) timing) ^ ")" |
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
449 |
(*^ "\n" ^ string_of_reports reports*) |
35380
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents:
35325
diff
changeset
|
450 |
in |
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
451 |
"mutant of " ^ thm_name ^ ":\n" |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
452 |
^ unyxml (Syntax.string_of_term_global thy t) ^ "\n" ^ space_implode "; " (map string_of_mtd_result results) |
35380
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents:
35325
diff
changeset
|
453 |
end |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35092
diff
changeset
|
454 |
|
34965 | 455 |
fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = |
456 |
thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^ |
|
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
457 |
Syntax.string_of_term_global thy t ^ "\n" ^ |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35092
diff
changeset
|
458 |
cat_lines (map (string_of_mutant_subentry' thy thm_name) mutant_subentries) ^ "\n" |
34965 | 459 |
|
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
460 |
fun theoryfile_string_of_mutant_subentry thy thm_name (i, (t, results)) = |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
461 |
"lemma " ^ thm_name ^ "_" ^ string_of_int (i + 1) ^ ":\n" ^ |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
462 |
"\"" ^ unyxml (Syntax.string_of_term_global thy t) ^ |
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
463 |
"\" \nquickcheck\noops\n" |
36255
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
464 |
|
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
465 |
fun theoryfile_string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
466 |
"subsubsection {* mutants of " ^ thm_name ^ " *}\n\n" ^ |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
467 |
cat_lines (map_index |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
468 |
(theoryfile_string_of_mutant_subentry thy thm_name) mutant_subentries) ^ "\n" |
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
bulwahn
parents:
35625
diff
changeset
|
469 |
|
34965 | 470 |
(* subentry -> string *) |
471 |
fun string_for_subentry (mtd_name, genuine_cex, potential_cex, no_cex, donno, |
|
472 |
timeout, error) = |
|
41491 | 473 |
" " ^ mtd_name ^ ": " ^ string_of_int genuine_cex ^ "+ " ^ |
474 |
string_of_int potential_cex ^ "= " ^ string_of_int no_cex ^ "- " ^ |
|
475 |
string_of_int donno ^ "? " ^ string_of_int timeout ^ "T " ^ |
|
476 |
string_of_int error ^ "!" |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
477 |
|
34965 | 478 |
(* entry -> string *) |
479 |
fun string_for_entry (thm_name, exec, subentries) = |
|
480 |
thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ":\n" ^ |
|
481 |
cat_lines (map string_for_subentry subentries) ^ "\n" |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
482 |
|
34965 | 483 |
(* report -> string *) |
484 |
fun string_for_report report = cat_lines (map string_for_entry report) |
|
485 |
||
486 |
(* string -> report -> unit *) |
|
487 |
fun write_report file_name = |
|
488 |
File.write (Path.explode file_name) o string_for_report |
|
489 |
||
490 |
(* theory -> mtd list -> thm list -> string -> unit *) |
|
46454
d72ab6bf6e6d
making num_mutations a configuration that can be changed with the mutabelle bash command
bulwahn
parents:
46453
diff
changeset
|
491 |
fun mutate_theorems_and_write_report thy (num_mutations, max_mutants) mtds thms file_name = |
34965 | 492 |
let |
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39555
diff
changeset
|
493 |
val _ = Output.urgent_message "Starting Mutabelle..." |
42361 | 494 |
val ctxt = Proof_Context.init_global thy |
34965 | 495 |
val path = Path.explode file_name |
496 |
(* for normal report: *) |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
497 |
(* |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
498 |
val (gen_create_entry, gen_string_for_entry) = (create_entry, string_for_entry) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
499 |
*) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
500 |
(* for detailled report: *) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
501 |
val (gen_create_entry, gen_string_for_entry) = (create_detailed_entry, string_of_detailed_entry thy) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
502 |
(* for theory creation: *) |
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
503 |
(*val (gen_create_entry, gen_string_for_entry) = (create_detailed_entry, theoryfile_string_of_detailed_entry thy)*) |
34965 | 504 |
in |
505 |
File.write path ( |
|
506 |
"Mutation options = " ^ |
|
507 |
"max_mutants: " ^ string_of_int max_mutants ^ |
|
508 |
"; num_mutations: " ^ string_of_int num_mutations ^ "\n" ^ |
|
509 |
"QC options = " ^ |
|
510 |
(*"quickcheck_generator: " ^ quickcheck_generator ^ ";*) |
|
40653
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
bulwahn
parents:
40381
diff
changeset
|
511 |
"size: " ^ string_of_int (Config.get ctxt Quickcheck.size) ^ |
43244 | 512 |
"; iterations: " ^ string_of_int (Config.get ctxt Quickcheck.iterations) ^ "\n" ^ |
513 |
"Isabelle environment = ISABELLE_GHC: " ^ getenv "ISABELLE_GHC" ^ "\n"); |
|
46454
d72ab6bf6e6d
making num_mutations a configuration that can be changed with the mutabelle bash command
bulwahn
parents:
46453
diff
changeset
|
514 |
map (File.append path o gen_string_for_entry o mutate_theorem gen_create_entry thy (num_mutations, max_mutants) mtds) thms; |
34965 | 515 |
() |
516 |
end |
|
517 |
||
518 |
end; |