| author | wenzelm | 
| Mon, 15 Apr 2013 12:03:16 +0200 | |
| changeset 51707 | 21d7933de1eb | 
| parent 51160 | 599ff65b85e2 | 
| child 52639 | df830310e550 | 
| 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: 
40381diff
changeset | 12 | datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error | Solved | Unsolved | 
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 13 | type timings = (string * int) list | 
| 34965 | 14 | |
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
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: 
35092diff
changeset | 16 | |
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40932diff
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: 
46453diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
changeset | 92 | |
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 93 | type timings = (string * int) list | 
| 34965 | 94 | |
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
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: 
35092diff
changeset | 96 | |
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
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: 
40381diff
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: 
40381diff
changeset | 106 | (** quickcheck **) | 
| 34965 | 107 | |
| 51092 | 108 | fun invoke_quickcheck change_options thy t = | 
| 46452 
e4f1cda51df6
increase timeout to 30 seconds; changing mutabelle script
 bulwahn parents: 
46434diff
changeset | 109 | TimeLimit.timeLimit (seconds 30.0) | 
| 34965 | 110 | (fn _ => | 
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
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: 
44845diff
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: 
46326diff
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: 
44845diff
changeset | 114 | [] => error "No active testers for quickcheck" | 
| 45428 
aa35c9454a95
quickcheck invocations in mutabelle must not catch codegenerator errors internally
 bulwahn parents: 
45397diff
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: 
44845diff
changeset | 116 | | _ => error "Multiple active testers for quickcheck" | 
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 117 | in | 
| 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 118 | case Quickcheck.counterexample_of result of | 
| 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 119 | NONE => (NoCex, Quickcheck.timings_of result) | 
| 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 120 | | SOME _ => (GenuineCex, Quickcheck.timings_of result) | 
| 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
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: 
40381diff
changeset | 125 | fun quickcheck_mtd change_options quickcheck_generator = | 
| 51092 | 126 |   ("quickcheck_" ^ quickcheck_generator, invoke_quickcheck change_options)
 | 
| 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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
changeset | 133 | in | 
| 43030 | 134 | case Solve_Direct.solve_direct state of | 
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 135 | (true, _) => (Solved, []) | 
| 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
changeset | 146 | in | 
| 46641 | 147 | case Try0.try0 (SOME (seconds 5.0)) ([], [], [], []) state of | 
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 148 | true => (Solved, []) | 
| 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
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: 
40381diff
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: 
40381diff
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: 
40381diff
changeset | 153 | |
| 40971 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 154 | (** sledgehammer **) | 
| 40974 | 155 | (* | 
| 40971 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 156 | fun invoke_sledgehammer thy t = | 
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
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: 
40932diff
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: 
40932diff
changeset | 159 | (Solved, ([], NONE)) | 
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 160 | else | 
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 161 | (Unsolved, ([], NONE)) | 
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 162 | |
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
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: 
39555diff
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: 
47108diff
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: 
42032diff
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: 
40381diff
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: 
40381diff
changeset | 207 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38857diff
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: 
35324diff
changeset | 212 |   (*(@{const_name induct_equal}, "'a"),
 | 
| 
4123977b469d
adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
 bulwahn parents: 
35324diff
changeset | 213 |   (@{const_name induct_implies}, "'a"),
 | 
| 
4123977b469d
adding ROOT.ML to HOL-Mutabelle session; uncommenting HOL.induct constants in Mutabelle session
 bulwahn parents: 
35324diff
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: 
35625diff
changeset | 217 |   (@{const_name "dummy_pattern"}, "'a::{}"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 218 |   (@{const_name "HOL.simp_implies"}, "prop => prop => prop"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 219 |   (@{const_name "bot_fun_inst.bot_fun"}, "'a"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 220 |   (@{const_name "top_fun_inst.top_fun"}, "'a"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 221 |   (@{const_name "Pure.term"}, "'a"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 222 |   (@{const_name "top_class.top"}, "'a"),
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
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: 
34974diff
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 = | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
51126diff
changeset | 233 |  [@{const_name "TYPE"}, @{const_name Datatype.dsum}, @{const_name Datatype.usum}]
 | 
| 34965 | 234 | |
| 235 | fun is_forbidden_theorem (s, th) = | |
| 236 | 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: 
46641diff
changeset | 237 | 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: 
36610diff
changeset | 238 | exists (member (op =) forbidden_consts) consts orelse | 
| 46711 
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
 wenzelm parents: 
46641diff
changeset | 239 | length (Long_Name.explode s) <> 2 orelse | 
| 
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
 wenzelm parents: 
46641diff
changeset | 240 | String.isPrefix "type_definition" (List.last (Long_Name.explode s)) orelse | 
| 34965 | 241 | 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: 
40381diff
changeset | 242 | String.isSuffix "_raw" s orelse | 
| 46711 
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
 wenzelm parents: 
46641diff
changeset | 243 | String.isPrefix "term_of" (List.last (Long_Name.explode s)) | 
| 34965 | 244 | end | 
| 245 | ||
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 246 | val forbidden_mutant_constnames = | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 247 | ["HOL.induct_equal", | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 248 | "HOL.induct_implies", | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 249 | "HOL.induct_conj", | 
| 46314 
f6532c597300
adding some more forbidden constant names for the mutated conjecture generation
 bulwahn parents: 
46310diff
changeset | 250 | "HOL.induct_forall", | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 251 |  @{const_name undefined},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 252 |  @{const_name default},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 253 |  @{const_name dummy_pattern},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 254 |  @{const_name "HOL.simp_implies"},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 255 |  @{const_name "bot_fun_inst.bot_fun"},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 256 |  @{const_name "top_fun_inst.top_fun"},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 257 |  @{const_name "Pure.term"},
 | 
| 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 258 |  @{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: 
40381diff
changeset | 259 |  (*@{const_name "HOL.equal"},*)
 | 
| 40971 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 260 |  @{const_name "Quotient.Quot_True"},
 | 
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 261 |  @{const_name "equal_fun_inst.equal_fun"},
 | 
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 262 |  @{const_name "equal_bool_inst.equal_bool"},
 | 
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 263 |  @{const_name "ord_fun_inst.less_eq_fun"},
 | 
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 264 |  @{const_name "ord_fun_inst.less_fun"},
 | 
| 
6604115019bf
adding filtering, sytactic welltyping, and sledgehammer method in mutabelle
 bulwahn parents: 
40932diff
changeset | 265 |  @{const_name Meson.skolem},
 | 
| 43111 | 266 |  @{const_name ATP.fequal},
 | 
| 46326 | 267 |  @{const_name ATP.fEx},
 | 
| 42435 
c3d880b13aa9
adding two further code-generator internal constants to the blacklist of Mutabelle
 bulwahn parents: 
42361diff
changeset | 268 |  @{const_name transfer_morphism},
 | 
| 
c3d880b13aa9
adding two further code-generator internal constants to the blacklist of Mutabelle
 bulwahn parents: 
42361diff
changeset | 269 |  @{const_name enum_prod_inst.enum_all_prod},
 | 
| 46314 
f6532c597300
adding some more forbidden constant names for the mutated conjecture generation
 bulwahn parents: 
46310diff
changeset | 270 |  @{const_name enum_prod_inst.enum_ex_prod},
 | 
| 51126 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 haftmann parents: 
51092diff
changeset | 271 |  @{const_name Quickcheck_Random.catch_match},
 | 
| 46434 
6d2af424d0f8
adding some forbidden constant names for mutabelle
 bulwahn parents: 
46433diff
changeset | 272 |  @{const_name Quickcheck_Exhaustive.unknown},
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46711diff
changeset | 273 |  @{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: 
40381diff
changeset | 274 |  (*@{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: 
40381diff
changeset | 275 | |
| 
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
 bulwahn parents: 
40381diff
changeset | 276 | 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: 
40381diff
changeset | 277 | [ | 
| 
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
 bulwahn parents: 
40381diff
changeset | 278 |    (@{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: 
40381diff
changeset | 279 |    (@{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: 
40381diff
changeset | 280 |    (@{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: 
40381diff
changeset | 281 |    (@{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: 
40381diff
changeset | 282 |    (@{const_name "Groups.times_class.times"}, @{typ "prop => prop => prop"}),
 | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
43883diff
changeset | 283 |    (@{const_name "Fields.inverse_class.divide"}, @{typ "prop => prop => prop"}),
 | 
| 44845 | 284 |    (@{const_name "Lattices.inf_class.inf"}, @{typ "prop => prop => prop"}),
 | 
| 285 |    (@{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: 
40381diff
changeset | 286 |    (@{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: 
40381diff
changeset | 287 |    (@{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: 
40381diff
changeset | 288 |    (@{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: 
40381diff
changeset | 289 |    (@{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: 
40381diff
changeset | 290 |    (@{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: 
40381diff
changeset | 291 |    (@{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: 
40381diff
changeset | 292 |    (@{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: 
40381diff
changeset | 293 |    (@{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: 
40381diff
changeset | 294 |    (@{const_name "Groups.one_class.one"}, @{typ "bool => prop"}),
 | 
| 46326 | 295 |    (@{const_name "Groups.zero_class.zero"},@{typ "bool => prop"}),
 | 
| 296 |    (@{const_name "equal_class.equal"},@{typ "bool => bool => bool"})]
 | |
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 297 | |
| 34965 | 298 | fun is_forbidden_mutant t = | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 299 | 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: 
40381diff
changeset | 300 | 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: 
40381diff
changeset | 301 | val consts = Term.add_consts t [] | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 302 | 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: 
40381diff
changeset | 303 | 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: 
40381diff
changeset | 304 | 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: 
40381diff
changeset | 305 | 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: 
40381diff
changeset | 306 | exists (member (op =) forbidden_mutant_consts) consts | 
| 34965 | 307 | end | 
| 308 | ||
| 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: 
40381diff
changeset | 309 | (* 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: 
40381diff
changeset | 310 | |
| 40248 | 311 | 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: 
40381diff
changeset | 312 | let | 
| 42361 | 313 | 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: 
40381diff
changeset | 314 | in | 
| 46452 
e4f1cda51df6
increase timeout to 30 seconds; changing mutabelle script
 bulwahn parents: 
46434diff
changeset | 315 | 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: 
44845diff
changeset | 316 | (Quickcheck.test_terms | 
| 45165 
f4896c792316
adding testing of quickcheck narrowing with finite types to mutabelle script; modified is_executable in mutabelle_extra
 bulwahn parents: 
45159diff
changeset | 317 | ((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: 
45159diff
changeset | 318 | Config.put Quickcheck.finite_types true #> | 
| 41106 | 319 | 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: 
40381diff
changeset | 320 | 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: 
44845diff
changeset | 321 | (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: 
44845diff
changeset | 322 | (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: 
40381diff
changeset | 323 | end | 
| 40248 | 324 | |
| 34965 | 325 | fun is_executable_thm thy th = is_executable_term thy (prop_of th) | 
| 326 | ||
| 327 | val freezeT = | |
| 328 | map_types (map_type_tvar (fn ((a, i), S) => | |
| 329 | TFree (if i = 0 then a else a ^ "_" ^ string_of_int i, S))) | |
| 330 | ||
| 331 | fun thms_of all thy = | |
| 332 | filter | |
| 333 | (fn th => (all orelse Context.theory_name (theory_of_thm th) = Context.theory_name thy) | |
| 334 | (* andalso is_executable_thm thy th *)) | |
| 335 | (map snd (filter_out is_forbidden_theorem (Mutabelle.all_unconcealed_thms_of thy))) | |
| 336 | ||
| 49441 
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
 bulwahn parents: 
47715diff
changeset | 337 | fun elapsed_time description e = | 
| 
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
 bulwahn parents: 
47715diff
changeset | 338 |   let val ({elapsed, ...}, result) = Timing.timing e ()
 | 
| 
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
 bulwahn parents: 
47715diff
changeset | 339 | 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: 
40381diff
changeset | 340 | (* | 
| 
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
 bulwahn parents: 
40381diff
changeset | 341 | 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: 
40381diff
changeset | 342 | 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: 
40381diff
changeset | 343 |     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: 
40381diff
changeset | 344 | 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: 
40381diff
changeset | 345 | 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: 
40381diff
changeset | 346 |     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: 
40381diff
changeset | 347 | 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: 
40381diff
changeset | 348 | *) | 
| 34965 | 349 | fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t = | 
| 350 | let | |
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 351 |     val _ = Output.urgent_message ("Invoking " ^ mtd_name)
 | 
| 49441 
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
 bulwahn parents: 
47715diff
changeset | 352 | val (res, timing) = elapsed_time "total time" | 
| 
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
 bulwahn parents: 
47715diff
changeset | 353 | (fn () => case try (invoke_mtd thy) t of | 
| 51092 | 354 | SOME (res, _) => res | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 355 |         | 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: 
47715diff
changeset | 356 | Error)) | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 357 |     val _ = Output.urgent_message (" Done")
 | 
| 49441 
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
 bulwahn parents: 
47715diff
changeset | 358 | in (res, [timing]) end | 
| 34965 | 359 | |
| 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: 
40381diff
changeset | 360 | (* 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: 
40381diff
changeset | 361 | |
| 34965 | 362 | fun create_detailed_entry thy thm exec mutants mtds = | 
| 363 | let | |
| 364 | fun create_mutant_subentry mutant = (mutant, | |
| 365 | map (fn (mtd_name, invoke_mtd) => | |
| 366 | (mtd_name, safe_invoke_mtd thy (mtd_name, invoke_mtd) mutant)) mtds) | |
| 367 | in | |
| 36743 
ce2297415b54
prefer Thm.get_name_hint, which is closer to a user-space idea of "theorem name";
 wenzelm parents: 
36692diff
changeset | 368 | (Thm.get_name_hint thm, exec, prop_of thm, map create_mutant_subentry mutants) | 
| 34965 | 369 | end | 
| 370 | ||
| 371 | (* (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: 
46453diff
changeset | 372 | fun mutate_theorem create_entry thy (num_mutations, max_mutants) mtds thm = | 
| 34965 | 373 | let | 
| 374 | val exec = is_executable_thm thy thm | |
| 43277 | 375 | val _ = tracing (if exec then "EXEC" else "NOEXEC") | 
| 34965 | 376 | val mutants = | 
| 377 | (if num_mutations = 0 then | |
| 378 | [Thm.prop_of thm] | |
| 379 | else | |
| 380 | Mutabelle.mutate_mix (Thm.prop_of thm) thy comms forbidden | |
| 381 | 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: 
40381diff
changeset | 382 |              |> tap (fn muts => tracing ("mutants: " ^ string_of_int (length muts)))
 | 
| 34965 | 383 | |> filter_out is_forbidden_mutant | 
| 384 | val mutants = | |
| 385 | if exec then | |
| 386 | let | |
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 387 |           val _ = Output.urgent_message ("BEFORE PARTITION OF " ^
 | 
| 41491 | 388 | string_of_int (length mutants) ^ " MUTANTS") | 
| 34965 | 389 | val (execs, noexecs) = List.partition (is_executable_term thy) (take_random (20 * max_mutants) mutants) | 
| 41491 | 390 |           val _ = tracing ("AFTER PARTITION (" ^ string_of_int (length execs) ^
 | 
| 391 | " vs " ^ string_of_int (length noexecs) ^ ")") | |
| 34965 | 392 | in | 
| 393 | execs @ take_random (Int.max (0, max_mutants - length execs)) noexecs | |
| 394 | end | |
| 395 | else | |
| 396 | mutants | |
| 397 | val mutants = mutants | |
| 398 | |> map Mutabelle.freeze |> map freezeT | |
| 399 | (* |> filter (not o is_forbidden_mutant) *) | |
| 42429 | 400 | |> map_filter (try (Sign.cert_term thy)) | 
| 401 | |> filter (is_some o try (Thm.cterm_of thy)) | |
| 402 | |> 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: 
40932diff
changeset | 403 | |> take_random max_mutants | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 404 |     val _ = map (fn t => Output.urgent_message ("MUTANT: " ^ Syntax.string_of_term_global thy t)) mutants
 | 
| 34965 | 405 | in | 
| 406 | create_entry thy thm exec mutants mtds | |
| 407 | end | |
| 408 | ||
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 409 | (* string -> string *) | 
| 39555 
ccb223a4d49c
added XML.content_of convenience -- cover XML.body, which is the general situation;
 wenzelm parents: 
39324diff
changeset | 410 | val unyxml = XML.content_of o YXML.parse_body | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 411 | |
| 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: 
35092diff
changeset | 412 | fun string_of_mutant_subentry' thy thm_name (t, results) = | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 413 | let | 
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 414 |    (* fun string_of_report (Quickcheck.Report {iterations = i, raised_match_errors = e,
 | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 415 | satisfied_assms = s, positive_concl_tests = p}) = | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 416 | "errors: " ^ string_of_int e ^ "; conclusion tests: " ^ string_of_int p | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 417 | fun string_of_reports NONE = "" | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 418 | | string_of_reports (SOME reports) = | 
| 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 419 | cat_lines (map (fn (size, [report]) => | 
| 42089 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 420 | "size " ^ string_of_int size ^ ": " ^ string_of_report report) (rev reports))*) | 
| 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 bulwahn parents: 
42032diff
changeset | 421 | 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: 
40381diff
changeset | 422 | mtd_name ^ ": " ^ string_of_outcome outcome | 
| 49441 
0ae4216a1783
recording elapsed time in mutabelle for more detailed evaluation
 bulwahn parents: 
47715diff
changeset | 423 |       ^ "(" ^ 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: 
35625diff
changeset | 424 | (*^ "\n" ^ string_of_reports reports*) | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35325diff
changeset | 425 | in | 
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 426 | "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: 
40381diff
changeset | 427 | ^ 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: 
35325diff
changeset | 428 | 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: 
35092diff
changeset | 429 | |
| 34965 | 430 | fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = | 
| 431 | thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^ | |
| 36255 
f8b3381e1437
tuning mutabelle; adding output of mutant theoryfile for interactive evaluation
 bulwahn parents: 
35625diff
changeset | 432 | 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: 
35092diff
changeset | 433 | cat_lines (map (string_of_mutant_subentry' thy thm_name) mutant_subentries) ^ "\n" | 
| 34965 | 434 | |
| 435 | (* subentry -> string *) | |
| 436 | fun string_for_subentry (mtd_name, genuine_cex, potential_cex, no_cex, donno, | |
| 437 | timeout, error) = | |
| 41491 | 438 | " " ^ mtd_name ^ ": " ^ string_of_int genuine_cex ^ "+ " ^ | 
| 439 | string_of_int potential_cex ^ "= " ^ string_of_int no_cex ^ "- " ^ | |
| 440 | string_of_int donno ^ "? " ^ string_of_int timeout ^ "T " ^ | |
| 441 | 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: 
40381diff
changeset | 442 | |
| 34965 | 443 | (* entry -> string *) | 
| 444 | fun string_for_entry (thm_name, exec, subentries) = | |
| 445 | thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ":\n" ^ | |
| 446 | 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: 
40381diff
changeset | 447 | |
| 34965 | 448 | (* report -> string *) | 
| 449 | fun string_for_report report = cat_lines (map string_for_entry report) | |
| 450 | ||
| 451 | (* string -> report -> unit *) | |
| 452 | fun write_report file_name = | |
| 453 | File.write (Path.explode file_name) o string_for_report | |
| 454 | ||
| 455 | (* 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: 
46453diff
changeset | 456 | fun mutate_theorems_and_write_report thy (num_mutations, max_mutants) mtds thms file_name = | 
| 34965 | 457 | let | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
39555diff
changeset | 458 | val _ = Output.urgent_message "Starting Mutabelle..." | 
| 42361 | 459 | val ctxt = Proof_Context.init_global thy | 
| 34965 | 460 | val path = Path.explode file_name | 
| 461 | (* 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: 
40381diff
changeset | 462 | (* | 
| 
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
 bulwahn parents: 
40381diff
changeset | 463 | 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: 
40381diff
changeset | 464 | *) | 
| 
d921c97bdbd8
adding AFP tests to Mutabelle_Extra; adopting mutabelle to recent quickcheck changes; filtering strange mutants; adding solvers to mutabelle; restructuring mutabelle
 bulwahn parents: 
40381diff
changeset | 465 | (* 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: 
40381diff
changeset | 466 | 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: 
40381diff
changeset | 467 | (* 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: 
40381diff
changeset | 468 | (*val (gen_create_entry, gen_string_for_entry) = (create_detailed_entry, theoryfile_string_of_detailed_entry thy)*) | 
| 34965 | 469 | in | 
| 470 | File.write path ( | |
| 471 | "Mutation options = " ^ | |
| 472 | "max_mutants: " ^ string_of_int max_mutants ^ | |
| 473 | "; num_mutations: " ^ string_of_int num_mutations ^ "\n" ^ | |
| 474 | "QC options = " ^ | |
| 475 | (*"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: 
40381diff
changeset | 476 | "size: " ^ string_of_int (Config.get ctxt Quickcheck.size) ^ | 
| 43244 | 477 | "; iterations: " ^ string_of_int (Config.get ctxt Quickcheck.iterations) ^ "\n" ^ | 
| 478 | "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: 
46453diff
changeset | 479 | map (File.append path o gen_string_for_entry o mutate_theorem gen_create_entry thy (num_mutations, max_mutants) mtds) thms; | 
| 34965 | 480 | () | 
| 481 | end | |
| 482 | ||
| 483 | end; |